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