doc-src/IsarRef/Thy/HOL_Specific.thy
author wenzelm
Wed, 08 Jun 2011 10:24:07 +0200
changeset 44141 bc72c1ccc89e
parent 43881 665623e695ea
parent 44134 41394a61cca9
child 44440 36ba44fe0781
permissions -rw-r--r--
merged
     1 theory HOL_Specific
     2 imports Base Main
     3 begin
     4 
     5 chapter {* Isabelle/HOL \label{ch:hol} *}
     6 
     7 section {* Higher-Order Logic *}
     8 
     9 text {* Isabelle/HOL is based on Higher-Order Logic, a polymorphic
    10   version of Church's Simple Theory of Types.  HOL can be best
    11   understood as a simply-typed version of classical set theory.  The
    12   logic was first implemented in Gordon's HOL system
    13   \cite{mgordon-hol}.  It extends Church's original logic
    14   \cite{church40} by explicit type variables (naive polymorphism) and
    15   a sound axiomatization scheme for new types based on subsets of
    16   existing types.
    17 
    18   Andrews's book \cite{andrews86} is a full description of the
    19   original Church-style higher-order logic, with proofs of correctness
    20   and completeness wrt.\ certain set-theoretic interpretations.  The
    21   particular extensions of Gordon-style HOL are explained semantically
    22   in two chapters of the 1993 HOL book \cite{pitts93}.
    23 
    24   Experience with HOL over decades has demonstrated that higher-order
    25   logic is widely applicable in many areas of mathematics and computer
    26   science.  In a sense, Higher-Order Logic is simpler than First-Order
    27   Logic, because there are fewer restrictions and special cases.  Note
    28   that HOL is \emph{weaker} than FOL with axioms for ZF set theory,
    29   which is traditionally considered the standard foundation of regular
    30   mathematics, but for most applications this does not matter.  If you
    31   prefer ML to Lisp, you will probably prefer HOL to ZF.
    32 
    33   \medskip The syntax of HOL follows @{text "\<lambda>"}-calculus and
    34   functional programming.  Function application is curried.  To apply
    35   the function @{text f} of type @{text "\<tau>\<^sub>1 \<Rightarrow> \<tau>\<^sub>2 \<Rightarrow> \<tau>\<^sub>3"} to the
    36   arguments @{text a} and @{text b} in HOL, you simply write @{text "f
    37   a b"} (as in ML or Haskell).  There is no ``apply'' operator; the
    38   existing application of the Pure @{text "\<lambda>"}-calculus is re-used.
    39   Note that in HOL @{text "f (a, b)"} means ``@{text "f"} applied to
    40   the pair @{text "(a, b)"} (which is notation for @{text "Pair a
    41   b"}).  The latter typically introduces extra formal efforts that can
    42   be avoided by currying functions by default.  Explicit tuples are as
    43   infrequent in HOL formalizations as in good ML or Haskell programs.
    44 
    45   \medskip Isabelle/HOL has a distinct feel, compared to other
    46   object-logics like Isabelle/ZF.  It identifies object-level types
    47   with meta-level types, taking advantage of the default
    48   type-inference mechanism of Isabelle/Pure.  HOL fully identifies
    49   object-level functions with meta-level functions, with native
    50   abstraction and application.
    51 
    52   These identifications allow Isabelle to support HOL particularly
    53   nicely, but they also mean that HOL requires some sophistication
    54   from the user.  In particular, an understanding of Hindley-Milner
    55   type-inference with type-classes, which are both used extensively in
    56   the standard libraries and applications.  Beginners can set
    57   @{attribute show_types} or even @{attribute show_sorts} to get more
    58   explicit information about the result of type-inference.  *}
    59 
    60 
    61 section {* Inductive and coinductive definitions \label{sec:hol-inductive} *}
    62 
    63 text {* An \emph{inductive definition} specifies the least predicate
    64   or set @{text R} closed under given rules: applying a rule to
    65   elements of @{text R} yields a result within @{text R}.  For
    66   example, a structural operational semantics is an inductive
    67   definition of an evaluation relation.
    68 
    69   Dually, a \emph{coinductive definition} specifies the greatest
    70   predicate or set @{text R} that is consistent with given rules:
    71   every element of @{text R} can be seen as arising by applying a rule
    72   to elements of @{text R}.  An important example is using
    73   bisimulation relations to formalise equivalence of processes and
    74   infinite data structures.
    75   
    76   Both inductive and coinductive definitions are based on the
    77   Knaster-Tarski fixed-point theorem for complete lattices.  The
    78   collection of introduction rules given by the user determines a
    79   functor on subsets of set-theoretic relations.  The required
    80   monotonicity of the recursion scheme is proven as a prerequisite to
    81   the fixed-point definition and the resulting consequences.  This
    82   works by pushing inclusion through logical connectives and any other
    83   operator that might be wrapped around recursive occurrences of the
    84   defined relation: there must be a monotonicity theorem of the form
    85   @{text "A \<le> B \<Longrightarrow> \<M> A \<le> \<M> B"}, for each premise @{text "\<M> R t"} in an
    86   introduction rule.  The default rule declarations of Isabelle/HOL
    87   already take care of most common situations.
    88 
    89   \begin{matharray}{rcl}
    90     @{command_def (HOL) "inductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
    91     @{command_def (HOL) "inductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
    92     @{command_def (HOL) "coinductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
    93     @{command_def (HOL) "coinductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
    94     @{attribute_def (HOL) mono} & : & @{text attribute} \\
    95   \end{matharray}
    96 
    97   @{rail "
    98     (@@{command (HOL) inductive} | @@{command (HOL) inductive_set} |
    99       @@{command (HOL) coinductive} | @@{command (HOL) coinductive_set})
   100     @{syntax target}? \\
   101     @{syntax \"fixes\"} (@'for' @{syntax \"fixes\"})? (@'where' clauses)? \\
   102     (@'monos' @{syntax thmrefs})?
   103     ;
   104     clauses: (@{syntax thmdecl}? @{syntax prop} + '|')
   105     ;
   106     @@{attribute (HOL) mono} (() | 'add' | 'del')
   107   "}
   108 
   109   \begin{description}
   110 
   111   \item @{command (HOL) "inductive"} and @{command (HOL)
   112   "coinductive"} define (co)inductive predicates from the introduction
   113   rules.
   114 
   115   The propositions given as @{text "clauses"} in the @{keyword
   116   "where"} part are either rules of the usual @{text "\<And>/\<Longrightarrow>"} format
   117   (with arbitrary nesting), or equalities using @{text "\<equiv>"}.  The
   118   latter specifies extra-logical abbreviations in the sense of
   119   @{command_ref abbreviation}.  Introducing abstract syntax
   120   simultaneously with the actual introduction rules is occasionally
   121   useful for complex specifications.
   122 
   123   The optional @{keyword "for"} part contains a list of parameters of
   124   the (co)inductive predicates that remain fixed throughout the
   125   definition, in contrast to arguments of the relation that may vary
   126   in each occurrence within the given @{text "clauses"}.
   127 
   128   The optional @{keyword "monos"} declaration contains additional
   129   \emph{monotonicity theorems}, which are required for each operator
   130   applied to a recursive set in the introduction rules.
   131 
   132   \item @{command (HOL) "inductive_set"} and @{command (HOL)
   133   "coinductive_set"} are wrappers for to the previous commands for
   134   native HOL predicates.  This allows to define (co)inductive sets,
   135   where multiple arguments are simulated via tuples.
   136 
   137   \item @{attribute (HOL) mono} declares monotonicity rules in the
   138   context.  These rule are involved in the automated monotonicity
   139   proof of the above inductive and coinductive definitions.
   140 
   141   \end{description}
   142 *}
   143 
   144 
   145 subsection {* Derived rules *}
   146 
   147 text {* A (co)inductive definition of @{text R} provides the following
   148   main theorems:
   149 
   150   \begin{description}
   151 
   152   \item @{text R.intros} is the list of introduction rules as proven
   153   theorems, for the recursive predicates (or sets).  The rules are
   154   also available individually, using the names given them in the
   155   theory file;
   156 
   157   \item @{text R.cases} is the case analysis (or elimination) rule;
   158 
   159   \item @{text R.induct} or @{text R.coinduct} is the (co)induction
   160   rule.
   161 
   162   \end{description}
   163 
   164   When several predicates @{text "R\<^sub>1, \<dots>, R\<^sub>n"} are
   165   defined simultaneously, the list of introduction rules is called
   166   @{text "R\<^sub>1_\<dots>_R\<^sub>n.intros"}, the case analysis rules are
   167   called @{text "R\<^sub>1.cases, \<dots>, R\<^sub>n.cases"}, and the list
   168   of mutual induction rules is called @{text
   169   "R\<^sub>1_\<dots>_R\<^sub>n.inducts"}.
   170 *}
   171 
   172 
   173 subsection {* Monotonicity theorems *}
   174 
   175 text {* The context maintains a default set of theorems that are used
   176   in monotonicity proofs.  New rules can be declared via the
   177   @{attribute (HOL) mono} attribute.  See the main Isabelle/HOL
   178   sources for some examples.  The general format of such monotonicity
   179   theorems is as follows:
   180 
   181   \begin{itemize}
   182 
   183   \item Theorems of the form @{text "A \<le> B \<Longrightarrow> \<M> A \<le> \<M> B"}, for proving
   184   monotonicity of inductive definitions whose introduction rules have
   185   premises involving terms such as @{text "\<M> R t"}.
   186 
   187   \item Monotonicity theorems for logical operators, which are of the
   188   general form @{text "(\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> (\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> \<longrightarrow> \<dots>"}.  For example, in
   189   the case of the operator @{text "\<or>"}, the corresponding theorem is
   190   \[
   191   \infer{@{text "P\<^sub>1 \<or> P\<^sub>2 \<longrightarrow> Q\<^sub>1 \<or> Q\<^sub>2"}}{@{text "P\<^sub>1 \<longrightarrow> Q\<^sub>1"} & @{text "P\<^sub>2 \<longrightarrow> Q\<^sub>2"}}
   192   \]
   193 
   194   \item De Morgan style equations for reasoning about the ``polarity''
   195   of expressions, e.g.
   196   \[
   197   @{prop "\<not> \<not> P \<longleftrightarrow> P"} \qquad\qquad
   198   @{prop "\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q"}
   199   \]
   200 
   201   \item Equations for reducing complex operators to more primitive
   202   ones whose monotonicity can easily be proved, e.g.
   203   \[
   204   @{prop "(P \<longrightarrow> Q) \<longleftrightarrow> \<not> P \<or> Q"} \qquad\qquad
   205   @{prop "Ball A P \<equiv> \<forall>x. x \<in> A \<longrightarrow> P x"}
   206   \]
   207 
   208   \end{itemize}
   209 *}
   210 
   211 subsubsection {* Examples *}
   212 
   213 text {* The finite powerset operator can be defined inductively like this: *}
   214 
   215 inductive_set Fin :: "'a set \<Rightarrow> 'a set set" for A :: "'a set"
   216 where
   217   empty: "{} \<in> Fin A"
   218 | insert: "a \<in> A \<Longrightarrow> B \<in> Fin A \<Longrightarrow> insert a B \<in> Fin A"
   219 
   220 text {* The accessible part of a relation is defined as follows: *}
   221 
   222 inductive acc :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
   223   for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool"  (infix "\<prec>" 50)
   224 where acc: "(\<And>y. y \<prec> x \<Longrightarrow> acc r y) \<Longrightarrow> acc r x"
   225 
   226 text {* Common logical connectives can be easily characterized as
   227 non-recursive inductive definitions with parameters, but without
   228 arguments. *}
   229 
   230 inductive AND for A B :: bool
   231 where "A \<Longrightarrow> B \<Longrightarrow> AND A B"
   232 
   233 inductive OR for A B :: bool
   234 where "A \<Longrightarrow> OR A B"
   235   | "B \<Longrightarrow> OR A B"
   236 
   237 inductive EXISTS for B :: "'a \<Rightarrow> bool"
   238 where "B a \<Longrightarrow> EXISTS B"
   239 
   240 text {* Here the @{text "cases"} or @{text "induct"} rules produced by
   241   the @{command inductive} package coincide with the expected
   242   elimination rules for Natural Deduction.  Already in the original
   243   article by Gerhard Gentzen \cite{Gentzen:1935} there is a hint that
   244   each connective can be characterized by its introductions, and the
   245   elimination can be constructed systematically. *}
   246 
   247 
   248 section {* Recursive functions \label{sec:recursion} *}
   249 
   250 text {*
   251   \begin{matharray}{rcl}
   252     @{command_def (HOL) "primrec"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   253     @{command_def (HOL) "fun"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   254     @{command_def (HOL) "function"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
   255     @{command_def (HOL) "termination"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
   256   \end{matharray}
   257 
   258   @{rail "
   259     @@{command (HOL) primrec} @{syntax target}? @{syntax \"fixes\"} @'where' equations
   260     ;
   261     (@@{command (HOL) fun} | @@{command (HOL) function}) @{syntax target}? functionopts?
   262       @{syntax \"fixes\"} \\ @'where' equations
   263     ;
   264 
   265     equations: (@{syntax thmdecl}? @{syntax prop} + '|')
   266     ;
   267     functionopts: '(' (('sequential' | 'domintros') + ',') ')'
   268     ;
   269     @@{command (HOL) termination} @{syntax term}?
   270   "}
   271 
   272   \begin{description}
   273 
   274   \item @{command (HOL) "primrec"} defines primitive recursive
   275   functions over datatypes (see also @{command_ref (HOL) datatype} and
   276   @{command_ref (HOL) rep_datatype}).  The given @{text equations}
   277   specify reduction rules that are produced by instantiating the
   278   generic combinator for primitive recursion that is available for
   279   each datatype.
   280 
   281   Each equation needs to be of the form:
   282 
   283   @{text [display] "f x\<^sub>1 \<dots> x\<^sub>m (C y\<^sub>1 \<dots> y\<^sub>k) z\<^sub>1 \<dots> z\<^sub>n = rhs"}
   284 
   285   such that @{text C} is a datatype constructor, @{text rhs} contains
   286   only the free variables on the left-hand side (or from the context),
   287   and all recursive occurrences of @{text "f"} in @{text "rhs"} are of
   288   the form @{text "f \<dots> y\<^sub>i \<dots>"} for some @{text i}.  At most one
   289   reduction rule for each constructor can be given.  The order does
   290   not matter.  For missing constructors, the function is defined to
   291   return a default value, but this equation is made difficult to
   292   access for users.
   293 
   294   The reduction rules are declared as @{attribute simp} by default,
   295   which enables standard proof methods like @{method simp} and
   296   @{method auto} to normalize expressions of @{text "f"} applied to
   297   datatype constructions, by simulating symbolic computation via
   298   rewriting.
   299 
   300   \item @{command (HOL) "function"} defines functions by general
   301   wellfounded recursion. A detailed description with examples can be
   302   found in \cite{isabelle-function}. The function is specified by a
   303   set of (possibly conditional) recursive equations with arbitrary
   304   pattern matching. The command generates proof obligations for the
   305   completeness and the compatibility of patterns.
   306 
   307   The defined function is considered partial, and the resulting
   308   simplification rules (named @{text "f.psimps"}) and induction rule
   309   (named @{text "f.pinduct"}) are guarded by a generated domain
   310   predicate @{text "f_dom"}. The @{command (HOL) "termination"}
   311   command can then be used to establish that the function is total.
   312 
   313   \item @{command (HOL) "fun"} is a shorthand notation for ``@{command
   314   (HOL) "function"}~@{text "(sequential)"}, followed by automated
   315   proof attempts regarding pattern matching and termination.  See
   316   \cite{isabelle-function} for further details.
   317 
   318   \item @{command (HOL) "termination"}~@{text f} commences a
   319   termination proof for the previously defined function @{text f}.  If
   320   this is omitted, the command refers to the most recent function
   321   definition.  After the proof is closed, the recursive equations and
   322   the induction principle is established.
   323 
   324   \end{description}
   325 
   326   Recursive definitions introduced by the @{command (HOL) "function"}
   327   command accommodate reasoning by induction (cf.\ @{method induct}):
   328   rule @{text "f.induct"} refers to a specific induction rule, with
   329   parameters named according to the user-specified equations. Cases
   330   are numbered starting from 1.  For @{command (HOL) "primrec"}, the
   331   induction principle coincides with structural recursion on the
   332   datatype where the recursion is carried out.
   333 
   334   The equations provided by these packages may be referred later as
   335   theorem list @{text "f.simps"}, where @{text f} is the (collective)
   336   name of the functions defined.  Individual equations may be named
   337   explicitly as well.
   338 
   339   The @{command (HOL) "function"} command accepts the following
   340   options.
   341 
   342   \begin{description}
   343 
   344   \item @{text sequential} enables a preprocessor which disambiguates
   345   overlapping patterns by making them mutually disjoint.  Earlier
   346   equations take precedence over later ones.  This allows to give the
   347   specification in a format very similar to functional programming.
   348   Note that the resulting simplification and induction rules
   349   correspond to the transformed specification, not the one given
   350   originally. This usually means that each equation given by the user
   351   may result in several theorems.  Also note that this automatic
   352   transformation only works for ML-style datatype patterns.
   353 
   354   \item @{text domintros} enables the automated generation of
   355   introduction rules for the domain predicate. While mostly not
   356   needed, they can be helpful in some proofs about partial functions.
   357 
   358   \end{description}
   359 *}
   360 
   361 subsubsection {* Example: evaluation of expressions *}
   362 
   363 text {* Subsequently, we define mutual datatypes for arithmetic and
   364   boolean expressions, and use @{command primrec} for evaluation
   365   functions that follow the same recursive structure. *}
   366 
   367 datatype 'a aexp =
   368     IF "'a bexp"  "'a aexp"  "'a aexp"
   369   | Sum "'a aexp"  "'a aexp"
   370   | Diff "'a aexp"  "'a aexp"
   371   | Var 'a
   372   | Num nat
   373 and 'a bexp =
   374     Less "'a aexp"  "'a aexp"
   375   | And "'a bexp"  "'a bexp"
   376   | Neg "'a bexp"
   377 
   378 
   379 text {* \medskip Evaluation of arithmetic and boolean expressions *}
   380 
   381 primrec evala :: "('a \<Rightarrow> nat) \<Rightarrow> 'a aexp \<Rightarrow> nat"
   382   and evalb :: "('a \<Rightarrow> nat) \<Rightarrow> 'a bexp \<Rightarrow> bool"
   383 where
   384   "evala env (IF b a1 a2) = (if evalb env b then evala env a1 else evala env a2)"
   385 | "evala env (Sum a1 a2) = evala env a1 + evala env a2"
   386 | "evala env (Diff a1 a2) = evala env a1 - evala env a2"
   387 | "evala env (Var v) = env v"
   388 | "evala env (Num n) = n"
   389 | "evalb env (Less a1 a2) = (evala env a1 < evala env a2)"
   390 | "evalb env (And b1 b2) = (evalb env b1 \<and> evalb env b2)"
   391 | "evalb env (Neg b) = (\<not> evalb env b)"
   392 
   393 text {* Since the value of an expression depends on the value of its
   394   variables, the functions @{const evala} and @{const evalb} take an
   395   additional parameter, an \emph{environment} that maps variables to
   396   their values.
   397 
   398   \medskip Substitution on expressions can be defined similarly.  The
   399   mapping @{text f} of type @{typ "'a \<Rightarrow> 'a aexp"} given as a
   400   parameter is lifted canonically on the types @{typ "'a aexp"} and
   401   @{typ "'a bexp"}, respectively.
   402 *}
   403 
   404 primrec substa :: "('a \<Rightarrow> 'b aexp) \<Rightarrow> 'a aexp \<Rightarrow> 'b aexp"
   405   and substb :: "('a \<Rightarrow> 'b aexp) \<Rightarrow> 'a bexp \<Rightarrow> 'b bexp"
   406 where
   407   "substa f (IF b a1 a2) = IF (substb f b) (substa f a1) (substa f a2)"
   408 | "substa f (Sum a1 a2) = Sum (substa f a1) (substa f a2)"
   409 | "substa f (Diff a1 a2) = Diff (substa f a1) (substa f a2)"
   410 | "substa f (Var v) = f v"
   411 | "substa f (Num n) = Num n"
   412 | "substb f (Less a1 a2) = Less (substa f a1) (substa f a2)"
   413 | "substb f (And b1 b2) = And (substb f b1) (substb f b2)"
   414 | "substb f (Neg b) = Neg (substb f b)"
   415 
   416 text {* In textbooks about semantics one often finds substitution
   417   theorems, which express the relationship between substitution and
   418   evaluation.  For @{typ "'a aexp"} and @{typ "'a bexp"}, we can prove
   419   such a theorem by mutual induction, followed by simplification.
   420 *}
   421 
   422 lemma subst_one:
   423   "evala env (substa (Var (v := a')) a) = evala (env (v := evala env a')) a"
   424   "evalb env (substb (Var (v := a')) b) = evalb (env (v := evala env a')) b"
   425   by (induct a and b) simp_all
   426 
   427 lemma subst_all:
   428   "evala env (substa s a) = evala (\<lambda>x. evala env (s x)) a"
   429   "evalb env (substb s b) = evalb (\<lambda>x. evala env (s x)) b"
   430   by (induct a and b) simp_all
   431 
   432 
   433 subsubsection {* Example: a substitution function for terms *}
   434 
   435 text {* Functions on datatypes with nested recursion are also defined
   436   by mutual primitive recursion. *}
   437 
   438 datatype ('a, 'b) "term" = Var 'a | App 'b "('a, 'b) term list"
   439 
   440 text {* A substitution function on type @{typ "('a, 'b) term"} can be
   441   defined as follows, by working simultaneously on @{typ "('a, 'b)
   442   term list"}: *}
   443 
   444 primrec subst_term :: "('a \<Rightarrow> ('a, 'b) term) \<Rightarrow> ('a, 'b) term \<Rightarrow> ('a, 'b) term" and
   445   subst_term_list :: "('a \<Rightarrow> ('a, 'b) term) \<Rightarrow> ('a, 'b) term list \<Rightarrow> ('a, 'b) term list"
   446 where
   447   "subst_term f (Var a) = f a"
   448 | "subst_term f (App b ts) = App b (subst_term_list f ts)"
   449 | "subst_term_list f [] = []"
   450 | "subst_term_list f (t # ts) = subst_term f t # subst_term_list f ts"
   451 
   452 text {* The recursion scheme follows the structure of the unfolded
   453   definition of type @{typ "('a, 'b) term"}.  To prove properties of this
   454   substitution function, mutual induction is needed:
   455 *}
   456 
   457 lemma "subst_term (subst_term f1 \<circ> f2) t = subst_term f1 (subst_term f2 t)" and
   458   "subst_term_list (subst_term f1 \<circ> f2) ts = subst_term_list f1 (subst_term_list f2 ts)"
   459   by (induct t and ts) simp_all
   460 
   461 
   462 subsubsection {* Example: a map function for infinitely branching trees *}
   463 
   464 text {* Defining functions on infinitely branching datatypes by
   465   primitive recursion is just as easy.
   466 *}
   467 
   468 datatype 'a tree = Atom 'a | Branch "nat \<Rightarrow> 'a tree"
   469 
   470 primrec map_tree :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a tree \<Rightarrow> 'b tree"
   471 where
   472   "map_tree f (Atom a) = Atom (f a)"
   473 | "map_tree f (Branch ts) = Branch (\<lambda>x. map_tree f (ts x))"
   474 
   475 text {* Note that all occurrences of functions such as @{text ts}
   476   above must be applied to an argument.  In particular, @{term
   477   "map_tree f \<circ> ts"} is not allowed here. *}
   478 
   479 text {* Here is a simple composition lemma for @{term map_tree}: *}
   480 
   481 lemma "map_tree g (map_tree f t) = map_tree (g \<circ> f) t"
   482   by (induct t) simp_all
   483 
   484 
   485 subsection {* Proof methods related to recursive definitions *}
   486 
   487 text {*
   488   \begin{matharray}{rcl}
   489     @{method_def (HOL) pat_completeness} & : & @{text method} \\
   490     @{method_def (HOL) relation} & : & @{text method} \\
   491     @{method_def (HOL) lexicographic_order} & : & @{text method} \\
   492     @{method_def (HOL) size_change} & : & @{text method} \\
   493   \end{matharray}
   494 
   495   @{rail "
   496     @@{method (HOL) relation} @{syntax term}
   497     ;
   498     @@{method (HOL) lexicographic_order} (@{syntax clasimpmod} * )
   499     ;
   500     @@{method (HOL) size_change} ( orders (@{syntax clasimpmod} * ) )
   501     ;
   502     orders: ( 'max' | 'min' | 'ms' ) *
   503   "}
   504 
   505   \begin{description}
   506 
   507   \item @{method (HOL) pat_completeness} is a specialized method to
   508   solve goals regarding the completeness of pattern matching, as
   509   required by the @{command (HOL) "function"} package (cf.\
   510   \cite{isabelle-function}).
   511 
   512   \item @{method (HOL) relation}~@{text R} introduces a termination
   513   proof using the relation @{text R}.  The resulting proof state will
   514   contain goals expressing that @{text R} is wellfounded, and that the
   515   arguments of recursive calls decrease with respect to @{text R}.
   516   Usually, this method is used as the initial proof step of manual
   517   termination proofs.
   518 
   519   \item @{method (HOL) "lexicographic_order"} attempts a fully
   520   automated termination proof by searching for a lexicographic
   521   combination of size measures on the arguments of the function. The
   522   method accepts the same arguments as the @{method auto} method,
   523   which it uses internally to prove local descents.  The @{syntax
   524   clasimpmod} modifiers are accepted (as for @{method auto}).
   525 
   526   In case of failure, extensive information is printed, which can help
   527   to analyse the situation (cf.\ \cite{isabelle-function}).
   528 
   529   \item @{method (HOL) "size_change"} also works on termination goals,
   530   using a variation of the size-change principle, together with a
   531   graph decomposition technique (see \cite{krauss_phd} for details).
   532   Three kinds of orders are used internally: @{text max}, @{text min},
   533   and @{text ms} (multiset), which is only available when the theory
   534   @{text Multiset} is loaded. When no order kinds are given, they are
   535   tried in order. The search for a termination proof uses SAT solving
   536   internally.
   537 
   538   For local descent proofs, the @{syntax clasimpmod} modifiers are
   539   accepted (as for @{method auto}).
   540 
   541   \end{description}
   542 *}
   543 
   544 
   545 subsection {* Functions with explicit partiality *}
   546 
   547 text {*
   548   \begin{matharray}{rcl}
   549     @{command_def (HOL) "partial_function"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   550     @{attribute_def (HOL) "partial_function_mono"} & : & @{text attribute} \\
   551   \end{matharray}
   552 
   553   @{rail "
   554     @@{command (HOL) partial_function} @{syntax target}?
   555       '(' @{syntax nameref} ')' @{syntax \"fixes\"} \\
   556       @'where' @{syntax thmdecl}? @{syntax prop}
   557   "}
   558 
   559   \begin{description}
   560 
   561   \item @{command (HOL) "partial_function"}~@{text "(mode)"} defines
   562   recursive functions based on fixpoints in complete partial
   563   orders. No termination proof is required from the user or
   564   constructed internally. Instead, the possibility of non-termination
   565   is modelled explicitly in the result type, which contains an
   566   explicit bottom element.
   567 
   568   Pattern matching and mutual recursion are currently not supported.
   569   Thus, the specification consists of a single function described by a
   570   single recursive equation.
   571 
   572   There are no fixed syntactic restrictions on the body of the
   573   function, but the induced functional must be provably monotonic
   574   wrt.\ the underlying order.  The monotonicitity proof is performed
   575   internally, and the definition is rejected when it fails. The proof
   576   can be influenced by declaring hints using the
   577   @{attribute (HOL) partial_function_mono} attribute.
   578 
   579   The mandatory @{text mode} argument specifies the mode of operation
   580   of the command, which directly corresponds to a complete partial
   581   order on the result type. By default, the following modes are
   582   defined:
   583 
   584   \begin{description}
   585   \item @{text option} defines functions that map into the @{type
   586   option} type. Here, the value @{term None} is used to model a
   587   non-terminating computation. Monotonicity requires that if @{term
   588   None} is returned by a recursive call, then the overall result
   589   must also be @{term None}. This is best achieved through the use of
   590   the monadic operator @{const "Option.bind"}.
   591 
   592   \item @{text tailrec} defines functions with an arbitrary result
   593   type and uses the slightly degenerated partial order where @{term
   594   "undefined"} is the bottom element.  Now, monotonicity requires that
   595   if @{term undefined} is returned by a recursive call, then the
   596   overall result must also be @{term undefined}. In practice, this is
   597   only satisfied when each recursive call is a tail call, whose result
   598   is directly returned. Thus, this mode of operation allows the
   599   definition of arbitrary tail-recursive functions.
   600   \end{description}
   601 
   602   Experienced users may define new modes by instantiating the locale
   603   @{const "partial_function_definitions"} appropriately.
   604 
   605   \item @{attribute (HOL) partial_function_mono} declares rules for
   606   use in the internal monononicity proofs of partial function
   607   definitions.
   608 
   609   \end{description}
   610 
   611 *}
   612 
   613 
   614 subsection {* Old-style recursive function definitions (TFL) *}
   615 
   616 text {*
   617   The old TFL commands @{command (HOL) "recdef"} and @{command (HOL)
   618   "recdef_tc"} for defining recursive are mostly obsolete; @{command
   619   (HOL) "function"} or @{command (HOL) "fun"} should be used instead.
   620 
   621   \begin{matharray}{rcl}
   622     @{command_def (HOL) "recdef"} & : & @{text "theory \<rightarrow> theory)"} \\
   623     @{command_def (HOL) "recdef_tc"}@{text "\<^sup>*"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
   624   \end{matharray}
   625 
   626   @{rail "
   627     @@{command (HOL) recdef} ('(' @'permissive' ')')? \\
   628       @{syntax name} @{syntax term} (@{syntax prop} +) hints?
   629     ;
   630     recdeftc @{syntax thmdecl}? tc
   631     ;
   632     hints: '(' @'hints' ( recdefmod * ) ')'
   633     ;
   634     recdefmod: (('recdef_simp' | 'recdef_cong' | 'recdef_wf')
   635       (() | 'add' | 'del') ':' @{syntax thmrefs}) | @{syntax clasimpmod}
   636     ;
   637     tc: @{syntax nameref} ('(' @{syntax nat} ')')?
   638   "}
   639 
   640   \begin{description}
   641 
   642   \item @{command (HOL) "recdef"} defines general well-founded
   643   recursive functions (using the TFL package), see also
   644   \cite{isabelle-HOL}.  The ``@{text "(permissive)"}'' option tells
   645   TFL to recover from failed proof attempts, returning unfinished
   646   results.  The @{text recdef_simp}, @{text recdef_cong}, and @{text
   647   recdef_wf} hints refer to auxiliary rules to be used in the internal
   648   automated proof process of TFL.  Additional @{syntax clasimpmod}
   649   declarations may be given to tune the context of the Simplifier
   650   (cf.\ \secref{sec:simplifier}) and Classical reasoner (cf.\
   651   \secref{sec:classical}).
   652 
   653   \item @{command (HOL) "recdef_tc"}~@{text "c (i)"} recommences the
   654   proof for leftover termination condition number @{text i} (default
   655   1) as generated by a @{command (HOL) "recdef"} definition of
   656   constant @{text c}.
   657 
   658   Note that in most cases, @{command (HOL) "recdef"} is able to finish
   659   its internal proofs without manual intervention.
   660 
   661   \end{description}
   662 
   663   \medskip Hints for @{command (HOL) "recdef"} may be also declared
   664   globally, using the following attributes.
   665 
   666   \begin{matharray}{rcl}
   667     @{attribute_def (HOL) recdef_simp} & : & @{text attribute} \\
   668     @{attribute_def (HOL) recdef_cong} & : & @{text attribute} \\
   669     @{attribute_def (HOL) recdef_wf} & : & @{text attribute} \\
   670   \end{matharray}
   671 
   672   @{rail "
   673     (@@{attribute (HOL) recdef_simp} | @@{attribute (HOL) recdef_cong} |
   674       @@{attribute (HOL) recdef_wf}) (() | 'add' | 'del')
   675   "}
   676 *}
   677 
   678 
   679 section {* Datatypes \label{sec:hol-datatype} *}
   680 
   681 text {*
   682   \begin{matharray}{rcl}
   683     @{command_def (HOL) "datatype"} & : & @{text "theory \<rightarrow> theory"} \\
   684     @{command_def (HOL) "rep_datatype"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
   685   \end{matharray}
   686 
   687   @{rail "
   688     @@{command (HOL) datatype} (spec + @'and')
   689     ;
   690     @@{command (HOL) rep_datatype} ('(' (@{syntax name} +) ')')? (@{syntax term} +)
   691     ;
   692 
   693     spec: @{syntax parname}? @{syntax typespec} @{syntax mixfix}? '=' (cons + '|')
   694     ;
   695     cons: @{syntax name} (@{syntax type} * ) @{syntax mixfix}?
   696   "}
   697 
   698   \begin{description}
   699 
   700   \item @{command (HOL) "datatype"} defines inductive datatypes in
   701   HOL.
   702 
   703   \item @{command (HOL) "rep_datatype"} represents existing types as
   704   datatypes.
   705 
   706   For foundational reasons, some basic types such as @{typ nat}, @{typ
   707   "'a \<times> 'b"}, @{typ "'a + 'b"}, @{typ bool} and @{typ unit} are
   708   introduced by more primitive means using @{command_ref typedef}.  To
   709   recover the rich infrastructure of @{command datatype} (e.g.\ rules
   710   for @{method cases} and @{method induct} and the primitive recursion
   711   combinators), such types may be represented as actual datatypes
   712   later.  This is done by specifying the constructors of the desired
   713   type, and giving a proof of the induction rule, distinctness and
   714   injectivity of constructors.
   715 
   716   For example, see @{file "~~/src/HOL/Sum_Type.thy"} for the
   717   representation of the primitive sum type as fully-featured datatype.
   718 
   719   \end{description}
   720 
   721   The generated rules for @{method induct} and @{method cases} provide
   722   case names according to the given constructors, while parameters are
   723   named after the types (see also \secref{sec:cases-induct}).
   724 
   725   See \cite{isabelle-HOL} for more details on datatypes, but beware of
   726   the old-style theory syntax being used there!  Apart from proper
   727   proof methods for case-analysis and induction, there are also
   728   emulations of ML tactics @{method (HOL) case_tac} and @{method (HOL)
   729   induct_tac} available, see \secref{sec:hol-induct-tac}; these admit
   730   to refer directly to the internal structure of subgoals (including
   731   internally bound parameters).
   732 *}
   733 
   734 
   735 subsubsection {* Examples *}
   736 
   737 text {* We define a type of finite sequences, with slightly different
   738   names than the existing @{typ "'a list"} that is already in @{theory
   739   Main}: *}
   740 
   741 datatype 'a seq = Empty | Seq 'a "'a seq"
   742 
   743 text {* We can now prove some simple lemma by structural induction: *}
   744 
   745 lemma "Seq x xs \<noteq> xs"
   746 proof (induct xs arbitrary: x)
   747   case Empty
   748   txt {* This case can be proved using the simplifier: the freeness
   749     properties of the datatype are already declared as @{attribute
   750     simp} rules. *}
   751   show "Seq x Empty \<noteq> Empty"
   752     by simp
   753 next
   754   case (Seq y ys)
   755   txt {* The step case is proved similarly. *}
   756   show "Seq x (Seq y ys) \<noteq> Seq y ys"
   757     using `Seq y ys \<noteq> ys` by simp
   758 qed
   759 
   760 text {* Here is a more succinct version of the same proof: *}
   761 
   762 lemma "Seq x xs \<noteq> xs"
   763   by (induct xs arbitrary: x) simp_all
   764 
   765 
   766 section {* Records \label{sec:hol-record} *}
   767 
   768 text {*
   769   In principle, records merely generalize the concept of tuples, where
   770   components may be addressed by labels instead of just position.  The
   771   logical infrastructure of records in Isabelle/HOL is slightly more
   772   advanced, though, supporting truly extensible record schemes.  This
   773   admits operations that are polymorphic with respect to record
   774   extension, yielding ``object-oriented'' effects like (single)
   775   inheritance.  See also \cite{NaraschewskiW-TPHOLs98} for more
   776   details on object-oriented verification and record subtyping in HOL.
   777 *}
   778 
   779 
   780 subsection {* Basic concepts *}
   781 
   782 text {*
   783   Isabelle/HOL supports both \emph{fixed} and \emph{schematic} records
   784   at the level of terms and types.  The notation is as follows:
   785 
   786   \begin{center}
   787   \begin{tabular}{l|l|l}
   788     & record terms & record types \\ \hline
   789     fixed & @{text "\<lparr>x = a, y = b\<rparr>"} & @{text "\<lparr>x :: A, y :: B\<rparr>"} \\
   790     schematic & @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} &
   791       @{text "\<lparr>x :: A, y :: B, \<dots> :: M\<rparr>"} \\
   792   \end{tabular}
   793   \end{center}
   794 
   795   \noindent The ASCII representation of @{text "\<lparr>x = a\<rparr>"} is @{text
   796   "(| x = a |)"}.
   797 
   798   A fixed record @{text "\<lparr>x = a, y = b\<rparr>"} has field @{text x} of value
   799   @{text a} and field @{text y} of value @{text b}.  The corresponding
   800   type is @{text "\<lparr>x :: A, y :: B\<rparr>"}, assuming that @{text "a :: A"}
   801   and @{text "b :: B"}.
   802 
   803   A record scheme like @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} contains fields
   804   @{text x} and @{text y} as before, but also possibly further fields
   805   as indicated by the ``@{text "\<dots>"}'' notation (which is actually part
   806   of the syntax).  The improper field ``@{text "\<dots>"}'' of a record
   807   scheme is called the \emph{more part}.  Logically it is just a free
   808   variable, which is occasionally referred to as ``row variable'' in
   809   the literature.  The more part of a record scheme may be
   810   instantiated by zero or more further components.  For example, the
   811   previous scheme may get instantiated to @{text "\<lparr>x = a, y = b, z =
   812   c, \<dots> = m'\<rparr>"}, where @{text m'} refers to a different more part.
   813   Fixed records are special instances of record schemes, where
   814   ``@{text "\<dots>"}'' is properly terminated by the @{text "() :: unit"}
   815   element.  In fact, @{text "\<lparr>x = a, y = b\<rparr>"} is just an abbreviation
   816   for @{text "\<lparr>x = a, y = b, \<dots> = ()\<rparr>"}.
   817 
   818   \medskip Two key observations make extensible records in a simply
   819   typed language like HOL work out:
   820 
   821   \begin{enumerate}
   822 
   823   \item the more part is internalized, as a free term or type
   824   variable,
   825 
   826   \item field names are externalized, they cannot be accessed within
   827   the logic as first-class values.
   828 
   829   \end{enumerate}
   830 
   831   \medskip In Isabelle/HOL record types have to be defined explicitly,
   832   fixing their field names and types, and their (optional) parent
   833   record.  Afterwards, records may be formed using above syntax, while
   834   obeying the canonical order of fields as given by their declaration.
   835   The record package provides several standard operations like
   836   selectors and updates.  The common setup for various generic proof
   837   tools enable succinct reasoning patterns.  See also the Isabelle/HOL
   838   tutorial \cite{isabelle-hol-book} for further instructions on using
   839   records in practice.
   840 *}
   841 
   842 
   843 subsection {* Record specifications *}
   844 
   845 text {*
   846   \begin{matharray}{rcl}
   847     @{command_def (HOL) "record"} & : & @{text "theory \<rightarrow> theory"} \\
   848   \end{matharray}
   849 
   850   @{rail "
   851     @@{command (HOL) record} @{syntax typespec_sorts} '=' \\
   852       (@{syntax type} '+')? (@{syntax constdecl} +)
   853   "}
   854 
   855   \begin{description}
   856 
   857   \item @{command (HOL) "record"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t = \<tau> + c\<^sub>1 :: \<sigma>\<^sub>1
   858   \<dots> c\<^sub>n :: \<sigma>\<^sub>n"} defines extensible record type @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"},
   859   derived from the optional parent record @{text "\<tau>"} by adding new
   860   field components @{text "c\<^sub>i :: \<sigma>\<^sub>i"} etc.
   861 
   862   The type variables of @{text "\<tau>"} and @{text "\<sigma>\<^sub>i"} need to be
   863   covered by the (distinct) parameters @{text "\<alpha>\<^sub>1, \<dots>,
   864   \<alpha>\<^sub>m"}.  Type constructor @{text t} has to be new, while @{text
   865   \<tau>} needs to specify an instance of an existing record type.  At
   866   least one new field @{text "c\<^sub>i"} has to be specified.
   867   Basically, field names need to belong to a unique record.  This is
   868   not a real restriction in practice, since fields are qualified by
   869   the record name internally.
   870 
   871   The parent record specification @{text \<tau>} is optional; if omitted
   872   @{text t} becomes a root record.  The hierarchy of all records
   873   declared within a theory context forms a forest structure, i.e.\ a
   874   set of trees starting with a root record each.  There is no way to
   875   merge multiple parent records!
   876 
   877   For convenience, @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} is made a
   878   type abbreviation for the fixed record type @{text "\<lparr>c\<^sub>1 ::
   879   \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n\<rparr>"}, likewise is @{text
   880   "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m, \<zeta>) t_scheme"} made an abbreviation for
   881   @{text "\<lparr>c\<^sub>1 :: \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n, \<dots> ::
   882   \<zeta>\<rparr>"}.
   883 
   884   \end{description}
   885 *}
   886 
   887 
   888 subsection {* Record operations *}
   889 
   890 text {*
   891   Any record definition of the form presented above produces certain
   892   standard operations.  Selectors and updates are provided for any
   893   field, including the improper one ``@{text more}''.  There are also
   894   cumulative record constructor functions.  To simplify the
   895   presentation below, we assume for now that @{text "(\<alpha>\<^sub>1, \<dots>,
   896   \<alpha>\<^sub>m) t"} is a root record with fields @{text "c\<^sub>1 ::
   897   \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n"}.
   898 
   899   \medskip \textbf{Selectors} and \textbf{updates} are available for
   900   any field (including ``@{text more}''):
   901 
   902   \begin{matharray}{lll}
   903     @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
   904     @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\<sigma>\<^sub>i \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
   905   \end{matharray}
   906 
   907   There is special syntax for application of updates: @{text "r\<lparr>x :=
   908   a\<rparr>"} abbreviates term @{text "x_update a r"}.  Further notation for
   909   repeated updates is also available: @{text "r\<lparr>x := a\<rparr>\<lparr>y := b\<rparr>\<lparr>z :=
   910   c\<rparr>"} may be written @{text "r\<lparr>x := a, y := b, z := c\<rparr>"}.  Note that
   911   because of postfix notation the order of fields shown here is
   912   reverse than in the actual term.  Since repeated updates are just
   913   function applications, fields may be freely permuted in @{text "\<lparr>x
   914   := a, y := b, z := c\<rparr>"}, as far as logical equality is concerned.
   915   Thus commutativity of independent updates can be proven within the
   916   logic for any two fields, but not as a general theorem.
   917 
   918   \medskip The \textbf{make} operation provides a cumulative record
   919   constructor function:
   920 
   921   \begin{matharray}{lll}
   922     @{text "t.make"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
   923   \end{matharray}
   924 
   925   \medskip We now reconsider the case of non-root records, which are
   926   derived of some parent.  In general, the latter may depend on
   927   another parent as well, resulting in a list of \emph{ancestor
   928   records}.  Appending the lists of fields of all ancestors results in
   929   a certain field prefix.  The record package automatically takes care
   930   of this by lifting operations over this context of ancestor fields.
   931   Assuming that @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} has ancestor
   932   fields @{text "b\<^sub>1 :: \<rho>\<^sub>1, \<dots>, b\<^sub>k :: \<rho>\<^sub>k"},
   933   the above record operations will get the following types:
   934 
   935   \medskip
   936   \begin{tabular}{lll}
   937     @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
   938     @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\<sigma>\<^sub>i \<Rightarrow>
   939       \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow>
   940       \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
   941     @{text "t.make"} & @{text "::"} & @{text "\<rho>\<^sub>1 \<Rightarrow> \<dots> \<rho>\<^sub>k \<Rightarrow> \<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow>
   942       \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
   943   \end{tabular}
   944   \medskip
   945 
   946   \noindent Some further operations address the extension aspect of a
   947   derived record scheme specifically: @{text "t.fields"} produces a
   948   record fragment consisting of exactly the new fields introduced here
   949   (the result may serve as a more part elsewhere); @{text "t.extend"}
   950   takes a fixed record and adds a given more part; @{text
   951   "t.truncate"} restricts a record scheme to a fixed record.
   952 
   953   \medskip
   954   \begin{tabular}{lll}
   955     @{text "t.fields"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
   956     @{text "t.extend"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr> \<Rightarrow>
   957       \<zeta> \<Rightarrow> \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>"} \\
   958     @{text "t.truncate"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
   959   \end{tabular}
   960   \medskip
   961 
   962   \noindent Note that @{text "t.make"} and @{text "t.fields"} coincide
   963   for root records.
   964 *}
   965 
   966 
   967 subsection {* Derived rules and proof tools *}
   968 
   969 text {*
   970   The record package proves several results internally, declaring
   971   these facts to appropriate proof tools.  This enables users to
   972   reason about record structures quite conveniently.  Assume that
   973   @{text t} is a record type as specified above.
   974 
   975   \begin{enumerate}
   976 
   977   \item Standard conversions for selectors or updates applied to
   978   record constructor terms are made part of the default Simplifier
   979   context; thus proofs by reduction of basic operations merely require
   980   the @{method simp} method without further arguments.  These rules
   981   are available as @{text "t.simps"}, too.
   982 
   983   \item Selectors applied to updated records are automatically reduced
   984   by an internal simplification procedure, which is also part of the
   985   standard Simplifier setup.
   986 
   987   \item Inject equations of a form analogous to @{prop "(x, y) = (x',
   988   y') \<equiv> x = x' \<and> y = y'"} are declared to the Simplifier and Classical
   989   Reasoner as @{attribute iff} rules.  These rules are available as
   990   @{text "t.iffs"}.
   991 
   992   \item The introduction rule for record equality analogous to @{text
   993   "x r = x r' \<Longrightarrow> y r = y r' \<dots> \<Longrightarrow> r = r'"} is declared to the Simplifier,
   994   and as the basic rule context as ``@{attribute intro}@{text "?"}''.
   995   The rule is called @{text "t.equality"}.
   996 
   997   \item Representations of arbitrary record expressions as canonical
   998   constructor terms are provided both in @{method cases} and @{method
   999   induct} format (cf.\ the generic proof methods of the same name,
  1000   \secref{sec:cases-induct}).  Several variations are available, for
  1001   fixed records, record schemes, more parts etc.
  1002 
  1003   The generic proof methods are sufficiently smart to pick the most
  1004   sensible rule according to the type of the indicated record
  1005   expression: users just need to apply something like ``@{text "(cases
  1006   r)"}'' to a certain proof problem.
  1007 
  1008   \item The derived record operations @{text "t.make"}, @{text
  1009   "t.fields"}, @{text "t.extend"}, @{text "t.truncate"} are \emph{not}
  1010   treated automatically, but usually need to be expanded by hand,
  1011   using the collective fact @{text "t.defs"}.
  1012 
  1013   \end{enumerate}
  1014 *}
  1015 
  1016 
  1017 subsubsection {* Examples *}
  1018 
  1019 text {* See @{file "~~/src/HOL/ex/Records.thy"}, for example. *}
  1020 
  1021 
  1022 section {* Adhoc tuples *}
  1023 
  1024 text {*
  1025   \begin{matharray}{rcl}
  1026     @{attribute_def (HOL) split_format}@{text "\<^sup>*"} & : & @{text attribute} \\
  1027   \end{matharray}
  1028 
  1029   @{rail "
  1030     @@{attribute (HOL) split_format} ('(' 'complete' ')')?
  1031   "}
  1032 
  1033   \begin{description}
  1034 
  1035   \item @{attribute (HOL) split_format}\ @{text "(complete)"} causes
  1036   arguments in function applications to be represented canonically
  1037   according to their tuple type structure.
  1038 
  1039   Note that this operation tends to invent funny names for new local
  1040   parameters introduced.
  1041 
  1042   \end{description}
  1043 *}
  1044 
  1045 
  1046 section {* Typedef axiomatization \label{sec:hol-typedef} *}
  1047 
  1048 text {* A Gordon/HOL-style type definition is a certain axiom scheme
  1049   that identifies a new type with a subset of an existing type.  More
  1050   precisely, the new type is defined by exhibiting an existing type
  1051   @{text \<tau>}, a set @{text "A :: \<tau> set"}, and a theorem that proves
  1052   @{prop "\<exists>x. x \<in> A"}.  Thus @{text A} is a non-empty subset of @{text
  1053   \<tau>}, and the new type denotes this subset.  New functions are
  1054   postulated that establish an isomorphism between the new type and
  1055   the subset.  In general, the type @{text \<tau>} may involve type
  1056   variables @{text "\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n"} which means that the type definition
  1057   produces a type constructor @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t"} depending on
  1058   those type arguments.
  1059 
  1060   The axiomatization can be considered a ``definition'' in the sense
  1061   of the particular set-theoretic interpretation of HOL
  1062   \cite{pitts93}, where the universe of types is required to be
  1063   downwards-closed wrt.\ arbitrary non-empty subsets.  Thus genuinely
  1064   new types introduced by @{command "typedef"} stay within the range
  1065   of HOL models by construction.  Note that @{command_ref
  1066   type_synonym} from Isabelle/Pure merely introduces syntactic
  1067   abbreviations, without any logical significance.
  1068   
  1069   \begin{matharray}{rcl}
  1070     @{command_def (HOL) "typedef"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
  1071   \end{matharray}
  1072 
  1073   @{rail "
  1074     @@{command (HOL) typedef} alt_name? abs_type '=' rep_set
  1075     ;
  1076 
  1077     alt_name: '(' (@{syntax name} | @'open' | @'open' @{syntax name}) ')'
  1078     ;
  1079     abs_type: @{syntax typespec_sorts} @{syntax mixfix}?
  1080     ;
  1081     rep_set: @{syntax term} (@'morphisms' @{syntax name} @{syntax name})?
  1082   "}
  1083 
  1084   \begin{description}
  1085 
  1086   \item @{command (HOL) "typedef"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t = A"}
  1087   axiomatizes a type definition in the background theory of the
  1088   current context, depending on a non-emptiness result of the set
  1089   @{text A} that needs to be proven here.  The set @{text A} may
  1090   contain type variables @{text "\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n"} as specified on the LHS,
  1091   but no term variables.
  1092 
  1093   Even though a local theory specification, the newly introduced type
  1094   constructor cannot depend on parameters or assumptions of the
  1095   context: this is structurally impossible in HOL.  In contrast, the
  1096   non-emptiness proof may use local assumptions in unusual situations,
  1097   which could result in different interpretations in target contexts:
  1098   the meaning of the bijection between the representing set @{text A}
  1099   and the new type @{text t} may then change in different application
  1100   contexts.
  1101 
  1102   By default, @{command (HOL) "typedef"} defines both a type
  1103   constructor @{text t} for the new type, and a term constant @{text
  1104   t} for the representing set within the old type.  Use the ``@{text
  1105   "(open)"}'' option to suppress a separate constant definition
  1106   altogether.  The injection from type to set is called @{text Rep_t},
  1107   its inverse @{text Abs_t}, unless explicit @{keyword (HOL)
  1108   "morphisms"} specification provides alternative names.
  1109 
  1110   The core axiomatization uses the locale predicate @{const
  1111   type_definition} as defined in Isabelle/HOL.  Various basic
  1112   consequences of that are instantiated accordingly, re-using the
  1113   locale facts with names derived from the new type constructor.  Thus
  1114   the generic @{thm type_definition.Rep} is turned into the specific
  1115   @{text "Rep_t"}, for example.
  1116 
  1117   Theorems @{thm type_definition.Rep}, @{thm
  1118   type_definition.Rep_inverse}, and @{thm type_definition.Abs_inverse}
  1119   provide the most basic characterization as a corresponding
  1120   injection/surjection pair (in both directions).  The derived rules
  1121   @{thm type_definition.Rep_inject} and @{thm
  1122   type_definition.Abs_inject} provide a more convenient version of
  1123   injectivity, suitable for automated proof tools (e.g.\ in
  1124   declarations involving @{attribute simp} or @{attribute iff}).
  1125   Furthermore, the rules @{thm type_definition.Rep_cases}~/ @{thm
  1126   type_definition.Rep_induct}, and @{thm type_definition.Abs_cases}~/
  1127   @{thm type_definition.Abs_induct} provide alternative views on
  1128   surjectivity.  These rules are already declared as set or type rules
  1129   for the generic @{method cases} and @{method induct} methods,
  1130   respectively.
  1131 
  1132   An alternative name for the set definition (and other derived
  1133   entities) may be specified in parentheses; the default is to use
  1134   @{text t} directly.
  1135 
  1136   \end{description}
  1137 
  1138   \begin{warn}
  1139   If you introduce a new type axiomatically, i.e.\ via @{command_ref
  1140   typedecl} and @{command_ref axiomatization}, the minimum requirement
  1141   is that it has a non-empty model, to avoid immediate collapse of the
  1142   HOL logic.  Moreover, one needs to demonstrate that the
  1143   interpretation of such free-form axiomatizations can coexist with
  1144   that of the regular @{command_def typedef} scheme, and any extension
  1145   that other people might have introduced elsewhere (e.g.\ in HOLCF
  1146   \cite{MuellerNvOS99}).
  1147   \end{warn}
  1148 *}
  1149 
  1150 subsubsection {* Examples *}
  1151 
  1152 text {* Type definitions permit the introduction of abstract data
  1153   types in a safe way, namely by providing models based on already
  1154   existing types.  Given some abstract axiomatic description @{text P}
  1155   of a type, this involves two steps:
  1156 
  1157   \begin{enumerate}
  1158 
  1159   \item Find an appropriate type @{text \<tau>} and subset @{text A} which
  1160   has the desired properties @{text P}, and make a type definition
  1161   based on this representation.
  1162 
  1163   \item Prove that @{text P} holds for @{text \<tau>} by lifting @{text P}
  1164   from the representation.
  1165 
  1166   \end{enumerate}
  1167 
  1168   You can later forget about the representation and work solely in
  1169   terms of the abstract properties @{text P}.
  1170 
  1171   \medskip The following trivial example pulls a three-element type
  1172   into existence within the formal logical environment of HOL. *}
  1173 
  1174 typedef three = "{(True, True), (True, False), (False, True)}"
  1175   by blast
  1176 
  1177 definition "One = Abs_three (True, True)"
  1178 definition "Two = Abs_three (True, False)"
  1179 definition "Three = Abs_three (False, True)"
  1180 
  1181 lemma three_distinct: "One \<noteq> Two"  "One \<noteq> Three"  "Two \<noteq> Three"
  1182   by (simp_all add: One_def Two_def Three_def Abs_three_inject three_def)
  1183 
  1184 lemma three_cases:
  1185   fixes x :: three obtains "x = One" | "x = Two" | "x = Three"
  1186   by (cases x) (auto simp: One_def Two_def Three_def Abs_three_inject three_def)
  1187 
  1188 text {* Note that such trivial constructions are better done with
  1189   derived specification mechanisms such as @{command datatype}: *}
  1190 
  1191 datatype three' = One' | Two' | Three'
  1192 
  1193 text {* This avoids re-doing basic definitions and proofs from the
  1194   primitive @{command typedef} above. *}
  1195 
  1196 
  1197 section {* Functorial structure of types *}
  1198 
  1199 text {*
  1200   \begin{matharray}{rcl}
  1201     @{command_def (HOL) "enriched_type"} & : & @{text "local_theory \<rightarrow> proof(prove)"}
  1202   \end{matharray}
  1203 
  1204   @{rail "
  1205     @@{command (HOL) enriched_type} (@{syntax name} ':')? @{syntax term}
  1206     ;
  1207   "}
  1208 
  1209   \begin{description}
  1210 
  1211   \item @{command (HOL) "enriched_type"}~@{text "prefix: m"} allows to
  1212   prove and register properties about the functorial structure of type
  1213   constructors.  These properties then can be used by other packages
  1214   to deal with those type constructors in certain type constructions.
  1215   Characteristic theorems are noted in the current local theory.  By
  1216   default, they are prefixed with the base name of the type
  1217   constructor, an explicit prefix can be given alternatively.
  1218 
  1219   The given term @{text "m"} is considered as \emph{mapper} for the
  1220   corresponding type constructor and must conform to the following
  1221   type pattern:
  1222 
  1223   \begin{matharray}{lll}
  1224     @{text "m"} & @{text "::"} &
  1225       @{text "\<sigma>\<^isub>1 \<Rightarrow> \<dots> \<sigma>\<^isub>k \<Rightarrow> (\<^vec>\<alpha>\<^isub>n) t \<Rightarrow> (\<^vec>\<beta>\<^isub>n) t"} \\
  1226   \end{matharray}
  1227 
  1228   \noindent where @{text t} is the type constructor, @{text
  1229   "\<^vec>\<alpha>\<^isub>n"} and @{text "\<^vec>\<beta>\<^isub>n"} are distinct
  1230   type variables free in the local theory and @{text "\<sigma>\<^isub>1"},
  1231   \ldots, @{text "\<sigma>\<^isub>k"} is a subsequence of @{text "\<alpha>\<^isub>1 \<Rightarrow>
  1232   \<beta>\<^isub>1"}, @{text "\<beta>\<^isub>1 \<Rightarrow> \<alpha>\<^isub>1"}, \ldots,
  1233   @{text "\<alpha>\<^isub>n \<Rightarrow> \<beta>\<^isub>n"}, @{text "\<beta>\<^isub>n \<Rightarrow>
  1234   \<alpha>\<^isub>n"}.
  1235 
  1236   \end{description}
  1237 *}
  1238 
  1239 
  1240 section {* Arithmetic proof support *}
  1241 
  1242 text {*
  1243   \begin{matharray}{rcl}
  1244     @{method_def (HOL) arith} & : & @{text method} \\
  1245     @{attribute_def (HOL) arith} & : & @{text attribute} \\
  1246     @{attribute_def (HOL) arith_split} & : & @{text attribute} \\
  1247   \end{matharray}
  1248 
  1249   The @{method (HOL) arith} method decides linear arithmetic problems
  1250   (on types @{text nat}, @{text int}, @{text real}).  Any current
  1251   facts are inserted into the goal before running the procedure.
  1252 
  1253   The @{attribute (HOL) arith} attribute declares facts that are
  1254   always supplied to the arithmetic provers implicitly.
  1255 
  1256   The @{attribute (HOL) arith_split} attribute declares case split
  1257   rules to be expanded before @{method (HOL) arith} is invoked.
  1258 
  1259   Note that a simpler (but faster) arithmetic prover is
  1260   already invoked by the Simplifier.
  1261 *}
  1262 
  1263 
  1264 section {* Intuitionistic proof search *}
  1265 
  1266 text {*
  1267   \begin{matharray}{rcl}
  1268     @{method_def (HOL) iprover} & : & @{text method} \\
  1269   \end{matharray}
  1270 
  1271   @{rail "
  1272     @@{method (HOL) iprover} ( @{syntax rulemod} * )
  1273   "}
  1274 
  1275   The @{method (HOL) iprover} method performs intuitionistic proof
  1276   search, depending on specifically declared rules from the context,
  1277   or given as explicit arguments.  Chained facts are inserted into the
  1278   goal before commencing proof search.
  1279 
  1280   Rules need to be classified as @{attribute (Pure) intro},
  1281   @{attribute (Pure) elim}, or @{attribute (Pure) dest}; here the
  1282   ``@{text "!"}'' indicator refers to ``safe'' rules, which may be
  1283   applied aggressively (without considering back-tracking later).
  1284   Rules declared with ``@{text "?"}'' are ignored in proof search (the
  1285   single-step @{method (Pure) rule} method still observes these).  An
  1286   explicit weight annotation may be given as well; otherwise the
  1287   number of rule premises will be taken into account here.
  1288 *}
  1289 
  1290 
  1291 section {* Coherent Logic *}
  1292 
  1293 text {*
  1294   \begin{matharray}{rcl}
  1295     @{method_def (HOL) "coherent"} & : & @{text method} \\
  1296   \end{matharray}
  1297 
  1298   @{rail "
  1299     @@{method (HOL) coherent} @{syntax thmrefs}?
  1300   "}
  1301 
  1302   The @{method (HOL) coherent} method solves problems of
  1303   \emph{Coherent Logic} \cite{Bezem-Coquand:2005}, which covers
  1304   applications in confluence theory, lattice theory and projective
  1305   geometry.  See @{file "~~/src/HOL/ex/Coherent.thy"} for some
  1306   examples.
  1307 *}
  1308 
  1309 
  1310 section {* Proving propositions *}
  1311 
  1312 text {*
  1313   In addition to the standard proof methods, a number of diagnosis
  1314   tools search for proofs and provide an Isar proof snippet on success.
  1315   These tools are available via the following commands.
  1316 
  1317   \begin{matharray}{rcl}
  1318     @{command_def (HOL) "solve_direct"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
  1319     @{command_def (HOL) "try"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
  1320     @{command_def (HOL) "try_methods"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
  1321     @{command_def (HOL) "sledgehammer"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
  1322     @{command_def (HOL) "sledgehammer_params"} & : & @{text "theory \<rightarrow> theory"}
  1323   \end{matharray}
  1324 
  1325   @{rail "
  1326     @@{command (HOL) try}
  1327     ;
  1328 
  1329     @@{command (HOL) try_methods} ( ( ( 'simp' | 'intro' | 'elim' | 'dest' ) ':' @{syntax thmrefs} ) + ) ?
  1330       @{syntax nat}?
  1331     ;
  1332 
  1333     @@{command (HOL) sledgehammer} ( '[' args ']' )? facts? @{syntax nat}?
  1334     ;
  1335 
  1336     @@{command (HOL) sledgehammer_params} ( ( '[' args ']' ) ? )
  1337     ;
  1338 
  1339     args: ( @{syntax name} '=' value + ',' )
  1340     ;
  1341 
  1342     facts: '(' ( ( ( ( 'add' | 'del' ) ':' ) ? @{syntax thmrefs} ) + ) ? ')'
  1343     ;
  1344   "} % FIXME check args "value"
  1345 
  1346   \begin{description}
  1347 
  1348   \item @{command (HOL) "solve_direct"} checks whether the current subgoals can
  1349     be solved directly by an existing theorem. Duplicate lemmas can be detected
  1350     in this way.
  1351 
  1352   \item @{command (HOL) "try_methods"} attempts to prove a subgoal using a combination
  1353     of standard proof methods (@{text auto}, @{text simp}, @{text blast}, etc.).
  1354     Additional facts supplied via @{text "simp:"}, @{text "intro:"},
  1355     @{text "elim:"}, and @{text "dest:"} are passed to the appropriate proof
  1356     methods.
  1357 
  1358   \item @{command (HOL) "sledgehammer"} attempts to prove a subgoal using external
  1359     automatic provers (resolution provers and SMT solvers). See the Sledgehammer
  1360     manual \cite{isabelle-sledgehammer} for details.
  1361 
  1362   \item @{command (HOL) "sledgehammer_params"} changes
  1363     @{command (HOL) "sledgehammer"} configuration options persistently.
  1364 
  1365   \end{description}
  1366 *}
  1367 
  1368 
  1369 section {* Checking and refuting propositions *}
  1370 
  1371 text {*
  1372   Identifying incorrect propositions usually involves evaluation of
  1373   particular assignments and systematic counterexample search.  This
  1374   is supported by the following commands.
  1375 
  1376   \begin{matharray}{rcl}
  1377     @{command_def (HOL) "value"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1378     @{command_def (HOL) "quickcheck"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
  1379     @{command_def (HOL) "refute"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
  1380     @{command_def (HOL) "nitpick"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow>"} \\
  1381     @{command_def (HOL) "quickcheck_params"} & : & @{text "theory \<rightarrow> theory"} \\
  1382     @{command_def (HOL) "refute_params"} & : & @{text "theory \<rightarrow> theory"} \\
  1383     @{command_def (HOL) "nitpick_params"} & : & @{text "theory \<rightarrow> theory"}
  1384   \end{matharray}
  1385 
  1386   @{rail "
  1387     @@{command (HOL) value} ( '[' name ']' )? modes? @{syntax term}
  1388     ;
  1389 
  1390     (@@{command (HOL) quickcheck} | @@{command (HOL) refute} | @@{command (HOL) nitpick})
  1391       ( '[' args ']' )? @{syntax nat}?
  1392     ;
  1393 
  1394     (@@{command (HOL) quickcheck_params} | @@{command (HOL) refute_params} |
  1395       @@{command (HOL) nitpick_params}) ( '[' args ']' )?
  1396     ;
  1397 
  1398     modes: '(' (@{syntax name} +) ')'
  1399     ;
  1400 
  1401     args: ( @{syntax name} '=' value + ',' )
  1402     ;
  1403   "} % FIXME check "value"
  1404 
  1405   \begin{description}
  1406 
  1407   \item @{command (HOL) "value"}~@{text t} evaluates and prints a
  1408     term; optionally @{text modes} can be specified, which are
  1409     appended to the current print mode; see \secref{sec:print-modes}.
  1410     Internally, the evaluation is performed by registered evaluators,
  1411     which are invoked sequentially until a result is returned.
  1412     Alternatively a specific evaluator can be selected using square
  1413     brackets; typical evaluators use the current set of code equations
  1414     to normalize and include @{text simp} for fully symbolic
  1415     evaluation using the simplifier, @{text nbe} for
  1416     \emph{normalization by evaluation} and \emph{code} for code
  1417     generation in SML.
  1418 
  1419   \item @{command (HOL) "quickcheck"} tests the current goal for
  1420     counterexamples using a series of assignments for its
  1421     free variables; by default the first subgoal is tested, an other
  1422     can be selected explicitly using an optional goal index.
  1423     Assignments can be chosen exhausting the search space upto a given
  1424     size or using a fixed number of random assignments in the search space.
  1425     By default, quickcheck uses exhaustive testing.
  1426     A number of configuration options are supported for
  1427     @{command (HOL) "quickcheck"}, notably:
  1428 
  1429     \begin{description}
  1430 
  1431     \item[@{text tester}] specifies how to explore the search space
  1432       (e.g. exhaustive or random).
  1433       An unknown configuration option is treated as an argument to tester,
  1434       making @{text "tester ="} optional.
  1435     \item[@{text size}] specifies the maximum size of the search space
  1436     for assignment values.
  1437 
  1438     \item[@{text eval}] takes a term or a list of terms and evaluates
  1439       these terms under the variable assignment found by quickcheck.
  1440 
  1441     \item[@{text iterations}] sets how many sets of assignments are
  1442     generated for each particular size.
  1443 
  1444     \item[@{text no_assms}] specifies whether assumptions in
  1445     structured proofs should be ignored.
  1446 
  1447     \item[@{text timeout}] sets the time limit in seconds.
  1448 
  1449     \item[@{text default_type}] sets the type(s) generally used to
  1450     instantiate type variables.
  1451 
  1452     \item[@{text report}] if set quickcheck reports how many tests
  1453     fulfilled the preconditions.
  1454 
  1455     \item[@{text quiet}] if not set quickcheck informs about the
  1456     current size for assignment values.
  1457 
  1458     \item[@{text expect}] can be used to check if the user's
  1459     expectation was met (@{text no_expectation}, @{text
  1460     no_counterexample}, or @{text counterexample}).
  1461 
  1462     \end{description}
  1463 
  1464     These option can be given within square brackets.
  1465 
  1466   \item @{command (HOL) "quickcheck_params"} changes
  1467     @{command (HOL) "quickcheck"} configuration options persistently.
  1468 
  1469   \item @{command (HOL) "refute"} tests the current goal for
  1470     counterexamples using a reduction to SAT. The following configuration
  1471     options are supported:
  1472 
  1473     \begin{description}
  1474 
  1475     \item[@{text minsize}] specifies the minimum size (cardinality) of the
  1476       models to search for.
  1477 
  1478     \item[@{text maxsize}] specifies the maximum size (cardinality) of the
  1479       models to search for. Nonpositive values mean $\infty$.
  1480 
  1481     \item[@{text maxvars}] specifies the maximum number of Boolean variables
  1482     to use when transforming the term into a propositional formula.
  1483     Nonpositive values mean $\infty$.
  1484 
  1485     \item[@{text satsolver}] specifies the SAT solver to use.
  1486 
  1487     \item[@{text no_assms}] specifies whether assumptions in
  1488     structured proofs should be ignored.
  1489 
  1490     \item[@{text maxtime}] sets the time limit in seconds.
  1491 
  1492     \item[@{text expect}] can be used to check if the user's
  1493     expectation was met (@{text genuine}, @{text potential},
  1494     @{text none}, or @{text unknown}).
  1495 
  1496     \end{description}
  1497 
  1498     These option can be given within square brackets.
  1499 
  1500   \item @{command (HOL) "refute_params"} changes
  1501     @{command (HOL) "refute"} configuration options persistently.
  1502 
  1503   \item @{command (HOL) "nitpick"} tests the current goal for counterexamples
  1504     using a reduction to first-order relational logic. See the Nitpick manual
  1505     \cite{isabelle-nitpick} for details.
  1506 
  1507   \item @{command (HOL) "nitpick_params"} changes
  1508     @{command (HOL) "nitpick"} configuration options persistently.
  1509 
  1510   \end{description}
  1511 *}
  1512 
  1513 
  1514 section {* Unstructured case analysis and induction \label{sec:hol-induct-tac} *}
  1515 
  1516 text {*
  1517   The following tools of Isabelle/HOL support cases analysis and
  1518   induction in unstructured tactic scripts; see also
  1519   \secref{sec:cases-induct} for proper Isar versions of similar ideas.
  1520 
  1521   \begin{matharray}{rcl}
  1522     @{method_def (HOL) case_tac}@{text "\<^sup>*"} & : & @{text method} \\
  1523     @{method_def (HOL) induct_tac}@{text "\<^sup>*"} & : & @{text method} \\
  1524     @{method_def (HOL) ind_cases}@{text "\<^sup>*"} & : & @{text method} \\
  1525     @{command_def (HOL) "inductive_cases"}@{text "\<^sup>*"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
  1526   \end{matharray}
  1527 
  1528   @{rail "
  1529     @@{method (HOL) case_tac} @{syntax goal_spec}? @{syntax term} rule?
  1530     ;
  1531     @@{method (HOL) induct_tac} @{syntax goal_spec}? (@{syntax insts} * @'and') rule?
  1532     ;
  1533     @@{method (HOL) ind_cases} (@{syntax prop}+) (@'for' (@{syntax name}+))?
  1534     ;
  1535     @@{command (HOL) inductive_cases} (@{syntax thmdecl}? (@{syntax prop}+) + @'and')
  1536     ;
  1537 
  1538     rule: 'rule' ':' @{syntax thmref}
  1539   "}
  1540 
  1541   \begin{description}
  1542 
  1543   \item @{method (HOL) case_tac} and @{method (HOL) induct_tac} admit
  1544   to reason about inductive types.  Rules are selected according to
  1545   the declarations by the @{attribute cases} and @{attribute induct}
  1546   attributes, cf.\ \secref{sec:cases-induct}.  The @{command (HOL)
  1547   datatype} package already takes care of this.
  1548 
  1549   These unstructured tactics feature both goal addressing and dynamic
  1550   instantiation.  Note that named rule cases are \emph{not} provided
  1551   as would be by the proper @{method cases} and @{method induct} proof
  1552   methods (see \secref{sec:cases-induct}).  Unlike the @{method
  1553   induct} method, @{method induct_tac} does not handle structured rule
  1554   statements, only the compact object-logic conclusion of the subgoal
  1555   being addressed.
  1556 
  1557   \item @{method (HOL) ind_cases} and @{command (HOL)
  1558   "inductive_cases"} provide an interface to the internal @{ML_text
  1559   mk_cases} operation.  Rules are simplified in an unrestricted
  1560   forward manner.
  1561 
  1562   While @{method (HOL) ind_cases} is a proof method to apply the
  1563   result immediately as elimination rules, @{command (HOL)
  1564   "inductive_cases"} provides case split theorems at the theory level
  1565   for later use.  The @{keyword "for"} argument of the @{method (HOL)
  1566   ind_cases} method allows to specify a list of variables that should
  1567   be generalized before applying the resulting rule.
  1568 
  1569   \end{description}
  1570 *}
  1571 
  1572 
  1573 section {* Executable code *}
  1574 
  1575 text {* For validation purposes, it is often useful to \emph{execute}
  1576   specifications.  In principle, execution could be simulated by
  1577   Isabelle's inference kernel, i.e. by a combination of resolution and
  1578   simplification.  Unfortunately, this approach is rather inefficient.
  1579   A more efficient way of executing specifications is to translate
  1580   them into a functional programming language such as ML.
  1581 
  1582   Isabelle provides two generic frameworks to support code generation
  1583   from executable specifications.  Isabelle/HOL instantiates these
  1584   mechanisms in a way that is amenable to end-user applications.
  1585 *}
  1586 
  1587 
  1588 subsection {* The new code generator (F. Haftmann) *}
  1589 
  1590 text {* This framework generates code from functional programs
  1591   (including overloading using type classes) to SML \cite{SML}, OCaml
  1592   \cite{OCaml}, Haskell \cite{haskell-revised-report} and Scala
  1593   \cite{scala-overview-tech-report}.  Conceptually, code generation is
  1594   split up in three steps: \emph{selection} of code theorems,
  1595   \emph{translation} into an abstract executable view and
  1596   \emph{serialization} to a specific \emph{target language}.
  1597   Inductive specifications can be executed using the predicate
  1598   compiler which operates within HOL.  See \cite{isabelle-codegen} for
  1599   an introduction.
  1600 
  1601   \begin{matharray}{rcl}
  1602     @{command_def (HOL) "export_code"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1603     @{attribute_def (HOL) code} & : & @{text attribute} \\
  1604     @{command_def (HOL) "code_abort"} & : & @{text "theory \<rightarrow> theory"} \\
  1605     @{command_def (HOL) "code_datatype"} & : & @{text "theory \<rightarrow> theory"} \\
  1606     @{command_def (HOL) "print_codesetup"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1607     @{attribute_def (HOL) code_inline} & : & @{text attribute} \\
  1608     @{attribute_def (HOL) code_post} & : & @{text attribute} \\
  1609     @{command_def (HOL) "print_codeproc"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1610     @{command_def (HOL) "code_thms"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1611     @{command_def (HOL) "code_deps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
  1612     @{command_def (HOL) "code_const"} & : & @{text "theory \<rightarrow> theory"} \\
  1613     @{command_def (HOL) "code_type"} & : & @{text "theory \<rightarrow> theory"} \\
  1614     @{command_def (HOL) "code_class"} & : & @{text "theory \<rightarrow> theory"} \\
  1615     @{command_def (HOL) "code_instance"} & : & @{text "theory \<rightarrow> theory"} \\
  1616     @{command_def (HOL) "code_reserved"} & : & @{text "theory \<rightarrow> theory"} \\
  1617     @{command_def (HOL) "code_monad"} & : & @{text "theory \<rightarrow> theory"} \\
  1618     @{command_def (HOL) "code_include"} & : & @{text "theory \<rightarrow> theory"} \\
  1619     @{command_def (HOL) "code_modulename"} & : & @{text "theory \<rightarrow> theory"} \\
  1620     @{command_def (HOL) "code_reflect"} & : & @{text "theory \<rightarrow> theory"}
  1621   \end{matharray}
  1622 
  1623   @{rail "
  1624     @@{command (HOL) export_code} ( constexpr + ) \\
  1625        ( ( @'in' target ( @'module_name' @{syntax string} ) ? \\
  1626         ( @'file' ( @{syntax string} | '-' ) ) ? ( '(' args ')' ) ?) + ) ?
  1627     ;
  1628 
  1629     const: @{syntax term}
  1630     ;
  1631 
  1632     constexpr: ( const | 'name._' | '_' )
  1633     ;
  1634 
  1635     typeconstructor: @{syntax nameref}
  1636     ;
  1637 
  1638     class: @{syntax nameref}
  1639     ;
  1640 
  1641     target: 'SML' | 'OCaml' | 'Haskell' | 'Scala'
  1642     ;
  1643 
  1644     @@{attribute (HOL) code} ( 'del' | 'abstype' | 'abstract' )?
  1645     ;
  1646 
  1647     @@{command (HOL) code_abort} ( const + )
  1648     ;
  1649 
  1650     @@{command (HOL) code_datatype} ( const + )
  1651     ;
  1652 
  1653     @@{attribute (HOL) code_inline} ( 'del' ) ?
  1654     ;
  1655 
  1656     @@{attribute (HOL) code_post} ( 'del' ) ?
  1657     ;
  1658 
  1659     @@{command (HOL) code_thms} ( constexpr + ) ?
  1660     ;
  1661 
  1662     @@{command (HOL) code_deps} ( constexpr + ) ?
  1663     ;
  1664 
  1665     @@{command (HOL) code_const} (const + @'and') \\
  1666       ( ( '(' target ( syntax ? + @'and' ) ')' ) + )
  1667     ;
  1668 
  1669     @@{command (HOL) code_type} (typeconstructor + @'and') \\
  1670       ( ( '(' target ( syntax ? + @'and' ) ')' ) + )
  1671     ;
  1672 
  1673     @@{command (HOL) code_class} (class + @'and') \\
  1674       ( ( '(' target \\ ( @{syntax string} ? + @'and' ) ')' ) + )
  1675     ;
  1676 
  1677     @@{command (HOL) code_instance} (( typeconstructor '::' class ) + @'and') \\
  1678       ( ( '(' target ( '-' ? + @'and' ) ')' ) + )
  1679     ;
  1680 
  1681     @@{command (HOL) code_reserved} target ( @{syntax string} + )
  1682     ;
  1683 
  1684     @@{command (HOL) code_monad} const const target
  1685     ;
  1686 
  1687     @@{command (HOL) code_include} target ( @{syntax string} ( @{syntax string} | '-') )
  1688     ;
  1689 
  1690     @@{command (HOL) code_modulename} target ( ( @{syntax string} @{syntax string} ) + )
  1691     ;
  1692 
  1693     @@{command (HOL) code_reflect} @{syntax string} \\
  1694       ( @'datatypes' ( @{syntax string} '=' ( '_' | ( @{syntax string} + '|' ) + @'and' ) ) ) ? \\
  1695       ( @'functions' ( @{syntax string} + ) ) ? ( @'file' @{syntax string} ) ?
  1696     ;
  1697 
  1698     syntax: @{syntax string} | ( @'infix' | @'infixl' | @'infixr' ) @{syntax nat} @{syntax string}
  1699   "}
  1700 
  1701   \begin{description}
  1702 
  1703   \item @{command (HOL) "export_code"} generates code for a given list
  1704   of constants in the specified target language(s).  If no
  1705   serialization instruction is given, only abstract code is generated
  1706   internally.
  1707 
  1708   Constants may be specified by giving them literally, referring to
  1709   all executable contants within a certain theory by giving @{text
  1710   "name.*"}, or referring to \emph{all} executable constants currently
  1711   available by giving @{text "*"}.
  1712 
  1713   By default, for each involved theory one corresponding name space
  1714   module is generated.  Alternativly, a module name may be specified
  1715   after the @{keyword "module_name"} keyword; then \emph{all} code is
  1716   placed in this module.
  1717 
  1718   For \emph{SML}, \emph{OCaml} and \emph{Scala} the file specification
  1719   refers to a single file; for \emph{Haskell}, it refers to a whole
  1720   directory, where code is generated in multiple files reflecting the
  1721   module hierarchy.  Omitting the file specification denotes standard
  1722   output.
  1723 
  1724   Serializers take an optional list of arguments in parentheses.  For
  1725   \emph{SML} and \emph{OCaml}, ``@{text no_signatures}`` omits
  1726   explicit module signatures.
  1727 
  1728   For \emph{Haskell} a module name prefix may be given using the
  1729   ``@{text "root:"}'' argument; ``@{text string_classes}'' adds a
  1730   ``@{verbatim "deriving (Read, Show)"}'' clause to each appropriate
  1731   datatype declaration.
  1732 
  1733   \item @{attribute (HOL) code} explicitly selects (or with option
  1734   ``@{text "del"}'' deselects) a code equation for code generation.
  1735   Usually packages introducing code equations provide a reasonable
  1736   default setup for selection.  Variants @{text "code abstype"} and
  1737   @{text "code abstract"} declare abstract datatype certificates or
  1738   code equations on abstract datatype representations respectively.
  1739 
  1740   \item @{command (HOL) "code_abort"} declares constants which are not
  1741   required to have a definition by means of code equations; if needed
  1742   these are implemented by program abort instead.
  1743 
  1744   \item @{command (HOL) "code_datatype"} specifies a constructor set
  1745   for a logical type.
  1746 
  1747   \item @{command (HOL) "print_codesetup"} gives an overview on
  1748   selected code equations and code generator datatypes.
  1749 
  1750   \item @{attribute (HOL) code_inline} declares (or with option
  1751   ``@{text "del"}'' removes) inlining theorems which are applied as
  1752   rewrite rules to any code equation during preprocessing.
  1753 
  1754   \item @{attribute (HOL) code_post} declares (or with option ``@{text
  1755   "del"}'' removes) theorems which are applied as rewrite rules to any
  1756   result of an evaluation.
  1757 
  1758   \item @{command (HOL) "print_codeproc"} prints the setup of the code
  1759   generator preprocessor.
  1760 
  1761   \item @{command (HOL) "code_thms"} prints a list of theorems
  1762   representing the corresponding program containing all given
  1763   constants after preprocessing.
  1764 
  1765   \item @{command (HOL) "code_deps"} visualizes dependencies of
  1766   theorems representing the corresponding program containing all given
  1767   constants after preprocessing.
  1768 
  1769   \item @{command (HOL) "code_const"} associates a list of constants
  1770   with target-specific serializations; omitting a serialization
  1771   deletes an existing serialization.
  1772 
  1773   \item @{command (HOL) "code_type"} associates a list of type
  1774   constructors with target-specific serializations; omitting a
  1775   serialization deletes an existing serialization.
  1776 
  1777   \item @{command (HOL) "code_class"} associates a list of classes
  1778   with target-specific class names; omitting a serialization deletes
  1779   an existing serialization.  This applies only to \emph{Haskell}.
  1780 
  1781   \item @{command (HOL) "code_instance"} declares a list of type
  1782   constructor / class instance relations as ``already present'' for a
  1783   given target.  Omitting a ``@{text "-"}'' deletes an existing
  1784   ``already present'' declaration.  This applies only to
  1785   \emph{Haskell}.
  1786 
  1787   \item @{command (HOL) "code_reserved"} declares a list of names as
  1788   reserved for a given target, preventing it to be shadowed by any
  1789   generated code.
  1790 
  1791   \item @{command (HOL) "code_monad"} provides an auxiliary mechanism
  1792   to generate monadic code for Haskell.
  1793 
  1794   \item @{command (HOL) "code_include"} adds arbitrary named content
  1795   (``include'') to generated code.  A ``@{text "-"}'' as last argument
  1796   will remove an already added ``include''.
  1797 
  1798   \item @{command (HOL) "code_modulename"} declares aliasings from one
  1799   module name onto another.
  1800 
  1801   \item @{command (HOL) "code_reflect"} without a ``@{text "file"}''
  1802   argument compiles code into the system runtime environment and
  1803   modifies the code generator setup that future invocations of system
  1804   runtime code generation referring to one of the ``@{text
  1805   "datatypes"}'' or ``@{text "functions"}'' entities use these precompiled
  1806   entities.  With a ``@{text "file"}'' argument, the corresponding code
  1807   is generated into that specified file without modifying the code
  1808   generator setup.
  1809 
  1810   \end{description}
  1811 *}
  1812 
  1813 
  1814 subsection {* The old code generator (S. Berghofer) *}
  1815 
  1816 text {* This framework generates code from both functional and
  1817   relational programs to SML, as explained below.
  1818 
  1819   \begin{matharray}{rcl}
  1820     @{command_def "code_module"} & : & @{text "theory \<rightarrow> theory"} \\
  1821     @{command_def "code_library"} & : & @{text "theory \<rightarrow> theory"} \\
  1822     @{command_def "consts_code"} & : & @{text "theory \<rightarrow> theory"} \\
  1823     @{command_def "types_code"} & : & @{text "theory \<rightarrow> theory"} \\
  1824     @{attribute_def code} & : & @{text attribute} \\
  1825   \end{matharray}
  1826 
  1827   @{rail "
  1828   ( @@{command code_module} | @@{command code_library} ) modespec? @{syntax name}? \\
  1829     ( @'file' name ) ? ( @'imports' ( @{syntax name} + ) ) ? \\
  1830     @'contains' ( ( @{syntax name} '=' @{syntax term} ) + | @{syntax term} + )
  1831   ;
  1832 
  1833   modespec: '(' ( @{syntax name} * ) ')'
  1834   ;
  1835 
  1836   @@{command (HOL) consts_code} (codespec +)
  1837   ;
  1838 
  1839   codespec: const template attachment ?
  1840   ;
  1841 
  1842   @@{command (HOL) types_code} (tycodespec +)
  1843   ;
  1844 
  1845   tycodespec: @{syntax name} template attachment ?
  1846   ;
  1847 
  1848   const: @{syntax term}
  1849   ;
  1850 
  1851   template: '(' @{syntax string} ')'
  1852   ;
  1853 
  1854   attachment: 'attach' modespec? '{' @{syntax text} '}'
  1855   ;
  1856 
  1857   @@{attribute code} name?
  1858   "}
  1859 *}
  1860 
  1861 
  1862 subsubsection {* Invoking the code generator *}
  1863 
  1864 text {* The code generator is invoked via the @{command code_module}
  1865   and @{command code_library} commands, which correspond to
  1866   \emph{incremental} and \emph{modular} code generation, respectively.
  1867 
  1868   \begin{description}
  1869 
  1870   \item [Modular] For each theory, an ML structure is generated,
  1871   containing the code generated from the constants defined in this
  1872   theory.
  1873 
  1874   \item [Incremental] All the generated code is emitted into the same
  1875   structure.  This structure may import code from previously generated
  1876   structures, which can be specified via @{keyword "imports"}.
  1877   Moreover, the generated structure may also be referred to in later
  1878   invocations of the code generator.
  1879 
  1880   \end{description}
  1881 
  1882   After the @{command code_module} and @{command code_library}
  1883   keywords, the user may specify an optional list of ``modes'' in
  1884   parentheses. These can be used to instruct the code generator to
  1885   emit additional code for special purposes, e.g.\ functions for
  1886   converting elements of generated datatypes to Isabelle terms, or
  1887   test data generators. The list of modes is followed by a module
  1888   name.  The module name is optional for modular code generation, but
  1889   must be specified for incremental code generation.
  1890 
  1891   The code can either be written to a file, in which case a file name
  1892   has to be specified after the @{keyword "file"} keyword, or be loaded
  1893   directly into Isabelle's ML environment. In the latter case, the
  1894   @{command ML} theory command can be used to inspect the results
  1895   interactively, for example.
  1896 
  1897   The terms from which to generate code can be specified after the
  1898   @{keyword "contains"} keyword, either as a list of bindings, or just
  1899   as a list of terms. In the latter case, the code generator just
  1900   produces code for all constants and types occuring in the term, but
  1901   does not bind the compiled terms to ML identifiers.
  1902 
  1903   Here is an example:
  1904 *}
  1905 
  1906 code_module Test
  1907 contains test = "foldl op + (0 :: int) [1, 2, 3, 4, 5]"
  1908 
  1909 text {* \noindent This binds the result of compiling the given term to
  1910   the ML identifier @{ML Test.test}.  *}
  1911 
  1912 ML {* @{assert} (Test.test = 15) *}
  1913 
  1914 
  1915 subsubsection {* Configuring the code generator *}
  1916 
  1917 text {* When generating code for a complex term, the code generator
  1918   recursively calls itself for all subterms.  When it arrives at a
  1919   constant, the default strategy of the code generator is to look up
  1920   its definition and try to generate code for it.  Constants which
  1921   have no definitions that are immediately executable, may be
  1922   associated with a piece of ML code manually using the @{command_ref
  1923   consts_code} command.  It takes a list whose elements consist of a
  1924   constant (given in usual term syntax -- an explicit type constraint
  1925   accounts for overloading), and a mixfix template describing the ML
  1926   code. The latter is very much the same as the mixfix templates used
  1927   when declaring new constants.  The most notable difference is that
  1928   terms may be included in the ML template using antiquotation
  1929   brackets @{verbatim "{"}@{verbatim "*"}~@{text "..."}~@{verbatim
  1930   "*"}@{verbatim "}"}.
  1931 
  1932   A similar mechanism is available for types: @{command_ref
  1933   types_code} associates type constructors with specific ML code.
  1934 
  1935   For example, the following declarations copied from @{file
  1936   "~~/src/HOL/Product_Type.thy"} describe how the product type of
  1937   Isabelle/HOL should be compiled to ML.  *}
  1938 
  1939 typedecl ('a, 'b) prod
  1940 consts Pair :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) prod"
  1941 
  1942 types_code prod  ("(_ */ _)")
  1943 consts_code Pair  ("(_,/ _)")
  1944 
  1945 text {* Sometimes, the code associated with a constant or type may
  1946   need to refer to auxiliary functions, which have to be emitted when
  1947   the constant is used. Code for such auxiliary functions can be
  1948   declared using @{keyword "attach"}. For example, the @{const wfrec}
  1949   function can be implemented as follows:
  1950 *}
  1951 
  1952 consts_code wfrec  ("\<module>wfrec?")  (* FIXME !? *)
  1953 attach {* fun wfrec f x = f (wfrec f) x *}
  1954 
  1955 text {* If the code containing a call to @{const wfrec} resides in an
  1956   ML structure different from the one containing the function
  1957   definition attached to @{const wfrec}, the name of the ML structure
  1958   (followed by a ``@{text "."}'')  is inserted in place of ``@{text
  1959   "\<module>"}'' in the above template.  The ``@{text "?"}''  means that
  1960   the code generator should ignore the first argument of @{const
  1961   wfrec}, i.e.\ the termination relation, which is usually not
  1962   executable.
  1963 
  1964   \medskip Another possibility of configuring the code generator is to
  1965   register theorems to be used for code generation. Theorems can be
  1966   registered via the @{attribute code} attribute. It takes an optional
  1967   name as an argument, which indicates the format of the
  1968   theorem. Currently supported formats are equations (this is the
  1969   default when no name is specified) and horn clauses (this is
  1970   indicated by the name \texttt{ind}). The left-hand sides of
  1971   equations may only contain constructors and distinct variables,
  1972   whereas horn clauses must have the same format as introduction rules
  1973   of inductive definitions.
  1974 
  1975   The following example specifies three equations from which to
  1976   generate code for @{term "op <"} on natural numbers (see also
  1977   @{"file" "~~/src/HOL/Nat.thy"}).  *}
  1978 
  1979 lemma [code]: "(Suc m < Suc n) = (m < n)"
  1980   and [code]: "((n::nat) < 0) = False"
  1981   and [code]: "(0 < Suc n) = True" by simp_all
  1982 
  1983 
  1984 subsubsection {* Specific HOL code generators *}
  1985 
  1986 text {* The basic code generator framework offered by Isabelle/Pure
  1987   has already been extended with additional code generators for
  1988   specific HOL constructs. These include datatypes, recursive
  1989   functions and inductive relations. The code generator for inductive
  1990   relations can handle expressions of the form @{text "(t\<^sub>1, \<dots>, t\<^sub>n) \<in>
  1991   r"}, where @{text "r"} is an inductively defined relation. If at
  1992   least one of the @{text "t\<^sub>i"} is a dummy pattern ``@{text "_"}'',
  1993   the above expression evaluates to a sequence of possible answers. If
  1994   all of the @{text "t\<^sub>i"} are proper terms, the expression evaluates
  1995   to a boolean value.
  1996 
  1997   The following example demonstrates this for beta-reduction on lambda
  1998   terms (see also @{"file" "~~/src/HOL/Proofs/Lambda/Lambda.thy"}).
  1999 *}
  2000 
  2001 datatype dB =
  2002     Var nat
  2003   | App dB dB  (infixl "\<degree>" 200)
  2004   | Abs dB
  2005 
  2006 primrec lift :: "dB \<Rightarrow> nat \<Rightarrow> dB"
  2007 where
  2008     "lift (Var i) k = (if i < k then Var i else Var (i + 1))"
  2009   | "lift (s \<degree> t) k = lift s k \<degree> lift t k"
  2010   | "lift (Abs s) k = Abs (lift s (k + 1))"
  2011 
  2012 primrec subst :: "dB \<Rightarrow> dB \<Rightarrow> nat \<Rightarrow> dB"  ("_[_'/_]" [300, 0, 0] 300)
  2013 where
  2014     "(Var i)[s/k] =
  2015       (if k < i then Var (i - 1) else if i = k then s else Var i)"
  2016   | "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
  2017   | "(Abs t)[s/k] = Abs (t[lift s 0 / k+1])"
  2018 
  2019 inductive beta :: "dB \<Rightarrow> dB \<Rightarrow> bool"  (infixl "\<rightarrow>\<^sub>\<beta>" 50)
  2020 where
  2021     beta: "Abs s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
  2022   | appL: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
  2023   | appR: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
  2024   | abs: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> Abs s \<rightarrow>\<^sub>\<beta> Abs t"
  2025 
  2026 code_module Test
  2027 contains
  2028   test1 = "Abs (Var 0) \<degree> Var 0 \<rightarrow>\<^sub>\<beta> Var 0"
  2029   test2 = "Abs (Abs (Var 0 \<degree> Var 0) \<degree> (Abs (Var 0) \<degree> Var 0)) \<rightarrow>\<^sub>\<beta> _"
  2030 
  2031 text {*
  2032   In the above example, @{ML Test.test1} evaluates to a boolean,
  2033   whereas @{ML Test.test2} is a lazy sequence whose elements can be
  2034   inspected separately.
  2035 *}
  2036 
  2037 ML {* @{assert} Test.test1 *}
  2038 ML {* val results = DSeq.list_of Test.test2 *}
  2039 ML {* @{assert} (length results = 2) *}
  2040 
  2041 text {*
  2042   \medskip The theory underlying the HOL code generator is described
  2043   more detailed in \cite{Berghofer-Nipkow:2002}. More examples that
  2044   illustrate the usage of the code generator can be found e.g.\ in
  2045   @{"file" "~~/src/HOL/MicroJava/J/JListExample.thy"} and @{"file"
  2046   "~~/src/HOL/MicroJava/JVM/JVMListExample.thy"}.
  2047 *}
  2048 
  2049 
  2050 section {* Definition by specification \label{sec:hol-specification} *}
  2051 
  2052 text {*
  2053   \begin{matharray}{rcl}
  2054     @{command_def (HOL) "specification"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
  2055     @{command_def (HOL) "ax_specification"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
  2056   \end{matharray}
  2057 
  2058   @{rail "
  2059   (@@{command (HOL) specification} | @@{command (HOL) ax_specification})
  2060     '(' (decl +) ')' \\ (@{syntax thmdecl}? @{syntax prop} +)
  2061   ;
  2062   decl: ((@{syntax name} ':')? @{syntax term} '(' @'overloaded' ')'?)
  2063   "}
  2064 
  2065   \begin{description}
  2066 
  2067   \item @{command (HOL) "specification"}~@{text "decls \<phi>"} sets up a
  2068   goal stating the existence of terms with the properties specified to
  2069   hold for the constants given in @{text decls}.  After finishing the
  2070   proof, the theory will be augmented with definitions for the given
  2071   constants, as well as with theorems stating the properties for these
  2072   constants.
  2073 
  2074   \item @{command (HOL) "ax_specification"}~@{text "decls \<phi>"} sets up
  2075   a goal stating the existence of terms with the properties specified
  2076   to hold for the constants given in @{text decls}.  After finishing
  2077   the proof, the theory will be augmented with axioms expressing the
  2078   properties given in the first place.
  2079 
  2080   \item @{text decl} declares a constant to be defined by the
  2081   specification given.  The definition for the constant @{text c} is
  2082   bound to the name @{text c_def} unless a theorem name is given in
  2083   the declaration.  Overloaded constants should be declared as such.
  2084 
  2085   \end{description}
  2086 
  2087   Whether to use @{command (HOL) "specification"} or @{command (HOL)
  2088   "ax_specification"} is to some extent a matter of style.  @{command
  2089   (HOL) "specification"} introduces no new axioms, and so by
  2090   construction cannot introduce inconsistencies, whereas @{command
  2091   (HOL) "ax_specification"} does introduce axioms, but only after the
  2092   user has explicitly proven it to be safe.  A practical issue must be
  2093   considered, though: After introducing two constants with the same
  2094   properties using @{command (HOL) "specification"}, one can prove
  2095   that the two constants are, in fact, equal.  If this might be a
  2096   problem, one should use @{command (HOL) "ax_specification"}.
  2097 *}
  2098 
  2099 end