ProgTutorial/First_Steps.thy
changeset 474 683fb6e468b1
parent 471 f65b9f14d5de
child 477 141751cab5b2
equal deleted inserted replaced
473:fea1b7ce5c4a 474:683fb6e468b1
    55   @{text "\<verbclose>"}\isanewline
    55   @{text "\<verbclose>"}\isanewline
    56   @{text "> 7"}\smallskip
    56   @{text "> 7"}\smallskip
    57   \end{graybox}
    57   \end{graybox}
    58   \end{isabelle}
    58   \end{isabelle}
    59 
    59 
    60   If you work with ProofGeneral then like normal Isabelle scripts 
    60   If you work with ProofGeneral then like normal Isabelle scripts
    61   \isacommand{ML}-commands can be evaluated by
    61   \isacommand{ML}-commands can be evaluated by using the advance and
    62   using the advance and undo buttons of your Isabelle environment. The code
    62   undo buttons of your Isabelle environment.  If you work with the
    63   inside the \isacommand{ML}-command can also contain value and function
    63   Jedit GUI, then you just have to hover the cursor over the code 
    64   bindings, for example
    64   and you see the evaluated result in the ``Output'' window.
    65 *}
    65 
    66 
    66   As mentioned in the Introduction, we will drop the \isacommand{ML}~@{text
    67 ML %gray {* 
       
    68   val r = Unsynchronized.ref 0
       
    69   fun f n = n + 1 
       
    70 *}
       
    71 
       
    72 text {*
       
    73   and even those can be undone when the proof script is retracted.  As
       
    74   mentioned in the Introduction, we will drop the \isacommand{ML}~@{text
       
    75   "\<verbopen> \<dots> \<verbclose>"} scaffolding whenever we show code. The lines
    67   "\<verbopen> \<dots> \<verbclose>"} scaffolding whenever we show code. The lines
    76   prefixed with @{text [quotes] ">"} are not part of the code, rather they
    68   prefixed with @{text [quotes] ">"} are not part of the code, rather they
    77   indicate what the response is when the code is evaluated.  There are also
    69   indicate what the response is when the code is evaluated.  There are also
    78   the commands \isacommand{ML\_val} and \isacommand{ML\_prf} for including
    70   the commands \isacommand{ML\_val} and \isacommand{ML\_prf} for including
    79   ML-code. The first evaluates the given code, but any effect on the theory,
    71   ML-code. The first evaluates the given code, but any effect on the theory,
   138   code. This can be done in a ``quick-and-dirty'' fashion using the function
   130   code. This can be done in a ``quick-and-dirty'' fashion using the function
   139   @{ML_ind writeln in Output}. For example
   131   @{ML_ind writeln in Output}. For example
   140 
   132 
   141   @{ML_response_fake [display,gray] "writeln \"any string\"" "\"any string\""}
   133   @{ML_response_fake [display,gray] "writeln \"any string\"" "\"any string\""}
   142 
   134 
   143   will print out @{text [quotes] "any string"} inside the response buffer of
   135   will print out @{text [quotes] "any string"} inside the response buffer.  
   144   Isabelle.  This function expects a string as argument. If you develop under
   136   This function expects a string as argument. If you develop under
   145   PolyML, then there is a convenient, though again ``quick-and-dirty'', method
   137   PolyML, then there is a convenient, though again ``quick-and-dirty'', method
   146   for converting values into strings, namely the antiquotation 
   138   for converting values into strings, namely the antiquotation 
   147   @{text "@{make_string}"}:
   139   @{text "@{make_string}"}:
   148 
   140 
   149   @{ML_response_fake [display,gray] "writeln (@{make_string} 1)" "\"1\""} 
   141   @{ML_response_fake [display,gray] "writeln (@{make_string} 1)" "\"1\""} 
   160 At command \"ML\"."}
   152 At command \"ML\"."}
   161 
   153 
   162   This function raises the exception @{text ERROR}, which will then 
   154   This function raises the exception @{text ERROR}, which will then 
   163   be displayed by the infrastructure. Note that this exception is meant 
   155   be displayed by the infrastructure. Note that this exception is meant 
   164   for ``user-level'' error messages seen by the ``end-user''. 
   156   for ``user-level'' error messages seen by the ``end-user''. 
   165   
       
   166   For messages where you want to indicate a genuine program error, then
   157   For messages where you want to indicate a genuine program error, then
   167   use the exception @{text Fail}. Here the infrastructure indicates that this 
   158   use the exception @{text Fail}. 
   168   is a low-level exception, and also prints the source position of the ML 
       
   169   raise statement. 
       
   170 
   159 
   171   Most often you want to inspect data of Isabelle's basic data structures,
   160   Most often you want to inspect data of Isabelle's basic data structures,
   172   namely @{ML_type term}, @{ML_type typ}, @{ML_type cterm}, @{ML_type ctyp}
   161   namely @{ML_type term}, @{ML_type typ}, @{ML_type cterm}, @{ML_type ctyp}
   173   and @{ML_type thm}. Isabelle contains elaborate pretty-printing functions,
   162   and @{ML_type thm}. Isabelle contains elaborate pretty-printing functions,
   174   which we will explain in more detail in Section \ref{sec:pretty}. For now
   163   which we will explain in more detail in Section \ref{sec:pretty}. For now
   322   together, like a warning message consisting of a term and its type, you
   311   together, like a warning message consisting of a term and its type, you
   323   should try to print these parcels together in a single string. Therefore do
   312   should try to print these parcels together in a single string. Therefore do
   324   \emph{not} print out information as
   313   \emph{not} print out information as
   325 
   314 
   326 @{ML_response_fake [display,gray]
   315 @{ML_response_fake [display,gray]
   327 "writeln \"First half,\"; 
   316 "pwriteln (Pretty.str \"First half,\"); 
   328 writeln \"and second half.\""
   317 pwriteln (Pretty.str \"and second half.\")"
   329 "First half,
   318 "First half,
   330 and second half."}
   319 and second half."}
   331 
   320 
   332   but as a single string with appropriate formatting. For example
   321   but as a single string with appropriate formatting. For example
   333 
   322 
   334 @{ML_response_fake [display,gray]
   323 @{ML_response_fake [display,gray]
   335 "writeln (\"First half,\" ^ \"\\n\" ^ \"and second half.\")"
   324 "pwriteln (Pretty.str (\"First half,\" ^ \"\\n\" ^ \"and second half.\"))"
   336 "First half,
   325 "First half,
   337 and second half."}
   326 and second half."}
   338   
   327   
   339   To ease this kind of string manipulations, there are a number
   328   To ease this kind of string manipulations, there are a number
   340   of library functions in Isabelle. For example,  the function 
   329   of library functions in Isabelle. For example,  the function 
   341   @{ML_ind cat_lines in Library} concatenates a list of strings 
   330   @{ML_ind cat_lines in Library} concatenates a list of strings 
   342   and inserts newlines in between each element. 
   331   and inserts newlines in between each element. 
   343 
   332 
   344   @{ML_response_fake [display, gray]
   333   @{ML_response_fake [display, gray]
   345   "writeln (cat_lines [\"foo\", \"bar\"])"
   334   "pwriteln (Pretty.str (cat_lines [\"foo\", \"bar\"]))"
   346   "foo
   335   "foo
   347 bar"}
   336 bar"}
   348 
   337 
   349   Section \ref{sec:pretty} will explain the infrastructure that Isabelle 
   338   Section \ref{sec:pretty} will explain the infrastructure that Isabelle 
   350   provides for more elaborate pretty printing. 
   339   provides for more elaborate pretty printing. 
   500   then increment-by-two, followed by increment-by-three. Again, the reverse
   489   then increment-by-two, followed by increment-by-three. Again, the reverse
   501   function composition allows you to read the code top-down. This combinator
   490   function composition allows you to read the code top-down. This combinator
   502   is often used for setup functions inside the
   491   is often used for setup functions inside the
   503   \isacommand{setup}-command. These functions have to be of type @{ML_type
   492   \isacommand{setup}-command. These functions have to be of type @{ML_type
   504   "theory -> theory"}. More than one such setup function can be composed with
   493   "theory -> theory"}. More than one such setup function can be composed with
   505   @{ML "#>"}.\footnote{give example} 
   494   @{ML "#>"}.\footnote{TBD: give example} 
   506   
   495   
   507   The remaining combinators we describe in this section add convenience for the
   496   The remaining combinators we describe in this section add convenience for the
   508   ``waterfall method'' of writing functions. The combinator @{ML_ind tap in
   497   ``waterfall method'' of writing functions. The combinator @{ML_ind tap in
   509   Basics} allows you to get hold of an intermediate result (to do some
   498   Basics} allows you to get hold of an intermediate result (to do some
   510   side-calculations or print out an intermediate result, for instance). The function
   499   side-calculations or print out an intermediate result, for instance). The function
   511  *}
   500  *}
   512 
   501 
   513 ML %linenosgray{*fun inc_by_three x =
   502 ML %linenosgray{*fun inc_by_three x =
   514      x |> (fn x => x + 1)
   503      x |> (fn x => x + 1)
   515        |> tap (fn x => writeln (@{make_string} x))
   504        |> tap (fn x => pwriteln (Pretty.str (@{make_string} x)))
   516        |> (fn x => x + 2)*}
   505        |> (fn x => x + 2)*}
   517 
   506 
   518 text {* 
   507 text {* 
   519   increments the argument first by @{text "1"} and then by @{text "2"}. In the
   508   increments the argument first by @{text "1"} and then by @{text "2"}. In the
   520   middle (Line 3), however, it uses @{ML tap} for printing the ``plus-one''
   509   middle (Line 3), however, it uses @{ML tap} for printing the ``plus-one''
   786 "(?P \<and> ?Q) = (?Q \<and> ?P)
   775 "(?P \<and> ?Q) = (?Q \<and> ?P)
   787 (?P \<and> ?Q \<and> ?R) = (?Q \<and> ?P \<and> ?R)
   776 (?P \<and> ?Q \<and> ?R) = (?Q \<and> ?P \<and> ?R)
   788 ((?P \<and> ?Q) \<and> ?R) = (?P \<and> ?Q \<and> ?R)"}  
   777 ((?P \<and> ?Q) \<and> ?R) = (?P \<and> ?Q \<and> ?R)"}  
   789 
   778 
   790   The thm-antiquotations can also be used for manipulating theorems. For 
   779   The thm-antiquotations can also be used for manipulating theorems. For 
   791   example, if you need the version of te theorem @{thm [source] refl} that 
   780   example, if you need the version of the theorem @{thm [source] refl} that 
   792   has a meta-equality instead of an equality, you can write
   781   has a meta-equality instead of an equality, you can write
   793 
   782 
   794 @{ML_response_fake [display,gray] 
   783 @{ML_response_fake [display,gray] 
   795   "@{thm refl[THEN eq_reflection]}"
   784   "@{thm refl[THEN eq_reflection]}"
   796   "?x \<equiv> ?x"}
   785   "?x \<equiv> ?x"}
   862          |> ML_Syntax.atomic
   851          |> ML_Syntax.atomic
   863 in
   852 in
   864   ML_Antiquote.inline @{binding "term_pat"} (parser >> term_pat)
   853   ML_Antiquote.inline @{binding "term_pat"} (parser >> term_pat)
   865 end*}
   854 end*}
   866 
   855 
       
   856 text {*
       
   857   To use it you also have to install it using \isacommand{setup} like so
       
   858 *}
       
   859 
   867 setup{* term_pat_setup *}
   860 setup{* term_pat_setup *}
   868 
   861 
   869 text {*
   862 text {*
   870   The parser in Line 2 provides us with a context and a string; this string is
   863   The parser in Line 2 provides us with a context and a string; this string is
   871   transformed into a term using the function @{ML_ind read_term_pattern in
   864   transformed into a term using the function @{ML_ind read_term_pattern in
   876   @{ML_response_fake [display,gray]
   869   @{ML_response_fake [display,gray]
   877   "@{term_pat \"Suc (?x::nat)\"}"
   870   "@{term_pat \"Suc (?x::nat)\"}"
   878   "Const (\"Suc\", \"nat \<Rightarrow> nat\") $ Var ((\"x\", 0), \"nat\")"}
   871   "Const (\"Suc\", \"nat \<Rightarrow> nat\") $ Var ((\"x\", 0), \"nat\")"}
   879 
   872 
   880   which shows the internal representation of the term @{text "Suc ?x"}. Similarly
   873   which shows the internal representation of the term @{text "Suc ?x"}. Similarly
   881   we can write an antiquotation for type patterns.
   874   we can write an antiquotation for type patterns. Its code is
   882 *}
   875 *}
   883 
   876 
   884 ML{*val type_pat_setup = 
   877 ML{*val type_pat_setup = 
   885 let
   878 let
   886   val parser = Args.context -- Scan.lift Args.name_source
   879   val parser = Args.context -- Scan.lift Args.name_source
   891          |> ML_Syntax.atomic
   884          |> ML_Syntax.atomic
   892 in
   885 in
   893   ML_Antiquote.inline @{binding "typ_pat"} (parser >> typ_pat)
   886   ML_Antiquote.inline @{binding "typ_pat"} (parser >> typ_pat)
   894 end*}
   887 end*}
   895 
   888 
       
   889 text {*
       
   890   which can be installed with
       
   891 *}
       
   892 
   896 setup{* type_pat_setup *}
   893 setup{* type_pat_setup *}
   897 
   894 
   898 text {*
   895 text {*
       
   896   However, a word of warning is in order: Introducing new antiquotations
       
   897   should be done only after careful deliberations. They can make your 
       
   898   code harder to read, than making it easier. 
       
   899 
   899   \begin{readmore}
   900   \begin{readmore}
   900   The file @{ML_file "Pure/ML/ml_antiquote.ML"} contains the the definitions
   901   The file @{ML_file "Pure/ML/ml_antiquote.ML"} contains the the definitions
   901   for most antiquotations. Most of the basic operations on ML-syntax are implemented 
   902   for most antiquotations. Most of the basic operations on ML-syntax are implemented 
   902   in @{ML_file "Pure/ML/ml_syntax.ML"}.
   903   in @{ML_file "Pure/ML/ml_syntax.ML"}.
   903   \end{readmore}
   904   \end{readmore}
   917   In fact, it allows one to inject and to project data of \emph{arbitrary} type. This is
   918   In fact, it allows one to inject and to project data of \emph{arbitrary} type. This is
   918   in contrast to datatypes, which only allow injection and projection of data for
   919   in contrast to datatypes, which only allow injection and projection of data for
   919   some \emph{fixed} collection of types. In light of the conventional wisdom cited
   920   some \emph{fixed} collection of types. In light of the conventional wisdom cited
   920   above it is important to keep in mind that the universal type does not
   921   above it is important to keep in mind that the universal type does not
   921   destroy type-safety of ML: storing and accessing the data can only be done
   922   destroy type-safety of ML: storing and accessing the data can only be done
   922   in a type-safe manner.
   923   in a type-safe manner...though run-time checks are needed for that.
   923 
   924 
   924   \begin{readmore}
   925   \begin{readmore}
   925   In Isabelle the universal type is implemented as the type @{ML_type
   926   In Isabelle the universal type is implemented as the type @{ML_type
   926   Universal.universal} in the file
   927   Universal.universal} in the file
   927   @{ML_file "Pure/ML-Systems/universal.ML"}.
   928   @{ML_file "Pure/ML-Systems/universal.ML"}.
   948   then store them in a @{ML_type "Universal.universal list"} as follows:
   949   then store them in a @{ML_type "Universal.universal list"} as follows:
   949 *}
   950 *}
   950 
   951 
   951 ML{*val foo_list = 
   952 ML{*val foo_list = 
   952 let
   953 let
   953   val thirteen  = inject_int 13
   954   val thirteen = inject_int 13
   954   val truth_val = inject_bool true
   955   val truth_val = inject_bool true
   955 in
   956 in
   956   [thirteen, truth_val]
   957   [thirteen, truth_val]
   957 end*}
   958 end*}
   958 
   959 
  1113 
  1114 
  1114   @{ML_response_fake [display,gray]
  1115   @{ML_response_fake [display,gray]
  1115   "WrongRefData.get @{theory}"
  1116   "WrongRefData.get @{theory}"
  1116   "ref [1, 1]"}
  1117   "ref [1, 1]"}
  1117 
  1118 
  1118   Now imagine how often you go backwards and forwards in your proof scripts. 
  1119   Now imagine how often you go backwards and forwards in your proof 
       
  1120   scripts.\footnote{The same problem can be triggered in the Jedit GUI by
       
  1121   making the parser to go over and over again over the \isacommand{setup} command.} 
  1119   By using references in Isabelle code, you are bound to cause all
  1122   By using references in Isabelle code, you are bound to cause all
  1120   hell to break loose. Therefore observe the coding convention: 
  1123   hell to break loose. Therefore observe the coding convention: 
  1121   Do not use references for storing data!
  1124   Do not use references for storing data!
  1122 
  1125 
  1123   \begin{readmore}
  1126   \begin{readmore}
  1147 
  1150 
  1148 ML{*fun update trm = Data.map (fn trms => trm::trms)
  1151 ML{*fun update trm = Data.map (fn trms => trm::trms)
  1149 
  1152 
  1150 fun print ctxt =
  1153 fun print ctxt =
  1151   case (Data.get ctxt) of
  1154   case (Data.get ctxt) of
  1152     [] => writeln "Empty!"
  1155     [] => pwriteln (Pretty.str "Empty!")
  1153   | trms => pwriteln (pretty_terms ctxt trms)*}
  1156   | trms => pwriteln (pretty_terms ctxt trms)*}
  1154 
  1157 
  1155 text {*
  1158 text {*
  1156   Next we start with the context generated by the antiquotation 
  1159   Next we start with the context generated by the antiquotation 
  1157   @{text "@{context}"} and update it in various ways. 
  1160   @{text "@{context}"} and update it in various ways. 
  1310 end"
  1313 end"
  1311   "(\"some string\", \"foo\", \"bar\")"}
  1314   "(\"some string\", \"foo\", \"bar\")"}
  1312 
  1315 
  1313   A concrete example for a configuration value is 
  1316   A concrete example for a configuration value is 
  1314   @{ML_ind simp_trace in Raw_Simplifier}, which switches on trace information
  1317   @{ML_ind simp_trace in Raw_Simplifier}, which switches on trace information
  1315   in the simplifier. This can be used as in the following proof
  1318   in the simplifier. This can be used for example in the following proof
  1316 *}
  1319 *}
  1317 
  1320 
  1318 lemma
  1321 lemma
  1319   shows "(False \<or> True) \<and> True"
  1322   shows "(False \<or> True) \<and> True"
  1320 proof (rule conjI)
  1323 proof (rule conjI)