ProgTutorial/Advanced.thy
changeset 565 cecd7a941885
parent 562 daf404920ab9
child 567 f7c97e64cc2a
equal deleted inserted replaced
564:6e2479089226 565:cecd7a941885
     1 theory Advanced
     1 theory Advanced
     2 imports Base First_Steps
     2 imports Base First_Steps
     3 begin
     3 begin
     4 
     4 
     5 
     5 
     6 chapter {* Advanced Isabelle\label{chp:advanced} *}
     6 chapter \<open>Advanced Isabelle\label{chp:advanced}\<close>
     7 
     7 
     8 text {*
     8 text \<open>
     9    \begin{flushright}
     9    \begin{flushright}
    10   {\em All things are difficult before they are easy.} \\[1ex]
    10   {\em All things are difficult before they are easy.} \\[1ex]
    11   proverb\\[2ex]
    11   proverb\\[2ex]
    12   {\em Programming is not just an act of telling a computer what 
    12   {\em Programming is not just an act of telling a computer what 
    13   to do: it is also an act of telling other programmers what you 
    13   to do: it is also an act of telling other programmers what you 
    25   can be seen as the ``long-term memory'' of Isabelle. There is usually only
    25   can be seen as the ``long-term memory'' of Isabelle. There is usually only
    26   one theory active at each moment. Proof contexts and local theories, on the
    26   one theory active at each moment. Proof contexts and local theories, on the
    27   other hand, store local data for a task at hand. They act like the
    27   other hand, store local data for a task at hand. They act like the
    28   ``short-term memory'' and there can be many of them that are active in
    28   ``short-term memory'' and there can be many of them that are active in
    29   parallel.
    29   parallel.
    30 *}
    30 \<close>
    31 
    31 
    32 section {* Theories\label{sec:theories} *}
    32 section \<open>Theories\label{sec:theories}\<close>
    33 
    33 
    34 text {*
    34 text \<open>
    35   Theories, as said above, are the most basic layer of abstraction in
    35   Theories, as said above, are the most basic layer of abstraction in
    36   Isabelle. They record information about definitions, syntax declarations, axioms,
    36   Isabelle. They record information about definitions, syntax declarations, axioms,
    37   theorems and much more.  For example, if a definition is made, it
    37   theorems and much more.  For example, if a definition is made, it
    38   must be stored in a theory in order to be usable later on. Similar
    38   must be stored in a theory in order to be usable later on. Similar
    39   with proofs: once a proof is finished, the proved theorem needs to
    39   with proofs: once a proof is finished, the proved theorem needs to
    41   usable. All relevant data of a theory can be queried with the
    41   usable. All relevant data of a theory can be queried with the
    42   Isabelle command \isacommand{print\_theory}.
    42   Isabelle command \isacommand{print\_theory}.
    43 
    43 
    44   \begin{isabelle}
    44   \begin{isabelle}
    45   \isacommand{print\_theory}\\
    45   \isacommand{print\_theory}\\
    46   @{text "> names: Pure Code_Generator HOL \<dots>"}\\
    46   \<open>> names: Pure Code_Generator HOL \<dots>\<close>\\
    47   @{text "> classes: Inf < type \<dots>"}\\
    47   \<open>> classes: Inf < type \<dots>\<close>\\
    48   @{text "> default sort: type"}\\
    48   \<open>> default sort: type\<close>\\
    49   @{text "> syntactic types: #prop \<dots>"}\\
    49   \<open>> syntactic types: #prop \<dots>\<close>\\
    50   @{text "> logical types: 'a \<times> 'b \<dots>"}\\
    50   \<open>> logical types: 'a \<times> 'b \<dots>\<close>\\
    51   @{text "> type arities: * :: (random, random) random \<dots>"}\\
    51   \<open>> type arities: * :: (random, random) random \<dots>\<close>\\
    52   @{text "> logical constants: == :: 'a \<Rightarrow> 'a \<Rightarrow> prop \<dots>"}\\
    52   \<open>> logical constants: == :: 'a \<Rightarrow> 'a \<Rightarrow> prop \<dots>\<close>\\
    53   @{text "> abbreviations: \<dots>"}\\
    53   \<open>> abbreviations: \<dots>\<close>\\
    54   @{text "> axioms: \<dots>"}\\
    54   \<open>> axioms: \<dots>\<close>\\
    55   @{text "> oracles: \<dots>"}\\
    55   \<open>> oracles: \<dots>\<close>\\
    56   @{text "> definitions: \<dots>"}\\
    56   \<open>> definitions: \<dots>\<close>\\
    57   @{text "> theorems: \<dots>"}
    57   \<open>> theorems: \<dots>\<close>
    58   \end{isabelle}
    58   \end{isabelle}
    59 
    59 
    60   Functions acting on theories often end with the suffix @{text "_global"},
    60   Functions acting on theories often end with the suffix \<open>_global\<close>,
    61   for example the function @{ML read_term_global in Syntax} in the structure
    61   for example the function @{ML read_term_global in Syntax} in the structure
    62   @{ML_struct Syntax}. The reason is to set them syntactically apart from
    62   @{ML_struct Syntax}. The reason is to set them syntactically apart from
    63   functions acting on contexts or local theories, which will be discussed in
    63   functions acting on contexts or local theories, which will be discussed in
    64   the next sections. There is a tendency amongst Isabelle developers to prefer
    64   the next sections. There is a tendency amongst Isabelle developers to prefer
    65   ``non-global'' operations, because they have some advantages, as we will also
    65   ``non-global'' operations, because they have some advantages, as we will also
    75   stored.  This is a fundamental principle in Isabelle. A similar situation
    75   stored.  This is a fundamental principle in Isabelle. A similar situation
    76   arises with declaring a constant, which can be done on the ML-level with 
    76   arises with declaring a constant, which can be done on the ML-level with 
    77   function @{ML_ind declare_const in Sign} from the structure @{ML_struct
    77   function @{ML_ind declare_const in Sign} from the structure @{ML_struct
    78   Sign}. To see how \isacommand{setup} works, consider the following code:
    78   Sign}. To see how \isacommand{setup} works, consider the following code:
    79 
    79 
    80 *}  
    80 \<close>  
    81 
    81 
    82 ML %grayML{*let
    82 ML %grayML\<open>let
    83   val thy = @{theory}
    83   val thy = @{theory}
    84   val bar_const = ((@{binding "BAR"}, @{typ "nat"}), NoSyn)
    84   val bar_const = ((@{binding "BAR"}, @{typ "nat"}), NoSyn)
    85 in 
    85 in 
    86   Sign.declare_const @{context} bar_const thy  
    86   Sign.declare_const @{context} bar_const thy  
    87 end*}
    87 end\<close>
    88 
    88 
    89 text {*
    89 text \<open>
    90   If you simply run this code\footnote{Recall that ML-code needs to be enclosed in
    90   If you simply run this code\footnote{Recall that ML-code needs to be enclosed in
    91   \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"}.} with the
    91   \isacommand{ML}~\<open>\<verbopen> \<dots> \<verbclose>\<close>.} with the
    92   intention of declaring a constant @{text "BAR"} having type @{typ nat}, then 
    92   intention of declaring a constant \<open>BAR\<close> having type @{typ nat}, then 
    93   indeed you obtain a theory as result. But if you query the
    93   indeed you obtain a theory as result. But if you query the
    94   constant on the Isabelle level using the command \isacommand{term}
    94   constant on the Isabelle level using the command \isacommand{term}
    95 
    95 
    96   \begin{isabelle}
    96   \begin{isabelle}
    97   \isacommand{term}~@{text BAR}\\
    97   \isacommand{term}~\<open>BAR\<close>\\
    98   @{text "> \"BAR\" :: \"'a\""}
    98   \<open>> "BAR" :: "'a"\<close>
    99   \end{isabelle}
    99   \end{isabelle}
   100 
   100 
   101   you can see that you do \emph{not} obtain the expected constant of type @{typ nat}, but a free 
   101   you can see that you do \emph{not} obtain the expected constant of type @{typ nat}, but a free 
   102   variable (printed in blue) of polymorphic type. The problem is that the 
   102   variable (printed in blue) of polymorphic type. The problem is that the 
   103   ML-expression above did not ``register'' the declaration with the current theory. 
   103   ML-expression above did not ``register'' the declaration with the current theory. 
   104   This is what the command \isacommand{setup} is for. The constant is properly 
   104   This is what the command \isacommand{setup} is for. The constant is properly 
   105   declared with
   105   declared with
   106 *}
   106 \<close>
   107 
   107 
   108 setup %gray {* fn thy => 
   108 setup %gray \<open>fn thy => 
   109 let
   109 let
   110   val bar_const = ((@{binding "BAR"}, @{typ "nat"}), NoSyn)
   110   val bar_const = ((@{binding "BAR"}, @{typ "nat"}), NoSyn)
   111   val (_, thy') = Sign.declare_const @{context} bar_const thy
   111   val (_, thy') = Sign.declare_const @{context} bar_const thy
   112 in 
   112 in 
   113   thy'
   113   thy'
   114 end *}
   114 end\<close>
   115 
   115 
   116 text {* 
   116 text \<open>
   117   where the declaration is actually applied to the current theory and
   117   where the declaration is actually applied to the current theory and
   118   
   118   
   119   \begin{isabelle}
   119   \begin{isabelle}
   120   \isacommand{term}~@{text [quotes] "BAR"}\\
   120   \isacommand{term}~@{text [quotes] "BAR"}\\
   121   @{text "> \"BAR\" :: \"nat\""}
   121   \<open>> "BAR" :: "nat"\<close>
   122   \end{isabelle}
   122   \end{isabelle}
   123 
   123 
   124   now returns a (black) constant with the type @{typ nat}, as expected.
   124   now returns a (black) constant with the type @{typ nat}, as expected.
   125 
   125 
   126   In a sense, \isacommand{setup} can be seen as a transaction that
   126   In a sense, \isacommand{setup} can be seen as a transaction that
   127   takes the current theory @{text thy}, applies an operation, and
   127   takes the current theory \<open>thy\<close>, applies an operation, and
   128   produces a new current theory @{text thy'}. This means that we have
   128   produces a new current theory \<open>thy'\<close>. This means that we have
   129   to be careful to apply operations always to the most current theory,
   129   to be careful to apply operations always to the most current theory,
   130   not to a \emph{stale} one. Consider again the function inside the
   130   not to a \emph{stale} one. Consider again the function inside the
   131   \isacommand{setup}-command:
   131   \isacommand{setup}-command:
   132 
   132 
   133   \begin{isabelle}
   133   \begin{isabelle}
   134   \begin{graybox}
   134   \begin{graybox}
   135   \isacommand{setup}~@{text "\<verbopen>"} @{ML
   135   \isacommand{setup}~\<open>\<verbopen>\<close> @{ML
   136 "fn thy => 
   136 "fn thy => 
   137 let
   137 let
   138   val bar_const = ((@{binding \"BAR\"}, @{typ \"nat\"}), NoSyn)
   138   val bar_const = ((@{binding \"BAR\"}, @{typ \"nat\"}), NoSyn)
   139   val (_, thy') = Sign.declare_const @{context} bar_const thy
   139   val (_, thy') = Sign.declare_const @{context} bar_const thy
   140 in
   140 in
   141   thy
   141   thy
   142 end"}~@{text "\<verbclose>"}\isanewline
   142 end"}~\<open>\<verbclose>\<close>\isanewline
   143   @{text "> ERROR: \"Stale theory encountered\""}
   143   \<open>> ERROR: "Stale theory encountered"\<close>
   144   \end{graybox}
   144   \end{graybox}
   145   \end{isabelle}
   145   \end{isabelle}
   146 
   146 
   147   This time we erroneously return the original theory @{text thy}, instead of
   147   This time we erroneously return the original theory \<open>thy\<close>, instead of
   148   the modified one @{text thy'}. Such buggy code will always result into 
   148   the modified one \<open>thy'\<close>. Such buggy code will always result into 
   149   a runtime error message about stale theories.
   149   a runtime error message about stale theories.
   150 
   150 
   151   \begin{readmore}
   151   \begin{readmore}
   152   Most of the functions about theories are implemented in
   152   Most of the functions about theories are implemented in
   153   @{ML_file "Pure/theory.ML"} and @{ML_file "Pure/global_theory.ML"}.
   153   @{ML_file "Pure/theory.ML"} and @{ML_file "Pure/global_theory.ML"}.
   154   \end{readmore}
   154   \end{readmore}
   155 *}
   155 \<close>
   156 
   156 
   157 section {* Contexts *}
   157 section \<open>Contexts\<close>
   158 
   158 
   159 text {*
   159 text \<open>
   160   Contexts are arguably more important than theories, even though they only 
   160   Contexts are arguably more important than theories, even though they only 
   161   contain ``short-term memory data''. The reason is that a vast number of
   161   contain ``short-term memory data''. The reason is that a vast number of
   162   functions in Isabelle depend in one way or another on contexts. Even such
   162   functions in Isabelle depend in one way or another on contexts. Even such
   163   mundane operations like printing out a term make essential use of contexts.
   163   mundane operations like printing out a term make essential use of contexts.
   164   For this consider the following contrived proof-snippet whose purpose is to 
   164   For this consider the following contrived proof-snippet whose purpose is to 
   165   fix two variables:
   165   fix two variables:
   166 *}
   166 \<close>
   167 
   167 
   168 lemma "True"
   168 lemma "True"
   169 proof -
   169 proof -
   170   ML_prf {* Variable.dest_fixes @{context} *} 
   170   ML_prf \<open>Variable.dest_fixes @{context}\<close> 
   171   fix x y  
   171   fix x y  
   172   ML_prf {* Variable.dest_fixes @{context} *}(*<*)oops(*>*)
   172   ML_prf \<open>Variable.dest_fixes @{context}\<close>(*<*)oops(*>*)
   173 
   173 
   174 text {*
   174 text \<open>
   175   The interesting point is that we injected ML-code before and after
   175   The interesting point is that we injected ML-code before and after
   176   the variables are fixed. For this remember that ML-code inside a proof
   176   the variables are fixed. For this remember that ML-code inside a proof
   177   needs to be enclosed inside \isacommand{ML\_prf}~@{text "\<verbopen> \<dots> \<verbclose>"},
   177   needs to be enclosed inside \isacommand{ML\_prf}~\<open>\<verbopen> \<dots> \<verbclose>\<close>,
   178   not \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"}. The function 
   178   not \isacommand{ML}~\<open>\<verbopen> \<dots> \<verbclose>\<close>. The function 
   179   @{ML_ind dest_fixes in Variable} from the structure @{ML_struct Variable} takes 
   179   @{ML_ind dest_fixes in Variable} from the structure @{ML_struct Variable} takes 
   180   a context and returns all its currently fixed variable (names). That 
   180   a context and returns all its currently fixed variable (names). That 
   181   means a context has a dataslot containing information about fixed variables.
   181   means a context has a dataslot containing information about fixed variables.
   182   The ML-antiquotation @{text "@{context}"} points to the context that is
   182   The ML-antiquotation \<open>@{context}\<close> points to the context that is
   183   active at that point of the theory. Consequently, in the first call to 
   183   active at that point of the theory. Consequently, in the first call to 
   184   @{ML dest_fixes in Variable} this dataslot is  empty; in the second it is 
   184   @{ML dest_fixes in Variable} this dataslot is  empty; in the second it is 
   185   filled with @{text x} and @{text y}. What is interesting is that contexts
   185   filled with \<open>x\<close> and \<open>y\<close>. What is interesting is that contexts
   186   can be stacked. For this consider the following proof fragment:
   186   can be stacked. For this consider the following proof fragment:
   187 *}
   187 \<close>
   188 
   188 
   189 lemma "True"
   189 lemma "True"
   190 proof -
   190 proof -
   191   fix x y
   191   fix x y
   192   { fix z w
   192   { fix z w
   193     ML_prf {* Variable.dest_fixes @{context} *} 
   193     ML_prf \<open>Variable.dest_fixes @{context}\<close> 
   194   }
   194   }
   195   ML_prf {* Variable.dest_fixes @{context} *}(*<*)oops(*>*)
   195   ML_prf \<open>Variable.dest_fixes @{context}\<close>(*<*)oops(*>*)
   196 
   196 
   197 text {*
   197 text \<open>
   198   Here the first time we call @{ML dest_fixes in Variable} we have four fixes variables;
   198   Here the first time we call @{ML dest_fixes in Variable} we have four fixes variables;
   199   the second time we get only the fixes variables @{text x} and @{text y} as answer, since
   199   the second time we get only the fixes variables \<open>x\<close> and \<open>y\<close> as answer, since
   200   @{text z} and @{text w} are not anymore in the scope of the context. 
   200   \<open>z\<close> and \<open>w\<close> are not anymore in the scope of the context. 
   201   This means the curly-braces act on the Isabelle level as opening and closing statements 
   201   This means the curly-braces act on the Isabelle level as opening and closing statements 
   202   for a context. The above proof fragment corresponds roughly to the following ML-code
   202   for a context. The above proof fragment corresponds roughly to the following ML-code
   203 *}
   203 \<close>
   204 
   204 
   205 ML %grayML{*val ctxt0 = @{context};
   205 ML %grayML\<open>val ctxt0 = @{context};
   206 val ([x, y], ctxt1) = Variable.add_fixes ["x", "y"] ctxt0;
   206 val ([x, y], ctxt1) = Variable.add_fixes ["x", "y"] ctxt0;
   207 val ([z, w], ctxt2) = Variable.add_fixes ["z", "w"] ctxt1*}
   207 val ([z, w], ctxt2) = Variable.add_fixes ["z", "w"] ctxt1\<close>
   208 
   208 
   209 text {*
   209 text \<open>
   210   where the function @{ML_ind add_fixes in Variable} fixes a list of variables
   210   where the function @{ML_ind add_fixes in Variable} fixes a list of variables
   211   specified by strings. Let us come back to the point about printing terms. Consider
   211   specified by strings. Let us come back to the point about printing terms. Consider
   212   printing out the term \mbox{@{text "(x, y, z, w)"}} using our function @{ML_ind pretty_term}.
   212   printing out the term \mbox{\<open>(x, y, z, w)\<close>} using our function @{ML_ind pretty_term}.
   213   This function takes a term and a context as argument. Notice how the printing
   213   This function takes a term and a context as argument. Notice how the printing
   214   of the term changes according to which context is used.
   214   of the term changes according to which context is used.
   215 
   215 
   216   \begin{isabelle}
   216   \begin{isabelle}
   217   \begin{graybox}
   217   \begin{graybox}
   222     [ pretty_term ctxt0 trm,
   222     [ pretty_term ctxt0 trm,
   223       pretty_term ctxt1 trm,
   223       pretty_term ctxt1 trm,
   224       pretty_term ctxt2 trm ])
   224       pretty_term ctxt2 trm ])
   225 end"}\\
   225 end"}\\
   226   \setlength{\fboxsep}{0mm}
   226   \setlength{\fboxsep}{0mm}
   227   @{text ">"}~@{text "("}\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text x}}}@{text ","}~%
   227   \<open>>\<close>~\<open>(\<close>\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>x\<close>}}\<open>,\<close>~%
   228   \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text y}}}@{text ","}~%
   228   \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>y\<close>}}\<open>,\<close>~%
   229   \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text z}}}@{text ","}~%
   229   \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>z\<close>}}\<open>,\<close>~%
   230   \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text w}}}@{text ")"}\\
   230   \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>w\<close>}}\<open>)\<close>\\
   231   @{text ">"}~@{text "("}\colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{@{text x}}}@{text ","}~%
   231   \<open>>\<close>~\<open>(\<close>\colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{\<open>x\<close>}}\<open>,\<close>~%
   232   \colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{@{text y}}}@{text ","}~%
   232   \colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{\<open>y\<close>}}\<open>,\<close>~%
   233   \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text z}}}@{text ","}~%
   233   \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>z\<close>}}\<open>,\<close>~%
   234   \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text w}}}@{text ")"}\\
   234   \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>w\<close>}}\<open>)\<close>\\
   235   @{text ">"}~@{text "("}\colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{@{text x}}}@{text ","}~%
   235   \<open>>\<close>~\<open>(\<close>\colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{\<open>x\<close>}}\<open>,\<close>~%
   236   \colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{@{text y}}}@{text ","}~%
   236   \colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{\<open>y\<close>}}\<open>,\<close>~%
   237   \colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{@{text z}}}@{text ","}~%
   237   \colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{\<open>z\<close>}}\<open>,\<close>~%
   238   \colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{@{text w}}}@{text ")"}
   238   \colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{\<open>w\<close>}}\<open>)\<close>
   239   \end{graybox}
   239   \end{graybox}
   240   \end{isabelle}
   240   \end{isabelle}
   241 
   241 
   242 
   242 
   243   The term we are printing is in all three cases the same---a tuple containing
   243   The term we are printing is in all three cases the same---a tuple containing
   244   four free variables. As you can see, however, in case of @{ML "ctxt0"} all
   244   four free variables. As you can see, however, in case of @{ML "ctxt0"} all
   245   variables are highlighted indicating a problem, while in case of @{ML
   245   variables are highlighted indicating a problem, while in case of @{ML
   246   "ctxt1"} @{text x} and @{text y} are printed as normal (blue) free
   246   "ctxt1"} \<open>x\<close> and \<open>y\<close> are printed as normal (blue) free
   247   variables, but not @{text z} and @{text w}. In the last case all variables
   247   variables, but not \<open>z\<close> and \<open>w\<close>. In the last case all variables
   248   are printed as expected. The point of this example is that the context
   248   are printed as expected. The point of this example is that the context
   249   contains the information which variables are fixed, and designates all other
   249   contains the information which variables are fixed, and designates all other
   250   free variables as being alien or faulty.  Therefore the highlighting. 
   250   free variables as being alien or faulty.  Therefore the highlighting. 
   251   While this seems like a minor detail, the concept of making the context aware 
   251   While this seems like a minor detail, the concept of making the context aware 
   252   of fixed variables is actually quite useful. For example it prevents us from 
   252   of fixed variables is actually quite useful. For example it prevents us from 
   264   @{ML_response_fake [gray, display]
   264   @{ML_response_fake [gray, display]
   265   "@{context}
   265   "@{context}
   266 |> Variable.variant_fixes [\"y\", \"y\", \"z\"]" 
   266 |> Variable.variant_fixes [\"y\", \"y\", \"z\"]" 
   267   "([\"y\", \"ya\", \"z\"], ...)"}
   267   "([\"y\", \"ya\", \"z\"], ...)"}
   268 
   268 
   269   Now a fresh variant for the second occurence of @{text y} is created
   269   Now a fresh variant for the second occurence of \<open>y\<close> is created
   270   avoiding any clash. In this way we can also create fresh free variables
   270   avoiding any clash. In this way we can also create fresh free variables
   271   that avoid any clashes with fixed variables. In Line~3 below we fix
   271   that avoid any clashes with fixed variables. In Line~3 below we fix
   272   the variable @{text x} in the context @{text ctxt1}. Next we want to
   272   the variable \<open>x\<close> in the context \<open>ctxt1\<close>. Next we want to
   273   create two fresh variables of type @{typ nat} as variants of the
   273   create two fresh variables of type @{typ nat} as variants of the
   274   string @{text [quotes] "x"} (Lines 6 and 7).
   274   string @{text [quotes] "x"} (Lines 6 and 7).
   275 
   275 
   276   @{ML_response_fake [display, gray, linenos]
   276   @{ML_response_fake [display, gray, linenos]
   277   "let
   277   "let
   283    Variable.variant_frees ctxt1 [] frees)
   283    Variable.variant_frees ctxt1 [] frees)
   284 end"
   284 end"
   285   "([(\"x\", \"nat\"), (\"xa\", \"nat\")], 
   285   "([(\"x\", \"nat\"), (\"xa\", \"nat\")], 
   286  [(\"xa\", \"nat\"), (\"xb\", \"nat\")])"}
   286  [(\"xa\", \"nat\"), (\"xb\", \"nat\")])"}
   287 
   287 
   288   As you can see, if we create the fresh variables with the context @{text ctxt0},
   288   As you can see, if we create the fresh variables with the context \<open>ctxt0\<close>,
   289   then we obtain @{text "x"} and @{text "xa"}; but in @{text ctxt1} we obtain @{text "xa"}
   289   then we obtain \<open>x\<close> and \<open>xa\<close>; but in \<open>ctxt1\<close> we obtain \<open>xa\<close>
   290   and @{text "xb"} avoiding @{text x}, which was fixed in this context.
   290   and \<open>xb\<close> avoiding \<open>x\<close>, which was fixed in this context.
   291 
   291 
   292   Often one has the problem that given some terms, create fresh variables
   292   Often one has the problem that given some terms, create fresh variables
   293   avoiding any variable occurring in those terms. For this you can use the
   293   avoiding any variable occurring in those terms. For this you can use the
   294   function @{ML_ind declare_term in Variable} from the structure @{ML_struct Variable}.
   294   function @{ML_ind declare_term in Variable} from the structure @{ML_struct Variable}.
   295 
   295 
   300 in
   300 in
   301   Variable.variant_frees ctxt1 [] frees
   301   Variable.variant_frees ctxt1 [] frees
   302 end"
   302 end"
   303   "[(\"xb\", \"nat\"), (\"xc\", \"nat\")]"}
   303   "[(\"xb\", \"nat\"), (\"xc\", \"nat\")]"}
   304 
   304 
   305   The result is @{text xb} and @{text xc} for the names of the fresh
   305   The result is \<open>xb\<close> and \<open>xc\<close> for the names of the fresh
   306   variables, since @{text x} and @{text xa} occur in the term we declared. 
   306   variables, since \<open>x\<close> and \<open>xa\<close> occur in the term we declared. 
   307   Note that @{ML_ind declare_term in Variable} does not fix the
   307   Note that @{ML_ind declare_term in Variable} does not fix the
   308   variables; it just makes them ``known'' to the context. You can see
   308   variables; it just makes them ``known'' to the context. You can see
   309   that if you print out a declared term.
   309   that if you print out a declared term.
   310 
   310 
   311   \begin{isabelle}
   311   \begin{isabelle}
   315   val ctxt1 = Variable.declare_term trm @{context}
   315   val ctxt1 = Variable.declare_term trm @{context}
   316 in
   316 in
   317   pwriteln (pretty_term ctxt1 trm)
   317   pwriteln (pretty_term ctxt1 trm)
   318 end"}\\
   318 end"}\\
   319   \setlength{\fboxsep}{0mm}
   319   \setlength{\fboxsep}{0mm}
   320   @{text ">"}~\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text P}}}~%
   320   \<open>>\<close>~\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>P\<close>}}~%
   321   \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text x}}}~%
   321   \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>x\<close>}}~%
   322   \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text y}}}~%
   322   \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>y\<close>}}~%
   323   \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text z}}}
   323   \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>z\<close>}}
   324   \end{graybox}
   324   \end{graybox}
   325   \end{isabelle}
   325   \end{isabelle}
   326 
   326 
   327   All variables are highligted, indicating that they are not
   327   All variables are highligted, indicating that they are not
   328   fixed. However, declaring a term is helpful when parsing terms using
   328   fixed. However, declaring a term is helpful when parsing terms using
   337   (Syntax.read_term ctxt0 \"x\", 
   337   (Syntax.read_term ctxt0 \"x\", 
   338    Syntax.read_term ctxt1 \"x\") 
   338    Syntax.read_term ctxt1 \"x\") 
   339 end"
   339 end"
   340   "(Free (\"x\", \"'a\"), Free (\"x\", \"nat\"))"}
   340   "(Free (\"x\", \"'a\"), Free (\"x\", \"nat\"))"}
   341   
   341   
   342   Parsing the string in the context @{text "ctxt0"} results in a free variable 
   342   Parsing the string in the context \<open>ctxt0\<close> results in a free variable 
   343   with a default polymorphic type, but in case of @{text "ctxt1"} we obtain a
   343   with a default polymorphic type, but in case of \<open>ctxt1\<close> we obtain a
   344   free variable of type @{typ nat} as previously declared. Which
   344   free variable of type @{typ nat} as previously declared. Which
   345   type the parsed term receives depends upon the last declaration that
   345   type the parsed term receives depends upon the last declaration that
   346   is made, as the next example illustrates.
   346   is made, as the next example illustrates.
   347 
   347 
   348   @{ML_response_fake [gray, display]
   348   @{ML_response_fake [gray, display]
   370   |> pretty_term ctxt0
   370   |> pretty_term ctxt0
   371   |> pwriteln
   371   |> pwriteln
   372 end"}
   372 end"}
   373   \end{linenos}
   373   \end{linenos}
   374   \setlength{\fboxsep}{0mm}
   374   \setlength{\fboxsep}{0mm}
   375   @{text ">"}~\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text P}}}~%
   375   \<open>>\<close>~\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>P\<close>}}~%
   376   @{text "?x ?y ?z"}
   376   \<open>?x ?y ?z\<close>
   377   \end{graybox}
   377   \end{graybox}
   378   \end{isabelle}
   378   \end{isabelle}
   379 
   379 
   380   In Line 3 we fix the variables @{term x}, @{term y} and @{term z} in
   380   In Line 3 we fix the variables @{term x}, @{term y} and @{term z} in
   381   context @{text ctxt1}.  The function @{ML_ind export_terms in
   381   context \<open>ctxt1\<close>.  The function @{ML_ind export_terms in
   382   Variable} from the structure @{ML_struct Variable} can be used to transfer
   382   Variable} from the structure @{ML_struct Variable} can be used to transfer
   383   terms between contexts. Transferring means to turn all (free)
   383   terms between contexts. Transferring means to turn all (free)
   384   variables that are fixed in one context, but not in the other, into
   384   variables that are fixed in one context, but not in the other, into
   385   schematic variables. In our example, we are transferring the term
   385   schematic variables. In our example, we are transferring the term
   386   @{text "P x y z"} from context @{text "ctxt1"} to @{text "ctxt0"},
   386   \<open>P x y z\<close> from context \<open>ctxt1\<close> to \<open>ctxt0\<close>,
   387   which means @{term x}, @{term y} and @{term z} become schematic
   387   which means @{term x}, @{term y} and @{term z} become schematic
   388   variables (as can be seen by the leading question marks in the result). 
   388   variables (as can be seen by the leading question marks in the result). 
   389   Note that the variable @{text P} stays a free variable, since it not fixed in
   389   Note that the variable \<open>P\<close> stays a free variable, since it not fixed in
   390   @{text ctxt1}; it is even highlighed, because @{text "ctxt0"} does
   390   \<open>ctxt1\<close>; it is even highlighed, because \<open>ctxt0\<close> does
   391   not know about it. Note also that in Line 6 we had to use the
   391   not know about it. Note also that in Line 6 we had to use the
   392   function @{ML_ind singleton}, because the function @{ML_ind
   392   function @{ML_ind singleton}, because the function @{ML_ind
   393   export_terms in Variable} normally works over lists of terms.
   393   export_terms in Variable} normally works over lists of terms.
   394 
   394 
   395   The case of transferring theorems is even more useful. The reason is
   395   The case of transferring theorems is even more useful. The reason is
   409 in
   409 in
   410   singleton (Proof_Context.export ctxt1 ctxt0) foo_thm
   410   singleton (Proof_Context.export ctxt1 ctxt0) foo_thm
   411 end"
   411 end"
   412   "?P ?x ?y ?z ?x ?y ?z"}
   412   "?P ?x ?y ?z ?x ?y ?z"}
   413 
   413 
   414   Since we fixed all variables in @{text ctxt1}, in the exported
   414   Since we fixed all variables in \<open>ctxt1\<close>, in the exported
   415   result all of them are schematic. The great point of contexts is
   415   result all of them are schematic. The great point of contexts is
   416   that exporting from one to another is not just restricted to
   416   that exporting from one to another is not just restricted to
   417   variables, but also works with assumptions. For this we can use the
   417   variables, but also works with assumptions. For this we can use the
   418   function @{ML_ind export in Assumption} from the structure
   418   function @{ML_ind export in Assumption} from the structure
   419   @{ML_struct Assumption}. Consider the following code.
   419   @{ML_struct Assumption}. Consider the following code.
   427   Assumption.export false ctxt1 ctxt0 eq'
   427   Assumption.export false ctxt1 ctxt0 eq'
   428 end"
   428 end"
   429   "x \<equiv> y \<Longrightarrow> y \<equiv> x"}
   429   "x \<equiv> y \<Longrightarrow> y \<equiv> x"}
   430   
   430   
   431   The function @{ML_ind add_assumes in Assumption} from the structure
   431   The function @{ML_ind add_assumes in Assumption} from the structure
   432   @{ML_struct Assumption} adds the assumption \mbox{@{text "x \<equiv> y"}}
   432   @{ML_struct Assumption} adds the assumption \mbox{\<open>x \<equiv> y\<close>}
   433   to the context @{text ctxt1} (Line 3). This function expects a list
   433   to the context \<open>ctxt1\<close> (Line 3). This function expects a list
   434   of @{ML_type cterm}s and returns them as theorems, together with the
   434   of @{ML_type cterm}s and returns them as theorems, together with the
   435   new context in which they are ``assumed''. In Line 4 we use the
   435   new context in which they are ``assumed''. In Line 4 we use the
   436   function @{ML_ind symmetric in Thm} from the structure @{ML_struct
   436   function @{ML_ind symmetric in Thm} from the structure @{ML_struct
   437   Thm} in order to obtain the symmetric version of the assumed
   437   Thm} in order to obtain the symmetric version of the assumed
   438   meta-equality. Now exporting the theorem @{text "eq'"} from @{text
   438   meta-equality. Now exporting the theorem \<open>eq'\<close> from \<open>ctxt1\<close> to \<open>ctxt0\<close> means @{term "y \<equiv> x"} will be prefixed with
   439   ctxt1} to @{text ctxt0} means @{term "y \<equiv> x"} will be prefixed with
       
   440   the assumed theorem. The boolean flag in @{ML_ind export in
   439   the assumed theorem. The boolean flag in @{ML_ind export in
   441   Assumption} indicates whether the assumptions should be marked with
   440   Assumption} indicates whether the assumptions should be marked with
   442   the goal marker (see Section~\ref{sec:basictactics}). In normal
   441   the goal marker (see Section~\ref{sec:basictactics}). In normal
   443   circumstances this is not necessary and so should be set to @{ML
   442   circumstances this is not necessary and so should be set to @{ML
   444   false}.  The result of the export is then the theorem \mbox{@{term
   443   false}.  The result of the export is then the theorem \mbox{@{term
   460   val eq' = Thm.symmetric eq
   459   val eq' = Thm.symmetric eq
   461 in
   460 in
   462   Proof_Context.export ctxt1 ctxt0 [eq']
   461   Proof_Context.export ctxt1 ctxt0 [eq']
   463 end"
   462 end"
   464   "[?x \<equiv> ?y \<Longrightarrow> ?y \<equiv> ?x]"}
   463   "[?x \<equiv> ?y \<Longrightarrow> ?y \<equiv> ?x]"}
   465 *}
   464 \<close>
   466 
   465 
   467 
   466 
   468 
   467 
   469 text {*
   468 text \<open>
   470 
   469 
   471 *}
   470 \<close>
   472 
   471 
   473 
   472 
   474 (*
   473 (*
   475 ML %grayML{*Proof_Context.debug := true*}
   474 ML %grayML{*Proof_Context.debug := true*}
   476 ML %grayML{*Proof_Context.verbose := true*}
   475 ML %grayML{*Proof_Context.verbose := true*}
   503   thm this
   502   thm this
   504 
   503 
   505 oops
   504 oops
   506 *)
   505 *)
   507 
   506 
   508 section {* Local Theories and Local Setups\label{sec:local} (TBD) *}
   507 section \<open>Local Theories and Local Setups\label{sec:local} (TBD)\<close>
   509 
   508 
   510 text {*
   509 text \<open>
   511   In contrast to an ordinary theory, which simply consists of a type
   510   In contrast to an ordinary theory, which simply consists of a type
   512   signature, as well as tables for constants, axioms and theorems, a local
   511   signature, as well as tables for constants, axioms and theorems, a local
   513   theory contains additional context information, such as locally fixed
   512   theory contains additional context information, such as locally fixed
   514   variables and local assumptions that may be used by the package. The type
   513   variables and local assumptions that may be used by the package. The type
   515   @{ML_type local_theory} is identical to the type of \emph{proof contexts}
   514   @{ML_type local_theory} is identical to the type of \emph{proof contexts}
   522    A similar command is \isacommand{local\_setup}, which expects a function
   521    A similar command is \isacommand{local\_setup}, which expects a function
   523   of type @{ML_type "local_theory -> local_theory"}. Later on we will also
   522   of type @{ML_type "local_theory -> local_theory"}. Later on we will also
   524   use the commands \isacommand{method\_setup} for installing methods in the
   523   use the commands \isacommand{method\_setup} for installing methods in the
   525   current theory and \isacommand{simproc\_setup} for adding new simprocs to
   524   current theory and \isacommand{simproc\_setup} for adding new simprocs to
   526   the current simpset.
   525   the current simpset.
   527 *}
   526 \<close>
   528 
   527 
   529 
   528 
   530 section {* Morphisms (TBD) *}
   529 section \<open>Morphisms (TBD)\<close>
   531 
   530 
   532 text {*
   531 text \<open>
   533   Morphisms are arbitrary transformations over terms, types, theorems and bindings.
   532   Morphisms are arbitrary transformations over terms, types, theorems and bindings.
   534   They can be constructed using the function @{ML_ind morphism in Morphism},
   533   They can be constructed using the function @{ML_ind morphism in Morphism},
   535   which expects a record with functions of type
   534   which expects a record with functions of type
   536 
   535 
   537   \begin{isabelle}
   536   \begin{isabelle}
   538   \begin{tabular}{rl}
   537   \begin{tabular}{rl}
   539   @{text "binding:"} & @{text "binding -> binding"}\\
   538   \<open>binding:\<close> & \<open>binding -> binding\<close>\\
   540   @{text "typ:"}     & @{text "typ -> typ"}\\
   539   \<open>typ:\<close>     & \<open>typ -> typ\<close>\\
   541   @{text "term:"}    & @{text "term -> term"}\\
   540   \<open>term:\<close>    & \<open>term -> term\<close>\\
   542   @{text "fact:"}    & @{text "thm list -> thm list"}
   541   \<open>fact:\<close>    & \<open>thm list -> thm list\<close>
   543   \end{tabular}
   542   \end{tabular}
   544   \end{isabelle}
   543   \end{isabelle}
   545 
   544 
   546   The simplest morphism is the  @{ML_ind identity in Morphism}-morphism defined as
   545   The simplest morphism is the  @{ML_ind identity in Morphism}-morphism defined as
   547 *}
   546 \<close>
   548 
   547 
   549 ML %grayML{*val identity = Morphism.morphism "" {binding = [], typ = [], term = [], fact = []}*}
   548 ML %grayML\<open>val identity = Morphism.morphism "" {binding = [], typ = [], term = [], fact = []}\<close>
   550   
   549   
   551 text {*
   550 text \<open>
   552   Morphisms can be composed with the function @{ML_ind "$>" in Morphism}
   551   Morphisms can be composed with the function @{ML_ind "$>" in Morphism}
   553 *}
   552 \<close>
   554 
   553 
   555 ML %grayML{*fun trm_phi (Free (x, T)) = Var ((x, 0), T) 
   554 ML %grayML\<open>fun trm_phi (Free (x, T)) = Var ((x, 0), T) 
   556   | trm_phi (Abs (x, T, t)) = Abs (x, T, trm_phi t)
   555   | trm_phi (Abs (x, T, t)) = Abs (x, T, trm_phi t)
   557   | trm_phi (t $ s) = (trm_phi t) $ (trm_phi s)
   556   | trm_phi (t $ s) = (trm_phi t) $ (trm_phi s)
   558   | trm_phi t = t*}
   557   | trm_phi t = t\<close>
   559 
   558 
   560 ML %grayML{*val phi = Morphism.term_morphism "foo" trm_phi*}
   559 ML %grayML\<open>val phi = Morphism.term_morphism "foo" trm_phi\<close>
   561 
   560 
   562 ML %grayML{*Morphism.term phi @{term "P x y"}*}
   561 ML %grayML\<open>Morphism.term phi @{term "P x y"}\<close>
   563 
   562 
   564 text {*
   563 text \<open>
   565   @{ML_ind term_morphism in Morphism}
   564   @{ML_ind term_morphism in Morphism}
   566 
   565 
   567   @{ML_ind term in Morphism},
   566   @{ML_ind term in Morphism},
   568   @{ML_ind thm in Morphism}
   567   @{ML_ind thm in Morphism}
   569 
   568 
   570   \begin{readmore}
   569   \begin{readmore}
   571   Morphisms are implemented in the file @{ML_file "Pure/morphism.ML"}.
   570   Morphisms are implemented in the file @{ML_file "Pure/morphism.ML"}.
   572   \end{readmore}
   571   \end{readmore}
   573 *}
   572 \<close>
   574 
   573 
   575 section {* Misc (TBD) *}
   574 section \<open>Misc (TBD)\<close>
   576 
   575 
   577 text {* 
   576 text \<open>
   578 FIXME: association lists:
   577 FIXME: association lists:
   579 @{ML_file "Pure/General/alist.ML"}
   578 @{ML_file "Pure/General/alist.ML"}
   580 
   579 
   581 FIXME: calling the ML-compiler
   580 FIXME: calling the ML-compiler
   582 
   581 
   583 *}
   582 \<close>
   584 
   583 
   585 section {* What Is In an Isabelle Name? (TBD) *}
   584 section \<open>What Is In an Isabelle Name? (TBD)\<close>
   586 
   585 
   587 text {*
   586 text \<open>
   588   On the ML-level of Isabelle, you often have to work with qualified names.
   587   On the ML-level of Isabelle, you often have to work with qualified names.
   589   These are strings with some additional information, such as positional
   588   These are strings with some additional information, such as positional
   590   information and qualifiers. Such qualified names can be generated with the
   589   information and qualifiers. Such qualified names can be generated with the
   591   antiquotation @{text "@{binding \<dots>}"}. For example
   590   antiquotation \<open>@{binding \<dots>}\<close>. For example
   592 
   591 
   593   @{ML_response [display,gray]
   592   @{ML_response [display,gray]
   594   "@{binding \"name\"}"
   593   "@{binding \"name\"}"
   595   "name"}
   594   "name"}
   596 
   595 
   597   An example where a qualified name is needed is the function 
   596   An example where a qualified name is needed is the function 
   598   @{ML_ind define in Local_Theory}.  This function is used below to define 
   597   @{ML_ind define in Local_Theory}.  This function is used below to define 
   599   the constant @{term "TrueConj"} as the conjunction @{term "True \<and> True"}.
   598   the constant @{term "TrueConj"} as the conjunction @{term "True \<and> True"}.
   600 *}
   599 \<close>
   601 
   600 
   602 local_setup %gray {* 
   601 local_setup %gray \<open>
   603   Local_Theory.define ((@{binding "TrueConj"}, NoSyn), 
   602   Local_Theory.define ((@{binding "TrueConj"}, NoSyn), 
   604       ((@{binding "TrueConj_def"}, []), @{term "True \<and> True"})) #> snd *}
   603       ((@{binding "TrueConj_def"}, []), @{term "True \<and> True"})) #> snd\<close>
   605 
   604 
   606 text {* 
   605 text \<open>
   607   Now querying the definition you obtain:
   606   Now querying the definition you obtain:
   608 
   607 
   609   \begin{isabelle}
   608   \begin{isabelle}
   610   \isacommand{thm}~@{text "TrueConj_def"}\\
   609   \isacommand{thm}~\<open>TrueConj_def\<close>\\
   611   @{text "> "}~@{thm TrueConj_def}
   610   \<open>> \<close>~@{thm TrueConj_def}
   612   \end{isabelle}
   611   \end{isabelle}
   613 
   612 
   614   \begin{readmore}
   613   \begin{readmore}
   615   The basic operations on bindings are implemented in 
   614   The basic operations on bindings are implemented in 
   616   @{ML_file "Pure/General/binding.ML"}.
   615   @{ML_file "Pure/General/binding.ML"}.
   620   \footnote{\bf FIXME give a pointer to \isacommand{local\_setup}; if not, then explain
   619   \footnote{\bf FIXME give a pointer to \isacommand{local\_setup}; if not, then explain
   621   why @{ML snd} is needed.}
   620   why @{ML snd} is needed.}
   622   \footnote{\bf FIXME: There should probably a separate section on binding, long-names
   621   \footnote{\bf FIXME: There should probably a separate section on binding, long-names
   623   and sign.}
   622   and sign.}
   624 
   623 
   625 *}
   624 \<close>
   626 
   625 
   627 
   626 
   628 ML {* Sign.intern_type @{theory} "list" *}
   627 ML \<open>Sign.intern_type @{theory} "list"\<close>
   629 ML {* Sign.intern_const @{theory} "prod_fun" *}
   628 ML \<open>Sign.intern_const @{theory} "prod_fun"\<close>
   630 
   629 
   631 text {*
   630 text \<open>
   632   \footnote{\bf FIXME: Explain the following better; maybe put in a separate
   631   \footnote{\bf FIXME: Explain the following better; maybe put in a separate
   633   section and link with the comment in the antiquotation section.}
   632   section and link with the comment in the antiquotation section.}
   634 
   633 
   635   Occasionally you have to calculate what the ``base'' name of a given
   634   Occasionally you have to calculate what the ``base'' name of a given
   636   constant is. For this you can use the function @{ML_ind  Long_Name.base_name}. For example:
   635   constant is. For this you can use the function @{ML_ind  Long_Name.base_name}. For example:
   639 
   638 
   640   \begin{readmore}
   639   \begin{readmore}
   641   Functions about naming are implemented in @{ML_file "Pure/General/name_space.ML"};
   640   Functions about naming are implemented in @{ML_file "Pure/General/name_space.ML"};
   642   functions about signatures in @{ML_file "Pure/sign.ML"}.
   641   functions about signatures in @{ML_file "Pure/sign.ML"}.
   643   \end{readmore}
   642   \end{readmore}
   644 *}
   643 \<close>
   645 
   644 
   646 text {* 
   645 text \<open>
   647   @{ML_ind "Binding.name_of"} returns the string without markup
   646   @{ML_ind "Binding.name_of"} returns the string without markup
   648 
   647 
   649   @{ML_ind "Binding.concealed"} 
   648   @{ML_ind "Binding.concealed"} 
   650 *}
   649 \<close>
   651 
   650 
   652 section {* Concurrency (TBD) *}
   651 section \<open>Concurrency (TBD)\<close>
   653 
   652 
   654 text {*
   653 text \<open>
   655   @{ML_ind prove_future in Goal}
   654   @{ML_ind prove_future in Goal}
   656   @{ML_ind future_result in Goal}
   655   @{ML_ind future_result in Goal}
   657 *}
   656 \<close>
   658 
   657 
   659 section {* Parse and Print Translations (TBD) *}
   658 section \<open>Parse and Print Translations (TBD)\<close>
   660 
   659 
   661 section {* Summary *}
   660 section \<open>Summary\<close>
   662 
   661 
   663 text {*
   662 text \<open>
   664   TBD
   663   TBD
   665 *}
   664 \<close>
   666 
   665 
   667 end
   666 end