ProgTutorial/Essential.thy
changeset 572 438703674711
parent 569 f875a25aa72d
child 573 321e220a6baa
equal deleted inserted replaced
571:95b42288294e 572:438703674711
    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
    95   \ref{sec:printing}), but none for beta-contractions. However you can avoid
   104   \ref{sec:printing}), but none for beta-contractions. However you can avoid
    96   the beta-contractions by switching off abbreviations using the configuration
   105   the beta-contractions by switching off abbreviations using the configuration
    97   value @{ML_ind show_abbrevs in Syntax}. For example
   106   value @{ML_ind show_abbrevs in Syntax}. For example
    98 
   107 
    99 
   108 
   100   @{ML_matchresult_fake [display, gray]
   109   @{ML_response [display, gray]
   101   \<open>let 
   110   \<open>let 
   102   val redex = @{term "(\<lambda>(x::int) (y::int). x)"} 
   111   val redex = @{term "(\<lambda>(x::int) (y::int). x)"} 
   103   val arg1 = @{term "1::int"} 
   112   val arg1 = @{term "1::int"} 
   104   val arg2 = @{term "2::int"}
   113   val arg2 = @{term "2::int"}
   105   val ctxt = Config.put show_abbrevs false @{context}
   114   val ctxt = Config.put show_abbrevs false @{context}
   112   Isabelle makes a distinction between \emph{free} variables (term-constructor
   121   Isabelle makes a distinction between \emph{free} variables (term-constructor
   113   @{ML Free} and written on the user level in blue colour) and
   122   @{ML Free} and written on the user level in blue colour) and
   114   \emph{schematic} variables (term-constructor @{ML Var} and written with a
   123   \emph{schematic} variables (term-constructor @{ML Var} and written with a
   115   leading question mark). Consider the following two examples
   124   leading question mark). Consider the following two examples
   116   
   125   
   117   @{ML_matchresult_fake [display, gray]
   126   @{ML_response [display, gray]
   118 \<open>let
   127 \<open>let
   119   val v1 = Var (("x", 3), @{typ bool})
   128   val v1 = Var (("x", 3), @{typ bool})
   120   val v2 = Var (("x1", 3), @{typ bool})
   129   val v2 = Var (("x1", 3), @{typ bool})
   121   val v3 = Free ("x", @{typ bool})
   130   val v3 = Free ("x", @{typ bool})
   122 in
   131 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 
   202   is needed whenever a term is constructed that will be proved as a theorem. 
   211   is needed whenever a term is constructed that will be proved as a theorem. 
   203 
   212 
   204   As already seen above, types can be constructed using the antiquotation 
   213   As already seen above, types can be constructed using the antiquotation 
   205   \<open>@{typ _}\<close>. For example:
   214   \<open>@{typ _}\<close>. For example:
   206 
   215 
   207   @{ML_matchresult_fake [display,gray] \<open>@{typ "bool \<Rightarrow> nat"}\<close> \<open>bool \<Rightarrow> nat\<close>}
   216   @{ML_response [display,gray] \<open>@{typ "bool \<Rightarrow> nat"}\<close> \<open>bool \<Rightarrow> nat\<close>}
   208 
   217 
   209   The corresponding datatype is
   218   The corresponding datatype is
   210 \<close>
   219 \<close>
   211   
   220   
   212 ML_val %grayML\<open>datatype typ =
   221 ML_val %grayML\<open>datatype typ =
   292 
   301 
   293 text \<open>
   302 text \<open>
   294   After that the types for booleans, lists and so on are printed out again 
   303   After that the types for booleans, lists and so on are printed out again 
   295   the standard Isabelle way.
   304   the standard Isabelle way.
   296 
   305 
   297   @{ML_matchresult_fake [display, gray]
   306   @{ML_response [display, gray]
   298   \<open>@{typ "bool"};
   307   \<open>(@{typ "bool"},
   299 @{typ "'a list"}\<close>
   308 @{typ "'a list"})\<close>
   300   \<open>"bool"
   309   \<open>("bool",
   301 "'a List.list"\<close>}
   310 "'a list")\<close>}
   302 
   311 
   303   \begin{readmore}
   312   \begin{readmore}
   304   Types are described in detail in \isccite{sec:types}. Their
   313   Types are described in detail in \isccite{sec:types}. Their
   305   definition and many useful operations are implemented 
   314   definition and many useful operations are implemented 
   306   in @{ML_file "Pure/type.ML"}.
   315   in @{ML_file "Pure/type.ML"}.
   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>
   464 text \<open>
   473 text \<open>
   465   As said above, after calling @{ML strip_alls}, you obtain a term with loose
   474   As said above, after calling @{ML strip_alls}, you obtain a term with loose
   466   bound variables. With the function @{ML subst_bounds}, you can replace these
   475   bound variables. With the function @{ML subst_bounds}, you can replace these
   467   loose @{ML_ind Bound in Term}s with the stripped off variables.
   476   loose @{ML_ind Bound in Term}s with the stripped off variables.
   468 
   477 
   469   @{ML_matchresult_fake [display, gray, linenos]
   478   @{ML_response [display, gray, linenos]
   470   \<open>let
   479   \<open>let
   471   val (vrs, trm) = strip_alls @{term "\<forall>x y. x = (y::bool)"}
   480   val (vrs, trm) = strip_alls @{term "\<forall>x y. x = (y::bool)"}
   472 in 
   481 in 
   473   subst_bounds (rev vrs, trm) 
   482   subst_bounds (rev vrs, trm) 
   474   |> pretty_term @{context}
   483   |> pretty_term @{context}
   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}
   517   between the first two quantified variables.
   526   between the first two quantified variables.
   518 
   527 
   519   The second function, @{ML_ind loose_bvar1 in Text}, tests whether a term
   528   The second function, @{ML_ind loose_bvar1 in Text}, tests whether a term
   520   contains a loose bound of a certain index. For example
   529   contains a loose bound of a certain index. For example
   521 
   530 
   522   @{ML_matchresult_fake [gray,display]
   531   @{ML_matchresult [gray,display]
   523 \<open>let
   532 \<open>let
   524   val body = snd (strip_alls @{term "\<forall>x y. x = (y::bool)"})
   533   val body = snd (strip_alls @{term "\<forall>x y. x = (y::bool)"})
   525 in
   534 in
   526   [loose_bvar1 (body, 0),
   535   [loose_bvar1 (body, 0),
   527    loose_bvar1 (body, 1),
   536    loose_bvar1 (body, 1),
   532   There are also many convenient functions that construct specific HOL-terms
   541   There are also many convenient functions that construct specific HOL-terms
   533   in the structure @{ML_structure HOLogic}. For example @{ML_ind mk_eq in
   542   in the structure @{ML_structure HOLogic}. For example @{ML_ind mk_eq in
   534   HOLogic} constructs an equality out of two terms.  The types needed in this
   543   HOLogic} constructs an equality out of two terms.  The types needed in this
   535   equality are calculated from the type of the arguments. For example
   544   equality are calculated from the type of the arguments. For example
   536 
   545 
   537 @{ML_matchresult_fake [gray,display]
   546 @{ML_response [gray,display]
   538 \<open>let
   547 \<open>let
   539   val eq = HOLogic.mk_eq (@{term "True"}, @{term "False"})
   548   val eq = HOLogic.mk_eq (@{term "True"}, @{term "False"})
   540 in
   549 in
   541   eq |> pretty_term @{context}
   550   eq |> pretty_term @{context}
   542      |> pwriteln
   551      |> pwriteln
   672    | _ => ty)\<close>
   681    | _ => ty)\<close>
   673 
   682 
   674 text \<open>
   683 text \<open>
   675   Here is an example:
   684   Here is an example:
   676 
   685 
   677 @{ML_matchresult_fake [display,gray] 
   686 @{ML_response [display,gray] 
   678 \<open>map_types nat_to_int @{term "a = (1::nat)"}\<close> 
   687 \<open>map_types nat_to_int @{term "a = (1::nat)"}\<close> 
   679 \<open>Const ("op =", "int \<Rightarrow> int \<Rightarrow> bool")
   688 \<open>Const ("HOL.eq", "int \<Rightarrow> int \<Rightarrow> bool")
   680            $ Free ("a", "int") $ Const ("HOL.one_class.one", "int")\<close>}
   689            $ Free ("a", "int") $ Const ("Groups.one_class.one", "int")\<close>}
   681 
   690 
   682   If you want to obtain the list of free type-variables of a term, you
   691   If you want to obtain the list of free type-variables of a term, you
   683   can use the function @{ML_ind  add_tfrees in Term} 
   692   can use the function @{ML_ind  add_tfrees in Term} 
   684   (similarly @{ML_ind  add_tvars in Term} for the schematic type-variables). 
   693   (similarly @{ML_ind  add_tvars in Term} for the schematic type-variables). 
   685   One would expect that such functions
   694   One would expect that such functions
   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]
   837   use in what follows our own pretty-printing function from
   846   use in what follows our own pretty-printing function from
   838   Figure~\ref{fig:prettyenv} for @{ML_type "Type.tyenv"}s. For the type
   847   Figure~\ref{fig:prettyenv} for @{ML_type "Type.tyenv"}s. For the type
   839   environment in the example this function prints out the more legible:
   848   environment in the example this function prints out the more legible:
   840 
   849 
   841 
   850 
   842   @{ML_matchresult_fake [display, gray]
   851   @{ML_response [display, gray]
   843   \<open>pretty_tyenv @{context} tyenv_unif\<close>
   852   \<open>pretty_tyenv @{context} tyenv_unif\<close>
   844   \<open>[?'a := ?'b list, ?'b := nat]\<close>}
   853   \<open>[?'a := ?'b list, ?'b := nat]\<close>}
   845 
   854 
   846   The way the unification function @{ML typ_unify in Sign} is implemented 
   855   The way the unification function @{ML typ_unify in Sign} is implemented 
   847   using an initial type environment and initial index makes it easy to
   856   using an initial type environment and initial index makes it easy to
   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 
   897 
   912 
   898   Let us now return to the unification problem \<open>?'a * ?'b\<close> and 
   913   Let us now return to the unification problem \<open>?'a * ?'b\<close> and 
   899   \<open>?'b list * nat\<close> from the beginning of this section, and the 
   914   \<open>?'b list * nat\<close> from the beginning of this section, and the 
   900   calculated type environment @{ML tyenv_unif}:
   915   calculated type environment @{ML tyenv_unif}:
   901 
   916 
   902   @{ML_matchresult_fake [display, gray]
   917   @{ML_response [display, gray]
   903   \<open>pretty_tyenv @{context} tyenv_unif\<close>
   918   \<open>pretty_tyenv @{context} tyenv_unif\<close>
   904   \<open>[?'a := ?'b list, ?'b := nat]\<close>}
   919   \<open>[?'a := ?'b list, ?'b := nat]\<close>}
   905 
   920 
   906   Observe that the type environment which the function @{ML typ_unify in
   921   Observe that the type environment which the function @{ML typ_unify in
   907   Sign} returns is \emph{not} an instantiation in fully solved form: while \<open>?'b\<close> is instantiated to @{typ nat}, this is not propagated to the
   922   Sign} returns is \emph{not} an instantiation in fully solved form: while \<open>?'b\<close> is instantiated to @{typ nat}, this is not propagated to the
   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}
  1048   alpha-equivalence, but this kind of matching should be used with care, since
  1063   alpha-equivalence, but this kind of matching should be used with care, since
  1049   it is not clear whether the result is meaningful. A meaningful example is
  1064   it is not clear whether the result is meaningful. A meaningful example is
  1050   matching \<open>\<lambda>x. P x\<close> against the pattern \<open>\<lambda>y. ?X y\<close>. In this
  1065   matching \<open>\<lambda>x. P x\<close> against the pattern \<open>\<lambda>y. ?X y\<close>. In this
  1051   case, first-order matching produces \<open>[?X := P]\<close>.
  1066   case, first-order matching produces \<open>[?X := P]\<close>.
  1052 
  1067 
  1053   @{ML_matchresult_fake [display, gray]
  1068   @{ML_response [display, gray]
  1054   \<open>let
  1069   \<open>let
  1055   val fo_pat = @{term_pat "\<lambda>y. (?X::nat\<Rightarrow>bool) y"}
  1070   val fo_pat = @{term_pat "\<lambda>y. (?X::nat\<Rightarrow>bool) y"}
  1056   val trm = @{term "\<lambda>x. (P::nat\<Rightarrow>bool) x"} 
  1071   val trm = @{term "\<lambda>x. (P::nat\<Rightarrow>bool) x"} 
  1057   val init = (Vartab.empty, Vartab.empty) 
  1072   val init = (Vartab.empty, Vartab.empty) 
  1058 in
  1073 in
  1099   record of type @{ML_type Envir.env} containing a max-index, a type environment 
  1114   record of type @{ML_type Envir.env} containing a max-index, a type environment 
  1100   and a term environment). The corresponding functions are @{ML_ind match in Pattern}
  1115   and a term environment). The corresponding functions are @{ML_ind match in Pattern}
  1101   and @{ML_ind unify in Pattern}, both implemented in the structure
  1116   and @{ML_ind unify in Pattern}, both implemented in the structure
  1102   @{ML_structure Pattern}. An example for higher-order pattern unification is
  1117   @{ML_structure Pattern}. An example for higher-order pattern unification is
  1103 
  1118 
  1104   @{ML_matchresult_fake [display, gray]
  1119   @{ML_response [display, gray]
  1105   \<open>let
  1120   \<open>let
  1106   val trm1 = @{term_pat "\<lambda>x y. g (?X y x) (f (?Y x))"}
  1121   val trm1 = @{term_pat "\<lambda>x y. g (?X y x) (f (?Y x))"}
  1107   val trm2 = @{term_pat "\<lambda>u v. g u (f u)"}
  1122   val trm2 = @{term_pat "\<lambda>u v. g u (f u)"}
  1108   val init = Envir.empty 0
  1123   val init = Envir.empty 0
  1109   val env = Pattern.unify (Context.Proof @{context}) (trm1, trm2) init (* FIXME: Reference to generic context *)
  1124   val env = Pattern.unify (Context.Proof @{context}) (trm1, trm2) init (* FIXME: Reference to generic 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 
  1179   also tries first to solve the problem by higher-order pattern
  1196   also tries first to solve the problem by higher-order pattern
  1180   unification. Only in case of failure full higher-order unification is
  1197   unification. Only in case of failure full higher-order unification is
  1181   called. This function has a built-in bound, which can be accessed and
  1198   called. This function has a built-in bound, which can be accessed and
  1182   manipulated as a configuration value. For example
  1199   manipulated as a configuration value. For example
  1183 
  1200 
  1184   @{ML_matchresult_fake [display,gray]
  1201   @{ML_matchresult [display,gray]
  1185   \<open>Config.get_global @{theory} (Unify.search_bound)\<close>
  1202   \<open>Config.get_global @{theory} (Unify.search_bound)\<close>
  1186   \<open>Int 60\<close>}
  1203   \<open>60\<close>}
  1187 
  1204 
  1188   If this bound is reached during unification, Isabelle prints out the
  1205   If this bound is reached during unification, Isabelle prints out the
  1189   warning message @{text [quotes] "Unification bound exceeded"} and
  1206   warning message @{text [quotes] "Unification bound exceeded"} and
  1190   plenty of diagnostic information (sometimes annoyingly plenty of 
  1207   plenty of diagnostic information (sometimes annoyingly plenty of 
  1191   information). 
  1208   information). 
  1264   environments (Line 8). As a simple example we instantiate the
  1281   environments (Line 8). As a simple example we instantiate the
  1265   @{thm [source] spec} rule so that its conclusion is of the form 
  1282   @{thm [source] spec} rule so that its conclusion is of the form 
  1266   @{term "Q True"}. 
  1283   @{term "Q True"}. 
  1267  
  1284  
  1268 
  1285 
  1269   @{ML_matchresult_fake [gray,display,linenos] 
  1286   @{ML_response [gray,display,linenos] 
  1270   \<open>let  
  1287   \<open>let  
  1271   val pat = Logic.strip_imp_concl (Thm.prop_of @{thm spec})
  1288   val pat = Logic.strip_imp_concl (Thm.prop_of @{thm spec})
  1272   val trm = @{term "Trueprop (Q True)"}
  1289   val trm = @{term "Trueprop (Q True)"}
  1273   val inst = matcher_inst @{context} pat trm 1
  1290   val inst = matcher_inst @{context} pat trm 1
  1274 in
  1291 in
  1359 
  1376 
  1360   To calculate the type, this function traverses the whole term and will
  1377   To calculate the type, this function traverses the whole term and will
  1361   detect any typing inconsistency. For example changing the type of the variable 
  1378   detect any typing inconsistency. For example changing the type of the variable 
  1362   @{term "x"} from @{typ "nat"} to @{typ "int"} will result in the error message: 
  1379   @{term "x"} from @{typ "nat"} to @{typ "int"} will result in the error message: 
  1363 
  1380 
  1364   @{ML_matchresult_fake [display,gray] 
  1381   @{ML_response [display,gray] 
  1365   \<open>type_of (@{term "f::nat \<Rightarrow> bool"} $ @{term "x::int"})\<close> 
  1382   \<open>type_of (@{term "f::nat \<Rightarrow> bool"} $ @{term "x::int"})\<close> 
  1366   \<open>*** Exception- TYPE ("type_of: type mismatch in application" \<dots>\<close>}
  1383   \<open>exception TYPE raised \<dots>\<close>}
  1367 
  1384 
  1368   Since the complete traversal might sometimes be too costly and
  1385   Since the complete traversal might sometimes be too costly and
  1369   not necessary, there is the function @{ML_ind fastype_of in Term}, which 
  1386   not necessary, there is the function @{ML_ind fastype_of in Term}, which 
  1370   also returns the type of a term.
  1387   also returns the type of a term.
  1371 
  1388 
  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)
  1544   inserts necessary \<open>Trueprop\<close>s.
  1561   inserts necessary \<open>Trueprop\<close>s.
  1545 
  1562 
  1546   If we print out the value of @{ML my_thm} then we see only the 
  1563   If we print out the value of @{ML my_thm} then we see only the 
  1547   final statement of the theorem.
  1564   final statement of the theorem.
  1548 
  1565 
  1549   @{ML_matchresult_fake [display, gray]
  1566   @{ML_response [display, gray]
  1550   \<open>pwriteln (pretty_thm @{context} my_thm)\<close>
  1567   \<open>pwriteln (pretty_thm @{context} my_thm)\<close>
  1551   \<open>\<lbrakk>\<And>x. P x \<Longrightarrow> Q x; P t\<rbrakk> \<Longrightarrow> Q t\<close>}
  1568   \<open>\<lbrakk>\<And>x. P x \<Longrightarrow> Q x; P t\<rbrakk> \<Longrightarrow> Q t\<close>}
  1552 
  1569 
  1553   However, internally the code-snippet constructs the following 
  1570   However, internally the code-snippet constructs the following 
  1554   proof.
  1571   proof.
  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>, 
  1901   @{ML_structure Skip_Proof} that allows us to turn any proposition into a theorem.
  1918   @{ML_structure Skip_Proof} that allows us to turn any proposition into a theorem.
  1902   Potentially making the system unsound.  This is sometimes useful for developing 
  1919   Potentially making the system unsound.  This is sometimes useful for developing 
  1903   purposes, or when explicit proof construction should be omitted due to performace 
  1920   purposes, or when explicit proof construction should be omitted due to performace 
  1904   reasons. An example of this function is as follows:
  1921   reasons. An example of this function is as follows:
  1905 
  1922 
  1906   @{ML_matchresult_fake [display, gray]
  1923   @{ML_response [display, gray]
  1907   \<open>Skip_Proof.make_thm @{theory} @{prop "True = False"}\<close>
  1924   \<open>Skip_Proof.make_thm @{theory} @{prop "True = False"}\<close>
  1908   \<open>True = False\<close>}
  1925   \<open>True = False\<close>}
  1909 
  1926 
  1910   \begin{readmore}
  1927   \begin{readmore}
  1911   Functions that setup goal states and prove theorems are implemented in 
  1928   Functions that setup goal states and prove theorems are implemented in 
  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}.
  1942        @{thm foo_data})]) #> snd\<close>
  1959        @{thm foo_data})]) #> snd\<close>
  1943 
  1960 
  1944 text \<open>
  1961 text \<open>
  1945   The data of the theorem @{thm [source] foo_data'} is then as follows:
  1962   The data of the theorem @{thm [source] foo_data'} is then as follows:
  1946 
  1963 
  1947   @{ML_matchresult_fake [display,gray]
  1964   @{ML_matchresult [display,gray]
  1948   \<open>Thm.get_tags @{thm foo_data'}\<close>
  1965   \<open>Thm.get_tags @{thm foo_data'}\<close>
  1949   \<open>[("name", "General.foo_data'"), ("kind", "lemma"), 
  1966   \<open>[("name", "Essential.foo_data'"), 
  1950  ("case_names", "foo_case_one;foo_case_two")]\<close>}
  1967  ("case_names", "foo_case_one;foo_case_two"), ("kind", "theorem")]\<close>}
  1951 
  1968 
  1952   You can observe the case names of this lemma on the user level when using 
  1969   You can observe the case names of this lemma on the user level when using 
  1953   the proof methods \<open>cases\<close> and \<open>induct\<close>. In the proof below
  1970   the proof methods \<open>cases\<close> and \<open>induct\<close>. In the proof below
  1954 \<close>
  1971 \<close>
  1955 
  1972 
  1999   structure @{ML_structure Proofterm}.
  2016   structure @{ML_structure Proofterm}.
  2000 \<close>
  2017 \<close>
  2001 
  2018 
  2002 ML %grayML\<open>fun pthms_of (PBody {thms, ...}) = map #2 thms 
  2019 ML %grayML\<open>fun pthms_of (PBody {thms, ...}) = map #2 thms 
  2003 val get_names = (map Proofterm.thm_node_name) o pthms_of 
  2020 val get_names = (map Proofterm.thm_node_name) o pthms_of 
  2004 val get_pbodies = map (Future.join o Proofterm.thm_node_body) o pthms_of\<close>
  2021 val get_pbodies = map (Future.join o Proofterm.thm_node_body) o pthms_of
       
  2022 fun get_all_names thm =
       
  2023 let
       
  2024   (*proof body with digest*)
       
  2025   val body = Proofterm.strip_thm (Thm.proof_body_of thm);
       
  2026   (*all theorems used in the graph of nested proofs*)
       
  2027 in Proofterm.fold_body_thms
       
  2028      (fn {name, ...} => insert (op =) name) [body] []
       
  2029 end\<close>
  2005 
  2030 
  2006 text \<open>
  2031 text \<open>
  2007   To see what their purpose is, consider the following three short proofs.
  2032   To see what their purpose is, consider the following three short proofs.
  2008 \<close>
  2033 \<close>
  2009 
  2034 ML \<open>
       
  2035 
       
  2036 \<close>
  2010 lemma my_conjIa:
  2037 lemma my_conjIa:
  2011 shows "A \<and> B \<Longrightarrow> A \<and> B"
  2038 shows "A \<and> B \<Longrightarrow> A \<and> B"
  2012 apply(rule conjI)
  2039 apply(rule conjI)
  2013 apply(drule conjunct1)
  2040 apply(drule conjunct1)
  2014 apply(assumption)
  2041 apply(assumption)
  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}
  2406 lemma trueI_2[my_thms]: "True" by simp
  2418 lemma trueI_2[my_thms]: "True" by simp
  2407 
  2419 
  2408 text \<open>
  2420 text \<open>
  2409   then you can see it is added to the initially empty list.
  2421   then you can see it is added to the initially empty list.
  2410 
  2422 
  2411   @{ML_matchresult_fake [display,gray]
  2423   @{ML_response [display,gray]
  2412   \<open>MyThms.get @{context}\<close> 
  2424   \<open>MyThms.get @{context}\<close> 
  2413   \<open>["True"]\<close>}
  2425   \<open>["True"]\<close>}
  2414 
  2426 
  2415   You can also add theorems using the command \isacommand{declare}.
  2427   You can also add theorems using the command \isacommand{declare}.
  2416 \<close>
  2428 \<close>
  2420 text \<open>
  2432 text \<open>
  2421   With this attribute, the \<open>add\<close> operation is the default and does 
  2433   With this attribute, the \<open>add\<close> operation is the default and does 
  2422   not need to be explicitly given. These three declarations will cause the 
  2434   not need to be explicitly given. These three declarations will cause the 
  2423   theorem list to be updated as:
  2435   theorem list to be updated as:
  2424 
  2436 
  2425   @{ML_matchresult_fake [display,gray]
  2437   @{ML_response [display,gray]
  2426   \<open>MyThms.get @{context}\<close>
  2438   \<open>MyThms.get @{context}\<close>
  2427   \<open>["True", "Suc (Suc 0) = 2"]\<close>}
  2439   \<open>["True", "Suc (Suc 0) = 2"]\<close>}
  2428 
  2440 
  2429   The theorem @{thm [source] trueI_2} only appears once, since the 
  2441   The theorem @{thm [source] trueI_2} only appears once, since the 
  2430   function @{ML_ind  add_thm in Thm} tests for duplicates, before extending
  2442   function @{ML_ind  add_thm in Thm} tests for duplicates, before extending
  2433 
  2445 
  2434 declare test[my_thms del]
  2446 declare test[my_thms del]
  2435 
  2447 
  2436 text \<open>After this, the theorem list is again: 
  2448 text \<open>After this, the theorem list is again: 
  2437 
  2449 
  2438   @{ML_matchresult_fake [display,gray]
  2450   @{ML_response [display,gray]
  2439   \<open>MyThms.get @{context}\<close>
  2451   \<open>MyThms.get @{context}\<close>
  2440   \<open>["True"]\<close>}
  2452   \<open>["True"]\<close>}
  2441 
  2453 
  2442   We used in this example two functions declared as @{ML_ind
  2454   We used in this example two functions declared as @{ML_ind
  2443   declaration_attribute in Thm}, but there can be any number of them. We just
  2455   declaration_attribute in Thm}, but there can be any number of them. We just
  2480   follows we will use the following wrapper function for printing a pretty
  2492   follows we will use the following wrapper function for printing a pretty
  2481   type:
  2493   type:
  2482 \<close>
  2494 \<close>
  2483 
  2495 
  2484 ML %grayML\<open>fun pprint prt = tracing (Pretty.string_of prt)\<close>
  2496 ML %grayML\<open>fun pprint prt = tracing (Pretty.string_of prt)\<close>
  2485 
  2497 ML %invisible\<open>val pprint = YXML.content_of o Pretty.string_of\<close>
  2486 text \<open>
  2498 text \<open>
  2487   The point of the pretty-printing infrastructure is to give hints about how to
  2499   The point of the pretty-printing infrastructure is to give hints about how to
  2488   layout text and let Isabelle do the actual layout. Let us first explain
  2500   layout text and let Isabelle do the actual layout. Let us first explain
  2489   how you can insert places where a line break can occur. For this assume the
  2501   how you can insert places where a line break can occur. For this assume the
  2490   following function that replicates a string n times:
  2502   following function that replicates a string n times:
  2510 aaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fo
  2522 aaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fo
  2511 oooooooooooooobaaaaaaaaaaaar\<close>}
  2523 oooooooooooooobaaaaaaaaaaaar\<close>}
  2512 
  2524 
  2513   We obtain the same if we just use the function @{ML pprint}.
  2525   We obtain the same if we just use the function @{ML pprint}.
  2514 
  2526 
  2515 @{ML_matchresult_fake [display,gray]
  2527 @{ML_response [display,gray]
  2516 \<open>pprint (Pretty.str test_str)\<close>
  2528 \<open>pprint (Pretty.str test_str)\<close>
  2517 \<open>fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar foooooooooo
  2529 \<open>fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar foooooooooo
  2518 ooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaa
  2530 ooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaa
  2519 aaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fo
  2531 aaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fo
  2520 oooooooooooooobaaaaaaaaaaaar\<close>}
  2532 oooooooooooooobaaaaaaaaaaaar\<close>}
  2524   function @{ML_ind breaks in Pretty}, which expects a list of pretty types
  2536   function @{ML_ind breaks in Pretty}, which expects a list of pretty types
  2525   and inserts a possible line break in between every two elements in this
  2537   and inserts a possible line break in between every two elements in this
  2526   list. To print this list of pretty types as a single string, we concatenate
  2538   list. To print this list of pretty types as a single string, we concatenate
  2527   them with the function @{ML_ind blk in Pretty} as follows:
  2539   them with the function @{ML_ind blk in Pretty} as follows:
  2528 
  2540 
  2529 @{ML_matchresult_fake [display,gray]
  2541 @{ML_response [display,gray]
  2530 \<open>let
  2542 \<open>let
  2531   val ptrs = map Pretty.str (space_explode " " test_str)
  2543   val ptrs = map Pretty.str (space_explode " " test_str)
  2532 in
  2544 in
  2533   pprint (Pretty.blk (0, Pretty.breaks ptrs))
  2545   pprint (Pretty.blk (0, Pretty.breaks ptrs))
  2534 end\<close>
  2546 end\<close>
  2540   Here the layout of @{ML test_str} is much more pleasing to the 
  2552   Here the layout of @{ML test_str} is much more pleasing to the 
  2541   eye. The @{ML \<open>0\<close>} in @{ML_ind  blk in Pretty} stands for no hanging 
  2553   eye. The @{ML \<open>0\<close>} in @{ML_ind  blk in Pretty} stands for no hanging 
  2542   indentation of the printed string. You can increase the indentation 
  2554   indentation of the printed string. You can increase the indentation 
  2543   and obtain
  2555   and obtain
  2544   
  2556   
  2545 @{ML_matchresult_fake [display,gray]
  2557 @{ML_response [display,gray]
  2546 \<open>let
  2558 \<open>let
  2547   val ptrs = map Pretty.str (space_explode " " test_str)
  2559   val ptrs = map Pretty.str (space_explode " " test_str)
  2548 in
  2560 in
  2549   pprint (Pretty.blk (3, Pretty.breaks ptrs))
  2561   pprint (Pretty.blk (3, Pretty.breaks ptrs))
  2550 end\<close>
  2562 end\<close>
  2555 
  2567 
  2556   where starting from the second line the indent is 3. If you want
  2568   where starting from the second line the indent is 3. If you want
  2557   that every line starts with the same indent, you can use the
  2569   that every line starts with the same indent, you can use the
  2558   function @{ML_ind  indent in Pretty} as follows:
  2570   function @{ML_ind  indent in Pretty} as follows:
  2559 
  2571 
  2560 @{ML_matchresult_fake [display,gray]
  2572 @{ML_response [display,gray]
  2561 \<open>let
  2573 \<open>let
  2562   val ptrs = map Pretty.str (space_explode " " test_str)
  2574   val ptrs = map Pretty.str (space_explode " " test_str)
  2563 in
  2575 in
  2564   pprint (Pretty.indent 10 (Pretty.blk (0, Pretty.breaks ptrs)))
  2576   pprint (Pretty.indent 10 (Pretty.blk (0, Pretty.breaks ptrs)))
  2565 end\<close>
  2577 end\<close>
  2570 
  2582 
  2571   If you want to print out a list of items separated by commas and 
  2583   If you want to print out a list of items separated by commas and 
  2572   have the linebreaks handled properly, you can use the function 
  2584   have the linebreaks handled properly, you can use the function 
  2573   @{ML_ind  commas in Pretty}. For example
  2585   @{ML_ind  commas in Pretty}. For example
  2574 
  2586 
  2575 @{ML_matchresult_fake [display,gray]
  2587 @{ML_response [display,gray]
  2576 \<open>let
  2588 \<open>let
  2577   val ptrs = map (Pretty.str o string_of_int) (99998 upto 100020)
  2589   val ptrs = map (Pretty.str o string_of_int) (99998 upto 100020)
  2578 in
  2590 in
  2579   pprint (Pretty.blk (0, Pretty.commas ptrs))
  2591   pprint (Pretty.blk (0, Pretty.commas ptrs))
  2580 end\<close>
  2592 end\<close>
  2588   @{ML_ind  enum in Pretty}. For example
  2600   @{ML_ind  enum in Pretty}. For example
  2589 \<close>
  2601 \<close>
  2590 
  2602 
  2591 text \<open>
  2603 text \<open>
  2592   
  2604   
  2593 @{ML_matchresult_fake [display,gray]
  2605 @{ML_response [display,gray]
  2594 \<open>let
  2606 \<open>let
  2595   val ptrs = map (Pretty.str o string_of_int) (99998 upto 100020)
  2607   val ptrs = map (Pretty.str o string_of_int) (99998 upto 100020)
  2596 in
  2608 in
  2597   pprint (Pretty.enum "," "{" "}" ptrs)
  2609   pprint (Pretty.enum "," "{" "}" ptrs)
  2598 end\<close>
  2610 end\<close>
  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>
  2679   The point of the pretty-printing functions is to conveninetly obtain 
  2700   The point of the pretty-printing functions is to conveninetly obtain 
  2680   a lay-out of terms and types that is pleasing to the eye. If we print
  2701   a lay-out of terms and types that is pleasing to the eye. If we print
  2681   out the the terms produced by the the function @{ML de_bruijn} from 
  2702   out the the terms produced by the the function @{ML de_bruijn} from 
  2682   exercise~\ref{ex:debruijn} we obtain the following: 
  2703   exercise~\ref{ex:debruijn} we obtain the following: 
  2683 
  2704 
  2684   @{ML_matchresult_fake [display,gray]
  2705   @{ML_response [display,gray]
  2685   \<open>pprint (Syntax.pretty_term  @{context} (de_bruijn 4))\<close>
  2706   \<open>pprint (Syntax.pretty_term  @{context} (de_bruijn 4))\<close>
  2686   \<open>(P 3 = P 4 \<longrightarrow> P 4 \<and> P 3 \<and> P 2 \<and> P 1) \<and>
  2707   \<open>(P 3 = P 4 \<longrightarrow> P 4 \<and> P 3 \<and> P 2 \<and> P 1) \<and>
  2687 (P 2 = P 3 \<longrightarrow> P 4 \<and> P 3 \<and> P 2 \<and> P 1) \<and>
  2708 (P 2 = P 3 \<longrightarrow> P 4 \<and> P 3 \<and> P 2 \<and> P 1) \<and>
  2688 (P 1 = P 2 \<longrightarrow> P 4 \<and> P 3 \<and> P 2 \<and> P 1) \<and> 
  2709 (P 1 = P 2 \<longrightarrow> P 4 \<and> P 3 \<and> P 2 \<and> P 1) \<and> 
  2689 (P 1 = P 4 \<longrightarrow> P 4 \<and> P 3 \<and> P 2 \<and> P 1) \<longrightarrow>
  2710 (P 1 = P 4 \<longrightarrow> P 4 \<and> P 3 \<and> P 2 \<and> P 1) \<longrightarrow>