|     43 | Var of indexname * typ  |     43 | Var of indexname * typ  | 
|     44 | Bound of int  |     44 | Bound of int  | 
|     45 | Abs of string * typ * term  |     45 | Abs of string * typ * term  | 
|     46 | $ of term * term\<close> |     46 | $ of term * term\<close> | 
|     47  |     47  | 
|         |     48 ML \<open> | 
|         |     49 let  | 
|         |     50   val redex = @{term "(\<lambda>(x::int) (y::int). x)"}  | 
|         |     51   val arg1 = @{term "1::int"}  | 
|         |     52   val arg2 = @{term "2::int"} | 
|         |     53 in | 
|         |     54   pretty_term @{context} (redex $ arg1 $ arg2) | 
|         |     55 end\<close> | 
|         |     56  | 
|     48 text \<open> |     57 text \<open> | 
|     49   This datatype implements Church-style lambda-terms, where types are |     58   This datatype implements Church-style lambda-terms, where types are | 
|     50   explicitly recorded in variables, constants and abstractions.  The |     59   explicitly recorded in variables, constants and abstractions.  The | 
|     51   important point of having terms is that you can pattern-match against them; |     60   important point of having terms is that you can pattern-match against them; | 
|     52   this cannot be done with certified terms. As can be seen in Line 5, |     61   this cannot be done with certified terms. As can be seen in Line 5, | 
|     53   terms use the usual de Bruijn index mechanism for representing bound |     62   terms use the usual de Bruijn index mechanism for representing bound | 
|     54   variables.  For example in |     63   variables.  For example in | 
|     55  |     64  | 
|     56   @{ML_matchresult_fake [display, gray] |     65   @{ML_response [display, gray] | 
|     57   \<open>@{term "\<lambda>x y. x y"}\<close> |     66   \<open>@{term "\<lambda>x y. x y"}\<close> | 
|     58   \<open>Abs ("x", "'a \<Rightarrow> 'b", Abs ("y", "'a", Bound 1 $ Bound 0))\<close>} |     67   \<open>Abs ("x", "'a \<Rightarrow> 'b", Abs ("y", "'a", Bound 1 $ Bound 0))\<close>} | 
|     59  |     68  | 
|     60   the indices refer to the number of Abstractions (@{ML Abs}) that we need to |     69   the indices refer to the number of Abstractions (@{ML Abs}) that we need to | 
|     61   skip until we hit the @{ML Abs} that binds the corresponding |     70   skip until we hit the @{ML Abs} that binds the corresponding | 
|     67   term-constructor @{ML $}. |     76   term-constructor @{ML $}. | 
|     68  |     77  | 
|     69   Be careful if you pretty-print terms. Consider pretty-printing the abstraction |     78   Be careful if you pretty-print terms. Consider pretty-printing the abstraction | 
|     70   term shown above: |     79   term shown above: | 
|     71  |     80  | 
|     72   @{ML_matchresult_fake [display, gray] |     81   @{ML_response [display, gray] | 
|     73 \<open>@{term "\<lambda>x y. x y"} |     82 \<open>@{term "\<lambda>x y. x y"} | 
|     74   |> pretty_term @{context} |     83   |> pretty_term @{context} | 
|     75   |> pwriteln\<close> |     84   |> pwriteln\<close> | 
|     76   \<open>\<lambda>x. x\<close>} |     85   \<open>\<lambda>x. x\<close>} | 
|     77  |     86  | 
|     78   This is one common source for puzzlement in Isabelle, which has  |     87   This is one common source for puzzlement in Isabelle, which has  | 
|     79   tacitly eta-contracted the output. You obtain a similar result  |     88   tacitly eta-contracted the output. You obtain a similar result  | 
|     80   with beta-redexes |     89   with beta-redexes | 
|     81  |     90  | 
|     82   @{ML_matchresult_fake [display, gray] |     91   @{ML_response [display, gray] | 
|     83 \<open>let  |     92 \<open>let  | 
|     84   val redex = @{term "(\<lambda>(x::int) (y::int). x)"}  |     93   val redex = @{term "(\<lambda>(x::int) (y::int). x)"}  | 
|     85   val arg1 = @{term "1::int"}  |     94   val arg1 = @{term "1::int"}  | 
|     86   val arg2 = @{term "2::int"} |     95   val arg2 = @{term "2::int"} | 
|     87 in |     96 in | 
|    142   \end{readmore} |    151   \end{readmore} | 
|    143    |    152    | 
|    144   Constructing terms via antiquotations has the advantage that only typable |    153   Constructing terms via antiquotations has the advantage that only typable | 
|    145   terms can be constructed. For example |    154   terms can be constructed. For example | 
|    146  |    155  | 
|    147   @{ML_matchresult_fake_both [display,gray] |    156   @{ML_response [display,gray] | 
|    148   \<open>@{term "x x"}\<close> |    157   \<open>@{term "x x"}\<close> | 
|    149   \<open>Type unification failed: Occurs check!\<close>} |    158   \<open>Type unification failed: Occurs check!\<close>} | 
|    150  |    159  | 
|    151   raises a typing error, while it is perfectly ok to construct the term |    160   raises a typing error, while it is perfectly ok to construct the term | 
|    152   with the raw ML-constructors: |    161   with the raw ML-constructors: | 
|    153  |    162  | 
|    154   @{ML_matchresult_fake [display,gray]  |    163   @{ML_response [display,gray]  | 
|    155 \<open>let |    164 \<open>let | 
|    156   val omega = Free ("x", @{typ "nat \<Rightarrow> nat"}) $ Free ("x", @{typ nat}) |    165   val omega = Free ("x", @{typ "nat \<Rightarrow> nat"}) $ Free ("x", @{typ nat}) | 
|    157 in  |    166 in  | 
|    158   pwriteln (pretty_term @{context} omega) |    167   pwriteln (pretty_term @{context} omega) | 
|    159 end\<close> |    168 end\<close> | 
|    160   \<open>x x\<close>} |    169   "x x"} | 
|    161    |    170    | 
|    162   Sometimes the internal representation of terms can be surprisingly different |    171   Sometimes the internal representation of terms can be surprisingly different | 
|    163   from what you see at the user-level, because the layers of |    172   from what you see at the user-level, because the layers of | 
|    164   parsing/type-checking/pretty printing can be quite elaborate.  |    173   parsing/type-checking/pretty printing can be quite elaborate.  | 
|    165  |    174  | 
|    356   constructing terms. One is the function @{ML_ind list_comb in Term}, which |    365   constructing terms. One is the function @{ML_ind list_comb in Term}, which | 
|    357   takes as argument a term and a list of terms, and produces as output the |    366   takes as argument a term and a list of terms, and produces as output the | 
|    358   term list applied to the term. For example |    367   term list applied to the term. For example | 
|    359  |    368  | 
|    360  |    369  | 
|    361 @{ML_matchresult_fake [display,gray] |    370 @{ML_response [display,gray] | 
|    362 \<open>let |    371 \<open>let | 
|    363   val trm = @{term "P::bool \<Rightarrow> bool \<Rightarrow> bool"} |    372   val trm = @{term "P::bool \<Rightarrow> bool \<Rightarrow> bool"} | 
|    364   val args = [@{term "True"}, @{term "False"}] |    373   val args = [@{term "True"}, @{term "False"}] | 
|    365 in |    374 in | 
|    366   list_comb (trm, args) |    375   list_comb (trm, args) | 
|    367 end\<close> |    376 end\<close> | 
|    368 \<open>Free ("P", "bool \<Rightarrow> bool \<Rightarrow> bool")  |    377 \<open>Free ("P", "bool \<Rightarrow> bool \<Rightarrow> bool")  | 
|    369    $ Const ("True", "bool") $ Const ("False", "bool")\<close>} |    378    $ Const ("HOL.True", "bool") $ Const ("HOL.False", "bool")\<close>} | 
|    370  |    379  | 
|    371   Another handy function is @{ML_ind lambda in Term}, which abstracts a variable  |    380   Another handy function is @{ML_ind lambda in Term}, which abstracts a variable  | 
|    372   in a term. For example |    381   in a term. For example | 
|    373  |    382  | 
|    374   @{ML_matchresult_fake [display,gray] |    383   @{ML_response [display,gray] | 
|    375 \<open>let |    384 \<open>let | 
|    376   val x_nat = @{term "x::nat"} |    385   val x_nat = @{term "x::nat"} | 
|    377   val trm = @{term "(P::nat \<Rightarrow> bool) x"} |    386   val trm = @{term "(P::nat \<Rightarrow> bool) x"} | 
|    378 in |    387 in | 
|    379   lambda x_nat trm |    388   lambda x_nat trm | 
|    380 end\<close> |    389 end\<close> | 
|    381   \<open>Abs ("x", "Nat.nat", Free ("P", "bool \<Rightarrow> bool") $ Bound 0)\<close>} |    390   \<open>Abs ("x", "nat", Free ("P", "nat \<Rightarrow> bool") $ Bound 0)\<close>} | 
|    382  |    391  | 
|    383   In this example, @{ML lambda} produces a de Bruijn index (i.e.\ @{ML \<open>Bound 0\<close>}),  |    392   In this example, @{ML lambda} produces a de Bruijn index (i.e.\ @{ML \<open>Bound 0\<close>}),  | 
|    384   and an abstraction, where it also records the type of the abstracted |    393   and an abstraction, where it also records the type of the abstracted | 
|    385   variable and for printing purposes also its name.  Note that because of the |    394   variable and for printing purposes also its name.  Note that because of the | 
|    386   typing annotation on \<open>P\<close>, the variable \<open>x\<close> in \<open>P x\<close> |    395   typing annotation on \<open>P\<close>, the variable \<open>x\<close> in \<open>P x\<close> | 
|    387   is of the same type as the abstracted variable. If it is of different type, |    396   is of the same type as the abstracted variable. If it is of different type, | 
|    388   as in |    397   as in | 
|    389  |    398  | 
|    390   @{ML_matchresult_fake [display,gray] |    399   @{ML_response [display,gray] | 
|    391 \<open>let |    400 \<open>let | 
|    392   val x_int = @{term "x::int"} |    401   val x_int = @{term "x::int"} | 
|    393   val trm = @{term "(P::nat \<Rightarrow> bool) x"} |    402   val trm = @{term "(P::nat \<Rightarrow> bool) x"} | 
|    394 in |    403 in | 
|    395   lambda x_int trm |    404   lambda x_int trm | 
|    404   There is also the function @{ML_ind subst_free in Term} with which terms can be |    413   There is also the function @{ML_ind subst_free in Term} with which terms can be | 
|    405   replaced by other terms. For example below, we will replace in @{term |    414   replaced by other terms. For example below, we will replace in @{term | 
|    406   "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0 x"} the subterm @{term "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0"} by |    415   "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0 x"} the subterm @{term "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0"} by | 
|    407   @{term y}, and @{term x} by @{term True}. |    416   @{term y}, and @{term x} by @{term True}. | 
|    408  |    417  | 
|    409   @{ML_matchresult_fake [display,gray] |    418   @{ML_response [display,gray] | 
|    410 \<open>let |    419 \<open>let | 
|    411   val sub1 = (@{term "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0"}, @{term "y::nat \<Rightarrow> nat"}) |    420   val sub1 = (@{term "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0"}, @{term "y::nat \<Rightarrow> nat"}) | 
|    412   val sub2 = (@{term "x::nat"}, @{term "True"}) |    421   val sub2 = (@{term "x::nat"}, @{term "True"}) | 
|    413   val trm = @{term "((f::nat \<Rightarrow> nat \<Rightarrow> nat) 0) x"} |    422   val trm = @{term "((f::nat \<Rightarrow> nat \<Rightarrow> nat) 0) x"} | 
|    414 in |    423 in | 
|    415   subst_free [sub1, sub2] trm |    424   subst_free [sub1, sub2] trm | 
|    416 end\<close> |    425 end\<close> | 
|    417   \<open>Free ("y", "nat \<Rightarrow> nat") $ Const ("True", "bool")\<close>} |    426   \<open>Free ("y", "nat \<Rightarrow> nat") $ Const ("HOL.True", "bool")\<close>} | 
|    418  |    427  | 
|    419   As can be seen, @{ML subst_free} does not take typability into account. |    428   As can be seen, @{ML subst_free} does not take typability into account. | 
|    420   However it takes alpha-equivalence into account: |    429   However it takes alpha-equivalence into account: | 
|    421  |    430  | 
|    422   @{ML_matchresult_fake [display, gray] |    431   @{ML_response [display, gray] | 
|    423 \<open>let |    432 \<open>let | 
|    424   val sub = (@{term "(\<lambda>y::nat. y)"}, @{term "x::nat"}) |    433   val sub = (@{term "(\<lambda>y::nat. y)"}, @{term "x::nat"}) | 
|    425   val trm = @{term "(\<lambda>x::nat. x)"} |    434   val trm = @{term "(\<lambda>x::nat. x)"} | 
|    426 in |    435 in | 
|    427   subst_free [sub] trm |    436   subst_free [sub] trm | 
|    444  |    453  | 
|    445 text \<open> |    454 text \<open> | 
|    446   The function returns a pair consisting of the stripped off variables and |    455   The function returns a pair consisting of the stripped off variables and | 
|    447   the body of the universal quantification. For example |    456   the body of the universal quantification. For example | 
|    448  |    457  | 
|    449   @{ML_matchresult_fake [display, gray] |    458   @{ML_response [display, gray] | 
|    450   \<open>strip_alls @{term "\<forall>x y. x = (y::bool)"}\<close> |    459   \<open>strip_alls @{term "\<forall>x y. x = (y::bool)"}\<close> | 
|    451 \<open>([Free ("x", "bool"), Free ("y", "bool")], |    460 \<open>([Free ("x", "bool"), Free ("y", "bool")], | 
|    452   Const ("op =", _) $ Bound 1 $ Bound 0)\<close>} |    461   Const ("HOL.eq",\<dots>) $ Bound 1 $ Bound 0)\<close>} | 
|    453  |    462  | 
|    454   Note that we produced in the body two dangling de Bruijn indices.  |    463   Note that we produced in the body two dangling de Bruijn indices.  | 
|    455   Later on we will also use the inverse function that |    464   Later on we will also use the inverse function that | 
|    456   builds the quantification from a body and a list of (free) variables. |    465   builds the quantification from a body and a list of (free) variables. | 
|    457 \<close> |    466 \<close> | 
|    486   name is by no means unique. If clashes need to be avoided, then we should |    495   name is by no means unique. If clashes need to be avoided, then we should | 
|    487   use the function @{ML_ind dest_abs in Term}, which returns the body where |    496   use the function @{ML_ind dest_abs in Term}, which returns the body where | 
|    488   the loose de Bruijn index is replaced by a unique free variable. For example |    497   the loose de Bruijn index is replaced by a unique free variable. For example | 
|    489  |    498  | 
|    490  |    499  | 
|    491   @{ML_matchresult_fake [display,gray] |    500   @{ML_response [display,gray] | 
|    492   \<open>let |    501   \<open>let | 
|    493   val body = Bound 0 $ Free ("x", @{typ nat}) |    502   val body = Bound 0 $ Free ("x", @{typ nat}) | 
|    494 in |    503 in | 
|    495   Term.dest_abs ("x", @{typ "nat \<Rightarrow> bool"}, body) |    504   Term.dest_abs ("x", @{typ "nat \<Rightarrow> bool"}, body) | 
|    496 end\<close> |    505 end\<close> | 
|    497   \<open>("xa", Free ("xa", "Nat.nat \<Rightarrow> bool") $ Free ("x", "Nat.nat"))\<close>} |    506   \<open>("xa", Free ("xa", "nat \<Rightarrow> bool") $ Free ("x", "nat"))\<close>} | 
|    498  |    507  | 
|    499   Sometimes it is necessary to manipulate de Bruijn indices in terms directly. |    508   Sometimes it is necessary to manipulate de Bruijn indices in terms directly. | 
|    500   There are many functions to do this. We describe only two. The first, |    509   There are many functions to do this. We describe only two. The first, | 
|    501   @{ML_ind incr_boundvars in Term}, increases by an integer the indices  |    510   @{ML_ind incr_boundvars in Term}, increases by an integer the indices  | 
|    502   of the loose bound variables in a term. In the code below |    511   of the loose bound variables in a term. In the code below | 
|    503  |    512  | 
|    504 @{ML_matchresult_fake [display,gray] |    513 @{ML_response [display,gray] | 
|    505 \<open>@{term "\<forall>x y z u. z = u"} |    514 \<open>@{term "\<forall>x y z u. z = u"} | 
|    506   |> strip_alls |    515   |> strip_alls | 
|    507   ||> incr_boundvars 2 |    516   ||> incr_boundvars 2 | 
|    508   |> build_alls |    517   |> build_alls | 
|    509   |> pretty_term @{context} |    518   |> pretty_term @{context} | 
|    796   below. In case of failure, @{ML typ_unify in Sign} will raise the exception |    805   below. In case of failure, @{ML typ_unify in Sign} will raise the exception | 
|    797   \<open>TUNIFY\<close>.  We can print out the resulting type environment bound to  |    806   \<open>TUNIFY\<close>.  We can print out the resulting type environment bound to  | 
|    798   @{ML tyenv_unif} with the built-in function @{ML_ind dest in Vartab} from the |    807   @{ML tyenv_unif} with the built-in function @{ML_ind dest in Vartab} from the | 
|    799   structure @{ML_structure Vartab}. |    808   structure @{ML_structure Vartab}. | 
|    800  |    809  | 
|    801   @{ML_matchresult_fake [display,gray] |    810   @{ML_response [display,gray] | 
|    802   \<open>Vartab.dest tyenv_unif\<close> |    811   \<open>Vartab.dest tyenv_unif\<close> | 
|    803   \<open>[(("'a", 0), (["HOL.type"], "?'b List.list")),  |    812   \<open>[(("'a", 0), (["HOL.type"], "?'b list")),  | 
|    804  (("'b", 0), (["HOL.type"], "nat"))]\<close>}  |    813  (("'b", 0), (["HOL.type"], "nat"))]\<close>}  | 
|    805 \<close> |    814 \<close> | 
|    806  |    815  | 
|    807 text_raw \<open> |    816 text_raw \<open> | 
|    808   \begin{figure}[t] |    817   \begin{figure}[t] | 
|    875   val ty2 = @{typ_pat "?'b::s2"} |    884   val ty2 = @{typ_pat "?'b::s2"} | 
|    876 in |    885 in | 
|    877   Sign.typ_unify @{theory} (ty1, ty2) (Vartab.empty, 0)  |    886   Sign.typ_unify @{theory} (ty1, ty2) (Vartab.empty, 0)  | 
|    878 end\<close> |    887 end\<close> | 
|    879  |    888  | 
|         |    889 declare[[show_sorts]] | 
|         |    890  | 
|    880 text \<open> |    891 text \<open> | 
|    881   To print out the result type environment we switch on the printing  |    892   To print out the result type environment we switch on the printing  | 
|    882   of sort information by setting @{ML_ind show_sorts in Syntax} to  |    893   of sort information by setting @{ML_ind show_sorts in Syntax} to  | 
|    883   true. This allows us to inspect the typing environment. |    894   true. This allows us to inspect the typing environment. | 
|    884  |    895  | 
|    885   @{ML_matchresult_fake [display,gray] |    896   @{ML_response [display,gray] | 
|    886   \<open>pretty_tyenv @{context} tyenv\<close> |    897   \<open>pretty_tyenv @{context} tyenv\<close> | 
|    887   \<open>[?'a::s1 := ?'a1::{s1, s2}, ?'b::s2 := ?'a1::{s1, s2}]\<close>} |    898   \<open>[?'a::s1 := ?'a1::{s1, s2}, ?'b::s2 := ?'a1::{s1, s2}]\<close>} | 
|    888  |    899 \<close> | 
|         |    900  | 
|         |    901 declare[[show_sorts=false]] | 
|         |    902  | 
|         |    903 text\<open> | 
|    889   As can be seen, the type variables \<open>?'a\<close> and \<open>?'b\<close> are instantiated |    904   As can be seen, the type variables \<open>?'a\<close> and \<open>?'b\<close> are instantiated | 
|    890   with a new type variable \<open>?'a1\<close> with sort \<open>{s1, s2}\<close>. Since a new |    905   with a new type variable \<open>?'a1\<close> with sort \<open>{s1, s2}\<close>. Since a new | 
|    891   type variable has been introduced the @{ML index}, originally being \<open>0\<close>,  |    906   type variable has been introduced the @{ML index}, originally being \<open>0\<close>,  | 
|    892   has been increased to \<open>1\<close>. |    907   has been increased to \<open>1\<close>. | 
|    893  |    908  | 
|    909   called an instantiation in \emph{triangular form}. These triangular  |    924   called an instantiation in \emph{triangular form}. These triangular  | 
|    910   instantiations, or triangular type environments, are used because of  |    925   instantiations, or triangular type environments, are used because of  | 
|    911   performance reasons. To apply such a type environment to a type, say \<open>?'a * |    926   performance reasons. To apply such a type environment to a type, say \<open>?'a * | 
|    912   ?'b\<close>, you should use the function @{ML_ind norm_type in Envir}: |    927   ?'b\<close>, you should use the function @{ML_ind norm_type in Envir}: | 
|    913  |    928  | 
|    914   @{ML_matchresult_fake [display, gray] |    929   @{ML_response [display, gray] | 
|    915   \<open>Envir.norm_type tyenv_unif @{typ_pat "?'a * ?'b"}\<close> |    930   \<open>Envir.norm_type tyenv_unif @{typ_pat "?'a * ?'b"}\<close> | 
|    916   \<open>nat list * nat\<close>} |    931   \<open>nat list \<times> nat\<close>} | 
|    917  |    932  | 
|    918   Matching of types can be done with the function @{ML_ind typ_match in Sign} |    933   Matching of types can be done with the function @{ML_ind typ_match in Sign} | 
|    919   also from the structure @{ML_structure Sign}. This function returns a @{ML_type |    934   also from the structure @{ML_structure Sign}. This function returns a @{ML_type | 
|    920   Type.tyenv} as well, but might raise the exception \<open>TYPE_MATCH\<close> in case |    935   Type.tyenv} as well, but might raise the exception \<open>TYPE_MATCH\<close> in case | 
|    921   of failure. For example |    936   of failure. For example | 
|    929 end\<close> |    944 end\<close> | 
|    930  |    945  | 
|    931 text \<open> |    946 text \<open> | 
|    932   Printing out the calculated matcher gives |    947   Printing out the calculated matcher gives | 
|    933  |    948  | 
|    934   @{ML_matchresult_fake [display,gray] |    949   @{ML_response [display,gray] | 
|    935   \<open>pretty_tyenv @{context} tyenv_match\<close> |    950   \<open>pretty_tyenv @{context} tyenv_match\<close> | 
|    936   \<open>[?'a := bool list, ?'b := nat]\<close>} |    951   \<open>[?'a := bool list, ?'b := nat]\<close>} | 
|    937    |    952    | 
|    938   Unlike unification, which uses the function @{ML norm_type in Envir},  |    953   Unlike unification, which uses the function @{ML norm_type in Envir},  | 
|    939   applying the matcher to a type needs to be done with the function |    954   applying the matcher to a type needs to be done with the function | 
|    940   @{ML_ind subst_type in Envir}. For example |    955   @{ML_ind subst_type in Envir}. For example | 
|    941  |    956  | 
|    942   @{ML_matchresult_fake [display, gray] |    957   @{ML_response [display, gray] | 
|    943   \<open>Envir.subst_type tyenv_match @{typ_pat "?'a * ?'b"}\<close> |    958   \<open>Envir.subst_type tyenv_match @{typ_pat "?'a * ?'b"}\<close> | 
|    944   \<open>bool list * nat\<close>} |    959   \<open>bool list \<times> nat\<close>} | 
|    945  |    960  | 
|    946   Be careful to observe the difference: always use |    961   Be careful to observe the difference: always use | 
|    947   @{ML subst_type in Envir} for matchers and @{ML norm_type in Envir}  |    962   @{ML subst_type in Envir} for matchers and @{ML norm_type in Envir}  | 
|    948   for unifiers. To show the difference, let us calculate the  |    963   for unifiers. To show the difference, let us calculate the  | 
|    949   following matcher: |    964   following matcher: | 
|    958  |    973  | 
|    959 text \<open> |    974 text \<open> | 
|    960   Now @{ML tyenv_unif} is equal to @{ML tyenv_match'}. If we apply  |    975   Now @{ML tyenv_unif} is equal to @{ML tyenv_match'}. If we apply  | 
|    961   @{ML norm_type in Envir} to the type \<open>?'a * ?'b\<close> we obtain |    976   @{ML norm_type in Envir} to the type \<open>?'a * ?'b\<close> we obtain | 
|    962  |    977  | 
|    963   @{ML_matchresult_fake [display, gray] |    978   @{ML_response [display, gray] | 
|    964   \<open>Envir.norm_type tyenv_match' @{typ_pat "?'a * ?'b"}\<close> |    979   \<open>Envir.norm_type tyenv_match' @{typ_pat "?'a * ?'b"}\<close> | 
|    965   \<open>nat list * nat\<close>} |    980   \<open>nat list \<times> nat\<close>} | 
|    966  |    981  | 
|    967   which does not solve the matching problem, and if |    982   which does not solve the matching problem, and if | 
|    968   we apply @{ML subst_type in Envir} to the same type we obtain |    983   we apply @{ML subst_type in Envir} to the same type we obtain | 
|    969  |    984  | 
|    970   @{ML_matchresult_fake [display, gray] |    985   @{ML_response [display, gray] | 
|    971   \<open>Envir.subst_type tyenv_unif @{typ_pat "?'a * ?'b"}\<close> |    986   \<open>Envir.subst_type tyenv_unif @{typ_pat "?'a * ?'b"}\<close> | 
|    972   \<open>?'b list * nat\<close>} |    987   \<open>?'b list \<times> nat\<close>} | 
|    973    |    988    | 
|    974   which does not solve the unification problem. |    989   which does not solve the unification problem. | 
|    975  |    990  | 
|    976   \begin{readmore} |    991   \begin{readmore} | 
|    977   Unification and matching for types is implemented |    992   Unification and matching for types is implemented | 
|   1021 text \<open> |   1036 text \<open> | 
|   1022   In this example we annotated types explicitly because then  |   1037   In this example we annotated types explicitly because then  | 
|   1023   the type environment is empty and can be ignored. The  |   1038   the type environment is empty and can be ignored. The  | 
|   1024   resulting term environment is |   1039   resulting term environment is | 
|   1025  |   1040  | 
|   1026   @{ML_matchresult_fake [display, gray] |   1041   @{ML_response [display, gray] | 
|   1027   \<open>pretty_env @{context} fo_env\<close> |   1042   \<open>pretty_env @{context} fo_env\<close> | 
|   1028   \<open>[?X := P, ?Y := \<lambda>a b. Q b a]\<close>} |   1043   \<open>[?X := P, ?Y := \<lambda>a b. Q b a]\<close>} | 
|   1029  |   1044  | 
|   1030   The matcher can be applied to a term using the function @{ML_ind subst_term |   1045   The matcher can be applied to a term using the function @{ML_ind subst_term | 
|   1031   in Envir} (remember the same convention for types applies to terms: @{ML |   1046   in Envir} (remember the same convention for types applies to terms: @{ML | 
|   1032   subst_term in Envir} is for matchers and @{ML norm_term in Envir} for |   1047   subst_term in Envir} is for matchers and @{ML norm_term in Envir} for | 
|   1033   unifiers). The function @{ML subst_term in Envir} expects a type environment, |   1048   unifiers). The function @{ML subst_term in Envir} expects a type environment, | 
|   1034   which is set to empty in the example below, and a term environment. |   1049   which is set to empty in the example below, and a term environment. | 
|   1035  |   1050  | 
|   1036   @{ML_matchresult_fake [display, gray] |   1051   @{ML_response [display, gray] | 
|   1037   \<open>let |   1052   \<open>let | 
|   1038   val trm = @{term_pat "(?X::(nat\<Rightarrow>nat\<Rightarrow>nat)\<Rightarrow>bool) ?Y"} |   1053   val trm = @{term_pat "(?X::(nat\<Rightarrow>nat\<Rightarrow>nat)\<Rightarrow>bool) ?Y"} | 
|   1039 in |   1054 in | 
|   1040   Envir.subst_term (Vartab.empty, fo_env) trm |   1055   Envir.subst_term (Vartab.empty, fo_env) trm | 
|   1041   |> pretty_term @{context} |   1056   |> pretty_term @{context} | 
|   1150 text \<open> |   1165 text \<open> | 
|   1151   \footnote{\bf FIXME: what is the list of term pairs in the unifier: flex-flex pairs?} |   1166   \footnote{\bf FIXME: what is the list of term pairs in the unifier: flex-flex pairs?} | 
|   1152  |   1167  | 
|   1153   We can print them out as follows. |   1168   We can print them out as follows. | 
|   1154  |   1169  | 
|   1155   @{ML_matchresult_fake [display, gray] |   1170   @{ML_response [display, gray] | 
|   1156   \<open>pretty_env @{context} (Envir.term_env un1); |   1171   \<open>pretty_env @{context} (Envir.term_env un1)\<close> | 
|   1157 pretty_env @{context} (Envir.term_env un2); |   1172   \<open>[?X := \<lambda>a. a, ?Y := f a]\<close>} | 
|   1158 pretty_env @{context} (Envir.term_env un3)\<close> |   1173   @{ML_response [display, gray] | 
|   1159   \<open>[?X := \<lambda>a. a, ?Y := f a] |   1174   \<open>pretty_env @{context} (Envir.term_env un2)\<close> | 
|   1160 [?X := f, ?Y := a] |   1175   \<open>[?X := f, ?Y := a]\<close>} | 
|   1161 [?X := \<lambda>b. f a]\<close>} |   1176   @{ML_response [display, gray] | 
|         |   1177   \<open>pretty_env @{context} (Envir.term_env un3)\<close> | 
|         |   1178   \<open>[?X := \<lambda>b. f a]\<close>} | 
|   1162  |   1179  | 
|   1163  |   1180  | 
|   1164   In case of failure the function @{ML_ind unifiers in Unify} does not raise |   1181   In case of failure the function @{ML_ind unifiers in Unify} does not raise | 
|   1165   an exception, rather returns the empty sequence. For example |   1182   an exception, rather returns the empty sequence. For example | 
|   1166  |   1183  | 
|   1385   information is redundant. A short-cut is to use the ``place-holder''  |   1402   information is redundant. A short-cut is to use the ``place-holder''  | 
|   1386   type @{ML_ind dummyT in Term} and then let type-inference figure out the  |   1403   type @{ML_ind dummyT in Term} and then let type-inference figure out the  | 
|   1387   complete type. The type inference can be invoked with the function |   1404   complete type. The type inference can be invoked with the function | 
|   1388   @{ML_ind check_term in Syntax}. An example is as follows: |   1405   @{ML_ind check_term in Syntax}. An example is as follows: | 
|   1389  |   1406  | 
|   1390   @{ML_matchresult_fake [display,gray]  |   1407   @{ML_response [display,gray]  | 
|   1391 \<open>let |   1408 \<open>let | 
|   1392   val c = Const (@{const_name "plus"}, dummyT)  |   1409   val c = Const (@{const_name "plus"}, dummyT)  | 
|   1393   val o = @{term "1::nat"}  |   1410   val o = @{term "1::nat"}  | 
|   1394   val v = Free ("x", dummyT) |   1411   val v = Free ("x", dummyT) | 
|   1395 in    |   1412 in    | 
|   1396   Syntax.check_term @{context} (c $ o $ v) |   1413   Syntax.check_term @{context} (c $ o $ v) | 
|   1397 end\<close> |   1414 end\<close> | 
|   1398 \<open>Const ("HOL.plus_class.plus", "nat \<Rightarrow> nat \<Rightarrow> nat") $ |   1415 \<open>Const ("Groups.plus_class.plus", "nat \<Rightarrow> nat \<Rightarrow> nat") $ | 
|   1399   Const ("HOL.one_class.one", "nat") $ Free ("x", "nat")\<close>} |   1416   Const ("Groups.one_class.one", "nat") $ Free ("x", "nat")\<close>} | 
|   1400  |   1417  | 
|   1401   Instead of giving explicitly the type for the constant \<open>plus\<close> and the free  |   1418   Instead of giving explicitly the type for the constant \<open>plus\<close> and the free  | 
|   1402   variable \<open>x\<close>, type-inference fills in the missing information. |   1419   variable \<open>x\<close>, type-inference fills in the missing information. | 
|   1403  |   1420  | 
|   1404   \begin{readmore} |   1421   \begin{readmore} | 
|   1430   be constructed via ``official interfaces''. |   1447   be constructed via ``official interfaces''. | 
|   1431  |   1448  | 
|   1432   Certification is always relative to a context. For example you can  |   1449   Certification is always relative to a context. For example you can  | 
|   1433   write: |   1450   write: | 
|   1434  |   1451  | 
|   1435   @{ML_matchresult_fake [display,gray]  |   1452   @{ML_response [display,gray]  | 
|   1436   \<open>Thm.cterm_of @{context} @{term "(a::nat) + b = c"}\<close>  |   1453   \<open>Thm.cterm_of @{context} @{term "(a::nat) + b = c"}\<close>  | 
|   1437   \<open>a + b = c\<close>} |   1454   \<open>a + b = c\<close>} | 
|   1438  |   1455  | 
|   1439   This can also be written with an antiquotation: |   1456   This can also be written with an antiquotation: | 
|   1440  |   1457  | 
|   1441   @{ML_matchresult_fake [display,gray]  |   1458   @{ML_response [display,gray]  | 
|   1442   \<open>@{cterm "(a::nat) + b = c"}\<close>  |   1459   \<open>@{cterm "(a::nat) + b = c"}\<close>  | 
|   1443   \<open>a + b = c\<close>} |   1460   \<open>a + b = c\<close>} | 
|   1444  |   1461  | 
|   1445   Attempting to obtain the certified term for |   1462   Attempting to obtain the certified term for | 
|   1446  |   1463  | 
|   1447   @{ML_matchresult_fake_both [display,gray]  |   1464   @{ML_response [display,gray]  | 
|   1448   \<open>@{cterm "1 + True"}\<close>  |   1465   \<open>@{cterm "1 + True"}\<close>  | 
|   1449   \<open>Type unification failed \<dots>\<close>} |   1466   \<open>Type unification failed\<dots>\<close>} | 
|   1450  |   1467  | 
|   1451   yields an error (since the term is not typable). A slightly more elaborate |   1468   yields an error (since the term is not typable). A slightly more elaborate | 
|   1452   example that type-checks is: |   1469   example that type-checks is: | 
|   1453  |   1470  | 
|   1454 @{ML_matchresult_fake [display,gray]  |   1471 @{ML_response [display,gray]  | 
|   1455 \<open>let |   1472 \<open>let | 
|   1456   val natT = @{typ "nat"} |   1473   val natT = @{typ "nat"} | 
|   1457   val zero = @{term "0::nat"} |   1474   val zero = @{term "0::nat"} | 
|   1458   val plus = Const (@{const_name plus}, [natT, natT] ---> natT) |   1475   val plus = Const (@{const_name plus}, [natT, natT] ---> natT) | 
|   1459 in |   1476 in | 
|   1463  |   1480  | 
|   1464   In Isabelle not just terms need to be certified, but also types. For example,  |   1481   In Isabelle not just terms need to be certified, but also types. For example,  | 
|   1465   you obtain the certified type for the Isabelle type @{typ "nat \<Rightarrow> bool"} on  |   1482   you obtain the certified type for the Isabelle type @{typ "nat \<Rightarrow> bool"} on  | 
|   1466   the ML-level as follows: |   1483   the ML-level as follows: | 
|   1467  |   1484  | 
|   1468   @{ML_matchresult_fake [display,gray] |   1485   @{ML_response [display,gray] | 
|   1469   \<open>Thm.ctyp_of @{context} (@{typ nat} --> @{typ bool})\<close> |   1486   \<open>Thm.ctyp_of @{context} (@{typ nat} --> @{typ bool})\<close> | 
|   1470   \<open>nat \<Rightarrow> bool\<close>} |   1487   \<open>nat \<Rightarrow> bool\<close>} | 
|   1471  |   1488  | 
|   1472   or with the antiquotation: |   1489   or with the antiquotation: | 
|   1473  |   1490  | 
|   1474   @{ML_matchresult_fake [display,gray] |   1491   @{ML_response [display,gray] | 
|   1475   \<open>@{ctyp "nat \<Rightarrow> bool"}\<close> |   1492   \<open>@{ctyp "nat \<Rightarrow> bool"}\<close> | 
|   1476   \<open>nat \<Rightarrow> bool\<close>} |   1493   \<open>nat \<Rightarrow> bool\<close>} | 
|   1477  |   1494  | 
|   1478   Since certified terms are, unlike terms, abstract objects, we cannot |   1495   Since certified terms are, unlike terms, abstract objects, we cannot | 
|   1479   pattern-match against them. However, we can construct them. For example |   1496   pattern-match against them. However, we can construct them. For example | 
|   1480   the function @{ML_ind apply in Thm} produces a certified application. |   1497   the function @{ML_ind apply in Thm} produces a certified application. | 
|   1481  |   1498  | 
|   1482   @{ML_matchresult_fake [display,gray] |   1499   @{ML_response [display,gray] | 
|   1483   \<open>Thm.apply @{cterm "P::nat \<Rightarrow> bool"} @{cterm "3::nat"}\<close> |   1500   \<open>Thm.apply @{cterm "P::nat \<Rightarrow> bool"} @{cterm "3::nat"}\<close> | 
|   1484   \<open>P 3\<close>} |   1501   \<open>P 3\<close>} | 
|   1485  |   1502  | 
|   1486   Similarly the function @{ML_ind list_comb in Drule} from the structure @{ML_structure Drule} |   1503   Similarly the function @{ML_ind list_comb in Drule} from the structure @{ML_structure Drule} | 
|   1487   applies a list of @{ML_type cterm}s. |   1504   applies a list of @{ML_type cterm}s. | 
|   1488  |   1505  | 
|   1489   @{ML_matchresult_fake [display,gray] |   1506   @{ML_response [display,gray] | 
|   1490   \<open>let |   1507   \<open>let | 
|   1491   val chead = @{cterm "P::unit \<Rightarrow> nat \<Rightarrow> bool"} |   1508   val chead = @{cterm "P::unit \<Rightarrow> nat \<Rightarrow> bool"} | 
|   1492   val cargs = [@{cterm "()"}, @{cterm "3::nat"}] |   1509   val cargs = [@{cterm "()"}, @{cterm "3::nat"}] | 
|   1493 in |   1510 in | 
|   1494   Drule.list_comb (chead, cargs) |   1511   Drule.list_comb (chead, cargs) | 
|   1723  |   1740  | 
|   1724   Many functions destruct theorems into @{ML_type cterm}s. For example |   1741   Many functions destruct theorems into @{ML_type cterm}s. For example | 
|   1725   the functions @{ML_ind lhs_of in Thm} and @{ML_ind rhs_of in Thm} return  |   1742   the functions @{ML_ind lhs_of in Thm} and @{ML_ind rhs_of in Thm} return  | 
|   1726   the left and right-hand side, respectively, of a meta-equality. |   1743   the left and right-hand side, respectively, of a meta-equality. | 
|   1727  |   1744  | 
|   1728   @{ML_matchresult_fake [display,gray] |   1745   @{ML_response [display,gray] | 
|   1729   \<open>let |   1746   \<open>let | 
|   1730   val eq = @{thm True_def} |   1747   val eq = @{thm True_def} | 
|   1731 in |   1748 in | 
|   1732   (Thm.lhs_of eq, Thm.rhs_of eq)  |   1749   (Thm.lhs_of eq, Thm.rhs_of eq)  | 
|   1733   |> apply2 (Pretty.string_of o (pretty_cterm @{context})) |   1750   |> apply2 (YXML.content_of o Pretty.string_of o (pretty_cterm @{context})) | 
|   1734 end\<close> |   1751 end\<close> | 
|   1735   \<open>(True, (\<lambda>x. x) = (\<lambda>x. x))\<close>} |   1752   \<open>("True", "(\<lambda>x. x) = (\<lambda>x. x)")\<close>} | 
|   1736  |   1753  | 
|   1737   Other function produce terms that can be pattern-matched.  |   1754   Other function produce terms that can be pattern-matched.  | 
|   1738   Suppose the following two theorems. |   1755   Suppose the following two theorems. | 
|   1739 \<close> |   1756 \<close> | 
|   1740  |   1757  | 
|   1776   Section~\ref{sec:simplifier}, the simplifier  |   1793   Section~\ref{sec:simplifier}, the simplifier  | 
|   1777   can be used to work directly over theorems, for example to unfold definitions. To show |   1794   can be used to work directly over theorems, for example to unfold definitions. To show | 
|   1778   this, we build the theorem @{term "True \<equiv> True"} (Line 1) and then  |   1795   this, we build the theorem @{term "True \<equiv> True"} (Line 1) and then  | 
|   1779   unfold the constant @{term "True"} according to its definition (Line 2). |   1796   unfold the constant @{term "True"} according to its definition (Line 2). | 
|   1780  |   1797  | 
|   1781   @{ML_matchresult_fake [display,gray,linenos] |   1798   @{ML_response [display,gray,linenos] | 
|   1782   \<open>Thm.reflexive @{cterm "True"} |   1799   \<open>Thm.reflexive @{cterm "True"} | 
|   1783   |> Simplifier.rewrite_rule @{context} [@{thm True_def}] |   1800   |> Simplifier.rewrite_rule @{context} [@{thm True_def}] | 
|   1784   |> pretty_thm @{context} |   1801   |> pretty_thm @{context} | 
|   1785   |> pwriteln\<close> |   1802   |> pwriteln\<close> | 
|   1786   \<open>(\<lambda>x. x) = (\<lambda>x. x) \<equiv> (\<lambda>x. x) = (\<lambda>x. x)\<close>} |   1803   \<open>(\<lambda>x. x) = (\<lambda>x. x) \<equiv> (\<lambda>x. x) = (\<lambda>x. x)\<close>} | 
|   1793   from the meta logic are more convenient for reasoning. Therefore there are |   1810   from the meta logic are more convenient for reasoning. Therefore there are | 
|   1794   some build in functions which help with these transformations. The function  |   1811   some build in functions which help with these transformations. The function  | 
|   1795   @{ML_ind rulify in Object_Logic}  |   1812   @{ML_ind rulify in Object_Logic}  | 
|   1796   replaces all object connectives by equivalents in the meta logic. For example |   1813   replaces all object connectives by equivalents in the meta logic. For example | 
|   1797  |   1814  | 
|   1798   @{ML_matchresult_fake [display, gray] |   1815   @{ML_response [display, gray] | 
|   1799   \<open>Object_Logic.rulify @{context} @{thm foo_test2}\<close> |   1816   \<open>Object_Logic.rulify @{context} @{thm foo_test2}\<close> | 
|   1800   \<open>\<lbrakk>?A; ?B\<rbrakk> \<Longrightarrow> ?C\<close>} |   1817   \<open>\<lbrakk>?A; ?B\<rbrakk> \<Longrightarrow> ?C\<close>} | 
|   1801  |   1818  | 
|   1802   The transformation in the other direction can be achieved with function |   1819   The transformation in the other direction can be achieved with function | 
|   1803   @{ML_ind atomize in Object_Logic} and the following code. |   1820   @{ML_ind atomize in Object_Logic} and the following code. | 
|   1804  |   1821  | 
|   1805   @{ML_matchresult_fake [display,gray] |   1822   @{ML_response [display,gray] | 
|   1806   \<open>let |   1823   \<open>let | 
|   1807   val thm = @{thm foo_test1} |   1824   val thm = @{thm foo_test1} | 
|   1808   val meta_eq = Object_Logic.atomize @{context} (Thm.cprop_of thm) |   1825   val meta_eq = Object_Logic.atomize @{context} (Thm.cprop_of thm) | 
|   1809 in |   1826 in | 
|   1810   Raw_Simplifier.rewrite_rule @{context} [meta_eq] thm |   1827   Raw_Simplifier.rewrite_rule @{context} [meta_eq] thm | 
|   1835 end\<close> |   1852 end\<close> | 
|   1836  |   1853  | 
|   1837 text \<open> |   1854 text \<open> | 
|   1838   This function produces for the theorem @{thm [source] list.induct} |   1855   This function produces for the theorem @{thm [source] list.induct} | 
|   1839  |   1856  | 
|   1840   @{ML_matchresult_fake [display, gray] |   1857   @{ML_response [display, gray] | 
|   1841   \<open>atomize_thm @{context} @{thm list.induct}\<close> |   1858   \<open>atomize_thm @{context} @{thm list.induct}\<close> | 
|   1842   \<open>\<forall>P list. P [] \<longrightarrow> (\<forall>a list. P list \<longrightarrow> P (a # list)) \<longrightarrow> P list\<close>} |   1859   \<open>"\<forall>P list. P [] \<longrightarrow> (\<forall>x1 x2. P x2 \<longrightarrow> P (x1 # x2)) \<longrightarrow> P list"\<close>} | 
|   1843  |   1860  | 
|   1844   Theorems can also be produced from terms by giving an explicit proof.  |   1861   Theorems can also be produced from terms by giving an explicit proof.  | 
|   1845   One way to achieve this is by using the function @{ML_ind prove in Goal} |   1862   One way to achieve this is by using the function @{ML_ind prove in Goal} | 
|   1846   in the structure @{ML_structure Goal}. For example below we use this function |   1863   in the structure @{ML_structure Goal}. For example below we use this function | 
|   1847   to prove the term @{term "P \<Longrightarrow> P"}. |   1864   to prove the term @{term "P \<Longrightarrow> P"}. | 
|   1848    |   1865    | 
|   1849   @{ML_matchresult_fake [display,gray] |   1866   @{ML_response [display,gray] | 
|   1850   \<open>let |   1867   \<open>let | 
|   1851   val trm = @{term "P \<Longrightarrow> P::bool"} |   1868   val trm = @{term "P \<Longrightarrow> P::bool"} | 
|   1852   val tac = K (assume_tac @{context} 1) |   1869   val tac = K (assume_tac @{context} 1) | 
|   1853 in |   1870 in | 
|   1854   Goal.prove @{context} ["P"] [] trm tac |   1871   Goal.prove @{context} ["P"] [] trm tac | 
|   1866   using the function @{ML_ind prove_common in Goal}. For this let us define the |   1883   using the function @{ML_ind prove_common in Goal}. For this let us define the | 
|   1867   following three helper functions. |   1884   following three helper functions. | 
|   1868 \<close> |   1885 \<close> | 
|   1869  |   1886  | 
|   1870 ML %grayML\<open>fun rep_goals i = replicate i @{prop "f x = f x"} |   1887 ML %grayML\<open>fun rep_goals i = replicate i @{prop "f x = f x"} | 
|   1871 fun rep_tacs i = replicate i (resolve_tac @{context} [@{thm refl}]) |   1888 fun rep_tacs ctxt i = replicate i (resolve_tac ctxt [@{thm refl}]) | 
|   1872  |   1889  | 
|   1873 fun multi_test ctxt i = |   1890 fun multi_test ctxt i = | 
|   1874   Goal.prove_common ctxt NONE ["f", "x"] [] (rep_goals i)  |   1891   Goal.prove_common ctxt NONE ["f", "x"] [] (rep_goals i)  | 
|   1875     (K ((Goal.conjunction_tac THEN' RANGE (rep_tacs i)) 1))\<close> |   1892     (K ((Goal.conjunction_tac THEN' RANGE (rep_tacs ctxt i)) 1))\<close> | 
|   1876  |   1893  | 
|   1877 text \<open> |   1894 text \<open> | 
|   1878   With them we can now produce three theorem instances of the  |   1895   With them we can now produce three theorem instances of the  | 
|   1879   proposition. |   1896   proposition. | 
|   1880  |   1897  | 
|   1881   @{ML_matchresult_fake [display, gray] |   1898   @{ML_response [display, gray] | 
|   1882   \<open>multi_test @{context} 3\<close> |   1899   \<open>multi_test @{context} 3\<close> | 
|   1883   \<open>["?f ?x = ?f ?x", "?f ?x = ?f ?x", "?f ?x = ?f ?x"]\<close>} |   1900   \<open>["?f ?x = ?f ?x", "?f ?x = ?f ?x", "?f ?x = ?f ?x"]\<close>} | 
|   1884  |   1901  | 
|   1885   However you should be careful with @{ML prove_common in Goal} and very |   1902   However you should be careful with @{ML prove_common in Goal} and very | 
|   1886   large goals. If you increase the counter in the code above to \<open>3000\<close>,  |   1903   large goals. If you increase the counter in the code above to \<open>3000\<close>,  | 
|   1917   kind, the names for cases and so on. This data is stored in a string-string |   1934   kind, the names for cases and so on. This data is stored in a string-string | 
|   1918   list and can be retrieved with the function @{ML_ind get_tags in |   1935   list and can be retrieved with the function @{ML_ind get_tags in | 
|   1919   Thm}. Assume you prove the following lemma. |   1936   Thm}. Assume you prove the following lemma. | 
|   1920 \<close> |   1937 \<close> | 
|   1921  |   1938  | 
|   1922 lemma foo_data:  |   1939 theorem foo_data:  | 
|   1923   shows "P \<Longrightarrow> P \<Longrightarrow> P" by assumption |   1940   shows "P \<Longrightarrow> P \<Longrightarrow> P" by assumption | 
|   1924  |   1941  | 
|   1925 text \<open> |   1942 text \<open> | 
|   1926   The auxiliary data of this lemma can be retrieved using the function  |   1943   The auxiliary data of this lemma can be retrieved using the function  | 
|   1927   @{ML_ind get_tags in Thm}. So far the the auxiliary data of this lemma is  |   1944   @{ML_ind get_tags in Thm}. So far the the auxiliary data of this lemma is  | 
|   1928  |   1945  | 
|   1929   @{ML_matchresult_fake [display,gray] |   1946   @{ML_matchresult [display,gray] | 
|   1930   \<open>Thm.get_tags @{thm foo_data}\<close> |   1947   \<open>Thm.get_tags @{thm foo_data}\<close> | 
|   1931   \<open>[("name", "General.foo_data"), ("kind", "lemma")]\<close>} |   1948   \<open>[("name", "Essential.foo_data"), ("kind", "theorem")]\<close>} | 
|   1932  |   1949  | 
|   1933   consisting of a name and a kind.  When we store lemmas in the theorem database,  |   1950   consisting of a name and a kind.  When we store lemmas in the theorem database,  | 
|   1934   we might want to explicitly extend this data by attaching case names to the  |   1951   we might want to explicitly extend this data by attaching case names to the  | 
|   1935   two premises of the lemma.  This can be done with the function @{ML_ind name in Rule_Cases} |   1952   two premises of the lemma.  This can be done with the function @{ML_ind name in Rule_Cases} | 
|   1936   from the structure @{ML_structure Rule_Cases}. |   1953   from the structure @{ML_structure Rule_Cases}. | 
|   2024 lemma my_conjIc: |   2051 lemma my_conjIc: | 
|   2025 shows "A \<and> B \<Longrightarrow> A \<and> B" |   2052 shows "A \<and> B \<Longrightarrow> A \<and> B" | 
|   2026 apply(auto) |   2053 apply(auto) | 
|   2027 done |   2054 done | 
|   2028  |   2055  | 
|   2029  |   2056 ML "Proofterm.proofs" | 
|         |   2057 ML \<open>@{thm my_conjIa} | 
|         |   2058   |> get_all_names\<close> | 
|   2030 text \<open> |   2059 text \<open> | 
|   2031   While the information about which theorems are used is obvious in |   2060   While the information about which theorems are used is obvious in | 
|   2032   the first two proofs, it is not obvious in the third, because of the |   2061   the first two proofs, it is not obvious in the third, because of the | 
|   2033   \<open>auto\<close>-step.  Fortunately, ``behind'' every proved theorem is |   2062   \<open>auto\<close>-step.  Fortunately, ``behind'' every proved theorem is | 
|   2034   a proof-tree that records all theorems that are employed for |   2063   a proof-tree that records all theorems that are employed for | 
|   2035   establishing this theorem.  We can traverse this proof-tree |   2064   establishing this theorem.  We can traverse this proof-tree | 
|   2036   extracting this information.  Let us first extract the name of the |   2065   extracting this information.  Let us first extract the name of the | 
|   2037   established theorem. This can be done with |   2066   established theorem. This can be done with | 
|   2038  |   2067  | 
|   2039   @{ML_matchresult_fake [display,gray] |   2068   @{ML_matchresult [display,gray] | 
|   2040   \<open>@{thm my_conjIa} |   2069   \<open>@{thm my_conjIa} | 
|   2041   |> Thm.proof_body_of |   2070   |> Thm.proof_body_of | 
|   2042   |> get_names\<close> |   2071   |> get_names\<close> | 
|   2043   \<open>["Essential.my_conjIa"]\<close>} |   2072   \<open>["Essential.my_conjIa"]\<close>} | 
|   2044  |   2073  | 
|   2046   we established the theorem @{thm [source] my_conjIa}. The function @{ML_ind |   2075   we established the theorem @{thm [source] my_conjIa}. The function @{ML_ind | 
|   2047   proof_body_of in Thm} returns a part of the data that is stored in a |   2076   proof_body_of in Thm} returns a part of the data that is stored in a | 
|   2048   theorem.  Notice that the first proof above references |   2077   theorem.  Notice that the first proof above references | 
|   2049   three theorems, namely @{thm [source] conjI}, @{thm [source] conjunct1}  |   2078   three theorems, namely @{thm [source] conjI}, @{thm [source] conjunct1}  | 
|   2050   and @{thm [source] conjunct2}. We can obtain them by descending into the |   2079   and @{thm [source] conjunct2}. We can obtain them by descending into the | 
|   2051   first level of the proof-tree, as follows. |   2080   proof-tree. The function @{ML get_all_names} recursively selects all names. | 
|   2052  |   2081  | 
|   2053   @{ML_matchresult_fake [display,gray] |   2082   @{ML_response [display,gray] | 
|   2054   \<open>@{thm my_conjIa} |   2083   \<open>@{thm my_conjIa} | 
|   2055   |> Thm.proof_body_of |   2084   |> get_all_names |> sort string_ord\<close> | 
|   2056   |> get_pbodies |   2085   \<open>["", "HOL.All_def", "HOL.FalseE", "HOL.False_def", "HOL.TrueI", "HOL.True_def",  | 
|   2057   |> map get_names |   2086    "HOL.True_or_False", "HOL.allI", "HOL.and_def", "HOL.conjI", | 
|   2058   |> List.concat\<close> |   2087    "HOL.conjunct1", "HOL.conjunct2", "HOL.disjE", "HOL.eqTrueE", "HOL.eqTrueI",  | 
|   2059   \<open>["HOL.conjunct2", "HOL.conjunct1", "HOL.conjI", "Pure.protectD",  |   2088    "HOL.ext", "HOL.fun_cong", "HOL.iffD1", "HOL.iffD2", "HOL.iffI", | 
|   2060   "Pure.protectI"]\<close>} |   2089    "HOL.impI", "HOL.mp", "HOL.or_def", "HOL.refl", "HOL.rev_iffD1",  | 
|         |   2090    "HOL.rev_iffD2", "HOL.spec", "HOL.ssubst", "HOL.subst", "HOL.sym", | 
|         |   2091    "Pure.protectD", "Pure.protectI"]\<close>} | 
|   2061  |   2092  | 
|   2062   The theorems @{thm [source] Pure.protectD} and @{thm [source] |   2093   The theorems @{thm [source] Pure.protectD} and @{thm [source] | 
|   2063   Pure.protectI} that are internal theorems are always part of a |   2094   Pure.protectI} that are internal theorems are always part of a | 
|   2064   proof in Isabelle. Note also that applications of \<open>assumption\<close> do not |   2095   proof in Isabelle. The other theorems are the theorems used in the proofs of the theorems | 
|         |   2096   @{thm [source] conjunct1}, @{thm [source] conjunct2} and @{thm [source] conjI}.  | 
|         |   2097  | 
|         |   2098   Note also that applications of \<open>assumption\<close> do not | 
|   2065   count as a separate theorem, as you can see in the following code. |   2099   count as a separate theorem, as you can see in the following code. | 
|   2066  |   2100  | 
|   2067   @{ML_matchresult_fake [display,gray] |   2101   @{ML_matchresult [display,gray] | 
|   2068   \<open>@{thm my_conjIb} |   2102   \<open>@{thm my_conjIb} | 
|   2069   |> Thm.proof_body_of |   2103   |> get_all_names |> sort string_ord\<close> | 
|   2070   |> get_pbodies |   2104   \<open>["", "Pure.protectD", "Pure.protectI"]\<close>} | 
|   2071   |> map get_names |   2105  | 
|   2072   |> List.concat\<close> |         | 
|   2073   \<open>["Pure.protectD", "Pure.protectI"]\<close>} |         | 
|   2074  |   2106  | 
|   2075   Interestingly, but not surprisingly, the proof of @{thm [source] |   2107   Interestingly, but not surprisingly, the proof of @{thm [source] | 
|   2076   my_conjIc} procceeds quite differently from @{thm [source] my_conjIa} |   2108   my_conjIc} procceeds quite differently from @{thm [source] my_conjIa} | 
|   2077   and @{thm [source] my_conjIb}, as can be seen by the theorems that |   2109   and @{thm [source] my_conjIb}, as can be seen by the theorems that | 
|   2078   are returned for @{thm [source] my_conjIc}. |   2110   are returned for @{thm [source] my_conjIc}. | 
|   2079  |   2111   | 
|   2080   @{ML_matchresult_fake [display,gray] |   2112   @{ML_response [display,gray] | 
|   2081   \<open>@{thm my_conjIc} |   2113   \<open>@{thm my_conjIc} | 
|   2082   |> Thm.proof_body_of |   2114   |> get_all_names\<close> | 
|   2083   |> get_pbodies |   2115   \<open>["HOL.simp_thms_25", "Pure.conjunctionD1", "Pure.conjunctionD2", "Pure.conjunctionI",  | 
|   2084   |> map get_names |   2116     "HOL.rev_mp", "HOL.exI", "HOL.allE", "HOL.exE",\<dots>]\<close>} | 
|   2085   |> List.concat\<close> |   2117  | 
|   2086   \<open>["HOL.Eq_TrueI", "HOL.simp_thms_25", "HOL.eq_reflection", |         | 
|   2087   "HOL.conjunct2", "HOL.conjunct1", "HOL.TrueI", "Pure.protectD", |         | 
|   2088   "Pure.protectI"]\<close>} |         | 
|   2089  |         | 
|   2090   Of course we can also descend into the second level of the tree  |         | 
|   2091   ``behind'' @{thm [source] my_conjIa} say, which |         | 
|   2092   means we obtain the theorems that are used in order to prove |         | 
|   2093   @{thm [source] conjunct1}, @{thm [source] conjunct2} and @{thm [source] conjI}. |         | 
|   2094  |         | 
|   2095   @{ML_matchresult_fake [display, gray] |         | 
|   2096   \<open>@{thm my_conjIa} |         | 
|   2097   |> Thm.proof_body_of |         | 
|   2098   |> get_pbodies |         | 
|   2099   |> map get_pbodies |         | 
|   2100   |> (map o map) get_names |         | 
|   2101   |> List.concat |         | 
|   2102   |> List.concat |         | 
|   2103   |> duplicates (op=)\<close> |         | 
|   2104   \<open>["HOL.spec", "HOL.and_def", "HOL.mp", "HOL.impI", "Pure.protectD", |         | 
|   2105   "Pure.protectI"]\<close>} |         | 
|   2106  |   2118  | 
|   2107   \begin{exercise} |   2119   \begin{exercise} | 
|   2108   Have a look at the theorems that are used when a lemma is ``proved'' |   2120   Have a look at the theorems that are used when a lemma is ``proved'' | 
|   2109   by \isacommand{sorry}?  |   2121   by \isacommand{sorry}?  | 
|   2110   \end{exercise} |   2122   \end{exercise} | 
|   2627   to insert a space (of length 1) before the  |   2639   to insert a space (of length 1) before the  | 
|   2628   @{text [quotes] "and"}. This space is also a place where a line break  |   2640   @{text [quotes] "and"}. This space is also a place where a line break  | 
|   2629   can occur. We do the same after the @{text [quotes] "and"}. This gives you |   2641   can occur. We do the same after the @{text [quotes] "and"}. This gives you | 
|   2630   for example |   2642   for example | 
|   2631  |   2643  | 
|   2632 @{ML_matchresult_fake [display,gray] |   2644 @{ML_response [display,gray] | 
|   2633 \<open>let |   2645 \<open>let | 
|   2634   val ptrs1 = map (Pretty.str o string_of_int) (1 upto 22) |   2646   val ptrs1 = map (Pretty.str o string_of_int) (1 upto 22) | 
|         |   2647 in | 
|         |   2648   pprint (Pretty.blk (0, and_list ptrs1)) | 
|         |   2649 end\<close> | 
|         |   2650 \<open>1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21  | 
|         |   2651 and 22\<close>} | 
|         |   2652 @{ML_response [display,gray] | 
|         |   2653 \<open>let | 
|   2635   val ptrs2 = map (Pretty.str o string_of_int) (10 upto 28) |   2654   val ptrs2 = map (Pretty.str o string_of_int) (10 upto 28) | 
|   2636 in |   2655 in | 
|   2637   pprint (Pretty.blk (0, and_list ptrs1)); |         | 
|   2638   pprint (Pretty.blk (0, and_list ptrs2)) |   2656   pprint (Pretty.blk (0, and_list ptrs2)) | 
|   2639 end\<close> |   2657 end\<close> | 
|   2640 \<open>1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21  |   2658 \<open>10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26, 27 and | 
|   2641 and 22 |         | 
|   2642 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26, 27 and |         | 
|   2643 28\<close>} |   2659 28\<close>} | 
|   2644  |   2660  | 
|   2645   Like @{ML blk in Pretty}, the function @{ML_ind chunks in Pretty} prints out  |   2661   Like @{ML blk in Pretty}, the function @{ML_ind chunks in Pretty} prints out  | 
|   2646   a list of items, but automatically inserts forced breaks between each item. |   2662   a list of items, but automatically inserts forced breaks between each item. | 
|   2647   Compare |   2663   Compare | 
|   2648  |   2664  | 
|   2649   @{ML_matchresult_fake [display,gray] |   2665   @{ML_response [display,gray] | 
|   2650   \<open>let |   2666   \<open>let | 
|   2651   val a_and_b = [Pretty.str "a", Pretty.str "b"] |   2667   val a_and_b = [Pretty.str "a", Pretty.str "b"] | 
|   2652 in |   2668 in | 
|   2653   pprint (Pretty.blk (0, a_and_b)); |   2669   pprint (Pretty.blk (0, a_and_b)) | 
|         |   2670 end\<close> | 
|         |   2671 \<open>ab\<close>} | 
|         |   2672  @{ML_response [display,gray] | 
|         |   2673   \<open>let | 
|         |   2674   val a_and_b = [Pretty.str "a", Pretty.str "b"] | 
|         |   2675 in | 
|   2654   pprint (Pretty.chunks a_and_b) |   2676   pprint (Pretty.chunks a_and_b) | 
|   2655 end\<close> |   2677 end\<close> | 
|   2656 \<open>ab |   2678 \<open>a | 
|   2657 a |         | 
|   2658 b\<close>} |   2679 b\<close>} | 
|   2659  |   2680  | 
|   2660   The function @{ML_ind big_list in Pretty} helps with printing long lists. |   2681   The function @{ML_ind big_list in Pretty} helps with printing long lists. | 
|   2661   It is used for example in the command \isacommand{print\_theorems}. An |   2682   It is used for example in the command \isacommand{print\_theorems}. An | 
|   2662   example is as follows. |   2683   example is as follows. | 
|   2663  |   2684  | 
|   2664   @{ML_matchresult_fake [display,gray] |   2685   @{ML_response [display,gray] | 
|   2665   \<open>let |   2686   \<open>let | 
|   2666   val pstrs = map (Pretty.str o string_of_int) (4 upto 10) |   2687   val pstrs = map (Pretty.str o string_of_int) (4 upto 10) | 
|   2667 in |   2688 in | 
|   2668   pprint (Pretty.big_list "header" pstrs) |   2689   pprint (Pretty.big_list "header" pstrs) | 
|   2669 end\<close> |   2690 end\<close> |