ProgTutorial/Recipes/Antiquotes.thy
changeset 569 f875a25aa72d
parent 565 cecd7a941885
child 573 321e220a6baa
equal deleted inserted replaced
568:be23597e81db 569:f875a25aa72d
    29   and also allows you to keep documents in sync with other code, for example
    29   and also allows you to keep documents in sync with other code, for example
    30   Isabelle.
    30   Isabelle.
    31 
    31 
    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] "@{ML_checked \"a_piece_of_code\"}"}
    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
    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
    38   "ML_Context.eval_source_in"} in Line 7 below). The complete code of the
       
    39   document antiquotation is as follows:
    38   document antiquotation is as follows:
    40 
    39 
    41 \<close>
    40 \<close>
    42 ML \<open>Input.pos_of\<close>
    41 ML \<open>Input.pos_of\<close>
    43 ML%linenosgray\<open>fun ml_enclose bg en source =
    42 ML%linenosgray\<open>fun ml_enclose bg en source =
    56 
    55 
    57 setup \<open>ml_checked_setup\<close>
    56 setup \<open>ml_checked_setup\<close>
    58 
    57 
    59 
    58 
    60 text \<open>
    59 text \<open>
    61   The parser @{ML "(Scan.lift Args.name)"} in Line 7 parses a string, in this
    60   The parser @{ML \<open>(Scan.lift Args.name)\<close>} in Line 7 parses a string, in this
    62   case the code, and then calls the function @{ML output_ml}. As mentioned
    61   case the code, and then calls the function @{ML output_ml}. As mentioned
    63   before, the parsed code is sent to the ML-compiler in Line 4 using the
    62   before, the parsed code is sent to the ML-compiler in Line 4 using the
    64   function @{ML ml_val}, which constructs the appropriate ML-expression, and
    63   function @{ML ml_val}, which constructs the appropriate ML-expression, and
    65   using @{ML "eval_in" in ML_Context}, which calls the compiler.  If the code is
    64   using @{ML \<open>eval_in\<close> in ML_Context}, which calls the compiler.  If the code is
    66   ``approved'' by the compiler, then the output function @{ML "output" in
    65   ``approved'' by the compiler, then the output function @{ML \<open>output\<close> in
    67   Document_Antiquotation} in the next line pretty prints the code. This function expects
    66   Document_Antiquotation} in the next line pretty prints the code. This function expects
    68   that the code is a list of (pretty)strings where each string correspond to a
    67   that the code is a list of (pretty)strings where each string correspond to a
    69   line in the output. Therefore the use of @{ML "(space_explode \"\\n\" txt)"
    68   line in the output. Therefore the use of @{ML \<open>(space_explode "\\n" txt)\<close>
    70   for txt} which produces such a list according to linebreaks.  There are a
    69   for txt} which produces such a list according to linebreaks.  There are a
    71   number of options for antiquotations that are observed by the function 
    70   number of options for antiquotations that are observed by the function 
    72   @{ML "output" 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> 
    73   and \<open>[quotes]\<close>). The function @{ML "antiquotation_raw" in Thy_Output} in 
    72   and \<open>[quotes]\<close>). The function @{ML \<open>antiquotation_raw\<close> in Thy_Output} in 
    74   Line 7 sets up the new document antiquotation.
    73   Line 7 sets up the new document antiquotation.
    75 
    74 
    76   \begin{readmore}
    75   \begin{readmore}
    77   For more information about options of document antiquotations see \rsccite{sec:antiq}).
    76   For more information about options of document antiquotations see \rsccite{sec:antiq}).
    78   \end{readmore}
    77   \end{readmore}
    79 
    78 
    80   Since we used the argument @{ML "Position.none"}, the compiler cannot give specific 
    79   Since we used the argument @{ML \<open>Position.none\<close>}, the compiler cannot give specific 
    81   information about the line number, in case an error is detected. We 
    80   information about the line number, in case an error is detected. We 
    82   can improve the code above slightly by writing 
    81   can improve the code above slightly by writing 
    83 \<close>
    82 \<close>
    84 (* FIXME: remove
    83 (* FIXME: remove
    85 ML%linenosgray{*fun output_ml ctxt (code_txt, pos) =
    84 ML%linenosgray{*fun output_ml ctxt (code_txt, pos) =
   114   that specifies what the result of the ML-code should be and checks the
   113   that specifies what the result of the ML-code should be and checks the
   115   consistency of the actual result with the given pattern. For this we are
   114   consistency of the actual result with the given pattern. For this we are
   116   going to implement the document antiquotation:
   115   going to implement the document antiquotation:
   117 
   116 
   118   
   117   
   119   @{text [display] "@{ML_resp \"a_piece_of_code\" \"a_pattern\"}"}
   118   @{text [display] \<open>@{ML_resp "a_piece_of_code" "a_pattern"}\<close>}
   120   
   119   
   121   To add some convenience and also to deal with large outputs, the user can
   120   To add some convenience and also to deal with large outputs, the user can
   122   give a partial specification by using ellipses. For example \<open>(\<dots>, \<dots>)\<close>
   121   give a partial specification by using ellipses. For example \<open>(\<dots>, \<dots>)\<close>
   123   for specifying a pair.  In order to check consistency between the pattern
   122   for specifying a pair.  In order to check consistency between the pattern
   124   and the output of the code, we have to change the ML-expression that is sent 
   123   and the output of the code, we have to change the ML-expression that is sent 
   191   In comparison with \<open>ML_checked\<close>, we only changed the line about 
   190   In comparison with \<open>ML_checked\<close>, we only changed the line about 
   192   the compiler (Line~2), the lines about
   191   the compiler (Line~2), the lines about
   193   the output (Lines 4 to 7) and the parser in the setup (Line 11). Now 
   192   the output (Lines 4 to 7) and the parser in the setup (Line 11). Now 
   194   you can write
   193   you can write
   195  
   194  
   196   @{text [display] "@{ML_resp [display] \"true andalso false\" \"false\"}"}
   195   @{text [display] \<open>@{ML_resp [display] "true andalso false" "false"}\<close>}
   197 
   196 
   198   to obtain
   197   to obtain
   199 
   198 
   200   @{ML_resp [display] "true andalso false" "false"} 
   199   @{ML_resp [display] "true andalso false" "false"} 
   201 
   200 
   202   or 
   201   or 
   203 
   202 
   204   @{text [display] "@{ML_resp [display] \"let val i = 3 in (i * i, \"foo\") end\" \"(9, \<dots>)\"}"}
   203   @{text [display] \<open>@{ML_resp [display] "let val i = 3 in (i * i, "foo") end" "(9, \<dots>)"}\<close>}
   205   
   204   
   206   to obtain
   205   to obtain
   207 
   206 
   208   @{ML_resp [display] "let val i = 3 in (i * i, \"foo\") end" "(9, _)"} 
   207   @{ML_resp [display] "let val i = 3 in (i * i, \"foo\") end" "(9, _)"} 
   209 
   208