products/Sources/formale Sprachen/Isabelle/Doc/Isar_Ref image not shown  

Quellcode-Bibliothek

© Kompilation durch diese Firma

[Weder Korrektheit noch Funktionsfähigkeit der Software werden zugesichert.]

Datei: HOL_Specific.thy   Sprache: Isabelle

Original von: Isabelle©

(*:maxLineLen=78:*)

theory HOL_Specific
  imports
    Main
    "HOL-Library.Old_Datatype"
    "HOL-Library.Old_Recdef"
    "HOL-Library.Adhoc_Overloading"
    "HOL-Library.Dlist"
    "HOL-Library.FSet"
    Base
begin


chapter \<open>Higher-Order Logic\<close>

text \<open>
  Isabelle/HOL is based on Higher-Order Logic, a polymorphic version of
  Church's Simple Theory of Types. HOL can be best understood as a
  simply-typed version of classical set theory. The logic was first
  implemented in Gordon's HOL system @{cite "mgordon-hol"}. It extends
  Church's original logic @{cite "church40"} by explicit type variables (naive
  polymorphism) and a sound axiomatization scheme for new types based on
  subsets of existing types.

  Andrews's book @{cite andrews86} is a full description of the original
  Church-style higher-order logic, with proofs of correctness and completeness
  wrt.\ certain set-theoretic interpretations. The particular extensions of
  Gordon-style HOL are explained semantically in two chapters of the 1993 HOL
  book @{cite pitts93}.

  Experience with HOL over decades has demonstrated that higher-order logic is
  widely applicable in many areas of mathematics and computer science. In a
  sense, Higher-Order Logic is simpler than First-Order Logic, because there
  are fewer restrictions and special cases. Note that HOL is \<^emph>\<open>weaker\<close> than
  FOL with axioms for ZF set theory, which is traditionally considered the
  standard foundation of regular mathematics, but for most applications this
  does not matter. If you prefer ML to Lisp, you will probably prefer HOL to
  ZF.

  \<^medskip> The syntax of HOL follows \<open>\<lambda>\<close>-calculus and functional programming.
  Function application is curried. To apply the function \<open>f\<close> of type \<open>\<tau>\<^sub>1 \<Rightarrow>
  \<tau>\<^sub>2 \<Rightarrow> \<tau>\<^sub>3\<close> to the arguments \<open>a\<close> and \<open>b\<close> in HOL, you simply write \<open>f a b\<close> (as
  in ML or Haskell). There is no ``apply'' operator; the existing application
  of the Pure \<open>\<lambda>\<close>-calculus is re-used. Note that in HOL \<open>f (a, b)\<close> means ``\<open>f\<close>
  applied to the pair \<open>(a, b)\<close> (which is notation for \<open>Pair a b\<close>). The latter
  typically introduces extra formal efforts that can be avoided by currying
  functions by default. Explicit tuples are as infrequent in HOL
  formalizations as in good ML or Haskell programs.

  \<^medskip> Isabelle/HOL has a distinct feel, compared to other object-logics like
  Isabelle/ZF. It identifies object-level types with meta-level types, taking
  advantage of the default type-inference mechanism of Isabelle/Pure. HOL
  fully identifies object-level functions with meta-level functions, with
  native abstraction and application.

  These identifications allow Isabelle to support HOL particularly nicely, but
  they also mean that HOL requires some sophistication from the user. In
  particular, an understanding of Hindley-Milner type-inference with
  type-classes, which are both used extensively in the standard libraries and
  applications.
\<close>


chapter \<open>Derived specification elements\<close>

section \<open>Inductive and coinductive definitions \label{sec:hol-inductive}\<close>

text \<open>
  \begin{matharray}{rcl}
    @{command_def (HOL) "inductive"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
    @{command_def (HOL) "inductive_set"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
    @{command_def (HOL) "coinductive"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
    @{command_def (HOL) "coinductive_set"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
    @{command_def "print_inductives"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\
    @{attribute_def (HOL) mono} & : & \<open>attribute\<close> \\
  \end{matharray}

  An \<^emph>\<open>inductive definition\<close> specifies the least predicate or set \<open>R\<close> closed
  under given rules: applying a rule to elements of \<open>R\<close> yields a result within
  \<open>R\<close>. For example, a structural operational semantics is an inductive
  definition of an evaluation relation.

  Dually, a \<^emph>\<open>coinductive definition\<close> specifies the greatest predicate or set
  \<open>R\<close> that is consistent with given rules: every element of \<open>R\<close> can be seen as
  arising by applying a rule to elements of \<open>R\<close>. An important example is using
  bisimulation relations to formalise equivalence of processes and infinite
  data structures.

  Both inductive and coinductive definitions are based on the Knaster-Tarski
  fixed-point theorem for complete lattices. The collection of introduction
  rules given by the user determines a functor on subsets of set-theoretic
  relations. The required monotonicity of the recursion scheme is proven as a
  prerequisite to the fixed-point definition and the resulting consequences.
  This works by pushing inclusion through logical connectives and any other
  operator that might be wrapped around recursive occurrences of the defined
  relation: there must be a monotonicity theorem of the form \<open>A \<le> B \<Longrightarrow> \<M> A \<le> \<M>
  B\<close>, for each premise \<open>\<M> R t\<close> in an introduction rule. The default rule
  declarations of Isabelle/HOL already take care of most common situations.

  \<^rail>\<open>
    (@@{command (HOL) inductive} | @@{command (HOL) inductive_set} |
      @@{command (HOL) coinductive} | @@{command (HOL) coinductive_set})
      @{syntax vars} @{syntax for_fixes} \<newline>
      (@'where' @{syntax multi_specs})? (@'monos' @{syntax thms})?
    ;
    @@{command print_inductives} ('!'?)
    ;
    @@{attribute (HOL) mono} (() | 'add' | 'del')
  \<close>

  \<^descr> @{command (HOL) "inductive"} and @{command (HOL) "coinductive"} define
  (co)inductive predicates from the introduction rules.

  The propositions given as \<open>clauses\<close> in the @{keyword "where"} part are
  either rules of the usual \<open>\<And>/\<Longrightarrow>\<close> format (with arbitrary nesting), or
  equalities using \<open>\<equiv>\<close>. The latter specifies extra-logical abbreviations in
  the sense of @{command_ref abbreviation}. Introducing abstract syntax
  simultaneously with the actual introduction rules is occasionally useful for
  complex specifications.

  The optional @{keyword "for"} part contains a list of parameters of the
  (co)inductive predicates that remain fixed throughout the definitionin
  contrast to arguments of the relation that may vary in each occurrence
  within the given \<open>clauses\<close>.

  The optional @{keyword "monos"declaration contains additional
  \<^emph>\<open>monotonicity theorems\<close>, which are required for each operator applied to a
  recursive set in the introduction rules.

  \<^descr> @{command (HOL) "inductive_set"} and @{command (HOL) "coinductive_set"}
  are wrappers for to the previous commands for native HOL predicates. This
  allows to define (co)inductive sets, where multiple arguments are simulated
  via tuples.

  \<^descr> @{command "print_inductives"} prints (co)inductive definitions and
  monotonicity rules; the ``\<open>!\<close>'' option indicates extra verbosity.

  \<^descr> @{attribute (HOL) mono} declares monotonicity rules in the context. These
  rule are involved in the automated monotonicity proof of the above inductive
  and coinductive definitions.
\<close>


subsection \<open>Derived rules\<close>

text \<open>
  A (co)inductive definition of \<open>R\<close> provides the following main theorems:

  \<^descr> \<open>R.intros\<close> is the list of introduction rules as proven theorems, for the
  recursive predicates (or sets). The rules are also available individually,
  using the names given them in the theory file;

  \<^descr> \<open>R.cases\<close> is the case analysis (or elimination) rule;

  \<^descr> \<open>R.induct\<close> or \<open>R.coinduct\<close> is the (co)induction rule;

  \<^descr> \<open>R.simps\<close> is the equation unrolling the fixpoint of the predicate one
  step.


  When several predicates \<open>R\<^sub>1, \<dots>, R\<^sub>n\<close> are defined simultaneously, the list
  of introduction rules is called \<open>R\<^sub>1_\<dots>_R\<^sub>n.intros\<close>, the case analysis rules
  are called \<open>R\<^sub>1.cases, \<dots>, R\<^sub>n.cases\<close>, and the list of mutual induction rules
  is called \<open>R\<^sub>1_\<dots>_R\<^sub>n.inducts\<close>.
\<close>


subsection \<open>Monotonicity theorems\<close>

text \<open>
  The context maintains a default set of theorems that are used in
  monotonicity proofs. New rules can be declared via the @{attribute (HOL)
  mono} attribute. See the main Isabelle/HOL sources for some examples. The
  general format of such monotonicity theorems is as follows:

  \<^item> Theorems of the form \<open>A \<le> B \<Longrightarrow> \<M> A \<le> \<M> B\<close>, for proving monotonicity of
  inductive definitions whose introduction rules have premises involving terms
  such as \<open>\<M> R t\<close>.

  \<^item> Monotonicity theorems for logical operators, which are of the general form
  \<open>(\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> (\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> \<longrightarrow> \<dots>\<close>. For example, in the case of the operator \<open>\<or>\<close>,
  the corresponding theorem is
  \[
  \infer{\<open>P\<^sub>1 \<or> P\<^sub>2 \<longrightarrow> Q\<^sub>1 \<or> Q\<^sub>2\<close>}{\<open>P\<^sub>1 \<longrightarrow> Q\<^sub>1\<close> & \<open>P\<^sub>2 \<longrightarrow> Q\<^sub>2\<close>}
  \]

  \<^item> De Morgan style equations for reasoning about the ``polarity'' of
  expressions, e.g.
  \[
  \<^prop>\<open>\<not> \<not> P \<longleftrightarrow> P\<close> \qquad\qquad
  \<^prop>\<open>\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q\<close>
  \]

  \<^item> Equations for reducing complex operators to more primitive ones whose
  monotonicity can easily be proved, e.g.
  \[
  \<^prop>\<open>(P \<longrightarrow> Q) \<longleftrightarrow> \<not> P \<or> Q\<close> \qquad\qquad
  \<^prop>\<open>Ball A P \<equiv> \<forall>x. x \<in> A \<longrightarrow> P x\<close>
  \]
\<close>


subsubsection \<open>Examples\<close>

text \<open>The finite powerset operator can be defined inductively like this:\<close>

(*<*)experiment begin(*>*)
inductive_set Fin :: "'a set \ 'a set set" for A :: "'a set"
where
  empty: "{} \ Fin A"
| insert: "a \ A \ B \ Fin A \ insert a B \ Fin A"

text \<open>The accessible part of a relation is defined as follows:\<close>

inductive acc :: "('a \ 'a \ bool) \ 'a \ bool"
  for r :: "'a \ 'a \ bool" (infix "\" 50)
where acc: "(\y. y \ x \ acc r y) \ acc r x"
(*<*)end(*>*)

text \<open>
  Common logical connectives can be easily characterized as non-recursive
  inductive definitions with parameters, but without arguments.
\<close>

(*<*)experiment begin(*>*)
inductive AND for A B :: bool
where "A \ B \ AND A B"

inductive OR for A B :: bool
where "A \ OR A B"
  | "B \ OR A B"

inductive EXISTS for B :: "'a \ bool"
where "B a \ EXISTS B"
(*<*)end(*>*)

text \<open>
  Here the \<open>cases\<close> or \<open>induct\<close> rules produced by the @{command inductive}
  package coincide with the expected elimination rules for Natural Deduction.
  Already in the original article by Gerhard Gentzen @{cite "Gentzen:1935"}
  there is a hint that each connective can be characterized by its
  introductions, and the elimination can be constructed systematically.
\<close>


section \<open>Recursive functions \label{sec:recursion}\<close>

text \<open>
  \begin{matharray}{rcl}
    @{command_def (HOL) "primrec"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
    @{command_def (HOL) "fun"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
    @{command_def (HOL) "function"} & : & \<open>local_theory \<rightarrow> proof(prove)\<close> \\
    @{command_def (HOL) "termination"} & : & \<open>local_theory \<rightarrow> proof(prove)\<close> \\
    @{command_def (HOL) "fun_cases"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
  \end{matharray}

  \<^rail>\<open>
    @@{command (HOL) primrec} @{syntax specification}
    ;
    (@@{command (HOL) fun} | @@{command (HOL) function}) opts? @{syntax specification}
    ;
    opts: '(' (('sequential' | 'domintros') + ','')'
    ;
    @@{command (HOL) termination} @{syntax term}?
    ;
    @@{command (HOL) fun_cases} (@{syntax thmdecl}? @{syntax prop} + @'and')
  \<close>

  \<^descr> @{command (HOL) "primrec"} defines primitive recursive functions over
  datatypes (see also @{command_ref (HOL) datatype}). The given \<open>equations\<close>
  specify reduction rules that are produced by instantiating the generic
  combinator for primitive recursion that is available for each datatype.

  Each equation needs to be of the form:

  @{text [display] "f x\<^sub>1 \ x\<^sub>m (C y\<^sub>1 \ y\<^sub>k) z\<^sub>1 \ z\<^sub>n = rhs"}

  such that \<open>C\<close> is a datatype constructor, \<open>rhs\<close> contains only the free
  variables on the left-hand side (or from the context), and all recursive
  occurrences of \<open>f\<close> in \<open>rhs\<close> are of the form \<open>f \<dots> y\<^sub>i \<dots>\<close> for some \<open>i\<close>. At
  most one reduction rule for each constructor can be given. The order does
  not matter. For missing constructors, the function is defined to return a
  default value, but this equation is made difficult to access for users.

  The reduction rules are declared as @{attribute simp} by default, which
  enables standard proof methods like @{method simp} and @{method auto} to
  normalize expressions of \<open>f\<close> applied to datatype constructions, by
  simulating symbolic computation via rewriting.

  \<^descr> @{command (HOL) "function"} defines functions by general wellfounded
  recursion. A detailed description with examples can be found in @{cite
  "isabelle-function"}. The function is specified by a set of (possibly
  conditional) recursive equations with arbitrary pattern matching. The
  command generates proof obligations for the completeness and the
  compatibility of patterns.

  The defined function is considered partial, and the resulting simplification
  rules (named \<open>f.psimps\<close>) and induction rule (named \<open>f.pinduct\<close>) are guarded
  by a generated domain predicate \<open>f_dom\<close>. The @{command (HOL) "termination"}
  command can then be used to establish that the function is total.

  \<^descr> @{command (HOL) "fun"} is a shorthand notation for ``@{command (HOL)
  "function"}~\<open>(sequential)\<close>'', followed by automated proof attempts regarding
  pattern matching and termination. See @{cite "isabelle-function"for
  further details.

  \<^descr> @{command (HOL) "termination"}~\<open>f\<close> commences a termination proof for the
  previously defined function \<open>f\<close>. If this is omitted, the command refers to
  the most recent function definition. After the proof is closed, the
  recursive equations and the induction principle is established.

  \<^descr> @{command (HOL) "fun_cases"} generates specialized elimination rules for
  function equations. It expects one or more function equations and produces
  rules that eliminate the given equalities, following the cases given in the
  function definition.


  Recursive definitions introduced by the @{command (HOL) "function"} command
  accommodate reasoning by induction (cf.\ @{method induct}): rule \<open>f.induct\<close>
  refers to a specific induction rule, with parameters named according to the
  user-specified equations. Cases are numbered starting from 1. For @{command
  (HOL) "primrec"}, the induction principle coincides with structural
  recursion on the datatype where the recursion is carried out.

  The equations provided by these packages may be referred later as theorem
  list \<open>f.simps\<close>, where \<open>f\<close> is the (collective) name of the functions defined.
  Individual equations may be named explicitly as well.

  The @{command (HOL) "function"} command accepts the following options.

  \<^descr> \<open>sequential\<close> enables a preprocessor which disambiguates overlapping
  patterns by making them mutually disjoint. Earlier equations take precedence
  over later ones. This allows to give the specification in a format very
  similar to functional programming. Note that the resulting simplification
  and induction rules correspond to the transformed specification, not the one
  given originally. This usually means that each equation given by the user
  may result in several theoremsAlso note that this automatic transformation
  only works for ML-style datatype patterns.

  \<^descr> \<open>domintros\<close> enables the automated generation of introduction rules for the
  domain predicate. While mostly not needed, they can be helpful in some
  proofs about partial functions.
\<close>


subsubsection \<open>Example: evaluation of expressions\<close>

text \<open>
  Subsequently, we define mutual datatypes for arithmetic and boolean
  expressions, and use @{command primrecfor evaluation functions that follow
  the same recursive structure.
\<close>

(*<*)experiment begin(*>*)
datatype 'a aexp =
    IF "'a bexp"  "'a aexp"  "'a aexp"
  | Sum "'a aexp"  "'a aexp"
  | Diff "'a aexp"  "'a aexp"
  | Var 'a
  | Num nat
and 'a bexp =
    Less "'a aexp"  "'a aexp"
  | And "'a bexp"  "'a bexp"
  | Neg "'a bexp"

text \<open>\<^medskip> Evaluation of arithmetic and boolean expressions\<close>

primrec evala :: "('a \ nat) \ 'a aexp \ nat"
  and evalb :: "('a \ nat) \ 'a bexp \ bool"
where
  "evala env (IF b a1 a2) = (if evalb env b then evala env a1 else evala env a2)"
"evala env (Sum a1 a2) = evala env a1 + evala env a2"
"evala env (Diff a1 a2) = evala env a1 - evala env a2"
"evala env (Var v) = env v"
"evala env (Num n) = n"
"evalb env (Less a1 a2) = (evala env a1 < evala env a2)"
"evalb env (And b1 b2) = (evalb env b1 \ evalb env b2)"
"evalb env (Neg b) = (\ evalb env b)"

text \<open>
  Since the value of an expression depends on the value of its variables, the
  functions \<^const>\<open>evala\<close> and \<^const>\<open>evalb\<close> take an additional parameter, an
  \<^emph>\<open>environment\<close> that maps variables to their values.

  \<^medskip>
  Substitution on expressions can be defined similarly. The mapping \<open>f\<close> of
  type \<^typ>\<open>'a \<Rightarrow> 'a aexp\<close> given as a parameter is lifted canonically on the
  types \<^typ>\<open>'a aexp\<close> and \<^typ>\<open>'a bexp\<close>, respectively.
\<close>

primrec substa :: "('a \ 'b aexp) \ 'a aexp \ 'b aexp"
  and substb :: "('a \ 'b aexp) \ 'a bexp \ 'b bexp"
where
  "substa f (IF b a1 a2) = IF (substb f b) (substa f a1) (substa f a2)"
"substa f (Sum a1 a2) = Sum (substa f a1) (substa f a2)"
"substa f (Diff a1 a2) = Diff (substa f a1) (substa f a2)"
"substa f (Var v) = f v"
"substa f (Num n) = Num n"
"substb f (Less a1 a2) = Less (substa f a1) (substa f a2)"
"substb f (And b1 b2) = And (substb f b1) (substb f b2)"
"substb f (Neg b) = Neg (substb f b)"

text \<open>
  In textbooks about semantics one often finds substitution theorems, which
  express the relationship between substitution and evaluation. For \<^typ>\<open>'a
  aexp\<close> and \<^typ>\<open>'a bexp\<close>, we can prove such a theorem by mutual
  induction, followed by simplification.
\<close>

lemma subst_one:
  "evala env (substa (Var (v := a')) a) = evala (env (v := evala env a')) a"
  "evalb env (substb (Var (v := a')) b) = evalb (env (v := evala env a')) b"
  by (induct a and b) simp_all

lemma subst_all:
  "evala env (substa s a) = evala (\x. evala env (s x)) a"
  "evalb env (substb s b) = evalb (\x. evala env (s x)) b"
  by (induct a and b) simp_all
(*<*)end(*>*)


subsubsection \<open>Example: a substitution function for terms\<close>

text \<open>Functions on datatypes with nested recursion are also defined
  by mutual primitive recursion.\<close>

(*<*)experiment begin(*>*)
datatype ('a, 'b) "term" = Var 'a | App '"('a, 'b) term list"

text \<open>
  A substitution function on type \<^typ>\<open>('a, 'b) term\<close> can be defined as
  follows, by working simultaneously on \<^typ>\<open>('a, 'b) term list\<close>:
\<close>

primrec subst_term :: "('a \ ('a, 'b) term) \ ('a, 'b) term \ ('a, 'b) term" and
  subst_term_list :: "('a \ ('a, 'b) term) \ ('a, 'b) term list \ ('a, 'b) term list"
where
  "subst_term f (Var a) = f a"
"subst_term f (App b ts) = App b (subst_term_list f ts)"
"subst_term_list f [] = []"
"subst_term_list f (t # ts) = subst_term f t # subst_term_list f ts"

text \<open>
  The recursion scheme follows the structure of the unfolded definition of
  type \<^typ>\<open>('a, 'b) term\<close>. To prove properties of this substitution
  function, mutual induction is needed:
\<close>

lemma "subst_term (subst_term f1 \ f2) t =
    subst_term f1 (subst_term f2 t)" and
  "subst_term_list (subst_term f1 \ f2) ts =
    subst_term_list f1 (subst_term_list f2 ts)"
  by (induct t and ts rule: subst_term.induct subst_term_list.induct) simp_all
(*<*)end(*>*)


subsubsection \<open>Example: a map function for infinitely branching trees\<close>

text \<open>Defining functions on infinitely branching datatypes by primitive
  recursion is just as easy.\<close>

(*<*)experiment begin(*>*)
datatype 'a tree = Atom 'a | Branch "nat \ 'a tree"

primrec map_tree :: "('a \ 'b) \ 'a tree \ 'b tree"
where
  "map_tree f (Atom a) = Atom (f a)"
"map_tree f (Branch ts) = Branch (\x. map_tree f (ts x))"

text \<open>
  Note that all occurrences of functions such as \<open>ts\<close> above must be applied to
  an argument. In particular, \<^term>\<open>map_tree f \<circ> ts\<close> is not allowed here.

  \<^medskip>
  Here is a simple composition lemma for \<^term>\<open>map_tree\<close>:
\<close>

lemma "map_tree g (map_tree f t) = map_tree (g \ f) t"
  by (induct t) simp_all
(*<*)end(*>*)


subsection \<open>Proof methods related to recursive definitions\<close>

text \<open>
  \begin{matharray}{rcl}
    @{method_def (HOL) pat_completeness} & : & \<open>method\<close> \\
    @{method_def (HOL) relation} & : & \<open>method\<close> \\
    @{method_def (HOL) lexicographic_order} & : & \<open>method\<close> \\
    @{method_def (HOL) size_change} & : & \<open>method\<close> \\
    @{attribute_def (HOL) termination_simp} & : & \<open>attribute\<close> \\
    @{method_def (HOL) induction_schema} & : & \<open>method\<close> \\
  \end{matharray}

  \<^rail>\<open>
    @@{method (HOL) relation} @{syntax term}
    ;
    @@{method (HOL) lexicographic_order} (@{syntax clasimpmod} * )
    ;
    @@{method (HOL) size_change} ( orders (@{syntax clasimpmod} * ) )
    ;
    @@{method (HOL) induction_schema}
    ;
    orders: ( 'max' | 'min' | 'ms' ) *
  \<close>

  \<^descr> @{method (HOL) pat_completeness} is a specialized method to solve goals
  regarding the completeness of pattern matching, as required by the @{command
  (HOL) "function"} package (cf.\ @{cite "isabelle-function"}).

  \<^descr> @{method (HOL) relation}~\<open>R\<close> introduces a termination proof using the
  relation \<open>R\<close>. The resulting proof state will contain goals expressing that
  \<open>R\<close> is wellfounded, and that the arguments of recursive calls decrease with
  respect to \<open>R\<close>. Usually, this method is used as the initial proof step of
  manual termination proofs.

  \<^descr> @{method (HOL) "lexicographic_order"} attempts a fully automated
  termination proof by searching for a lexicographic combination of size
  measures on the arguments of the function. The method accepts the same
  arguments as the @{method auto} method, which it uses internally to prove
  local descents. The @{syntax clasimpmod} modifiers are accepted (as for
  @{method auto}).

  In case of failure, extensive information is printed, which can help to
  analyse the situation (cf.\ @{cite "isabelle-function"}).

  \<^descr> @{method (HOL) "size_change"} also works on termination goals, using a
  variation of the size-change principle, together with a graph decomposition
  technique (see @{cite krauss_phd} for details). Three kinds of orders are
  used internally: \<open>max\<close>, \<open>min\<close>, and \<open>ms\<close> (multiset), which is only available
  when the theory \<open>Multiset\<close> is loaded. When no order kinds are given, they
  are tried in order. The search for a termination proof uses SAT solving
  internally.

  For local descent proofs, the @{syntax clasimpmod} modifiers are accepted
  (as for @{method auto}).

  \<^descr> @{attribute (HOL) termination_simp} declares extra rules for the
  simplifier, when invoked in termination proofs. This can be useful, e.g.,
  for special rules involving size estimations.

  \<^descr> @{method (HOL) induction_schema} derives user-specified induction rules
  from well-founded induction and completeness of patterns. This factors out
  some operations that are done internally by the function package and makes
  them available separately. See \<^file>\<open>~~/src/HOL/Examples/Induction_Schema.thy\<close> for
  examples.
\<close>


subsection \<open>Functions with explicit partiality\<close>

text \<open>
  \begin{matharray}{rcl}
    @{command_def (HOL) "partial_function"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
    @{attribute_def (HOL) "partial_function_mono"} & : & \<open>attribute\<close> \\
  \end{matharray}

  \<^rail>\<open>
    @@{command (HOL) partial_function} '(' @{syntax name} ')'
      @{syntax specification}
  \<close>

  \<^descr> @{command (HOL) "partial_function"}~\<open>(mode)\<close> defines recursive functions
  based on fixpoints in complete partial orders. No termination proof is
  required from the user or constructed internally. Instead, the possibility
  of non-termination is modelled explicitly in the result type, which contains
  an explicit bottom element.

  Pattern matching and mutual recursion are currently not supported. Thus, the
  specification consists of a single function described by a single recursive
  equation.

  There are no fixed syntactic restrictions on the body of the function, but
  the induced functional must be provably monotonic wrt.\ the underlying
  order. The monotonicity proof is performed internally, and the definition is
  rejected when it fails. The proof can be influenced by declaring hints using
  the @{attribute (HOL) partial_function_mono} attribute.

  The mandatory \<open>mode\<close> argument specifies the mode of operation of the
  command, which directly corresponds to a complete partial order on the
  result type. By default, the following modes are defined:

    \<^descr> \<open>option\<close> defines functions that map into the \<^type>\<open>option\<close> type. Here,
    the value \<^term>\<open>None\<close> is used to model a non-terminating computation.
    Monotonicity requires that if \<^term>\<open>None\<close> is returned by a recursive
    call, then the overall result must also be \<^term>\<open>None\<close>. This is best
    achieved through the use of the monadic operator \<^const>\<open>Option.bind\<close>.

    \<^descr> \<open>tailrec\<close> defines functions with an arbitrary result type and uses the
    slightly degenerated partial order where \<^term>\<open>undefined\<close> is the bottom
    element. Now, monotonicity requires that if \<^term>\<open>undefined\<close> is returned
    by a recursive call, then the overall result must also be \<^term>\<open>undefined\<close>. In practice, this is only satisfied when each recursive call
    is a tail call, whose result is directly returned. Thus, this mode of
    operation allows the definition of arbitrary tail-recursive functions.

  Experienced users may define new modes by instantiating the locale \<^const>\<open>partial_function_definitions\<close> appropriately.

  \<^descr> @{attribute (HOL) partial_function_mono} declares rules for use in the
  internal monotonicity proofs of partial function definitions.
\<close>


subsection \<open>Old-style recursive function definitions (TFL)\<close>

text \<open>
  \begin{matharray}{rcl}
    @{command_def (HOL) "recdef"} & : & \<open>theory \<rightarrow> theory)\<close> \\
  \end{matharray}

  The old TFL command @{command (HOL) "recdef"for defining recursive is
  mostly obsolete; @{command (HOL) "function"} or @{command (HOL) "fun"}
  should be used instead.

  \<^rail>\<open>
    @@{command (HOL) recdef} ('(' @'permissive' ')')? \<newline>
      @{syntax name} @{syntax term} (@{syntax prop} +) hints?
    ;
    hints'(' @'hints' ( recdefmod * ) ')'
    ;
    recdefmod: (('recdef_simp' | 'recdef_cong' | 'recdef_wf')
      (() | 'add' | 'del'':' @{syntax thms}) | @{syntax clasimpmod}
  \<close>

  \<^descr> @{command (HOL) "recdef"} defines general well-founded recursive functions
  (using the TFL package). The ``\<open>(permissive)\<close>'' option tells TFL to recover
  from failed proof attempts, returning unfinished results. The \<open>recdef_simp\<close>,
  \<open>recdef_cong\<close>, and \<open>recdef_wf\<close> hints refer to auxiliary rules to be used in
  the internal automated proof process of TFL. Additional @{syntax clasimpmod}
  declarations may be given to tune the context of the Simplifier (cf.\
  \secref{sec:simplifier}) and Classical reasoner (cf.\
  \secref{sec:classical}).


  \<^medskip>
  Hints for @{command (HOL) "recdef"} may be also declared globally, using the
  following attributes.

  \begin{matharray}{rcl}
    @{attribute_def (HOL) recdef_simp} & : & \<open>attribute\<close> \\
    @{attribute_def (HOL) recdef_cong} & : & \<open>attribute\<close> \\
    @{attribute_def (HOL) recdef_wf} & : & \<open>attribute\<close> \\
  \end{matharray}

  \<^rail>\<open>
    (@@{attribute (HOL) recdef_simp} | @@{attribute (HOL) recdef_cong} |
      @@{attribute (HOL) recdef_wf}) (() | 'add' | 'del')
  \<close>
\<close>


section \<open>Adhoc overloading of constants\<close>

text \<open>
  \begin{tabular}{rcll}
  @{command_def "adhoc_overloading"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
  @{command_def "no_adhoc_overloading"} & : & \<open>local_theory \<rightarrow> local_theory\<close> \\
  @{attribute_def "show_variants"} & : & \<open>attribute\<close> & default \<open>false\<close> \\
  \end{tabular}

  \<^medskip>
  Adhoc overloading allows to overload a constant depending on its type.
  Typically this involves the introduction of an uninterpreted constant (used
  for input and outputand the addition of some variants (used internally).
  For examples see \<^file>\<open>~~/src/HOL/Examples/Adhoc_Overloading_Examples.thy\<close> and
  \<^file>\<open>~~/src/HOL/Library/Monad_Syntax.thy\<close>.

  \<^rail>\<open>
    (@@{command adhoc_overloading} | @@{command no_adhoc_overloading})
      (@{syntax name} (@{syntax term} + ) + @'and')
  \<close>

  \<^descr> @{command "adhoc_overloading"}~\<open>c v\<^sub>1 ... v\<^sub>n\<close> associates variants with an
  existing constant.

  \<^descr> @{command "no_adhoc_overloading"} is similar to @{command
  "adhoc_overloading"}, but removes the specified variants from the present
  context.

  \<^descr> @{attribute "show_variants"} controls printing of variants of overloaded
  constants. If enabled, the internally used variants are printed instead of
  their respective overloaded constants. This is occasionally useful to check
  whether the system agrees with a user's expectations about derived variants.
\<close>


section \<open>Definition by specification \label{sec:hol-specification}\<close>

text \<open>
  \begin{matharray}{rcl}
    @{command_def (HOL) "specification"} & : & \<open>theory \<rightarrow> proof(prove)\<close> \\
  \end{matharray}

  \<^rail>\<open>
    @@{command (HOL) specification'(' (decl +) ')' \<newline>
      (@{syntax thmdecl}? @{syntax prop} +)
    ;
    decl: (@{syntax name} ':')? @{syntax term} ('(' @'overloaded' ')')?
  \<close>

  \<^descr> @{command (HOL) "specification"}~\<open>decls \<phi>\<close> sets up a goal stating the
  existence of terms with the properties specified to hold for the constants
  given in \<open>decls\<close>. After finishing the proof, the theory will be augmented
  with definitions for the given constants, as well as with theorems stating
  the properties for these constants.

  \<open>decl\<close> declares a constant to be defined by the specification given. The
  definition for the constant \<open>c\<close> is bound to the name \<open>c_def\<close> unless a
  theorem name is given in the declarationOverloaded constants should be
  declared as such.
\<close>


section \<open>Old-style datatypes \label{sec:hol-datatype}\<close>

text \<open>
  \begin{matharray}{rcl}
    @{command_def (HOL) "old_rep_datatype"} & : & \<open>theory \<rightarrow> proof(prove)\<close> \\
  \end{matharray}

  \<^rail>\<open>
    @@{command (HOL) old_rep_datatype} ('(' (@{syntax name} +) ')')? (@{syntax term} +)
    ;

    spec: @{syntax typespec_sorts} @{syntax mixfix}? '=' (cons + '|')
    ;
    cons: @{syntax name} (@{syntax type} * ) @{syntax mixfix}?
  \<close>

  \<^descr> @{command (HOL) "old_rep_datatype"} represents existing types as
  old-style datatypes.


  These commands are mostly obsolete; @{command (HOL) "datatype"} should be
  used instead.

  See @{cite "isabelle-datatypes"for more details on datatypes. Apart from
  proper proof methods for case analysis and induction, there are also
  emulations of ML tactics @{method (HOL) case_tac} and @{method (HOL)
  induct_tac} available, see \secref{sec:hol-induct-tac}; these admit to refer
  directly to the internal structure of subgoals (including internally bound
  parameters).
\<close>


subsubsection \<open>Examples\<close>

text \<open>
  We define a type of finite sequences, with slightly different names than the
  existing \<^typ>\<open>'a list\<close> that is already in \<^theory>\<open>Main\<close>:
\<close>

(*<*)experiment begin(*>*)
datatype 'a seq = Empty | Seq '"'a seq"

text \<open>We can now prove some simple lemma by structural induction:\<close>

lemma "Seq x xs \ xs"
proof (induct xs arbitrary: x)
  case Empty
  txt \<open>This case can be proved using the simplifier: the freeness
    properties of the datatype are already declared as @{attribute
    simp} rules.\<close>
  show "Seq x Empty \ Empty"
    by simp
next
  case (Seq y ys)
  txt \<open>The step case is proved similarly.\<close>
  show "Seq x (Seq y ys) \ Seq y ys"
    using \<open>Seq y ys \<noteq> ys\<close> by simp
qed

text \<open>Here is a more succinct version of the same proof:\<close>

lemma "Seq x xs \ xs"
  by (induct xs arbitrary: x) simp_all
(*<*)end(*>*)


section \<open>Records \label{sec:hol-record}\<close>

text \<open>
  In principle, records merely generalize the concept of tuples, where
  components may be addressed by labels instead of just position. The logical
  infrastructure of records in Isabelle/HOL is slightly more advanced, though,
  supporting truly extensible record schemes. This admits operations that are
  polymorphic with respect to record extension, yielding ``object-oriented''
  effects like (single) inheritance. See also @{cite "NaraschewskiW-TPHOLs98"}
  for more details on object-oriented verification and record subtyping in
  HOL.
\<close>


subsection \<open>Basic concepts\<close>

text \<open>
  Isabelle/HOL supports both \<^emph>\<open>fixed\<close> and \<^emph>\<open>schematic\<close> records at the level of
  terms and types. The notation is as follows:

  \begin{center}
  \begin{tabular}{l|l|l}
    & record terms & record types \\ \hline
    fixed & \<open>\<lparr>x = a, y = b\<rparr>\<close> & \<open>\<lparr>x :: A, y :: B\<rparr>\<close> \\
    schematic & \<open>\<lparr>x = a, y = b, \<dots> = m\<rparr>\<close> &
      \<open>\<lparr>x :: A, y :: B, \<dots> :: M\<rparr>\<close> \\
  \end{tabular}
  \end{center}

  The ASCII representation of \<open>\<lparr>x = a\<rparr>\<close> is \<open>(| x = a |)\<close>.

  A fixed record \<open>\<lparr>x = a, y = b\<rparr>\<close> has field \<open>x\<close> of value \<open>a\<close> and field \<open>y\<close> of
  value \<open>b\<close>. The corresponding type is \<open>\<lparr>x :: A, y :: B\<rparr>\<close>, assuming that \<open>a ::
  A\<close> and \<open>b :: B\<close>.

  A record scheme like \<open>\<lparr>x = a, y = b, \<dots> = m\<rparr>\<close> contains fields \<open>x\<close> and \<open>y\<close> as
  before, but also possibly further fields as indicated by the ``\<open>\<dots>\<close>''
  notation (which is actually part of the syntax). The improper field ``\<open>\<dots>\<close>''
  of a record scheme is called the \<^emph>\<open>more part\<close>. Logically it is just a free
  variable, which is occasionally referred to as ``row variable'' in the
  literature. The more part of a record scheme may be instantiated by zero or
  more further components. For example, the previous scheme may get
  instantiated to \<open>\<lparr>x = a, y = b, z = c, \<dots> = m'\<rparr>\<close>, where \<open>m'\<close> refers to a
  different more part. Fixed records are special instances of record schemes,
  where ``\<open>\<dots>\<close>'' is properly terminated by the \<open>() :: unit\<close> element. In fact,
  \<open>\<lparr>x = a, y = b\<rparr>\<close> is just an abbreviation for \<open>\<lparr>x = a, y = b, \<dots> = ()\<rparr>\<close>.

  \<^medskip>
  Two key observations make extensible records in a simply typed language like
  HOL work out:

  \<^enum> the more part is internalized, as a free term or type variable,

  \<^enum> field names are externalized, they cannot be accessed within the logic as
  first-class values.


  \<^medskip>
  In Isabelle/HOL record types have to be defined explicitly, fixing their
  field names and typesand their (optional) parent record. Afterwards,
  records may be formed using above syntax, while obeying the canonical order
  of fields as given by their declaration. The record package provides several
  standard operations like selectors and updates. The common setup for various
  generic proof tools enable succinct reasoning patterns. See also the
  Isabelle/HOL tutorial @{cite "isabelle-hol-book"for further instructions
  on using records in practice.
\<close>


subsection \<open>Record specifications\<close>

text \<open>
  \begin{matharray}{rcl}
    @{command_def (HOL) "record"} & : & \<open>theory \<rightarrow> theory\<close> \\
    @{command_def (HOL) "print_record"} & : & \<open>context \<rightarrow>\<close> \\
  \end{matharray}

  \<^rail>\<open>
    @@{command (HOL) record} @{syntax "overloaded"}? @{syntax typespec_sorts} '=' \<newline>
      (@{syntax type} '+')? (constdecl +)
    ;
    constdecl: @{syntax name} '::' @{syntax type} @{syntax mixfix}?
    ;
    @@{command (HOL) print_record} modes? @{syntax typespec_sorts}
    ;
    modes: '(' (@{syntax name} +) ')'
  \<close>

  \<^descr> @{command (HOL) "record"}~\<open>(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t = \<tau> + c\<^sub>1 :: \<sigma>\<^sub>1 \<dots> c\<^sub>n :: \<sigma>\<^sub>n\<close>
  defines extensible record type \<open>(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t\<close>, derived from the optional
  parent record \<open>\<tau>\<close> by adding new field components \<open>c\<^sub>i :: \<sigma>\<^sub>i\<close> etc.

  The type variables of \<open>\<tau>\<close> and \<open>\<sigma>\<^sub>i\<close> need to be covered by the (distinct)
  parameters \<open>\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m\<close>. Type constructor \<open>t\<close> has to be new, while \<open>\<tau>\<close>
  needs to specify an instance of an existing record type. At least one new
  field \<open>c\<^sub>i\<close> has to be specified. Basically, field names need to belong to a
  unique record. This is not a real restriction in practice, since fields are
  qualified by the record name internally.

  The parent record specification \<open>\<tau>\<close> is optional; if omitted \<open>t\<close> becomes a
  root record. The hierarchy of all records declared within a theory context
  forms a forest structure, i.e.\ a set of trees starting with a root record
  each. There is no way to merge multiple parent records!

  For convenience, \<open>(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t\<close> is made a type abbreviation for the fixed
  record type \<open>\<lparr>c\<^sub>1 :: \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n\<rparr>\<close>, likewise is \<open>(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m, \<zeta>)
  t_scheme\<close> made an abbreviation for \<open>\<lparr>c\<^sub>1 :: \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n, \<dots> :: \<zeta>\<rparr>\<close>.

  \<^descr> @{command (HOL) "print_record"}~\<open>(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t\<close> prints the definition of
  record \<open>(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t\<close>. Optionally \<open>modes\<close> can be specified, which are
  appended to the current print mode; see \secref{sec:print-modes}.
\<close>


subsection \<open>Record operations\<close>

text \<open>
  Any record definition of the form presented above produces certain standard
  operations. Selectors and updates are provided for any field, including the
  improper one ``\<open>more\<close>''. There are also cumulative record constructor
  functions. To simplify the presentation below, we assume for now that \<open>(\<alpha>\<^sub>1,
  \<dots>, \<alpha>\<^sub>m) t\<close> is a root record with fields \<open>c\<^sub>1 :: \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n\<close>.

  \<^medskip>
  \<^bold>\<open>Selectors\<close> and \<^bold>\<open>updates\<close> are available for any
  field (including ``\<open>more\<close>''):

  \begin{matharray}{lll}
    \<open>c\<^sub>i\<close> & \<open>::\<close> & \<open>\<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i\<close> \\
    \<open>c\<^sub>i_update\<close> & \<open>::\<close> & \<open>\<sigma>\<^sub>i \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>\<close> \\
  \end{matharray}

  There is special syntax for application of updates: \<open>r\<lparr>x := a\<rparr>\<close> abbreviates
  term \<open>x_update a r\<close>. Further notation for repeated updates is also
  available: \<open>r\<lparr>x := a\<rparr>\<lparr>y := b\<rparr>\<lparr>z := c\<rparr>\<close> may be written \<open>r\<lparr>x := a, y := b, z
  := c\<rparr>\<close>. Note that because of postfix notation the order of fields shown here
  is reverse than in the actual term. Since repeated updates are just function
  applications, fields may be freely permuted in \<open>\<lparr>x := a, y := b, z := c\<rparr>\<close>,
  as far as logical equality is concerned. Thus commutativity of independent
  updates can be proven within the logic for any two fields, but not as a
  general theorem.

  \<^medskip>
  The \<^bold>\<open>make\<close> operation provides a cumulative record constructor function:

  \begin{matharray}{lll}
    \<open>t.make\<close> & \<open>::\<close> & \<open>\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>\<close> \\
  \end{matharray}

  \<^medskip>
  We now reconsider the case of non-root records, which are derived of some
  parent. In general, the latter may depend on another parent as well,
  resulting in a list of \<^emph>\<open>ancestor records\<close>. Appending the lists of fields of
  all ancestors results in a certain field prefix. The record package
  automatically takes care of this by lifting operations over this context of
  ancestor fields. Assuming that \<open>(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t\<close> has ancestor fields \<open>b\<^sub>1 ::
  \<rho>\<^sub>1, \<dots>, b\<^sub>k :: \<rho>\<^sub>k\<close>, the above record operations will get the following
  types:

  \<^medskip>
  \begin{tabular}{lll}
    \<open>c\<^sub>i\<close> & \<open>::\<close> & \<open>\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i\<close> \\
    \<open>c\<^sub>i_update\<close> & \<open>::\<close> & \<open>\<sigma>\<^sub>i \<Rightarrow>
      \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow>
      \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>\<close> \\
    \<open>t.make\<close> & \<open>::\<close> & \<open>\<rho>\<^sub>1 \<Rightarrow> \<dots> \<rho>\<^sub>k \<Rightarrow> \<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow>
      \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr>\<close> \\
  \end{tabular}
  \<^medskip>

  Some further operations address the extension aspect of a derived record
  scheme specifically: \<open>t.fields\<close> produces a record fragment consisting of
  exactly the new fields introduced here (the result may serve as a more part
  elsewhere); \<open>t.extend\<close> takes a fixed record and adds a given more part;
  \<open>t.truncate\<close> restricts a record scheme to a fixed record.

  \<^medskip>
  \begin{tabular}{lll}
    \<open>t.fields\<close> & \<open>::\<close> & \<open>\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>\<close> \\
    \<open>t.extend\<close> & \<open>::\<close> & \<open>\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr> \<Rightarrow>
      \<zeta> \<Rightarrow> \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr>\<close> \\
    \<open>t.truncate\<close> & \<open>::\<close> & \<open>\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>\<rparr>\<close> \\
  \end{tabular}
  \<^medskip>

  Note that \<open>t.make\<close> and \<open>t.fields\<close> coincide for root records.
\<close>


subsection \<open>Derived rules and proof tools\<close>

text \<open>
  The record package proves several results internally, declaring these facts
  to appropriate proof tools. This enables users to reason about record
  structures quite conveniently. Assume that \<open>t\<close> is a record type as specified
  above.

  \<^enum> Standard conversions for selectors or updates applied to record
  constructor terms are made part of the default Simplifier contextthus
  proofs by reduction of basic operations merely require the @{method simp}
  method without further arguments. These rules are available as \<open>t.simps\<close>,
  too.

  \<^enum> Selectors applied to updated records are automatically reduced by an
  internal simplification procedure, which is also part of the standard
  Simplifier setup.

  \<^enum> Inject equations of a form analogous to \<^prop>\<open>(x, y) = (x', y') \<equiv> x = x'
  \<and> y = y'\<close> are declared to the Simplifier and Classical Reasoner as
  @{attribute iff} rules. These rules are available as \<open>t.iffs\<close>.

  \<^enum> The introduction rule for record equality analogous to \<open>x r = x r' \<Longrightarrow> y r =
  y r' \ \ r = r'\ is declared to the Simplifier, and as the basic rule
  context as ``@{attribute intro}\<open>?\<close>''. The rule is called \<open>t.equality\<close>.

  \<^enum> Representations of arbitrary record expressions as canonical constructor
  terms are provided both in @{method cases} and @{method induct} format (cf.\
  the generic proof methods of the same name, \secref{sec:cases-induct}).
  Several variations are available, for fixed records, record schemes, more
  parts etc.

  The generic proof methods are sufficiently smart to pick the most sensible
  rule according to the type of the indicated record expression: users just
  need to apply something like ``\<open>(cases r)\<close>'' to a certain proof problem.

  \<^enum> The derived record operations \<open>t.make\<close>, \<open>t.fields\<close>, \<open>t.extend\<close>,
  \<open>t.truncate\<close> are \<^emph>\<open>not\<close> treated automatically, but usually need to be
  expanded by hand, using the collective fact \<open>t.defs\<close>.
\<close>


subsubsection \<open>Examples\<close>

text \<open>See \<^file>\<open>~~/src/HOL/Examples/Records.thy\<close>, for example.\<close>


section \<open>Semantic subtype definitions \label{sec:hol-typedef}\<close>

text \<open>
  \begin{matharray}{rcl}
    @{command_def (HOL) "typedef"} & : & \<open>local_theory \<rightarrow> proof(prove)\<close> \\
  \end{matharray}

  A type definition identifies a new type with a non-empty subset of an
  existing type. More precisely, the new type is defined by exhibiting an
  existing type \<open>\<tau>\<close>, a set \<open>A :: \<tau> set\<close>, and proving \<^prop>\<open>\<exists>x. x \<in> A\<close>. Thus
  \<open>A\<close> is a non-empty subset of \<open>\<tau>\<close>, and the new type denotes this subset. New
  functions are postulated that establish an isomorphism between the new type
  and the subset. In general, the type \<open>\<tau>\<close> may involve type variables \<open>\<alpha>\<^sub>1, \<dots>,
  \<alpha>\<^sub>n\<close> which means that the type definition produces a type constructor \<open>(\<alpha>\<^sub>1,
  \<dots>, \<alpha>\<^sub>n) t\<close> depending on those type arguments.

  \<^rail>\<open>
    @@{command (HOL) typedef} @{syntax "overloaded"}? abs_type '=' rep_set
    ;
    @{syntax_def "overloaded"}: ('(' @'overloaded' ')')
    ;
    abs_type: @{syntax typespec_sorts} @{syntax mixfix}?
    ;
    rep_set: @{syntax term} (@'morphisms' @{syntax name} @{syntax name})?
  \<close>

  To understand the concept of type definition better, we need to recount its
  somewhat complex history. The HOL logic goes back to the ``Simple Theory of
  Types'' (STT) of A. Church @{cite "church40"}, which is further explained in
  the book by P. Andrews @{cite "andrews86"}. The overview article by W.
  Farmer @{cite "Farmer:2008"} points out the ``seven virtues'' of this
  relatively simple family of logics. STT has only ground types, without
  polymorphism and without type definitions.

  \<^medskip>
  M. Gordon @{cite "Gordon:1985:HOL"} augmented Church's STT by adding
  schematic polymorphism (type variables and type constructors) and a facility
  to introduce new types as semantic subtypes from existing types. This
  genuine extension of the logic was explained semantically by A. Pitts in the
  book of the original Cambridge HOL88 system @{cite "pitts93"}. Type
  definitions work in this setting, because the general model-theory of STT is
  restricted to models that ensure that the universe of type interpretations
  is closed by forming subsets (via predicates taken from the logic).

  \<^medskip>
  Isabelle/HOL goes beyond Gordon-style HOL by admitting overloaded constant
  definitions @{cite "Wenzel:1997:TPHOL" and "Haftmann-Wenzel:2006:classes"},
  which are actually a concept of Isabelle/Pure and do not depend on
  particular set-theoretic semantics of HOL. Over many years, there was no
  formal checking of semantic type definitions in Isabelle/HOL versus
  syntactic constant definitions in Isabelle/Pure. So the @{command typedef}
  command was described as ``axiomatic'' in the sense of
  \secref{sec:axiomatizations}, only with some local checks of the given type
  and its representing set.

  Recent clarification of overloading in the HOL logic proper @{cite
  "Kuncar-Popescu:2015"} demonstrates how the dissimilar concepts of constant
  definitions versus type definitions may be understood uniformly. This
  requires an interpretation of Isabelle/HOL that substantially reforms the
  set-theoretic model of A. Pitts @{cite "pitts93"}, by taking a schematic
  view on polymorphism and interpreting only ground types in the set-theoretic
  sense of HOL88. Moreover, type-constructors may be explicitly overloaded,
  e.g.\ by making the subset depend on type-class parameters (cf.\
  \secref{sec:class}). This is semantically like a dependent type: the meaning
  relies on the operations provided by different type-class instances.

  \<^descr> @{command (HOL) "typedef"}~\<open>(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t = A\<close> defines a new type \<open>(\<alpha>\<^sub>1,
  \<dots>, \<alpha>\<^sub>n) t\<close> from the set \<open>A\<close> over an existing type. The set \<open>A\<close> may contain
  type variables \<open>\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n\<close> as specified on the LHS, but no term variables.
  Non-emptiness of \<open>A\<close> needs to be proven on the spot, in order to turn the
  internal conditional characterization into usable theorems.

  The ``\<open>(overloaded)\<close>'' option allows the @{command "typedef"} specification
  to depend on constants that are not (yet) specified and thus left open as
  parameters, e.g.\ type-class parameters.

  Within a local theory specification, the newly introduced type constructor
  cannot depend on parameters or assumptions of the context: this is
  syntactically impossible in HOL. The non-emptiness proof may formally depend
  on local assumptions, but this has little practical relevance.

  For @{command (HOL) "typedef"}~\<open>t = A\<close> the newly introduced type \<open>t\<close> is
  accompanied by a pair of morphisms to relate it to the representing set over
  the old type. By default, the injection from type to set is called \<open>Rep_t\<close>
  and its inverse \<open>Abs_t\<close>: An explicit @{keyword (HOL) "morphisms"}
  specification allows to provide alternative names.

  The logical characterization of @{command typedefuses the predicate of
  locale \<^const>\<open>type_definition\<close> that is defined in Isabelle/HOL. Various
  basic consequences of that are instantiated accordingly, re-using the locale
  facts with names derived from the new type constructor. Thus the generic
  theorem @{thm type_definition.Rep} is turned into the specific \<open>Rep_t\<close>, for
  example.

  Theorems @{thm type_definition.Rep}, @{thm type_definition.Rep_inverse}, and
  @{thm type_definition.Abs_inverse} provide the most basic characterization
  as a corresponding injection/surjection pair (in both directions). The
  derived rules @{thm type_definition.Rep_inject} and @{thm
  type_definition.Abs_inject} provide a more convenient version of
  injectivity, suitable for automated proof tools (e.g.\ in declarations
  involving @{attribute simp} or @{attribute iff}). Furthermore, the rules
  @{thm type_definition.Rep_cases}~/ @{thm type_definition.Rep_induct}, and
  @{thm type_definition.Abs_cases}~/ @{thm type_definition.Abs_induct} provide
  alternative views on surjectivity. These rules are already declared as set
  or type rules for the generic @{method cases} and @{method induct} methods,
  respectively.
\<close>


subsubsection \<open>Examples\<close>

text \<open>
  The following trivial example pulls a three-element type into existence
  within the formal logical environment of Isabelle/HOL.\<close>

(*<*)experiment begin(*>*)
typedef three = "{(True, True), (True, False), (False, True)}"
  by blast

definition "One = Abs_three (True, True)"
definition "Two = Abs_three (True, False)"
definition "Three = Abs_three (False, True)"

lemma three_distinct: "One \ Two" "One \ Three" "Two \ Three"
  by (simp_all add: One_def Two_def Three_def Abs_three_inject)

lemma three_cases:
  fixes x :: three obtains "x = One" | "x = Two" | "x = Three"
  by (cases x) (auto simp: One_def Two_def Three_def Abs_three_inject)
(*<*)end(*>*)

text \<open>Note that such trivial constructions are better done with
  derived specification mechanisms such as @{command datatype}:\<close>

(*<*)experiment begin(*>*)
datatype three = One | Two | Three
(*<*)end(*>*)

text \<open>This avoids re-doing basic definitions and proofs from the
  primitive @{command typedef} above.\<close>



section \<open>Functorial structure of types\<close>

text \<open>
  \begin{matharray}{rcl}
    @{command_def (HOL) "functor"} & : & \<open>local_theory \<rightarrow> proof(prove)\<close>
  \end{matharray}

  \<^rail>\<open>
    @@{command (HOL) functor} (@{syntax name} ':')? @{syntax term}
  \<close>

  \<^descr> @{command (HOL) "functor"}~\<open>prefix: m\<close> allows to prove and register
  properties about the functorial structure of type constructors. These
  properties then can be used by other packages to deal with those type
  constructors in certain type constructions. Characteristic theorems are
  noted in the current local theoryBy default, they are prefixed with the
  base name of the type constructor, an explicit prefix can be given
  alternatively.

  The given term \<open>m\<close> is considered as \<^emph>\<open>mapper\<close> for the corresponding type
  constructor and must conform to the following type pattern:

  \begin{matharray}{lll}
    \<open>m\<close> & \<open>::\<close> &
      \<open>\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>k \<Rightarrow> (\<^vec>\<alpha>\<^sub>n) t \<Rightarrow> (\<^vec>\<beta>\<^sub>n) t\<close> \\
  \end{matharray}

  where \<open>t\<close> is the type constructor, \<open>\<^vec>\<alpha>\<^sub>n\<close> and \<open>\<^vec>\<beta>\<^sub>n\<close> are
  distinct type variables free in the local theory and \<open>\<sigma>\<^sub>1\<close>, \ldots, \<open>\<sigma>\<^sub>k\<close> is
  a subsequence of \<open>\<alpha>\<^sub>1 \<Rightarrow> \<beta>\<^sub>1\<close>, \<open>\<beta>\<^sub>1 \<Rightarrow> \<alpha>\<^sub>1\<close>, \ldots, \<open>\<alpha>\<^sub>n \<Rightarrow> \<beta>\<^sub>n\<close>, \<open>\<beta>\<^sub>n \<Rightarrow> \<alpha>\<^sub>n\<close>.
\<close>


section \<open>Quotient types with lifting and transfer\<close>

text \<open>
  The quotient package defines a new quotient type given a raw type and a
  partial equivalence relation (\secref{sec:quotient-type}). The package also
  historically includes automation for transporting definitions and theorems
  (\secref{sec:old-quotient}), but most of this automation was superseded by
  the Lifting (\secref{sec:lifting}) and Transfer (\secref{sec:transfer})
  packages.
\<close>


subsection \<open>Quotient type definition \label{sec:quotient-type}\<close>

text \<open>
  \begin{matharray}{rcl}
    @{command_def (HOL) "quotient_type"} & : & \<open>local_theory \<rightarrow> proof(prove)\<close>\\
  \end{matharray}

  \<^rail>\<open>
    @@{command (HOL) quotient_type} @{syntax "overloaded"}? \<newline>
      @{syntax typespec} @{syntax mixfix}? '=' quot_type \<newline>
      quot_morphisms? quot_parametric?
    ;
    quot_type: @{syntax type} '/' ('partial' ':')? @{syntax term}
    ;
    quot_morphisms: @'morphisms' @{syntax name} @{syntax name}
    ;
    quot_parametric: @'parametric' @{syntax thm}
  \<close>

  \<^descr> @{command (HOL) "quotient_type"} defines a new quotient type \<open>\<tau>\<close>. The
  injection from a quotient type to a raw type is called \<open>rep_\<tau>\<close>, its inverse
  \<open>abs_\<tau>\<close> unless explicit @{keyword (HOL) "morphisms"} specification provides
  alternative names. @{command (HOL) "quotient_type"} requires the user to
  prove that the relation is an equivalence relation (predicate \<open>equivp\<close>),
  unless the user specifies explicitly \<open>partial\<close> in which case the obligation
  is \<open>part_equivp\<close>. A quotient defined with \<open>partial\<close> is weaker in the sense
  that less things can be proved automatically.

  The command internally proves a Quotient theorem and sets up the Lifting
  package by the command @{command (HOL) setup_lifting}. Thus the Lifting and
  Transfer packages can be used also with quotient types defined by @{command
  (HOL) "quotient_type"} without any extra set-up. The parametricity theorem
  for the equivalence relation R can be provided as an extra argument of the
  command and is passed to the corresponding internal call of @{command (HOL)
  setup_lifting}. This theorem allows the Lifting package to generate a
  stronger transfer rule for equality.
\<close>


subsection \<open>Lifting package \label{sec:lifting}\<close>

text \<open>
  The Lifting package allows users to lift terms of the raw type to the
  abstract type, which is a necessary step in building a library for an
  abstract type. Lifting defines a new constant by combining coercion
  functions (\<^term>\<open>Abs\<close> and \<^term>\<open>Rep\<close>) with the raw term. It also proves an
  appropriate transfer rule for the Transfer (\secref{sec:transfer}) package
  andif possible, an equation for the code generator.

  The Lifting package provides two main commands: @{command (HOL)
  "setup_lifting"for initializing the package to work with a new type, and
  @{command (HOL) "lift_definition"for lifting constants. The Lifting
  package works with all four kinds of type abstraction: type copies,
  subtypes, total quotients and partial quotients.

  Theoretical background can be found in @{cite
  "Huffman-Kuncar:2013:lifting_transfer"}.

  \begin{matharray}{rcl}
    @{command_def (HOL) "setup_lifting"} & : & \<open>local_theory \<rightarrow> local_theory\<close>\\
    @{command_def (HOL) "lift_definition"} & : & \<open>local_theory \<rightarrow> proof(prove)\<close>\\
    @{command_def (HOL) "lifting_forget"} & : & \<open>local_theory \<rightarrow> local_theory\<close>\\
    @{command_def (HOL) "lifting_update"} & : & \<open>local_theory \<rightarrow> local_theory\<close>\\
    @{command_def (HOL) "print_quot_maps"} & : & \<open>context \<rightarrow>\<close>\\
    @{command_def (HOL) "print_quotients"} & : & \<open>context \<rightarrow>\<close>\\
    @{attribute_def (HOL) "quot_map"} & : & \<open>attribute\<close> \\
    @{attribute_def (HOL) "relator_eq_onp"} & : & \<open>attribute\<close> \\
    @{attribute_def (HOL) "relator_mono"} & : & \<open>attribute\<close> \\
    @{attribute_def (HOL) "relator_distr"} & : & \<open>attribute\<close> \\
    @{attribute_def (HOL) "quot_del"} & : & \<open>attribute\<close> \\
    @{attribute_def (HOL) "lifting_restore"} & : & \<open>attribute\<close> \\
  \end{matharray}

  \<^rail>\<open>
    @@{command (HOL) setup_lifting} @{syntax thm} @{syntax thm}? \<newline>
      (@'parametric' @{syntax thm})?
    ;
    @@{command (HOL) lift_definition} ('(' 'code_dt' ')')? \<newline>
      @{syntax name} '::' @{syntax type} @{syntax mixfix}? 'is' @{syntax term\<newline>
      (@'parametric' (@{syntax thm}+))?
    ;
    @@{command (HOL) lifting_forget} @{syntax name}
    ;
    @@{command (HOL) lifting_update} @{syntax name}
    ;
    @@{attribute (HOL) lifting_restore}
      @{syntax thm} (@{syntax thm} @{syntax thm})?
  \<close>

  \<^descr> @{command (HOL) "setup_lifting"} Sets up the Lifting package to work with
  a user-defined type. The command supports two modes.

    \<^enum> The first one is a low-level mode when the user must provide as a first
    argument of @{command (HOL) "setup_lifting"} a quotient theorem \<^term>\<open>Quotient R Abs Rep T\<close>. The package configures a transfer rule for
    equality, a domain transfer rules and sets up the @{command_def (HOL)
    "lift_definition"} command to work with the abstract type. An optional
    theorem \<^term>\<open>reflp R\<close>, which certifies that the equivalence relation R
    is total, can be provided as a second argument. This allows the package to
    generate stronger transfer rules. And finally, the parametricity theorem
    for \<^term>\<open>R\<close> can be provided as a third argument. This allows the package
    to generate a stronger transfer rule for equality.

    Users generally will not prove the \<open>Quotient\<close> theorem manually for new
    types, as special commands exist to automate the process.

    \<^enum> When a new subtype is defined by @{command (HOL) typedef}, @{command
    (HOL) "lift_definition"} can be used in its second mode, where only the
    \<^term>\<open>type_definition\<close> theorem \<^term>\<open>type_definition Rep Abs A\<close> is
    used as an argument of the command. The command internally proves the
    corresponding \<^term>\<open>Quotient\<close> theorem and registers it with @{command
    (HOL) setup_lifting} using its first mode.

--> --------------------

--> maximum size reached

--> --------------------

¤ Dauer der Verarbeitung: 0.28 Sekunden  (vorverarbeitet)  ¤





aufgebrochen in jeweils 16 Zeichen
Quellennavigators
aufgebrochen in jeweils 16 Zeichen
sprechenden Kalenders

in der Quellcodebibliothek suchen




Haftungshinweis

Die Informationen auf dieser Webseite wurden nach bestem Wissen sorgfältig zusammengestellt. Es wird jedoch weder Vollständigkeit, noch Richtigkeit, noch Qualität der bereit gestellten Informationen zugesichert.


Bemerkung:

Die farbliche Syntaxdarstellung ist noch experimentell.


Bot Zugriff