closer correspondence of document and session names, while maintaining document names for external reference
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Implementation/Base.thy Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,7 @@
+theory Base
+imports Main
+begin
+
+ML_file "../antiquote_setup.ML"
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Implementation/Eq.thy Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,126 @@
+theory Eq
+imports Base
+begin
+
+chapter {* Equational reasoning *}
+
+text {* Equality is one of the most fundamental concepts of
+ mathematics. The Isabelle/Pure logic (\chref{ch:logic}) provides a
+ builtin relation @{text "\<equiv> :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> prop"} that expresses equality
+ of arbitrary terms (or propositions) at the framework level, as
+ expressed by certain basic inference rules (\secref{sec:eq-rules}).
+
+ Equational reasoning means to replace equals by equals, using
+ reflexivity and transitivity to form chains of replacement steps,
+ and congruence rules to access sub-structures. Conversions
+ (\secref{sec:conv}) provide a convenient framework to compose basic
+ equational steps to build specific equational reasoning tools.
+
+ Higher-order matching is able to provide suitable instantiations for
+ giving equality rules, which leads to the versatile concept of
+ @{text "\<lambda>"}-term rewriting (\secref{sec:rewriting}). Internally
+ this is based on the general-purpose Simplifier engine of Isabelle,
+ which is more specific and more efficient than plain conversions.
+
+ Object-logics usually introduce specific notions of equality or
+ equivalence, and relate it with the Pure equality. This enables to
+ re-use the Pure tools for equational reasoning for particular
+ object-logic connectives as well.
+*}
+
+
+section {* Basic equality rules \label{sec:eq-rules} *}
+
+text {* Isabelle/Pure uses @{text "\<equiv>"} for equality of arbitrary
+ terms, which includes equivalence of propositions of the logical
+ framework. The conceptual axiomatization of the constant @{text "\<equiv>
+ :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> prop"} is given in \figref{fig:pure-equality}. The
+ inference kernel presents slightly different equality rules, which
+ may be understood as derived rules from this minimal axiomatization.
+ The Pure theory also provides some theorems that express the same
+ reasoning schemes as theorems that can be composed like object-level
+ rules as explained in \secref{sec:obj-rules}.
+
+ For example, @{ML Thm.symmetric} as Pure inference is an ML function
+ that maps a theorem @{text "th"} stating @{text "t \<equiv> u"} to one
+ stating @{text "u \<equiv> t"}. In contrast, @{thm [source]
+ Pure.symmetric} as Pure theorem expresses the same reasoning in
+ declarative form. If used like @{text "th [THEN Pure.symmetric]"}
+ in Isar source notation, it achieves a similar effect as the ML
+ inference function, although the rule attribute @{attribute THEN} or
+ ML operator @{ML "op RS"} involve the full machinery of higher-order
+ unification (modulo @{text "\<beta>\<eta>"}-conversion) and lifting of @{text
+ "\<And>/\<Longrightarrow>"} contexts. *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML Thm.reflexive: "cterm -> thm"} \\
+ @{index_ML Thm.symmetric: "thm -> thm"} \\
+ @{index_ML Thm.transitive: "thm -> thm -> thm"} \\
+ @{index_ML Thm.abstract_rule: "string -> cterm -> thm -> thm"} \\
+ @{index_ML Thm.combination: "thm -> thm -> thm"} \\[0.5ex]
+ @{index_ML Thm.equal_intr: "thm -> thm -> thm"} \\
+ @{index_ML Thm.equal_elim: "thm -> thm -> thm"} \\
+ \end{mldecls}
+
+ See also @{file "~~/src/Pure/thm.ML" } for further description of
+ these inference rules, and a few more for primitive @{text "\<beta>"} and
+ @{text "\<eta>"} conversions. Note that @{text "\<alpha>"} conversion is
+ implicit due to the representation of terms with de-Bruijn indices
+ (\secref{sec:terms}). *}
+
+
+section {* Conversions \label{sec:conv} *}
+
+text {*
+ %FIXME
+
+ The classic article that introduces the concept of conversion (for
+ Cambridge LCF) is \cite{paulson:1983}.
+*}
+
+
+section {* Rewriting \label{sec:rewriting} *}
+
+text {* Rewriting normalizes a given term (theorem or goal) by
+ replacing instances of given equalities @{text "t \<equiv> u"} in subterms.
+ Rewriting continues until no rewrites are applicable to any subterm.
+ This may be used to unfold simple definitions of the form @{text "f
+ x\<^sub>1 \<dots> x\<^sub>n \<equiv> u"}, but is slightly more general than that.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML rewrite_rule: "Proof.context -> thm list -> thm -> thm"} \\
+ @{index_ML rewrite_goals_rule: "Proof.context -> thm list -> thm -> thm"} \\
+ @{index_ML rewrite_goal_tac: "Proof.context -> thm list -> int -> tactic"} \\
+ @{index_ML rewrite_goals_tac: "Proof.context -> thm list -> tactic"} \\
+ @{index_ML fold_goals_tac: "Proof.context -> thm list -> tactic"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML rewrite_rule}~@{text "ctxt rules thm"} rewrites the whole
+ theorem by the given rules.
+
+ \item @{ML rewrite_goals_rule}~@{text "ctxt rules thm"} rewrites the
+ outer premises of the given theorem. Interpreting the same as a
+ goal state (\secref{sec:tactical-goals}) it means to rewrite all
+ subgoals (in the same manner as @{ML rewrite_goals_tac}).
+
+ \item @{ML rewrite_goal_tac}~@{text "ctxt rules i"} rewrites subgoal
+ @{text "i"} by the given rewrite rules.
+
+ \item @{ML rewrite_goals_tac}~@{text "ctxt rules"} rewrites all subgoals
+ by the given rewrite rules.
+
+ \item @{ML fold_goals_tac}~@{text "ctxt rules"} essentially uses @{ML
+ rewrite_goals_tac} with the symmetric form of each member of @{text
+ "rules"}, re-ordered to fold longer expression first. This supports
+ to idea to fold primitive definitions that appear in expended form
+ in the proof state.
+
+ \end{description}
+*}
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Implementation/Integration.thy Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,291 @@
+theory Integration
+imports Base
+begin
+
+chapter {* System integration *}
+
+section {* Isar toplevel \label{sec:isar-toplevel} *}
+
+text {* The Isar toplevel may be considered the central hub of the
+ Isabelle/Isar system, where all key components and sub-systems are
+ integrated into a single read-eval-print loop of Isar commands,
+ which also incorporates the underlying ML compiler.
+
+ Isabelle/Isar departs from the original ``LCF system architecture''
+ where ML was really The Meta Language for defining theories and
+ conducting proofs. Instead, ML now only serves as the
+ implementation language for the system (and user extensions), while
+ the specific Isar toplevel supports the concepts of theory and proof
+ development natively. This includes the graph structure of theories
+ and the block structure of proofs, support for unlimited undo,
+ facilities for tracing, debugging, timing, profiling etc.
+
+ \medskip The toplevel maintains an implicit state, which is
+ transformed by a sequence of transitions -- either interactively or
+ in batch-mode.
+
+ The toplevel state is a disjoint sum of empty @{text toplevel}, or
+ @{text theory}, or @{text proof}. On entering the main Isar loop we
+ start with an empty toplevel. A theory is commenced by giving a
+ @{text \<THEORY>} header; within a theory we may issue theory
+ commands such as @{text \<DEFINITION>}, or state a @{text
+ \<THEOREM>} to be proven. Now we are within a proof state, with a
+ rich collection of Isar proof commands for structured proof
+ composition, or unstructured proof scripts. When the proof is
+ concluded we get back to the theory, which is then updated by
+ storing the resulting fact. Further theory declarations or theorem
+ statements with proofs may follow, until we eventually conclude the
+ theory development by issuing @{text \<END>}. The resulting theory
+ is then stored within the theory database and we are back to the
+ empty toplevel.
+
+ In addition to these proper state transformations, there are also
+ some diagnostic commands for peeking at the toplevel state without
+ modifying it (e.g.\ \isakeyword{thm}, \isakeyword{term},
+ \isakeyword{print-cases}).
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type Toplevel.state} \\
+ @{index_ML_exception Toplevel.UNDEF} \\
+ @{index_ML Toplevel.is_toplevel: "Toplevel.state -> bool"} \\
+ @{index_ML Toplevel.theory_of: "Toplevel.state -> theory"} \\
+ @{index_ML Toplevel.proof_of: "Toplevel.state -> Proof.state"} \\
+ @{index_ML Toplevel.timing: "bool Unsynchronized.ref"} \\
+ @{index_ML Toplevel.profiling: "int Unsynchronized.ref"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type Toplevel.state} represents Isar toplevel
+ states, which are normally manipulated through the concept of
+ toplevel transitions only (\secref{sec:toplevel-transition}). Also
+ note that a raw toplevel state is subject to the same linearity
+ restrictions as a theory context (cf.~\secref{sec:context-theory}).
+
+ \item @{ML Toplevel.UNDEF} is raised for undefined toplevel
+ operations. Many operations work only partially for certain cases,
+ since @{ML_type Toplevel.state} is a sum type.
+
+ \item @{ML Toplevel.is_toplevel}~@{text "state"} checks for an empty
+ toplevel state.
+
+ \item @{ML Toplevel.theory_of}~@{text "state"} selects the
+ background theory of @{text "state"}, raises @{ML Toplevel.UNDEF}
+ for an empty toplevel state.
+
+ \item @{ML Toplevel.proof_of}~@{text "state"} selects the Isar proof
+ state if available, otherwise raises @{ML Toplevel.UNDEF}.
+
+ \item @{ML "Toplevel.timing := true"} makes the toplevel print timing
+ information for each Isar command being executed.
+
+ \item @{ML Toplevel.profiling}~@{ML_text ":="}~@{text "n"} controls
+ low-level profiling of the underlying ML runtime system. For
+ Poly/ML, @{text "n = 1"} means time and @{text "n = 2"} space
+ profiling.
+
+ \end{description}
+*}
+
+text %mlantiq {*
+ \begin{matharray}{rcl}
+ @{ML_antiquotation_def "Isar.state"} & : & @{text ML_antiquotation} \\
+ \end{matharray}
+
+ \begin{description}
+
+ \item @{text "@{Isar.state}"} refers to Isar toplevel state at that
+ point --- as abstract value.
+
+ This only works for diagnostic ML commands, such as @{command
+ ML_val} or @{command ML_command}.
+
+ \end{description}
+*}
+
+
+subsection {* Toplevel transitions \label{sec:toplevel-transition} *}
+
+text {*
+ An Isar toplevel transition consists of a partial function on the
+ toplevel state, with additional information for diagnostics and
+ error reporting: there are fields for command name, source position,
+ optional source text, as well as flags for interactive-only commands
+ (which issue a warning in batch-mode), printing of result state,
+ etc.
+
+ The operational part is represented as the sequential union of a
+ list of partial functions, which are tried in turn until the first
+ one succeeds. This acts like an outer case-expression for various
+ alternative state transitions. For example, \isakeyword{qed} works
+ differently for a local proofs vs.\ the global ending of the main
+ proof.
+
+ Toplevel transitions are composed via transition transformers.
+ Internally, Isar commands are put together from an empty transition
+ extended by name and source position. It is then left to the
+ individual command parser to turn the given concrete syntax into a
+ suitable transition transformer that adjoins actual operations on a
+ theory or proof state etc.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML Toplevel.print: "Toplevel.transition -> Toplevel.transition"} \\
+ @{index_ML Toplevel.keep: "(Toplevel.state -> unit) ->
+ Toplevel.transition -> Toplevel.transition"} \\
+ @{index_ML Toplevel.theory: "(theory -> theory) ->
+ Toplevel.transition -> Toplevel.transition"} \\
+ @{index_ML Toplevel.theory_to_proof: "(theory -> Proof.state) ->
+ Toplevel.transition -> Toplevel.transition"} \\
+ @{index_ML Toplevel.proof: "(Proof.state -> Proof.state) ->
+ Toplevel.transition -> Toplevel.transition"} \\
+ @{index_ML Toplevel.proofs: "(Proof.state -> Proof.state Seq.result Seq.seq) ->
+ Toplevel.transition -> Toplevel.transition"} \\
+ @{index_ML Toplevel.end_proof: "(bool -> Proof.state -> Proof.context) ->
+ Toplevel.transition -> Toplevel.transition"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML Toplevel.print}~@{text "tr"} sets the print flag, which
+ causes the toplevel loop to echo the result state (in interactive
+ mode).
+
+ \item @{ML Toplevel.keep}~@{text "tr"} adjoins a diagnostic
+ function.
+
+ \item @{ML Toplevel.theory}~@{text "tr"} adjoins a theory
+ transformer.
+
+ \item @{ML Toplevel.theory_to_proof}~@{text "tr"} adjoins a global
+ goal function, which turns a theory into a proof state. The theory
+ may be changed before entering the proof; the generic Isar goal
+ setup includes an argument that specifies how to apply the proven
+ result to the theory, when the proof is finished.
+
+ \item @{ML Toplevel.proof}~@{text "tr"} adjoins a deterministic
+ proof command, with a singleton result.
+
+ \item @{ML Toplevel.proofs}~@{text "tr"} adjoins a general proof
+ command, with zero or more result states (represented as a lazy
+ list).
+
+ \item @{ML Toplevel.end_proof}~@{text "tr"} adjoins a concluding
+ proof command, that returns the resulting theory, after storing the
+ resulting facts in the context etc.
+
+ \end{description}
+*}
+
+
+section {* Theory database \label{sec:theory-database} *}
+
+text {*
+ The theory database maintains a collection of theories, together
+ with some administrative information about their original sources,
+ which are held in an external store (i.e.\ some directory within the
+ regular file system).
+
+ The theory database is organized as a directed acyclic graph;
+ entries are referenced by theory name. Although some additional
+ interfaces allow to include a directory specification as well, this
+ is only a hint to the underlying theory loader. The internal theory
+ name space is flat!
+
+ Theory @{text A} is associated with the main theory file @{text
+ A}\verb,.thy,, which needs to be accessible through the theory
+ loader path. Any number of additional ML source files may be
+ associated with each theory, by declaring these dependencies in the
+ theory header as @{text \<USES>}, and loading them consecutively
+ within the theory context. The system keeps track of incoming ML
+ sources and associates them with the current theory.
+
+ The basic internal actions of the theory database are @{text
+ "update"} and @{text "remove"}:
+
+ \begin{itemize}
+
+ \item @{text "update A"} introduces a link of @{text "A"} with a
+ @{text "theory"} value of the same name; it asserts that the theory
+ sources are now consistent with that value;
+
+ \item @{text "remove A"} deletes entry @{text "A"} from the theory
+ database.
+
+ \end{itemize}
+
+ These actions are propagated to sub- or super-graphs of a theory
+ entry as expected, in order to preserve global consistency of the
+ state of all loaded theories with the sources of the external store.
+ This implies certain causalities between actions: @{text "update"}
+ or @{text "remove"} of an entry will @{text "remove"} all
+ descendants.
+
+ \medskip There are separate user-level interfaces to operate on the
+ theory database directly or indirectly. The primitive actions then
+ just happen automatically while working with the system. In
+ particular, processing a theory header @{text "\<THEORY> A
+ \<IMPORTS> B\<^sub>1 \<dots> B\<^sub>n \<BEGIN>"} ensures that the
+ sub-graph of the collective imports @{text "B\<^sub>1 \<dots> B\<^sub>n"}
+ is up-to-date, too. Earlier theories are reloaded as required, with
+ @{text update} actions proceeding in topological order according to
+ theory dependencies. There may be also a wave of implied @{text
+ remove} actions for derived theory nodes until a stable situation
+ is achieved eventually.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML use_thy: "string -> unit"} \\
+ @{index_ML use_thys: "string list -> unit"} \\
+ @{index_ML Thy_Info.get_theory: "string -> theory"} \\
+ @{index_ML Thy_Info.remove_thy: "string -> unit"} \\[1ex]
+ @{index_ML Thy_Info.register_thy: "theory -> unit"} \\[1ex]
+ @{ML_text "datatype action = Update | Remove"} \\
+ @{index_ML Thy_Info.add_hook: "(Thy_Info.action -> string -> unit) -> unit"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML use_thy}~@{text A} ensures that theory @{text A} is fully
+ up-to-date wrt.\ the external file store, reloading outdated
+ ancestors as required. In batch mode, the simultaneous @{ML
+ use_thys} should be used exclusively.
+
+ \item @{ML use_thys} is similar to @{ML use_thy}, but handles
+ several theories simultaneously. Thus it acts like processing the
+ import header of a theory, without performing the merge of the
+ result. By loading a whole sub-graph of theories like that, the
+ intrinsic parallelism can be exploited by the system, to speedup
+ loading.
+
+ \item @{ML Thy_Info.get_theory}~@{text A} retrieves the theory value
+ presently associated with name @{text A}. Note that the result
+ might be outdated.
+
+ \item @{ML Thy_Info.remove_thy}~@{text A} deletes theory @{text A} and all
+ descendants from the theory database.
+
+ \item @{ML Thy_Info.register_thy}~@{text "text thy"} registers an
+ existing theory value with the theory loader database and updates
+ source version information according to the current file-system
+ state.
+
+ \item @{ML "Thy_Info.add_hook"}~@{text f} registers function @{text
+ f} as a hook for theory database actions. The function will be
+ invoked with the action and theory name being involved; thus derived
+ actions may be performed in associated system components, e.g.\
+ maintaining the state of an editor for the theory sources.
+
+ The kind and order of actions occurring in practice depends both on
+ user interactions and the internal process of resolving theory
+ imports. Hooks should not rely on a particular policy here! Any
+ exceptions raised by the hook are ignored.
+
+ \end{description}
+*}
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Implementation/Isar.thy Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,585 @@
+theory Isar
+imports Base
+begin
+
+chapter {* Isar language elements *}
+
+text {* The Isar proof language (see also
+ \cite[\S2]{isabelle-isar-ref}) consists of three main categories of
+ language elements as follows.
+
+ \begin{enumerate}
+
+ \item Proof \emph{commands} define the primary language of
+ transactions of the underlying Isar/VM interpreter. Typical
+ examples are @{command "fix"}, @{command "assume"}, @{command
+ "show"}, @{command "proof"}, and @{command "qed"}.
+
+ Composing proof commands according to the rules of the Isar/VM leads
+ to expressions of structured proof text, such that both the machine
+ and the human reader can give it a meaning as formal reasoning.
+
+ \item Proof \emph{methods} define a secondary language of mixed
+ forward-backward refinement steps involving facts and goals.
+ Typical examples are @{method rule}, @{method unfold}, and @{method
+ simp}.
+
+ Methods can occur in certain well-defined parts of the Isar proof
+ language, say as arguments to @{command "proof"}, @{command "qed"},
+ or @{command "by"}.
+
+ \item \emph{Attributes} define a tertiary language of small
+ annotations to theorems being defined or referenced. Attributes can
+ modify both the context and the theorem.
+
+ Typical examples are @{attribute intro} (which affects the context),
+ and @{attribute symmetric} (which affects the theorem).
+
+ \end{enumerate}
+*}
+
+
+section {* Proof commands *}
+
+text {* A \emph{proof command} is state transition of the Isar/VM
+ proof interpreter.
+
+ In principle, Isar proof commands could be defined in user-space as
+ well. The system is built like that in the first place: one part of
+ the commands are primitive, the other part is defined as derived
+ elements. Adding to the genuine structured proof language requires
+ profound understanding of the Isar/VM machinery, though, so this is
+ beyond the scope of this manual.
+
+ What can be done realistically is to define some diagnostic commands
+ that inspect the general state of the Isar/VM, and report some
+ feedback to the user. Typically this involves checking of the
+ linguistic \emph{mode} of a proof state, or peeking at the pending
+ goals (if available).
+
+ Another common application is to define a toplevel command that
+ poses a problem to the user as Isar proof state and processes the
+ final result relatively to the context. Thus a proof can be
+ incorporated into the context of some user-space tool, without
+ modifying the Isar proof language itself. *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type Proof.state} \\
+ @{index_ML Proof.assert_forward: "Proof.state -> Proof.state"} \\
+ @{index_ML Proof.assert_chain: "Proof.state -> Proof.state"} \\
+ @{index_ML Proof.assert_backward: "Proof.state -> Proof.state"} \\
+ @{index_ML Proof.simple_goal: "Proof.state -> {context: Proof.context, goal: thm}"} \\
+ @{index_ML Proof.goal: "Proof.state ->
+ {context: Proof.context, facts: thm list, goal: thm}"} \\
+ @{index_ML Proof.raw_goal: "Proof.state ->
+ {context: Proof.context, facts: thm list, goal: thm}"} \\
+ @{index_ML Proof.theorem: "Method.text option ->
+ (thm list list -> Proof.context -> Proof.context) ->
+ (term * term list) list list -> Proof.context -> Proof.state"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type Proof.state} represents Isar proof states.
+ This is a block-structured configuration with proof context,
+ linguistic mode, and optional goal. The latter consists of goal
+ context, goal facts (``@{text "using"}''), and tactical goal state
+ (see \secref{sec:tactical-goals}).
+
+ The general idea is that the facts shall contribute to the
+ refinement of some parts of the tactical goal --- how exactly is
+ defined by the proof method that is applied in that situation.
+
+ \item @{ML Proof.assert_forward}, @{ML Proof.assert_chain}, @{ML
+ Proof.assert_backward} are partial identity functions that fail
+ unless a certain linguistic mode is active, namely ``@{text
+ "proof(state)"}'', ``@{text "proof(chain)"}'', ``@{text
+ "proof(prove)"}'', respectively (using the terminology of
+ \cite{isabelle-isar-ref}).
+
+ It is advisable study the implementations of existing proof commands
+ for suitable modes to be asserted.
+
+ \item @{ML Proof.simple_goal}~@{text "state"} returns the structured
+ Isar goal (if available) in the form seen by ``simple'' methods
+ (like @{method simp} or @{method blast}). The Isar goal facts are
+ already inserted as premises into the subgoals, which are presented
+ individually as in @{ML Proof.goal}.
+
+ \item @{ML Proof.goal}~@{text "state"} returns the structured Isar
+ goal (if available) in the form seen by regular methods (like
+ @{method rule}). The auxiliary internal encoding of Pure
+ conjunctions is split into individual subgoals as usual.
+
+ \item @{ML Proof.raw_goal}~@{text "state"} returns the structured
+ Isar goal (if available) in the raw internal form seen by ``raw''
+ methods (like @{method induct}). This form is rarely appropriate
+ for dignostic tools; @{ML Proof.simple_goal} or @{ML Proof.goal}
+ should be used in most situations.
+
+ \item @{ML Proof.theorem}~@{text "before_qed after_qed statement ctxt"}
+ initializes a toplevel Isar proof state within a given context.
+
+ The optional @{text "before_qed"} method is applied at the end of
+ the proof, just before extracting the result (this feature is rarely
+ used).
+
+ The @{text "after_qed"} continuation receives the extracted result
+ in order to apply it to the final context in a suitable way (e.g.\
+ storing named facts). Note that at this generic level the target
+ context is specified as @{ML_type Proof.context}, but the usual
+ wrapping of toplevel proofs into command transactions will provide a
+ @{ML_type local_theory} here (\chref{ch:local-theory}). This
+ affects the way how results are stored.
+
+ The @{text "statement"} is given as a nested list of terms, each
+ associated with optional @{keyword "is"} patterns as usual in the
+ Isar source language. The original nested list structure over terms
+ is turned into one over theorems when @{text "after_qed"} is
+ invoked.
+
+ \end{description}
+*}
+
+
+text %mlantiq {*
+ \begin{matharray}{rcl}
+ @{ML_antiquotation_def "Isar.goal"} & : & @{text ML_antiquotation} \\
+ \end{matharray}
+
+ \begin{description}
+
+ \item @{text "@{Isar.goal}"} refers to the regular goal state (if
+ available) of the current proof state managed by the Isar toplevel
+ --- as abstract value.
+
+ This only works for diagnostic ML commands, such as @{command
+ ML_val} or @{command ML_command}.
+
+ \end{description}
+*}
+
+text %mlex {* The following example peeks at a certain goal configuration. *}
+
+notepad
+begin
+ have A and B and C
+ ML_val {*
+ val n = Thm.nprems_of (#goal @{Isar.goal});
+ @{assert} (n = 3);
+ *}
+ oops
+
+text {* Here we see 3 individual subgoals in the same way as regular
+ proof methods would do. *}
+
+
+section {* Proof methods *}
+
+text {* A @{text "method"} is a function @{text "context \<rightarrow> thm\<^sup>* \<rightarrow> goal
+ \<rightarrow> (cases \<times> goal)\<^sup>*\<^sup>*"} that operates on the full Isar goal
+ configuration with context, goal facts, and tactical goal state and
+ enumerates possible follow-up goal states, with the potential
+ addition of named extensions of the proof context (\emph{cases}).
+ The latter feature is rarely used.
+
+ This means a proof method is like a structurally enhanced tactic
+ (cf.\ \secref{sec:tactics}). The well-formedness conditions for
+ tactics need to hold for methods accordingly, with the following
+ additions.
+
+ \begin{itemize}
+
+ \item Goal addressing is further limited either to operate either
+ uniformly on \emph{all} subgoals, or specifically on the
+ \emph{first} subgoal.
+
+ Exception: old-style tactic emulations that are embedded into the
+ method space, e.g.\ @{method rule_tac}.
+
+ \item A non-trivial method always needs to make progress: an
+ identical follow-up goal state has to be avoided.\footnote{This
+ enables the user to write method expressions like @{text "meth\<^sup>+"}
+ without looping, while the trivial do-nothing case can be recovered
+ via @{text "meth\<^sup>?"}.}
+
+ Exception: trivial stuttering steps, such as ``@{method -}'' or
+ @{method succeed}.
+
+ \item Goal facts passed to the method must not be ignored. If there
+ is no sensible use of facts outside the goal state, facts should be
+ inserted into the subgoals that are addressed by the method.
+
+ \end{itemize}
+
+ \medskip Syntactically, the language of proof methods appears as
+ arguments to Isar commands like @{command "by"} or @{command apply}.
+ User-space additions are reasonably easy by plugging suitable
+ method-valued parser functions into the framework, using the
+ @{command method_setup} command, for example.
+
+ To get a better idea about the range of possibilities, consider the
+ following Isar proof schemes. This is the general form of
+ structured proof text:
+
+ \medskip
+ \begin{tabular}{l}
+ @{command from}~@{text "facts\<^sub>1"}~@{command have}~@{text "props"}~@{command using}~@{text "facts\<^sub>2"} \\
+ @{command proof}~@{text "(initial_method)"} \\
+ \quad@{text "body"} \\
+ @{command qed}~@{text "(terminal_method)"} \\
+ \end{tabular}
+ \medskip
+
+ The goal configuration consists of @{text "facts\<^sub>1"} and
+ @{text "facts\<^sub>2"} appended in that order, and various @{text
+ "props"} being claimed. The @{text "initial_method"} is invoked
+ with facts and goals together and refines the problem to something
+ that is handled recursively in the proof @{text "body"}. The @{text
+ "terminal_method"} has another chance to finish any remaining
+ subgoals, but it does not see the facts of the initial step.
+
+ \medskip This pattern illustrates unstructured proof scripts:
+
+ \medskip
+ \begin{tabular}{l}
+ @{command have}~@{text "props"} \\
+ \quad@{command using}~@{text "facts\<^sub>1"}~@{command apply}~@{text "method\<^sub>1"} \\
+ \quad@{command apply}~@{text "method\<^sub>2"} \\
+ \quad@{command using}~@{text "facts\<^sub>3"}~@{command apply}~@{text "method\<^sub>3"} \\
+ \quad@{command done} \\
+ \end{tabular}
+ \medskip
+
+ The @{text "method\<^sub>1"} operates on the original claim while
+ using @{text "facts\<^sub>1"}. Since the @{command apply} command
+ structurally resets the facts, the @{text "method\<^sub>2"} will
+ operate on the remaining goal state without facts. The @{text
+ "method\<^sub>3"} will see again a collection of @{text
+ "facts\<^sub>3"} that has been inserted into the script explicitly.
+
+ \medskip Empirically, any Isar proof method can be categorized as
+ follows.
+
+ \begin{enumerate}
+
+ \item \emph{Special method with cases} with named context additions
+ associated with the follow-up goal state.
+
+ Example: @{method "induct"}, which is also a ``raw'' method since it
+ operates on the internal representation of simultaneous claims as
+ Pure conjunction (\secref{sec:logic-aux}), instead of separate
+ subgoals (\secref{sec:tactical-goals}).
+
+ \item \emph{Structured method} with strong emphasis on facts outside
+ the goal state.
+
+ Example: @{method "rule"}, which captures the key ideas behind
+ structured reasoning in Isar in purest form.
+
+ \item \emph{Simple method} with weaker emphasis on facts, which are
+ inserted into subgoals to emulate old-style tactical as
+ ``premises''.
+
+ Examples: @{method "simp"}, @{method "blast"}, @{method "auto"}.
+
+ \item \emph{Old-style tactic emulation} with detailed numeric goal
+ addressing and explicit references to entities of the internal goal
+ state (which are otherwise invisible from proper Isar proof text).
+ The naming convention @{text "foo_tac"} makes this special
+ non-standard status clear.
+
+ Example: @{method "rule_tac"}.
+
+ \end{enumerate}
+
+ When implementing proof methods, it is advisable to study existing
+ implementations carefully and imitate the typical ``boiler plate''
+ for context-sensitive parsing and further combinators to wrap-up
+ tactic expressions as methods.\footnote{Aliases or abbreviations of
+ the standard method combinators should be avoided. Note that from
+ Isabelle99 until Isabelle2009 the system did provide various odd
+ combinations of method wrappers that made user applications more
+ complicated than necessary.}
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type Proof.method} \\
+ @{index_ML METHOD_CASES: "(thm list -> cases_tactic) -> Proof.method"} \\
+ @{index_ML METHOD: "(thm list -> tactic) -> Proof.method"} \\
+ @{index_ML SIMPLE_METHOD: "tactic -> Proof.method"} \\
+ @{index_ML SIMPLE_METHOD': "(int -> tactic) -> Proof.method"} \\
+ @{index_ML Method.insert_tac: "thm list -> int -> tactic"} \\
+ @{index_ML Method.setup: "binding -> (Proof.context -> Proof.method) context_parser ->
+ string -> theory -> theory"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type Proof.method} represents proof methods as
+ abstract type.
+
+ \item @{ML METHOD_CASES}~@{text "(fn facts => cases_tactic)"} wraps
+ @{text cases_tactic} depending on goal facts as proof method with
+ cases; the goal context is passed via method syntax.
+
+ \item @{ML METHOD}~@{text "(fn facts => tactic)"} wraps @{text
+ tactic} depending on goal facts as regular proof method; the goal
+ context is passed via method syntax.
+
+ \item @{ML SIMPLE_METHOD}~@{text "tactic"} wraps a tactic that
+ addresses all subgoals uniformly as simple proof method. Goal facts
+ are already inserted into all subgoals before @{text "tactic"} is
+ applied.
+
+ \item @{ML SIMPLE_METHOD'}~@{text "tactic"} wraps a tactic that
+ addresses a specific subgoal as simple proof method that operates on
+ subgoal 1. Goal facts are inserted into the subgoal then the @{text
+ "tactic"} is applied.
+
+ \item @{ML Method.insert_tac}~@{text "facts i"} inserts @{text
+ "facts"} into subgoal @{text "i"}. This is convenient to reproduce
+ part of the @{ML SIMPLE_METHOD} or @{ML SIMPLE_METHOD'} wrapping
+ within regular @{ML METHOD}, for example.
+
+ \item @{ML Method.setup}~@{text "name parser description"} provides
+ the functionality of the Isar command @{command method_setup} as ML
+ function.
+
+ \end{description}
+*}
+
+text %mlex {* See also @{command method_setup} in
+ \cite{isabelle-isar-ref} which includes some abstract examples.
+
+ \medskip The following toy examples illustrate how the goal facts
+ and state are passed to proof methods. The pre-defined proof method
+ called ``@{method tactic}'' wraps ML source of type @{ML_type
+ tactic} (abstracted over @{ML_text facts}). This allows immediate
+ experimentation without parsing of concrete syntax. *}
+
+notepad
+begin
+ assume a: A and b: B
+
+ have "A \<and> B"
+ apply (tactic {* rtac @{thm conjI} 1 *})
+ using a apply (tactic {* resolve_tac facts 1 *})
+ using b apply (tactic {* resolve_tac facts 1 *})
+ done
+
+ have "A \<and> B"
+ using a and b
+ ML_val "@{Isar.goal}"
+ apply (tactic {* Method.insert_tac facts 1 *})
+ apply (tactic {* (rtac @{thm conjI} THEN_ALL_NEW atac) 1 *})
+ done
+end
+
+text {* \medskip The next example implements a method that simplifies
+ the first subgoal by rewrite rules given as arguments. *}
+
+method_setup my_simp = {*
+ Attrib.thms >> (fn thms => fn ctxt =>
+ SIMPLE_METHOD' (fn i =>
+ CHANGED (asm_full_simp_tac
+ (put_simpset HOL_basic_ss ctxt addsimps thms) i)))
+*} "rewrite subgoal by given rules"
+
+text {* The concrete syntax wrapping of @{command method_setup} always
+ passes-through the proof context at the end of parsing, but it is
+ not used in this example.
+
+ The @{ML Attrib.thms} parser produces a list of theorems from the
+ usual Isar syntax involving attribute expressions etc.\ (syntax
+ category @{syntax thmrefs}) \cite{isabelle-isar-ref}. The resulting
+ @{ML_text thms} are added to @{ML HOL_basic_ss} which already
+ contains the basic Simplifier setup for HOL.
+
+ The tactic @{ML asm_full_simp_tac} is the one that is also used in
+ method @{method simp} by default. The extra wrapping by the @{ML
+ CHANGED} tactical ensures progress of simplification: identical goal
+ states are filtered out explicitly to make the raw tactic conform to
+ standard Isar method behaviour.
+
+ \medskip Method @{method my_simp} can be used in Isar proofs like
+ this:
+*}
+
+notepad
+begin
+ fix a b c
+ assume a: "a = b"
+ assume b: "b = c"
+ have "a = c" by (my_simp a b)
+end
+
+text {* Here is a similar method that operates on all subgoals,
+ instead of just the first one. *}
+
+method_setup my_simp_all = {*
+ Attrib.thms >> (fn thms => fn ctxt =>
+ SIMPLE_METHOD
+ (CHANGED
+ (ALLGOALS (asm_full_simp_tac
+ (put_simpset HOL_basic_ss ctxt addsimps thms)))))
+*} "rewrite all subgoals by given rules"
+
+notepad
+begin
+ fix a b c
+ assume a: "a = b"
+ assume b: "b = c"
+ have "a = c" and "c = b" by (my_simp_all a b)
+end
+
+text {* \medskip Apart from explicit arguments, common proof methods
+ typically work with a default configuration provided by the context.
+ As a shortcut to rule management we use a cheap solution via functor
+ @{ML_functor Named_Thms} (see also @{file
+ "~~/src/Pure/Tools/named_thms.ML"}). *}
+
+ML {*
+ structure My_Simps =
+ Named_Thms
+ (val name = @{binding my_simp} val description = "my_simp rule")
+*}
+setup My_Simps.setup
+
+text {* This provides ML access to a list of theorems in canonical
+ declaration order via @{ML My_Simps.get}. The user can add or
+ delete rules via the attribute @{attribute my_simp}. The actual
+ proof method is now defined as before, but we append the explicit
+ arguments and the rules from the context. *}
+
+method_setup my_simp' = {*
+ Attrib.thms >> (fn thms => fn ctxt =>
+ SIMPLE_METHOD' (fn i =>
+ CHANGED (asm_full_simp_tac
+ (put_simpset HOL_basic_ss ctxt
+ addsimps (thms @ My_Simps.get ctxt)) i)))
+*} "rewrite subgoal by given rules and my_simp rules from the context"
+
+text {*
+ \medskip Method @{method my_simp'} can be used in Isar proofs
+ like this:
+*}
+
+notepad
+begin
+ fix a b c
+ assume [my_simp]: "a \<equiv> b"
+ assume [my_simp]: "b \<equiv> c"
+ have "a \<equiv> c" by my_simp'
+end
+
+text {* \medskip The @{method my_simp} variants defined above are
+ ``simple'' methods, i.e.\ the goal facts are merely inserted as goal
+ premises by the @{ML SIMPLE_METHOD'} or @{ML SIMPLE_METHOD} wrapper.
+ For proof methods that are similar to the standard collection of
+ @{method simp}, @{method blast}, @{method fast}, @{method auto}
+ there is little more that can be done.
+
+ Note that using the primary goal facts in the same manner as the
+ method arguments obtained via concrete syntax or the context does
+ not meet the requirement of ``strong emphasis on facts'' of regular
+ proof methods, because rewrite rules as used above can be easily
+ ignored. A proof text ``@{command using}~@{text "foo"}~@{command
+ "by"}~@{text "my_simp"}'' where @{text "foo"} is not used would
+ deceive the reader.
+
+ \medskip The technical treatment of rules from the context requires
+ further attention. Above we rebuild a fresh @{ML_type simpset} from
+ the arguments and \emph{all} rules retrieved from the context on
+ every invocation of the method. This does not scale to really large
+ collections of rules, which easily emerges in the context of a big
+ theory library, for example.
+
+ This is an inherent limitation of the simplistic rule management via
+ functor @{ML_functor Named_Thms}, because it lacks tool-specific
+ storage and retrieval. More realistic applications require
+ efficient index-structures that organize theorems in a customized
+ manner, such as a discrimination net that is indexed by the
+ left-hand sides of rewrite rules. For variations on the Simplifier,
+ re-use of the existing type @{ML_type simpset} is adequate, but
+ scalability would require it be maintained statically within the
+ context data, not dynamically on each tool invocation. *}
+
+
+section {* Attributes \label{sec:attributes} *}
+
+text {* An \emph{attribute} is a function @{text "context \<times> thm \<rightarrow>
+ context \<times> thm"}, which means both a (generic) context and a theorem
+ can be modified simultaneously. In practice this mixed form is very
+ rare, instead attributes are presented either as \emph{declaration
+ attribute:} @{text "thm \<rightarrow> context \<rightarrow> context"} or \emph{rule
+ attribute:} @{text "context \<rightarrow> thm \<rightarrow> thm"}.
+
+ Attributes can have additional arguments via concrete syntax. There
+ is a collection of context-sensitive parsers for various logical
+ entities (types, terms, theorems). These already take care of
+ applying morphisms to the arguments when attribute expressions are
+ moved into a different context (see also \secref{sec:morphisms}).
+
+ When implementing declaration attributes, it is important to operate
+ exactly on the variant of the generic context that is provided by
+ the system, which is either global theory context or local proof
+ context. In particular, the background theory of a local context
+ must not be modified in this situation! *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type attribute} \\
+ @{index_ML Thm.rule_attribute: "(Context.generic -> thm -> thm) -> attribute"} \\
+ @{index_ML Thm.declaration_attribute: "
+ (thm -> Context.generic -> Context.generic) -> attribute"} \\
+ @{index_ML Attrib.setup: "binding -> attribute context_parser ->
+ string -> theory -> theory"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type attribute} represents attributes as concrete
+ type alias.
+
+ \item @{ML Thm.rule_attribute}~@{text "(fn context => rule)"} wraps
+ a context-dependent rule (mapping on @{ML_type thm}) as attribute.
+
+ \item @{ML Thm.declaration_attribute}~@{text "(fn thm => decl)"}
+ wraps a theorem-dependent declaration (mapping on @{ML_type
+ Context.generic}) as attribute.
+
+ \item @{ML Attrib.setup}~@{text "name parser description"} provides
+ the functionality of the Isar command @{command attribute_setup} as
+ ML function.
+
+ \end{description}
+*}
+
+text %mlantiq {*
+ \begin{matharray}{rcl}
+ @{ML_antiquotation_def attributes} & : & @{text ML_antiquotation} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{ML_antiquotation attributes} attributes
+ \<close>}
+
+ \begin{description}
+
+ \item @{text "@{attributes [\<dots>]}"} embeds attribute source
+ representation into the ML text, which is particularly useful with
+ declarations like @{ML Local_Theory.note}. Attribute names are
+ internalized at compile time, but the source is unevaluated. This
+ means attributes with formal arguments (types, terms, theorems) may
+ be subject to odd effects of dynamic scoping!
+
+ \end{description}
+*}
+
+text %mlex {* See also @{command attribute_setup} in
+ \cite{isabelle-isar-ref} which includes some abstract examples. *}
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Implementation/Local_Theory.thy Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,168 @@
+theory Local_Theory
+imports Base
+begin
+
+chapter {* Local theory specifications \label{ch:local-theory} *}
+
+text {*
+ A \emph{local theory} combines aspects of both theory and proof
+ context (cf.\ \secref{sec:context}), such that definitional
+ specifications may be given relatively to parameters and
+ assumptions. A local theory is represented as a regular proof
+ context, augmented by administrative data about the \emph{target
+ context}.
+
+ The target is usually derived from the background theory by adding
+ local @{text "\<FIX>"} and @{text "\<ASSUME>"} elements, plus
+ suitable modifications of non-logical context data (e.g.\ a special
+ type-checking discipline). Once initialized, the target is ready to
+ absorb definitional primitives: @{text "\<DEFINE>"} for terms and
+ @{text "\<NOTE>"} for theorems. Such definitions may get
+ transformed in a target-specific way, but the programming interface
+ hides such details.
+
+ Isabelle/Pure provides target mechanisms for locales, type-classes,
+ type-class instantiations, and general overloading. In principle,
+ users can implement new targets as well, but this rather arcane
+ discipline is beyond the scope of this manual. In contrast,
+ implementing derived definitional packages to be used within a local
+ theory context is quite easy: the interfaces are even simpler and
+ more abstract than the underlying primitives for raw theories.
+
+ Many definitional packages for local theories are available in
+ Isabelle. Although a few old packages only work for global
+ theories, the standard way of implementing definitional packages in
+ Isabelle is via the local theory interface.
+*}
+
+
+section {* Definitional elements *}
+
+text {*
+ There are separate elements @{text "\<DEFINE> c \<equiv> t"} for terms, and
+ @{text "\<NOTE> b = thm"} for theorems. Types are treated
+ implicitly, according to Hindley-Milner discipline (cf.\
+ \secref{sec:variables}). These definitional primitives essentially
+ act like @{text "let"}-bindings within a local context that may
+ already contain earlier @{text "let"}-bindings and some initial
+ @{text "\<lambda>"}-bindings. Thus we gain \emph{dependent definitions}
+ that are relative to an initial axiomatic context. The following
+ diagram illustrates this idea of axiomatic elements versus
+ definitional elements:
+
+ \begin{center}
+ \begin{tabular}{|l|l|l|}
+ \hline
+ & @{text "\<lambda>"}-binding & @{text "let"}-binding \\
+ \hline
+ types & fixed @{text "\<alpha>"} & arbitrary @{text "\<beta>"} \\
+ terms & @{text "\<FIX> x :: \<tau>"} & @{text "\<DEFINE> c \<equiv> t"} \\
+ theorems & @{text "\<ASSUME> a: A"} & @{text "\<NOTE> b = \<^BG>B\<^EN>"} \\
+ \hline
+ \end{tabular}
+ \end{center}
+
+ A user package merely needs to produce suitable @{text "\<DEFINE>"}
+ and @{text "\<NOTE>"} elements according to the application. For
+ example, a package for inductive definitions might first @{text
+ "\<DEFINE>"} a certain predicate as some fixed-point construction,
+ then @{text "\<NOTE>"} a proven result about monotonicity of the
+ functor involved here, and then produce further derived concepts via
+ additional @{text "\<DEFINE>"} and @{text "\<NOTE>"} elements.
+
+ The cumulative sequence of @{text "\<DEFINE>"} and @{text "\<NOTE>"}
+ produced at package runtime is managed by the local theory
+ infrastructure by means of an \emph{auxiliary context}. Thus the
+ system holds up the impression of working within a fully abstract
+ situation with hypothetical entities: @{text "\<DEFINE> c \<equiv> t"}
+ always results in a literal fact @{text "\<^BG>c \<equiv> t\<^EN>"}, where
+ @{text "c"} is a fixed variable @{text "c"}. The details about
+ global constants, name spaces etc. are handled internally.
+
+ So the general structure of a local theory is a sandwich of three
+ layers:
+
+ \begin{center}
+ \framebox{\quad auxiliary context \quad\framebox{\quad target context \quad\framebox{\quad background theory\quad}}}
+ \end{center}
+
+ When a definitional package is finished, the auxiliary context is
+ reset to the target context. The target now holds definitions for
+ terms and theorems that stem from the hypothetical @{text
+ "\<DEFINE>"} and @{text "\<NOTE>"} elements, transformed by the
+ particular target policy (see \cite[\S4--5]{Haftmann-Wenzel:2009}
+ for details). *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type local_theory: Proof.context} \\
+ @{index_ML Named_Target.init: "(local_theory -> local_theory) ->
+ string -> theory -> local_theory"} \\[1ex]
+ @{index_ML Local_Theory.define: "(binding * mixfix) * (Attrib.binding * term) ->
+ local_theory -> (term * (string * thm)) * local_theory"} \\
+ @{index_ML Local_Theory.note: "Attrib.binding * thm list ->
+ local_theory -> (string * thm list) * local_theory"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type local_theory} represents local theories.
+ Although this is merely an alias for @{ML_type Proof.context}, it is
+ semantically a subtype of the same: a @{ML_type local_theory} holds
+ target information as special context data. Subtyping means that
+ any value @{text "lthy:"}~@{ML_type local_theory} can be also used
+ with operations on expecting a regular @{text "ctxt:"}~@{ML_type
+ Proof.context}.
+
+ \item @{ML Named_Target.init}~@{text "before_exit name thy"}
+ initializes a local theory derived from the given background theory.
+ An empty name refers to a \emph{global theory} context, and a
+ non-empty name refers to a @{command locale} or @{command class}
+ context (a fully-qualified internal name is expected here). This is
+ useful for experimentation --- normally the Isar toplevel already
+ takes care to initialize the local theory context. The given @{text
+ "before_exit"} function is invoked before leaving the context; in
+ most situations plain identity @{ML I} is sufficient.
+
+ \item @{ML Local_Theory.define}~@{text "((b, mx), (a, rhs))
+ lthy"} defines a local entity according to the specification that is
+ given relatively to the current @{text "lthy"} context. In
+ particular the term of the RHS may refer to earlier local entities
+ from the auxiliary context, or hypothetical parameters from the
+ target context. The result is the newly defined term (which is
+ always a fixed variable with exactly the same name as specified for
+ the LHS), together with an equational theorem that states the
+ definition as a hypothetical fact.
+
+ Unless an explicit name binding is given for the RHS, the resulting
+ fact will be called @{text "b_def"}. Any given attributes are
+ applied to that same fact --- immediately in the auxiliary context
+ \emph{and} in any transformed versions stemming from target-specific
+ policies or any later interpretations of results from the target
+ context (think of @{command locale} and @{command interpretation},
+ for example). This means that attributes should be usually plain
+ declarations such as @{attribute simp}, while non-trivial rules like
+ @{attribute simplified} are better avoided.
+
+ \item @{ML Local_Theory.note}~@{text "(a, ths) lthy"} is
+ analogous to @{ML Local_Theory.define}, but defines facts instead of
+ terms. There is also a slightly more general variant @{ML
+ Local_Theory.notes} that defines several facts (with attribute
+ expressions) simultaneously.
+
+ This is essentially the internal version of the @{command lemmas}
+ command, or @{command declare} if an empty name binding is given.
+
+ \end{description}
+*}
+
+
+section {* Morphisms and declarations \label{sec:morphisms} *}
+
+text {*
+ %FIXME
+
+ See also \cite{Chaieb-Wenzel:2007}.
+*}
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Implementation/Logic.thy Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,1462 @@
+theory Logic
+imports Base
+begin
+
+chapter {* Primitive logic \label{ch:logic} *}
+
+text {*
+ The logical foundations of Isabelle/Isar are that of the Pure logic,
+ which has been introduced as a Natural Deduction framework in
+ \cite{paulson700}. This is essentially the same logic as ``@{text
+ "\<lambda>HOL"}'' in the more abstract setting of Pure Type Systems (PTS)
+ \cite{Barendregt-Geuvers:2001}, although there are some key
+ differences in the specific treatment of simple types in
+ Isabelle/Pure.
+
+ Following type-theoretic parlance, the Pure logic consists of three
+ levels of @{text "\<lambda>"}-calculus with corresponding arrows, @{text
+ "\<Rightarrow>"} for syntactic function space (terms depending on terms), @{text
+ "\<And>"} for universal quantification (proofs depending on terms), and
+ @{text "\<Longrightarrow>"} for implication (proofs depending on proofs).
+
+ Derivations are relative to a logical theory, which declares type
+ constructors, constants, and axioms. Theory declarations support
+ schematic polymorphism, which is strictly speaking outside the
+ logic.\footnote{This is the deeper logical reason, why the theory
+ context @{text "\<Theta>"} is separate from the proof context @{text "\<Gamma>"}
+ of the core calculus: type constructors, term constants, and facts
+ (proof constants) may involve arbitrary type schemes, but the type
+ of a locally fixed term parameter is also fixed!}
+*}
+
+
+section {* Types \label{sec:types} *}
+
+text {*
+ The language of types is an uninterpreted order-sorted first-order
+ algebra; types are qualified by ordered type classes.
+
+ \medskip A \emph{type class} is an abstract syntactic entity
+ declared in the theory context. The \emph{subclass relation} @{text
+ "c\<^sub>1 \<subseteq> c\<^sub>2"} is specified by stating an acyclic
+ generating relation; the transitive closure is maintained
+ internally. The resulting relation is an ordering: reflexive,
+ transitive, and antisymmetric.
+
+ A \emph{sort} is a list of type classes written as @{text "s = {c\<^sub>1,
+ \<dots>, c\<^sub>m}"}, it represents symbolic intersection. Notationally, the
+ curly braces are omitted for singleton intersections, i.e.\ any
+ class @{text "c"} may be read as a sort @{text "{c}"}. The ordering
+ on type classes is extended to sorts according to the meaning of
+ intersections: @{text "{c\<^sub>1, \<dots> c\<^sub>m} \<subseteq> {d\<^sub>1, \<dots>, d\<^sub>n}"} iff @{text
+ "\<forall>j. \<exists>i. c\<^sub>i \<subseteq> d\<^sub>j"}. The empty intersection @{text "{}"} refers to
+ the universal sort, which is the largest element wrt.\ the sort
+ order. Thus @{text "{}"} represents the ``full sort'', not the
+ empty one! The intersection of all (finitely many) classes declared
+ in the current theory is the least element wrt.\ the sort ordering.
+
+ \medskip A \emph{fixed type variable} is a pair of a basic name
+ (starting with a @{text "'"} character) and a sort constraint, e.g.\
+ @{text "('a, s)"} which is usually printed as @{text "\<alpha>\<^sub>s"}.
+ A \emph{schematic type variable} is a pair of an indexname and a
+ sort constraint, e.g.\ @{text "(('a, 0), s)"} which is usually
+ printed as @{text "?\<alpha>\<^sub>s"}.
+
+ Note that \emph{all} syntactic components contribute to the identity
+ of type variables: basic name, index, and sort constraint. The core
+ logic handles type variables with the same name but different sorts
+ as different, although the type-inference layer (which is outside
+ the core) rejects anything like that.
+
+ A \emph{type constructor} @{text "\<kappa>"} is a @{text "k"}-ary operator
+ on types declared in the theory. Type constructor application is
+ written postfix as @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>k)\<kappa>"}. For
+ @{text "k = 0"} the argument tuple is omitted, e.g.\ @{text "prop"}
+ instead of @{text "()prop"}. For @{text "k = 1"} the parentheses
+ are omitted, e.g.\ @{text "\<alpha> list"} instead of @{text "(\<alpha>)list"}.
+ Further notation is provided for specific constructors, notably the
+ right-associative infix @{text "\<alpha> \<Rightarrow> \<beta>"} instead of @{text "(\<alpha>,
+ \<beta>)fun"}.
+
+ The logical category \emph{type} is defined inductively over type
+ variables and type constructors as follows: @{text "\<tau> = \<alpha>\<^sub>s | ?\<alpha>\<^sub>s |
+ (\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>k)\<kappa>"}.
+
+ A \emph{type abbreviation} is a syntactic definition @{text
+ "(\<^vec>\<alpha>)\<kappa> = \<tau>"} of an arbitrary type expression @{text "\<tau>"} over
+ variables @{text "\<^vec>\<alpha>"}. Type abbreviations appear as type
+ constructors in the syntax, but are expanded before entering the
+ logical core.
+
+ A \emph{type arity} declares the image behavior of a type
+ constructor wrt.\ the algebra of sorts: @{text "\<kappa> :: (s\<^sub>1, \<dots>,
+ s\<^sub>k)s"} means that @{text "(\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>k)\<kappa>"} is
+ of sort @{text "s"} if every argument type @{text "\<tau>\<^sub>i"} is
+ of sort @{text "s\<^sub>i"}. Arity declarations are implicitly
+ completed, i.e.\ @{text "\<kappa> :: (\<^vec>s)c"} entails @{text "\<kappa> ::
+ (\<^vec>s)c'"} for any @{text "c' \<supseteq> c"}.
+
+ \medskip The sort algebra is always maintained as \emph{coregular},
+ which means that type arities are consistent with the subclass
+ relation: for any type constructor @{text "\<kappa>"}, and classes @{text
+ "c\<^sub>1 \<subseteq> c\<^sub>2"}, and arities @{text "\<kappa> ::
+ (\<^vec>s\<^sub>1)c\<^sub>1"} and @{text "\<kappa> ::
+ (\<^vec>s\<^sub>2)c\<^sub>2"} holds @{text "\<^vec>s\<^sub>1 \<subseteq>
+ \<^vec>s\<^sub>2"} component-wise.
+
+ The key property of a coregular order-sorted algebra is that sort
+ constraints can be solved in a most general fashion: for each type
+ constructor @{text "\<kappa>"} and sort @{text "s"} there is a most general
+ vector of argument sorts @{text "(s\<^sub>1, \<dots>, s\<^sub>k)"} such
+ that a type scheme @{text "(\<alpha>\<^bsub>s\<^sub>1\<^esub>, \<dots>,
+ \<alpha>\<^bsub>s\<^sub>k\<^esub>)\<kappa>"} is of sort @{text "s"}.
+ Consequently, type unification has most general solutions (modulo
+ equivalence of sorts), so type-inference produces primary types as
+ expected \cite{nipkow-prehofer}.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type class: string} \\
+ @{index_ML_type sort: "class list"} \\
+ @{index_ML_type arity: "string * sort list * sort"} \\
+ @{index_ML_type typ} \\
+ @{index_ML Term.map_atyps: "(typ -> typ) -> typ -> typ"} \\
+ @{index_ML Term.fold_atyps: "(typ -> 'a -> 'a) -> typ -> 'a -> 'a"} \\
+ \end{mldecls}
+ \begin{mldecls}
+ @{index_ML Sign.subsort: "theory -> sort * sort -> bool"} \\
+ @{index_ML Sign.of_sort: "theory -> typ * sort -> bool"} \\
+ @{index_ML Sign.add_type: "Proof.context -> binding * int * mixfix -> theory -> theory"} \\
+ @{index_ML Sign.add_type_abbrev: "Proof.context ->
+ binding * string list * typ -> theory -> theory"} \\
+ @{index_ML Sign.primitive_class: "binding * class list -> theory -> theory"} \\
+ @{index_ML Sign.primitive_classrel: "class * class -> theory -> theory"} \\
+ @{index_ML Sign.primitive_arity: "arity -> theory -> theory"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type class} represents type classes.
+
+ \item Type @{ML_type sort} represents sorts, i.e.\ finite
+ intersections of classes. The empty list @{ML "[]: sort"} refers to
+ the empty class intersection, i.e.\ the ``full sort''.
+
+ \item Type @{ML_type arity} represents type arities. A triple
+ @{text "(\<kappa>, \<^vec>s, s) : arity"} represents @{text "\<kappa> ::
+ (\<^vec>s)s"} as described above.
+
+ \item Type @{ML_type typ} represents types; this is a datatype with
+ constructors @{ML TFree}, @{ML TVar}, @{ML Type}.
+
+ \item @{ML Term.map_atyps}~@{text "f \<tau>"} applies the mapping @{text
+ "f"} to all atomic types (@{ML TFree}, @{ML TVar}) occurring in
+ @{text "\<tau>"}.
+
+ \item @{ML Term.fold_atyps}~@{text "f \<tau>"} iterates the operation
+ @{text "f"} over all occurrences of atomic types (@{ML TFree}, @{ML
+ TVar}) in @{text "\<tau>"}; the type structure is traversed from left to
+ right.
+
+ \item @{ML Sign.subsort}~@{text "thy (s\<^sub>1, s\<^sub>2)"}
+ tests the subsort relation @{text "s\<^sub>1 \<subseteq> s\<^sub>2"}.
+
+ \item @{ML Sign.of_sort}~@{text "thy (\<tau>, s)"} tests whether type
+ @{text "\<tau>"} is of sort @{text "s"}.
+
+ \item @{ML Sign.add_type}~@{text "ctxt (\<kappa>, k, mx)"} declares a
+ new type constructors @{text "\<kappa>"} with @{text "k"} arguments and
+ optional mixfix syntax.
+
+ \item @{ML Sign.add_type_abbrev}~@{text "ctxt (\<kappa>, \<^vec>\<alpha>, \<tau>)"}
+ defines a new type abbreviation @{text "(\<^vec>\<alpha>)\<kappa> = \<tau>"}.
+
+ \item @{ML Sign.primitive_class}~@{text "(c, [c\<^sub>1, \<dots>,
+ c\<^sub>n])"} declares a new class @{text "c"}, together with class
+ relations @{text "c \<subseteq> c\<^sub>i"}, for @{text "i = 1, \<dots>, n"}.
+
+ \item @{ML Sign.primitive_classrel}~@{text "(c\<^sub>1,
+ c\<^sub>2)"} declares the class relation @{text "c\<^sub>1 \<subseteq>
+ c\<^sub>2"}.
+
+ \item @{ML Sign.primitive_arity}~@{text "(\<kappa>, \<^vec>s, s)"} declares
+ the arity @{text "\<kappa> :: (\<^vec>s)s"}.
+
+ \end{description}
+*}
+
+text %mlantiq {*
+ \begin{matharray}{rcl}
+ @{ML_antiquotation_def "class"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "sort"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "type_name"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "type_abbrev"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "nonterminal"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "typ"} & : & @{text ML_antiquotation} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{ML_antiquotation class} nameref
+ ;
+ @@{ML_antiquotation sort} sort
+ ;
+ (@@{ML_antiquotation type_name} |
+ @@{ML_antiquotation type_abbrev} |
+ @@{ML_antiquotation nonterminal}) nameref
+ ;
+ @@{ML_antiquotation typ} type
+ \<close>}
+
+ \begin{description}
+
+ \item @{text "@{class c}"} inlines the internalized class @{text
+ "c"} --- as @{ML_type string} literal.
+
+ \item @{text "@{sort s}"} inlines the internalized sort @{text "s"}
+ --- as @{ML_type "string list"} literal.
+
+ \item @{text "@{type_name c}"} inlines the internalized type
+ constructor @{text "c"} --- as @{ML_type string} literal.
+
+ \item @{text "@{type_abbrev c}"} inlines the internalized type
+ abbreviation @{text "c"} --- as @{ML_type string} literal.
+
+ \item @{text "@{nonterminal c}"} inlines the internalized syntactic
+ type~/ grammar nonterminal @{text "c"} --- as @{ML_type string}
+ literal.
+
+ \item @{text "@{typ \<tau>}"} inlines the internalized type @{text "\<tau>"}
+ --- as constructor term for datatype @{ML_type typ}.
+
+ \end{description}
+*}
+
+
+section {* Terms \label{sec:terms} *}
+
+text {*
+ The language of terms is that of simply-typed @{text "\<lambda>"}-calculus
+ with de-Bruijn indices for bound variables (cf.\ \cite{debruijn72}
+ or \cite{paulson-ml2}), with the types being determined by the
+ corresponding binders. In contrast, free variables and constants
+ have an explicit name and type in each occurrence.
+
+ \medskip A \emph{bound variable} is a natural number @{text "b"},
+ which accounts for the number of intermediate binders between the
+ variable occurrence in the body and its binding position. For
+ example, the de-Bruijn term @{text "\<lambda>\<^bsub>bool\<^esub>. \<lambda>\<^bsub>bool\<^esub>. 1 \<and> 0"} would
+ correspond to @{text "\<lambda>x\<^bsub>bool\<^esub>. \<lambda>y\<^bsub>bool\<^esub>. x \<and> y"} in a named
+ representation. Note that a bound variable may be represented by
+ different de-Bruijn indices at different occurrences, depending on
+ the nesting of abstractions.
+
+ A \emph{loose variable} is a bound variable that is outside the
+ scope of local binders. The types (and names) for loose variables
+ can be managed as a separate context, that is maintained as a stack
+ of hypothetical binders. The core logic operates on closed terms,
+ without any loose variables.
+
+ A \emph{fixed variable} is a pair of a basic name and a type, e.g.\
+ @{text "(x, \<tau>)"} which is usually printed @{text "x\<^sub>\<tau>"} here. A
+ \emph{schematic variable} is a pair of an indexname and a type,
+ e.g.\ @{text "((x, 0), \<tau>)"} which is likewise printed as @{text
+ "?x\<^sub>\<tau>"}.
+
+ \medskip A \emph{constant} is a pair of a basic name and a type,
+ e.g.\ @{text "(c, \<tau>)"} which is usually printed as @{text "c\<^sub>\<tau>"}
+ here. Constants are declared in the context as polymorphic families
+ @{text "c :: \<sigma>"}, meaning that all substitution instances @{text
+ "c\<^sub>\<tau>"} for @{text "\<tau> = \<sigma>\<vartheta>"} are valid.
+
+ The vector of \emph{type arguments} of constant @{text "c\<^sub>\<tau>"} wrt.\
+ the declaration @{text "c :: \<sigma>"} is defined as the codomain of the
+ matcher @{text "\<vartheta> = {?\<alpha>\<^sub>1 \<mapsto> \<tau>\<^sub>1, \<dots>, ?\<alpha>\<^sub>n \<mapsto> \<tau>\<^sub>n}"} presented in
+ canonical order @{text "(\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>n)"}, corresponding to the
+ left-to-right occurrences of the @{text "\<alpha>\<^sub>i"} in @{text "\<sigma>"}.
+ Within a given theory context, there is a one-to-one correspondence
+ between any constant @{text "c\<^sub>\<tau>"} and the application @{text "c(\<tau>\<^sub>1,
+ \<dots>, \<tau>\<^sub>n)"} of its type arguments. For example, with @{text "plus :: \<alpha>
+ \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>"}, the instance @{text "plus\<^bsub>nat \<Rightarrow> nat \<Rightarrow> nat\<^esub>"} corresponds to
+ @{text "plus(nat)"}.
+
+ Constant declarations @{text "c :: \<sigma>"} may contain sort constraints
+ for type variables in @{text "\<sigma>"}. These are observed by
+ type-inference as expected, but \emph{ignored} by the core logic.
+ This means the primitive logic is able to reason with instances of
+ polymorphic constants that the user-level type-checker would reject
+ due to violation of type class restrictions.
+
+ \medskip An \emph{atomic term} is either a variable or constant.
+ The logical category \emph{term} is defined inductively over atomic
+ terms, with abstraction and application as follows: @{text "t = b |
+ x\<^sub>\<tau> | ?x\<^sub>\<tau> | c\<^sub>\<tau> | \<lambda>\<^sub>\<tau>. t | t\<^sub>1 t\<^sub>2"}. Parsing and printing takes care of
+ converting between an external representation with named bound
+ variables. Subsequently, we shall use the latter notation instead
+ of internal de-Bruijn representation.
+
+ The inductive relation @{text "t :: \<tau>"} assigns a (unique) type to a
+ term according to the structure of atomic terms, abstractions, and
+ applicatins:
+ \[
+ \infer{@{text "a\<^sub>\<tau> :: \<tau>"}}{}
+ \qquad
+ \infer{@{text "(\<lambda>x\<^sub>\<tau>. t) :: \<tau> \<Rightarrow> \<sigma>"}}{@{text "t :: \<sigma>"}}
+ \qquad
+ \infer{@{text "t u :: \<sigma>"}}{@{text "t :: \<tau> \<Rightarrow> \<sigma>"} & @{text "u :: \<tau>"}}
+ \]
+ A \emph{well-typed term} is a term that can be typed according to these rules.
+
+ Typing information can be omitted: type-inference is able to
+ reconstruct the most general type of a raw term, while assigning
+ most general types to all of its variables and constants.
+ Type-inference depends on a context of type constraints for fixed
+ variables, and declarations for polymorphic constants.
+
+ The identity of atomic terms consists both of the name and the type
+ component. This means that different variables @{text
+ "x\<^bsub>\<tau>\<^sub>1\<^esub>"} and @{text "x\<^bsub>\<tau>\<^sub>2\<^esub>"} may become the same after
+ type instantiation. Type-inference rejects variables of the same
+ name, but different types. In contrast, mixed instances of
+ polymorphic constants occur routinely.
+
+ \medskip The \emph{hidden polymorphism} of a term @{text "t :: \<sigma>"}
+ is the set of type variables occurring in @{text "t"}, but not in
+ its type @{text "\<sigma>"}. This means that the term implicitly depends
+ on type arguments that are not accounted in the result type, i.e.\
+ there are different type instances @{text "t\<vartheta> :: \<sigma>"} and
+ @{text "t\<vartheta>' :: \<sigma>"} with the same type. This slightly
+ pathological situation notoriously demands additional care.
+
+ \medskip A \emph{term abbreviation} is a syntactic definition @{text
+ "c\<^sub>\<sigma> \<equiv> t"} of a closed term @{text "t"} of type @{text "\<sigma>"},
+ without any hidden polymorphism. A term abbreviation looks like a
+ constant in the syntax, but is expanded before entering the logical
+ core. Abbreviations are usually reverted when printing terms, using
+ @{text "t \<rightarrow> c\<^sub>\<sigma>"} as rules for higher-order rewriting.
+
+ \medskip Canonical operations on @{text "\<lambda>"}-terms include @{text
+ "\<alpha>\<beta>\<eta>"}-conversion: @{text "\<alpha>"}-conversion refers to capture-free
+ renaming of bound variables; @{text "\<beta>"}-conversion contracts an
+ abstraction applied to an argument term, substituting the argument
+ in the body: @{text "(\<lambda>x. b)a"} becomes @{text "b[a/x]"}; @{text
+ "\<eta>"}-conversion contracts vacuous application-abstraction: @{text
+ "\<lambda>x. f x"} becomes @{text "f"}, provided that the bound variable
+ does not occur in @{text "f"}.
+
+ Terms are normally treated modulo @{text "\<alpha>"}-conversion, which is
+ implicit in the de-Bruijn representation. Names for bound variables
+ in abstractions are maintained separately as (meaningless) comments,
+ mostly for parsing and printing. Full @{text "\<alpha>\<beta>\<eta>"}-conversion is
+ commonplace in various standard operations (\secref{sec:obj-rules})
+ that are based on higher-order unification and matching.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type term} \\
+ @{index_ML_op "aconv": "term * term -> bool"} \\
+ @{index_ML Term.map_types: "(typ -> typ) -> term -> term"} \\
+ @{index_ML Term.fold_types: "(typ -> 'a -> 'a) -> term -> 'a -> 'a"} \\
+ @{index_ML Term.map_aterms: "(term -> term) -> term -> term"} \\
+ @{index_ML Term.fold_aterms: "(term -> 'a -> 'a) -> term -> 'a -> 'a"} \\
+ \end{mldecls}
+ \begin{mldecls}
+ @{index_ML fastype_of: "term -> typ"} \\
+ @{index_ML lambda: "term -> term -> term"} \\
+ @{index_ML betapply: "term * term -> term"} \\
+ @{index_ML incr_boundvars: "int -> term -> term"} \\
+ @{index_ML Sign.declare_const: "Proof.context ->
+ (binding * typ) * mixfix -> theory -> term * theory"} \\
+ @{index_ML Sign.add_abbrev: "string -> binding * term ->
+ theory -> (term * term) * theory"} \\
+ @{index_ML Sign.const_typargs: "theory -> string * typ -> typ list"} \\
+ @{index_ML Sign.const_instance: "theory -> string * typ list -> typ"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type term} represents de-Bruijn terms, with comments
+ in abstractions, and explicitly named free variables and constants;
+ this is a datatype with constructors @{index_ML Bound}, @{index_ML
+ Free}, @{index_ML Var}, @{index_ML Const}, @{index_ML Abs},
+ @{index_ML_op "$"}.
+
+ \item @{text "t"}~@{ML_text aconv}~@{text "u"} checks @{text
+ "\<alpha>"}-equivalence of two terms. This is the basic equality relation
+ on type @{ML_type term}; raw datatype equality should only be used
+ for operations related to parsing or printing!
+
+ \item @{ML Term.map_types}~@{text "f t"} applies the mapping @{text
+ "f"} to all types occurring in @{text "t"}.
+
+ \item @{ML Term.fold_types}~@{text "f t"} iterates the operation
+ @{text "f"} over all occurrences of types in @{text "t"}; the term
+ structure is traversed from left to right.
+
+ \item @{ML Term.map_aterms}~@{text "f t"} applies the mapping @{text
+ "f"} to all atomic terms (@{ML Bound}, @{ML Free}, @{ML Var}, @{ML
+ Const}) occurring in @{text "t"}.
+
+ \item @{ML Term.fold_aterms}~@{text "f t"} iterates the operation
+ @{text "f"} over all occurrences of atomic terms (@{ML Bound}, @{ML
+ Free}, @{ML Var}, @{ML Const}) in @{text "t"}; the term structure is
+ traversed from left to right.
+
+ \item @{ML fastype_of}~@{text "t"} determines the type of a
+ well-typed term. This operation is relatively slow, despite the
+ omission of any sanity checks.
+
+ \item @{ML lambda}~@{text "a b"} produces an abstraction @{text
+ "\<lambda>a. b"}, where occurrences of the atomic term @{text "a"} in the
+ body @{text "b"} are replaced by bound variables.
+
+ \item @{ML betapply}~@{text "(t, u)"} produces an application @{text
+ "t u"}, with topmost @{text "\<beta>"}-conversion if @{text "t"} is an
+ abstraction.
+
+ \item @{ML incr_boundvars}~@{text "j"} increments a term's dangling
+ bound variables by the offset @{text "j"}. This is required when
+ moving a subterm into a context where it is enclosed by a different
+ number of abstractions. Bound variables with a matching abstraction
+ are unaffected.
+
+ \item @{ML Sign.declare_const}~@{text "ctxt ((c, \<sigma>), mx)"} declares
+ a new constant @{text "c :: \<sigma>"} with optional mixfix syntax.
+
+ \item @{ML Sign.add_abbrev}~@{text "print_mode (c, t)"}
+ introduces a new term abbreviation @{text "c \<equiv> t"}.
+
+ \item @{ML Sign.const_typargs}~@{text "thy (c, \<tau>)"} and @{ML
+ Sign.const_instance}~@{text "thy (c, [\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>n])"}
+ convert between two representations of polymorphic constants: full
+ type instance vs.\ compact type arguments form.
+
+ \end{description}
+*}
+
+text %mlantiq {*
+ \begin{matharray}{rcl}
+ @{ML_antiquotation_def "const_name"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "const_abbrev"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "const"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "term"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "prop"} & : & @{text ML_antiquotation} \\
+ \end{matharray}
+
+ @{rail \<open>
+ (@@{ML_antiquotation const_name} |
+ @@{ML_antiquotation const_abbrev}) nameref
+ ;
+ @@{ML_antiquotation const} ('(' (type + ',') ')')?
+ ;
+ @@{ML_antiquotation term} term
+ ;
+ @@{ML_antiquotation prop} prop
+ \<close>}
+
+ \begin{description}
+
+ \item @{text "@{const_name c}"} inlines the internalized logical
+ constant name @{text "c"} --- as @{ML_type string} literal.
+
+ \item @{text "@{const_abbrev c}"} inlines the internalized
+ abbreviated constant name @{text "c"} --- as @{ML_type string}
+ literal.
+
+ \item @{text "@{const c(\<^vec>\<tau>)}"} inlines the internalized
+ constant @{text "c"} with precise type instantiation in the sense of
+ @{ML Sign.const_instance} --- as @{ML Const} constructor term for
+ datatype @{ML_type term}.
+
+ \item @{text "@{term t}"} inlines the internalized term @{text "t"}
+ --- as constructor term for datatype @{ML_type term}.
+
+ \item @{text "@{prop \<phi>}"} inlines the internalized proposition
+ @{text "\<phi>"} --- as constructor term for datatype @{ML_type term}.
+
+ \end{description}
+*}
+
+
+section {* Theorems \label{sec:thms} *}
+
+text {*
+ A \emph{proposition} is a well-typed term of type @{text "prop"}, a
+ \emph{theorem} is a proven proposition (depending on a context of
+ hypotheses and the background theory). Primitive inferences include
+ plain Natural Deduction rules for the primary connectives @{text
+ "\<And>"} and @{text "\<Longrightarrow>"} of the framework. There is also a builtin
+ notion of equality/equivalence @{text "\<equiv>"}.
+*}
+
+
+subsection {* Primitive connectives and rules \label{sec:prim-rules} *}
+
+text {*
+ The theory @{text "Pure"} contains constant declarations for the
+ primitive connectives @{text "\<And>"}, @{text "\<Longrightarrow>"}, and @{text "\<equiv>"} of
+ the logical framework, see \figref{fig:pure-connectives}. The
+ derivability judgment @{text "A\<^sub>1, \<dots>, A\<^sub>n \<turnstile> B"} is
+ defined inductively by the primitive inferences given in
+ \figref{fig:prim-rules}, with the global restriction that the
+ hypotheses must \emph{not} contain any schematic variables. The
+ builtin equality is conceptually axiomatized as shown in
+ \figref{fig:pure-equality}, although the implementation works
+ directly with derived inferences.
+
+ \begin{figure}[htb]
+ \begin{center}
+ \begin{tabular}{ll}
+ @{text "all :: (\<alpha> \<Rightarrow> prop) \<Rightarrow> prop"} & universal quantification (binder @{text "\<And>"}) \\
+ @{text "\<Longrightarrow> :: prop \<Rightarrow> prop \<Rightarrow> prop"} & implication (right associative infix) \\
+ @{text "\<equiv> :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> prop"} & equality relation (infix) \\
+ \end{tabular}
+ \caption{Primitive connectives of Pure}\label{fig:pure-connectives}
+ \end{center}
+ \end{figure}
+
+ \begin{figure}[htb]
+ \begin{center}
+ \[
+ \infer[@{text "(axiom)"}]{@{text "\<turnstile> A"}}{@{text "A \<in> \<Theta>"}}
+ \qquad
+ \infer[@{text "(assume)"}]{@{text "A \<turnstile> A"}}{}
+ \]
+ \[
+ \infer[@{text "(\<And>\<hyphen>intro)"}]{@{text "\<Gamma> \<turnstile> \<And>x. B[x]"}}{@{text "\<Gamma> \<turnstile> B[x]"} & @{text "x \<notin> \<Gamma>"}}
+ \qquad
+ \infer[@{text "(\<And>\<hyphen>elim)"}]{@{text "\<Gamma> \<turnstile> B[a]"}}{@{text "\<Gamma> \<turnstile> \<And>x. B[x]"}}
+ \]
+ \[
+ \infer[@{text "(\<Longrightarrow>\<hyphen>intro)"}]{@{text "\<Gamma> - A \<turnstile> A \<Longrightarrow> B"}}{@{text "\<Gamma> \<turnstile> B"}}
+ \qquad
+ \infer[@{text "(\<Longrightarrow>\<hyphen>elim)"}]{@{text "\<Gamma>\<^sub>1 \<union> \<Gamma>\<^sub>2 \<turnstile> B"}}{@{text "\<Gamma>\<^sub>1 \<turnstile> A \<Longrightarrow> B"} & @{text "\<Gamma>\<^sub>2 \<turnstile> A"}}
+ \]
+ \caption{Primitive inferences of Pure}\label{fig:prim-rules}
+ \end{center}
+ \end{figure}
+
+ \begin{figure}[htb]
+ \begin{center}
+ \begin{tabular}{ll}
+ @{text "\<turnstile> (\<lambda>x. b[x]) a \<equiv> b[a]"} & @{text "\<beta>"}-conversion \\
+ @{text "\<turnstile> x \<equiv> x"} & reflexivity \\
+ @{text "\<turnstile> x \<equiv> y \<Longrightarrow> P x \<Longrightarrow> P y"} & substitution \\
+ @{text "\<turnstile> (\<And>x. f x \<equiv> g x) \<Longrightarrow> f \<equiv> g"} & extensionality \\
+ @{text "\<turnstile> (A \<Longrightarrow> B) \<Longrightarrow> (B \<Longrightarrow> A) \<Longrightarrow> A \<equiv> B"} & logical equivalence \\
+ \end{tabular}
+ \caption{Conceptual axiomatization of Pure equality}\label{fig:pure-equality}
+ \end{center}
+ \end{figure}
+
+ The introduction and elimination rules for @{text "\<And>"} and @{text
+ "\<Longrightarrow>"} are analogous to formation of dependently typed @{text
+ "\<lambda>"}-terms representing the underlying proof objects. Proof terms
+ are irrelevant in the Pure logic, though; they cannot occur within
+ propositions. The system provides a runtime option to record
+ explicit proof terms for primitive inferences, see also
+ \secref{sec:proof-terms}. Thus all three levels of @{text
+ "\<lambda>"}-calculus become explicit: @{text "\<Rightarrow>"} for terms, and @{text
+ "\<And>/\<Longrightarrow>"} for proofs (cf.\ \cite{Berghofer-Nipkow:2000:TPHOL}).
+
+ Observe that locally fixed parameters (as in @{text
+ "\<And>\<hyphen>intro"}) need not be recorded in the hypotheses, because
+ the simple syntactic types of Pure are always inhabitable.
+ ``Assumptions'' @{text "x :: \<tau>"} for type-membership are only
+ present as long as some @{text "x\<^sub>\<tau>"} occurs in the statement
+ body.\footnote{This is the key difference to ``@{text "\<lambda>HOL"}'' in
+ the PTS framework \cite{Barendregt-Geuvers:2001}, where hypotheses
+ @{text "x : A"} are treated uniformly for propositions and types.}
+
+ \medskip The axiomatization of a theory is implicitly closed by
+ forming all instances of type and term variables: @{text "\<turnstile>
+ A\<vartheta>"} holds for any substitution instance of an axiom
+ @{text "\<turnstile> A"}. By pushing substitutions through derivations
+ inductively, we also get admissible @{text "generalize"} and @{text
+ "instantiate"} rules as shown in \figref{fig:subst-rules}.
+
+ \begin{figure}[htb]
+ \begin{center}
+ \[
+ \infer{@{text "\<Gamma> \<turnstile> B[?\<alpha>]"}}{@{text "\<Gamma> \<turnstile> B[\<alpha>]"} & @{text "\<alpha> \<notin> \<Gamma>"}}
+ \quad
+ \infer[\quad@{text "(generalize)"}]{@{text "\<Gamma> \<turnstile> B[?x]"}}{@{text "\<Gamma> \<turnstile> B[x]"} & @{text "x \<notin> \<Gamma>"}}
+ \]
+ \[
+ \infer{@{text "\<Gamma> \<turnstile> B[\<tau>]"}}{@{text "\<Gamma> \<turnstile> B[?\<alpha>]"}}
+ \quad
+ \infer[\quad@{text "(instantiate)"}]{@{text "\<Gamma> \<turnstile> B[t]"}}{@{text "\<Gamma> \<turnstile> B[?x]"}}
+ \]
+ \caption{Admissible substitution rules}\label{fig:subst-rules}
+ \end{center}
+ \end{figure}
+
+ Note that @{text "instantiate"} does not require an explicit
+ side-condition, because @{text "\<Gamma>"} may never contain schematic
+ variables.
+
+ In principle, variables could be substituted in hypotheses as well,
+ but this would disrupt the monotonicity of reasoning: deriving
+ @{text "\<Gamma>\<vartheta> \<turnstile> B\<vartheta>"} from @{text "\<Gamma> \<turnstile> B"} is
+ correct, but @{text "\<Gamma>\<vartheta> \<supseteq> \<Gamma>"} does not necessarily hold:
+ the result belongs to a different proof context.
+
+ \medskip An \emph{oracle} is a function that produces axioms on the
+ fly. Logically, this is an instance of the @{text "axiom"} rule
+ (\figref{fig:prim-rules}), but there is an operational difference.
+ The system always records oracle invocations within derivations of
+ theorems by a unique tag.
+
+ Axiomatizations should be limited to the bare minimum, typically as
+ part of the initial logical basis of an object-logic formalization.
+ Later on, theories are usually developed in a strictly definitional
+ fashion, by stating only certain equalities over new constants.
+
+ A \emph{simple definition} consists of a constant declaration @{text
+ "c :: \<sigma>"} together with an axiom @{text "\<turnstile> c \<equiv> t"}, where @{text "t
+ :: \<sigma>"} is a closed term without any hidden polymorphism. The RHS
+ may depend on further defined constants, but not @{text "c"} itself.
+ Definitions of functions may be presented as @{text "c \<^vec>x \<equiv>
+ t"} instead of the puristic @{text "c \<equiv> \<lambda>\<^vec>x. t"}.
+
+ An \emph{overloaded definition} consists of a collection of axioms
+ for the same constant, with zero or one equations @{text
+ "c((\<^vec>\<alpha>)\<kappa>) \<equiv> t"} for each type constructor @{text "\<kappa>"} (for
+ distinct variables @{text "\<^vec>\<alpha>"}). The RHS may mention
+ previously defined constants as above, or arbitrary constants @{text
+ "d(\<alpha>\<^sub>i)"} for some @{text "\<alpha>\<^sub>i"} projected from @{text
+ "\<^vec>\<alpha>"}. Thus overloaded definitions essentially work by
+ primitive recursion over the syntactic structure of a single type
+ argument. See also \cite[\S4.3]{Haftmann-Wenzel:2006:classes}.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML Logic.all: "term -> term -> term"} \\
+ @{index_ML Logic.mk_implies: "term * term -> term"} \\
+ \end{mldecls}
+ \begin{mldecls}
+ @{index_ML_type ctyp} \\
+ @{index_ML_type cterm} \\
+ @{index_ML Thm.ctyp_of: "theory -> typ -> ctyp"} \\
+ @{index_ML Thm.cterm_of: "theory -> term -> cterm"} \\
+ @{index_ML Thm.apply: "cterm -> cterm -> cterm"} \\
+ @{index_ML Thm.lambda: "cterm -> cterm -> cterm"} \\
+ @{index_ML Thm.all: "cterm -> cterm -> cterm"} \\
+ @{index_ML Drule.mk_implies: "cterm * cterm -> cterm"} \\
+ \end{mldecls}
+ \begin{mldecls}
+ @{index_ML_type thm} \\
+ @{index_ML Thm.peek_status: "thm -> {oracle: bool, unfinished: bool, failed: bool}"} \\
+ @{index_ML Thm.transfer: "theory -> thm -> thm"} \\
+ @{index_ML Thm.assume: "cterm -> thm"} \\
+ @{index_ML Thm.forall_intr: "cterm -> thm -> thm"} \\
+ @{index_ML Thm.forall_elim: "cterm -> thm -> thm"} \\
+ @{index_ML Thm.implies_intr: "cterm -> thm -> thm"} \\
+ @{index_ML Thm.implies_elim: "thm -> thm -> thm"} \\
+ @{index_ML Thm.generalize: "string list * string list -> int -> thm -> thm"} \\
+ @{index_ML Thm.instantiate: "(ctyp * ctyp) list * (cterm * cterm) list -> thm -> thm"} \\
+ @{index_ML Thm.add_axiom: "Proof.context ->
+ binding * term -> theory -> (string * thm) * theory"} \\
+ @{index_ML Thm.add_oracle: "binding * ('a -> cterm) -> theory ->
+ (string * ('a -> thm)) * theory"} \\
+ @{index_ML Thm.add_def: "Proof.context -> bool -> bool ->
+ binding * term -> theory -> (string * thm) * theory"} \\
+ \end{mldecls}
+ \begin{mldecls}
+ @{index_ML Theory.add_deps: "Proof.context -> string ->
+ string * typ -> (string * typ) list -> theory -> theory"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML Thm.peek_status}~@{text "thm"} informs about the current
+ status of the derivation object behind the given theorem. This is a
+ snapshot of a potentially ongoing (parallel) evaluation of proofs.
+ The three Boolean values indicate the following: @{verbatim oracle}
+ if the finished part contains some oracle invocation; @{verbatim
+ unfinished} if some future proofs are still pending; @{verbatim
+ failed} if some future proof has failed, rendering the theorem
+ invalid!
+
+ \item @{ML Logic.all}~@{text "a B"} produces a Pure quantification
+ @{text "\<And>a. B"}, where occurrences of the atomic term @{text "a"} in
+ the body proposition @{text "B"} are replaced by bound variables.
+ (See also @{ML lambda} on terms.)
+
+ \item @{ML Logic.mk_implies}~@{text "(A, B)"} produces a Pure
+ implication @{text "A \<Longrightarrow> B"}.
+
+ \item Types @{ML_type ctyp} and @{ML_type cterm} represent certified
+ types and terms, respectively. These are abstract datatypes that
+ guarantee that its values have passed the full well-formedness (and
+ well-typedness) checks, relative to the declarations of type
+ constructors, constants etc.\ in the background theory. The
+ abstract types @{ML_type ctyp} and @{ML_type cterm} are part of the
+ same inference kernel that is mainly responsible for @{ML_type thm}.
+ Thus syntactic operations on @{ML_type ctyp} and @{ML_type cterm}
+ are located in the @{ML_structure Thm} module, even though theorems are
+ not yet involved at that stage.
+
+ \item @{ML Thm.ctyp_of}~@{text "thy \<tau>"} and @{ML
+ Thm.cterm_of}~@{text "thy t"} explicitly checks types and terms,
+ respectively. This also involves some basic normalizations, such
+ expansion of type and term abbreviations from the theory context.
+ Full re-certification is relatively slow and should be avoided in
+ tight reasoning loops.
+
+ \item @{ML Thm.apply}, @{ML Thm.lambda}, @{ML Thm.all}, @{ML
+ Drule.mk_implies} etc.\ compose certified terms (or propositions)
+ incrementally. This is equivalent to @{ML Thm.cterm_of} after
+ unchecked @{ML_op "$"}, @{ML lambda}, @{ML Logic.all}, @{ML
+ Logic.mk_implies} etc., but there can be a big difference in
+ performance when large existing entities are composed by a few extra
+ constructions on top. There are separate operations to decompose
+ certified terms and theorems to produce certified terms again.
+
+ \item Type @{ML_type thm} represents proven propositions. This is
+ an abstract datatype that guarantees that its values have been
+ constructed by basic principles of the @{ML_structure Thm} module.
+ Every @{ML_type thm} value refers its background theory,
+ cf.\ \secref{sec:context-theory}.
+
+ \item @{ML Thm.transfer}~@{text "thy thm"} transfers the given
+ theorem to a \emph{larger} theory, see also \secref{sec:context}.
+ This formal adjustment of the background context has no logical
+ significance, but is occasionally required for formal reasons, e.g.\
+ when theorems that are imported from more basic theories are used in
+ the current situation.
+
+ \item @{ML Thm.assume}, @{ML Thm.forall_intr}, @{ML
+ Thm.forall_elim}, @{ML Thm.implies_intr}, and @{ML Thm.implies_elim}
+ correspond to the primitive inferences of \figref{fig:prim-rules}.
+
+ \item @{ML Thm.generalize}~@{text "(\<^vec>\<alpha>, \<^vec>x)"}
+ corresponds to the @{text "generalize"} rules of
+ \figref{fig:subst-rules}. Here collections of type and term
+ variables are generalized simultaneously, specified by the given
+ basic names.
+
+ \item @{ML Thm.instantiate}~@{text "(\<^vec>\<alpha>\<^sub>s,
+ \<^vec>x\<^sub>\<tau>)"} corresponds to the @{text "instantiate"} rules
+ of \figref{fig:subst-rules}. Type variables are substituted before
+ term variables. Note that the types in @{text "\<^vec>x\<^sub>\<tau>"}
+ refer to the instantiated versions.
+
+ \item @{ML Thm.add_axiom}~@{text "ctxt (name, A)"} declares an
+ arbitrary proposition as axiom, and retrieves it as a theorem from
+ the resulting theory, cf.\ @{text "axiom"} in
+ \figref{fig:prim-rules}. Note that the low-level representation in
+ the axiom table may differ slightly from the returned theorem.
+
+ \item @{ML Thm.add_oracle}~@{text "(binding, oracle)"} produces a named
+ oracle rule, essentially generating arbitrary axioms on the fly,
+ cf.\ @{text "axiom"} in \figref{fig:prim-rules}.
+
+ \item @{ML Thm.add_def}~@{text "ctxt unchecked overloaded (name, c
+ \<^vec>x \<equiv> t)"} states a definitional axiom for an existing constant
+ @{text "c"}. Dependencies are recorded via @{ML Theory.add_deps},
+ unless the @{text "unchecked"} option is set. Note that the
+ low-level representation in the axiom table may differ slightly from
+ the returned theorem.
+
+ \item @{ML Theory.add_deps}~@{text "ctxt name c\<^sub>\<tau> \<^vec>d\<^sub>\<sigma>"}
+ declares dependencies of a named specification for constant @{text
+ "c\<^sub>\<tau>"}, relative to existing specifications for constants @{text
+ "\<^vec>d\<^sub>\<sigma>"}.
+
+ \end{description}
+*}
+
+
+text %mlantiq {*
+ \begin{matharray}{rcl}
+ @{ML_antiquotation_def "ctyp"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "cterm"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "cprop"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "thm"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "thms"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "lemma"} & : & @{text ML_antiquotation} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{ML_antiquotation ctyp} typ
+ ;
+ @@{ML_antiquotation cterm} term
+ ;
+ @@{ML_antiquotation cprop} prop
+ ;
+ @@{ML_antiquotation thm} thmref
+ ;
+ @@{ML_antiquotation thms} thmrefs
+ ;
+ @@{ML_antiquotation lemma} ('(' @'open' ')')? ((prop +) + @'and') \<newline>
+ @'by' method method?
+ \<close>}
+
+ \begin{description}
+
+ \item @{text "@{ctyp \<tau>}"} produces a certified type wrt.\ the
+ current background theory --- as abstract value of type @{ML_type
+ ctyp}.
+
+ \item @{text "@{cterm t}"} and @{text "@{cprop \<phi>}"} produce a
+ certified term wrt.\ the current background theory --- as abstract
+ value of type @{ML_type cterm}.
+
+ \item @{text "@{thm a}"} produces a singleton fact --- as abstract
+ value of type @{ML_type thm}.
+
+ \item @{text "@{thms a}"} produces a general fact --- as abstract
+ value of type @{ML_type "thm list"}.
+
+ \item @{text "@{lemma \<phi> by meth}"} produces a fact that is proven on
+ the spot according to the minimal proof, which imitates a terminal
+ Isar proof. The result is an abstract value of type @{ML_type thm}
+ or @{ML_type "thm list"}, depending on the number of propositions
+ given here.
+
+ The internal derivation object lacks a proper theorem name, but it
+ is formally closed, unless the @{text "(open)"} option is specified
+ (this may impact performance of applications with proof terms).
+
+ Since ML antiquotations are always evaluated at compile-time, there
+ is no run-time overhead even for non-trivial proofs. Nonetheless,
+ the justification is syntactically limited to a single @{command
+ "by"} step. More complex Isar proofs should be done in regular
+ theory source, before compiling the corresponding ML text that uses
+ the result.
+
+ \end{description}
+
+*}
+
+
+subsection {* Auxiliary connectives \label{sec:logic-aux} *}
+
+text {* Theory @{text "Pure"} provides a few auxiliary connectives
+ that are defined on top of the primitive ones, see
+ \figref{fig:pure-aux}. These special constants are useful in
+ certain internal encodings, and are normally not directly exposed to
+ the user.
+
+ \begin{figure}[htb]
+ \begin{center}
+ \begin{tabular}{ll}
+ @{text "conjunction :: prop \<Rightarrow> prop \<Rightarrow> prop"} & (infix @{text "&&&"}) \\
+ @{text "\<turnstile> A &&& B \<equiv> (\<And>C. (A \<Longrightarrow> B \<Longrightarrow> C) \<Longrightarrow> C)"} \\[1ex]
+ @{text "prop :: prop \<Rightarrow> prop"} & (prefix @{text "#"}, suppressed) \\
+ @{text "#A \<equiv> A"} \\[1ex]
+ @{text "term :: \<alpha> \<Rightarrow> prop"} & (prefix @{text "TERM"}) \\
+ @{text "term x \<equiv> (\<And>A. A \<Longrightarrow> A)"} \\[1ex]
+ @{text "type :: \<alpha> itself"} & (prefix @{text "TYPE"}) \\
+ @{text "(unspecified)"} \\
+ \end{tabular}
+ \caption{Definitions of auxiliary connectives}\label{fig:pure-aux}
+ \end{center}
+ \end{figure}
+
+ The introduction @{text "A \<Longrightarrow> B \<Longrightarrow> A &&& B"}, and eliminations
+ (projections) @{text "A &&& B \<Longrightarrow> A"} and @{text "A &&& B \<Longrightarrow> B"} are
+ available as derived rules. Conjunction allows to treat
+ simultaneous assumptions and conclusions uniformly, e.g.\ consider
+ @{text "A \<Longrightarrow> B \<Longrightarrow> C &&& D"}. In particular, the goal mechanism
+ represents multiple claims as explicit conjunction internally, but
+ this is refined (via backwards introduction) into separate sub-goals
+ before the user commences the proof; the final result is projected
+ into a list of theorems using eliminations (cf.\
+ \secref{sec:tactical-goals}).
+
+ The @{text "prop"} marker (@{text "#"}) makes arbitrarily complex
+ propositions appear as atomic, without changing the meaning: @{text
+ "\<Gamma> \<turnstile> A"} and @{text "\<Gamma> \<turnstile> #A"} are interchangeable. See
+ \secref{sec:tactical-goals} for specific operations.
+
+ The @{text "term"} marker turns any well-typed term into a derivable
+ proposition: @{text "\<turnstile> TERM t"} holds unconditionally. Although
+ this is logically vacuous, it allows to treat terms and proofs
+ uniformly, similar to a type-theoretic framework.
+
+ The @{text "TYPE"} constructor is the canonical representative of
+ the unspecified type @{text "\<alpha> itself"}; it essentially injects the
+ language of types into that of terms. There is specific notation
+ @{text "TYPE(\<tau>)"} for @{text "TYPE\<^bsub>\<tau> itself\<^esub>"}.
+ Although being devoid of any particular meaning, the term @{text
+ "TYPE(\<tau>)"} accounts for the type @{text "\<tau>"} within the term
+ language. In particular, @{text "TYPE(\<alpha>)"} may be used as formal
+ argument in primitive definitions, in order to circumvent hidden
+ polymorphism (cf.\ \secref{sec:terms}). For example, @{text "c
+ TYPE(\<alpha>) \<equiv> A[\<alpha>]"} defines @{text "c :: \<alpha> itself \<Rightarrow> prop"} in terms of
+ a proposition @{text "A"} that depends on an additional type
+ argument, which is essentially a predicate on types.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML Conjunction.intr: "thm -> thm -> thm"} \\
+ @{index_ML Conjunction.elim: "thm -> thm * thm"} \\
+ @{index_ML Drule.mk_term: "cterm -> thm"} \\
+ @{index_ML Drule.dest_term: "thm -> cterm"} \\
+ @{index_ML Logic.mk_type: "typ -> term"} \\
+ @{index_ML Logic.dest_type: "term -> typ"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML Conjunction.intr} derives @{text "A &&& B"} from @{text
+ "A"} and @{text "B"}.
+
+ \item @{ML Conjunction.elim} derives @{text "A"} and @{text "B"}
+ from @{text "A &&& B"}.
+
+ \item @{ML Drule.mk_term} derives @{text "TERM t"}.
+
+ \item @{ML Drule.dest_term} recovers term @{text "t"} from @{text
+ "TERM t"}.
+
+ \item @{ML Logic.mk_type}~@{text "\<tau>"} produces the term @{text
+ "TYPE(\<tau>)"}.
+
+ \item @{ML Logic.dest_type}~@{text "TYPE(\<tau>)"} recovers the type
+ @{text "\<tau>"}.
+
+ \end{description}
+*}
+
+
+subsection {* Sort hypotheses *}
+
+text {* Type variables are decorated with sorts, as explained in
+ \secref{sec:types}. This constrains type instantiation to certain
+ ranges of types: variable @{text "\<alpha>\<^sub>s"} may only be assigned to types
+ @{text "\<tau>"} that belong to sort @{text "s"}. Within the logic, sort
+ constraints act like implicit preconditions on the result @{text
+ "\<lparr>\<alpha>\<^sub>1 : s\<^sub>1\<rparr>, \<dots>, \<lparr>\<alpha>\<^sub>n : s\<^sub>n\<rparr>, \<Gamma> \<turnstile> \<phi>"} where the type variables @{text
+ "\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n"} cover the propositions @{text "\<Gamma>"}, @{text "\<phi>"}, as
+ well as the proof of @{text "\<Gamma> \<turnstile> \<phi>"}.
+
+ These \emph{sort hypothesis} of a theorem are passed monotonically
+ through further derivations. They are redundant, as long as the
+ statement of a theorem still contains the type variables that are
+ accounted here. The logical significance of sort hypotheses is
+ limited to the boundary case where type variables disappear from the
+ proposition, e.g.\ @{text "\<lparr>\<alpha>\<^sub>s : s\<rparr> \<turnstile> \<phi>"}. Since such dangling type
+ variables can be renamed arbitrarily without changing the
+ proposition @{text "\<phi>"}, the inference kernel maintains sort
+ hypotheses in anonymous form @{text "s \<turnstile> \<phi>"}.
+
+ In most practical situations, such extra sort hypotheses may be
+ stripped in a final bookkeeping step, e.g.\ at the end of a proof:
+ they are typically left over from intermediate reasoning with type
+ classes that can be satisfied by some concrete type @{text "\<tau>"} of
+ sort @{text "s"} to replace the hypothetical type variable @{text
+ "\<alpha>\<^sub>s"}. *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML Thm.extra_shyps: "thm -> sort list"} \\
+ @{index_ML Thm.strip_shyps: "thm -> thm"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML Thm.extra_shyps}~@{text "thm"} determines the extraneous
+ sort hypotheses of the given theorem, i.e.\ the sorts that are not
+ present within type variables of the statement.
+
+ \item @{ML Thm.strip_shyps}~@{text "thm"} removes any extraneous
+ sort hypotheses that can be witnessed from the type signature.
+
+ \end{description}
+*}
+
+text %mlex {* The following artificial example demonstrates the
+ derivation of @{prop False} with a pending sort hypothesis involving
+ a logically empty sort. *}
+
+class empty =
+ assumes bad: "\<And>(x::'a) y. x \<noteq> y"
+
+theorem (in empty) false: False
+ using bad by blast
+
+ML {*
+ @{assert} (Thm.extra_shyps @{thm false} = [@{sort empty}])
+*}
+
+text {* Thanks to the inference kernel managing sort hypothesis
+ according to their logical significance, this example is merely an
+ instance of \emph{ex falso quodlibet consequitur} --- not a collapse
+ of the logical framework! *}
+
+
+section {* Object-level rules \label{sec:obj-rules} *}
+
+text {*
+ The primitive inferences covered so far mostly serve foundational
+ purposes. User-level reasoning usually works via object-level rules
+ that are represented as theorems of Pure. Composition of rules
+ involves \emph{backchaining}, \emph{higher-order unification} modulo
+ @{text "\<alpha>\<beta>\<eta>"}-conversion of @{text "\<lambda>"}-terms, and so-called
+ \emph{lifting} of rules into a context of @{text "\<And>"} and @{text
+ "\<Longrightarrow>"} connectives. Thus the full power of higher-order Natural
+ Deduction in Isabelle/Pure becomes readily available.
+*}
+
+
+subsection {* Hereditary Harrop Formulae *}
+
+text {*
+ The idea of object-level rules is to model Natural Deduction
+ inferences in the style of Gentzen \cite{Gentzen:1935}, but we allow
+ arbitrary nesting similar to \cite{extensions91}. The most basic
+ rule format is that of a \emph{Horn Clause}:
+ \[
+ \infer{@{text "A"}}{@{text "A\<^sub>1"} & @{text "\<dots>"} & @{text "A\<^sub>n"}}
+ \]
+ where @{text "A, A\<^sub>1, \<dots>, A\<^sub>n"} are atomic propositions
+ of the framework, usually of the form @{text "Trueprop B"}, where
+ @{text "B"} is a (compound) object-level statement. This
+ object-level inference corresponds to an iterated implication in
+ Pure like this:
+ \[
+ @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> A"}
+ \]
+ As an example consider conjunction introduction: @{text "A \<Longrightarrow> B \<Longrightarrow> A \<and>
+ B"}. Any parameters occurring in such rule statements are
+ conceptionally treated as arbitrary:
+ \[
+ @{text "\<And>x\<^sub>1 \<dots> x\<^sub>m. A\<^sub>1 x\<^sub>1 \<dots> x\<^sub>m \<Longrightarrow> \<dots> A\<^sub>n x\<^sub>1 \<dots> x\<^sub>m \<Longrightarrow> A x\<^sub>1 \<dots> x\<^sub>m"}
+ \]
+
+ Nesting of rules means that the positions of @{text "A\<^sub>i"} may
+ again hold compound rules, not just atomic propositions.
+ Propositions of this format are called \emph{Hereditary Harrop
+ Formulae} in the literature \cite{Miller:1991}. Here we give an
+ inductive characterization as follows:
+
+ \medskip
+ \begin{tabular}{ll}
+ @{text "\<^bold>x"} & set of variables \\
+ @{text "\<^bold>A"} & set of atomic propositions \\
+ @{text "\<^bold>H = \<And>\<^bold>x\<^sup>*. \<^bold>H\<^sup>* \<Longrightarrow> \<^bold>A"} & set of Hereditary Harrop Formulas \\
+ \end{tabular}
+ \medskip
+
+ Thus we essentially impose nesting levels on propositions formed
+ from @{text "\<And>"} and @{text "\<Longrightarrow>"}. At each level there is a prefix
+ of parameters and compound premises, concluding an atomic
+ proposition. Typical examples are @{text "\<longrightarrow>"}-introduction @{text
+ "(A \<Longrightarrow> B) \<Longrightarrow> A \<longrightarrow> B"} or mathematical induction @{text "P 0 \<Longrightarrow> (\<And>n. P n
+ \<Longrightarrow> P (Suc n)) \<Longrightarrow> P n"}. Even deeper nesting occurs in well-founded
+ induction @{text "(\<And>x. (\<And>y. y \<prec> x \<Longrightarrow> P y) \<Longrightarrow> P x) \<Longrightarrow> P x"}, but this
+ already marks the limit of rule complexity that is usually seen in
+ practice.
+
+ \medskip Regular user-level inferences in Isabelle/Pure always
+ maintain the following canonical form of results:
+
+ \begin{itemize}
+
+ \item Normalization by @{text "(A \<Longrightarrow> (\<And>x. B x)) \<equiv> (\<And>x. A \<Longrightarrow> B x)"},
+ which is a theorem of Pure, means that quantifiers are pushed in
+ front of implication at each level of nesting. The normal form is a
+ Hereditary Harrop Formula.
+
+ \item The outermost prefix of parameters is represented via
+ schematic variables: instead of @{text "\<And>\<^vec>x. \<^vec>H \<^vec>x
+ \<Longrightarrow> A \<^vec>x"} we have @{text "\<^vec>H ?\<^vec>x \<Longrightarrow> A ?\<^vec>x"}.
+ Note that this representation looses information about the order of
+ parameters, and vacuous quantifiers vanish automatically.
+
+ \end{itemize}
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML Simplifier.norm_hhf: "Proof.context -> thm -> thm"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML Simplifier.norm_hhf}~@{text "ctxt thm"} normalizes the given
+ theorem according to the canonical form specified above. This is
+ occasionally helpful to repair some low-level tools that do not
+ handle Hereditary Harrop Formulae properly.
+
+ \end{description}
+*}
+
+
+subsection {* Rule composition *}
+
+text {*
+ The rule calculus of Isabelle/Pure provides two main inferences:
+ @{inference resolution} (i.e.\ back-chaining of rules) and
+ @{inference assumption} (i.e.\ closing a branch), both modulo
+ higher-order unification. There are also combined variants, notably
+ @{inference elim_resolution} and @{inference dest_resolution}.
+
+ To understand the all-important @{inference resolution} principle,
+ we first consider raw @{inference_def composition} (modulo
+ higher-order unification with substitution @{text "\<vartheta>"}):
+ \[
+ \infer[(@{inference_def composition})]{@{text "\<^vec>A\<vartheta> \<Longrightarrow> C\<vartheta>"}}
+ {@{text "\<^vec>A \<Longrightarrow> B"} & @{text "B' \<Longrightarrow> C"} & @{text "B\<vartheta> = B'\<vartheta>"}}
+ \]
+ Here the conclusion of the first rule is unified with the premise of
+ the second; the resulting rule instance inherits the premises of the
+ first and conclusion of the second. Note that @{text "C"} can again
+ consist of iterated implications. We can also permute the premises
+ of the second rule back-and-forth in order to compose with @{text
+ "B'"} in any position (subsequently we shall always refer to
+ position 1 w.l.o.g.).
+
+ In @{inference composition} the internal structure of the common
+ part @{text "B"} and @{text "B'"} is not taken into account. For
+ proper @{inference resolution} we require @{text "B"} to be atomic,
+ and explicitly observe the structure @{text "\<And>\<^vec>x. \<^vec>H
+ \<^vec>x \<Longrightarrow> B' \<^vec>x"} of the premise of the second rule. The
+ idea is to adapt the first rule by ``lifting'' it into this context,
+ by means of iterated application of the following inferences:
+ \[
+ \infer[(@{inference_def imp_lift})]{@{text "(\<^vec>H \<Longrightarrow> \<^vec>A) \<Longrightarrow> (\<^vec>H \<Longrightarrow> B)"}}{@{text "\<^vec>A \<Longrightarrow> B"}}
+ \]
+ \[
+ \infer[(@{inference_def all_lift})]{@{text "(\<And>\<^vec>x. \<^vec>A (?\<^vec>a \<^vec>x)) \<Longrightarrow> (\<And>\<^vec>x. B (?\<^vec>a \<^vec>x))"}}{@{text "\<^vec>A ?\<^vec>a \<Longrightarrow> B ?\<^vec>a"}}
+ \]
+ By combining raw composition with lifting, we get full @{inference
+ resolution} as follows:
+ \[
+ \infer[(@{inference_def resolution})]
+ {@{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> \<^vec>A (?\<^vec>a \<^vec>x))\<vartheta> \<Longrightarrow> C\<vartheta>"}}
+ {\begin{tabular}{l}
+ @{text "\<^vec>A ?\<^vec>a \<Longrightarrow> B ?\<^vec>a"} \\
+ @{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> B' \<^vec>x) \<Longrightarrow> C"} \\
+ @{text "(\<lambda>\<^vec>x. B (?\<^vec>a \<^vec>x))\<vartheta> = B'\<vartheta>"} \\
+ \end{tabular}}
+ \]
+
+ Continued resolution of rules allows to back-chain a problem towards
+ more and sub-problems. Branches are closed either by resolving with
+ a rule of 0 premises, or by producing a ``short-circuit'' within a
+ solved situation (again modulo unification):
+ \[
+ \infer[(@{inference_def assumption})]{@{text "C\<vartheta>"}}
+ {@{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> A \<^vec>x) \<Longrightarrow> C"} & @{text "A\<vartheta> = H\<^sub>i\<vartheta>"}~~\text{(for some~@{text i})}}
+ \]
+
+ %FIXME @{inference_def elim_resolution}, @{inference_def dest_resolution}
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_op "RSN": "thm * (int * thm) -> thm"} \\
+ @{index_ML_op "RS": "thm * thm -> thm"} \\
+
+ @{index_ML_op "RLN": "thm list * (int * thm list) -> thm list"} \\
+ @{index_ML_op "RL": "thm list * thm list -> thm list"} \\
+
+ @{index_ML_op "MRS": "thm list * thm -> thm"} \\
+ @{index_ML_op "OF": "thm * thm list -> thm"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{text "rule\<^sub>1 RSN (i, rule\<^sub>2)"} resolves the conclusion of
+ @{text "rule\<^sub>1"} with the @{text i}-th premise of @{text "rule\<^sub>2"},
+ according to the @{inference resolution} principle explained above.
+ Unless there is precisely one resolvent it raises exception @{ML
+ THM}.
+
+ This corresponds to the rule attribute @{attribute THEN} in Isar
+ source language.
+
+ \item @{text "rule\<^sub>1 RS rule\<^sub>2"} abbreviates @{text "rule\<^sub>1 RSN (1,
+ rule\<^sub>2)"}.
+
+ \item @{text "rules\<^sub>1 RLN (i, rules\<^sub>2)"} joins lists of rules. For
+ every @{text "rule\<^sub>1"} in @{text "rules\<^sub>1"} and @{text "rule\<^sub>2"} in
+ @{text "rules\<^sub>2"}, it resolves the conclusion of @{text "rule\<^sub>1"} with
+ the @{text "i"}-th premise of @{text "rule\<^sub>2"}, accumulating multiple
+ results in one big list. Note that such strict enumerations of
+ higher-order unifications can be inefficient compared to the lazy
+ variant seen in elementary tactics like @{ML resolve_tac}.
+
+ \item @{text "rules\<^sub>1 RL rules\<^sub>2"} abbreviates @{text "rules\<^sub>1 RLN (1,
+ rules\<^sub>2)"}.
+
+ \item @{text "[rule\<^sub>1, \<dots>, rule\<^sub>n] MRS rule"} resolves @{text "rule\<^sub>i"}
+ against premise @{text "i"} of @{text "rule"}, for @{text "i = n, \<dots>,
+ 1"}. By working from right to left, newly emerging premises are
+ concatenated in the result, without interfering.
+
+ \item @{text "rule OF rules"} is an alternative notation for @{text
+ "rules MRS rule"}, which makes rule composition look more like
+ function application. Note that the argument @{text "rules"} need
+ not be atomic.
+
+ This corresponds to the rule attribute @{attribute OF} in Isar
+ source language.
+
+ \end{description}
+*}
+
+
+section {* Proof terms \label{sec:proof-terms} *}
+
+text {* The Isabelle/Pure inference kernel can record the proof of
+ each theorem as a proof term that contains all logical inferences in
+ detail. Rule composition by resolution (\secref{sec:obj-rules}) and
+ type-class reasoning is broken down to primitive rules of the
+ logical framework. The proof term can be inspected by a separate
+ proof-checker, for example.
+
+ According to the well-known \emph{Curry-Howard isomorphism}, a proof
+ can be viewed as a @{text "\<lambda>"}-term. Following this idea, proofs in
+ Isabelle are internally represented by a datatype similar to the one
+ for terms described in \secref{sec:terms}. On top of these
+ syntactic terms, two more layers of @{text "\<lambda>"}-calculus are added,
+ which correspond to @{text "\<And>x :: \<alpha>. B x"} and @{text "A \<Longrightarrow> B"}
+ according to the propositions-as-types principle. The resulting
+ 3-level @{text "\<lambda>"}-calculus resembles ``@{text "\<lambda>HOL"}'' in the
+ more abstract setting of Pure Type Systems (PTS)
+ \cite{Barendregt-Geuvers:2001}, if some fine points like schematic
+ polymorphism and type classes are ignored.
+
+ \medskip\emph{Proof abstractions} of the form @{text "\<^bold>\<lambda>x :: \<alpha>. prf"}
+ or @{text "\<^bold>\<lambda>p : A. prf"} correspond to introduction of @{text
+ "\<And>"}/@{text "\<Longrightarrow>"}, and \emph{proof applications} of the form @{text
+ "p \<cdot> t"} or @{text "p \<bullet> q"} correspond to elimination of @{text
+ "\<And>"}/@{text "\<Longrightarrow>"}. Actual types @{text "\<alpha>"}, propositions @{text
+ "A"}, and terms @{text "t"} might be suppressed and reconstructed
+ from the overall proof term.
+
+ \medskip Various atomic proofs indicate special situations within
+ the proof construction as follows.
+
+ A \emph{bound proof variable} is a natural number @{text "b"} that
+ acts as de-Bruijn index for proof term abstractions.
+
+ A \emph{minimal proof} ``@{text "?"}'' is a dummy proof term. This
+ indicates some unrecorded part of the proof.
+
+ @{text "Hyp A"} refers to some pending hypothesis by giving its
+ proposition. This indicates an open context of implicit hypotheses,
+ similar to loose bound variables or free variables within a term
+ (\secref{sec:terms}).
+
+ An \emph{axiom} or \emph{oracle} @{text "a : A[\<^vec>\<tau>]"} refers
+ some postulated @{text "proof constant"}, which is subject to
+ schematic polymorphism of theory content, and the particular type
+ instantiation may be given explicitly. The vector of types @{text
+ "\<^vec>\<tau>"} refers to the schematic type variables in the generic
+ proposition @{text "A"} in canonical order.
+
+ A \emph{proof promise} @{text "a : A[\<^vec>\<tau>]"} is a placeholder
+ for some proof of polymorphic proposition @{text "A"}, with explicit
+ type instantiation as given by the vector @{text "\<^vec>\<tau>"}, as
+ above. Unlike axioms or oracles, proof promises may be
+ \emph{fulfilled} eventually, by substituting @{text "a"} by some
+ particular proof @{text "q"} at the corresponding type instance.
+ This acts like Hindley-Milner @{text "let"}-polymorphism: a generic
+ local proof definition may get used at different type instances, and
+ is replaced by the concrete instance eventually.
+
+ A \emph{named theorem} wraps up some concrete proof as a closed
+ formal entity, in the manner of constant definitions for proof
+ terms. The \emph{proof body} of such boxed theorems involves some
+ digest about oracles and promises occurring in the original proof.
+ This allows the inference kernel to manage this critical information
+ without the full overhead of explicit proof terms.
+*}
+
+
+subsection {* Reconstructing and checking proof terms *}
+
+text {* Fully explicit proof terms can be large, but most of this
+ information is redundant and can be reconstructed from the context.
+ Therefore, the Isabelle/Pure inference kernel records only
+ \emph{implicit} proof terms, by omitting all typing information in
+ terms, all term and type labels of proof abstractions, and some
+ argument terms of applications @{text "p \<cdot> t"} (if possible).
+
+ There are separate operations to reconstruct the full proof term
+ later on, using \emph{higher-order pattern unification}
+ \cite{nipkow-patterns,Berghofer-Nipkow:2000:TPHOL}.
+
+ The \emph{proof checker} expects a fully reconstructed proof term,
+ and can turn it into a theorem by replaying its primitive inferences
+ within the kernel. *}
+
+
+subsection {* Concrete syntax of proof terms *}
+
+text {* The concrete syntax of proof terms is a slight extension of
+ the regular inner syntax of Isabelle/Pure \cite{isabelle-isar-ref}.
+ Its main syntactic category @{syntax (inner) proof} is defined as
+ follows:
+
+ \begin{center}
+ \begin{supertabular}{rclr}
+
+ @{syntax_def (inner) proof} & = & @{verbatim Lam} @{text params} @{verbatim "."} @{text proof} \\
+ & @{text "|"} & @{text "\<^bold>\<lambda>"} @{text "params"} @{verbatim "."} @{text proof} \\
+ & @{text "|"} & @{text proof} @{verbatim "%"} @{text any} \\
+ & @{text "|"} & @{text proof} @{text "\<cdot>"} @{text any} \\
+ & @{text "|"} & @{text proof} @{verbatim "%%"} @{text proof} \\
+ & @{text "|"} & @{text proof} @{text "\<bullet>"} @{text proof} \\
+ & @{text "|"} & @{text "id | longid"} \\
+ \\
+
+ @{text param} & = & @{text idt} \\
+ & @{text "|"} & @{text idt} @{verbatim ":"} @{text prop} \\
+ & @{text "|"} & @{verbatim "("} @{text param} @{verbatim ")"} \\
+ \\
+
+ @{text params} & = & @{text param} \\
+ & @{text "|"} & @{text param} @{text params} \\
+
+ \end{supertabular}
+ \end{center}
+
+ Implicit term arguments in partial proofs are indicated by ``@{text
+ "_"}''. Type arguments for theorems and axioms may be specified
+ using @{text "p \<cdot> TYPE(type)"} (they must appear before any other
+ term argument of a theorem or axiom, but may be omitted altogether).
+
+ \medskip There are separate read and print operations for proof
+ terms, in order to avoid conflicts with the regular term language.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type proof} \\
+ @{index_ML_type proof_body} \\
+ @{index_ML proofs: "int Unsynchronized.ref"} \\
+ @{index_ML Reconstruct.reconstruct_proof:
+ "theory -> term -> proof -> proof"} \\
+ @{index_ML Reconstruct.expand_proof: "theory ->
+ (string * term option) list -> proof -> proof"} \\
+ @{index_ML Proof_Checker.thm_of_proof: "theory -> proof -> thm"} \\
+ @{index_ML Proof_Syntax.read_proof: "theory -> bool -> bool -> string -> proof"} \\
+ @{index_ML Proof_Syntax.pretty_proof: "Proof.context -> proof -> Pretty.T"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type proof} represents proof terms; this is a
+ datatype with constructors @{index_ML Abst}, @{index_ML AbsP},
+ @{index_ML_op "%"}, @{index_ML_op "%%"}, @{index_ML PBound},
+ @{index_ML MinProof}, @{index_ML Hyp}, @{index_ML PAxm}, @{index_ML
+ Oracle}, @{index_ML Promise}, @{index_ML PThm} as explained above.
+ %FIXME OfClass (!?)
+
+ \item Type @{ML_type proof_body} represents the nested proof
+ information of a named theorem, consisting of a digest of oracles
+ and named theorem over some proof term. The digest only covers the
+ directly visible part of the proof: in order to get the full
+ information, the implicit graph of nested theorems needs to be
+ traversed (e.g.\ using @{ML Proofterm.fold_body_thms}).
+
+ \item @{ML Thm.proof_of}~@{text "thm"} and @{ML
+ Thm.proof_body_of}~@{text "thm"} produce the proof term or proof
+ body (with digest of oracles and theorems) from a given theorem.
+ Note that this involves a full join of internal futures that fulfill
+ pending proof promises, and thus disrupts the natural bottom-up
+ construction of proofs by introducing dynamic ad-hoc dependencies.
+ Parallel performance may suffer by inspecting proof terms at
+ run-time.
+
+ \item @{ML proofs} specifies the detail of proof recording within
+ @{ML_type thm} values produced by the inference kernel: @{ML 0}
+ records only the names of oracles, @{ML 1} records oracle names and
+ propositions, @{ML 2} additionally records full proof terms.
+ Officially named theorems that contribute to a result are recorded
+ in any case.
+
+ \item @{ML Reconstruct.reconstruct_proof}~@{text "thy prop prf"}
+ turns the implicit proof term @{text "prf"} into a full proof of the
+ given proposition.
+
+ Reconstruction may fail if @{text "prf"} is not a proof of @{text
+ "prop"}, or if it does not contain sufficient information for
+ reconstruction. Failure may only happen for proofs that are
+ constructed manually, but not for those produced automatically by
+ the inference kernel.
+
+ \item @{ML Reconstruct.expand_proof}~@{text "thy [thm\<^sub>1, \<dots>, thm\<^sub>n]
+ prf"} expands and reconstructs the proofs of all specified theorems,
+ with the given (full) proof. Theorems that are not unique specified
+ via their name may be disambiguated by giving their proposition.
+
+ \item @{ML Proof_Checker.thm_of_proof}~@{text "thy prf"} turns the
+ given (full) proof into a theorem, by replaying it using only
+ primitive rules of the inference kernel.
+
+ \item @{ML Proof_Syntax.read_proof}~@{text "thy b\<^sub>1 b\<^sub>2 s"} reads in a
+ proof term. The Boolean flags indicate the use of sort and type
+ information. Usually, typing information is left implicit and is
+ inferred during proof reconstruction. %FIXME eliminate flags!?
+
+ \item @{ML Proof_Syntax.pretty_proof}~@{text "ctxt prf"}
+ pretty-prints the given proof term.
+
+ \end{description}
+*}
+
+text %mlex {* Detailed proof information of a theorem may be retrieved
+ as follows: *}
+
+lemma ex: "A \<and> B \<longrightarrow> B \<and> A"
+proof
+ assume "A \<and> B"
+ then obtain B and A ..
+ then show "B \<and> A" ..
+qed
+
+ML_val {*
+ (*proof body with digest*)
+ val body = Proofterm.strip_thm (Thm.proof_body_of @{thm ex});
+
+ (*proof term only*)
+ val prf = Proofterm.proof_of body;
+ Pretty.writeln (Proof_Syntax.pretty_proof @{context} prf);
+
+ (*all theorems used in the graph of nested proofs*)
+ val all_thms =
+ Proofterm.fold_body_thms
+ (fn (name, _, _) => insert (op =) name) [body] [];
+*}
+
+text {* The result refers to various basic facts of Isabelle/HOL:
+ @{thm [source] HOL.impI}, @{thm [source] HOL.conjE}, @{thm [source]
+ HOL.conjI} etc. The combinator @{ML Proofterm.fold_body_thms}
+ recursively explores the graph of the proofs of all theorems being
+ used here.
+
+ \medskip Alternatively, we may produce a proof term manually, and
+ turn it into a theorem as follows: *}
+
+ML_val {*
+ val thy = @{theory};
+ val prf =
+ Proof_Syntax.read_proof thy true false
+ "impI \<cdot> _ \<cdot> _ \<bullet> \
+ \ (\<^bold>\<lambda>H: _. \
+ \ conjE \<cdot> _ \<cdot> _ \<cdot> _ \<bullet> H \<bullet> \
+ \ (\<^bold>\<lambda>(H: _) Ha: _. conjI \<cdot> _ \<cdot> _ \<bullet> Ha \<bullet> H))";
+ val thm =
+ prf
+ |> Reconstruct.reconstruct_proof thy @{prop "A \<and> B \<longrightarrow> B \<and> A"}
+ |> Proof_Checker.thm_of_proof thy
+ |> Drule.export_without_context;
+*}
+
+text {* \medskip See also @{file "~~/src/HOL/Proofs/ex/XML_Data.thy"}
+ for further examples, with export and import of proof terms via
+ XML/ML data representation.
+*}
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Implementation/ML.thy Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,2116 @@
+theory "ML"
+imports Base
+begin
+
+chapter {* Isabelle/ML *}
+
+text {* Isabelle/ML is best understood as a certain culture based on
+ Standard ML. Thus it is not a new programming language, but a
+ certain way to use SML at an advanced level within the Isabelle
+ environment. This covers a variety of aspects that are geared
+ towards an efficient and robust platform for applications of formal
+ logic with fully foundational proof construction --- according to
+ the well-known \emph{LCF principle}. There is specific
+ infrastructure with library modules to address the needs of this
+ difficult task. For example, the raw parallel programming model of
+ Poly/ML is presented as considerably more abstract concept of
+ \emph{future values}, which is then used to augment the inference
+ kernel, proof interpreter, and theory loader accordingly.
+
+ The main aspects of Isabelle/ML are introduced below. These
+ first-hand explanations should help to understand how proper
+ Isabelle/ML is to be read and written, and to get access to the
+ wealth of experience that is expressed in the source text and its
+ history of changes.\footnote{See
+ @{url "http://isabelle.in.tum.de/repos/isabelle"} for the full
+ Mercurial history. There are symbolic tags to refer to official
+ Isabelle releases, as opposed to arbitrary \emph{tip} versions that
+ merely reflect snapshots that are never really up-to-date.} *}
+
+
+section {* Style and orthography *}
+
+text {* The sources of Isabelle/Isar are optimized for
+ \emph{readability} and \emph{maintainability}. The main purpose is
+ to tell an informed reader what is really going on and how things
+ really work. This is a non-trivial aim, but it is supported by a
+ certain style of writing Isabelle/ML that has emerged from long
+ years of system development.\footnote{See also the interesting style
+ guide for OCaml
+ @{url "http://caml.inria.fr/resources/doc/guides/guidelines.en.html"}
+ which shares many of our means and ends.}
+
+ The main principle behind any coding style is \emph{consistency}.
+ For a single author of a small program this merely means ``choose
+ your style and stick to it''. A complex project like Isabelle, with
+ long years of development and different contributors, requires more
+ standardization. A coding style that is changed every few years or
+ with every new contributor is no style at all, because consistency
+ is quickly lost. Global consistency is hard to achieve, though.
+ Nonetheless, one should always strive at least for local consistency
+ of modules and sub-systems, without deviating from some general
+ principles how to write Isabelle/ML.
+
+ In a sense, good coding style is like an \emph{orthography} for the
+ sources: it helps to read quickly over the text and see through the
+ main points, without getting distracted by accidental presentation
+ of free-style code.
+*}
+
+
+subsection {* Header and sectioning *}
+
+text {* Isabelle source files have a certain standardized header
+ format (with precise spacing) that follows ancient traditions
+ reaching back to the earliest versions of the system by Larry
+ Paulson. See @{file "~~/src/Pure/thm.ML"}, for example.
+
+ The header includes at least @{verbatim Title} and @{verbatim
+ Author} entries, followed by a prose description of the purpose of
+ the module. The latter can range from a single line to several
+ paragraphs of explanations.
+
+ The rest of the file is divided into sections, subsections,
+ subsubsections, paragraphs etc.\ using a simple layout via ML
+ comments as follows.
+
+\begin{verbatim}
+(*** section ***)
+
+(** subsection **)
+
+(* subsubsection *)
+
+(*short paragraph*)
+
+(*
+ long paragraph,
+ with more text
+*)
+\end{verbatim}
+
+ As in regular typography, there is some extra space \emph{before}
+ section headings that are adjacent to plain text (not other headings
+ as in the example above).
+
+ \medskip The precise wording of the prose text given in these
+ headings is chosen carefully to introduce the main theme of the
+ subsequent formal ML text.
+*}
+
+
+subsection {* Naming conventions *}
+
+text {* Since ML is the primary medium to express the meaning of the
+ source text, naming of ML entities requires special care.
+
+ \paragraph{Notation.} A name consists of 1--3 \emph{words} (rarely
+ 4, but not more) that are separated by underscore. There are three
+ variants concerning upper or lower case letters, which are used for
+ certain ML categories as follows:
+
+ \medskip
+ \begin{tabular}{lll}
+ variant & example & ML categories \\\hline
+ lower-case & @{ML_text foo_bar} & values, types, record fields \\
+ capitalized & @{ML_text Foo_Bar} & datatype constructors, structures, functors \\
+ upper-case & @{ML_text FOO_BAR} & special values, exception constructors, signatures \\
+ \end{tabular}
+ \medskip
+
+ For historical reasons, many capitalized names omit underscores,
+ e.g.\ old-style @{ML_text FooBar} instead of @{ML_text Foo_Bar}.
+ Genuine mixed-case names are \emph{not} used, because clear division
+ of words is essential for readability.\footnote{Camel-case was
+ invented to workaround the lack of underscore in some early
+ non-ASCII character sets. Later it became habitual in some language
+ communities that are now strong in numbers.}
+
+ A single (capital) character does not count as ``word'' in this
+ respect: some Isabelle/ML names are suffixed by extra markers like
+ this: @{ML_text foo_barT}.
+
+ Name variants are produced by adding 1--3 primes, e.g.\ @{ML_text
+ foo'}, @{ML_text foo''}, or @{ML_text foo'''}, but not @{ML_text
+ foo''''} or more. Decimal digits scale better to larger numbers,
+ e.g.\ @{ML_text foo0}, @{ML_text foo1}, @{ML_text foo42}.
+
+ \paragraph{Scopes.} Apart from very basic library modules, ML
+ structures are not ``opened'', but names are referenced with
+ explicit qualification, as in @{ML Syntax.string_of_term} for
+ example. When devising names for structures and their components it
+ is important aim at eye-catching compositions of both parts, because
+ this is how they are seen in the sources and documentation. For the
+ same reasons, aliases of well-known library functions should be
+ avoided.
+
+ Local names of function abstraction or case/let bindings are
+ typically shorter, sometimes using only rudiments of ``words'',
+ while still avoiding cryptic shorthands. An auxiliary function
+ called @{ML_text helper}, @{ML_text aux}, or @{ML_text f} is
+ considered bad style.
+
+ Example:
+
+ \begin{verbatim}
+ (* RIGHT *)
+
+ fun print_foo ctxt foo =
+ let
+ fun print t = ... Syntax.string_of_term ctxt t ...
+ in ... end;
+
+
+ (* RIGHT *)
+
+ fun print_foo ctxt foo =
+ let
+ val string_of_term = Syntax.string_of_term ctxt;
+ fun print t = ... string_of_term t ...
+ in ... end;
+
+
+ (* WRONG *)
+
+ val string_of_term = Syntax.string_of_term;
+
+ fun print_foo ctxt foo =
+ let
+ fun aux t = ... string_of_term ctxt t ...
+ in ... end;
+
+ \end{verbatim}
+
+
+ \paragraph{Specific conventions.} Here are some specific name forms
+ that occur frequently in the sources.
+
+ \begin{itemize}
+
+ \item A function that maps @{ML_text foo} to @{ML_text bar} is
+ called @{ML_text foo_to_bar} or @{ML_text bar_of_foo} (never
+ @{ML_text foo2bar}, @{ML_text bar_from_foo}, @{ML_text
+ bar_for_foo}, or @{ML_text bar4foo}).
+
+ \item The name component @{ML_text legacy} means that the operation
+ is about to be discontinued soon.
+
+ \item The name component @{ML_text old} means that this is historic
+ material that might disappear at some later stage.
+
+ \item The name component @{ML_text global} means that this works
+ with the background theory instead of the regular local context
+ (\secref{sec:context}), sometimes for historical reasons, sometimes
+ due a genuine lack of locality of the concept involved, sometimes as
+ a fall-back for the lack of a proper context in the application
+ code. Whenever there is a non-global variant available, the
+ application should be migrated to use it with a proper local
+ context.
+
+ \item Variables of the main context types of the Isabelle/Isar
+ framework (\secref{sec:context} and \chref{ch:local-theory}) have
+ firm naming conventions as follows:
+
+ \begin{itemize}
+
+ \item theories are called @{ML_text thy}, rarely @{ML_text theory}
+ (never @{ML_text thry})
+
+ \item proof contexts are called @{ML_text ctxt}, rarely @{ML_text
+ context} (never @{ML_text ctx})
+
+ \item generic contexts are called @{ML_text context}, rarely
+ @{ML_text ctxt}
+
+ \item local theories are called @{ML_text lthy}, except for local
+ theories that are treated as proof context (which is a semantic
+ super-type)
+
+ \end{itemize}
+
+ Variations with primed or decimal numbers are always possible, as
+ well as sematic prefixes like @{ML_text foo_thy} or @{ML_text
+ bar_ctxt}, but the base conventions above need to be preserved.
+ This allows to visualize the their data flow via plain regular
+ expressions in the editor.
+
+ \item The main logical entities (\secref{ch:logic}) have established
+ naming convention as follows:
+
+ \begin{itemize}
+
+ \item sorts are called @{ML_text S}
+
+ \item types are called @{ML_text T}, @{ML_text U}, or @{ML_text
+ ty} (never @{ML_text t})
+
+ \item terms are called @{ML_text t}, @{ML_text u}, or @{ML_text
+ tm} (never @{ML_text trm})
+
+ \item certified types are called @{ML_text cT}, rarely @{ML_text
+ T}, with variants as for types
+
+ \item certified terms are called @{ML_text ct}, rarely @{ML_text
+ t}, with variants as for terms (never @{ML_text ctrm})
+
+ \item theorems are called @{ML_text th}, or @{ML_text thm}
+
+ \end{itemize}
+
+ Proper semantic names override these conventions completely. For
+ example, the left-hand side of an equation (as a term) can be called
+ @{ML_text lhs} (not @{ML_text lhs_tm}). Or a term that is known
+ to be a variable can be called @{ML_text v} or @{ML_text x}.
+
+ \item Tactics (\secref{sec:tactics}) are sufficiently important to
+ have specific naming conventions. The name of a basic tactic
+ definition always has a @{ML_text "_tac"} suffix, the subgoal index
+ (if applicable) is always called @{ML_text i}, and the goal state
+ (if made explicit) is usually called @{ML_text st} instead of the
+ somewhat misleading @{ML_text thm}. Any other arguments are given
+ before the latter two, and the general context is given first.
+ Example:
+
+ \begin{verbatim}
+ fun my_tac ctxt arg1 arg2 i st = ...
+ \end{verbatim}
+
+ Note that the goal state @{ML_text st} above is rarely made
+ explicit, if tactic combinators (tacticals) are used as usual.
+
+ \end{itemize}
+*}
+
+
+subsection {* General source layout *}
+
+text {* The general Isabelle/ML source layout imitates regular
+ type-setting to some extent, augmented by the requirements for
+ deeply nested expressions that are commonplace in functional
+ programming.
+
+ \paragraph{Line length} is 80 characters according to ancient
+ standards, but we allow as much as 100 characters (not
+ more).\footnote{Readability requires to keep the beginning of a line
+ in view while watching its end. Modern wide-screen displays do not
+ change the way how the human brain works. Sources also need to be
+ printable on plain paper with reasonable font-size.} The extra 20
+ characters acknowledge the space requirements due to qualified
+ library references in Isabelle/ML.
+
+ \paragraph{White-space} is used to emphasize the structure of
+ expressions, following mostly standard conventions for mathematical
+ typesetting, as can be seen in plain {\TeX} or {\LaTeX}. This
+ defines positioning of spaces for parentheses, punctuation, and
+ infixes as illustrated here:
+
+ \begin{verbatim}
+ val x = y + z * (a + b);
+ val pair = (a, b);
+ val record = {foo = 1, bar = 2};
+ \end{verbatim}
+
+ Lines are normally broken \emph{after} an infix operator or
+ punctuation character. For example:
+
+ \begin{verbatim}
+ val x =
+ a +
+ b +
+ c;
+
+ val tuple =
+ (a,
+ b,
+ c);
+ \end{verbatim}
+
+ Some special infixes (e.g.\ @{ML_text "|>"}) work better at the
+ start of the line, but punctuation is always at the end.
+
+ Function application follows the tradition of @{text "\<lambda>"}-calculus,
+ not informal mathematics. For example: @{ML_text "f a b"} for a
+ curried function, or @{ML_text "g (a, b)"} for a tupled function.
+ Note that the space between @{ML_text g} and the pair @{ML_text
+ "(a, b)"} follows the important principle of
+ \emph{compositionality}: the layout of @{ML_text "g p"} does not
+ change when @{ML_text "p"} is refined to the concrete pair
+ @{ML_text "(a, b)"}.
+
+ \paragraph{Indentation} uses plain spaces, never hard
+ tabulators.\footnote{Tabulators were invented to move the carriage
+ of a type-writer to certain predefined positions. In software they
+ could be used as a primitive run-length compression of consecutive
+ spaces, but the precise result would depend on non-standardized
+ editor configuration.}
+
+ Each level of nesting is indented by 2 spaces, sometimes 1, very
+ rarely 4, never 8 or any other odd number.
+
+ Indentation follows a simple logical format that only depends on the
+ nesting depth, not the accidental length of the text that initiates
+ a level of nesting. Example:
+
+ \begin{verbatim}
+ (* RIGHT *)
+
+ if b then
+ expr1_part1
+ expr1_part2
+ else
+ expr2_part1
+ expr2_part2
+
+
+ (* WRONG *)
+
+ if b then expr1_part1
+ expr1_part2
+ else expr2_part1
+ expr2_part2
+ \end{verbatim}
+
+ The second form has many problems: it assumes a fixed-width font
+ when viewing the sources, it uses more space on the line and thus
+ makes it hard to observe its strict length limit (working against
+ \emph{readability}), it requires extra editing to adapt the layout
+ to changes of the initial text (working against
+ \emph{maintainability}) etc.
+
+ \medskip For similar reasons, any kind of two-dimensional or tabular
+ layouts, ASCII-art with lines or boxes of asterisks etc.\ should be
+ avoided.
+
+ \paragraph{Complex expressions} that consist of multi-clausal
+ function definitions, @{ML_text handle}, @{ML_text case},
+ @{ML_text let} (and combinations) require special attention. The
+ syntax of Standard ML is quite ambitious and admits a lot of
+ variance that can distort the meaning of the text.
+
+ Clauses of @{ML_text fun}, @{ML_text fn}, @{ML_text handle},
+ @{ML_text case} get extra indentation to indicate the nesting
+ clearly. Example:
+
+ \begin{verbatim}
+ (* RIGHT *)
+
+ fun foo p1 =
+ expr1
+ | foo p2 =
+ expr2
+
+
+ (* WRONG *)
+
+ fun foo p1 =
+ expr1
+ | foo p2 =
+ expr2
+ \end{verbatim}
+
+ Body expressions consisting of @{ML_text case} or @{ML_text let}
+ require care to maintain compositionality, to prevent loss of
+ logical indentation where it is especially important to see the
+ structure of the text. Example:
+
+ \begin{verbatim}
+ (* RIGHT *)
+
+ fun foo p1 =
+ (case e of
+ q1 => ...
+ | q2 => ...)
+ | foo p2 =
+ let
+ ...
+ in
+ ...
+ end
+
+
+ (* WRONG *)
+
+ fun foo p1 = case e of
+ q1 => ...
+ | q2 => ...
+ | foo p2 =
+ let
+ ...
+ in
+ ...
+ end
+ \end{verbatim}
+
+ Extra parentheses around @{ML_text case} expressions are optional,
+ but help to analyse the nesting based on character matching in the
+ editor.
+
+ \medskip There are two main exceptions to the overall principle of
+ compositionality in the layout of complex expressions.
+
+ \begin{enumerate}
+
+ \item @{ML_text "if"} expressions are iterated as if there would be
+ a multi-branch conditional in SML, e.g.
+
+ \begin{verbatim}
+ (* RIGHT *)
+
+ if b1 then e1
+ else if b2 then e2
+ else e3
+ \end{verbatim}
+
+ \item @{ML_text fn} abstractions are often layed-out as if they
+ would lack any structure by themselves. This traditional form is
+ motivated by the possibility to shift function arguments back and
+ forth wrt.\ additional combinators. Example:
+
+ \begin{verbatim}
+ (* RIGHT *)
+
+ fun foo x y = fold (fn z =>
+ expr)
+ \end{verbatim}
+
+ Here the visual appearance is that of three arguments @{ML_text x},
+ @{ML_text y}, @{ML_text z}.
+
+ \end{enumerate}
+
+ Such weakly structured layout should be use with great care. Here
+ are some counter-examples involving @{ML_text let} expressions:
+
+ \begin{verbatim}
+ (* WRONG *)
+
+ fun foo x = let
+ val y = ...
+ in ... end
+
+
+ (* WRONG *)
+
+ fun foo x = let
+ val y = ...
+ in ... end
+
+
+ (* WRONG *)
+
+ fun foo x =
+ let
+ val y = ...
+ in ... end
+ \end{verbatim}
+
+ \medskip In general the source layout is meant to emphasize the
+ structure of complex language expressions, not to pretend that SML
+ had a completely different syntax (say that of Haskell or Java).
+*}
+
+
+section {* SML embedded into Isabelle/Isar *}
+
+text {* ML and Isar are intertwined via an open-ended bootstrap
+ process that provides more and more programming facilities and
+ logical content in an alternating manner. Bootstrapping starts from
+ the raw environment of existing implementations of Standard ML
+ (mainly Poly/ML, but also SML/NJ).
+
+ Isabelle/Pure marks the point where the original ML toplevel is
+ superseded by the Isar toplevel that maintains a uniform context for
+ arbitrary ML values (see also \secref{sec:context}). This formal
+ environment holds ML compiler bindings, logical entities, and many
+ other things. Raw SML is never encountered again after the initial
+ bootstrap of Isabelle/Pure.
+
+ Object-logics like Isabelle/HOL are built within the
+ Isabelle/ML/Isar environment by introducing suitable theories with
+ associated ML modules, either inlined or as separate files. Thus
+ Isabelle/HOL is defined as a regular user-space application within
+ the Isabelle framework. Further add-on tools can be implemented in
+ ML within the Isar context in the same manner: ML is part of the
+ standard repertoire of Isabelle, and there is no distinction between
+ ``user'' and ``developer'' in this respect.
+*}
+
+
+subsection {* Isar ML commands *}
+
+text {* The primary Isar source language provides facilities to ``open
+ a window'' to the underlying ML compiler. Especially see the Isar
+ commands @{command_ref "ML_file"} and @{command_ref "ML"}: both work the
+ same way, only the source text is provided via a file vs.\ inlined,
+ respectively. Apart from embedding ML into the main theory
+ definition like that, there are many more commands that refer to ML
+ source, such as @{command_ref setup} or @{command_ref declaration}.
+ Even more fine-grained embedding of ML into Isar is encountered in
+ the proof method @{method_ref tactic}, which refines the pending
+ goal state via a given expression of type @{ML_type tactic}.
+*}
+
+text %mlex {* The following artificial example demonstrates some ML
+ toplevel declarations within the implicit Isar theory context. This
+ is regular functional programming without referring to logical
+ entities yet.
+*}
+
+ML {*
+ fun factorial 0 = 1
+ | factorial n = n * factorial (n - 1)
+*}
+
+text {* Here the ML environment is already managed by Isabelle, i.e.\
+ the @{ML factorial} function is not yet accessible in the preceding
+ paragraph, nor in a different theory that is independent from the
+ current one in the import hierarchy.
+
+ Removing the above ML declaration from the source text will remove
+ any trace of this definition as expected. The Isabelle/ML toplevel
+ environment is managed in a \emph{stateless} way: unlike the raw ML
+ toplevel there are no global side-effects involved
+ here.\footnote{Such a stateless compilation environment is also a
+ prerequisite for robust parallel compilation within independent
+ nodes of the implicit theory development graph.}
+
+ \medskip The next example shows how to embed ML into Isar proofs, using
+ @{command_ref "ML_prf"} instead of Instead of @{command_ref "ML"}.
+ As illustrated below, the effect on the ML environment is local to
+ the whole proof body, ignoring the block structure.
+*}
+
+notepad
+begin
+ ML_prf %"ML" {* val a = 1 *}
+ {
+ ML_prf %"ML" {* val b = a + 1 *}
+ } -- {* Isar block structure ignored by ML environment *}
+ ML_prf %"ML" {* val c = b + 1 *}
+end
+
+text {* By side-stepping the normal scoping rules for Isar proof
+ blocks, embedded ML code can refer to the different contexts and
+ manipulate corresponding entities, e.g.\ export a fact from a block
+ context.
+
+ \medskip Two further ML commands are useful in certain situations:
+ @{command_ref ML_val} and @{command_ref ML_command} are
+ \emph{diagnostic} in the sense that there is no effect on the
+ underlying environment, and can thus used anywhere (even outside a
+ theory). The examples below produce long strings of digits by
+ invoking @{ML factorial}: @{command ML_val} already takes care of
+ printing the ML toplevel result, but @{command ML_command} is silent
+ so we produce an explicit output message. *}
+
+ML_val {* factorial 100 *}
+ML_command {* writeln (string_of_int (factorial 100)) *}
+
+notepad
+begin
+ ML_val {* factorial 100 *}
+ ML_command {* writeln (string_of_int (factorial 100)) *}
+end
+
+
+subsection {* Compile-time context *}
+
+text {* Whenever the ML compiler is invoked within Isabelle/Isar, the
+ formal context is passed as a thread-local reference variable. Thus
+ ML code may access the theory context during compilation, by reading
+ or writing the (local) theory under construction. Note that such
+ direct access to the compile-time context is rare. In practice it
+ is typically done via some derived ML functions instead.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML ML_Context.the_generic_context: "unit -> Context.generic"} \\
+ @{index_ML "Context.>>": "(Context.generic -> Context.generic) -> unit"} \\
+ @{index_ML ML_Thms.bind_thms: "string * thm list -> unit"} \\
+ @{index_ML ML_Thms.bind_thm: "string * thm -> unit"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML "ML_Context.the_generic_context ()"} refers to the theory
+ context of the ML toplevel --- at compile time. ML code needs to
+ take care to refer to @{ML "ML_Context.the_generic_context ()"}
+ correctly. Recall that evaluation of a function body is delayed
+ until actual run-time.
+
+ \item @{ML "Context.>>"}~@{text f} applies context transformation
+ @{text f} to the implicit context of the ML toplevel.
+
+ \item @{ML ML_Thms.bind_thms}~@{text "(name, thms)"} stores a list of
+ theorems produced in ML both in the (global) theory context and the
+ ML toplevel, associating it with the provided name. Theorems are
+ put into a global ``standard'' format before being stored.
+
+ \item @{ML ML_Thms.bind_thm} is similar to @{ML ML_Thms.bind_thms} but refers to a
+ singleton fact.
+
+ \end{description}
+
+ It is important to note that the above functions are really
+ restricted to the compile time, even though the ML compiler is
+ invoked at run-time. The majority of ML code either uses static
+ antiquotations (\secref{sec:ML-antiq}) or refers to the theory or
+ proof context at run-time, by explicit functional abstraction.
+*}
+
+
+subsection {* Antiquotations \label{sec:ML-antiq} *}
+
+text {* A very important consequence of embedding SML into Isar is the
+ concept of \emph{ML antiquotation}. The standard token language of
+ ML is augmented by special syntactic entities of the following form:
+
+ @{rail \<open>
+ @{syntax_def antiquote}: '@{' nameref args '}'
+ \<close>}
+
+ Here @{syntax nameref} and @{syntax args} are regular outer syntax
+ categories \cite{isabelle-isar-ref}. Attributes and proof methods
+ use similar syntax.
+
+ \medskip A regular antiquotation @{text "@{name args}"} processes
+ its arguments by the usual means of the Isar source language, and
+ produces corresponding ML source text, either as literal
+ \emph{inline} text (e.g. @{text "@{term t}"}) or abstract
+ \emph{value} (e.g. @{text "@{thm th}"}). This pre-compilation
+ scheme allows to refer to formal entities in a robust manner, with
+ proper static scoping and with some degree of logical checking of
+ small portions of the code.
+*}
+
+
+subsection {* Printing ML values *}
+
+text {* The ML compiler knows about the structure of values according
+ to their static type, and can print them in the manner of the
+ toplevel loop, although the details are non-portable. The
+ antiquotations @{ML_antiquotation_def "make_string"} and
+ @{ML_antiquotation_def "print"} provide a quasi-portable way to
+ refer to this potential capability of the underlying ML system in
+ generic Isabelle/ML sources.
+
+ This is occasionally useful for diagnostic or demonstration
+ purposes. Note that production-quality tools require proper
+ user-level error messages. *}
+
+text %mlantiq {*
+ \begin{matharray}{rcl}
+ @{ML_antiquotation_def "make_string"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "print"} & : & @{text ML_antiquotation} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{ML_antiquotation make_string}
+ ;
+ @@{ML_antiquotation print} @{syntax name}?
+ \<close>}
+
+ \begin{description}
+
+ \item @{text "@{make_string}"} inlines a function to print arbitrary
+ values similar to the ML toplevel. The result is compiler dependent
+ and may fall back on "?" in certain situations.
+
+ \item @{text "@{print f}"} uses the ML function @{text "f: string ->
+ unit"} to output the result of @{text "@{make_string}"} above,
+ together with the source position of the antiquotation. The default
+ output function is @{ML writeln}.
+
+ \end{description}
+*}
+
+text %mlex {* The following artificial examples show how to produce
+ adhoc output of ML values for debugging purposes. *}
+
+ML {*
+ val x = 42;
+ val y = true;
+
+ writeln (@{make_string} {x = x, y = y});
+
+ @{print} {x = x, y = y};
+ @{print tracing} {x = x, y = y};
+*}
+
+
+section {* Canonical argument order \label{sec:canonical-argument-order} *}
+
+text {* Standard ML is a language in the tradition of @{text
+ "\<lambda>"}-calculus and \emph{higher-order functional programming},
+ similar to OCaml, Haskell, or Isabelle/Pure and HOL as logical
+ languages. Getting acquainted with the native style of representing
+ functions in that setting can save a lot of extra boiler-plate of
+ redundant shuffling of arguments, auxiliary abstractions etc.
+
+ Functions are usually \emph{curried}: the idea of turning arguments
+ of type @{text "\<tau>\<^sub>i"} (for @{text "i \<in> {1, \<dots> n}"}) into a result of
+ type @{text "\<tau>"} is represented by the iterated function space
+ @{text "\<tau>\<^sub>1 \<rightarrow> \<dots> \<rightarrow> \<tau>\<^sub>n \<rightarrow> \<tau>"}. This is isomorphic to the well-known
+ encoding via tuples @{text "\<tau>\<^sub>1 \<times> \<dots> \<times> \<tau>\<^sub>n \<rightarrow> \<tau>"}, but the curried
+ version fits more smoothly into the basic calculus.\footnote{The
+ difference is even more significant in higher-order logic, because
+ the redundant tuple structure needs to be accommodated by formal
+ reasoning.}
+
+ Currying gives some flexiblity due to \emph{partial application}. A
+ function @{text "f: \<tau>\<^sub>1 \<rightarrow> \<tau>\<^sub>2 \<rightarrow> \<tau>"} can be applied to @{text "x: \<tau>\<^sub>1"}
+ and the remaining @{text "(f x): \<tau>\<^sub>2 \<rightarrow> \<tau>"} passed to another function
+ etc. How well this works in practice depends on the order of
+ arguments. In the worst case, arguments are arranged erratically,
+ and using a function in a certain situation always requires some
+ glue code. Thus we would get exponentially many oppurtunities to
+ decorate the code with meaningless permutations of arguments.
+
+ This can be avoided by \emph{canonical argument order}, which
+ observes certain standard patterns and minimizes adhoc permutations
+ in their application. In Isabelle/ML, large portions of text can be
+ written without auxiliary operations like @{text "swap: \<alpha> \<times> \<beta> \<rightarrow> \<beta> \<times>
+ \<alpha>"} or @{text "C: (\<alpha> \<rightarrow> \<beta> \<rightarrow> \<gamma>) \<rightarrow> (\<beta> \<rightarrow> \<alpha> \<rightarrow> \<gamma>)"} (the latter not
+ present in the Isabelle/ML library).
+
+ \medskip The basic idea is that arguments that vary less are moved
+ further to the left than those that vary more. Two particularly
+ important categories of functions are \emph{selectors} and
+ \emph{updates}.
+
+ The subsequent scheme is based on a hypothetical set-like container
+ of type @{text "\<beta>"} that manages elements of type @{text "\<alpha>"}. Both
+ the names and types of the associated operations are canonical for
+ Isabelle/ML.
+
+ \begin{center}
+ \begin{tabular}{ll}
+ kind & canonical name and type \\\hline
+ selector & @{text "member: \<beta> \<rightarrow> \<alpha> \<rightarrow> bool"} \\
+ update & @{text "insert: \<alpha> \<rightarrow> \<beta> \<rightarrow> \<beta>"} \\
+ \end{tabular}
+ \end{center}
+
+ Given a container @{text "B: \<beta>"}, the partially applied @{text
+ "member B"} is a predicate over elements @{text "\<alpha> \<rightarrow> bool"}, and
+ thus represents the intended denotation directly. It is customary
+ to pass the abstract predicate to further operations, not the
+ concrete container. The argument order makes it easy to use other
+ combinators: @{text "forall (member B) list"} will check a list of
+ elements for membership in @{text "B"} etc. Often the explicit
+ @{text "list"} is pointless and can be contracted to @{text "forall
+ (member B)"} to get directly a predicate again.
+
+ In contrast, an update operation varies the container, so it moves
+ to the right: @{text "insert a"} is a function @{text "\<beta> \<rightarrow> \<beta>"} to
+ insert a value @{text "a"}. These can be composed naturally as
+ @{text "insert c \<circ> insert b \<circ> insert a"}. The slightly awkward
+ inversion of the composition order is due to conventional
+ mathematical notation, which can be easily amended as explained
+ below.
+*}
+
+
+subsection {* Forward application and composition *}
+
+text {* Regular function application and infix notation works best for
+ relatively deeply structured expressions, e.g.\ @{text "h (f x y + g
+ z)"}. The important special case of \emph{linear transformation}
+ applies a cascade of functions @{text "f\<^sub>n (\<dots> (f\<^sub>1 x))"}. This
+ becomes hard to read and maintain if the functions are themselves
+ given as complex expressions. The notation can be significantly
+ improved by introducing \emph{forward} versions of application and
+ composition as follows:
+
+ \medskip
+ \begin{tabular}{lll}
+ @{text "x |> f"} & @{text "\<equiv>"} & @{text "f x"} \\
+ @{text "(f #> g) x"} & @{text "\<equiv>"} & @{text "x |> f |> g"} \\
+ \end{tabular}
+ \medskip
+
+ This enables to write conveniently @{text "x |> f\<^sub>1 |> \<dots> |> f\<^sub>n"} or
+ @{text "f\<^sub>1 #> \<dots> #> f\<^sub>n"} for its functional abstraction over @{text
+ "x"}.
+
+ \medskip There is an additional set of combinators to accommodate
+ multiple results (via pairs) that are passed on as multiple
+ arguments (via currying).
+
+ \medskip
+ \begin{tabular}{lll}
+ @{text "(x, y) |-> f"} & @{text "\<equiv>"} & @{text "f x y"} \\
+ @{text "(f #-> g) x"} & @{text "\<equiv>"} & @{text "x |> f |-> g"} \\
+ \end{tabular}
+ \medskip
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_op "|> ": "'a * ('a -> 'b) -> 'b"} \\
+ @{index_ML_op "|-> ": "('c * 'a) * ('c -> 'a -> 'b) -> 'b"} \\
+ @{index_ML_op "#> ": "('a -> 'b) * ('b -> 'c) -> 'a -> 'c"} \\
+ @{index_ML_op "#-> ": "('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd"} \\
+ \end{mldecls}
+*}
+
+
+subsection {* Canonical iteration *}
+
+text {* As explained above, a function @{text "f: \<alpha> \<rightarrow> \<beta> \<rightarrow> \<beta>"} can be
+ understood as update on a configuration of type @{text "\<beta>"},
+ parametrized by arguments of type @{text "\<alpha>"}. Given @{text "a: \<alpha>"}
+ the partial application @{text "(f a): \<beta> \<rightarrow> \<beta>"} operates
+ homogeneously on @{text "\<beta>"}. This can be iterated naturally over a
+ list of parameters @{text "[a\<^sub>1, \<dots>, a\<^sub>n]"} as @{text "f a\<^sub>1 #> \<dots> #> f a\<^sub>n"}.
+ The latter expression is again a function @{text "\<beta> \<rightarrow> \<beta>"}.
+ It can be applied to an initial configuration @{text "b: \<beta>"} to
+ start the iteration over the given list of arguments: each @{text
+ "a"} in @{text "a\<^sub>1, \<dots>, a\<^sub>n"} is applied consecutively by updating a
+ cumulative configuration.
+
+ The @{text fold} combinator in Isabelle/ML lifts a function @{text
+ "f"} as above to its iterated version over a list of arguments.
+ Lifting can be repeated, e.g.\ @{text "(fold \<circ> fold) f"} iterates
+ over a list of lists as expected.
+
+ The variant @{text "fold_rev"} works inside-out over the list of
+ arguments, such that @{text "fold_rev f \<equiv> fold f \<circ> rev"} holds.
+
+ The @{text "fold_map"} combinator essentially performs @{text
+ "fold"} and @{text "map"} simultaneously: each application of @{text
+ "f"} produces an updated configuration together with a side-result;
+ the iteration collects all such side-results as a separate list.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML fold: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\
+ @{index_ML fold_rev: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\
+ @{index_ML fold_map: "('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML fold}~@{text f} lifts the parametrized update function
+ @{text "f"} to a list of parameters.
+
+ \item @{ML fold_rev}~@{text "f"} is similar to @{ML fold}~@{text
+ "f"}, but works inside-out.
+
+ \item @{ML fold_map}~@{text "f"} lifts the parametrized update
+ function @{text "f"} (with side-result) to a list of parameters and
+ cumulative side-results.
+
+ \end{description}
+
+ \begin{warn}
+ The literature on functional programming provides a multitude of
+ combinators called @{text "foldl"}, @{text "foldr"} etc. SML97
+ provides its own variations as @{ML List.foldl} and @{ML
+ List.foldr}, while the classic Isabelle library also has the
+ historic @{ML Library.foldl} and @{ML Library.foldr}. To avoid
+ unnecessary complication and confusion, all these historical
+ versions should be ignored, and @{ML fold} (or @{ML fold_rev}) used
+ exclusively.
+ \end{warn}
+*}
+
+text %mlex {* The following example shows how to fill a text buffer
+ incrementally by adding strings, either individually or from a given
+ list.
+*}
+
+ML {*
+ val s =
+ Buffer.empty
+ |> Buffer.add "digits: "
+ |> fold (Buffer.add o string_of_int) (0 upto 9)
+ |> Buffer.content;
+
+ @{assert} (s = "digits: 0123456789");
+*}
+
+text {* Note how @{ML "fold (Buffer.add o string_of_int)"} above saves
+ an extra @{ML "map"} over the given list. This kind of peephole
+ optimization reduces both the code size and the tree structures in
+ memory (``deforestation''), but it requires some practice to read
+ and write fluently.
+
+ \medskip The next example elaborates the idea of canonical
+ iteration, demonstrating fast accumulation of tree content using a
+ text buffer.
+*}
+
+ML {*
+ datatype tree = Text of string | Elem of string * tree list;
+
+ fun slow_content (Text txt) = txt
+ | slow_content (Elem (name, ts)) =
+ "<" ^ name ^ ">" ^
+ implode (map slow_content ts) ^
+ "</" ^ name ^ ">"
+
+ fun add_content (Text txt) = Buffer.add txt
+ | add_content (Elem (name, ts)) =
+ Buffer.add ("<" ^ name ^ ">") #>
+ fold add_content ts #>
+ Buffer.add ("</" ^ name ^ ">");
+
+ fun fast_content tree =
+ Buffer.empty |> add_content tree |> Buffer.content;
+*}
+
+text {* The slow part of @{ML slow_content} is the @{ML implode} of
+ the recursive results, because it copies previously produced strings
+ again.
+
+ The incremental @{ML add_content} avoids this by operating on a
+ buffer that is passed through in a linear fashion. Using @{ML_text
+ "#>"} and contraction over the actual buffer argument saves some
+ additional boiler-plate. Of course, the two @{ML "Buffer.add"}
+ invocations with concatenated strings could have been split into
+ smaller parts, but this would have obfuscated the source without
+ making a big difference in allocations. Here we have done some
+ peephole-optimization for the sake of readability.
+
+ Another benefit of @{ML add_content} is its ``open'' form as a
+ function on buffers that can be continued in further linear
+ transformations, folding etc. Thus it is more compositional than
+ the naive @{ML slow_content}. As realistic example, compare the
+ old-style @{ML "Term.maxidx_of_term: term -> int"} with the newer
+ @{ML "Term.maxidx_term: term -> int -> int"} in Isabelle/Pure.
+
+ Note that @{ML fast_content} above is only defined as example. In
+ many practical situations, it is customary to provide the
+ incremental @{ML add_content} only and leave the initialization and
+ termination to the concrete application by the user.
+*}
+
+
+section {* Message output channels \label{sec:message-channels} *}
+
+text {* Isabelle provides output channels for different kinds of
+ messages: regular output, high-volume tracing information, warnings,
+ and errors.
+
+ Depending on the user interface involved, these messages may appear
+ in different text styles or colours. The standard output for
+ terminal sessions prefixes each line of warnings by @{verbatim
+ "###"} and errors by @{verbatim "***"}, but leaves anything else
+ unchanged.
+
+ Messages are associated with the transaction context of the running
+ Isar command. This enables the front-end to manage commands and
+ resulting messages together. For example, after deleting a command
+ from a given theory document version, the corresponding message
+ output can be retracted from the display.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML writeln: "string -> unit"} \\
+ @{index_ML tracing: "string -> unit"} \\
+ @{index_ML warning: "string -> unit"} \\
+ @{index_ML error: "string -> 'a"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML writeln}~@{text "text"} outputs @{text "text"} as regular
+ message. This is the primary message output operation of Isabelle
+ and should be used by default.
+
+ \item @{ML tracing}~@{text "text"} outputs @{text "text"} as special
+ tracing message, indicating potential high-volume output to the
+ front-end (hundreds or thousands of messages issued by a single
+ command). The idea is to allow the user-interface to downgrade the
+ quality of message display to achieve higher throughput.
+
+ Note that the user might have to take special actions to see tracing
+ output, e.g.\ switch to a different output window. So this channel
+ should not be used for regular output.
+
+ \item @{ML warning}~@{text "text"} outputs @{text "text"} as
+ warning, which typically means some extra emphasis on the front-end
+ side (color highlighting, icons, etc.).
+
+ \item @{ML error}~@{text "text"} raises exception @{ML ERROR}~@{text
+ "text"} and thus lets the Isar toplevel print @{text "text"} on the
+ error channel, which typically means some extra emphasis on the
+ front-end side (color highlighting, icons, etc.).
+
+ This assumes that the exception is not handled before the command
+ terminates. Handling exception @{ML ERROR}~@{text "text"} is a
+ perfectly legal alternative: it means that the error is absorbed
+ without any message output.
+
+ \begin{warn}
+ The actual error channel is accessed via @{ML Output.error_message}, but
+ the old interaction protocol of Proof~General \emph{crashes} if that
+ function is used in regular ML code: error output and toplevel
+ command failure always need to coincide in classic TTY interaction.
+ \end{warn}
+
+ \end{description}
+
+ \begin{warn}
+ Regular Isabelle/ML code should output messages exclusively by the
+ official channels. Using raw I/O on \emph{stdout} or \emph{stderr}
+ instead (e.g.\ via @{ML TextIO.output}) is apt to cause problems in
+ the presence of parallel and asynchronous processing of Isabelle
+ theories. Such raw output might be displayed by the front-end in
+ some system console log, with a low chance that the user will ever
+ see it. Moreover, as a genuine side-effect on global process
+ channels, there is no proper way to retract output when Isar command
+ transactions are reset by the system.
+ \end{warn}
+
+ \begin{warn}
+ The message channels should be used in a message-oriented manner.
+ This means that multi-line output that logically belongs together is
+ issued by a \emph{single} invocation of @{ML writeln} etc.\ with the
+ functional concatenation of all message constituents.
+ \end{warn}
+*}
+
+text %mlex {* The following example demonstrates a multi-line
+ warning. Note that in some situations the user sees only the first
+ line, so the most important point should be made first.
+*}
+
+ML_command {*
+ warning (cat_lines
+ ["Beware the Jabberwock, my son!",
+ "The jaws that bite, the claws that catch!",
+ "Beware the Jubjub Bird, and shun",
+ "The frumious Bandersnatch!"]);
+*}
+
+
+section {* Exceptions \label{sec:exceptions} *}
+
+text {* The Standard ML semantics of strict functional evaluation
+ together with exceptions is rather well defined, but some delicate
+ points need to be observed to avoid that ML programs go wrong
+ despite static type-checking. Exceptions in Isabelle/ML are
+ subsequently categorized as follows.
+
+ \paragraph{Regular user errors.} These are meant to provide
+ informative feedback about malformed input etc.
+
+ The \emph{error} function raises the corresponding \emph{ERROR}
+ exception, with a plain text message as argument. \emph{ERROR}
+ exceptions can be handled internally, in order to be ignored, turned
+ into other exceptions, or cascaded by appending messages. If the
+ corresponding Isabelle/Isar command terminates with an \emph{ERROR}
+ exception state, the toplevel will print the result on the error
+ channel (see \secref{sec:message-channels}).
+
+ It is considered bad style to refer to internal function names or
+ values in ML source notation in user error messages.
+
+ Grammatical correctness of error messages can be improved by
+ \emph{omitting} final punctuation: messages are often concatenated
+ or put into a larger context (e.g.\ augmented with source position).
+ By not insisting in the final word at the origin of the error, the
+ system can perform its administrative tasks more easily and
+ robustly.
+
+ \paragraph{Program failures.} There is a handful of standard
+ exceptions that indicate general failure situations, or failures of
+ core operations on logical entities (types, terms, theorems,
+ theories, see \chref{ch:logic}).
+
+ These exceptions indicate a genuine breakdown of the program, so the
+ main purpose is to determine quickly what has happened where.
+ Traditionally, the (short) exception message would include the name
+ of an ML function, although this is no longer necessary, because the
+ ML runtime system prints a detailed source position of the
+ corresponding @{ML_text raise} keyword.
+
+ \medskip User modules can always introduce their own custom
+ exceptions locally, e.g.\ to organize internal failures robustly
+ without overlapping with existing exceptions. Exceptions that are
+ exposed in module signatures require extra care, though, and should
+ \emph{not} be introduced by default. Surprise by users of a module
+ can be often minimized by using plain user errors instead.
+
+ \paragraph{Interrupts.} These indicate arbitrary system events:
+ both the ML runtime system and the Isabelle/ML infrastructure signal
+ various exceptional situations by raising the special
+ \emph{Interrupt} exception in user code.
+
+ This is the one and only way that physical events can intrude an
+ Isabelle/ML program. Such an interrupt can mean out-of-memory,
+ stack overflow, timeout, internal signaling of threads, or the user
+ producing a console interrupt manually etc. An Isabelle/ML program
+ that intercepts interrupts becomes dependent on physical effects of
+ the environment. Even worse, exception handling patterns that are
+ too general by accident, e.g.\ by mispelled exception constructors,
+ will cover interrupts unintentionally and thus render the program
+ semantics ill-defined.
+
+ Note that the Interrupt exception dates back to the original SML90
+ language definition. It was excluded from the SML97 version to
+ avoid its malign impact on ML program semantics, but without
+ providing a viable alternative. Isabelle/ML recovers physical
+ interruptibility (which is an indispensable tool to implement
+ managed evaluation of command transactions), but requires user code
+ to be strictly transparent wrt.\ interrupts.
+
+ \begin{warn}
+ Isabelle/ML user code needs to terminate promptly on interruption,
+ without guessing at its meaning to the system infrastructure.
+ Temporary handling of interrupts for cleanup of global resources
+ etc.\ needs to be followed immediately by re-raising of the original
+ exception.
+ \end{warn}
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML try: "('a -> 'b) -> 'a -> 'b option"} \\
+ @{index_ML can: "('a -> 'b) -> 'a -> bool"} \\
+ @{index_ML_exception ERROR: string} \\
+ @{index_ML_exception Fail: string} \\
+ @{index_ML Exn.is_interrupt: "exn -> bool"} \\
+ @{index_ML reraise: "exn -> 'a"} \\
+ @{index_ML Runtime.exn_trace: "(unit -> 'a) -> 'a"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML try}~@{text "f x"} makes the partiality of evaluating
+ @{text "f x"} explicit via the option datatype. Interrupts are
+ \emph{not} handled here, i.e.\ this form serves as safe replacement
+ for the \emph{unsafe} version @{ML_text "(SOME"}~@{text "f
+ x"}~@{ML_text "handle _ => NONE)"} that is occasionally seen in
+ books about SML97, not Isabelle/ML.
+
+ \item @{ML can} is similar to @{ML try} with more abstract result.
+
+ \item @{ML ERROR}~@{text "msg"} represents user errors; this
+ exception is normally raised indirectly via the @{ML error} function
+ (see \secref{sec:message-channels}).
+
+ \item @{ML Fail}~@{text "msg"} represents general program failures.
+
+ \item @{ML Exn.is_interrupt} identifies interrupts robustly, without
+ mentioning concrete exception constructors in user code. Handled
+ interrupts need to be re-raised promptly!
+
+ \item @{ML reraise}~@{text "exn"} raises exception @{text "exn"}
+ while preserving its implicit position information (if possible,
+ depending on the ML platform).
+
+ \item @{ML Runtime.exn_trace}~@{ML_text "(fn () =>"}~@{text
+ "e"}@{ML_text ")"} evaluates expression @{text "e"} while printing
+ a full trace of its stack of nested exceptions (if possible,
+ depending on the ML platform).
+
+ Inserting @{ML Runtime.exn_trace} into ML code temporarily is
+ useful for debugging, but not suitable for production code.
+
+ \end{description}
+*}
+
+text %mlantiq {*
+ \begin{matharray}{rcl}
+ @{ML_antiquotation_def "assert"} & : & @{text ML_antiquotation} \\
+ \end{matharray}
+
+ \begin{description}
+
+ \item @{text "@{assert}"} inlines a function
+ @{ML_type "bool -> unit"} that raises @{ML Fail} if the argument is
+ @{ML false}. Due to inlining the source position of failed
+ assertions is included in the error output.
+
+ \end{description}
+*}
+
+
+section {* Strings of symbols \label{sec:symbols} *}
+
+text {* A \emph{symbol} constitutes the smallest textual unit in
+ Isabelle/ML --- raw ML characters are normally not encountered at
+ all! Isabelle strings consist of a sequence of symbols, represented
+ as a packed string or an exploded list of strings. Each symbol is
+ in itself a small string, which has either one of the following
+ forms:
+
+ \begin{enumerate}
+
+ \item a single ASCII character ``@{text "c"}'', for example
+ ``\verb,a,'',
+
+ \item a codepoint according to UTF8 (non-ASCII byte sequence),
+
+ \item a regular symbol ``\verb,\,\verb,<,@{text "ident"}\verb,>,'',
+ for example ``\verb,\,\verb,<alpha>,'',
+
+ \item a control symbol ``\verb,\,\verb,<^,@{text "ident"}\verb,>,'',
+ for example ``\verb,\,\verb,<^bold>,'',
+
+ \item a raw symbol ``\verb,\,\verb,<^raw:,@{text text}\verb,>,''
+ where @{text text} consists of printable characters excluding
+ ``\verb,.,'' and ``\verb,>,'', for example
+ ``\verb,\,\verb,<^raw:$\sum_{i = 1}^n$>,'',
+
+ \item a numbered raw control symbol ``\verb,\,\verb,<^raw,@{text
+ n}\verb,>, where @{text n} consists of digits, for example
+ ``\verb,\,\verb,<^raw42>,''.
+
+ \end{enumerate}
+
+ The @{text "ident"} syntax for symbol names is @{text "letter
+ (letter | digit)\<^sup>*"}, where @{text "letter = A..Za..z"} and @{text
+ "digit = 0..9"}. There are infinitely many regular symbols and
+ control symbols, but a fixed collection of standard symbols is
+ treated specifically. For example, ``\verb,\,\verb,<alpha>,'' is
+ classified as a letter, which means it may occur within regular
+ Isabelle identifiers.
+
+ The character set underlying Isabelle symbols is 7-bit ASCII, but
+ 8-bit character sequences are passed-through unchanged. Unicode/UCS
+ data in UTF-8 encoding is processed in a non-strict fashion, such
+ that well-formed code sequences are recognized
+ accordingly.\footnote{Note that ISO-Latin-1 differs from UTF-8 only
+ in some special punctuation characters that even have replacements
+ within the standard collection of Isabelle symbols. Text consisting
+ of ASCII plus accented letters can be processed in either encoding.}
+ Unicode provides its own collection of mathematical symbols, but
+ within the core Isabelle/ML world there is no link to the standard
+ collection of Isabelle regular symbols.
+
+ \medskip Output of Isabelle symbols depends on the print mode
+ \cite{isabelle-isar-ref}. For example, the standard {\LaTeX}
+ setup of the Isabelle document preparation system would present
+ ``\verb,\,\verb,<alpha>,'' as @{text "\<alpha>"}, and
+ ``\verb,\,\verb,<^bold>,\verb,\,\verb,<alpha>,'' as @{text "\<^bold>\<alpha>"}.
+ On-screen rendering usually works by mapping a finite subset of
+ Isabelle symbols to suitable Unicode characters.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type "Symbol.symbol": string} \\
+ @{index_ML Symbol.explode: "string -> Symbol.symbol list"} \\
+ @{index_ML Symbol.is_letter: "Symbol.symbol -> bool"} \\
+ @{index_ML Symbol.is_digit: "Symbol.symbol -> bool"} \\
+ @{index_ML Symbol.is_quasi: "Symbol.symbol -> bool"} \\
+ @{index_ML Symbol.is_blank: "Symbol.symbol -> bool"} \\
+ \end{mldecls}
+ \begin{mldecls}
+ @{index_ML_type "Symbol.sym"} \\
+ @{index_ML Symbol.decode: "Symbol.symbol -> Symbol.sym"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type "Symbol.symbol"} represents individual Isabelle
+ symbols.
+
+ \item @{ML "Symbol.explode"}~@{text "str"} produces a symbol list
+ from the packed form. This function supersedes @{ML
+ "String.explode"} for virtually all purposes of manipulating text in
+ Isabelle!\footnote{The runtime overhead for exploded strings is
+ mainly that of the list structure: individual symbols that happen to
+ be a singleton string do not require extra memory in Poly/ML.}
+
+ \item @{ML "Symbol.is_letter"}, @{ML "Symbol.is_digit"}, @{ML
+ "Symbol.is_quasi"}, @{ML "Symbol.is_blank"} classify standard
+ symbols according to fixed syntactic conventions of Isabelle, cf.\
+ \cite{isabelle-isar-ref}.
+
+ \item Type @{ML_type "Symbol.sym"} is a concrete datatype that
+ represents the different kinds of symbols explicitly, with
+ constructors @{ML "Symbol.Char"}, @{ML "Symbol.Sym"}, @{ML
+ "Symbol.UTF8"}, @{ML "Symbol.Ctrl"}, @{ML "Symbol.Raw"}.
+
+ \item @{ML "Symbol.decode"} converts the string representation of a
+ symbol into the datatype version.
+
+ \end{description}
+
+ \paragraph{Historical note.} In the original SML90 standard the
+ primitive ML type @{ML_type char} did not exists, and @{ML_text
+ "explode: string -> string list"} produced a list of singleton
+ strings like @{ML "raw_explode: string -> string list"} in
+ Isabelle/ML today. When SML97 came out, Isabelle did not adopt its
+ somewhat anachronistic 8-bit or 16-bit characters, but the idea of
+ exploding a string into a list of small strings was extended to
+ ``symbols'' as explained above. Thus Isabelle sources can refer to
+ an infinite store of user-defined symbols, without having to worry
+ about the multitude of Unicode encodings that have emerged over the
+ years. *}
+
+
+section {* Basic data types *}
+
+text {* The basis library proposal of SML97 needs to be treated with
+ caution. Many of its operations simply do not fit with important
+ Isabelle/ML conventions (like ``canonical argument order'', see
+ \secref{sec:canonical-argument-order}), others cause problems with
+ the parallel evaluation model of Isabelle/ML (such as @{ML
+ TextIO.print} or @{ML OS.Process.system}).
+
+ Subsequently we give a brief overview of important operations on
+ basic ML data types.
+*}
+
+
+subsection {* Characters *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type char} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type char} is \emph{not} used. The smallest textual
+ unit in Isabelle is represented as a ``symbol'' (see
+ \secref{sec:symbols}).
+
+ \end{description}
+*}
+
+
+subsection {* Strings *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type string} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type string} represents immutable vectors of 8-bit
+ characters. There are operations in SML to convert back and forth
+ to actual byte vectors, which are seldom used.
+
+ This historically important raw text representation is used for
+ Isabelle-specific purposes with the following implicit substructures
+ packed into the string content:
+
+ \begin{enumerate}
+
+ \item sequence of Isabelle symbols (see also \secref{sec:symbols}),
+ with @{ML Symbol.explode} as key operation;
+
+ \item XML tree structure via YXML (see also \cite{isabelle-sys}),
+ with @{ML YXML.parse_body} as key operation.
+
+ \end{enumerate}
+
+ Note that Isabelle/ML string literals may refer Isabelle symbols
+ like ``\verb,\,\verb,<alpha>,'' natively, \emph{without} escaping
+ the backslash. This is a consequence of Isabelle treating all
+ source text as strings of symbols, instead of raw characters.
+
+ \end{description}
+*}
+
+text %mlex {* The subsequent example illustrates the difference of
+ physical addressing of bytes versus logical addressing of symbols in
+ Isabelle strings.
+*}
+
+ML_val {*
+ val s = "\<A>";
+
+ @{assert} (length (Symbol.explode s) = 1);
+ @{assert} (size s = 4);
+*}
+
+text {* Note that in Unicode renderings of the symbol @{text "\<A>"},
+ variations of encodings like UTF-8 or UTF-16 pose delicate questions
+ about the multi-byte representations its codepoint, which is outside
+ of the 16-bit address space of the original Unicode standard from
+ the 1990-ies. In Isabelle/ML it is just ``\verb,\,\verb,<A>,''
+ literally, using plain ASCII characters beyond any doubts. *}
+
+
+subsection {* Integers *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type int} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type int} represents regular mathematical integers,
+ which are \emph{unbounded}. Overflow never happens in
+ practice.\footnote{The size limit for integer bit patterns in memory
+ is 64\,MB for 32-bit Poly/ML, and much higher for 64-bit systems.}
+ This works uniformly for all supported ML platforms (Poly/ML and
+ SML/NJ).
+
+ Literal integers in ML text are forced to be of this one true
+ integer type --- adhoc overloading of SML97 is disabled.
+
+ Structure @{ML_structure IntInf} of SML97 is obsolete and superseded by
+ @{ML_structure Int}. Structure @{ML_structure Integer} in @{file
+ "~~/src/Pure/General/integer.ML"} provides some additional
+ operations.
+
+ \end{description}
+*}
+
+
+subsection {* Time *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type Time.time} \\
+ @{index_ML seconds: "real -> Time.time"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type Time.time} represents time abstractly according
+ to the SML97 basis library definition. This is adequate for
+ internal ML operations, but awkward in concrete time specifications.
+
+ \item @{ML seconds}~@{text "s"} turns the concrete scalar @{text
+ "s"} (measured in seconds) into an abstract time value. Floating
+ point numbers are easy to use as configuration options in the
+ context (see \secref{sec:config-options}) or system preferences that
+ are maintained externally.
+
+ \end{description}
+*}
+
+
+subsection {* Options *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML Option.map: "('a -> 'b) -> 'a option -> 'b option"} \\
+ @{index_ML is_some: "'a option -> bool"} \\
+ @{index_ML is_none: "'a option -> bool"} \\
+ @{index_ML the: "'a option -> 'a"} \\
+ @{index_ML these: "'a list option -> 'a list"} \\
+ @{index_ML the_list: "'a option -> 'a list"} \\
+ @{index_ML the_default: "'a -> 'a option -> 'a"} \\
+ \end{mldecls}
+*}
+
+text {* Apart from @{ML Option.map} most other operations defined in
+ structure @{ML_structure Option} are alien to Isabelle/ML an never
+ used. The operations shown above are defined in @{file
+ "~~/src/Pure/General/basics.ML"}. *}
+
+
+subsection {* Lists *}
+
+text {* Lists are ubiquitous in ML as simple and light-weight
+ ``collections'' for many everyday programming tasks. Isabelle/ML
+ provides important additions and improvements over operations that
+ are predefined in the SML97 library. *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML cons: "'a -> 'a list -> 'a list"} \\
+ @{index_ML member: "('b * 'a -> bool) -> 'a list -> 'b -> bool"} \\
+ @{index_ML insert: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\
+ @{index_ML remove: "('b * 'a -> bool) -> 'b -> 'a list -> 'a list"} \\
+ @{index_ML update: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML cons}~@{text "x xs"} evaluates to @{text "x :: xs"}.
+
+ Tupled infix operators are a historical accident in Standard ML.
+ The curried @{ML cons} amends this, but it should be only used when
+ partial application is required.
+
+ \item @{ML member}, @{ML insert}, @{ML remove}, @{ML update} treat
+ lists as a set-like container that maintains the order of elements.
+ See @{file "~~/src/Pure/library.ML"} for the full specifications
+ (written in ML). There are some further derived operations like
+ @{ML union} or @{ML inter}.
+
+ Note that @{ML insert} is conservative about elements that are
+ already a @{ML member} of the list, while @{ML update} ensures that
+ the latest entry is always put in front. The latter discipline is
+ often more appropriate in declarations of context data
+ (\secref{sec:context-data}) that are issued by the user in Isar
+ source: later declarations take precedence over earlier ones.
+
+ \end{description}
+*}
+
+text %mlex {* Using canonical @{ML fold} together with @{ML cons} (or
+ similar standard operations) alternates the orientation of data.
+ The is quite natural and should not be altered forcible by inserting
+ extra applications of @{ML rev}. The alternative @{ML fold_rev} can
+ be used in the few situations, where alternation should be
+ prevented.
+*}
+
+ML {*
+ val items = [1, 2, 3, 4, 5, 6, 7, 8, 9, 10];
+
+ val list1 = fold cons items [];
+ @{assert} (list1 = rev items);
+
+ val list2 = fold_rev cons items [];
+ @{assert} (list2 = items);
+*}
+
+text {* The subsequent example demonstrates how to \emph{merge} two
+ lists in a natural way. *}
+
+ML {*
+ fun merge_lists eq (xs, ys) = fold_rev (insert eq) ys xs;
+*}
+
+text {* Here the first list is treated conservatively: only the new
+ elements from the second list are inserted. The inside-out order of
+ insertion via @{ML fold_rev} attempts to preserve the order of
+ elements in the result.
+
+ This way of merging lists is typical for context data
+ (\secref{sec:context-data}). See also @{ML merge} as defined in
+ @{file "~~/src/Pure/library.ML"}.
+*}
+
+
+subsection {* Association lists *}
+
+text {* The operations for association lists interpret a concrete list
+ of pairs as a finite function from keys to values. Redundant
+ representations with multiple occurrences of the same key are
+ implicitly normalized: lookup and update only take the first
+ occurrence into account.
+*}
+
+text {*
+ \begin{mldecls}
+ @{index_ML AList.lookup: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> 'c option"} \\
+ @{index_ML AList.defined: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> bool"} \\
+ @{index_ML AList.update: "('a * 'a -> bool) -> 'a * 'b -> ('a * 'b) list -> ('a * 'b) list"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML AList.lookup}, @{ML AList.defined}, @{ML AList.update}
+ implement the main ``framework operations'' for mappings in
+ Isabelle/ML, following standard conventions for their names and
+ types.
+
+ Note that a function called @{text lookup} is obliged to express its
+ partiality via an explicit option element. There is no choice to
+ raise an exception, without changing the name to something like
+ @{text "the_element"} or @{text "get"}.
+
+ The @{text "defined"} operation is essentially a contraction of @{ML
+ is_some} and @{text "lookup"}, but this is sufficiently frequent to
+ justify its independent existence. This also gives the
+ implementation some opportunity for peep-hole optimization.
+
+ \end{description}
+
+ Association lists are adequate as simple and light-weight
+ implementation of finite mappings in many practical situations. A
+ more heavy-duty table structure is defined in @{file
+ "~~/src/Pure/General/table.ML"}; that version scales easily to
+ thousands or millions of elements.
+*}
+
+
+subsection {* Unsynchronized references *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type "'a Unsynchronized.ref"} \\
+ @{index_ML Unsynchronized.ref: "'a -> 'a Unsynchronized.ref"} \\
+ @{index_ML "!": "'a Unsynchronized.ref -> 'a"} \\
+ @{index_ML_op ":=": "'a Unsynchronized.ref * 'a -> unit"} \\
+ \end{mldecls}
+*}
+
+text {* Due to ubiquitous parallelism in Isabelle/ML (see also
+ \secref{sec:multi-threading}), the mutable reference cells of
+ Standard ML are notorious for causing problems. In a highly
+ parallel system, both correctness \emph{and} performance are easily
+ degraded when using mutable data.
+
+ The unwieldy name of @{ML Unsynchronized.ref} for the constructor
+ for references in Isabelle/ML emphasizes the inconveniences caused by
+ mutability. Existing operations @{ML "!"} and @{ML_op ":="} are
+ unchanged, but should be used with special precautions, say in a
+ strictly local situation that is guaranteed to be restricted to
+ sequential evaluation --- now and in the future.
+
+ \begin{warn}
+ Never @{ML_text "open Unsynchronized"}, not even in a local scope!
+ Pretending that mutable state is no problem is a very bad idea.
+ \end{warn}
+*}
+
+
+section {* Thread-safe programming \label{sec:multi-threading} *}
+
+text {* Multi-threaded execution has become an everyday reality in
+ Isabelle since Poly/ML 5.2.1 and Isabelle2008. Isabelle/ML provides
+ implicit and explicit parallelism by default, and there is no way
+ for user-space tools to ``opt out''. ML programs that are purely
+ functional, output messages only via the official channels
+ (\secref{sec:message-channels}), and do not intercept interrupts
+ (\secref{sec:exceptions}) can participate in the multi-threaded
+ environment immediately without further ado.
+
+ More ambitious tools with more fine-grained interaction with the
+ environment need to observe the principles explained below.
+*}
+
+
+subsection {* Multi-threading with shared memory *}
+
+text {* Multiple threads help to organize advanced operations of the
+ system, such as real-time conditions on command transactions,
+ sub-components with explicit communication, general asynchronous
+ interaction etc. Moreover, parallel evaluation is a prerequisite to
+ make adequate use of the CPU resources that are available on
+ multi-core systems.\footnote{Multi-core computing does not mean that
+ there are ``spare cycles'' to be wasted. It means that the
+ continued exponential speedup of CPU performance due to ``Moore's
+ Law'' follows different rules: clock frequency has reached its peak
+ around 2005, and applications need to be parallelized in order to
+ avoid a perceived loss of performance. See also
+ \cite{Sutter:2005}.}
+
+ Isabelle/Isar exploits the inherent structure of theories and proofs
+ to support \emph{implicit parallelism} to a large extent. LCF-style
+ theorem provides almost ideal conditions for that, see also
+ \cite{Wenzel:2009}. This means, significant parts of theory and
+ proof checking is parallelized by default. In Isabelle2013, a
+ maximum speedup-factor of 3.5 on 4 cores and 6.5 on 8 cores can be
+ expected.
+
+ \medskip ML threads lack the memory protection of separate
+ processes, and operate concurrently on shared heap memory. This has
+ the advantage that results of independent computations are directly
+ available to other threads: abstract values can be passed without
+ copying or awkward serialization that is typically required for
+ separate processes.
+
+ To make shared-memory multi-threading work robustly and efficiently,
+ some programming guidelines need to be observed. While the ML
+ system is responsible to maintain basic integrity of the
+ representation of ML values in memory, the application programmer
+ needs to ensure that multi-threaded execution does not break the
+ intended semantics.
+
+ \begin{warn}
+ To participate in implicit parallelism, tools need to be
+ thread-safe. A single ill-behaved tool can affect the stability and
+ performance of the whole system.
+ \end{warn}
+
+ Apart from observing the principles of thread-safeness passively,
+ advanced tools may also exploit parallelism actively, e.g.\ by using
+ ``future values'' (\secref{sec:futures}) or the more basic library
+ functions for parallel list operations (\secref{sec:parlist}).
+
+ \begin{warn}
+ Parallel computing resources are managed centrally by the
+ Isabelle/ML infrastructure. User programs must not fork their own
+ ML threads to perform computations.
+ \end{warn}
+*}
+
+
+subsection {* Critical shared resources *}
+
+text {* Thread-safeness is mainly concerned about concurrent
+ read/write access to shared resources, which are outside the purely
+ functional world of ML. This covers the following in particular.
+
+ \begin{itemize}
+
+ \item Global references (or arrays), i.e.\ mutable memory cells that
+ persist over several invocations of associated
+ operations.\footnote{This is independent of the visibility of such
+ mutable values in the toplevel scope.}
+
+ \item Global state of the running Isabelle/ML process, i.e.\ raw I/O
+ channels, environment variables, current working directory.
+
+ \item Writable resources in the file-system that are shared among
+ different threads or external processes.
+
+ \end{itemize}
+
+ Isabelle/ML provides various mechanisms to avoid critical shared
+ resources in most situations. As last resort there are some
+ mechanisms for explicit synchronization. The following guidelines
+ help to make Isabelle/ML programs work smoothly in a concurrent
+ environment.
+
+ \begin{itemize}
+
+ \item Avoid global references altogether. Isabelle/Isar maintains a
+ uniform context that incorporates arbitrary data declared by user
+ programs (\secref{sec:context-data}). This context is passed as
+ plain value and user tools can get/map their own data in a purely
+ functional manner. Configuration options within the context
+ (\secref{sec:config-options}) provide simple drop-in replacements
+ for historic reference variables.
+
+ \item Keep components with local state information re-entrant.
+ Instead of poking initial values into (private) global references, a
+ new state record can be created on each invocation, and passed
+ through any auxiliary functions of the component. The state record
+ may well contain mutable references, without requiring any special
+ synchronizations, as long as each invocation gets its own copy.
+
+ \item Avoid raw output on @{text "stdout"} or @{text "stderr"}. The
+ Poly/ML library is thread-safe for each individual output operation,
+ but the ordering of parallel invocations is arbitrary. This means
+ raw output will appear on some system console with unpredictable
+ interleaving of atomic chunks.
+
+ Note that this does not affect regular message output channels
+ (\secref{sec:message-channels}). An official message is associated
+ with the command transaction from where it originates, independently
+ of other transactions. This means each running Isar command has
+ effectively its own set of message channels, and interleaving can
+ only happen when commands use parallelism internally (and only at
+ message boundaries).
+
+ \item Treat environment variables and the current working directory
+ of the running process as strictly read-only.
+
+ \item Restrict writing to the file-system to unique temporary files.
+ Isabelle already provides a temporary directory that is unique for
+ the running process, and there is a centralized source of unique
+ serial numbers in Isabelle/ML. Thus temporary files that are passed
+ to to some external process will be always disjoint, and thus
+ thread-safe.
+
+ \end{itemize}
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML File.tmp_path: "Path.T -> Path.T"} \\
+ @{index_ML serial_string: "unit -> string"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML File.tmp_path}~@{text "path"} relocates the base
+ component of @{text "path"} into the unique temporary directory of
+ the running Isabelle/ML process.
+
+ \item @{ML serial_string}~@{text "()"} creates a new serial number
+ that is unique over the runtime of the Isabelle/ML process.
+
+ \end{description}
+*}
+
+text %mlex {* The following example shows how to create unique
+ temporary file names.
+*}
+
+ML {*
+ val tmp1 = File.tmp_path (Path.basic ("foo" ^ serial_string ()));
+ val tmp2 = File.tmp_path (Path.basic ("foo" ^ serial_string ()));
+ @{assert} (tmp1 <> tmp2);
+*}
+
+
+subsection {* Explicit synchronization *}
+
+text {* Isabelle/ML also provides some explicit synchronization
+ mechanisms, for the rare situations where mutable shared resources
+ are really required. These are based on the synchronizations
+ primitives of Poly/ML, which have been adapted to the specific
+ assumptions of the concurrent Isabelle/ML environment. User code
+ must not use the Poly/ML primitives directly!
+
+ \medskip The most basic synchronization concept is a single
+ \emph{critical section} (also called ``monitor'' in the literature).
+ A thread that enters the critical section prevents all other threads
+ from doing the same. A thread that is already within the critical
+ section may re-enter it in an idempotent manner.
+
+ Such centralized locking is convenient, because it prevents
+ deadlocks by construction.
+
+ \medskip More fine-grained locking works via \emph{synchronized
+ variables}. An explicit state component is associated with
+ mechanisms for locking and signaling. There are operations to
+ await a condition, change the state, and signal the change to all
+ other waiting threads.
+
+ Here the synchronized access to the state variable is \emph{not}
+ re-entrant: direct or indirect nesting within the same thread will
+ cause a deadlock!
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML NAMED_CRITICAL: "string -> (unit -> 'a) -> 'a"} \\
+ @{index_ML CRITICAL: "(unit -> 'a) -> 'a"} \\
+ \end{mldecls}
+ \begin{mldecls}
+ @{index_ML_type "'a Synchronized.var"} \\
+ @{index_ML Synchronized.var: "string -> 'a -> 'a Synchronized.var"} \\
+ @{index_ML Synchronized.guarded_access: "'a Synchronized.var ->
+ ('a -> ('b * 'a) option) -> 'b"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML NAMED_CRITICAL}~@{text "name e"} evaluates @{text "e ()"}
+ within the central critical section of Isabelle/ML. No other thread
+ may do so at the same time, but non-critical parallel execution will
+ continue. The @{text "name"} argument is used for tracing and might
+ help to spot sources of congestion.
+
+ Entering the critical section without contention is very fast. Each
+ thread should stay within the critical section only very briefly,
+ otherwise parallel performance may degrade.
+
+ \item @{ML CRITICAL} is the same as @{ML NAMED_CRITICAL} with empty
+ name argument.
+
+ \item Type @{ML_type "'a Synchronized.var"} represents synchronized
+ variables with state of type @{ML_type 'a}.
+
+ \item @{ML Synchronized.var}~@{text "name x"} creates a synchronized
+ variable that is initialized with value @{text "x"}. The @{text
+ "name"} is used for tracing.
+
+ \item @{ML Synchronized.guarded_access}~@{text "var f"} lets the
+ function @{text "f"} operate within a critical section on the state
+ @{text "x"} as follows: if @{text "f x"} produces @{ML NONE}, it
+ continues to wait on the internal condition variable, expecting that
+ some other thread will eventually change the content in a suitable
+ manner; if @{text "f x"} produces @{ML SOME}~@{text "(y, x')"} it is
+ satisfied and assigns the new state value @{text "x'"}, broadcasts a
+ signal to all waiting threads on the associated condition variable,
+ and returns the result @{text "y"}.
+
+ \end{description}
+
+ There are some further variants of the @{ML
+ Synchronized.guarded_access} combinator, see @{file
+ "~~/src/Pure/Concurrent/synchronized.ML"} for details.
+*}
+
+text %mlex {* The following example implements a counter that produces
+ positive integers that are unique over the runtime of the Isabelle
+ process:
+*}
+
+ML {*
+ local
+ val counter = Synchronized.var "counter" 0;
+ in
+ fun next () =
+ Synchronized.guarded_access counter
+ (fn i =>
+ let val j = i + 1
+ in SOME (j, j) end);
+ end;
+*}
+
+ML {*
+ val a = next ();
+ val b = next ();
+ @{assert} (a <> b);
+*}
+
+text {* \medskip See @{file "~~/src/Pure/Concurrent/mailbox.ML"} how
+ to implement a mailbox as synchronized variable over a purely
+ functional queue. *}
+
+
+section {* Managed evaluation *}
+
+text {* Execution of Standard ML follows the model of strict
+ functional evaluation with optional exceptions. Evaluation happens
+ whenever some function is applied to (sufficiently many)
+ arguments. The result is either an explicit value or an implicit
+ exception.
+
+ \emph{Managed evaluation} in Isabelle/ML organizes expressions and
+ results to control certain physical side-conditions, to say more
+ specifically when and how evaluation happens. For example, the
+ Isabelle/ML library supports lazy evaluation with memoing, parallel
+ evaluation via futures, asynchronous evaluation via promises,
+ evaluation with time limit etc.
+
+ \medskip An \emph{unevaluated expression} is represented either as
+ unit abstraction @{verbatim "fn () => a"} of type
+ @{verbatim "unit -> 'a"} or as regular function
+ @{verbatim "fn a => b"} of type @{verbatim "'a -> 'b"}. Both forms
+ occur routinely, and special care is required to tell them apart ---
+ the static type-system of SML is only of limited help here.
+
+ The first form is more intuitive: some combinator @{text "(unit ->
+ 'a) -> 'a"} applies the given function to @{text "()"} to initiate
+ the postponed evaluation process. The second form is more flexible:
+ some combinator @{text "('a -> 'b) -> 'a -> 'b"} acts like a
+ modified form of function application; several such combinators may
+ be cascaded to modify a given function, before it is ultimately
+ applied to some argument.
+
+ \medskip \emph{Reified results} make the disjoint sum of regular
+ values versions exceptional situations explicit as ML datatype:
+ @{text "'a result = Res of 'a | Exn of exn"}. This is typically
+ used for administrative purposes, to store the overall outcome of an
+ evaluation process.
+
+ \emph{Parallel exceptions} aggregate reified results, such that
+ multiple exceptions are digested as a collection in canonical form
+ that identifies exceptions according to their original occurrence.
+ This is particular important for parallel evaluation via futures
+ \secref{sec:futures}, which are organized as acyclic graph of
+ evaluations that depend on other evaluations: exceptions stemming
+ from shared sub-graphs are exposed exactly once and in the order of
+ their original occurrence (e.g.\ when printed at the toplevel).
+ Interrupt counts as neutral element here: it is treated as minimal
+ information about some canceled evaluation process, and is absorbed
+ by the presence of regular program exceptions. *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type "'a Exn.result"} \\
+ @{index_ML Exn.capture: "('a -> 'b) -> 'a -> 'b Exn.result"} \\
+ @{index_ML Exn.interruptible_capture: "('a -> 'b) -> 'a -> 'b Exn.result"} \\
+ @{index_ML Exn.release: "'a Exn.result -> 'a"} \\
+ @{index_ML Par_Exn.release_all: "'a Exn.result list -> 'a list"} \\
+ @{index_ML Par_Exn.release_first: "'a Exn.result list -> 'a list"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type "'a Exn.result"} represents the disjoint sum of
+ ML results explicitly, with constructor @{ML Exn.Res} for regular
+ values and @{ML "Exn.Exn"} for exceptions.
+
+ \item @{ML Exn.capture}~@{text "f x"} manages the evaluation of
+ @{text "f x"} such that exceptions are made explicit as @{ML
+ "Exn.Exn"}. Note that this includes physical interrupts (see also
+ \secref{sec:exceptions}), so the same precautions apply to user
+ code: interrupts must not be absorbed accidentally!
+
+ \item @{ML Exn.interruptible_capture} is similar to @{ML
+ Exn.capture}, but interrupts are immediately re-raised as required
+ for user code.
+
+ \item @{ML Exn.release}~@{text "result"} releases the original
+ runtime result, exposing its regular value or raising the reified
+ exception.
+
+ \item @{ML Par_Exn.release_all}~@{text "results"} combines results
+ that were produced independently (e.g.\ by parallel evaluation). If
+ all results are regular values, that list is returned. Otherwise,
+ the collection of all exceptions is raised, wrapped-up as collective
+ parallel exception. Note that the latter prevents access to
+ individual exceptions by conventional @{verbatim "handle"} of SML.
+
+ \item @{ML Par_Exn.release_first} is similar to @{ML
+ Par_Exn.release_all}, but only the first exception that has occurred
+ in the original evaluation process is raised again, the others are
+ ignored. That single exception may get handled by conventional
+ means in SML.
+
+ \end{description}
+*}
+
+
+subsection {* Parallel skeletons \label{sec:parlist} *}
+
+text {*
+ Algorithmic skeletons are combinators that operate on lists in
+ parallel, in the manner of well-known @{text map}, @{text exists},
+ @{text forall} etc. Management of futures (\secref{sec:futures})
+ and their results as reified exceptions is wrapped up into simple
+ programming interfaces that resemble the sequential versions.
+
+ What remains is the application-specific problem to present
+ expressions with suitable \emph{granularity}: each list element
+ corresponds to one evaluation task. If the granularity is too
+ coarse, the available CPUs are not saturated. If it is too
+ fine-grained, CPU cycles are wasted due to the overhead of
+ organizing parallel processing. In the worst case, parallel
+ performance will be less than the sequential counterpart!
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML Par_List.map: "('a -> 'b) -> 'a list -> 'b list"} \\
+ @{index_ML Par_List.get_some: "('a -> 'b option) -> 'a list -> 'b option"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML Par_List.map}~@{text "f [x\<^sub>1, \<dots>, x\<^sub>n]"} is like @{ML
+ "map"}~@{text "f [x\<^sub>1, \<dots>, x\<^sub>n]"}, but the evaluation of @{text "f x\<^sub>i"}
+ for @{text "i = 1, \<dots>, n"} is performed in parallel.
+
+ An exception in any @{text "f x\<^sub>i"} cancels the overall evaluation
+ process. The final result is produced via @{ML
+ Par_Exn.release_first} as explained above, which means the first
+ program exception that happened to occur in the parallel evaluation
+ is propagated, and all other failures are ignored.
+
+ \item @{ML Par_List.get_some}~@{text "f [x\<^sub>1, \<dots>, x\<^sub>n]"} produces some
+ @{text "f x\<^sub>i"} that is of the form @{text "SOME y\<^sub>i"}, if that
+ exists, otherwise @{text "NONE"}. Thus it is similar to @{ML
+ Library.get_first}, but subject to a non-deterministic parallel
+ choice process. The first successful result cancels the overall
+ evaluation process; other exceptions are propagated as for @{ML
+ Par_List.map}.
+
+ This generic parallel choice combinator is the basis for derived
+ forms, such as @{ML Par_List.find_some}, @{ML Par_List.exists}, @{ML
+ Par_List.forall}.
+
+ \end{description}
+*}
+
+text %mlex {* Subsequently, the Ackermann function is evaluated in
+ parallel for some ranges of arguments. *}
+
+ML_val {*
+ fun ackermann 0 n = n + 1
+ | ackermann m 0 = ackermann (m - 1) 1
+ | ackermann m n = ackermann (m - 1) (ackermann m (n - 1));
+
+ Par_List.map (ackermann 2) (500 upto 1000);
+ Par_List.map (ackermann 3) (5 upto 10);
+*}
+
+
+subsection {* Lazy evaluation *}
+
+text {*
+ %FIXME
+
+ See also @{file "~~/src/Pure/Concurrent/lazy.ML"}.
+*}
+
+
+subsection {* Future values \label{sec:futures} *}
+
+text {*
+ %FIXME
+
+ See also @{file "~~/src/Pure/Concurrent/future.ML"}.
+*}
+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Implementation/Prelim.thy Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,1069 @@
+theory Prelim
+imports Base
+begin
+
+chapter {* Preliminaries *}
+
+section {* Contexts \label{sec:context} *}
+
+text {*
+ A logical context represents the background that is required for
+ formulating statements and composing proofs. It acts as a medium to
+ produce formal content, depending on earlier material (declarations,
+ results etc.).
+
+ For example, derivations within the Isabelle/Pure logic can be
+ described as a judgment @{text "\<Gamma> \<turnstile>\<^sub>\<Theta> \<phi>"}, which means that a
+ proposition @{text "\<phi>"} is derivable from hypotheses @{text "\<Gamma>"}
+ within the theory @{text "\<Theta>"}. There are logical reasons for
+ keeping @{text "\<Theta>"} and @{text "\<Gamma>"} separate: theories can be
+ liberal about supporting type constructors and schematic
+ polymorphism of constants and axioms, while the inner calculus of
+ @{text "\<Gamma> \<turnstile> \<phi>"} is strictly limited to Simple Type Theory (with
+ fixed type variables in the assumptions).
+
+ \medskip Contexts and derivations are linked by the following key
+ principles:
+
+ \begin{itemize}
+
+ \item Transfer: monotonicity of derivations admits results to be
+ transferred into a \emph{larger} context, i.e.\ @{text "\<Gamma> \<turnstile>\<^sub>\<Theta>
+ \<phi>"} implies @{text "\<Gamma>' \<turnstile>\<^sub>\<Theta>\<^sub>' \<phi>"} for contexts @{text "\<Theta>'
+ \<supseteq> \<Theta>"} and @{text "\<Gamma>' \<supseteq> \<Gamma>"}.
+
+ \item Export: discharge of hypotheses admits results to be exported
+ into a \emph{smaller} context, i.e.\ @{text "\<Gamma>' \<turnstile>\<^sub>\<Theta> \<phi>"}
+ implies @{text "\<Gamma> \<turnstile>\<^sub>\<Theta> \<Delta> \<Longrightarrow> \<phi>"} where @{text "\<Gamma>' \<supseteq> \<Gamma>"} and
+ @{text "\<Delta> = \<Gamma>' - \<Gamma>"}. Note that @{text "\<Theta>"} remains unchanged here,
+ only the @{text "\<Gamma>"} part is affected.
+
+ \end{itemize}
+
+ \medskip By modeling the main characteristics of the primitive
+ @{text "\<Theta>"} and @{text "\<Gamma>"} above, and abstracting over any
+ particular logical content, we arrive at the fundamental notions of
+ \emph{theory context} and \emph{proof context} in Isabelle/Isar.
+ These implement a certain policy to manage arbitrary \emph{context
+ data}. There is a strongly-typed mechanism to declare new kinds of
+ data at compile time.
+
+ The internal bootstrap process of Isabelle/Pure eventually reaches a
+ stage where certain data slots provide the logical content of @{text
+ "\<Theta>"} and @{text "\<Gamma>"} sketched above, but this does not stop there!
+ Various additional data slots support all kinds of mechanisms that
+ are not necessarily part of the core logic.
+
+ For example, there would be data for canonical introduction and
+ elimination rules for arbitrary operators (depending on the
+ object-logic and application), which enables users to perform
+ standard proof steps implicitly (cf.\ the @{text "rule"} method
+ \cite{isabelle-isar-ref}).
+
+ \medskip Thus Isabelle/Isar is able to bring forth more and more
+ concepts successively. In particular, an object-logic like
+ Isabelle/HOL continues the Isabelle/Pure setup by adding specific
+ components for automated reasoning (classical reasoner, tableau
+ prover, structured induction etc.) and derived specification
+ mechanisms (inductive predicates, recursive functions etc.). All of
+ this is ultimately based on the generic data management by theory
+ and proof contexts introduced here.
+*}
+
+
+subsection {* Theory context \label{sec:context-theory} *}
+
+text {* A \emph{theory} is a data container with explicit name and
+ unique identifier. Theories are related by a (nominal) sub-theory
+ relation, which corresponds to the dependency graph of the original
+ construction; each theory is derived from a certain sub-graph of
+ ancestor theories. To this end, the system maintains a set of
+ symbolic ``identification stamps'' within each theory.
+
+ The @{text "merge"} operation produces the least upper bound of two
+ theories, which actually degenerates into absorption of one theory
+ into the other (according to the nominal sub-theory relation).
+
+ The @{text "begin"} operation starts a new theory by importing
+ several parent theories and entering a special mode of nameless
+ incremental updates, until the final @{text "end"} operation is
+ performed.
+
+ \medskip The example in \figref{fig:ex-theory} below shows a theory
+ graph derived from @{text "Pure"}, with theory @{text "Length"}
+ importing @{text "Nat"} and @{text "List"}. The body of @{text
+ "Length"} consists of a sequence of updates, resulting in locally a
+ linear sub-theory relation for each intermediate step.
+
+ \begin{figure}[htb]
+ \begin{center}
+ \begin{tabular}{rcccl}
+ & & @{text "Pure"} \\
+ & & @{text "\<down>"} \\
+ & & @{text "FOL"} \\
+ & $\swarrow$ & & $\searrow$ & \\
+ @{text "Nat"} & & & & @{text "List"} \\
+ & $\searrow$ & & $\swarrow$ \\
+ & & @{text "Length"} \\
+ & & \multicolumn{3}{l}{~~@{keyword "begin"}} \\
+ & & $\vdots$~~ \\
+ & & \multicolumn{3}{l}{~~@{command "end"}} \\
+ \end{tabular}
+ \caption{A theory definition depending on ancestors}\label{fig:ex-theory}
+ \end{center}
+ \end{figure}
+
+ \medskip Derived formal entities may retain a reference to the
+ background theory in order to indicate the formal context from which
+ they were produced. This provides an immutable certificate of the
+ background theory. *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type theory} \\
+ @{index_ML Theory.eq_thy: "theory * theory -> bool"} \\
+ @{index_ML Theory.subthy: "theory * theory -> bool"} \\
+ @{index_ML Theory.merge: "theory * theory -> theory"} \\
+ @{index_ML Theory.begin_theory: "string * Position.T -> theory list -> theory"} \\
+ @{index_ML Theory.parents_of: "theory -> theory list"} \\
+ @{index_ML Theory.ancestors_of: "theory -> theory list"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type theory} represents theory contexts.
+
+ \item @{ML "Theory.eq_thy"}~@{text "(thy\<^sub>1, thy\<^sub>2)"} check strict
+ identity of two theories.
+
+ \item @{ML "Theory.subthy"}~@{text "(thy\<^sub>1, thy\<^sub>2)"} compares theories
+ according to the intrinsic graph structure of the construction.
+ This sub-theory relation is a nominal approximation of inclusion
+ (@{text "\<subseteq>"}) of the corresponding content (according to the
+ semantics of the ML modules that implement the data).
+
+ \item @{ML "Theory.merge"}~@{text "(thy\<^sub>1, thy\<^sub>2)"} absorbs one theory
+ into the other. This version of ad-hoc theory merge fails for
+ unrelated theories!
+
+ \item @{ML "Theory.begin_theory"}~@{text "name parents"} constructs
+ a new theory based on the given parents. This ML function is
+ normally not invoked directly.
+
+ \item @{ML "Theory.parents_of"}~@{text "thy"} returns the direct
+ ancestors of @{text thy}.
+
+ \item @{ML "Theory.ancestors_of"}~@{text "thy"} returns all
+ ancestors of @{text thy} (not including @{text thy} itself).
+
+ \end{description}
+*}
+
+text %mlantiq {*
+ \begin{matharray}{rcl}
+ @{ML_antiquotation_def "theory"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "theory_context"} & : & @{text ML_antiquotation} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{ML_antiquotation theory} nameref?
+ ;
+ @@{ML_antiquotation theory_context} nameref
+ \<close>}
+
+ \begin{description}
+
+ \item @{text "@{theory}"} refers to the background theory of the
+ current context --- as abstract value.
+
+ \item @{text "@{theory A}"} refers to an explicitly named ancestor
+ theory @{text "A"} of the background theory of the current context
+ --- as abstract value.
+
+ \item @{text "@{theory_context A}"} is similar to @{text "@{theory
+ A}"}, but presents the result as initial @{ML_type Proof.context}
+ (see also @{ML Proof_Context.init_global}).
+
+ \end{description}
+*}
+
+
+subsection {* Proof context \label{sec:context-proof} *}
+
+text {* A proof context is a container for pure data that refers to
+ the theory from which it is derived. The @{text "init"} operation
+ creates a proof context from a given theory. There is an explicit
+ @{text "transfer"} operation to force resynchronization with updates
+ to the background theory -- this is rarely required in practice.
+
+ Entities derived in a proof context need to record logical
+ requirements explicitly, since there is no separate context
+ identification or symbolic inclusion as for theories. For example,
+ hypotheses used in primitive derivations (cf.\ \secref{sec:thms})
+ are recorded separately within the sequent @{text "\<Gamma> \<turnstile> \<phi>"}, just to
+ make double sure. Results could still leak into an alien proof
+ context due to programming errors, but Isabelle/Isar includes some
+ extra validity checks in critical positions, notably at the end of a
+ sub-proof.
+
+ Proof contexts may be manipulated arbitrarily, although the common
+ discipline is to follow block structure as a mental model: a given
+ context is extended consecutively, and results are exported back
+ into the original context. Note that an Isar proof state models
+ block-structured reasoning explicitly, using a stack of proof
+ contexts internally. For various technical reasons, the background
+ theory of an Isar proof state must not be changed while the proof is
+ still under construction!
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type Proof.context} \\
+ @{index_ML Proof_Context.init_global: "theory -> Proof.context"} \\
+ @{index_ML Proof_Context.theory_of: "Proof.context -> theory"} \\
+ @{index_ML Proof_Context.transfer: "theory -> Proof.context -> Proof.context"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type Proof.context} represents proof contexts.
+
+ \item @{ML Proof_Context.init_global}~@{text "thy"} produces a proof
+ context derived from @{text "thy"}, initializing all data.
+
+ \item @{ML Proof_Context.theory_of}~@{text "ctxt"} selects the
+ background theory from @{text "ctxt"}.
+
+ \item @{ML Proof_Context.transfer}~@{text "thy ctxt"} promotes the
+ background theory of @{text "ctxt"} to the super theory @{text
+ "thy"}.
+
+ \end{description}
+*}
+
+text %mlantiq {*
+ \begin{matharray}{rcl}
+ @{ML_antiquotation_def "context"} & : & @{text ML_antiquotation} \\
+ \end{matharray}
+
+ \begin{description}
+
+ \item @{text "@{context}"} refers to \emph{the} context at
+ compile-time --- as abstract value. Independently of (local) theory
+ or proof mode, this always produces a meaningful result.
+
+ This is probably the most common antiquotation in interactive
+ experimentation with ML inside Isar.
+
+ \end{description}
+*}
+
+
+subsection {* Generic contexts \label{sec:generic-context} *}
+
+text {*
+ A generic context is the disjoint sum of either a theory or proof
+ context. Occasionally, this enables uniform treatment of generic
+ context data, typically extra-logical information. Operations on
+ generic contexts include the usual injections, partial selections,
+ and combinators for lifting operations on either component of the
+ disjoint sum.
+
+ Moreover, there are total operations @{text "theory_of"} and @{text
+ "proof_of"} to convert a generic context into either kind: a theory
+ can always be selected from the sum, while a proof context might
+ have to be constructed by an ad-hoc @{text "init"} operation, which
+ incurs a small runtime overhead.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type Context.generic} \\
+ @{index_ML Context.theory_of: "Context.generic -> theory"} \\
+ @{index_ML Context.proof_of: "Context.generic -> Proof.context"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type Context.generic} is the direct sum of @{ML_type
+ "theory"} and @{ML_type "Proof.context"}, with the datatype
+ constructors @{ML "Context.Theory"} and @{ML "Context.Proof"}.
+
+ \item @{ML Context.theory_of}~@{text "context"} always produces a
+ theory from the generic @{text "context"}, using @{ML
+ "Proof_Context.theory_of"} as required.
+
+ \item @{ML Context.proof_of}~@{text "context"} always produces a
+ proof context from the generic @{text "context"}, using @{ML
+ "Proof_Context.init_global"} as required (note that this re-initializes the
+ context data with each invocation).
+
+ \end{description}
+*}
+
+
+subsection {* Context data \label{sec:context-data} *}
+
+text {* The main purpose of theory and proof contexts is to manage
+ arbitrary (pure) data. New data types can be declared incrementally
+ at compile time. There are separate declaration mechanisms for any
+ of the three kinds of contexts: theory, proof, generic.
+
+ \paragraph{Theory data} declarations need to implement the following
+ SML signature:
+
+ \medskip
+ \begin{tabular}{ll}
+ @{text "\<type> T"} & representing type \\
+ @{text "\<val> empty: T"} & empty default value \\
+ @{text "\<val> extend: T \<rightarrow> T"} & re-initialize on import \\
+ @{text "\<val> merge: T \<times> T \<rightarrow> T"} & join on import \\
+ \end{tabular}
+ \medskip
+
+ The @{text "empty"} value acts as initial default for \emph{any}
+ theory that does not declare actual data content; @{text "extend"}
+ is acts like a unitary version of @{text "merge"}.
+
+ Implementing @{text "merge"} can be tricky. The general idea is
+ that @{text "merge (data\<^sub>1, data\<^sub>2)"} inserts those parts of @{text
+ "data\<^sub>2"} into @{text "data\<^sub>1"} that are not yet present, while
+ keeping the general order of things. The @{ML Library.merge}
+ function on plain lists may serve as canonical template.
+
+ Particularly note that shared parts of the data must not be
+ duplicated by naive concatenation, or a theory graph that is like a
+ chain of diamonds would cause an exponential blowup!
+
+ \paragraph{Proof context data} declarations need to implement the
+ following SML signature:
+
+ \medskip
+ \begin{tabular}{ll}
+ @{text "\<type> T"} & representing type \\
+ @{text "\<val> init: theory \<rightarrow> T"} & produce initial value \\
+ \end{tabular}
+ \medskip
+
+ The @{text "init"} operation is supposed to produce a pure value
+ from the given background theory and should be somehow
+ ``immediate''. Whenever a proof context is initialized, which
+ happens frequently, the the system invokes the @{text "init"}
+ operation of \emph{all} theory data slots ever declared. This also
+ means that one needs to be economic about the total number of proof
+ data declarations in the system, i.e.\ each ML module should declare
+ at most one, sometimes two data slots for its internal use.
+ Repeated data declarations to simulate a record type should be
+ avoided!
+
+ \paragraph{Generic data} provides a hybrid interface for both theory
+ and proof data. The @{text "init"} operation for proof contexts is
+ predefined to select the current data value from the background
+ theory.
+
+ \bigskip Any of the above data declarations over type @{text "T"}
+ result in an ML structure with the following signature:
+
+ \medskip
+ \begin{tabular}{ll}
+ @{text "get: context \<rightarrow> T"} \\
+ @{text "put: T \<rightarrow> context \<rightarrow> context"} \\
+ @{text "map: (T \<rightarrow> T) \<rightarrow> context \<rightarrow> context"} \\
+ \end{tabular}
+ \medskip
+
+ These other operations provide exclusive access for the particular
+ kind of context (theory, proof, or generic context). This interface
+ observes the ML discipline for types and scopes: there is no other
+ way to access the corresponding data slot of a context. By keeping
+ these operations private, an Isabelle/ML module may maintain
+ abstract values authentically. *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_functor Theory_Data} \\
+ @{index_ML_functor Proof_Data} \\
+ @{index_ML_functor Generic_Data} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML_functor Theory_Data}@{text "(spec)"} declares data for
+ type @{ML_type theory} according to the specification provided as
+ argument structure. The resulting structure provides data init and
+ access operations as described above.
+
+ \item @{ML_functor Proof_Data}@{text "(spec)"} is analogous to
+ @{ML_functor Theory_Data} for type @{ML_type Proof.context}.
+
+ \item @{ML_functor Generic_Data}@{text "(spec)"} is analogous to
+ @{ML_functor Theory_Data} for type @{ML_type Context.generic}.
+
+ \end{description}
+*}
+
+text %mlex {*
+ The following artificial example demonstrates theory
+ data: we maintain a set of terms that are supposed to be wellformed
+ wrt.\ the enclosing theory. The public interface is as follows:
+*}
+
+ML {*
+ signature WELLFORMED_TERMS =
+ sig
+ val get: theory -> term list
+ val add: term -> theory -> theory
+ end;
+*}
+
+text {* The implementation uses private theory data internally, and
+ only exposes an operation that involves explicit argument checking
+ wrt.\ the given theory. *}
+
+ML {*
+ structure Wellformed_Terms: WELLFORMED_TERMS =
+ struct
+
+ structure Terms = Theory_Data
+ (
+ type T = term Ord_List.T;
+ val empty = [];
+ val extend = I;
+ fun merge (ts1, ts2) =
+ Ord_List.union Term_Ord.fast_term_ord ts1 ts2;
+ );
+
+ val get = Terms.get;
+
+ fun add raw_t thy =
+ let
+ val t = Sign.cert_term thy raw_t;
+ in
+ Terms.map (Ord_List.insert Term_Ord.fast_term_ord t) thy
+ end;
+
+ end;
+*}
+
+text {* Type @{ML_type "term Ord_List.T"} is used for reasonably
+ efficient representation of a set of terms: all operations are
+ linear in the number of stored elements. Here we assume that users
+ of this module do not care about the declaration order, since that
+ data structure forces its own arrangement of elements.
+
+ Observe how the @{ML_text merge} operation joins the data slots of
+ the two constituents: @{ML Ord_List.union} prevents duplication of
+ common data from different branches, thus avoiding the danger of
+ exponential blowup. Plain list append etc.\ must never be used for
+ theory data merges!
+
+ \medskip Our intended invariant is achieved as follows:
+ \begin{enumerate}
+
+ \item @{ML Wellformed_Terms.add} only admits terms that have passed
+ the @{ML Sign.cert_term} check of the given theory at that point.
+
+ \item Wellformedness in the sense of @{ML Sign.cert_term} is
+ monotonic wrt.\ the sub-theory relation. So our data can move
+ upwards in the hierarchy (via extension or merges), and maintain
+ wellformedness without further checks.
+
+ \end{enumerate}
+
+ Note that all basic operations of the inference kernel (which
+ includes @{ML Sign.cert_term}) observe this monotonicity principle,
+ but other user-space tools don't. For example, fully-featured
+ type-inference via @{ML Syntax.check_term} (cf.\
+ \secref{sec:term-check}) is not necessarily monotonic wrt.\ the
+ background theory, since constraints of term constants can be
+ modified by later declarations, for example.
+
+ In most cases, user-space context data does not have to take such
+ invariants too seriously. The situation is different in the
+ implementation of the inference kernel itself, which uses the very
+ same data mechanisms for types, constants, axioms etc.
+*}
+
+
+subsection {* Configuration options \label{sec:config-options} *}
+
+text {* A \emph{configuration option} is a named optional value of
+ some basic type (Boolean, integer, string) that is stored in the
+ context. It is a simple application of general context data
+ (\secref{sec:context-data}) that is sufficiently common to justify
+ customized setup, which includes some concrete declarations for
+ end-users using existing notation for attributes (cf.\
+ \secref{sec:attributes}).
+
+ For example, the predefined configuration option @{attribute
+ show_types} controls output of explicit type constraints for
+ variables in printed terms (cf.\ \secref{sec:read-print}). Its
+ value can be modified within Isar text like this:
+*}
+
+declare [[show_types = false]]
+ -- {* declaration within (local) theory context *}
+
+notepad
+begin
+ note [[show_types = true]]
+ -- {* declaration within proof (forward mode) *}
+ term x
+
+ have "x = x"
+ using [[show_types = false]]
+ -- {* declaration within proof (backward mode) *}
+ ..
+end
+
+text {* Configuration options that are not set explicitly hold a
+ default value that can depend on the application context. This
+ allows to retrieve the value from another slot within the context,
+ or fall back on a global preference mechanism, for example.
+
+ The operations to declare configuration options and get/map their
+ values are modeled as direct replacements for historic global
+ references, only that the context is made explicit. This allows
+ easy configuration of tools, without relying on the execution order
+ as required for old-style mutable references. *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML Config.get: "Proof.context -> 'a Config.T -> 'a"} \\
+ @{index_ML Config.map: "'a Config.T -> ('a -> 'a) -> Proof.context -> Proof.context"} \\
+ @{index_ML Attrib.setup_config_bool: "binding -> (Context.generic -> bool) ->
+ bool Config.T"} \\
+ @{index_ML Attrib.setup_config_int: "binding -> (Context.generic -> int) ->
+ int Config.T"} \\
+ @{index_ML Attrib.setup_config_real: "binding -> (Context.generic -> real) ->
+ real Config.T"} \\
+ @{index_ML Attrib.setup_config_string: "binding -> (Context.generic -> string) ->
+ string Config.T"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML Config.get}~@{text "ctxt config"} gets the value of
+ @{text "config"} in the given context.
+
+ \item @{ML Config.map}~@{text "config f ctxt"} updates the context
+ by updating the value of @{text "config"}.
+
+ \item @{text "config ="}~@{ML Attrib.setup_config_bool}~@{text "name
+ default"} creates a named configuration option of type @{ML_type
+ bool}, with the given @{text "default"} depending on the application
+ context. The resulting @{text "config"} can be used to get/map its
+ value in a given context. There is an implicit update of the
+ background theory that registers the option as attribute with some
+ concrete syntax.
+
+ \item @{ML Attrib.config_int}, @{ML Attrib.config_real}, and @{ML
+ Attrib.config_string} work like @{ML Attrib.config_bool}, but for
+ types @{ML_type int} and @{ML_type string}, respectively.
+
+ \end{description}
+*}
+
+text %mlex {* The following example shows how to declare and use a
+ Boolean configuration option called @{text "my_flag"} with constant
+ default value @{ML false}. *}
+
+ML {*
+ val my_flag =
+ Attrib.setup_config_bool @{binding my_flag} (K false)
+*}
+
+text {* Now the user can refer to @{attribute my_flag} in
+ declarations, while ML tools can retrieve the current value from the
+ context via @{ML Config.get}. *}
+
+ML_val {* @{assert} (Config.get @{context} my_flag = false) *}
+
+declare [[my_flag = true]]
+
+ML_val {* @{assert} (Config.get @{context} my_flag = true) *}
+
+notepad
+begin
+ {
+ note [[my_flag = false]]
+ ML_val {* @{assert} (Config.get @{context} my_flag = false) *}
+ }
+ ML_val {* @{assert} (Config.get @{context} my_flag = true) *}
+end
+
+text {* Here is another example involving ML type @{ML_type real}
+ (floating-point numbers). *}
+
+ML {*
+ val airspeed_velocity =
+ Attrib.setup_config_real @{binding airspeed_velocity} (K 0.0)
+*}
+
+declare [[airspeed_velocity = 10]]
+declare [[airspeed_velocity = 9.9]]
+
+
+section {* Names \label{sec:names} *}
+
+text {* In principle, a name is just a string, but there are various
+ conventions for representing additional structure. For example,
+ ``@{text "Foo.bar.baz"}'' is considered as a long name consisting of
+ qualifier @{text "Foo.bar"} and base name @{text "baz"}. The
+ individual constituents of a name may have further substructure,
+ e.g.\ the string ``\verb,\,\verb,<alpha>,'' encodes as a single
+ symbol (\secref{sec:symbols}).
+
+ \medskip Subsequently, we shall introduce specific categories of
+ names. Roughly speaking these correspond to logical entities as
+ follows:
+ \begin{itemize}
+
+ \item Basic names (\secref{sec:basic-name}): free and bound
+ variables.
+
+ \item Indexed names (\secref{sec:indexname}): schematic variables.
+
+ \item Long names (\secref{sec:long-name}): constants of any kind
+ (type constructors, term constants, other concepts defined in user
+ space). Such entities are typically managed via name spaces
+ (\secref{sec:name-space}).
+
+ \end{itemize}
+*}
+
+
+subsection {* Basic names \label{sec:basic-name} *}
+
+text {*
+ A \emph{basic name} essentially consists of a single Isabelle
+ identifier. There are conventions to mark separate classes of basic
+ names, by attaching a suffix of underscores: one underscore means
+ \emph{internal name}, two underscores means \emph{Skolem name},
+ three underscores means \emph{internal Skolem name}.
+
+ For example, the basic name @{text "foo"} has the internal version
+ @{text "foo_"}, with Skolem versions @{text "foo__"} and @{text
+ "foo___"}, respectively.
+
+ These special versions provide copies of the basic name space, apart
+ from anything that normally appears in the user text. For example,
+ system generated variables in Isar proof contexts are usually marked
+ as internal, which prevents mysterious names like @{text "xaa"} to
+ appear in human-readable text.
+
+ \medskip Manipulating binding scopes often requires on-the-fly
+ renamings. A \emph{name context} contains a collection of already
+ used names. The @{text "declare"} operation adds names to the
+ context.
+
+ The @{text "invents"} operation derives a number of fresh names from
+ a given starting point. For example, the first three names derived
+ from @{text "a"} are @{text "a"}, @{text "b"}, @{text "c"}.
+
+ The @{text "variants"} operation produces fresh names by
+ incrementing tentative names as base-26 numbers (with digits @{text
+ "a..z"}) until all clashes are resolved. For example, name @{text
+ "foo"} results in variants @{text "fooa"}, @{text "foob"}, @{text
+ "fooc"}, \dots, @{text "fooaa"}, @{text "fooab"} etc.; each renaming
+ step picks the next unused variant from this sequence.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML Name.internal: "string -> string"} \\
+ @{index_ML Name.skolem: "string -> string"} \\
+ \end{mldecls}
+ \begin{mldecls}
+ @{index_ML_type Name.context} \\
+ @{index_ML Name.context: Name.context} \\
+ @{index_ML Name.declare: "string -> Name.context -> Name.context"} \\
+ @{index_ML Name.invent: "Name.context -> string -> int -> string list"} \\
+ @{index_ML Name.variant: "string -> Name.context -> string * Name.context"} \\
+ \end{mldecls}
+ \begin{mldecls}
+ @{index_ML Variable.names_of: "Proof.context -> Name.context"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML Name.internal}~@{text "name"} produces an internal name
+ by adding one underscore.
+
+ \item @{ML Name.skolem}~@{text "name"} produces a Skolem name by
+ adding two underscores.
+
+ \item Type @{ML_type Name.context} represents the context of already
+ used names; the initial value is @{ML "Name.context"}.
+
+ \item @{ML Name.declare}~@{text "name"} enters a used name into the
+ context.
+
+ \item @{ML Name.invent}~@{text "context name n"} produces @{text
+ "n"} fresh names derived from @{text "name"}.
+
+ \item @{ML Name.variant}~@{text "name context"} produces a fresh
+ variant of @{text "name"}; the result is declared to the context.
+
+ \item @{ML Variable.names_of}~@{text "ctxt"} retrieves the context
+ of declared type and term variable names. Projecting a proof
+ context down to a primitive name context is occasionally useful when
+ invoking lower-level operations. Regular management of ``fresh
+ variables'' is done by suitable operations of structure @{ML_structure
+ Variable}, which is also able to provide an official status of
+ ``locally fixed variable'' within the logical environment (cf.\
+ \secref{sec:variables}).
+
+ \end{description}
+*}
+
+text %mlex {* The following simple examples demonstrate how to produce
+ fresh names from the initial @{ML Name.context}. *}
+
+ML {*
+ val list1 = Name.invent Name.context "a" 5;
+ @{assert} (list1 = ["a", "b", "c", "d", "e"]);
+
+ val list2 =
+ #1 (fold_map Name.variant ["x", "x", "a", "a", "'a", "'a"] Name.context);
+ @{assert} (list2 = ["x", "xa", "a", "aa", "'a", "'aa"]);
+*}
+
+text {* \medskip The same works relatively to the formal context as
+ follows. *}
+
+locale ex = fixes a b c :: 'a
+begin
+
+ML {*
+ val names = Variable.names_of @{context};
+
+ val list1 = Name.invent names "a" 5;
+ @{assert} (list1 = ["d", "e", "f", "g", "h"]);
+
+ val list2 =
+ #1 (fold_map Name.variant ["x", "x", "a", "a", "'a", "'a"] names);
+ @{assert} (list2 = ["x", "xa", "aa", "ab", "'aa", "'ab"]);
+*}
+
+end
+
+
+subsection {* Indexed names \label{sec:indexname} *}
+
+text {*
+ An \emph{indexed name} (or @{text "indexname"}) is a pair of a basic
+ name and a natural number. This representation allows efficient
+ renaming by incrementing the second component only. The canonical
+ way to rename two collections of indexnames apart from each other is
+ this: determine the maximum index @{text "maxidx"} of the first
+ collection, then increment all indexes of the second collection by
+ @{text "maxidx + 1"}; the maximum index of an empty collection is
+ @{text "-1"}.
+
+ Occasionally, basic names are injected into the same pair type of
+ indexed names: then @{text "(x, -1)"} is used to encode the basic
+ name @{text "x"}.
+
+ \medskip Isabelle syntax observes the following rules for
+ representing an indexname @{text "(x, i)"} as a packed string:
+
+ \begin{itemize}
+
+ \item @{text "?x"} if @{text "x"} does not end with a digit and @{text "i = 0"},
+
+ \item @{text "?xi"} if @{text "x"} does not end with a digit,
+
+ \item @{text "?x.i"} otherwise.
+
+ \end{itemize}
+
+ Indexnames may acquire large index numbers after several maxidx
+ shifts have been applied. Results are usually normalized towards
+ @{text "0"} at certain checkpoints, notably at the end of a proof.
+ This works by producing variants of the corresponding basic name
+ components. For example, the collection @{text "?x1, ?x7, ?x42"}
+ becomes @{text "?x, ?xa, ?xb"}.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type indexname: "string * int"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type indexname} represents indexed names. This is
+ an abbreviation for @{ML_type "string * int"}. The second component
+ is usually non-negative, except for situations where @{text "(x,
+ -1)"} is used to inject basic names into this type. Other negative
+ indexes should not be used.
+
+ \end{description}
+*}
+
+
+subsection {* Long names \label{sec:long-name} *}
+
+text {* A \emph{long name} consists of a sequence of non-empty name
+ components. The packed representation uses a dot as separator, as
+ in ``@{text "A.b.c"}''. The last component is called \emph{base
+ name}, the remaining prefix is called \emph{qualifier} (which may be
+ empty). The qualifier can be understood as the access path to the
+ named entity while passing through some nested block-structure,
+ although our free-form long names do not really enforce any strict
+ discipline.
+
+ For example, an item named ``@{text "A.b.c"}'' may be understood as
+ a local entity @{text "c"}, within a local structure @{text "b"},
+ within a global structure @{text "A"}. In practice, long names
+ usually represent 1--3 levels of qualification. User ML code should
+ not make any assumptions about the particular structure of long
+ names!
+
+ The empty name is commonly used as an indication of unnamed
+ entities, or entities that are not entered into the corresponding
+ name space, whenever this makes any sense. The basic operations on
+ long names map empty names again to empty names.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML Long_Name.base_name: "string -> string"} \\
+ @{index_ML Long_Name.qualifier: "string -> string"} \\
+ @{index_ML Long_Name.append: "string -> string -> string"} \\
+ @{index_ML Long_Name.implode: "string list -> string"} \\
+ @{index_ML Long_Name.explode: "string -> string list"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML Long_Name.base_name}~@{text "name"} returns the base name
+ of a long name.
+
+ \item @{ML Long_Name.qualifier}~@{text "name"} returns the qualifier
+ of a long name.
+
+ \item @{ML Long_Name.append}~@{text "name\<^sub>1 name\<^sub>2"} appends two long
+ names.
+
+ \item @{ML Long_Name.implode}~@{text "names"} and @{ML
+ Long_Name.explode}~@{text "name"} convert between the packed string
+ representation and the explicit list form of long names.
+
+ \end{description}
+*}
+
+
+subsection {* Name spaces \label{sec:name-space} *}
+
+text {* A @{text "name space"} manages a collection of long names,
+ together with a mapping between partially qualified external names
+ and fully qualified internal names (in both directions). Note that
+ the corresponding @{text "intern"} and @{text "extern"} operations
+ are mostly used for parsing and printing only! The @{text
+ "declare"} operation augments a name space according to the accesses
+ determined by a given binding, and a naming policy from the context.
+
+ \medskip A @{text "binding"} specifies details about the prospective
+ long name of a newly introduced formal entity. It consists of a
+ base name, prefixes for qualification (separate ones for system
+ infrastructure and user-space mechanisms), a slot for the original
+ source position, and some additional flags.
+
+ \medskip A @{text "naming"} provides some additional details for
+ producing a long name from a binding. Normally, the naming is
+ implicit in the theory or proof context. The @{text "full"}
+ operation (and its variants for different context types) produces a
+ fully qualified internal name to be entered into a name space. The
+ main equation of this ``chemical reaction'' when binding new
+ entities in a context is as follows:
+
+ \medskip
+ \begin{tabular}{l}
+ @{text "binding + naming \<longrightarrow> long name + name space accesses"}
+ \end{tabular}
+
+ \bigskip As a general principle, there is a separate name space for
+ each kind of formal entity, e.g.\ fact, logical constant, type
+ constructor, type class. It is usually clear from the occurrence in
+ concrete syntax (or from the scope) which kind of entity a name
+ refers to. For example, the very same name @{text "c"} may be used
+ uniformly for a constant, type constructor, and type class.
+
+ There are common schemes to name derived entities systematically
+ according to the name of the main logical entity involved, e.g.\
+ fact @{text "c.intro"} for a canonical introduction rule related to
+ constant @{text "c"}. This technique of mapping names from one
+ space into another requires some care in order to avoid conflicts.
+ In particular, theorem names derived from a type constructor or type
+ class should get an additional suffix in addition to the usual
+ qualification. This leads to the following conventions for derived
+ names:
+
+ \medskip
+ \begin{tabular}{ll}
+ logical entity & fact name \\\hline
+ constant @{text "c"} & @{text "c.intro"} \\
+ type @{text "c"} & @{text "c_type.intro"} \\
+ class @{text "c"} & @{text "c_class.intro"} \\
+ \end{tabular}
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type binding} \\
+ @{index_ML Binding.empty: binding} \\
+ @{index_ML Binding.name: "string -> binding"} \\
+ @{index_ML Binding.qualify: "bool -> string -> binding -> binding"} \\
+ @{index_ML Binding.prefix: "bool -> string -> binding -> binding"} \\
+ @{index_ML Binding.conceal: "binding -> binding"} \\
+ @{index_ML Binding.print: "binding -> string"} \\
+ \end{mldecls}
+ \begin{mldecls}
+ @{index_ML_type Name_Space.naming} \\
+ @{index_ML Name_Space.default_naming: Name_Space.naming} \\
+ @{index_ML Name_Space.add_path: "string -> Name_Space.naming -> Name_Space.naming"} \\
+ @{index_ML Name_Space.full_name: "Name_Space.naming -> binding -> string"} \\
+ \end{mldecls}
+ \begin{mldecls}
+ @{index_ML_type Name_Space.T} \\
+ @{index_ML Name_Space.empty: "string -> Name_Space.T"} \\
+ @{index_ML Name_Space.merge: "Name_Space.T * Name_Space.T -> Name_Space.T"} \\
+ @{index_ML Name_Space.declare: "Context.generic -> bool ->
+ binding -> Name_Space.T -> string * Name_Space.T"} \\
+ @{index_ML Name_Space.intern: "Name_Space.T -> string -> string"} \\
+ @{index_ML Name_Space.extern: "Proof.context -> Name_Space.T -> string -> string"} \\
+ @{index_ML Name_Space.is_concealed: "Name_Space.T -> string -> bool"}
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type binding} represents the abstract concept of
+ name bindings.
+
+ \item @{ML Binding.empty} is the empty binding.
+
+ \item @{ML Binding.name}~@{text "name"} produces a binding with base
+ name @{text "name"}. Note that this lacks proper source position
+ information; see also the ML antiquotation @{ML_antiquotation
+ binding}.
+
+ \item @{ML Binding.qualify}~@{text "mandatory name binding"}
+ prefixes qualifier @{text "name"} to @{text "binding"}. The @{text
+ "mandatory"} flag tells if this name component always needs to be
+ given in name space accesses --- this is mostly @{text "false"} in
+ practice. Note that this part of qualification is typically used in
+ derived specification mechanisms.
+
+ \item @{ML Binding.prefix} is similar to @{ML Binding.qualify}, but
+ affects the system prefix. This part of extra qualification is
+ typically used in the infrastructure for modular specifications,
+ notably ``local theory targets'' (see also \chref{ch:local-theory}).
+
+ \item @{ML Binding.conceal}~@{text "binding"} indicates that the
+ binding shall refer to an entity that serves foundational purposes
+ only. This flag helps to mark implementation details of
+ specification mechanism etc. Other tools should not depend on the
+ particulars of concealed entities (cf.\ @{ML
+ Name_Space.is_concealed}).
+
+ \item @{ML Binding.print}~@{text "binding"} produces a string
+ representation for human-readable output, together with some formal
+ markup that might get used in GUI front-ends, for example.
+
+ \item Type @{ML_type Name_Space.naming} represents the abstract
+ concept of a naming policy.
+
+ \item @{ML Name_Space.default_naming} is the default naming policy.
+ In a theory context, this is usually augmented by a path prefix
+ consisting of the theory name.
+
+ \item @{ML Name_Space.add_path}~@{text "path naming"} augments the
+ naming policy by extending its path component.
+
+ \item @{ML Name_Space.full_name}~@{text "naming binding"} turns a
+ name binding (usually a basic name) into the fully qualified
+ internal name, according to the given naming policy.
+
+ \item Type @{ML_type Name_Space.T} represents name spaces.
+
+ \item @{ML Name_Space.empty}~@{text "kind"} and @{ML Name_Space.merge}~@{text
+ "(space\<^sub>1, space\<^sub>2)"} are the canonical operations for
+ maintaining name spaces according to theory data management
+ (\secref{sec:context-data}); @{text "kind"} is a formal comment
+ to characterize the purpose of a name space.
+
+ \item @{ML Name_Space.declare}~@{text "context strict binding
+ space"} enters a name binding as fully qualified internal name into
+ the name space, using the naming of the context.
+
+ \item @{ML Name_Space.intern}~@{text "space name"} internalizes a
+ (partially qualified) external name.
+
+ This operation is mostly for parsing! Note that fully qualified
+ names stemming from declarations are produced via @{ML
+ "Name_Space.full_name"} and @{ML "Name_Space.declare"}
+ (or their derivatives for @{ML_type theory} and
+ @{ML_type Proof.context}).
+
+ \item @{ML Name_Space.extern}~@{text "ctxt space name"} externalizes a
+ (fully qualified) internal name.
+
+ This operation is mostly for printing! User code should not rely on
+ the precise result too much.
+
+ \item @{ML Name_Space.is_concealed}~@{text "space name"} indicates
+ whether @{text "name"} refers to a strictly private entity that
+ other tools are supposed to ignore!
+
+ \end{description}
+*}
+
+text %mlantiq {*
+ \begin{matharray}{rcl}
+ @{ML_antiquotation_def "binding"} & : & @{text ML_antiquotation} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{ML_antiquotation binding} name
+ \<close>}
+
+ \begin{description}
+
+ \item @{text "@{binding name}"} produces a binding with base name
+ @{text "name"} and the source position taken from the concrete
+ syntax of this antiquotation. In many situations this is more
+ appropriate than the more basic @{ML Binding.name} function.
+
+ \end{description}
+*}
+
+text %mlex {* The following example yields the source position of some
+ concrete binding inlined into the text:
+*}
+
+ML {* Binding.pos_of @{binding here} *}
+
+text {* \medskip That position can be also printed in a message as
+ follows: *}
+
+ML_command {*
+ writeln
+ ("Look here" ^ Position.here (Binding.pos_of @{binding here}))
+*}
+
+text {* This illustrates a key virtue of formalized bindings as
+ opposed to raw specifications of base names: the system can use this
+ additional information for feedback given to the user (error
+ messages etc.).
+
+ \medskip The following example refers to its source position
+ directly, which is occasionally useful for experimentation and
+ diagnostic purposes: *}
+
+ML_command {*
+ warning ("Look here" ^ Position.here @{here})
+*}
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Implementation/Proof.thy Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,492 @@
+theory Proof
+imports Base
+begin
+
+chapter {* Structured proofs *}
+
+section {* Variables \label{sec:variables} *}
+
+text {*
+ Any variable that is not explicitly bound by @{text "\<lambda>"}-abstraction
+ is considered as ``free''. Logically, free variables act like
+ outermost universal quantification at the sequent level: @{text
+ "A\<^sub>1(x), \<dots>, A\<^sub>n(x) \<turnstile> B(x)"} means that the result
+ holds \emph{for all} values of @{text "x"}. Free variables for
+ terms (not types) can be fully internalized into the logic: @{text
+ "\<turnstile> B(x)"} and @{text "\<turnstile> \<And>x. B(x)"} are interchangeable, provided
+ that @{text "x"} does not occur elsewhere in the context.
+ Inspecting @{text "\<turnstile> \<And>x. B(x)"} more closely, we see that inside the
+ quantifier, @{text "x"} is essentially ``arbitrary, but fixed'',
+ while from outside it appears as a place-holder for instantiation
+ (thanks to @{text "\<And>"} elimination).
+
+ The Pure logic represents the idea of variables being either inside
+ or outside the current scope by providing separate syntactic
+ categories for \emph{fixed variables} (e.g.\ @{text "x"}) vs.\
+ \emph{schematic variables} (e.g.\ @{text "?x"}). Incidently, a
+ universal result @{text "\<turnstile> \<And>x. B(x)"} has the HHF normal form @{text
+ "\<turnstile> B(?x)"}, which represents its generality without requiring an
+ explicit quantifier. The same principle works for type variables:
+ @{text "\<turnstile> B(?\<alpha>)"} represents the idea of ``@{text "\<turnstile> \<forall>\<alpha>. B(\<alpha>)"}''
+ without demanding a truly polymorphic framework.
+
+ \medskip Additional care is required to treat type variables in a
+ way that facilitates type-inference. In principle, term variables
+ depend on type variables, which means that type variables would have
+ to be declared first. For example, a raw type-theoretic framework
+ would demand the context to be constructed in stages as follows:
+ @{text "\<Gamma> = \<alpha>: type, x: \<alpha>, a: A(x\<^sub>\<alpha>)"}.
+
+ We allow a slightly less formalistic mode of operation: term
+ variables @{text "x"} are fixed without specifying a type yet
+ (essentially \emph{all} potential occurrences of some instance
+ @{text "x\<^sub>\<tau>"} are fixed); the first occurrence of @{text "x"}
+ within a specific term assigns its most general type, which is then
+ maintained consistently in the context. The above example becomes
+ @{text "\<Gamma> = x: term, \<alpha>: type, A(x\<^sub>\<alpha>)"}, where type @{text
+ "\<alpha>"} is fixed \emph{after} term @{text "x"}, and the constraint
+ @{text "x :: \<alpha>"} is an implicit consequence of the occurrence of
+ @{text "x\<^sub>\<alpha>"} in the subsequent proposition.
+
+ This twist of dependencies is also accommodated by the reverse
+ operation of exporting results from a context: a type variable
+ @{text "\<alpha>"} is considered fixed as long as it occurs in some fixed
+ term variable of the context. For example, exporting @{text "x:
+ term, \<alpha>: type \<turnstile> x\<^sub>\<alpha> \<equiv> x\<^sub>\<alpha>"} produces in the first step @{text "x: term
+ \<turnstile> x\<^sub>\<alpha> \<equiv> x\<^sub>\<alpha>"} for fixed @{text "\<alpha>"}, and only in the second step
+ @{text "\<turnstile> ?x\<^sub>?\<^sub>\<alpha> \<equiv> ?x\<^sub>?\<^sub>\<alpha>"} for schematic @{text "?x"} and @{text "?\<alpha>"}.
+ The following Isar source text illustrates this scenario.
+*}
+
+notepad
+begin
+ {
+ fix x -- {* all potential occurrences of some @{text "x::\<tau>"} are fixed *}
+ {
+ have "x::'a \<equiv> x" -- {* implicit type assigment by concrete occurrence *}
+ by (rule reflexive)
+ }
+ thm this -- {* result still with fixed type @{text "'a"} *}
+ }
+ thm this -- {* fully general result for arbitrary @{text "?x::?'a"} *}
+end
+
+text {* The Isabelle/Isar proof context manages the details of term
+ vs.\ type variables, with high-level principles for moving the
+ frontier between fixed and schematic variables.
+
+ The @{text "add_fixes"} operation explictly declares fixed
+ variables; the @{text "declare_term"} operation absorbs a term into
+ a context by fixing new type variables and adding syntactic
+ constraints.
+
+ The @{text "export"} operation is able to perform the main work of
+ generalizing term and type variables as sketched above, assuming
+ that fixing variables and terms have been declared properly.
+
+ There @{text "import"} operation makes a generalized fact a genuine
+ part of the context, by inventing fixed variables for the schematic
+ ones. The effect can be reversed by using @{text "export"} later,
+ potentially with an extended context; the result is equivalent to
+ the original modulo renaming of schematic variables.
+
+ The @{text "focus"} operation provides a variant of @{text "import"}
+ for nested propositions (with explicit quantification): @{text
+ "\<And>x\<^sub>1 \<dots> x\<^sub>n. B(x\<^sub>1, \<dots>, x\<^sub>n)"} is
+ decomposed by inventing fixed variables @{text "x\<^sub>1, \<dots>,
+ x\<^sub>n"} for the body.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML Variable.add_fixes: "
+ string list -> Proof.context -> string list * Proof.context"} \\
+ @{index_ML Variable.variant_fixes: "
+ string list -> Proof.context -> string list * Proof.context"} \\
+ @{index_ML Variable.declare_term: "term -> Proof.context -> Proof.context"} \\
+ @{index_ML Variable.declare_constraints: "term -> Proof.context -> Proof.context"} \\
+ @{index_ML Variable.export: "Proof.context -> Proof.context -> thm list -> thm list"} \\
+ @{index_ML Variable.polymorphic: "Proof.context -> term list -> term list"} \\
+ @{index_ML Variable.import: "bool -> thm list -> Proof.context ->
+ (((ctyp * ctyp) list * (cterm * cterm) list) * thm list) * Proof.context"} \\
+ @{index_ML Variable.focus: "term -> Proof.context ->
+ ((string * (string * typ)) list * term) * Proof.context"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML Variable.add_fixes}~@{text "xs ctxt"} fixes term
+ variables @{text "xs"}, returning the resulting internal names. By
+ default, the internal representation coincides with the external
+ one, which also means that the given variables must not be fixed
+ already. There is a different policy within a local proof body: the
+ given names are just hints for newly invented Skolem variables.
+
+ \item @{ML Variable.variant_fixes} is similar to @{ML
+ Variable.add_fixes}, but always produces fresh variants of the given
+ names.
+
+ \item @{ML Variable.declare_term}~@{text "t ctxt"} declares term
+ @{text "t"} to belong to the context. This automatically fixes new
+ type variables, but not term variables. Syntactic constraints for
+ type and term variables are declared uniformly, though.
+
+ \item @{ML Variable.declare_constraints}~@{text "t ctxt"} declares
+ syntactic constraints from term @{text "t"}, without making it part
+ of the context yet.
+
+ \item @{ML Variable.export}~@{text "inner outer thms"} generalizes
+ fixed type and term variables in @{text "thms"} according to the
+ difference of the @{text "inner"} and @{text "outer"} context,
+ following the principles sketched above.
+
+ \item @{ML Variable.polymorphic}~@{text "ctxt ts"} generalizes type
+ variables in @{text "ts"} as far as possible, even those occurring
+ in fixed term variables. The default policy of type-inference is to
+ fix newly introduced type variables, which is essentially reversed
+ with @{ML Variable.polymorphic}: here the given terms are detached
+ from the context as far as possible.
+
+ \item @{ML Variable.import}~@{text "open thms ctxt"} invents fixed
+ type and term variables for the schematic ones occurring in @{text
+ "thms"}. The @{text "open"} flag indicates whether the fixed names
+ should be accessible to the user, otherwise newly introduced names
+ are marked as ``internal'' (\secref{sec:names}).
+
+ \item @{ML Variable.focus}~@{text B} decomposes the outermost @{text
+ "\<And>"} prefix of proposition @{text "B"}.
+
+ \end{description}
+*}
+
+text %mlex {* The following example shows how to work with fixed term
+ and type parameters and with type-inference. *}
+
+ML {*
+ (*static compile-time context -- for testing only*)
+ val ctxt0 = @{context};
+
+ (*locally fixed parameters -- no type assignment yet*)
+ val ([x, y], ctxt1) = ctxt0 |> Variable.add_fixes ["x", "y"];
+
+ (*t1: most general fixed type; t1': most general arbitrary type*)
+ val t1 = Syntax.read_term ctxt1 "x";
+ val t1' = singleton (Variable.polymorphic ctxt1) t1;
+
+ (*term u enforces specific type assignment*)
+ val u = Syntax.read_term ctxt1 "(x::nat) \<equiv> y";
+
+ (*official declaration of u -- propagates constraints etc.*)
+ val ctxt2 = ctxt1 |> Variable.declare_term u;
+ val t2 = Syntax.read_term ctxt2 "x"; (*x::nat is enforced*)
+*}
+
+text {* In the above example, the starting context is derived from the
+ toplevel theory, which means that fixed variables are internalized
+ literally: @{text "x"} is mapped again to @{text "x"}, and
+ attempting to fix it again in the subsequent context is an error.
+ Alternatively, fixed parameters can be renamed explicitly as
+ follows: *}
+
+ML {*
+ val ctxt0 = @{context};
+ val ([x1, x2, x3], ctxt1) =
+ ctxt0 |> Variable.variant_fixes ["x", "x", "x"];
+*}
+
+text {* The following ML code can now work with the invented names of
+ @{text x1}, @{text x2}, @{text x3}, without depending on
+ the details on the system policy for introducing these variants.
+ Recall that within a proof body the system always invents fresh
+ ``skolem constants'', e.g.\ as follows: *}
+
+notepad
+begin
+ ML_prf %"ML" {*
+ val ctxt0 = @{context};
+
+ val ([x1], ctxt1) = ctxt0 |> Variable.add_fixes ["x"];
+ val ([x2], ctxt2) = ctxt1 |> Variable.add_fixes ["x"];
+ val ([x3], ctxt3) = ctxt2 |> Variable.add_fixes ["x"];
+
+ val ([y1, y2], ctxt4) =
+ ctxt3 |> Variable.variant_fixes ["y", "y"];
+ *}
+end
+
+text {* In this situation @{ML Variable.add_fixes} and @{ML
+ Variable.variant_fixes} are very similar, but identical name
+ proposals given in a row are only accepted by the second version.
+ *}
+
+
+section {* Assumptions \label{sec:assumptions} *}
+
+text {*
+ An \emph{assumption} is a proposition that it is postulated in the
+ current context. Local conclusions may use assumptions as
+ additional facts, but this imposes implicit hypotheses that weaken
+ the overall statement.
+
+ Assumptions are restricted to fixed non-schematic statements, i.e.\
+ all generality needs to be expressed by explicit quantifiers.
+ Nevertheless, the result will be in HHF normal form with outermost
+ quantifiers stripped. For example, by assuming @{text "\<And>x :: \<alpha>. P
+ x"} we get @{text "\<And>x :: \<alpha>. P x \<turnstile> P ?x"} for schematic @{text "?x"}
+ of fixed type @{text "\<alpha>"}. Local derivations accumulate more and
+ more explicit references to hypotheses: @{text "A\<^sub>1, \<dots>,
+ A\<^sub>n \<turnstile> B"} where @{text "A\<^sub>1, \<dots>, A\<^sub>n"} needs to
+ be covered by the assumptions of the current context.
+
+ \medskip The @{text "add_assms"} operation augments the context by
+ local assumptions, which are parameterized by an arbitrary @{text
+ "export"} rule (see below).
+
+ The @{text "export"} operation moves facts from a (larger) inner
+ context into a (smaller) outer context, by discharging the
+ difference of the assumptions as specified by the associated export
+ rules. Note that the discharged portion is determined by the
+ difference of contexts, not the facts being exported! There is a
+ separate flag to indicate a goal context, where the result is meant
+ to refine an enclosing sub-goal of a structured proof state.
+
+ \medskip The most basic export rule discharges assumptions directly
+ by means of the @{text "\<Longrightarrow>"} introduction rule:
+ \[
+ \infer[(@{text "\<Longrightarrow>\<hyphen>intro"})]{@{text "\<Gamma> - A \<turnstile> A \<Longrightarrow> B"}}{@{text "\<Gamma> \<turnstile> B"}}
+ \]
+
+ The variant for goal refinements marks the newly introduced
+ premises, which causes the canonical Isar goal refinement scheme to
+ enforce unification with local premises within the goal:
+ \[
+ \infer[(@{text "#\<Longrightarrow>\<hyphen>intro"})]{@{text "\<Gamma> - A \<turnstile> #A \<Longrightarrow> B"}}{@{text "\<Gamma> \<turnstile> B"}}
+ \]
+
+ \medskip Alternative versions of assumptions may perform arbitrary
+ transformations on export, as long as the corresponding portion of
+ hypotheses is removed from the given facts. For example, a local
+ definition works by fixing @{text "x"} and assuming @{text "x \<equiv> t"},
+ with the following export rule to reverse the effect:
+ \[
+ \infer[(@{text "\<equiv>\<hyphen>expand"})]{@{text "\<Gamma> - (x \<equiv> t) \<turnstile> B t"}}{@{text "\<Gamma> \<turnstile> B x"}}
+ \]
+ This works, because the assumption @{text "x \<equiv> t"} was introduced in
+ a context with @{text "x"} being fresh, so @{text "x"} does not
+ occur in @{text "\<Gamma>"} here.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type Assumption.export} \\
+ @{index_ML Assumption.assume: "Proof.context -> cterm -> thm"} \\
+ @{index_ML Assumption.add_assms:
+ "Assumption.export ->
+ cterm list -> Proof.context -> thm list * Proof.context"} \\
+ @{index_ML Assumption.add_assumes: "
+ cterm list -> Proof.context -> thm list * Proof.context"} \\
+ @{index_ML Assumption.export: "bool -> Proof.context -> Proof.context -> thm -> thm"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type Assumption.export} represents arbitrary export
+ rules, which is any function of type @{ML_type "bool -> cterm list
+ -> thm -> thm"}, where the @{ML_type "bool"} indicates goal mode,
+ and the @{ML_type "cterm list"} the collection of assumptions to be
+ discharged simultaneously.
+
+ \item @{ML Assumption.assume}~@{text "ctxt A"} turns proposition @{text
+ "A"} into a primitive assumption @{text "A \<turnstile> A'"}, where the
+ conclusion @{text "A'"} is in HHF normal form.
+
+ \item @{ML Assumption.add_assms}~@{text "r As"} augments the context
+ by assumptions @{text "As"} with export rule @{text "r"}. The
+ resulting facts are hypothetical theorems as produced by the raw
+ @{ML Assumption.assume}.
+
+ \item @{ML Assumption.add_assumes}~@{text "As"} is a special case of
+ @{ML Assumption.add_assms} where the export rule performs @{text
+ "\<Longrightarrow>\<hyphen>intro"} or @{text "#\<Longrightarrow>\<hyphen>intro"}, depending on goal
+ mode.
+
+ \item @{ML Assumption.export}~@{text "is_goal inner outer thm"}
+ exports result @{text "thm"} from the the @{text "inner"} context
+ back into the @{text "outer"} one; @{text "is_goal = true"} means
+ this is a goal context. The result is in HHF normal form. Note
+ that @{ML "Proof_Context.export"} combines @{ML "Variable.export"}
+ and @{ML "Assumption.export"} in the canonical way.
+
+ \end{description}
+*}
+
+text %mlex {* The following example demonstrates how rules can be
+ derived by building up a context of assumptions first, and exporting
+ some local fact afterwards. We refer to @{theory Pure} equality
+ here for testing purposes.
+*}
+
+ML {*
+ (*static compile-time context -- for testing only*)
+ val ctxt0 = @{context};
+
+ val ([eq], ctxt1) =
+ ctxt0 |> Assumption.add_assumes [@{cprop "x \<equiv> y"}];
+ val eq' = Thm.symmetric eq;
+
+ (*back to original context -- discharges assumption*)
+ val r = Assumption.export false ctxt1 ctxt0 eq';
+*}
+
+text {* Note that the variables of the resulting rule are not
+ generalized. This would have required to fix them properly in the
+ context beforehand, and export wrt.\ variables afterwards (cf.\ @{ML
+ Variable.export} or the combined @{ML "Proof_Context.export"}). *}
+
+
+section {* Structured goals and results \label{sec:struct-goals} *}
+
+text {*
+ Local results are established by monotonic reasoning from facts
+ within a context. This allows common combinations of theorems,
+ e.g.\ via @{text "\<And>/\<Longrightarrow>"} elimination, resolution rules, or equational
+ reasoning, see \secref{sec:thms}. Unaccounted context manipulations
+ should be avoided, notably raw @{text "\<And>/\<Longrightarrow>"} introduction or ad-hoc
+ references to free variables or assumptions not present in the proof
+ context.
+
+ \medskip The @{text "SUBPROOF"} combinator allows to structure a
+ tactical proof recursively by decomposing a selected sub-goal:
+ @{text "(\<And>x. A(x) \<Longrightarrow> B(x)) \<Longrightarrow> \<dots>"} is turned into @{text "B(x) \<Longrightarrow> \<dots>"}
+ after fixing @{text "x"} and assuming @{text "A(x)"}. This means
+ the tactic needs to solve the conclusion, but may use the premise as
+ a local fact, for locally fixed variables.
+
+ The family of @{text "FOCUS"} combinators is similar to @{text
+ "SUBPROOF"}, but allows to retain schematic variables and pending
+ subgoals in the resulting goal state.
+
+ The @{text "prove"} operation provides an interface for structured
+ backwards reasoning under program control, with some explicit sanity
+ checks of the result. The goal context can be augmented by
+ additional fixed variables (cf.\ \secref{sec:variables}) and
+ assumptions (cf.\ \secref{sec:assumptions}), which will be available
+ as local facts during the proof and discharged into implications in
+ the result. Type and term variables are generalized as usual,
+ according to the context.
+
+ The @{text "obtain"} operation produces results by eliminating
+ existing facts by means of a given tactic. This acts like a dual
+ conclusion: the proof demonstrates that the context may be augmented
+ by parameters and assumptions, without affecting any conclusions
+ that do not mention these parameters. See also
+ \cite{isabelle-isar-ref} for the user-level @{command obtain} and
+ @{command guess} elements. Final results, which may not refer to
+ the parameters in the conclusion, need to exported explicitly into
+ the original context. *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML SUBPROOF: "(Subgoal.focus -> tactic) ->
+ Proof.context -> int -> tactic"} \\
+ @{index_ML Subgoal.FOCUS: "(Subgoal.focus -> tactic) ->
+ Proof.context -> int -> tactic"} \\
+ @{index_ML Subgoal.FOCUS_PREMS: "(Subgoal.focus -> tactic) ->
+ Proof.context -> int -> tactic"} \\
+ @{index_ML Subgoal.FOCUS_PARAMS: "(Subgoal.focus -> tactic) ->
+ Proof.context -> int -> tactic"} \\
+ @{index_ML Subgoal.focus: "Proof.context -> int -> thm -> Subgoal.focus * thm"} \\
+ @{index_ML Subgoal.focus_prems: "Proof.context -> int -> thm -> Subgoal.focus * thm"} \\
+ @{index_ML Subgoal.focus_params: "Proof.context -> int -> thm -> Subgoal.focus * thm"} \\
+ \end{mldecls}
+
+ \begin{mldecls}
+ @{index_ML Goal.prove: "Proof.context -> string list -> term list -> term ->
+ ({prems: thm list, context: Proof.context} -> tactic) -> thm"} \\
+ @{index_ML Goal.prove_multi: "Proof.context -> string list -> term list -> term list ->
+ ({prems: thm list, context: Proof.context} -> tactic) -> thm list"} \\
+ \end{mldecls}
+ \begin{mldecls}
+ @{index_ML Obtain.result: "(Proof.context -> tactic) -> thm list ->
+ Proof.context -> ((string * cterm) list * thm list) * Proof.context"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML SUBPROOF}~@{text "tac ctxt i"} decomposes the structure
+ of the specified sub-goal, producing an extended context and a
+ reduced goal, which needs to be solved by the given tactic. All
+ schematic parameters of the goal are imported into the context as
+ fixed ones, which may not be instantiated in the sub-proof.
+
+ \item @{ML Subgoal.FOCUS}, @{ML Subgoal.FOCUS_PREMS}, and @{ML
+ Subgoal.FOCUS_PARAMS} are similar to @{ML SUBPROOF}, but are
+ slightly more flexible: only the specified parts of the subgoal are
+ imported into the context, and the body tactic may introduce new
+ subgoals and schematic variables.
+
+ \item @{ML Subgoal.focus}, @{ML Subgoal.focus_prems}, @{ML
+ Subgoal.focus_params} extract the focus information from a goal
+ state in the same way as the corresponding tacticals above. This is
+ occasionally useful to experiment without writing actual tactics
+ yet.
+
+ \item @{ML Goal.prove}~@{text "ctxt xs As C tac"} states goal @{text
+ "C"} in the context augmented by fixed variables @{text "xs"} and
+ assumptions @{text "As"}, and applies tactic @{text "tac"} to solve
+ it. The latter may depend on the local assumptions being presented
+ as facts. The result is in HHF normal form.
+
+ \item @{ML Goal.prove_multi} is simular to @{ML Goal.prove}, but
+ states several conclusions simultaneously. The goal is encoded by
+ means of Pure conjunction; @{ML Goal.conjunction_tac} will turn this
+ into a collection of individual subgoals.
+
+ \item @{ML Obtain.result}~@{text "tac thms ctxt"} eliminates the
+ given facts using a tactic, which results in additional fixed
+ variables and assumptions in the context. Final results need to be
+ exported explicitly.
+
+ \end{description}
+*}
+
+text %mlex {* The following minimal example illustrates how to access
+ the focus information of a structured goal state. *}
+
+notepad
+begin
+ fix A B C :: "'a \<Rightarrow> bool"
+
+ have "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x"
+ ML_val
+ {*
+ val {goal, context = goal_ctxt, ...} = @{Isar.goal};
+ val (focus as {params, asms, concl, ...}, goal') =
+ Subgoal.focus goal_ctxt 1 goal;
+ val [A, B] = #prems focus;
+ val [(_, x)] = #params focus;
+ *}
+ oops
+
+text {* \medskip The next example demonstrates forward-elimination in
+ a local context, using @{ML Obtain.result}. *}
+
+notepad
+begin
+ assume ex: "\<exists>x. B x"
+
+ ML_prf %"ML" {*
+ val ctxt0 = @{context};
+ val (([(_, x)], [B]), ctxt1) = ctxt0
+ |> Obtain.result (fn _ => etac @{thm exE} 1) [@{thm ex}];
+ *}
+ ML_prf %"ML" {*
+ singleton (Proof_Context.export ctxt1 ctxt0) @{thm refl};
+ *}
+ ML_prf %"ML" {*
+ Proof_Context.export ctxt1 ctxt0 [Thm.reflexive x]
+ handle ERROR msg => (warning msg; []);
+ *}
+end
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Implementation/Syntax.thy Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,151 @@
+theory Syntax
+imports Base
+begin
+
+chapter {* Concrete syntax and type-checking *}
+
+text {* Pure @{text "\<lambda>"}-calculus as introduced in \chref{ch:logic} is
+ an adequate foundation for logical languages --- in the tradition of
+ \emph{higher-order abstract syntax} --- but end-users require
+ additional means for reading and printing of terms and types. This
+ important add-on outside the logical core is called \emph{inner
+ syntax} in Isabelle jargon, as opposed to the \emph{outer syntax} of
+ the theory and proof language (cf.\ \cite{isabelle-isar-ref}).
+
+ For example, according to \cite{church40} quantifiers are
+ represented as higher-order constants @{text "All :: ('a \<Rightarrow> bool) \<Rightarrow>
+ bool"} such that @{text "All (\<lambda>x::'a. B x)"} faithfully represents
+ the idea that is displayed as @{text "\<forall>x::'a. B x"} via @{keyword
+ "binder"} notation. Moreover, type-inference in the style of
+ Hindley-Milner \cite{hindleymilner} (and extensions) enables users
+ to write @{text "\<forall>x. B x"} concisely, when the type @{text "'a"} is
+ already clear from the context.\footnote{Type-inference taken to the
+ extreme can easily confuse users, though. Beginners often stumble
+ over unexpectedly general types inferred by the system.}
+
+ \medskip The main inner syntax operations are \emph{read} for
+ parsing together with type-checking, and \emph{pretty} for formatted
+ output. See also \secref{sec:read-print}.
+
+ Furthermore, the input and output syntax layers are sub-divided into
+ separate phases for \emph{concrete syntax} versus \emph{abstract
+ syntax}, see also \secref{sec:parse-unparse} and
+ \secref{sec:term-check}, respectively. This results in the
+ following decomposition of the main operations:
+
+ \begin{itemize}
+
+ \item @{text "read = parse; check"}
+
+ \item @{text "pretty = uncheck; unparse"}
+
+ \end{itemize}
+
+ Some specification package might thus intercept syntax processing at
+ a well-defined stage after @{text "parse"}, to a augment the
+ resulting pre-term before full type-reconstruction is performed by
+ @{text "check"}, for example. Note that the formal status of bound
+ variables, versus free variables, versus constants must not be
+ changed here! *}
+
+
+section {* Reading and pretty printing \label{sec:read-print} *}
+
+text {* Read and print operations are roughly dual to each other, such
+ that for the user @{text "s' = pretty (read s)"} looks similar to
+ the original source text @{text "s"}, but the details depend on many
+ side-conditions. There are also explicit options to control
+ suppressing of type information in the output. The default
+ configuration routinely looses information, so @{text "t' = read
+ (pretty t)"} might fail, produce a differently typed term, or a
+ completely different term in the face of syntactic overloading! *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML Syntax.read_typ: "Proof.context -> string -> typ"} \\
+ @{index_ML Syntax.read_term: "Proof.context -> string -> term"} \\
+ @{index_ML Syntax.read_prop: "Proof.context -> string -> term"} \\
+ @{index_ML Syntax.pretty_typ: "Proof.context -> typ -> Pretty.T"} \\
+ @{index_ML Syntax.pretty_term: "Proof.context -> term -> Pretty.T"} \\
+ \end{mldecls}
+
+ %FIXME description
+*}
+
+
+section {* Parsing and unparsing \label{sec:parse-unparse} *}
+
+text {* Parsing and unparsing converts between actual source text and
+ a certain \emph{pre-term} format, where all bindings and scopes are
+ resolved faithfully. Thus the names of free variables or constants
+ are already determined in the sense of the logical context, but type
+ information might is still missing. Pre-terms support an explicit
+ language of \emph{type constraints} that may be augmented by user
+ code to guide the later \emph{check} phase, for example.
+
+ Actual parsing is based on traditional lexical analysis and Earley
+ parsing for arbitrary context-free grammars. The user can specify
+ this via mixfix annotations. Moreover, there are \emph{syntax
+ translations} that can be augmented by the user, either
+ declaratively via @{command translations} or programmatically via
+ @{command parse_translation}, @{command print_translation} etc. The
+ final scope resolution is performed by the system, according to name
+ spaces for types, constants etc.\ determined by the context.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML Syntax.parse_typ: "Proof.context -> string -> typ"} \\
+ @{index_ML Syntax.parse_term: "Proof.context -> string -> term"} \\
+ @{index_ML Syntax.parse_prop: "Proof.context -> string -> term"} \\
+ @{index_ML Syntax.unparse_typ: "Proof.context -> typ -> Pretty.T"} \\
+ @{index_ML Syntax.unparse_term: "Proof.context -> term -> Pretty.T"} \\
+ \end{mldecls}
+
+ %FIXME description
+*}
+
+
+section {* Checking and unchecking \label{sec:term-check} *}
+
+text {* These operations define the transition from pre-terms and
+ fully-annotated terms in the sense of the logical core
+ (\chref{ch:logic}).
+
+ The \emph{check} phase is meant to subsume a variety of mechanisms
+ in the manner of ``type-inference'' or ``type-reconstruction'' or
+ ``type-improvement'', not just type-checking in the narrow sense.
+ The \emph{uncheck} phase is roughly dual, it prunes type-information
+ before pretty printing.
+
+ A typical add-on for the check/uncheck syntax layer is the @{command
+ abbreviation} mechanism. Here the user specifies syntactic
+ definitions that are managed by the system as polymorphic @{text
+ "let"} bindings. These are expanded during the @{text "check"}
+ phase, and contracted during the @{text "uncheck"} phase, without
+ affecting the type-assignment of the given terms.
+
+ \medskip The precise meaning of type checking depends on the context
+ --- additional check/uncheck plugins might be defined in user space!
+
+ For example, the @{command class} command defines a context where
+ @{text "check"} treats certain type instances of overloaded
+ constants according to the ``dictionary construction'' of its
+ logical foundation. This involves ``type improvement''
+ (specialization of slightly too general types) and replacement by
+ certain locale parameters. See also \cite{Haftmann-Wenzel:2009}.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML Syntax.check_typs: "Proof.context -> typ list -> typ list"} \\
+ @{index_ML Syntax.check_terms: "Proof.context -> term list -> term list"} \\
+ @{index_ML Syntax.check_props: "Proof.context -> term list -> term list"} \\
+ @{index_ML Syntax.uncheck_typs: "Proof.context -> typ list -> typ list"} \\
+ @{index_ML Syntax.uncheck_terms: "Proof.context -> term list -> term list"} \\
+ \end{mldecls}
+
+ %FIXME description
+*}
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Implementation/Tactic.thy Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,939 @@
+theory Tactic
+imports Base
+begin
+
+chapter {* Tactical reasoning *}
+
+text {* Tactical reasoning works by refining an initial claim in a
+ backwards fashion, until a solved form is reached. A @{text "goal"}
+ consists of several subgoals that need to be solved in order to
+ achieve the main statement; zero subgoals means that the proof may
+ be finished. A @{text "tactic"} is a refinement operation that maps
+ a goal to a lazy sequence of potential successors. A @{text
+ "tactical"} is a combinator for composing tactics. *}
+
+
+section {* Goals \label{sec:tactical-goals} *}
+
+text {*
+ Isabelle/Pure represents a goal as a theorem stating that the
+ subgoals imply the main goal: @{text "A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow>
+ C"}. The outermost goal structure is that of a Horn Clause: i.e.\
+ an iterated implication without any quantifiers\footnote{Recall that
+ outermost @{text "\<And>x. \<phi>[x]"} is always represented via schematic
+ variables in the body: @{text "\<phi>[?x]"}. These variables may get
+ instantiated during the course of reasoning.}. For @{text "n = 0"}
+ a goal is called ``solved''.
+
+ The structure of each subgoal @{text "A\<^sub>i"} is that of a
+ general Hereditary Harrop Formula @{text "\<And>x\<^sub>1 \<dots>
+ \<And>x\<^sub>k. H\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> H\<^sub>m \<Longrightarrow> B"}. Here @{text
+ "x\<^sub>1, \<dots>, x\<^sub>k"} are goal parameters, i.e.\
+ arbitrary-but-fixed entities of certain types, and @{text
+ "H\<^sub>1, \<dots>, H\<^sub>m"} are goal hypotheses, i.e.\ facts that may
+ be assumed locally. Together, this forms the goal context of the
+ conclusion @{text B} to be established. The goal hypotheses may be
+ again arbitrary Hereditary Harrop Formulas, although the level of
+ nesting rarely exceeds 1--2 in practice.
+
+ The main conclusion @{text C} is internally marked as a protected
+ proposition, which is represented explicitly by the notation @{text
+ "#C"} here. This ensures that the decomposition into subgoals and
+ main conclusion is well-defined for arbitrarily structured claims.
+
+ \medskip Basic goal management is performed via the following
+ Isabelle/Pure rules:
+
+ \[
+ \infer[@{text "(init)"}]{@{text "C \<Longrightarrow> #C"}}{} \qquad
+ \infer[@{text "(finish)"}]{@{text "C"}}{@{text "#C"}}
+ \]
+
+ \medskip The following low-level variants admit general reasoning
+ with protected propositions:
+
+ \[
+ \infer[@{text "(protect n)"}]{@{text "A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow> #C"}}{@{text "A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow> C"}}
+ \]
+ \[
+ \infer[@{text "(conclude)"}]{@{text "A \<Longrightarrow> \<dots> \<Longrightarrow> C"}}{@{text "A \<Longrightarrow> \<dots> \<Longrightarrow> #C"}}
+ \]
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML Goal.init: "cterm -> thm"} \\
+ @{index_ML Goal.finish: "Proof.context -> thm -> thm"} \\
+ @{index_ML Goal.protect: "int -> thm -> thm"} \\
+ @{index_ML Goal.conclude: "thm -> thm"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML "Goal.init"}~@{text C} initializes a tactical goal from
+ the well-formed proposition @{text C}.
+
+ \item @{ML "Goal.finish"}~@{text "ctxt thm"} checks whether theorem
+ @{text "thm"} is a solved goal (no subgoals), and concludes the
+ result by removing the goal protection. The context is only
+ required for printing error messages.
+
+ \item @{ML "Goal.protect"}~@{text "n thm"} protects the statement
+ of theorem @{text "thm"}. The parameter @{text n} indicates the
+ number of premises to be retained.
+
+ \item @{ML "Goal.conclude"}~@{text "thm"} removes the goal
+ protection, even if there are pending subgoals.
+
+ \end{description}
+*}
+
+
+section {* Tactics\label{sec:tactics} *}
+
+text {* A @{text "tactic"} is a function @{text "goal \<rightarrow> goal\<^sup>*\<^sup>*"} that
+ maps a given goal state (represented as a theorem, cf.\
+ \secref{sec:tactical-goals}) to a lazy sequence of potential
+ successor states. The underlying sequence implementation is lazy
+ both in head and tail, and is purely functional in \emph{not}
+ supporting memoing.\footnote{The lack of memoing and the strict
+ nature of SML requires some care when working with low-level
+ sequence operations, to avoid duplicate or premature evaluation of
+ results. It also means that modified runtime behavior, such as
+ timeout, is very hard to achieve for general tactics.}
+
+ An \emph{empty result sequence} means that the tactic has failed: in
+ a compound tactic expression other tactics might be tried instead,
+ or the whole refinement step might fail outright, producing a
+ toplevel error message in the end. When implementing tactics from
+ scratch, one should take care to observe the basic protocol of
+ mapping regular error conditions to an empty result; only serious
+ faults should emerge as exceptions.
+
+ By enumerating \emph{multiple results}, a tactic can easily express
+ the potential outcome of an internal search process. There are also
+ combinators for building proof tools that involve search
+ systematically, see also \secref{sec:tacticals}.
+
+ \medskip As explained before, a goal state essentially consists of a
+ list of subgoals that imply the main goal (conclusion). Tactics may
+ operate on all subgoals or on a particularly specified subgoal, but
+ must not change the main conclusion (apart from instantiating
+ schematic goal variables).
+
+ Tactics with explicit \emph{subgoal addressing} are of the form
+ @{text "int \<rightarrow> tactic"} and may be applied to a particular subgoal
+ (counting from 1). If the subgoal number is out of range, the
+ tactic should fail with an empty result sequence, but must not raise
+ an exception!
+
+ Operating on a particular subgoal means to replace it by an interval
+ of zero or more subgoals in the same place; other subgoals must not
+ be affected, apart from instantiating schematic variables ranging
+ over the whole goal state.
+
+ A common pattern of composing tactics with subgoal addressing is to
+ try the first one, and then the second one only if the subgoal has
+ not been solved yet. Special care is required here to avoid bumping
+ into unrelated subgoals that happen to come after the original
+ subgoal. Assuming that there is only a single initial subgoal is a
+ very common error when implementing tactics!
+
+ Tactics with internal subgoal addressing should expose the subgoal
+ index as @{text "int"} argument in full generality; a hardwired
+ subgoal 1 is not acceptable.
+
+ \medskip The main well-formedness conditions for proper tactics are
+ summarized as follows.
+
+ \begin{itemize}
+
+ \item General tactic failure is indicated by an empty result, only
+ serious faults may produce an exception.
+
+ \item The main conclusion must not be changed, apart from
+ instantiating schematic variables.
+
+ \item A tactic operates either uniformly on all subgoals, or
+ specifically on a selected subgoal (without bumping into unrelated
+ subgoals).
+
+ \item Range errors in subgoal addressing produce an empty result.
+
+ \end{itemize}
+
+ Some of these conditions are checked by higher-level goal
+ infrastructure (\secref{sec:struct-goals}); others are not checked
+ explicitly, and violating them merely results in ill-behaved tactics
+ experienced by the user (e.g.\ tactics that insist in being
+ applicable only to singleton goals, or prevent composition via
+ standard tacticals such as @{ML REPEAT}).
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type tactic: "thm -> thm Seq.seq"} \\
+ @{index_ML no_tac: tactic} \\
+ @{index_ML all_tac: tactic} \\
+ @{index_ML print_tac: "string -> tactic"} \\[1ex]
+ @{index_ML PRIMITIVE: "(thm -> thm) -> tactic"} \\[1ex]
+ @{index_ML SUBGOAL: "(term * int -> tactic) -> int -> tactic"} \\
+ @{index_ML CSUBGOAL: "(cterm * int -> tactic) -> int -> tactic"} \\
+ @{index_ML SELECT_GOAL: "tactic -> int -> tactic"} \\
+ @{index_ML PREFER_GOAL: "tactic -> int -> tactic"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type @{ML_type tactic} represents tactics. The
+ well-formedness conditions described above need to be observed. See
+ also @{file "~~/src/Pure/General/seq.ML"} for the underlying
+ implementation of lazy sequences.
+
+ \item Type @{ML_type "int -> tactic"} represents tactics with
+ explicit subgoal addressing, with well-formedness conditions as
+ described above.
+
+ \item @{ML no_tac} is a tactic that always fails, returning the
+ empty sequence.
+
+ \item @{ML all_tac} is a tactic that always succeeds, returning a
+ singleton sequence with unchanged goal state.
+
+ \item @{ML print_tac}~@{text "message"} is like @{ML all_tac}, but
+ prints a message together with the goal state on the tracing
+ channel.
+
+ \item @{ML PRIMITIVE}~@{text rule} turns a primitive inference rule
+ into a tactic with unique result. Exception @{ML THM} is considered
+ a regular tactic failure and produces an empty result; other
+ exceptions are passed through.
+
+ \item @{ML SUBGOAL}~@{text "(fn (subgoal, i) => tactic)"} is the
+ most basic form to produce a tactic with subgoal addressing. The
+ given abstraction over the subgoal term and subgoal number allows to
+ peek at the relevant information of the full goal state. The
+ subgoal range is checked as required above.
+
+ \item @{ML CSUBGOAL} is similar to @{ML SUBGOAL}, but passes the
+ subgoal as @{ML_type cterm} instead of raw @{ML_type term}. This
+ avoids expensive re-certification in situations where the subgoal is
+ used directly for primitive inferences.
+
+ \item @{ML SELECT_GOAL}~@{text "tac i"} confines a tactic to the
+ specified subgoal @{text "i"}. This rearranges subgoals and the
+ main goal protection (\secref{sec:tactical-goals}), while retaining
+ the syntactic context of the overall goal state (concerning
+ schematic variables etc.).
+
+ \item @{ML PREFER_GOAL}~@{text "tac i"} rearranges subgoals to put
+ @{text "i"} in front. This is similar to @{ML SELECT_GOAL}, but
+ without changing the main goal protection.
+
+ \end{description}
+*}
+
+
+subsection {* Resolution and assumption tactics \label{sec:resolve-assume-tac} *}
+
+text {* \emph{Resolution} is the most basic mechanism for refining a
+ subgoal using a theorem as object-level rule.
+ \emph{Elim-resolution} is particularly suited for elimination rules:
+ it resolves with a rule, proves its first premise by assumption, and
+ finally deletes that assumption from any new subgoals.
+ \emph{Destruct-resolution} is like elim-resolution, but the given
+ destruction rules are first turned into canonical elimination
+ format. \emph{Forward-resolution} is like destruct-resolution, but
+ without deleting the selected assumption. The @{text "r/e/d/f"}
+ naming convention is maintained for several different kinds of
+ resolution rules and tactics.
+
+ Assumption tactics close a subgoal by unifying some of its premises
+ against its conclusion.
+
+ \medskip All the tactics in this section operate on a subgoal
+ designated by a positive integer. Other subgoals might be affected
+ indirectly, due to instantiation of schematic variables.
+
+ There are various sources of non-determinism, the tactic result
+ sequence enumerates all possibilities of the following choices (if
+ applicable):
+
+ \begin{enumerate}
+
+ \item selecting one of the rules given as argument to the tactic;
+
+ \item selecting a subgoal premise to eliminate, unifying it against
+ the first premise of the rule;
+
+ \item unifying the conclusion of the subgoal to the conclusion of
+ the rule.
+
+ \end{enumerate}
+
+ Recall that higher-order unification may produce multiple results
+ that are enumerated here.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML resolve_tac: "thm list -> int -> tactic"} \\
+ @{index_ML eresolve_tac: "thm list -> int -> tactic"} \\
+ @{index_ML dresolve_tac: "thm list -> int -> tactic"} \\
+ @{index_ML forward_tac: "thm list -> int -> tactic"} \\
+ @{index_ML biresolve_tac: "(bool * thm) list -> int -> tactic"} \\[1ex]
+ @{index_ML assume_tac: "int -> tactic"} \\
+ @{index_ML eq_assume_tac: "int -> tactic"} \\[1ex]
+ @{index_ML match_tac: "thm list -> int -> tactic"} \\
+ @{index_ML ematch_tac: "thm list -> int -> tactic"} \\
+ @{index_ML dmatch_tac: "thm list -> int -> tactic"} \\
+ @{index_ML bimatch_tac: "(bool * thm) list -> int -> tactic"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML resolve_tac}~@{text "thms i"} refines the goal state
+ using the given theorems, which should normally be introduction
+ rules. The tactic resolves a rule's conclusion with subgoal @{text
+ i}, replacing it by the corresponding versions of the rule's
+ premises.
+
+ \item @{ML eresolve_tac}~@{text "thms i"} performs elim-resolution
+ with the given theorems, which are normally be elimination rules.
+
+ Note that @{ML "eresolve_tac [asm_rl]"} is equivalent to @{ML
+ assume_tac}, which facilitates mixing of assumption steps with
+ genuine eliminations.
+
+ \item @{ML dresolve_tac}~@{text "thms i"} performs
+ destruct-resolution with the given theorems, which should normally
+ be destruction rules. This replaces an assumption by the result of
+ applying one of the rules.
+
+ \item @{ML forward_tac} is like @{ML dresolve_tac} except that the
+ selected assumption is not deleted. It applies a rule to an
+ assumption, adding the result as a new assumption.
+
+ \item @{ML biresolve_tac}~@{text "brls i"} refines the proof state
+ by resolution or elim-resolution on each rule, as indicated by its
+ flag. It affects subgoal @{text "i"} of the proof state.
+
+ For each pair @{text "(flag, rule)"}, it applies resolution if the
+ flag is @{text "false"} and elim-resolution if the flag is @{text
+ "true"}. A single tactic call handles a mixture of introduction and
+ elimination rules, which is useful to organize the search process
+ systematically in proof tools.
+
+ \item @{ML assume_tac}~@{text i} attempts to solve subgoal @{text i}
+ by assumption (modulo higher-order unification).
+
+ \item @{ML eq_assume_tac} is similar to @{ML assume_tac}, but checks
+ only for immediate @{text "\<alpha>"}-convertibility instead of using
+ unification. It succeeds (with a unique next state) if one of the
+ assumptions is equal to the subgoal's conclusion. Since it does not
+ instantiate variables, it cannot make other subgoals unprovable.
+
+ \item @{ML match_tac}, @{ML ematch_tac}, @{ML dmatch_tac}, and @{ML
+ bimatch_tac} are similar to @{ML resolve_tac}, @{ML eresolve_tac},
+ @{ML dresolve_tac}, and @{ML biresolve_tac}, respectively, but do
+ not instantiate schematic variables in the goal state.%
+\footnote{Strictly speaking, matching means to treat the unknowns in the goal
+ state as constants, but these tactics merely discard unifiers that would
+ update the goal state. In rare situations (where the conclusion and
+ goal state have flexible terms at the same position), the tactic
+ will fail even though an acceptable unifier exists.}
+ These tactics were written for a specific application within the classical reasoner.
+
+ Flexible subgoals are not updated at will, but are left alone.
+ \end{description}
+*}
+
+
+subsection {* Explicit instantiation within a subgoal context *}
+
+text {* The main resolution tactics (\secref{sec:resolve-assume-tac})
+ use higher-order unification, which works well in many practical
+ situations despite its daunting theoretical properties.
+ Nonetheless, there are important problem classes where unguided
+ higher-order unification is not so useful. This typically involves
+ rules like universal elimination, existential introduction, or
+ equational substitution. Here the unification problem involves
+ fully flexible @{text "?P ?x"} schemes, which are hard to manage
+ without further hints.
+
+ By providing a (small) rigid term for @{text "?x"} explicitly, the
+ remaining unification problem is to assign a (large) term to @{text
+ "?P"}, according to the shape of the given subgoal. This is
+ sufficiently well-behaved in most practical situations.
+
+ \medskip Isabelle provides separate versions of the standard @{text
+ "r/e/d/f"} resolution tactics that allow to provide explicit
+ instantiations of unknowns of the given rule, wrt.\ terms that refer
+ to the implicit context of the selected subgoal.
+
+ An instantiation consists of a list of pairs of the form @{text
+ "(?x, t)"}, where @{text ?x} is a schematic variable occurring in
+ the given rule, and @{text t} is a term from the current proof
+ context, augmented by the local goal parameters of the selected
+ subgoal; cf.\ the @{text "focus"} operation described in
+ \secref{sec:variables}.
+
+ Entering the syntactic context of a subgoal is a brittle operation,
+ because its exact form is somewhat accidental, and the choice of
+ bound variable names depends on the presence of other local and
+ global names. Explicit renaming of subgoal parameters prior to
+ explicit instantiation might help to achieve a bit more robustness.
+
+ Type instantiations may be given as well, via pairs like @{text
+ "(?'a, \<tau>)"}. Type instantiations are distinguished from term
+ instantiations by the syntactic form of the schematic variable.
+ Types are instantiated before terms are. Since term instantiation
+ already performs simple type-inference, so explicit type
+ instantiations are seldom necessary.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML res_inst_tac: "Proof.context -> (indexname * string) list -> thm -> int -> tactic"} \\
+ @{index_ML eres_inst_tac: "Proof.context -> (indexname * string) list -> thm -> int -> tactic"} \\
+ @{index_ML dres_inst_tac: "Proof.context -> (indexname * string) list -> thm -> int -> tactic"} \\
+ @{index_ML forw_inst_tac: "Proof.context -> (indexname * string) list -> thm -> int -> tactic"} \\
+ @{index_ML subgoal_tac: "Proof.context -> string -> int -> tactic"} \\
+ @{index_ML thin_tac: "Proof.context -> string -> int -> tactic"} \\
+ @{index_ML rename_tac: "string list -> int -> tactic"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML res_inst_tac}~@{text "ctxt insts thm i"} instantiates the
+ rule @{text thm} with the instantiations @{text insts}, as described
+ above, and then performs resolution on subgoal @{text i}.
+
+ \item @{ML eres_inst_tac} is like @{ML res_inst_tac}, but performs
+ elim-resolution.
+
+ \item @{ML dres_inst_tac} is like @{ML res_inst_tac}, but performs
+ destruct-resolution.
+
+ \item @{ML forw_inst_tac} is like @{ML dres_inst_tac} except that
+ the selected assumption is not deleted.
+
+ \item @{ML subgoal_tac}~@{text "ctxt \<phi> i"} adds the proposition
+ @{text "\<phi>"} as local premise to subgoal @{text "i"}, and poses the
+ same as a new subgoal @{text "i + 1"} (in the original context).
+
+ \item @{ML thin_tac}~@{text "ctxt \<phi> i"} deletes the specified
+ premise from subgoal @{text i}. Note that @{text \<phi>} may contain
+ schematic variables, to abbreviate the intended proposition; the
+ first matching subgoal premise will be deleted. Removing useless
+ premises from a subgoal increases its readability and can make
+ search tactics run faster.
+
+ \item @{ML rename_tac}~@{text "names i"} renames the innermost
+ parameters of subgoal @{text i} according to the provided @{text
+ names} (which need to be distinct indentifiers).
+
+ \end{description}
+
+ For historical reasons, the above instantiation tactics take
+ unparsed string arguments, which makes them hard to use in general
+ ML code. The slightly more advanced @{ML Subgoal.FOCUS} combinator
+ of \secref{sec:struct-goals} allows to refer to internal goal
+ structure with explicit context management.
+*}
+
+
+subsection {* Rearranging goal states *}
+
+text {* In rare situations there is a need to rearrange goal states:
+ either the overall collection of subgoals, or the local structure of
+ a subgoal. Various administrative tactics allow to operate on the
+ concrete presentation these conceptual sets of formulae. *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML rotate_tac: "int -> int -> tactic"} \\
+ @{index_ML distinct_subgoals_tac: tactic} \\
+ @{index_ML flexflex_tac: tactic} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML rotate_tac}~@{text "n i"} rotates the premises of subgoal
+ @{text i} by @{text n} positions: from right to left if @{text n} is
+ positive, and from left to right if @{text n} is negative.
+
+ \item @{ML distinct_subgoals_tac} removes duplicate subgoals from a
+ proof state. This is potentially inefficient.
+
+ \item @{ML flexflex_tac} removes all flex-flex pairs from the proof
+ state by applying the trivial unifier. This drastic step loses
+ information. It is already part of the Isar infrastructure for
+ facts resulting from goals, and rarely needs to be invoked manually.
+
+ Flex-flex constraints arise from difficult cases of higher-order
+ unification. To prevent this, use @{ML res_inst_tac} to instantiate
+ some variables in a rule. Normally flex-flex constraints can be
+ ignored; they often disappear as unknowns get instantiated.
+
+ \end{description}
+*}
+
+
+subsection {* Raw composition: resolution without lifting *}
+
+text {*
+ Raw composition of two rules means resolving them without prior
+ lifting or renaming of unknowns. This low-level operation, which
+ underlies the resolution tactics, may occasionally be useful for
+ special effects. Schematic variables are not renamed by default, so
+ beware of clashes!
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML compose_tac: "(bool * thm * int) -> int -> tactic"} \\
+ @{index_ML Drule.compose: "thm * int * thm -> thm"} \\
+ @{index_ML_op COMP: "thm * thm -> thm"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML compose_tac}~@{text "(flag, rule, m) i"} refines subgoal
+ @{text "i"} using @{text "rule"}, without lifting. The @{text
+ "rule"} is taken to have the form @{text "\<psi>\<^sub>1 \<Longrightarrow> \<dots> \<psi>\<^sub>m \<Longrightarrow> \<psi>"}, where
+ @{text "\<psi>"} need not be atomic; thus @{text "m"} determines the
+ number of new subgoals. If @{text "flag"} is @{text "true"} then it
+ performs elim-resolution --- it solves the first premise of @{text
+ "rule"} by assumption and deletes that assumption.
+
+ \item @{ML Drule.compose}~@{text "(thm\<^sub>1, i, thm\<^sub>2)"} uses @{text "thm\<^sub>1"},
+ regarded as an atomic formula, to solve premise @{text "i"} of
+ @{text "thm\<^sub>2"}. Let @{text "thm\<^sub>1"} and @{text "thm\<^sub>2"} be @{text
+ "\<psi>"} and @{text "\<phi>\<^sub>1 \<Longrightarrow> \<dots> \<phi>\<^sub>n \<Longrightarrow> \<phi>"}. The unique @{text "s"} that
+ unifies @{text "\<psi>"} and @{text "\<phi>\<^sub>i"} yields the theorem @{text "(\<phi>\<^sub>1 \<Longrightarrow>
+ \<dots> \<phi>\<^sub>i\<^sub>-\<^sub>1 \<Longrightarrow> \<phi>\<^sub>i\<^sub>+\<^sub>1 \<Longrightarrow> \<dots> \<phi>\<^sub>n \<Longrightarrow> \<phi>)s"}. Multiple results are considered as
+ error (exception @{ML THM}).
+
+ \item @{text "thm\<^sub>1 COMP thm\<^sub>2"} is the same as @{text "Drule.compose
+ (thm\<^sub>1, 1, thm\<^sub>2)"}.
+
+ \end{description}
+
+ \begin{warn}
+ These low-level operations are stepping outside the structure
+ imposed by regular rule resolution. Used without understanding of
+ the consequences, they may produce results that cause problems with
+ standard rules and tactics later on.
+ \end{warn}
+*}
+
+
+section {* Tacticals \label{sec:tacticals} *}
+
+text {* A \emph{tactical} is a functional combinator for building up
+ complex tactics from simpler ones. Common tacticals perform
+ sequential composition, disjunctive choice, iteration, or goal
+ addressing. Various search strategies may be expressed via
+ tacticals.
+*}
+
+
+subsection {* Combining tactics *}
+
+text {* Sequential composition and alternative choices are the most
+ basic ways to combine tactics, similarly to ``@{verbatim ","}'' and
+ ``@{verbatim "|"}'' in Isar method notation. This corresponds to
+ @{ML_op "THEN"} and @{ML_op "ORELSE"} in ML, but there are further
+ possibilities for fine-tuning alternation of tactics such as @{ML_op
+ "APPEND"}. Further details become visible in ML due to explicit
+ subgoal addressing.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_op "THEN": "tactic * tactic -> tactic"} \\
+ @{index_ML_op "ORELSE": "tactic * tactic -> tactic"} \\
+ @{index_ML_op "APPEND": "tactic * tactic -> tactic"} \\
+ @{index_ML "EVERY": "tactic list -> tactic"} \\
+ @{index_ML "FIRST": "tactic list -> tactic"} \\[0.5ex]
+
+ @{index_ML_op "THEN'": "('a -> tactic) * ('a -> tactic) -> 'a -> tactic"} \\
+ @{index_ML_op "ORELSE'": "('a -> tactic) * ('a -> tactic) -> 'a -> tactic"} \\
+ @{index_ML_op "APPEND'": "('a -> tactic) * ('a -> tactic) -> 'a -> tactic"} \\
+ @{index_ML "EVERY'": "('a -> tactic) list -> 'a -> tactic"} \\
+ @{index_ML "FIRST'": "('a -> tactic) list -> 'a -> tactic"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{text "tac\<^sub>1"}~@{ML_op THEN}~@{text "tac\<^sub>2"} is the sequential
+ composition of @{text "tac\<^sub>1"} and @{text "tac\<^sub>2"}. Applied to a goal
+ state, it returns all states reachable in two steps by applying
+ @{text "tac\<^sub>1"} followed by @{text "tac\<^sub>2"}. First, it applies @{text
+ "tac\<^sub>1"} to the goal state, getting a sequence of possible next
+ states; then, it applies @{text "tac\<^sub>2"} to each of these and
+ concatenates the results to produce again one flat sequence of
+ states.
+
+ \item @{text "tac\<^sub>1"}~@{ML_op ORELSE}~@{text "tac\<^sub>2"} makes a choice
+ between @{text "tac\<^sub>1"} and @{text "tac\<^sub>2"}. Applied to a state, it
+ tries @{text "tac\<^sub>1"} and returns the result if non-empty; if @{text
+ "tac\<^sub>1"} fails then it uses @{text "tac\<^sub>2"}. This is a deterministic
+ choice: if @{text "tac\<^sub>1"} succeeds then @{text "tac\<^sub>2"} is excluded
+ from the result.
+
+ \item @{text "tac\<^sub>1"}~@{ML_op APPEND}~@{text "tac\<^sub>2"} concatenates the
+ possible results of @{text "tac\<^sub>1"} and @{text "tac\<^sub>2"}. Unlike
+ @{ML_op "ORELSE"} there is \emph{no commitment} to either tactic, so
+ @{ML_op "APPEND"} helps to avoid incompleteness during search, at
+ the cost of potential inefficiencies.
+
+ \item @{ML EVERY}~@{text "[tac\<^sub>1, \<dots>, tac\<^sub>n]"} abbreviates @{text
+ "tac\<^sub>1"}~@{ML_op THEN}~@{text "\<dots>"}~@{ML_op THEN}~@{text "tac\<^sub>n"}.
+ Note that @{ML "EVERY []"} is the same as @{ML all_tac}: it always
+ succeeds.
+
+ \item @{ML FIRST}~@{text "[tac\<^sub>1, \<dots>, tac\<^sub>n]"} abbreviates @{text
+ "tac\<^sub>1"}~@{ML_op ORELSE}~@{text "\<dots>"}~@{ML_op "ORELSE"}~@{text
+ "tac\<^sub>n"}. Note that @{ML "FIRST []"} is the same as @{ML no_tac}: it
+ always fails.
+
+ \item @{ML_op "THEN'"} is the lifted version of @{ML_op "THEN"}, for
+ tactics with explicit subgoal addressing. So @{text
+ "(tac\<^sub>1"}~@{ML_op THEN'}~@{text "tac\<^sub>2) i"} is the same as @{text
+ "(tac\<^sub>1 i"}~@{ML_op THEN}~@{text "tac\<^sub>2 i)"}.
+
+ The other primed tacticals work analogously.
+
+ \end{description}
+*}
+
+
+subsection {* Repetition tacticals *}
+
+text {* These tacticals provide further control over repetition of
+ tactics, beyond the stylized forms of ``@{verbatim "?"}'' and
+ ``@{verbatim "+"}'' in Isar method expressions. *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML "TRY": "tactic -> tactic"} \\
+ @{index_ML "REPEAT": "tactic -> tactic"} \\
+ @{index_ML "REPEAT1": "tactic -> tactic"} \\
+ @{index_ML "REPEAT_DETERM": "tactic -> tactic"} \\
+ @{index_ML "REPEAT_DETERM_N": "int -> tactic -> tactic"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML TRY}~@{text "tac"} applies @{text "tac"} to the goal
+ state and returns the resulting sequence, if non-empty; otherwise it
+ returns the original state. Thus, it applies @{text "tac"} at most
+ once.
+
+ Note that for tactics with subgoal addressing, the combinator can be
+ applied via functional composition: @{ML "TRY"}~@{ML_op o}~@{text
+ "tac"}. There is no need for @{verbatim TRY'}.
+
+ \item @{ML REPEAT}~@{text "tac"} applies @{text "tac"} to the goal
+ state and, recursively, to each element of the resulting sequence.
+ The resulting sequence consists of those states that make @{text
+ "tac"} fail. Thus, it applies @{text "tac"} as many times as
+ possible (including zero times), and allows backtracking over each
+ invocation of @{text "tac"}. @{ML REPEAT} is more general than @{ML
+ REPEAT_DETERM}, but requires more space.
+
+ \item @{ML REPEAT1}~@{text "tac"} is like @{ML REPEAT}~@{text "tac"}
+ but it always applies @{text "tac"} at least once, failing if this
+ is impossible.
+
+ \item @{ML REPEAT_DETERM}~@{text "tac"} applies @{text "tac"} to the
+ goal state and, recursively, to the head of the resulting sequence.
+ It returns the first state to make @{text "tac"} fail. It is
+ deterministic, discarding alternative outcomes.
+
+ \item @{ML REPEAT_DETERM_N}~@{text "n tac"} is like @{ML
+ REPEAT_DETERM}~@{text "tac"} but the number of repetitions is bound
+ by @{text "n"} (where @{ML "~1"} means @{text "\<infinity>"}).
+
+ \end{description}
+*}
+
+text %mlex {* The basic tactics and tacticals considered above follow
+ some algebraic laws:
+
+ \begin{itemize}
+
+ \item @{ML all_tac} is the identity element of the tactical @{ML_op
+ "THEN"}.
+
+ \item @{ML no_tac} is the identity element of @{ML_op "ORELSE"} and
+ @{ML_op "APPEND"}. Also, it is a zero element for @{ML_op "THEN"},
+ which means that @{text "tac"}~@{ML_op THEN}~@{ML no_tac} is
+ equivalent to @{ML no_tac}.
+
+ \item @{ML TRY} and @{ML REPEAT} can be expressed as (recursive)
+ functions over more basic combinators (ignoring some internal
+ implementation tricks):
+
+ \end{itemize}
+*}
+
+ML {*
+ fun TRY tac = tac ORELSE all_tac;
+ fun REPEAT tac st = ((tac THEN REPEAT tac) ORELSE all_tac) st;
+*}
+
+text {* If @{text "tac"} can return multiple outcomes then so can @{ML
+ REPEAT}~@{text "tac"}. @{ML REPEAT} uses @{ML_op "ORELSE"} and not
+ @{ML_op "APPEND"}, it applies @{text "tac"} as many times as
+ possible in each outcome.
+
+ \begin{warn}
+ Note the explicit abstraction over the goal state in the ML
+ definition of @{ML REPEAT}. Recursive tacticals must be coded in
+ this awkward fashion to avoid infinite recursion of eager functional
+ evaluation in Standard ML. The following attempt would make @{ML
+ REPEAT}~@{text "tac"} loop:
+ \end{warn}
+*}
+
+ML {*
+ (*BAD -- does not terminate!*)
+ fun REPEAT tac = (tac THEN REPEAT tac) ORELSE all_tac;
+*}
+
+
+subsection {* Applying tactics to subgoal ranges *}
+
+text {* Tactics with explicit subgoal addressing
+ @{ML_type "int -> tactic"} can be used together with tacticals that
+ act like ``subgoal quantifiers'': guided by success of the body
+ tactic a certain range of subgoals is covered. Thus the body tactic
+ is applied to \emph{all} subgoals, \emph{some} subgoal etc.
+
+ Suppose that the goal state has @{text "n \<ge> 0"} subgoals. Many of
+ these tacticals address subgoal ranges counting downwards from
+ @{text "n"} towards @{text "1"}. This has the fortunate effect that
+ newly emerging subgoals are concatenated in the result, without
+ interfering each other. Nonetheless, there might be situations
+ where a different order is desired. *}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML ALLGOALS: "(int -> tactic) -> tactic"} \\
+ @{index_ML SOMEGOAL: "(int -> tactic) -> tactic"} \\
+ @{index_ML FIRSTGOAL: "(int -> tactic) -> tactic"} \\
+ @{index_ML HEADGOAL: "(int -> tactic) -> tactic"} \\
+ @{index_ML REPEAT_SOME: "(int -> tactic) -> tactic"} \\
+ @{index_ML REPEAT_FIRST: "(int -> tactic) -> tactic"} \\
+ @{index_ML RANGE: "(int -> tactic) list -> int -> tactic"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML ALLGOALS}~@{text "tac"} is equivalent to @{text "tac
+ n"}~@{ML_op THEN}~@{text "\<dots>"}~@{ML_op THEN}~@{text "tac 1"}. It
+ applies the @{text tac} to all the subgoals, counting downwards.
+
+ \item @{ML SOMEGOAL}~@{text "tac"} is equivalent to @{text "tac
+ n"}~@{ML_op ORELSE}~@{text "\<dots>"}~@{ML_op ORELSE}~@{text "tac 1"}. It
+ applies @{text "tac"} to one subgoal, counting downwards.
+
+ \item @{ML FIRSTGOAL}~@{text "tac"} is equivalent to @{text "tac
+ 1"}~@{ML_op ORELSE}~@{text "\<dots>"}~@{ML_op ORELSE}~@{text "tac n"}. It
+ applies @{text "tac"} to one subgoal, counting upwards.
+
+ \item @{ML HEADGOAL}~@{text "tac"} is equivalent to @{text "tac 1"}.
+ It applies @{text "tac"} unconditionally to the first subgoal.
+
+ \item @{ML REPEAT_SOME}~@{text "tac"} applies @{text "tac"} once or
+ more to a subgoal, counting downwards.
+
+ \item @{ML REPEAT_FIRST}~@{text "tac"} applies @{text "tac"} once or
+ more to a subgoal, counting upwards.
+
+ \item @{ML RANGE}~@{text "[tac\<^sub>1, \<dots>, tac\<^sub>k] i"} is equivalent to
+ @{text "tac\<^sub>k (i + k - 1)"}~@{ML_op THEN}~@{text "\<dots>"}~@{ML_op
+ THEN}~@{text "tac\<^sub>1 i"}. It applies the given list of tactics to the
+ corresponding range of subgoals, counting downwards.
+
+ \end{description}
+*}
+
+
+subsection {* Control and search tacticals *}
+
+text {* A predicate on theorems @{ML_type "thm -> bool"} can test
+ whether a goal state enjoys some desirable property --- such as
+ having no subgoals. Tactics that search for satisfactory goal
+ states are easy to express. The main search procedures,
+ depth-first, breadth-first and best-first, are provided as
+ tacticals. They generate the search tree by repeatedly applying a
+ given tactic. *}
+
+
+text %mlref ""
+
+subsubsection {* Filtering a tactic's results *}
+
+text {*
+ \begin{mldecls}
+ @{index_ML FILTER: "(thm -> bool) -> tactic -> tactic"} \\
+ @{index_ML CHANGED: "tactic -> tactic"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML FILTER}~@{text "sat tac"} applies @{text "tac"} to the
+ goal state and returns a sequence consisting of those result goal
+ states that are satisfactory in the sense of @{text "sat"}.
+
+ \item @{ML CHANGED}~@{text "tac"} applies @{text "tac"} to the goal
+ state and returns precisely those states that differ from the
+ original state (according to @{ML Thm.eq_thm}). Thus @{ML
+ CHANGED}~@{text "tac"} always has some effect on the state.
+
+ \end{description}
+*}
+
+
+subsubsection {* Depth-first search *}
+
+text {*
+ \begin{mldecls}
+ @{index_ML DEPTH_FIRST: "(thm -> bool) -> tactic -> tactic"} \\
+ @{index_ML DEPTH_SOLVE: "tactic -> tactic"} \\
+ @{index_ML DEPTH_SOLVE_1: "tactic -> tactic"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML DEPTH_FIRST}~@{text "sat tac"} returns the goal state if
+ @{text "sat"} returns true. Otherwise it applies @{text "tac"},
+ then recursively searches from each element of the resulting
+ sequence. The code uses a stack for efficiency, in effect applying
+ @{text "tac"}~@{ML_op THEN}~@{ML DEPTH_FIRST}~@{text "sat tac"} to
+ the state.
+
+ \item @{ML DEPTH_SOLVE}@{text "tac"} uses @{ML DEPTH_FIRST} to
+ search for states having no subgoals.
+
+ \item @{ML DEPTH_SOLVE_1}~@{text "tac"} uses @{ML DEPTH_FIRST} to
+ search for states having fewer subgoals than the given state. Thus,
+ it insists upon solving at least one subgoal.
+
+ \end{description}
+*}
+
+
+subsubsection {* Other search strategies *}
+
+text {*
+ \begin{mldecls}
+ @{index_ML BREADTH_FIRST: "(thm -> bool) -> tactic -> tactic"} \\
+ @{index_ML BEST_FIRST: "(thm -> bool) * (thm -> int) -> tactic -> tactic"} \\
+ @{index_ML THEN_BEST_FIRST: "tactic -> (thm -> bool) * (thm -> int) -> tactic -> tactic"} \\
+ \end{mldecls}
+
+ These search strategies will find a solution if one exists.
+ However, they do not enumerate all solutions; they terminate after
+ the first satisfactory result from @{text "tac"}.
+
+ \begin{description}
+
+ \item @{ML BREADTH_FIRST}~@{text "sat tac"} uses breadth-first
+ search to find states for which @{text "sat"} is true. For most
+ applications, it is too slow.
+
+ \item @{ML BEST_FIRST}~@{text "(sat, dist) tac"} does a heuristic
+ search, using @{text "dist"} to estimate the distance from a
+ satisfactory state (in the sense of @{text "sat"}). It maintains a
+ list of states ordered by distance. It applies @{text "tac"} to the
+ head of this list; if the result contains any satisfactory states,
+ then it returns them. Otherwise, @{ML BEST_FIRST} adds the new
+ states to the list, and continues.
+
+ The distance function is typically @{ML size_of_thm}, which computes
+ the size of the state. The smaller the state, the fewer and simpler
+ subgoals it has.
+
+ \item @{ML THEN_BEST_FIRST}~@{text "tac\<^sub>0 (sat, dist) tac"} is like
+ @{ML BEST_FIRST}, except that the priority queue initially contains
+ the result of applying @{text "tac\<^sub>0"} to the goal state. This
+ tactical permits separate tactics for starting the search and
+ continuing the search.
+
+ \end{description}
+*}
+
+
+subsubsection {* Auxiliary tacticals for searching *}
+
+text {*
+ \begin{mldecls}
+ @{index_ML COND: "(thm -> bool) -> tactic -> tactic -> tactic"} \\
+ @{index_ML IF_UNSOLVED: "tactic -> tactic"} \\
+ @{index_ML SOLVE: "tactic -> tactic"} \\
+ @{index_ML DETERM: "tactic -> tactic"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML COND}~@{text "sat tac\<^sub>1 tac\<^sub>2"} applies @{text "tac\<^sub>1"} to
+ the goal state if it satisfies predicate @{text "sat"}, and applies
+ @{text "tac\<^sub>2"}. It is a conditional tactical in that only one of
+ @{text "tac\<^sub>1"} and @{text "tac\<^sub>2"} is applied to a goal state.
+ However, both @{text "tac\<^sub>1"} and @{text "tac\<^sub>2"} are evaluated
+ because ML uses eager evaluation.
+
+ \item @{ML IF_UNSOLVED}~@{text "tac"} applies @{text "tac"} to the
+ goal state if it has any subgoals, and simply returns the goal state
+ otherwise. Many common tactics, such as @{ML resolve_tac}, fail if
+ applied to a goal state that has no subgoals.
+
+ \item @{ML SOLVE}~@{text "tac"} applies @{text "tac"} to the goal
+ state and then fails iff there are subgoals left.
+
+ \item @{ML DETERM}~@{text "tac"} applies @{text "tac"} to the goal
+ state and returns the head of the resulting sequence. @{ML DETERM}
+ limits the search space by making its argument deterministic.
+
+ \end{description}
+*}
+
+
+subsubsection {* Predicates and functions useful for searching *}
+
+text {*
+ \begin{mldecls}
+ @{index_ML has_fewer_prems: "int -> thm -> bool"} \\
+ @{index_ML Thm.eq_thm: "thm * thm -> bool"} \\
+ @{index_ML Thm.eq_thm_prop: "thm * thm -> bool"} \\
+ @{index_ML size_of_thm: "thm -> int"} \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item @{ML has_fewer_prems}~@{text "n thm"} reports whether @{text
+ "thm"} has fewer than @{text "n"} premises.
+
+ \item @{ML Thm.eq_thm}~@{text "(thm\<^sub>1, thm\<^sub>2)"} reports whether @{text
+ "thm\<^sub>1"} and @{text "thm\<^sub>2"} are equal. Both theorems must have the
+ same conclusions, the same set of hypotheses, and the same set of sort
+ hypotheses. Names of bound variables are ignored as usual.
+
+ \item @{ML Thm.eq_thm_prop}~@{text "(thm\<^sub>1, thm\<^sub>2)"} reports whether
+ the propositions of @{text "thm\<^sub>1"} and @{text "thm\<^sub>2"} are equal.
+ Names of bound variables are ignored.
+
+ \item @{ML size_of_thm}~@{text "thm"} computes the size of @{text
+ "thm"}, namely the number of variables, constants and abstractions
+ in its conclusion. It may serve as a distance function for
+ @{ML BEST_FIRST}.
+
+ \end{description}
+*}
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Implementation/document/build Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,18 @@
+#!/usr/bin/env bash
+
+set -e
+
+FORMAT="$1"
+VARIANT="$2"
+
+"$ISABELLE_TOOL" logo Isar
+
+cp "$ISABELLE_HOME/src/Doc/iman.sty" .
+cp "$ISABELLE_HOME/src/Doc/extra.sty" .
+cp "$ISABELLE_HOME/src/Doc/isar.sty" .
+cp "$ISABELLE_HOME/src/Doc/ttbox.sty" .
+cp "$ISABELLE_HOME/src/Doc/underscore.sty" .
+cp "$ISABELLE_HOME/src/Doc/manual.bib" .
+
+"$ISABELLE_HOME/src/Doc/prepare_document" "$FORMAT"
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Implementation/document/root.tex Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,125 @@
+\documentclass[12pt,a4paper,fleqn]{report}
+\usepackage[T1]{fontenc}
+\usepackage{latexsym,graphicx}
+\usepackage[refpage]{nomencl}
+\usepackage{iman,extra,isar,proof}
+\usepackage[nohyphen,strings]{underscore}
+\usepackage{isabelle}
+\usepackage{isabellesym}
+\usepackage{railsetup}
+\usepackage{ttbox}
+\usepackage{supertabular}
+\usepackage{style}
+\usepackage{pdfsetup}
+
+
+\hyphenation{Isabelle}
+\hyphenation{Isar}
+
+\isadroptag{theory}
+\title{\includegraphics[scale=0.5]{isabelle_isar}
+ \\[4ex] The Isabelle/Isar Implementation}
+\author{\emph{Makarius Wenzel} \\[3ex]
+ With Contributions by
+ Stefan Berghofer, \\
+ Florian Haftmann
+ and Larry Paulson
+}
+
+\makeindex
+
+
+\begin{document}
+
+\maketitle
+
+\begin{abstract}
+ We describe the key concepts underlying the Isabelle/Isar
+ implementation, including ML references for the most important
+ functions. The aim is to give some insight into the overall system
+ architecture, and provide clues on implementing applications within
+ this framework.
+\end{abstract}
+
+\vspace*{2.5cm}
+\begin{quote}
+
+ {\small\em Isabelle was not designed; it evolved. Not everyone
+ likes this idea. Specification experts rightly abhor
+ trial-and-error programming. They suggest that no one should
+ write a program without first writing a complete formal
+ specification. But university departments are not software houses.
+ Programs like Isabelle are not products: when they have served
+ their purpose, they are discarded.}
+
+ Lawrence C. Paulson, ``Isabelle: The Next 700 Theorem Provers''
+
+ \vspace*{1cm}
+
+ {\small\em As I did 20 years ago, I still fervently believe that the
+ only way to make software secure, reliable, and fast is to make it
+ small. Fight features.}
+
+ Andrew S. Tanenbaum
+
+ \vspace*{1cm}
+
+ {\small\em One thing that UNIX does not need is more features. It is
+ successful in part because it has a small number of good ideas
+ that work well together. Merely adding features does not make it
+ easier for users to do things --- it just makes the manual
+ thicker. The right solution in the right place is always more
+ effective than haphazard hacking.}
+
+ Rob Pike and Brian W. Kernighan
+
+ \vspace*{1cm}
+
+ {\small\em If you look at software today, through the lens of the
+ history of engineering, it's certainly engineering of a sort--but
+ it's the kind of engineering that people without the concept of
+ the arch did. Most software today is very much like an Egyptian
+ pyramid with millions of bricks piled on top of each other, with
+ no structural integrity, but just done by brute force and
+ thousands of slaves.}
+
+ Alan Kay
+
+\end{quote}
+
+\thispagestyle{empty}\clearpage
+
+\pagenumbering{roman}
+\tableofcontents
+\listoffigures
+\clearfirst
+
+\setcounter{chapter}{-1}
+
+\input{ML.tex}
+\input{Prelim.tex}
+\input{Logic.tex}
+\input{Syntax.tex}
+\input{Tactic.tex}
+\input{Eq.tex}
+\input{Proof.tex}
+\input{Isar.tex}
+\input{Local_Theory.tex}
+\input{Integration.tex}
+
+\begingroup
+\tocentry{\bibname}
+\bibliographystyle{abbrv} \small\raggedright\frenchspacing
+\bibliography{manual}
+\endgroup
+
+\tocentry{\indexname}
+\printindex
+
+\end{document}
+
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: t
+%%% End:
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Implementation/document/style.sty Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,68 @@
+%% toc
+\newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1}
+\@mkboth{\MakeUppercase{#1}}{\MakeUppercase{#1}}}
+
+%% references
+\newcommand{\secref}[1]{\S\ref{#1}}
+\newcommand{\chref}[1]{chapter~\ref{#1}}
+\newcommand{\figref}[1]{figure~\ref{#1}}
+
+%% math
+\newcommand{\text}[1]{\mbox{#1}}
+\newcommand{\isasymvartheta}{\isamath{\theta}}
+\newcommand{\isactrlvec}[1]{\emph{$\vec{#1}$}}
+\newcommand{\isactrlBG}{\isacharbackquoteopen}
+\newcommand{\isactrlEN}{\isacharbackquoteclose}
+
+\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2}
+
+\pagestyle{headings}
+\sloppy
+\binperiod
+
+\parindent 0pt\parskip 0.5ex
+
+\renewcommand{\isadigit}[1]{\isamath{#1}}
+
+\newenvironment{mldecls}{\par\noindent\begingroup\footnotesize\def\isanewline{\\}\begin{tabular}{l}}{\end{tabular}\smallskip\endgroup}
+
+\isafoldtag{FIXME}
+
+\isakeeptag{mlref}
+\renewcommand{\isatagmlref}{\subsection*{\makebox[0pt][r]{\fbox{ML}~~}Reference}}
+\renewcommand{\endisatagmlref}{}
+
+\isakeeptag{mlantiq}
+\renewcommand{\isatagmlantiq}{\subsection*{\makebox[0pt][r]{\fbox{ML}~~}Antiquotations}}
+\renewcommand{\endisatagmlantiq}{}
+
+\isakeeptag{mlex}
+\renewcommand{\isatagmlex}{\subsection*{\makebox[0pt][r]{\fbox{ML}~~}Examples}}
+\renewcommand{\endisatagmlex}{}
+
+\renewcommand{\isatagML}{\begingroup\isabellestyle{default}\isastyle\def\isadigit##1{##1}}
+\renewcommand{\endisatagML}{\endgroup}
+
+\newcommand{\minorcmd}[1]{{\sf #1}}
+\newcommand{\isasymtype}{\minorcmd{type}}
+\newcommand{\isasymval}{\minorcmd{val}}
+
+\newcommand{\isasymFIX}{\isakeyword{fix}}
+\newcommand{\isasymASSUME}{\isakeyword{assume}}
+\newcommand{\isasymDEFINE}{\isakeyword{define}}
+\newcommand{\isasymNOTE}{\isakeyword{note}}
+\newcommand{\isasymGUESS}{\isakeyword{guess}}
+\newcommand{\isasymOBTAIN}{\isakeyword{obtain}}
+\newcommand{\isasymTHEORY}{\isakeyword{theory}}
+\newcommand{\isasymUSES}{\isakeyword{uses}}
+\newcommand{\isasymEND}{\isakeyword{end}}
+\newcommand{\isasymCONSTS}{\isakeyword{consts}}
+\newcommand{\isasymDEFS}{\isakeyword{defs}}
+\newcommand{\isasymTHEOREM}{\isakeyword{theorem}}
+\newcommand{\isasymDEFINITION}{\isakeyword{definition}}
+
+\isabellestyle{literal}
+
+\railtermfont{\isabellestyle{tt}}
+\railnontermfont{\isabellestyle{itunderscore}}
+\railnamefont{\isabellestyle{itunderscore}}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Isar-Ref/Base.thy Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,7 @@
+theory Base
+imports Pure
+begin
+
+ML_file "../antiquote_setup.ML"
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Isar-Ref/Document_Preparation.thy Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,589 @@
+theory Document_Preparation
+imports Base Main
+begin
+
+chapter {* Document preparation \label{ch:document-prep} *}
+
+text {* Isabelle/Isar provides a simple document preparation system
+ based on {PDF-\LaTeX}, with support for hyperlinks and bookmarks
+ within that format. This allows to produce papers, books, theses
+ etc.\ from Isabelle theory sources.
+
+ {\LaTeX} output is generated while processing a \emph{session} in
+ batch mode, as explained in the \emph{The Isabelle System Manual}
+ \cite{isabelle-sys}. The main Isabelle tools to get started with
+ document preparation are @{tool_ref mkroot} and @{tool_ref build}.
+
+ The classic Isabelle/HOL tutorial \cite{isabelle-hol-book} also
+ explains some aspects of theory presentation. *}
+
+
+section {* Markup commands \label{sec:markup} *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{command_def "header"} & : & @{text "toplevel \<rightarrow> toplevel"} \\[0.5ex]
+ @{command_def "chapter"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+ @{command_def "section"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+ @{command_def "subsection"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+ @{command_def "subsubsection"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+ @{command_def "text"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+ @{command_def "text_raw"} & : & @{text "local_theory \<rightarrow> local_theory"} \\[0.5ex]
+ @{command_def "sect"} & : & @{text "proof \<rightarrow> proof"} \\
+ @{command_def "subsect"} & : & @{text "proof \<rightarrow> proof"} \\
+ @{command_def "subsubsect"} & : & @{text "proof \<rightarrow> proof"} \\
+ @{command_def "txt"} & : & @{text "proof \<rightarrow> proof"} \\
+ @{command_def "txt_raw"} & : & @{text "proof \<rightarrow> proof"} \\
+ \end{matharray}
+
+ Markup commands provide a structured way to insert text into the
+ document generated from a theory. Each markup command takes a
+ single @{syntax text} argument, which is passed as argument to a
+ corresponding {\LaTeX} macro. The default macros provided by
+ @{file "~~/lib/texinputs/isabelle.sty"} can be redefined according
+ to the needs of the underlying document and {\LaTeX} styles.
+
+ Note that formal comments (\secref{sec:comments}) are similar to
+ markup commands, but have a different status within Isabelle/Isar
+ syntax.
+
+ @{rail \<open>
+ (@@{command chapter} | @@{command section} | @@{command subsection} |
+ @@{command subsubsection} | @@{command text}) @{syntax target}? @{syntax text}
+ ;
+ (@@{command header} | @@{command text_raw} | @@{command sect} | @@{command subsect} |
+ @@{command subsubsect} | @@{command txt} | @@{command txt_raw}) @{syntax text}
+ \<close>}
+
+ \begin{description}
+
+ \item @{command header} provides plain text markup just preceding
+ the formal beginning of a theory. The corresponding {\LaTeX} macro
+ is @{verbatim "\\isamarkupheader"}, which acts like @{command
+ section} by default.
+
+ \item @{command chapter}, @{command section}, @{command subsection},
+ and @{command subsubsection} mark chapter and section headings
+ within the main theory body or local theory targets. The
+ corresponding {\LaTeX} macros are @{verbatim "\\isamarkupchapter"},
+ @{verbatim "\\isamarkupsection"}, @{verbatim
+ "\\isamarkupsubsection"} etc.
+
+ \item @{command sect}, @{command subsect}, and @{command subsubsect}
+ mark section headings within proofs. The corresponding {\LaTeX}
+ macros are @{verbatim "\\isamarkupsect"}, @{verbatim
+ "\\isamarkupsubsect"} etc.
+
+ \item @{command text} and @{command txt} specify paragraphs of plain
+ text. This corresponds to a {\LaTeX} environment @{verbatim
+ "\\begin{isamarkuptext}"} @{text "\<dots>"} @{verbatim
+ "\\end{isamarkuptext}"} etc.
+
+ \item @{command text_raw} and @{command txt_raw} insert {\LaTeX}
+ source into the output, without additional markup. Thus the full
+ range of document manipulations becomes available, at the risk of
+ messing up document output.
+
+ \end{description}
+
+ Except for @{command "text_raw"} and @{command "txt_raw"}, the text
+ passed to any of the above markup commands may refer to formal
+ entities via \emph{document antiquotations}, see also
+ \secref{sec:antiq}. These are interpreted in the present theory or
+ proof context, or the named @{text "target"}.
+
+ \medskip The proof markup commands closely resemble those for theory
+ specifications, but have a different formal status and produce
+ different {\LaTeX} macros. The default definitions coincide for
+ analogous commands such as @{command section} and @{command sect}.
+*}
+
+
+section {* Document Antiquotations \label{sec:antiq} *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{antiquotation_def "theory"} & : & @{text antiquotation} \\
+ @{antiquotation_def "thm"} & : & @{text antiquotation} \\
+ @{antiquotation_def "lemma"} & : & @{text antiquotation} \\
+ @{antiquotation_def "prop"} & : & @{text antiquotation} \\
+ @{antiquotation_def "term"} & : & @{text antiquotation} \\
+ @{antiquotation_def term_type} & : & @{text antiquotation} \\
+ @{antiquotation_def typeof} & : & @{text antiquotation} \\
+ @{antiquotation_def const} & : & @{text antiquotation} \\
+ @{antiquotation_def abbrev} & : & @{text antiquotation} \\
+ @{antiquotation_def typ} & : & @{text antiquotation} \\
+ @{antiquotation_def type} & : & @{text antiquotation} \\
+ @{antiquotation_def class} & : & @{text antiquotation} \\
+ @{antiquotation_def "text"} & : & @{text antiquotation} \\
+ @{antiquotation_def goals} & : & @{text antiquotation} \\
+ @{antiquotation_def subgoals} & : & @{text antiquotation} \\
+ @{antiquotation_def prf} & : & @{text antiquotation} \\
+ @{antiquotation_def full_prf} & : & @{text antiquotation} \\
+ @{antiquotation_def ML} & : & @{text antiquotation} \\
+ @{antiquotation_def ML_op} & : & @{text antiquotation} \\
+ @{antiquotation_def ML_type} & : & @{text antiquotation} \\
+ @{antiquotation_def ML_structure} & : & @{text antiquotation} \\
+ @{antiquotation_def ML_functor} & : & @{text antiquotation} \\
+ @{antiquotation_def "file"} & : & @{text antiquotation} \\
+ @{antiquotation_def "url"} & : & @{text antiquotation} \\
+ \end{matharray}
+
+ The overall content of an Isabelle/Isar theory may alternate between
+ formal and informal text. The main body consists of formal
+ specification and proof commands, interspersed with markup commands
+ (\secref{sec:markup}) or document comments (\secref{sec:comments}).
+ The argument of markup commands quotes informal text to be printed
+ in the resulting document, but may again refer to formal entities
+ via \emph{document antiquotations}.
+
+ For example, embedding of ``@{text [source=false] "@{term [show_types] \"f x = a + x\"}"}''
+ within a text block makes
+ \isa{{\isacharparenleft}f{\isasymColon}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}a{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharplus}\ x} appear in the final {\LaTeX} document.
+
+ Antiquotations usually spare the author tedious typing of logical
+ entities in full detail. Even more importantly, some degree of
+ consistency-checking between the main body of formal text and its
+ informal explanation is achieved, since terms and types appearing in
+ antiquotations are checked within the current theory or proof
+ context.
+
+ %% FIXME less monolithic presentation, move to individual sections!?
+ @{rail \<open>
+ '@{' antiquotation '}'
+ ;
+ @{syntax_def antiquotation}:
+ @@{antiquotation theory} options @{syntax name} |
+ @@{antiquotation thm} options styles @{syntax thmrefs} |
+ @@{antiquotation lemma} options @{syntax prop} @'by' @{syntax method} @{syntax method}? |
+ @@{antiquotation prop} options styles @{syntax prop} |
+ @@{antiquotation term} options styles @{syntax term} |
+ @@{antiquotation (HOL) value} options styles @{syntax term} |
+ @@{antiquotation term_type} options styles @{syntax term} |
+ @@{antiquotation typeof} options styles @{syntax term} |
+ @@{antiquotation const} options @{syntax term} |
+ @@{antiquotation abbrev} options @{syntax term} |
+ @@{antiquotation typ} options @{syntax type} |
+ @@{antiquotation type} options @{syntax name} |
+ @@{antiquotation class} options @{syntax name} |
+ @@{antiquotation text} options @{syntax name}
+ ;
+ @{syntax antiquotation}:
+ @@{antiquotation goals} options |
+ @@{antiquotation subgoals} options |
+ @@{antiquotation prf} options @{syntax thmrefs} |
+ @@{antiquotation full_prf} options @{syntax thmrefs} |
+ @@{antiquotation ML} options @{syntax name} |
+ @@{antiquotation ML_op} options @{syntax name} |
+ @@{antiquotation ML_type} options @{syntax name} |
+ @@{antiquotation ML_structure} options @{syntax name} |
+ @@{antiquotation ML_functor} options @{syntax name} |
+ @@{antiquotation "file"} options @{syntax name} |
+ @@{antiquotation file_unchecked} options @{syntax name} |
+ @@{antiquotation url} options @{syntax name}
+ ;
+ options: '[' (option * ',') ']'
+ ;
+ option: @{syntax name} | @{syntax name} '=' @{syntax name}
+ ;
+ styles: '(' (style + ',') ')'
+ ;
+ style: (@{syntax name} +)
+ \<close>}
+
+ Note that the syntax of antiquotations may \emph{not} include source
+ comments @{verbatim "(*"}~@{text "\<dots>"}~@{verbatim "*)"} nor verbatim
+ text @{verbatim "{"}@{verbatim "*"}~@{text "\<dots>"}~@{verbatim
+ "*"}@{verbatim "}"}.
+
+ \begin{description}
+
+ \item @{text "@{theory A}"} prints the name @{text "A"}, which is
+ guaranteed to refer to a valid ancestor theory in the current
+ context.
+
+ \item @{text "@{thm a\<^sub>1 \<dots> a\<^sub>n}"} prints theorems @{text "a\<^sub>1 \<dots> a\<^sub>n"}.
+ Full fact expressions are allowed here, including attributes
+ (\secref{sec:syn-att}).
+
+ \item @{text "@{prop \<phi>}"} prints a well-typed proposition @{text
+ "\<phi>"}.
+
+ \item @{text "@{lemma \<phi> by m}"} proves a well-typed proposition
+ @{text "\<phi>"} by method @{text m} and prints the original @{text "\<phi>"}.
+
+ \item @{text "@{term t}"} prints a well-typed term @{text "t"}.
+
+ \item @{text "@{value t}"} evaluates a term @{text "t"} and prints
+ its result, see also @{command_ref (HOL) value}.
+
+ \item @{text "@{term_type t}"} prints a well-typed term @{text "t"}
+ annotated with its type.
+
+ \item @{text "@{typeof t}"} prints the type of a well-typed term
+ @{text "t"}.
+
+ \item @{text "@{const c}"} prints a logical or syntactic constant
+ @{text "c"}.
+
+ \item @{text "@{abbrev c x\<^sub>1 \<dots> x\<^sub>n}"} prints a constant abbreviation
+ @{text "c x\<^sub>1 \<dots> x\<^sub>n \<equiv> rhs"} as defined in the current context.
+
+ \item @{text "@{typ \<tau>}"} prints a well-formed type @{text "\<tau>"}.
+
+ \item @{text "@{type \<kappa>}"} prints a (logical or syntactic) type
+ constructor @{text "\<kappa>"}.
+
+ \item @{text "@{class c}"} prints a class @{text c}.
+
+ \item @{text "@{text s}"} prints uninterpreted source text @{text
+ s}. This is particularly useful to print portions of text according
+ to the Isabelle document style, without demanding well-formedness,
+ e.g.\ small pieces of terms that should not be parsed or
+ type-checked yet.
+
+ \item @{text "@{goals}"} prints the current \emph{dynamic} goal
+ state. This is mainly for support of tactic-emulation scripts
+ within Isar. Presentation of goal states does not conform to the
+ idea of human-readable proof documents!
+
+ When explaining proofs in detail it is usually better to spell out
+ the reasoning via proper Isar proof commands, instead of peeking at
+ the internal machine configuration.
+
+ \item @{text "@{subgoals}"} is similar to @{text "@{goals}"}, but
+ does not print the main goal.
+
+ \item @{text "@{prf a\<^sub>1 \<dots> a\<^sub>n}"} prints the (compact) proof terms
+ corresponding to the theorems @{text "a\<^sub>1 \<dots> a\<^sub>n"}. Note that this
+ requires proof terms to be switched on for the current logic
+ session.
+
+ \item @{text "@{full_prf a\<^sub>1 \<dots> a\<^sub>n}"} is like @{text "@{prf a\<^sub>1 \<dots>
+ a\<^sub>n}"}, but prints the full proof terms, i.e.\ also displays
+ information omitted in the compact proof term, which is denoted by
+ ``@{text _}'' placeholders there.
+
+ \item @{text "@{ML s}"}, @{text "@{ML_op s}"}, @{text "@{ML_type
+ s}"}, @{text "@{ML_structure s}"}, and @{text "@{ML_functor s}"}
+ check text @{text s} as ML value, infix operator, type, structure,
+ and functor respectively. The source is printed verbatim.
+
+ \item @{text "@{file path}"} checks that @{text "path"} refers to a
+ file (or directory) and prints it verbatim.
+
+ \item @{text "@{file_unchecked path}"} is like @{text "@{file
+ path}"}, but does not check the existence of the @{text "path"}
+ within the file-system.
+
+ \item @{text "@{url name}"} produces markup for the given URL, which
+ results in an active hyperlink within the text.
+
+ \end{description}
+*}
+
+
+subsection {* Styled antiquotations *}
+
+text {* The antiquotations @{text thm}, @{text prop} and @{text
+ term} admit an extra \emph{style} specification to modify the
+ printed result. A style is specified by a name with a possibly
+ empty number of arguments; multiple styles can be sequenced with
+ commas. The following standard styles are available:
+
+ \begin{description}
+
+ \item @{text lhs} extracts the first argument of any application
+ form with at least two arguments --- typically meta-level or
+ object-level equality, or any other binary relation.
+
+ \item @{text rhs} is like @{text lhs}, but extracts the second
+ argument.
+
+ \item @{text "concl"} extracts the conclusion @{text C} from a rule
+ in Horn-clause normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}.
+
+ \item @{text "prem"} @{text n} extract premise number
+ @{text "n"} from from a rule in Horn-clause
+ normal form @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> C"}
+
+ \end{description}
+*}
+
+
+subsection {* General options *}
+
+text {* The following options are available to tune the printed output
+ of antiquotations. Note that many of these coincide with system and
+ configuration options of the same names.
+
+ \begin{description}
+
+ \item @{antiquotation_option_def show_types}~@{text "= bool"} and
+ @{antiquotation_option_def show_sorts}~@{text "= bool"} control
+ printing of explicit type and sort constraints.
+
+ \item @{antiquotation_option_def show_structs}~@{text "= bool"}
+ controls printing of implicit structures.
+
+ \item @{antiquotation_option_def show_abbrevs}~@{text "= bool"}
+ controls folding of abbreviations.
+
+ \item @{antiquotation_option_def names_long}~@{text "= bool"} forces
+ names of types and constants etc.\ to be printed in their fully
+ qualified internal form.
+
+ \item @{antiquotation_option_def names_short}~@{text "= bool"}
+ forces names of types and constants etc.\ to be printed unqualified.
+ Note that internalizing the output again in the current context may
+ well yield a different result.
+
+ \item @{antiquotation_option_def names_unique}~@{text "= bool"}
+ determines whether the printed version of qualified names should be
+ made sufficiently long to avoid overlap with names declared further
+ back. Set to @{text false} for more concise output.
+
+ \item @{antiquotation_option_def eta_contract}~@{text "= bool"}
+ prints terms in @{text \<eta>}-contracted form.
+
+ \item @{antiquotation_option_def display}~@{text "= bool"} indicates
+ if the text is to be output as multi-line ``display material'',
+ rather than a small piece of text without line breaks (which is the
+ default).
+
+ In this mode the embedded entities are printed in the same style as
+ the main theory text.
+
+ \item @{antiquotation_option_def break}~@{text "= bool"} controls
+ line breaks in non-display material.
+
+ \item @{antiquotation_option_def quotes}~@{text "= bool"} indicates
+ if the output should be enclosed in double quotes.
+
+ \item @{antiquotation_option_def mode}~@{text "= name"} adds @{text
+ name} to the print mode to be used for presentation. Note that the
+ standard setup for {\LaTeX} output is already present by default,
+ including the modes @{text latex} and @{text xsymbols}.
+
+ \item @{antiquotation_option_def margin}~@{text "= nat"} and
+ @{antiquotation_option_def indent}~@{text "= nat"} change the margin
+ or indentation for pretty printing of display material.
+
+ \item @{antiquotation_option_def goals_limit}~@{text "= nat"}
+ determines the maximum number of subgoals to be printed (for goal-based
+ antiquotation).
+
+ \item @{antiquotation_option_def source}~@{text "= bool"} prints the
+ original source text of the antiquotation arguments, rather than its
+ internal representation. Note that formal checking of
+ @{antiquotation "thm"}, @{antiquotation "term"}, etc. is still
+ enabled; use the @{antiquotation "text"} antiquotation for unchecked
+ output.
+
+ Regular @{text "term"} and @{text "typ"} antiquotations with @{text
+ "source = false"} involve a full round-trip from the original source
+ to an internalized logical entity back to a source form, according
+ to the syntax of the current context. Thus the printed output is
+ not under direct control of the author, it may even fluctuate a bit
+ as the underlying theory is changed later on.
+
+ In contrast, @{antiquotation_option source}~@{text "= true"}
+ admits direct printing of the given source text, with the desirable
+ well-formedness check in the background, but without modification of
+ the printed text.
+
+ \end{description}
+
+ For boolean flags, ``@{text "name = true"}'' may be abbreviated as
+ ``@{text name}''. All of the above flags are disabled by default,
+ unless changed specifically for a logic session in the corresponding
+ @{verbatim "ROOT"} file. *}
+
+
+section {* Markup via command tags \label{sec:tags} *}
+
+text {* Each Isabelle/Isar command may be decorated by additional
+ presentation tags, to indicate some modification in the way it is
+ printed in the document.
+
+ @{rail \<open>
+ @{syntax_def tags}: ( tag * )
+ ;
+ tag: '%' (@{syntax ident} | @{syntax string})
+ \<close>}
+
+ Some tags are pre-declared for certain classes of commands, serving
+ as default markup if no tags are given in the text:
+
+ \medskip
+ \begin{tabular}{ll}
+ @{text "theory"} & theory begin/end \\
+ @{text "proof"} & all proof commands \\
+ @{text "ML"} & all commands involving ML code \\
+ \end{tabular}
+
+ \medskip The Isabelle document preparation system
+ \cite{isabelle-sys} allows tagged command regions to be presented
+ specifically, e.g.\ to fold proof texts, or drop parts of the text
+ completely.
+
+ For example ``@{command "by"}~@{text "%invisible auto"}'' causes
+ that piece of proof to be treated as @{text invisible} instead of
+ @{text "proof"} (the default), which may be shown or hidden
+ depending on the document setup. In contrast, ``@{command
+ "by"}~@{text "%visible auto"}'' forces this text to be shown
+ invariably.
+
+ Explicit tag specifications within a proof apply to all subsequent
+ commands of the same level of nesting. For example, ``@{command
+ "proof"}~@{text "%visible \<dots>"}~@{command "qed"}'' forces the whole
+ sub-proof to be typeset as @{text visible} (unless some of its parts
+ are tagged differently).
+
+ \medskip Command tags merely produce certain markup environments for
+ type-setting. The meaning of these is determined by {\LaTeX}
+ macros, as defined in @{file "~~/lib/texinputs/isabelle.sty"} or
+ by the document author. The Isabelle document preparation tools
+ also provide some high-level options to specify the meaning of
+ arbitrary tags to ``keep'', ``drop'', or ``fold'' the corresponding
+ parts of the text. Logic sessions may also specify ``document
+ versions'', where given tags are interpreted in some particular way.
+ Again see \cite{isabelle-sys} for further details.
+*}
+
+
+section {* Railroad diagrams *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{antiquotation_def "rail"} & : & @{text antiquotation} \\
+ \end{matharray}
+
+ @{rail \<open>
+ 'rail' (@{syntax string} | @{syntax cartouche})
+ \<close>}
+
+ The @{antiquotation rail} antiquotation allows to include syntax
+ diagrams into Isabelle documents. {\LaTeX} requires the style file
+ @{file "~~/lib/texinputs/pdfsetup.sty"}, which can be used via
+ @{verbatim "\\usepackage{pdfsetup}"} in @{verbatim "root.tex"}, for
+ example.
+
+ The rail specification language is quoted here as Isabelle @{syntax
+ string} or text @{syntax "cartouche"}; it has its own grammar given
+ below.
+
+ \begingroup
+ \def\isasymnewline{\isatext{\tt\isacharbackslash<newline>}}
+ @{rail \<open>
+ rule? + ';'
+ ;
+ rule: ((identifier | @{syntax antiquotation}) ':')? body
+ ;
+ body: concatenation + '|'
+ ;
+ concatenation: ((atom '?'?) +) (('*' | '+') atom?)?
+ ;
+ atom: '(' body? ')' | identifier |
+ '@'? (string | @{syntax antiquotation}) |
+ '\<newline>'
+ \<close>}
+ \endgroup
+
+ The lexical syntax of @{text "identifier"} coincides with that of
+ @{syntax ident} in regular Isabelle syntax, but @{text string} uses
+ single quotes instead of double quotes of the standard @{syntax
+ string} category.
+
+ Each @{text rule} defines a formal language (with optional name),
+ using a notation that is similar to EBNF or regular expressions with
+ recursion. The meaning and visual appearance of these rail language
+ elements is illustrated by the following representative examples.
+
+ \begin{itemize}
+
+ \item Empty @{verbatim "()"}
+
+ @{rail \<open>()\<close>}
+
+ \item Nonterminal @{verbatim "A"}
+
+ @{rail \<open>A\<close>}
+
+ \item Nonterminal via Isabelle antiquotation
+ @{verbatim "@{syntax method}"}
+
+ @{rail \<open>@{syntax method}\<close>}
+
+ \item Terminal @{verbatim "'xyz'"}
+
+ @{rail \<open>'xyz'\<close>}
+
+ \item Terminal in keyword style @{verbatim "@'xyz'"}
+
+ @{rail \<open>@'xyz'\<close>}
+
+ \item Terminal via Isabelle antiquotation
+ @{verbatim "@@{method rule}"}
+
+ @{rail \<open>@@{method rule}\<close>}
+
+ \item Concatenation @{verbatim "A B C"}
+
+ @{rail \<open>A B C\<close>}
+
+ \item Newline inside concatenation
+ @{verbatim "A B C \<newline> D E F"}
+
+ @{rail \<open>A B C \<newline> D E F\<close>}
+
+ \item Variants @{verbatim "A | B | C"}
+
+ @{rail \<open>A | B | C\<close>}
+
+ \item Option @{verbatim "A ?"}
+
+ @{rail \<open>A ?\<close>}
+
+ \item Repetition @{verbatim "A *"}
+
+ @{rail \<open>A *\<close>}
+
+ \item Repetition with separator @{verbatim "A * sep"}
+
+ @{rail \<open>A * sep\<close>}
+
+ \item Strict repetition @{verbatim "A +"}
+
+ @{rail \<open>A +\<close>}
+
+ \item Strict repetition with separator @{verbatim "A + sep"}
+
+ @{rail \<open>A + sep\<close>}
+
+ \end{itemize}
+*}
+
+
+section {* Draft presentation *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{command_def "display_drafts"}@{text "\<^sup>*"} & : & @{text "any \<rightarrow>"} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{command display_drafts} (@{syntax name} +)
+ \<close>}
+
+ \begin{description}
+
+ \item @{command "display_drafts"}~@{text paths} performs simple output of a
+ given list of raw source files. Only those symbols that do not require
+ additional {\LaTeX} packages are displayed properly, everything else is left
+ verbatim.
+
+ \end{description}
+*}
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Isar-Ref/First_Order_Logic.thy Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,520 @@
+
+header {* Example: First-Order Logic *}
+
+theory %visible First_Order_Logic
+imports Base (* FIXME Pure!? *)
+begin
+
+text {*
+ \noindent In order to commence a new object-logic within
+ Isabelle/Pure we introduce abstract syntactic categories @{text "i"}
+ for individuals and @{text "o"} for object-propositions. The latter
+ is embedded into the language of Pure propositions by means of a
+ separate judgment.
+*}
+
+typedecl i
+typedecl o
+
+judgment
+ Trueprop :: "o \<Rightarrow> prop" ("_" 5)
+
+text {*
+ \noindent Note that the object-logic judgement is implicit in the
+ syntax: writing @{prop A} produces @{term "Trueprop A"} internally.
+ From the Pure perspective this means ``@{prop A} is derivable in the
+ object-logic''.
+*}
+
+
+subsection {* Equational reasoning \label{sec:framework-ex-equal} *}
+
+text {*
+ Equality is axiomatized as a binary predicate on individuals, with
+ reflexivity as introduction, and substitution as elimination
+ principle. Note that the latter is particularly convenient in a
+ framework like Isabelle, because syntactic congruences are
+ implicitly produced by unification of @{term "B x"} against
+ expressions containing occurrences of @{term x}.
+*}
+
+axiomatization
+ equal :: "i \<Rightarrow> i \<Rightarrow> o" (infix "=" 50)
+where
+ refl [intro]: "x = x" and
+ subst [elim]: "x = y \<Longrightarrow> B x \<Longrightarrow> B y"
+
+text {*
+ \noindent Substitution is very powerful, but also hard to control in
+ full generality. We derive some common symmetry~/ transitivity
+ schemes of @{term equal} as particular consequences.
+*}
+
+theorem sym [sym]:
+ assumes "x = y"
+ shows "y = x"
+proof -
+ have "x = x" ..
+ with `x = y` show "y = x" ..
+qed
+
+theorem forw_subst [trans]:
+ assumes "y = x" and "B x"
+ shows "B y"
+proof -
+ from `y = x` have "x = y" ..
+ from this and `B x` show "B y" ..
+qed
+
+theorem back_subst [trans]:
+ assumes "B x" and "x = y"
+ shows "B y"
+proof -
+ from `x = y` and `B x`
+ show "B y" ..
+qed
+
+theorem trans [trans]:
+ assumes "x = y" and "y = z"
+ shows "x = z"
+proof -
+ from `y = z` and `x = y`
+ show "x = z" ..
+qed
+
+
+subsection {* Basic group theory *}
+
+text {*
+ As an example for equational reasoning we consider some bits of
+ group theory. The subsequent locale definition postulates group
+ operations and axioms; we also derive some consequences of this
+ specification.
+*}
+
+locale group =
+ fixes prod :: "i \<Rightarrow> i \<Rightarrow> i" (infix "\<circ>" 70)
+ and inv :: "i \<Rightarrow> i" ("(_\<inverse>)" [1000] 999)
+ and unit :: i ("1")
+ assumes assoc: "(x \<circ> y) \<circ> z = x \<circ> (y \<circ> z)"
+ and left_unit: "1 \<circ> x = x"
+ and left_inv: "x\<inverse> \<circ> x = 1"
+begin
+
+theorem right_inv: "x \<circ> x\<inverse> = 1"
+proof -
+ have "x \<circ> x\<inverse> = 1 \<circ> (x \<circ> x\<inverse>)" by (rule left_unit [symmetric])
+ also have "\<dots> = (1 \<circ> x) \<circ> x\<inverse>" by (rule assoc [symmetric])
+ also have "1 = (x\<inverse>)\<inverse> \<circ> x\<inverse>" by (rule left_inv [symmetric])
+ also have "\<dots> \<circ> x = (x\<inverse>)\<inverse> \<circ> (x\<inverse> \<circ> x)" by (rule assoc)
+ also have "x\<inverse> \<circ> x = 1" by (rule left_inv)
+ also have "((x\<inverse>)\<inverse> \<circ> \<dots>) \<circ> x\<inverse> = (x\<inverse>)\<inverse> \<circ> (1 \<circ> x\<inverse>)" by (rule assoc)
+ also have "1 \<circ> x\<inverse> = x\<inverse>" by (rule left_unit)
+ also have "(x\<inverse>)\<inverse> \<circ> \<dots> = 1" by (rule left_inv)
+ finally show "x \<circ> x\<inverse> = 1" .
+qed
+
+theorem right_unit: "x \<circ> 1 = x"
+proof -
+ have "1 = x\<inverse> \<circ> x" by (rule left_inv [symmetric])
+ also have "x \<circ> \<dots> = (x \<circ> x\<inverse>) \<circ> x" by (rule assoc [symmetric])
+ also have "x \<circ> x\<inverse> = 1" by (rule right_inv)
+ also have "\<dots> \<circ> x = x" by (rule left_unit)
+ finally show "x \<circ> 1 = x" .
+qed
+
+text {*
+ \noindent Reasoning from basic axioms is often tedious. Our proofs
+ work by producing various instances of the given rules (potentially
+ the symmetric form) using the pattern ``@{command have}~@{text
+ eq}~@{command "by"}~@{text "(rule r)"}'' and composing the chain of
+ results via @{command also}/@{command finally}. These steps may
+ involve any of the transitivity rules declared in
+ \secref{sec:framework-ex-equal}, namely @{thm trans} in combining
+ the first two results in @{thm right_inv} and in the final steps of
+ both proofs, @{thm forw_subst} in the first combination of @{thm
+ right_unit}, and @{thm back_subst} in all other calculational steps.
+
+ Occasional substitutions in calculations are adequate, but should
+ not be over-emphasized. The other extreme is to compose a chain by
+ plain transitivity only, with replacements occurring always in
+ topmost position. For example:
+*}
+
+(*<*)
+theorem "\<And>A. PROP A \<Longrightarrow> PROP A"
+proof -
+ assume [symmetric, defn]: "\<And>x y. (x \<equiv> y) \<equiv> Trueprop (x = y)"
+(*>*)
+ have "x \<circ> 1 = x \<circ> (x\<inverse> \<circ> x)" unfolding left_inv ..
+ also have "\<dots> = (x \<circ> x\<inverse>) \<circ> x" unfolding assoc ..
+ also have "\<dots> = 1 \<circ> x" unfolding right_inv ..
+ also have "\<dots> = x" unfolding left_unit ..
+ finally have "x \<circ> 1 = x" .
+(*<*)
+qed
+(*>*)
+
+text {*
+ \noindent Here we have re-used the built-in mechanism for unfolding
+ definitions in order to normalize each equational problem. A more
+ realistic object-logic would include proper setup for the Simplifier
+ (\secref{sec:simplifier}), the main automated tool for equational
+ reasoning in Isabelle. Then ``@{command unfolding}~@{thm
+ left_inv}~@{command ".."}'' would become ``@{command "by"}~@{text
+ "(simp only: left_inv)"}'' etc.
+*}
+
+end
+
+
+subsection {* Propositional logic \label{sec:framework-ex-prop} *}
+
+text {*
+ We axiomatize basic connectives of propositional logic: implication,
+ disjunction, and conjunction. The associated rules are modeled
+ after Gentzen's system of Natural Deduction \cite{Gentzen:1935}.
+*}
+
+axiomatization
+ imp :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<longrightarrow>" 25) where
+ impI [intro]: "(A \<Longrightarrow> B) \<Longrightarrow> A \<longrightarrow> B" and
+ impD [dest]: "(A \<longrightarrow> B) \<Longrightarrow> A \<Longrightarrow> B"
+
+axiomatization
+ disj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<or>" 30) where
+ disjI\<^sub>1 [intro]: "A \<Longrightarrow> A \<or> B" and
+ disjI\<^sub>2 [intro]: "B \<Longrightarrow> A \<or> B" and
+ disjE [elim]: "A \<or> B \<Longrightarrow> (A \<Longrightarrow> C) \<Longrightarrow> (B \<Longrightarrow> C) \<Longrightarrow> C"
+
+axiomatization
+ conj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<and>" 35) where
+ conjI [intro]: "A \<Longrightarrow> B \<Longrightarrow> A \<and> B" and
+ conjD\<^sub>1: "A \<and> B \<Longrightarrow> A" and
+ conjD\<^sub>2: "A \<and> B \<Longrightarrow> B"
+
+text {*
+ \noindent The conjunctive destructions have the disadvantage that
+ decomposing @{prop "A \<and> B"} involves an immediate decision which
+ component should be projected. The more convenient simultaneous
+ elimination @{prop "A \<and> B \<Longrightarrow> (A \<Longrightarrow> B \<Longrightarrow> C) \<Longrightarrow> C"} can be derived as
+ follows:
+*}
+
+theorem conjE [elim]:
+ assumes "A \<and> B"
+ obtains A and B
+proof
+ from `A \<and> B` show A by (rule conjD\<^sub>1)
+ from `A \<and> B` show B by (rule conjD\<^sub>2)
+qed
+
+text {*
+ \noindent Here is an example of swapping conjuncts with a single
+ intermediate elimination step:
+*}
+
+(*<*)
+lemma "\<And>A. PROP A \<Longrightarrow> PROP A"
+proof -
+(*>*)
+ assume "A \<and> B"
+ then obtain B and A ..
+ then have "B \<and> A" ..
+(*<*)
+qed
+(*>*)
+
+text {*
+ \noindent Note that the analogous elimination rule for disjunction
+ ``@{text "\<ASSUMES> A \<or> B \<OBTAINS> A \<BBAR> B"}'' coincides with
+ the original axiomatization of @{thm disjE}.
+
+ \medskip We continue propositional logic by introducing absurdity
+ with its characteristic elimination. Plain truth may then be
+ defined as a proposition that is trivially true.
+*}
+
+axiomatization
+ false :: o ("\<bottom>") where
+ falseE [elim]: "\<bottom> \<Longrightarrow> A"
+
+definition
+ true :: o ("\<top>") where
+ "\<top> \<equiv> \<bottom> \<longrightarrow> \<bottom>"
+
+theorem trueI [intro]: \<top>
+ unfolding true_def ..
+
+text {*
+ \medskip\noindent Now negation represents an implication towards
+ absurdity:
+*}
+
+definition
+ not :: "o \<Rightarrow> o" ("\<not> _" [40] 40) where
+ "\<not> A \<equiv> A \<longrightarrow> \<bottom>"
+
+theorem notI [intro]:
+ assumes "A \<Longrightarrow> \<bottom>"
+ shows "\<not> A"
+unfolding not_def
+proof
+ assume A
+ then show \<bottom> by (rule `A \<Longrightarrow> \<bottom>`)
+qed
+
+theorem notE [elim]:
+ assumes "\<not> A" and A
+ shows B
+proof -
+ from `\<not> A` have "A \<longrightarrow> \<bottom>" unfolding not_def .
+ from `A \<longrightarrow> \<bottom>` and `A` have \<bottom> ..
+ then show B ..
+qed
+
+
+subsection {* Classical logic *}
+
+text {*
+ Subsequently we state the principle of classical contradiction as a
+ local assumption. Thus we refrain from forcing the object-logic
+ into the classical perspective. Within that context, we may derive
+ well-known consequences of the classical principle.
+*}
+
+locale classical =
+ assumes classical: "(\<not> C \<Longrightarrow> C) \<Longrightarrow> C"
+begin
+
+theorem double_negation:
+ assumes "\<not> \<not> C"
+ shows C
+proof (rule classical)
+ assume "\<not> C"
+ with `\<not> \<not> C` show C ..
+qed
+
+theorem tertium_non_datur: "C \<or> \<not> C"
+proof (rule double_negation)
+ show "\<not> \<not> (C \<or> \<not> C)"
+ proof
+ assume "\<not> (C \<or> \<not> C)"
+ have "\<not> C"
+ proof
+ assume C then have "C \<or> \<not> C" ..
+ with `\<not> (C \<or> \<not> C)` show \<bottom> ..
+ qed
+ then have "C \<or> \<not> C" ..
+ with `\<not> (C \<or> \<not> C)` show \<bottom> ..
+ qed
+qed
+
+text {*
+ \noindent These examples illustrate both classical reasoning and
+ non-trivial propositional proofs in general. All three rules
+ characterize classical logic independently, but the original rule is
+ already the most convenient to use, because it leaves the conclusion
+ unchanged. Note that @{prop "(\<not> C \<Longrightarrow> C) \<Longrightarrow> C"} fits again into our
+ format for eliminations, despite the additional twist that the
+ context refers to the main conclusion. So we may write @{thm
+ classical} as the Isar statement ``@{text "\<OBTAINS> \<not> thesis"}''.
+ This also explains nicely how classical reasoning really works:
+ whatever the main @{text thesis} might be, we may always assume its
+ negation!
+*}
+
+end
+
+
+subsection {* Quantifiers \label{sec:framework-ex-quant} *}
+
+text {*
+ Representing quantifiers is easy, thanks to the higher-order nature
+ of the underlying framework. According to the well-known technique
+ introduced by Church \cite{church40}, quantifiers are operators on
+ predicates, which are syntactically represented as @{text "\<lambda>"}-terms
+ of type @{typ "i \<Rightarrow> o"}. Binder notation turns @{text "All (\<lambda>x. B
+ x)"} into @{text "\<forall>x. B x"} etc.
+*}
+
+axiomatization
+ All :: "(i \<Rightarrow> o) \<Rightarrow> o" (binder "\<forall>" 10) where
+ allI [intro]: "(\<And>x. B x) \<Longrightarrow> \<forall>x. B x" and
+ allD [dest]: "(\<forall>x. B x) \<Longrightarrow> B a"
+
+axiomatization
+ Ex :: "(i \<Rightarrow> o) \<Rightarrow> o" (binder "\<exists>" 10) where
+ exI [intro]: "B a \<Longrightarrow> (\<exists>x. B x)" and
+ exE [elim]: "(\<exists>x. B x) \<Longrightarrow> (\<And>x. B x \<Longrightarrow> C) \<Longrightarrow> C"
+
+text {*
+ \noindent The statement of @{thm exE} corresponds to ``@{text
+ "\<ASSUMES> \<exists>x. B x \<OBTAINS> x \<WHERE> B x"}'' in Isar. In the
+ subsequent example we illustrate quantifier reasoning involving all
+ four rules:
+*}
+
+theorem
+ assumes "\<exists>x. \<forall>y. R x y"
+ shows "\<forall>y. \<exists>x. R x y"
+proof -- {* @{text "\<forall>"} introduction *}
+ obtain x where "\<forall>y. R x y" using `\<exists>x. \<forall>y. R x y` .. -- {* @{text "\<exists>"} elimination *}
+ fix y have "R x y" using `\<forall>y. R x y` .. -- {* @{text "\<forall>"} destruction *}
+ then show "\<exists>x. R x y" .. -- {* @{text "\<exists>"} introduction *}
+qed
+
+
+subsection {* Canonical reasoning patterns *}
+
+text {*
+ The main rules of first-order predicate logic from
+ \secref{sec:framework-ex-prop} and \secref{sec:framework-ex-quant}
+ can now be summarized as follows, using the native Isar statement
+ format of \secref{sec:framework-stmt}.
+
+ \medskip
+ \begin{tabular}{l}
+ @{text "impI: \<ASSUMES> A \<Longrightarrow> B \<SHOWS> A \<longrightarrow> B"} \\
+ @{text "impD: \<ASSUMES> A \<longrightarrow> B \<AND> A \<SHOWS> B"} \\[1ex]
+
+ @{text "disjI\<^sub>1: \<ASSUMES> A \<SHOWS> A \<or> B"} \\
+ @{text "disjI\<^sub>2: \<ASSUMES> B \<SHOWS> A \<or> B"} \\
+ @{text "disjE: \<ASSUMES> A \<or> B \<OBTAINS> A \<BBAR> B"} \\[1ex]
+
+ @{text "conjI: \<ASSUMES> A \<AND> B \<SHOWS> A \<and> B"} \\
+ @{text "conjE: \<ASSUMES> A \<and> B \<OBTAINS> A \<AND> B"} \\[1ex]
+
+ @{text "falseE: \<ASSUMES> \<bottom> \<SHOWS> A"} \\
+ @{text "trueI: \<SHOWS> \<top>"} \\[1ex]
+
+ @{text "notI: \<ASSUMES> A \<Longrightarrow> \<bottom> \<SHOWS> \<not> A"} \\
+ @{text "notE: \<ASSUMES> \<not> A \<AND> A \<SHOWS> B"} \\[1ex]
+
+ @{text "allI: \<ASSUMES> \<And>x. B x \<SHOWS> \<forall>x. B x"} \\
+ @{text "allE: \<ASSUMES> \<forall>x. B x \<SHOWS> B a"} \\[1ex]
+
+ @{text "exI: \<ASSUMES> B a \<SHOWS> \<exists>x. B x"} \\
+ @{text "exE: \<ASSUMES> \<exists>x. B x \<OBTAINS> a \<WHERE> B a"}
+ \end{tabular}
+ \medskip
+
+ \noindent This essentially provides a declarative reading of Pure
+ rules as Isar reasoning patterns: the rule statements tells how a
+ canonical proof outline shall look like. Since the above rules have
+ already been declared as @{attribute (Pure) intro}, @{attribute
+ (Pure) elim}, @{attribute (Pure) dest} --- each according to its
+ particular shape --- we can immediately write Isar proof texts as
+ follows:
+*}
+
+(*<*)
+theorem "\<And>A. PROP A \<Longrightarrow> PROP A"
+proof -
+(*>*)
+
+ txt_raw {*\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
+
+ have "A \<longrightarrow> B"
+ proof
+ assume A
+ show B sorry %noproof
+ qed
+
+ txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
+
+ have "A \<longrightarrow> B" and A sorry %noproof
+ then have B ..
+
+ txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
+
+ have A sorry %noproof
+ then have "A \<or> B" ..
+
+ have B sorry %noproof
+ then have "A \<or> B" ..
+
+ txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
+
+ have "A \<or> B" sorry %noproof
+ then have C
+ proof
+ assume A
+ then show C sorry %noproof
+ next
+ assume B
+ then show C sorry %noproof
+ qed
+
+ txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
+
+ have A and B sorry %noproof
+ then have "A \<and> B" ..
+
+ txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
+
+ have "A \<and> B" sorry %noproof
+ then obtain A and B ..
+
+ txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
+
+ have "\<bottom>" sorry %noproof
+ then have A ..
+
+ txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
+
+ have "\<top>" ..
+
+ txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
+
+ have "\<not> A"
+ proof
+ assume A
+ then show "\<bottom>" sorry %noproof
+ qed
+
+ txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
+
+ have "\<not> A" and A sorry %noproof
+ then have B ..
+
+ txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
+
+ have "\<forall>x. B x"
+ proof
+ fix x
+ show "B x" sorry %noproof
+ qed
+
+ txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
+
+ have "\<forall>x. B x" sorry %noproof
+ then have "B a" ..
+
+ txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
+
+ have "\<exists>x. B x"
+ proof
+ show "B a" sorry %noproof
+ qed
+
+ txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*)
+
+ have "\<exists>x. B x" sorry %noproof
+ then obtain a where "B a" ..
+
+ txt_raw {*\end{minipage}*}
+
+(*<*)
+qed
+(*>*)
+
+text {*
+ \bigskip\noindent Of course, these proofs are merely examples. As
+ sketched in \secref{sec:framework-subproof}, there is a fair amount
+ of flexibility in expressing Pure deductions in Isar. Here the user
+ is asked to express himself adequately, aiming at proof texts of
+ literary quality.
+*}
+
+end %visible
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Isar-Ref/Framework.thy Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,1016 @@
+theory Framework
+imports Base Main
+begin
+
+chapter {* The Isabelle/Isar Framework \label{ch:isar-framework} *}
+
+text {*
+ Isabelle/Isar
+ \cite{Wenzel:1999:TPHOL,Wenzel-PhD,Nipkow-TYPES02,Wenzel-Paulson:2006,Wenzel:2006:Festschrift}
+ is intended as a generic framework for developing formal
+ mathematical documents with full proof checking. Definitions and
+ proofs are organized as theories. An assembly of theory sources may
+ be presented as a printed document; see also
+ \chref{ch:document-prep}.
+
+ The main objective of Isar is the design of a human-readable
+ structured proof language, which is called the ``primary proof
+ format'' in Isar terminology. Such a primary proof language is
+ somewhere in the middle between the extremes of primitive proof
+ objects and actual natural language. In this respect, Isar is a bit
+ more formalistic than Mizar
+ \cite{Trybulec:1993:MizarFeatures,Rudnicki:1992:MizarOverview,Wiedijk:1999:Mizar},
+ using logical symbols for certain reasoning schemes where Mizar
+ would prefer English words; see \cite{Wenzel-Wiedijk:2002} for
+ further comparisons of these systems.
+
+ So Isar challenges the traditional way of recording informal proofs
+ in mathematical prose, as well as the common tendency to see fully
+ formal proofs directly as objects of some logical calculus (e.g.\
+ @{text "\<lambda>"}-terms in a version of type theory). In fact, Isar is
+ better understood as an interpreter of a simple block-structured
+ language for describing the data flow of local facts and goals,
+ interspersed with occasional invocations of proof methods.
+ Everything is reduced to logical inferences internally, but these
+ steps are somewhat marginal compared to the overall bookkeeping of
+ the interpretation process. Thanks to careful design of the syntax
+ and semantics of Isar language elements, a formal record of Isar
+ instructions may later appear as an intelligible text to the
+ attentive reader.
+
+ The Isar proof language has emerged from careful analysis of some
+ inherent virtues of the existing logical framework of Isabelle/Pure
+ \cite{paulson-found,paulson700}, notably composition of higher-order
+ natural deduction rules, which is a generalization of Gentzen's
+ original calculus \cite{Gentzen:1935}. The approach of generic
+ inference systems in Pure is continued by Isar towards actual proof
+ texts.
+
+ Concrete applications require another intermediate layer: an
+ object-logic. Isabelle/HOL \cite{isa-tutorial} (simply-typed
+ set-theory) is being used most of the time; Isabelle/ZF
+ \cite{isabelle-ZF} is less extensively developed, although it would
+ probably fit better for classical mathematics.
+
+ \medskip In order to illustrate natural deduction in Isar, we shall
+ refer to the background theory and library of Isabelle/HOL. This
+ includes common notions of predicate logic, naive set-theory etc.\
+ using fairly standard mathematical notation. From the perspective
+ of generic natural deduction there is nothing special about the
+ logical connectives of HOL (@{text "\<and>"}, @{text "\<or>"}, @{text "\<forall>"},
+ @{text "\<exists>"}, etc.), only the resulting reasoning principles are
+ relevant to the user. There are similar rules available for
+ set-theory operators (@{text "\<inter>"}, @{text "\<union>"}, @{text "\<Inter>"}, @{text
+ "\<Union>"}, etc.), or any other theory developed in the library (lattice
+ theory, topology etc.).
+
+ Subsequently we briefly review fragments of Isar proof texts
+ corresponding directly to such general deduction schemes. The
+ examples shall refer to set-theory, to minimize the danger of
+ understanding connectives of predicate logic as something special.
+
+ \medskip The following deduction performs @{text "\<inter>"}-introduction,
+ working forwards from assumptions towards the conclusion. We give
+ both the Isar text, and depict the primitive rule involved, as
+ determined by unification of the problem against rules that are
+ declared in the library context.
+*}
+
+text_raw {*\medskip\begin{minipage}{0.6\textwidth}*}
+
+(*<*)
+notepad
+begin
+(*>*)
+ assume "x \<in> A" and "x \<in> B"
+ then have "x \<in> A \<inter> B" ..
+(*<*)
+end
+(*>*)
+
+text_raw {*\end{minipage}\begin{minipage}{0.4\textwidth}*}
+
+text {*
+ \infer{@{prop "x \<in> A \<inter> B"}}{@{prop "x \<in> A"} & @{prop "x \<in> B"}}
+*}
+
+text_raw {*\end{minipage}*}
+
+text {*
+ \medskip\noindent Note that @{command assume} augments the proof
+ context, @{command then} indicates that the current fact shall be
+ used in the next step, and @{command have} states an intermediate
+ goal. The two dots ``@{command ".."}'' refer to a complete proof of
+ this claim, using the indicated facts and a canonical rule from the
+ context. We could have been more explicit here by spelling out the
+ final proof step via the @{command "by"} command:
+*}
+
+(*<*)
+notepad
+begin
+(*>*)
+ assume "x \<in> A" and "x \<in> B"
+ then have "x \<in> A \<inter> B" by (rule IntI)
+(*<*)
+end
+(*>*)
+
+text {*
+ \noindent The format of the @{text "\<inter>"}-introduction rule represents
+ the most basic inference, which proceeds from given premises to a
+ conclusion, without any nested proof context involved.
+
+ The next example performs backwards introduction on @{term "\<Inter>\<A>"},
+ the intersection of all sets within a given set. This requires a
+ nested proof of set membership within a local context, where @{term
+ A} is an arbitrary-but-fixed member of the collection:
+*}
+
+text_raw {*\medskip\begin{minipage}{0.6\textwidth}*}
+
+(*<*)
+notepad
+begin
+(*>*)
+ have "x \<in> \<Inter>\<A>"
+ proof
+ fix A
+ assume "A \<in> \<A>"
+ show "x \<in> A" sorry %noproof
+ qed
+(*<*)
+end
+(*>*)
+
+text_raw {*\end{minipage}\begin{minipage}{0.4\textwidth}*}
+
+text {*
+ \infer{@{prop "x \<in> \<Inter>\<A>"}}{\infer*{@{prop "x \<in> A"}}{@{text "[A][A \<in> \<A>]"}}}
+*}
+
+text_raw {*\end{minipage}*}
+
+text {*
+ \medskip\noindent This Isar reasoning pattern again refers to the
+ primitive rule depicted above. The system determines it in the
+ ``@{command proof}'' step, which could have been spelt out more
+ explicitly as ``@{command proof}~@{text "(rule InterI)"}''. Note
+ that the rule involves both a local parameter @{term "A"} and an
+ assumption @{prop "A \<in> \<A>"} in the nested reasoning. This kind of
+ compound rule typically demands a genuine sub-proof in Isar, working
+ backwards rather than forwards as seen before. In the proof body we
+ encounter the @{command fix}-@{command assume}-@{command show}
+ outline of nested sub-proofs that is typical for Isar. The final
+ @{command show} is like @{command have} followed by an additional
+ refinement of the enclosing claim, using the rule derived from the
+ proof body.
+
+ \medskip The next example involves @{term "\<Union>\<A>"}, which can be
+ characterized as the set of all @{term "x"} such that @{prop "\<exists>A. x
+ \<in> A \<and> A \<in> \<A>"}. The elimination rule for @{prop "x \<in> \<Union>\<A>"} does
+ not mention @{text "\<exists>"} and @{text "\<and>"} at all, but admits to obtain
+ directly a local @{term "A"} such that @{prop "x \<in> A"} and @{prop "A
+ \<in> \<A>"} hold. This corresponds to the following Isar proof and
+ inference rule, respectively:
+*}
+
+text_raw {*\medskip\begin{minipage}{0.6\textwidth}*}
+
+(*<*)
+notepad
+begin
+(*>*)
+ assume "x \<in> \<Union>\<A>"
+ then have C
+ proof
+ fix A
+ assume "x \<in> A" and "A \<in> \<A>"
+ show C sorry %noproof
+ qed
+(*<*)
+end
+(*>*)
+
+text_raw {*\end{minipage}\begin{minipage}{0.4\textwidth}*}
+
+text {*
+ \infer{@{prop "C"}}{@{prop "x \<in> \<Union>\<A>"} & \infer*{@{prop "C"}~}{@{text "[A][x \<in> A, A \<in> \<A>]"}}}
+*}
+
+text_raw {*\end{minipage}*}
+
+text {*
+ \medskip\noindent Although the Isar proof follows the natural
+ deduction rule closely, the text reads not as natural as
+ anticipated. There is a double occurrence of an arbitrary
+ conclusion @{prop "C"}, which represents the final result, but is
+ irrelevant for now. This issue arises for any elimination rule
+ involving local parameters. Isar provides the derived language
+ element @{command obtain}, which is able to perform the same
+ elimination proof more conveniently:
+*}
+
+(*<*)
+notepad
+begin
+(*>*)
+ assume "x \<in> \<Union>\<A>"
+ then obtain A where "x \<in> A" and "A \<in> \<A>" ..
+(*<*)
+end
+(*>*)
+
+text {*
+ \noindent Here we avoid to mention the final conclusion @{prop "C"}
+ and return to plain forward reasoning. The rule involved in the
+ ``@{command ".."}'' proof is the same as before.
+*}
+
+
+section {* The Pure framework \label{sec:framework-pure} *}
+
+text {*
+ The Pure logic \cite{paulson-found,paulson700} is an intuitionistic
+ fragment of higher-order logic \cite{church40}. In type-theoretic
+ parlance, there are three levels of @{text "\<lambda>"}-calculus with
+ corresponding arrows @{text "\<Rightarrow>"}/@{text "\<And>"}/@{text "\<Longrightarrow>"}:
+
+ \medskip
+ \begin{tabular}{ll}
+ @{text "\<alpha> \<Rightarrow> \<beta>"} & syntactic function space (terms depending on terms) \\
+ @{text "\<And>x. B(x)"} & universal quantification (proofs depending on terms) \\
+ @{text "A \<Longrightarrow> B"} & implication (proofs depending on proofs) \\
+ \end{tabular}
+ \medskip
+
+ \noindent Here only the types of syntactic terms, and the
+ propositions of proof terms have been shown. The @{text
+ "\<lambda>"}-structure of proofs can be recorded as an optional feature of
+ the Pure inference kernel \cite{Berghofer-Nipkow:2000:TPHOL}, but
+ the formal system can never depend on them due to \emph{proof
+ irrelevance}.
+
+ On top of this most primitive layer of proofs, Pure implements a
+ generic calculus for nested natural deduction rules, similar to
+ \cite{Schroeder-Heister:1984}. Here object-logic inferences are
+ internalized as formulae over @{text "\<And>"} and @{text "\<Longrightarrow>"}.
+ Combining such rule statements may involve higher-order unification
+ \cite{paulson-natural}.
+*}
+
+
+subsection {* Primitive inferences *}
+
+text {*
+ Term syntax provides explicit notation for abstraction @{text "\<lambda>x ::
+ \<alpha>. b(x)"} and application @{text "b a"}, while types are usually
+ implicit thanks to type-inference; terms of type @{text "prop"} are
+ called propositions. Logical statements are composed via @{text "\<And>x
+ :: \<alpha>. B(x)"} and @{text "A \<Longrightarrow> B"}. Primitive reasoning operates on
+ judgments of the form @{text "\<Gamma> \<turnstile> \<phi>"}, with standard introduction
+ and elimination rules for @{text "\<And>"} and @{text "\<Longrightarrow>"} that refer to
+ fixed parameters @{text "x\<^sub>1, \<dots>, x\<^sub>m"} and hypotheses
+ @{text "A\<^sub>1, \<dots>, A\<^sub>n"} from the context @{text "\<Gamma>"};
+ the corresponding proof terms are left implicit. The subsequent
+ inference rules define @{text "\<Gamma> \<turnstile> \<phi>"} inductively, relative to a
+ collection of axioms:
+
+ \[
+ \infer{@{text "\<turnstile> A"}}{(@{text "A"} \text{~axiom})}
+ \qquad
+ \infer{@{text "A \<turnstile> A"}}{}
+ \]
+
+ \[
+ \infer{@{text "\<Gamma> \<turnstile> \<And>x. B(x)"}}{@{text "\<Gamma> \<turnstile> B(x)"} & @{text "x \<notin> \<Gamma>"}}
+ \qquad
+ \infer{@{text "\<Gamma> \<turnstile> B(a)"}}{@{text "\<Gamma> \<turnstile> \<And>x. B(x)"}}
+ \]
+
+ \[
+ \infer{@{text "\<Gamma> - A \<turnstile> A \<Longrightarrow> B"}}{@{text "\<Gamma> \<turnstile> B"}}
+ \qquad
+ \infer{@{text "\<Gamma>\<^sub>1 \<union> \<Gamma>\<^sub>2 \<turnstile> B"}}{@{text "\<Gamma>\<^sub>1 \<turnstile> A \<Longrightarrow> B"} & @{text "\<Gamma>\<^sub>2 \<turnstile> A"}}
+ \]
+
+ Furthermore, Pure provides a built-in equality @{text "\<equiv> :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow>
+ prop"} with axioms for reflexivity, substitution, extensionality,
+ and @{text "\<alpha>\<beta>\<eta>"}-conversion on @{text "\<lambda>"}-terms.
+
+ \medskip An object-logic introduces another layer on top of Pure,
+ e.g.\ with types @{text "i"} for individuals and @{text "o"} for
+ propositions, term constants @{text "Trueprop :: o \<Rightarrow> prop"} as
+ (implicit) derivability judgment and connectives like @{text "\<and> :: o
+ \<Rightarrow> o \<Rightarrow> o"} or @{text "\<forall> :: (i \<Rightarrow> o) \<Rightarrow> o"}, and axioms for object-level
+ rules such as @{text "conjI: A \<Longrightarrow> B \<Longrightarrow> A \<and> B"} or @{text "allI: (\<And>x. B
+ x) \<Longrightarrow> \<forall>x. B x"}. Derived object rules are represented as theorems of
+ Pure. After the initial object-logic setup, further axiomatizations
+ are usually avoided; plain definitions and derived principles are
+ used exclusively.
+*}
+
+
+subsection {* Reasoning with rules \label{sec:framework-resolution} *}
+
+text {*
+ Primitive inferences mostly serve foundational purposes. The main
+ reasoning mechanisms of Pure operate on nested natural deduction
+ rules expressed as formulae, using @{text "\<And>"} to bind local
+ parameters and @{text "\<Longrightarrow>"} to express entailment. Multiple
+ parameters and premises are represented by repeating these
+ connectives in a right-associative manner.
+
+ Since @{text "\<And>"} and @{text "\<Longrightarrow>"} commute thanks to the theorem
+ @{prop "(A \<Longrightarrow> (\<And>x. B x)) \<equiv> (\<And>x. A \<Longrightarrow> B x)"}, we may assume w.l.o.g.\
+ that rule statements always observe the normal form where
+ quantifiers are pulled in front of implications at each level of
+ nesting. This means that any Pure proposition may be presented as a
+ \emph{Hereditary Harrop Formula} \cite{Miller:1991} which is of the
+ form @{text "\<And>x\<^sub>1 \<dots> x\<^sub>m. H\<^sub>1 \<Longrightarrow> \<dots> H\<^sub>n \<Longrightarrow>
+ A"} for @{text "m, n \<ge> 0"}, and @{text "A"} atomic, and @{text
+ "H\<^sub>1, \<dots>, H\<^sub>n"} being recursively of the same format.
+ Following the convention that outermost quantifiers are implicit,
+ Horn clauses @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> A"} are a special
+ case of this.
+
+ For example, @{text "\<inter>"}-introduction rule encountered before is
+ represented as a Pure theorem as follows:
+ \[
+ @{text "IntI:"}~@{prop "x \<in> A \<Longrightarrow> x \<in> B \<Longrightarrow> x \<in> A \<inter> B"}
+ \]
+
+ \noindent This is a plain Horn clause, since no further nesting on
+ the left is involved. The general @{text "\<Inter>"}-introduction
+ corresponds to a Hereditary Harrop Formula with one additional level
+ of nesting:
+ \[
+ @{text "InterI:"}~@{prop "(\<And>A. A \<in> \<A> \<Longrightarrow> x \<in> A) \<Longrightarrow> x \<in> \<Inter>\<A>"}
+ \]
+
+ \medskip Goals are also represented as rules: @{text "A\<^sub>1 \<Longrightarrow>
+ \<dots> A\<^sub>n \<Longrightarrow> C"} states that the sub-goals @{text "A\<^sub>1, \<dots>,
+ A\<^sub>n"} entail the result @{text "C"}; for @{text "n = 0"} the
+ goal is finished. To allow @{text "C"} being a rule statement
+ itself, we introduce the protective marker @{text "# :: prop \<Rightarrow>
+ prop"}, which is defined as identity and hidden from the user. We
+ initialize and finish goal states as follows:
+
+ \[
+ \begin{array}{c@ {\qquad}c}
+ \infer[(@{inference_def init})]{@{text "C \<Longrightarrow> #C"}}{} &
+ \infer[(@{inference_def finish})]{@{text C}}{@{text "#C"}}
+ \end{array}
+ \]
+
+ \noindent Goal states are refined in intermediate proof steps until
+ a finished form is achieved. Here the two main reasoning principles
+ are @{inference resolution}, for back-chaining a rule against a
+ sub-goal (replacing it by zero or more sub-goals), and @{inference
+ assumption}, for solving a sub-goal (finding a short-circuit with
+ local assumptions). Below @{text "\<^vec>x"} stands for @{text
+ "x\<^sub>1, \<dots>, x\<^sub>n"} (@{text "n \<ge> 0"}).
+
+ \[
+ \infer[(@{inference_def resolution})]
+ {@{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> \<^vec>A (\<^vec>a \<^vec>x))\<vartheta> \<Longrightarrow> C\<vartheta>"}}
+ {\begin{tabular}{rl}
+ @{text "rule:"} &
+ @{text "\<^vec>A \<^vec>a \<Longrightarrow> B \<^vec>a"} \\
+ @{text "goal:"} &
+ @{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> B' \<^vec>x) \<Longrightarrow> C"} \\
+ @{text "goal unifier:"} &
+ @{text "(\<lambda>\<^vec>x. B (\<^vec>a \<^vec>x))\<vartheta> = B'\<vartheta>"} \\
+ \end{tabular}}
+ \]
+
+ \medskip
+
+ \[
+ \infer[(@{inference_def assumption})]{@{text "C\<vartheta>"}}
+ {\begin{tabular}{rl}
+ @{text "goal:"} &
+ @{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> A \<^vec>x) \<Longrightarrow> C"} \\
+ @{text "assm unifier:"} & @{text "A\<vartheta> = H\<^sub>i\<vartheta>"}~~\text{(for some~@{text "H\<^sub>i"})} \\
+ \end{tabular}}
+ \]
+
+ The following trace illustrates goal-oriented reasoning in
+ Isabelle/Pure:
+
+ {\footnotesize
+ \medskip
+ \begin{tabular}{r@ {\quad}l}
+ @{text "(A \<and> B \<Longrightarrow> B \<and> A) \<Longrightarrow> #(A \<and> B \<Longrightarrow> B \<and> A)"} & @{text "(init)"} \\
+ @{text "(A \<and> B \<Longrightarrow> B) \<Longrightarrow> (A \<and> B \<Longrightarrow> A) \<Longrightarrow> #\<dots>"} & @{text "(resolution B \<Longrightarrow> A \<Longrightarrow> B \<and> A)"} \\
+ @{text "(A \<and> B \<Longrightarrow> A \<and> B) \<Longrightarrow> (A \<and> B \<Longrightarrow> A) \<Longrightarrow> #\<dots>"} & @{text "(resolution A \<and> B \<Longrightarrow> B)"} \\
+ @{text "(A \<and> B \<Longrightarrow> A) \<Longrightarrow> #\<dots>"} & @{text "(assumption)"} \\
+ @{text "(A \<and> B \<Longrightarrow> A \<and> B) \<Longrightarrow> #\<dots>"} & @{text "(resolution A \<and> B \<Longrightarrow> A)"} \\
+ @{text "#\<dots>"} & @{text "(assumption)"} \\
+ @{text "A \<and> B \<Longrightarrow> B \<and> A"} & @{text "(finish)"} \\
+ \end{tabular}
+ \medskip
+ }
+
+ Compositions of @{inference assumption} after @{inference
+ resolution} occurs quite often, typically in elimination steps.
+ Traditional Isabelle tactics accommodate this by a combined
+ @{inference_def elim_resolution} principle. In contrast, Isar uses
+ a slightly more refined combination, where the assumptions to be
+ closed are marked explicitly, using again the protective marker
+ @{text "#"}:
+
+ \[
+ \infer[(@{inference refinement})]
+ {@{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> \<^vec>G' (\<^vec>a \<^vec>x))\<vartheta> \<Longrightarrow> C\<vartheta>"}}
+ {\begin{tabular}{rl}
+ @{text "sub\<hyphen>proof:"} &
+ @{text "\<^vec>G \<^vec>a \<Longrightarrow> B \<^vec>a"} \\
+ @{text "goal:"} &
+ @{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> B' \<^vec>x) \<Longrightarrow> C"} \\
+ @{text "goal unifier:"} &
+ @{text "(\<lambda>\<^vec>x. B (\<^vec>a \<^vec>x))\<vartheta> = B'\<vartheta>"} \\
+ @{text "assm unifiers:"} &
+ @{text "(\<lambda>\<^vec>x. G\<^sub>j (\<^vec>a \<^vec>x))\<vartheta> = #H\<^sub>i\<vartheta>"} \\
+ & \quad (for each marked @{text "G\<^sub>j"} some @{text "#H\<^sub>i"}) \\
+ \end{tabular}}
+ \]
+
+ \noindent Here the @{text "sub\<hyphen>proof"} rule stems from the
+ main @{command fix}-@{command assume}-@{command show} outline of
+ Isar (cf.\ \secref{sec:framework-subproof}): each assumption
+ indicated in the text results in a marked premise @{text "G"} above.
+ The marking enforces resolution against one of the sub-goal's
+ premises. Consequently, @{command fix}-@{command assume}-@{command
+ show} enables to fit the result of a sub-proof quite robustly into a
+ pending sub-goal, while maintaining a good measure of flexibility.
+*}
+
+
+section {* The Isar proof language \label{sec:framework-isar} *}
+
+text {*
+ Structured proofs are presented as high-level expressions for
+ composing entities of Pure (propositions, facts, and goals). The
+ Isar proof language allows to organize reasoning within the
+ underlying rule calculus of Pure, but Isar is not another logical
+ calculus!
+
+ Isar is an exercise in sound minimalism. Approximately half of the
+ language is introduced as primitive, the rest defined as derived
+ concepts. The following grammar describes the core language
+ (category @{text "proof"}), which is embedded into theory
+ specification elements such as @{command theorem}; see also
+ \secref{sec:framework-stmt} for the separate category @{text
+ "statement"}.
+
+ \medskip
+ \begin{tabular}{rcl}
+ @{text "theory\<hyphen>stmt"} & = & @{command "theorem"}~@{text "statement proof |"}~~@{command "definition"}~@{text "\<dots> | \<dots>"} \\[1ex]
+
+ @{text "proof"} & = & @{text "prfx\<^sup>*"}~@{command "proof"}~@{text "method\<^sup>? stmt\<^sup>*"}~@{command "qed"}~@{text "method\<^sup>?"} \\[1ex]
+
+ @{text prfx} & = & @{command "using"}~@{text "facts"} \\
+ & @{text "|"} & @{command "unfolding"}~@{text "facts"} \\
+
+ @{text stmt} & = & @{command "{"}~@{text "stmt\<^sup>*"}~@{command "}"} \\
+ & @{text "|"} & @{command "next"} \\
+ & @{text "|"} & @{command "note"}~@{text "name = facts"} \\
+ & @{text "|"} & @{command "let"}~@{text "term = term"} \\
+ & @{text "|"} & @{command "fix"}~@{text "var\<^sup>+"} \\
+ & @{text "|"} & @{command assume}~@{text "\<guillemotleft>inference\<guillemotright> name: props"} \\
+ & @{text "|"} & @{command "then"}@{text "\<^sup>?"}~@{text goal} \\
+ @{text goal} & = & @{command "have"}~@{text "name: props proof"} \\
+ & @{text "|"} & @{command "show"}~@{text "name: props proof"} \\
+ \end{tabular}
+
+ \medskip Simultaneous propositions or facts may be separated by the
+ @{keyword "and"} keyword.
+
+ \medskip The syntax for terms and propositions is inherited from
+ Pure (and the object-logic). A @{text "pattern"} is a @{text
+ "term"} with schematic variables, to be bound by higher-order
+ matching.
+
+ \medskip Facts may be referenced by name or proposition. For
+ example, the result of ``@{command have}~@{text "a: A \<langle>proof\<rangle>"}''
+ becomes available both as @{text "a"} and
+ \isacharbackquoteopen@{text "A"}\isacharbackquoteclose. Moreover,
+ fact expressions may involve attributes that modify either the
+ theorem or the background context. For example, the expression
+ ``@{text "a [OF b]"}'' refers to the composition of two facts
+ according to the @{inference resolution} inference of
+ \secref{sec:framework-resolution}, while ``@{text "a [intro]"}''
+ declares a fact as introduction rule in the context.
+
+ The special fact called ``@{fact this}'' always refers to the last
+ result, as produced by @{command note}, @{command assume}, @{command
+ have}, or @{command show}. Since @{command note} occurs
+ frequently together with @{command then} we provide some
+ abbreviations:
+
+ \medskip
+ \begin{tabular}{rcl}
+ @{command from}~@{text a} & @{text "\<equiv>"} & @{command note}~@{text a}~@{command then} \\
+ @{command with}~@{text a} & @{text "\<equiv>"} & @{command from}~@{text "a \<AND> this"} \\
+ \end{tabular}
+ \medskip
+
+ The @{text "method"} category is essentially a parameter and may be
+ populated later. Methods use the facts indicated by @{command
+ "then"} or @{command using}, and then operate on the goal state.
+ Some basic methods are predefined: ``@{method "-"}'' leaves the goal
+ unchanged, ``@{method this}'' applies the facts as rules to the
+ goal, ``@{method (Pure) "rule"}'' applies the facts to another rule and the
+ result to the goal (both ``@{method this}'' and ``@{method (Pure) rule}''
+ refer to @{inference resolution} of
+ \secref{sec:framework-resolution}). The secondary arguments to
+ ``@{method (Pure) rule}'' may be specified explicitly as in ``@{text "(rule
+ a)"}'', or picked from the context. In the latter case, the system
+ first tries rules declared as @{attribute (Pure) elim} or
+ @{attribute (Pure) dest}, followed by those declared as @{attribute
+ (Pure) intro}.
+
+ The default method for @{command proof} is ``@{method (Pure) rule}''
+ (arguments picked from the context), for @{command qed} it is
+ ``@{method "-"}''. Further abbreviations for terminal proof steps
+ are ``@{command "by"}~@{text "method\<^sub>1 method\<^sub>2"}'' for
+ ``@{command proof}~@{text "method\<^sub>1"}~@{command qed}~@{text
+ "method\<^sub>2"}'', and ``@{command ".."}'' for ``@{command
+ "by"}~@{method (Pure) rule}, and ``@{command "."}'' for ``@{command
+ "by"}~@{method this}''. The @{command unfolding} element operates
+ directly on the current facts and goal by applying equalities.
+
+ \medskip Block structure can be indicated explicitly by ``@{command
+ "{"}~@{text "\<dots>"}~@{command "}"}'', although the body of a sub-proof
+ already involves implicit nesting. In any case, @{command next}
+ jumps into the next section of a block, i.e.\ it acts like closing
+ an implicit block scope and opening another one; there is no direct
+ correspondence to subgoals here.
+
+ The remaining elements @{command fix} and @{command assume} build up
+ a local context (see \secref{sec:framework-context}), while
+ @{command show} refines a pending sub-goal by the rule resulting
+ from a nested sub-proof (see \secref{sec:framework-subproof}).
+ Further derived concepts will support calculational reasoning (see
+ \secref{sec:framework-calc}).
+*}
+
+
+subsection {* Context elements \label{sec:framework-context} *}
+
+text {*
+ In judgments @{text "\<Gamma> \<turnstile> \<phi>"} of the primitive framework, @{text "\<Gamma>"}
+ essentially acts like a proof context. Isar elaborates this idea
+ towards a higher-level notion, with additional information for
+ type-inference, term abbreviations, local facts, hypotheses etc.
+
+ The element @{command fix}~@{text "x :: \<alpha>"} declares a local
+ parameter, i.e.\ an arbitrary-but-fixed entity of a given type; in
+ results exported from the context, @{text "x"} may become anything.
+ The @{command assume}~@{text "\<guillemotleft>inference\<guillemotright>"} element provides a
+ general interface to hypotheses: ``@{command assume}~@{text
+ "\<guillemotleft>inference\<guillemotright> A"}'' produces @{text "A \<turnstile> A"} locally, while the
+ included inference tells how to discharge @{text A} from results
+ @{text "A \<turnstile> B"} later on. There is no user-syntax for @{text
+ "\<guillemotleft>inference\<guillemotright>"}, i.e.\ it may only occur internally when derived
+ commands are defined in ML.
+
+ At the user-level, the default inference for @{command assume} is
+ @{inference discharge} as given below. The additional variants
+ @{command presume} and @{command def} are defined as follows:
+
+ \medskip
+ \begin{tabular}{rcl}
+ @{command presume}~@{text A} & @{text "\<equiv>"} & @{command assume}~@{text "\<guillemotleft>weak\<hyphen>discharge\<guillemotright> A"} \\
+ @{command def}~@{text "x \<equiv> a"} & @{text "\<equiv>"} & @{command fix}~@{text x}~@{command assume}~@{text "\<guillemotleft>expansion\<guillemotright> x \<equiv> a"} \\
+ \end{tabular}
+ \medskip
+
+ \[
+ \infer[(@{inference_def discharge})]{@{text "\<strut>\<Gamma> - A \<turnstile> #A \<Longrightarrow> B"}}{@{text "\<strut>\<Gamma> \<turnstile> B"}}
+ \]
+ \[
+ \infer[(@{inference_def "weak\<hyphen>discharge"})]{@{text "\<strut>\<Gamma> - A \<turnstile> A \<Longrightarrow> B"}}{@{text "\<strut>\<Gamma> \<turnstile> B"}}
+ \]
+ \[
+ \infer[(@{inference_def expansion})]{@{text "\<strut>\<Gamma> - (x \<equiv> a) \<turnstile> B a"}}{@{text "\<strut>\<Gamma> \<turnstile> B x"}}
+ \]
+
+ \medskip Note that @{inference discharge} and @{inference
+ "weak\<hyphen>discharge"} differ in the marker for @{prop A}, which is
+ relevant when the result of a @{command fix}-@{command
+ assume}-@{command show} outline is composed with a pending goal,
+ cf.\ \secref{sec:framework-subproof}.
+
+ The most interesting derived context element in Isar is @{command
+ obtain} \cite[\S5.3]{Wenzel-PhD}, which supports generalized
+ elimination steps in a purely forward manner. The @{command obtain}
+ command takes a specification of parameters @{text "\<^vec>x"} and
+ assumptions @{text "\<^vec>A"} to be added to the context, together
+ with a proof of a case rule stating that this extension is
+ conservative (i.e.\ may be removed from closed results later on):
+
+ \medskip
+ \begin{tabular}{l}
+ @{text "\<langle>facts\<rangle>"}~~@{command obtain}~@{text "\<^vec>x \<WHERE> \<^vec>A \<^vec>x \<langle>proof\<rangle> \<equiv>"} \\[0.5ex]
+ \quad @{command have}~@{text "case: \<And>thesis. (\<And>\<^vec>x. \<^vec>A \<^vec>x \<Longrightarrow> thesis) \<Longrightarrow> thesis\<rangle>"} \\
+ \quad @{command proof}~@{method "-"} \\
+ \qquad @{command fix}~@{text thesis} \\
+ \qquad @{command assume}~@{text "[intro]: \<And>\<^vec>x. \<^vec>A \<^vec>x \<Longrightarrow> thesis"} \\
+ \qquad @{command show}~@{text thesis}~@{command using}~@{text "\<langle>facts\<rangle> \<langle>proof\<rangle>"} \\
+ \quad @{command qed} \\
+ \quad @{command fix}~@{text "\<^vec>x"}~@{command assume}~@{text "\<guillemotleft>elimination case\<guillemotright> \<^vec>A \<^vec>x"} \\
+ \end{tabular}
+ \medskip
+
+ \[
+ \infer[(@{inference elimination})]{@{text "\<Gamma> \<turnstile> B"}}{
+ \begin{tabular}{rl}
+ @{text "case:"} &
+ @{text "\<Gamma> \<turnstile> \<And>thesis. (\<And>\<^vec>x. \<^vec>A \<^vec>x \<Longrightarrow> thesis) \<Longrightarrow> thesis"} \\[0.2ex]
+ @{text "result:"} &
+ @{text "\<Gamma> \<union> \<^vec>A \<^vec>y \<turnstile> B"} \\[0.2ex]
+ \end{tabular}}
+ \]
+
+ \noindent Here the name ``@{text thesis}'' is a specific convention
+ for an arbitrary-but-fixed proposition; in the primitive natural
+ deduction rules shown before we have occasionally used @{text C}.
+ The whole statement of ``@{command obtain}~@{text x}~@{keyword
+ "where"}~@{text "A x"}'' may be read as a claim that @{text "A x"}
+ may be assumed for some arbitrary-but-fixed @{text "x"}. Also note
+ that ``@{command obtain}~@{text "A \<AND> B"}'' without parameters
+ is similar to ``@{command have}~@{text "A \<AND> B"}'', but the
+ latter involves multiple sub-goals.
+
+ \medskip The subsequent Isar proof texts explain all context
+ elements introduced above using the formal proof language itself.
+ After finishing a local proof within a block, we indicate the
+ exported result via @{command note}.
+*}
+
+(*<*)
+theorem True
+proof
+(*>*)
+ txt_raw {* \begin{minipage}[t]{0.45\textwidth} *}
+ {
+ fix x
+ have "B x" sorry %noproof
+ }
+ note `\<And>x. B x`
+ txt_raw {* \end{minipage}\quad\begin{minipage}[t]{0.45\textwidth} *}(*<*)next(*>*)
+ {
+ assume A
+ have B sorry %noproof
+ }
+ note `A \<Longrightarrow> B`
+ txt_raw {* \end{minipage}\\[3ex]\begin{minipage}[t]{0.45\textwidth} *}(*<*)next(*>*)
+ {
+ def x \<equiv> a
+ have "B x" sorry %noproof
+ }
+ note `B a`
+ txt_raw {* \end{minipage}\quad\begin{minipage}[t]{0.45\textwidth} *}(*<*)next(*>*)
+ {
+ obtain x where "A x" sorry %noproof
+ have B sorry %noproof
+ }
+ note `B`
+ txt_raw {* \end{minipage} *}
+(*<*)
+qed
+(*>*)
+
+text {*
+ \bigskip\noindent This illustrates the meaning of Isar context
+ elements without goals getting in between.
+*}
+
+subsection {* Structured statements \label{sec:framework-stmt} *}
+
+text {*
+ The category @{text "statement"} of top-level theorem specifications
+ is defined as follows:
+
+ \medskip
+ \begin{tabular}{rcl}
+ @{text "statement"} & @{text "\<equiv>"} & @{text "name: props \<AND> \<dots>"} \\
+ & @{text "|"} & @{text "context\<^sup>* conclusion"} \\[0.5ex]
+
+ @{text "context"} & @{text "\<equiv>"} & @{text "\<FIXES> vars \<AND> \<dots>"} \\
+ & @{text "|"} & @{text "\<ASSUMES> name: props \<AND> \<dots>"} \\
+
+ @{text "conclusion"} & @{text "\<equiv>"} & @{text "\<SHOWS> name: props \<AND> \<dots>"} \\
+ & @{text "|"} & @{text "\<OBTAINS> vars \<AND> \<dots> \<WHERE> name: props \<AND> \<dots>"} \\
+ & & \quad @{text "\<BBAR> \<dots>"} \\
+ \end{tabular}
+
+ \medskip\noindent A simple @{text "statement"} consists of named
+ propositions. The full form admits local context elements followed
+ by the actual conclusions, such as ``@{keyword "fixes"}~@{text
+ x}~@{keyword "assumes"}~@{text "A x"}~@{keyword "shows"}~@{text "B
+ x"}''. The final result emerges as a Pure rule after discharging
+ the context: @{prop "\<And>x. A x \<Longrightarrow> B x"}.
+
+ The @{keyword "obtains"} variant is another abbreviation defined
+ below; unlike @{command obtain} (cf.\
+ \secref{sec:framework-context}) there may be several ``cases''
+ separated by ``@{text "\<BBAR>"}'', each consisting of several
+ parameters (@{text "vars"}) and several premises (@{text "props"}).
+ This specifies multi-branch elimination rules.
+
+ \medskip
+ \begin{tabular}{l}
+ @{text "\<OBTAINS> \<^vec>x \<WHERE> \<^vec>A \<^vec>x \<BBAR> \<dots> \<equiv>"} \\[0.5ex]
+ \quad @{text "\<FIXES> thesis"} \\
+ \quad @{text "\<ASSUMES> [intro]: \<And>\<^vec>x. \<^vec>A \<^vec>x \<Longrightarrow> thesis \<AND> \<dots>"} \\
+ \quad @{text "\<SHOWS> thesis"} \\
+ \end{tabular}
+ \medskip
+
+ Presenting structured statements in such an ``open'' format usually
+ simplifies the subsequent proof, because the outer structure of the
+ problem is already laid out directly. E.g.\ consider the following
+ canonical patterns for @{text "\<SHOWS>"} and @{text "\<OBTAINS>"},
+ respectively:
+*}
+
+text_raw {*\begin{minipage}{0.5\textwidth}*}
+
+theorem
+ fixes x and y
+ assumes "A x" and "B y"
+ shows "C x y"
+proof -
+ from `A x` and `B y`
+ show "C x y" sorry %noproof
+qed
+
+text_raw {*\end{minipage}\begin{minipage}{0.5\textwidth}*}
+
+theorem
+ obtains x and y
+ where "A x" and "B y"
+proof -
+ have "A a" and "B b" sorry %noproof
+ then show thesis ..
+qed
+
+text_raw {*\end{minipage}*}
+
+text {*
+ \medskip\noindent Here local facts \isacharbackquoteopen@{text "A
+ x"}\isacharbackquoteclose\ and \isacharbackquoteopen@{text "B
+ y"}\isacharbackquoteclose\ are referenced immediately; there is no
+ need to decompose the logical rule structure again. In the second
+ proof the final ``@{command then}~@{command show}~@{text
+ thesis}~@{command ".."}'' involves the local rule case @{text "\<And>x
+ y. A x \<Longrightarrow> B y \<Longrightarrow> thesis"} for the particular instance of terms @{text
+ "a"} and @{text "b"} produced in the body.
+*}
+
+
+subsection {* Structured proof refinement \label{sec:framework-subproof} *}
+
+text {*
+ By breaking up the grammar for the Isar proof language, we may
+ understand a proof text as a linear sequence of individual proof
+ commands. These are interpreted as transitions of the Isar virtual
+ machine (Isar/VM), which operates on a block-structured
+ configuration in single steps. This allows users to write proof
+ texts in an incremental manner, and inspect intermediate
+ configurations for debugging.
+
+ The basic idea is analogous to evaluating algebraic expressions on a
+ stack machine: @{text "(a + b) \<cdot> c"} then corresponds to a sequence
+ of single transitions for each symbol @{text "(, a, +, b, ), \<cdot>, c"}.
+ In Isar the algebraic values are facts or goals, and the operations
+ are inferences.
+
+ \medskip The Isar/VM state maintains a stack of nodes, each node
+ contains the local proof context, the linguistic mode, and a pending
+ goal (optional). The mode determines the type of transition that
+ may be performed next, it essentially alternates between forward and
+ backward reasoning, with an intermediate stage for chained facts
+ (see \figref{fig:isar-vm}).
+
+ \begin{figure}[htb]
+ \begin{center}
+ \includegraphics[width=0.8\textwidth]{isar-vm}
+ \end{center}
+ \caption{Isar/VM modes}\label{fig:isar-vm}
+ \end{figure}
+
+ For example, in @{text "state"} mode Isar acts like a mathematical
+ scratch-pad, accepting declarations like @{command fix}, @{command
+ assume}, and claims like @{command have}, @{command show}. A goal
+ statement changes the mode to @{text "prove"}, which means that we
+ may now refine the problem via @{command unfolding} or @{command
+ proof}. Then we are again in @{text "state"} mode of a proof body,
+ which may issue @{command show} statements to solve pending
+ sub-goals. A concluding @{command qed} will return to the original
+ @{text "state"} mode one level upwards. The subsequent Isar/VM
+ trace indicates block structure, linguistic mode, goal state, and
+ inferences:
+*}
+
+text_raw {* \begingroup\footnotesize *}
+(*<*)notepad begin
+(*>*)
+ txt_raw {* \begin{minipage}[t]{0.18\textwidth} *}
+ have "A \<longrightarrow> B"
+ proof
+ assume A
+ show B
+ sorry %noproof
+ qed
+ txt_raw {* \end{minipage}\quad
+\begin{minipage}[t]{0.06\textwidth}
+@{text "begin"} \\
+\\
+\\
+@{text "begin"} \\
+@{text "end"} \\
+@{text "end"} \\
+\end{minipage}
+\begin{minipage}[t]{0.08\textwidth}
+@{text "prove"} \\
+@{text "state"} \\
+@{text "state"} \\
+@{text "prove"} \\
+@{text "state"} \\
+@{text "state"} \\
+\end{minipage}\begin{minipage}[t]{0.35\textwidth}
+@{text "(A \<longrightarrow> B) \<Longrightarrow> #(A \<longrightarrow> B)"} \\
+@{text "(A \<Longrightarrow> B) \<Longrightarrow> #(A \<longrightarrow> B)"} \\
+\\
+\\
+@{text "#(A \<longrightarrow> B)"} \\
+@{text "A \<longrightarrow> B"} \\
+\end{minipage}\begin{minipage}[t]{0.4\textwidth}
+@{text "(init)"} \\
+@{text "(resolution impI)"} \\
+\\
+\\
+@{text "(refinement #A \<Longrightarrow> B)"} \\
+@{text "(finish)"} \\
+\end{minipage} *}
+(*<*)
+end
+(*>*)
+text_raw {* \endgroup *}
+
+text {*
+ \noindent Here the @{inference refinement} inference from
+ \secref{sec:framework-resolution} mediates composition of Isar
+ sub-proofs nicely. Observe that this principle incorporates some
+ degree of freedom in proof composition. In particular, the proof
+ body allows parameters and assumptions to be re-ordered, or commuted
+ according to Hereditary Harrop Form. Moreover, context elements
+ that are not used in a sub-proof may be omitted altogether. For
+ example:
+*}
+
+text_raw {*\begin{minipage}{0.5\textwidth}*}
+
+(*<*)
+notepad
+begin
+(*>*)
+ have "\<And>x y. A x \<Longrightarrow> B y \<Longrightarrow> C x y"
+ proof -
+ fix x and y
+ assume "A x" and "B y"
+ show "C x y" sorry %noproof
+ qed
+
+txt_raw {*\end{minipage}\begin{minipage}{0.5\textwidth}*}
+
+(*<*)
+next
+(*>*)
+ have "\<And>x y. A x \<Longrightarrow> B y \<Longrightarrow> C x y"
+ proof -
+ fix x assume "A x"
+ fix y assume "B y"
+ show "C x y" sorry %noproof
+ qed
+
+txt_raw {*\end{minipage}\\[3ex]\begin{minipage}{0.5\textwidth}*}
+
+(*<*)
+next
+(*>*)
+ have "\<And>x y. A x \<Longrightarrow> B y \<Longrightarrow> C x y"
+ proof -
+ fix y assume "B y"
+ fix x assume "A x"
+ show "C x y" sorry
+ qed
+
+txt_raw {*\end{minipage}\begin{minipage}{0.5\textwidth}*}
+(*<*)
+next
+(*>*)
+ have "\<And>x y. A x \<Longrightarrow> B y \<Longrightarrow> C x y"
+ proof -
+ fix y assume "B y"
+ fix x
+ show "C x y" sorry
+ qed
+(*<*)
+end
+(*>*)
+
+text_raw {*\end{minipage}*}
+
+text {*
+ \medskip\noindent Such ``peephole optimizations'' of Isar texts are
+ practically important to improve readability, by rearranging
+ contexts elements according to the natural flow of reasoning in the
+ body, while still observing the overall scoping rules.
+
+ \medskip This illustrates the basic idea of structured proof
+ processing in Isar. The main mechanisms are based on natural
+ deduction rule composition within the Pure framework. In
+ particular, there are no direct operations on goal states within the
+ proof body. Moreover, there is no hidden automated reasoning
+ involved, just plain unification.
+*}
+
+
+subsection {* Calculational reasoning \label{sec:framework-calc} *}
+
+text {*
+ The existing Isar infrastructure is sufficiently flexible to support
+ calculational reasoning (chains of transitivity steps) as derived
+ concept. The generic proof elements introduced below depend on
+ rules declared as @{attribute trans} in the context. It is left to
+ the object-logic to provide a suitable rule collection for mixed
+ relations of @{text "="}, @{text "<"}, @{text "\<le>"}, @{text "\<subset>"},
+ @{text "\<subseteq>"} etc. Due to the flexibility of rule composition
+ (\secref{sec:framework-resolution}), substitution of equals by
+ equals is covered as well, even substitution of inequalities
+ involving monotonicity conditions; see also \cite[\S6]{Wenzel-PhD}
+ and \cite{Bauer-Wenzel:2001}.
+
+ The generic calculational mechanism is based on the observation that
+ rules such as @{text "trans:"}~@{prop "x = y \<Longrightarrow> y = z \<Longrightarrow> x = z"}
+ proceed from the premises towards the conclusion in a deterministic
+ fashion. Thus we may reason in forward mode, feeding intermediate
+ results into rules selected from the context. The course of
+ reasoning is organized by maintaining a secondary fact called
+ ``@{fact calculation}'', apart from the primary ``@{fact this}''
+ already provided by the Isar primitives. In the definitions below,
+ @{attribute OF} refers to @{inference resolution}
+ (\secref{sec:framework-resolution}) with multiple rule arguments,
+ and @{text "trans"} represents to a suitable rule from the context:
+
+ \begin{matharray}{rcl}
+ @{command "also"}@{text "\<^sub>0"} & \equiv & @{command "note"}~@{text "calculation = this"} \\
+ @{command "also"}@{text "\<^sub>n\<^sub>+\<^sub>1"} & \equiv & @{command "note"}~@{text "calculation = trans [OF calculation this]"} \\[0.5ex]
+ @{command "finally"} & \equiv & @{command "also"}~@{command "from"}~@{text calculation} \\
+ \end{matharray}
+
+ \noindent The start of a calculation is determined implicitly in the
+ text: here @{command also} sets @{fact calculation} to the current
+ result; any subsequent occurrence will update @{fact calculation} by
+ combination with the next result and a transitivity rule. The
+ calculational sequence is concluded via @{command finally}, where
+ the final result is exposed for use in a concluding claim.
+
+ Here is a canonical proof pattern, using @{command have} to
+ establish the intermediate results:
+*}
+
+(*<*)
+notepad
+begin
+(*>*)
+ have "a = b" sorry
+ also have "\<dots> = c" sorry
+ also have "\<dots> = d" sorry
+ finally have "a = d" .
+(*<*)
+end
+(*>*)
+
+text {*
+ \noindent The term ``@{text "\<dots>"}'' above is a special abbreviation
+ provided by the Isabelle/Isar syntax layer: it statically refers to
+ the right-hand side argument of the previous statement given in the
+ text. Thus it happens to coincide with relevant sub-expressions in
+ the calculational chain, but the exact correspondence is dependent
+ on the transitivity rules being involved.
+
+ \medskip Symmetry rules such as @{prop "x = y \<Longrightarrow> y = x"} are like
+ transitivities with only one premise. Isar maintains a separate
+ rule collection declared via the @{attribute sym} attribute, to be
+ used in fact expressions ``@{text "a [symmetric]"}'', or single-step
+ proofs ``@{command assume}~@{text "x = y"}~@{command then}~@{command
+ have}~@{text "y = x"}~@{command ".."}''.
+*}
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Isar-Ref/Generic.thy Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,2016 @@
+theory Generic
+imports Base Main
+begin
+
+chapter {* Generic tools and packages \label{ch:gen-tools} *}
+
+section {* Configuration options \label{sec:config} *}
+
+text {* Isabelle/Pure maintains a record of named configuration
+ options within the theory or proof context, with values of type
+ @{ML_type bool}, @{ML_type int}, @{ML_type real}, or @{ML_type
+ string}. Tools may declare options in ML, and then refer to these
+ values (relative to the context). Thus global reference variables
+ are easily avoided. The user may change the value of a
+ configuration option by means of an associated attribute of the same
+ name. This form of context declaration works particularly well with
+ commands such as @{command "declare"} or @{command "using"} like
+ this:
+*}
+
+declare [[show_main_goal = false]]
+
+notepad
+begin
+ note [[show_main_goal = true]]
+end
+
+text {* For historical reasons, some tools cannot take the full proof
+ context into account and merely refer to the background theory.
+ This is accommodated by configuration options being declared as
+ ``global'', which may not be changed within a local context.
+
+ \begin{matharray}{rcll}
+ @{command_def "print_options"} & : & @{text "context \<rightarrow>"} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @{syntax name} ('=' ('true' | 'false' | @{syntax int} | @{syntax float} | @{syntax name}))?
+ \<close>}
+
+ \begin{description}
+
+ \item @{command "print_options"} prints the available configuration
+ options, with names, types, and current values.
+
+ \item @{text "name = value"} as an attribute expression modifies the
+ named option, with the syntax of the value depending on the option's
+ type. For @{ML_type bool} the default value is @{text true}. Any
+ attempt to change a global option in a local context is ignored.
+
+ \end{description}
+*}
+
+
+section {* Basic proof tools *}
+
+subsection {* Miscellaneous methods and attributes \label{sec:misc-meth-att} *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{method_def unfold} & : & @{text method} \\
+ @{method_def fold} & : & @{text method} \\
+ @{method_def insert} & : & @{text method} \\[0.5ex]
+ @{method_def erule}@{text "\<^sup>*"} & : & @{text method} \\
+ @{method_def drule}@{text "\<^sup>*"} & : & @{text method} \\
+ @{method_def frule}@{text "\<^sup>*"} & : & @{text method} \\
+ @{method_def intro} & : & @{text method} \\
+ @{method_def elim} & : & @{text method} \\
+ @{method_def succeed} & : & @{text method} \\
+ @{method_def fail} & : & @{text method} \\
+ \end{matharray}
+
+ @{rail \<open>
+ (@@{method fold} | @@{method unfold} | @@{method insert}) @{syntax thmrefs}
+ ;
+ (@@{method erule} | @@{method drule} | @@{method frule})
+ ('(' @{syntax nat} ')')? @{syntax thmrefs}
+ ;
+ (@@{method intro} | @@{method elim}) @{syntax thmrefs}?
+ \<close>}
+
+ \begin{description}
+
+ \item @{method unfold}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} and @{method fold}~@{text
+ "a\<^sub>1 \<dots> a\<^sub>n"} expand (or fold back) the given definitions throughout
+ all goals; any chained facts provided are inserted into the goal and
+ subject to rewriting as well.
+
+ \item @{method insert}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} inserts theorems as facts
+ into all goals of the proof state. Note that current facts
+ indicated for forward chaining are ignored.
+
+ \item @{method erule}~@{text "a\<^sub>1 \<dots> a\<^sub>n"}, @{method
+ drule}~@{text "a\<^sub>1 \<dots> a\<^sub>n"}, and @{method frule}~@{text
+ "a\<^sub>1 \<dots> a\<^sub>n"} are similar to the basic @{method rule}
+ method (see \secref{sec:pure-meth-att}), but apply rules by
+ elim-resolution, destruct-resolution, and forward-resolution,
+ respectively \cite{isabelle-implementation}. The optional natural
+ number argument (default 0) specifies additional assumption steps to
+ be performed here.
+
+ Note that these methods are improper ones, mainly serving for
+ experimentation and tactic script emulation. Different modes of
+ basic rule application are usually expressed in Isar at the proof
+ language level, rather than via implicit proof state manipulations.
+ For example, a proper single-step elimination would be done using
+ the plain @{method rule} method, with forward chaining of current
+ facts.
+
+ \item @{method intro} and @{method elim} repeatedly refine some goal
+ by intro- or elim-resolution, after having inserted any chained
+ facts. Exactly the rules given as arguments are taken into account;
+ this allows fine-tuned decomposition of a proof problem, in contrast
+ to common automated tools.
+
+ \item @{method succeed} yields a single (unchanged) result; it is
+ the identity of the ``@{text ","}'' method combinator (cf.\
+ \secref{sec:proof-meth}).
+
+ \item @{method fail} yields an empty result sequence; it is the
+ identity of the ``@{text "|"}'' method combinator (cf.\
+ \secref{sec:proof-meth}).
+
+ \end{description}
+
+ \begin{matharray}{rcl}
+ @{attribute_def tagged} & : & @{text attribute} \\
+ @{attribute_def untagged} & : & @{text attribute} \\[0.5ex]
+ @{attribute_def THEN} & : & @{text attribute} \\
+ @{attribute_def unfolded} & : & @{text attribute} \\
+ @{attribute_def folded} & : & @{text attribute} \\
+ @{attribute_def abs_def} & : & @{text attribute} \\[0.5ex]
+ @{attribute_def rotated} & : & @{text attribute} \\
+ @{attribute_def (Pure) elim_format} & : & @{text attribute} \\
+ @{attribute_def no_vars}@{text "\<^sup>*"} & : & @{text attribute} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{attribute tagged} @{syntax name} @{syntax name}
+ ;
+ @@{attribute untagged} @{syntax name}
+ ;
+ @@{attribute THEN} ('[' @{syntax nat} ']')? @{syntax thmref}
+ ;
+ (@@{attribute unfolded} | @@{attribute folded}) @{syntax thmrefs}
+ ;
+ @@{attribute rotated} @{syntax int}?
+ \<close>}
+
+ \begin{description}
+
+ \item @{attribute tagged}~@{text "name value"} and @{attribute
+ untagged}~@{text name} add and remove \emph{tags} of some theorem.
+ Tags may be any list of string pairs that serve as formal comment.
+ The first string is considered the tag name, the second its value.
+ Note that @{attribute untagged} removes any tags of the same name.
+
+ \item @{attribute THEN}~@{text a} composes rules by resolution; it
+ resolves with the first premise of @{text a} (an alternative
+ position may be also specified). See also @{ML_op "RS"} in
+ \cite{isabelle-implementation}.
+
+ \item @{attribute unfolded}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} and @{attribute
+ folded}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} expand and fold back again the given
+ definitions throughout a rule.
+
+ \item @{attribute abs_def} turns an equation of the form @{prop "f x
+ y \<equiv> t"} into @{prop "f \<equiv> \<lambda>x y. t"}, which ensures that @{method
+ simp} or @{method unfold} steps always expand it. This also works
+ for object-logic equality.
+
+ \item @{attribute rotated}~@{text n} rotate the premises of a
+ theorem by @{text n} (default 1).
+
+ \item @{attribute (Pure) elim_format} turns a destruction rule into
+ elimination rule format, by resolving with the rule @{prop "PROP A \<Longrightarrow>
+ (PROP A \<Longrightarrow> PROP B) \<Longrightarrow> PROP B"}.
+
+ Note that the Classical Reasoner (\secref{sec:classical}) provides
+ its own version of this operation.
+
+ \item @{attribute no_vars} replaces schematic variables by free
+ ones; this is mainly for tuning output of pretty printed theorems.
+
+ \end{description}
+*}
+
+
+subsection {* Low-level equational reasoning *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{method_def subst} & : & @{text method} \\
+ @{method_def hypsubst} & : & @{text method} \\
+ @{method_def split} & : & @{text method} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{method subst} ('(' 'asm' ')')? \<newline> ('(' (@{syntax nat}+) ')')? @{syntax thmref}
+ ;
+ @@{method split} @{syntax thmrefs}
+ \<close>}
+
+ These methods provide low-level facilities for equational reasoning
+ that are intended for specialized applications only. Normally,
+ single step calculations would be performed in a structured text
+ (see also \secref{sec:calculation}), while the Simplifier methods
+ provide the canonical way for automated normalization (see
+ \secref{sec:simplifier}).
+
+ \begin{description}
+
+ \item @{method subst}~@{text eq} performs a single substitution step
+ using rule @{text eq}, which may be either a meta or object
+ equality.
+
+ \item @{method subst}~@{text "(asm) eq"} substitutes in an
+ assumption.
+
+ \item @{method subst}~@{text "(i \<dots> j) eq"} performs several
+ substitutions in the conclusion. The numbers @{text i} to @{text j}
+ indicate the positions to substitute at. Positions are ordered from
+ the top of the term tree moving down from left to right. For
+ example, in @{text "(a + b) + (c + d)"} there are three positions
+ where commutativity of @{text "+"} is applicable: 1 refers to @{text
+ "a + b"}, 2 to the whole term, and 3 to @{text "c + d"}.
+
+ If the positions in the list @{text "(i \<dots> j)"} are non-overlapping
+ (e.g.\ @{text "(2 3)"} in @{text "(a + b) + (c + d)"}) you may
+ assume all substitutions are performed simultaneously. Otherwise
+ the behaviour of @{text subst} is not specified.
+
+ \item @{method subst}~@{text "(asm) (i \<dots> j) eq"} performs the
+ substitutions in the assumptions. The positions refer to the
+ assumptions in order from left to right. For example, given in a
+ goal of the form @{text "P (a + b) \<Longrightarrow> P (c + d) \<Longrightarrow> \<dots>"}, position 1 of
+ commutativity of @{text "+"} is the subterm @{text "a + b"} and
+ position 2 is the subterm @{text "c + d"}.
+
+ \item @{method hypsubst} performs substitution using some
+ assumption; this only works for equations of the form @{text "x =
+ t"} where @{text x} is a free or bound variable.
+
+ \item @{method split}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} performs single-step case
+ splitting using the given rules. Splitting is performed in the
+ conclusion or some assumption of the subgoal, depending of the
+ structure of the rule.
+
+ Note that the @{method simp} method already involves repeated
+ application of split rules as declared in the current context, using
+ @{attribute split}, for example.
+
+ \end{description}
+*}
+
+
+subsection {* Further tactic emulations \label{sec:tactics} *}
+
+text {*
+ The following improper proof methods emulate traditional tactics.
+ These admit direct access to the goal state, which is normally
+ considered harmful! In particular, this may involve both numbered
+ goal addressing (default 1), and dynamic instantiation within the
+ scope of some subgoal.
+
+ \begin{warn}
+ Dynamic instantiations refer to universally quantified parameters
+ of a subgoal (the dynamic context) rather than fixed variables and
+ term abbreviations of a (static) Isar context.
+ \end{warn}
+
+ Tactic emulation methods, unlike their ML counterparts, admit
+ simultaneous instantiation from both dynamic and static contexts.
+ If names occur in both contexts goal parameters hide locally fixed
+ variables. Likewise, schematic variables refer to term
+ abbreviations, if present in the static context. Otherwise the
+ schematic variable is interpreted as a schematic variable and left
+ to be solved by unification with certain parts of the subgoal.
+
+ Note that the tactic emulation proof methods in Isabelle/Isar are
+ consistently named @{text foo_tac}. Note also that variable names
+ occurring on left hand sides of instantiations must be preceded by a
+ question mark if they coincide with a keyword or contain dots. This
+ is consistent with the attribute @{attribute "where"} (see
+ \secref{sec:pure-meth-att}).
+
+ \begin{matharray}{rcl}
+ @{method_def rule_tac}@{text "\<^sup>*"} & : & @{text method} \\
+ @{method_def erule_tac}@{text "\<^sup>*"} & : & @{text method} \\
+ @{method_def drule_tac}@{text "\<^sup>*"} & : & @{text method} \\
+ @{method_def frule_tac}@{text "\<^sup>*"} & : & @{text method} \\
+ @{method_def cut_tac}@{text "\<^sup>*"} & : & @{text method} \\
+ @{method_def thin_tac}@{text "\<^sup>*"} & : & @{text method} \\
+ @{method_def subgoal_tac}@{text "\<^sup>*"} & : & @{text method} \\
+ @{method_def rename_tac}@{text "\<^sup>*"} & : & @{text method} \\
+ @{method_def rotate_tac}@{text "\<^sup>*"} & : & @{text method} \\
+ @{method_def tactic}@{text "\<^sup>*"} & : & @{text method} \\
+ @{method_def raw_tactic}@{text "\<^sup>*"} & : & @{text method} \\
+ \end{matharray}
+
+ @{rail \<open>
+ (@@{method rule_tac} | @@{method erule_tac} | @@{method drule_tac} |
+ @@{method frule_tac} | @@{method cut_tac} | @@{method thin_tac}) @{syntax goal_spec}? \<newline>
+ ( dynamic_insts @'in' @{syntax thmref} | @{syntax thmrefs} )
+ ;
+ @@{method subgoal_tac} @{syntax goal_spec}? (@{syntax prop} +)
+ ;
+ @@{method rename_tac} @{syntax goal_spec}? (@{syntax name} +)
+ ;
+ @@{method rotate_tac} @{syntax goal_spec}? @{syntax int}?
+ ;
+ (@@{method tactic} | @@{method raw_tactic}) @{syntax text}
+ ;
+
+ dynamic_insts: ((@{syntax name} '=' @{syntax term}) + @'and')
+ \<close>}
+
+\begin{description}
+
+ \item @{method rule_tac} etc. do resolution of rules with explicit
+ instantiation. This works the same way as the ML tactics @{ML
+ res_inst_tac} etc. (see \cite{isabelle-implementation})
+
+ Multiple rules may be only given if there is no instantiation; then
+ @{method rule_tac} is the same as @{ML resolve_tac} in ML (see
+ \cite{isabelle-implementation}).
+
+ \item @{method cut_tac} inserts facts into the proof state as
+ assumption of a subgoal; instantiations may be given as well. Note
+ that the scope of schematic variables is spread over the main goal
+ statement and rule premises are turned into new subgoals. This is
+ in contrast to the regular method @{method insert} which inserts
+ closed rule statements.
+
+ \item @{method thin_tac}~@{text \<phi>} deletes the specified premise
+ from a subgoal. Note that @{text \<phi>} may contain schematic
+ variables, to abbreviate the intended proposition; the first
+ matching subgoal premise will be deleted. Removing useless premises
+ from a subgoal increases its readability and can make search tactics
+ run faster.
+
+ \item @{method subgoal_tac}~@{text "\<phi>\<^sub>1 \<dots> \<phi>\<^sub>n"} adds the propositions
+ @{text "\<phi>\<^sub>1 \<dots> \<phi>\<^sub>n"} as local premises to a subgoal, and poses the same
+ as new subgoals (in the original context).
+
+ \item @{method rename_tac}~@{text "x\<^sub>1 \<dots> x\<^sub>n"} renames parameters of a
+ goal according to the list @{text "x\<^sub>1, \<dots>, x\<^sub>n"}, which refers to the
+ \emph{suffix} of variables.
+
+ \item @{method rotate_tac}~@{text n} rotates the premises of a
+ subgoal by @{text n} positions: from right to left if @{text n} is
+ positive, and from left to right if @{text n} is negative; the
+ default value is 1.
+
+ \item @{method tactic}~@{text "text"} produces a proof method from
+ any ML text of type @{ML_type tactic}. Apart from the usual ML
+ environment and the current proof context, the ML code may refer to
+ the locally bound values @{ML_text facts}, which indicates any
+ current facts used for forward-chaining.
+
+ \item @{method raw_tactic} is similar to @{method tactic}, but
+ presents the goal state in its raw internal form, where simultaneous
+ subgoals appear as conjunction of the logical framework instead of
+ the usual split into several subgoals. While feature this is useful
+ for debugging of complex method definitions, it should not never
+ appear in production theories.
+
+ \end{description}
+*}
+
+
+section {* The Simplifier \label{sec:simplifier} *}
+
+text {* The Simplifier performs conditional and unconditional
+ rewriting and uses contextual information: rule declarations in the
+ background theory or local proof context are taken into account, as
+ well as chained facts and subgoal premises (``local assumptions'').
+ There are several general hooks that allow to modify the
+ simplification strategy, or incorporate other proof tools that solve
+ sub-problems, produce rewrite rules on demand etc.
+
+ The rewriting strategy is always strictly bottom up, except for
+ congruence rules, which are applied while descending into a term.
+ Conditions in conditional rewrite rules are solved recursively
+ before the rewrite rule is applied.
+
+ The default Simplifier setup of major object logics (HOL, HOLCF,
+ FOL, ZF) makes the Simplifier ready for immediate use, without
+ engaging into the internal structures. Thus it serves as
+ general-purpose proof tool with the main focus on equational
+ reasoning, and a bit more than that.
+*}
+
+
+subsection {* Simplification methods \label{sec:simp-meth} *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{method_def simp} & : & @{text method} \\
+ @{method_def simp_all} & : & @{text method} \\
+ \end{matharray}
+
+ @{rail \<open>
+ (@@{method simp} | @@{method simp_all}) opt? (@{syntax simpmod} * )
+ ;
+
+ opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use' | 'asm_lr' ) ')'
+ ;
+ @{syntax_def simpmod}: ('add' | 'del' | 'only' | 'split' (() | 'add' | 'del') |
+ 'cong' (() | 'add' | 'del')) ':' @{syntax thmrefs}
+ \<close>}
+
+ \begin{description}
+
+ \item @{method simp} invokes the Simplifier on the first subgoal,
+ after inserting chained facts as additional goal premises; further
+ rule declarations may be included via @{text "(simp add: facts)"}.
+ The proof method fails if the subgoal remains unchanged after
+ simplification.
+
+ Note that the original goal premises and chained facts are subject
+ to simplification themselves, while declarations via @{text
+ "add"}/@{text "del"} merely follow the policies of the object-logic
+ to extract rewrite rules from theorems, without further
+ simplification. This may lead to slightly different behavior in
+ either case, which might be required precisely like that in some
+ boundary situations to perform the intended simplification step!
+
+ \medskip The @{text only} modifier first removes all other rewrite
+ rules, looper tactics (including split rules), congruence rules, and
+ then behaves like @{text add}. Implicit solvers remain, which means
+ that trivial rules like reflexivity or introduction of @{text
+ "True"} are available to solve the simplified subgoals, but also
+ non-trivial tools like linear arithmetic in HOL. The latter may
+ lead to some surprise of the meaning of ``only'' in Isabelle/HOL
+ compared to English!
+
+ \medskip The @{text split} modifiers add or delete rules for the
+ Splitter (see also \secref{sec:simp-strategies} on the looper).
+ This works only if the Simplifier method has been properly setup to
+ include the Splitter (all major object logics such HOL, HOLCF, FOL,
+ ZF do this already).
+
+ There is also a separate @{method_ref split} method available for
+ single-step case splitting. The effect of repeatedly applying
+ @{text "(split thms)"} can be imitated by ``@{text "(simp only:
+ split: thms)"}''.
+
+ \medskip The @{text cong} modifiers add or delete Simplifier
+ congruence rules (see also \secref{sec:simp-rules}); the default is
+ to add.
+
+ \item @{method simp_all} is similar to @{method simp}, but acts on
+ all goals, working backwards from the last to the first one as usual
+ in Isabelle.\footnote{The order is irrelevant for goals without
+ schematic variables, so simplification might actually be performed
+ in parallel here.}
+
+ Chained facts are inserted into all subgoals, before the
+ simplification process starts. Further rule declarations are the
+ same as for @{method simp}.
+
+ The proof method fails if all subgoals remain unchanged after
+ simplification.
+
+ \end{description}
+
+ By default the Simplifier methods above take local assumptions fully
+ into account, using equational assumptions in the subsequent
+ normalization process, or simplifying assumptions themselves.
+ Further options allow to fine-tune the behavior of the Simplifier
+ in this respect, corresponding to a variety of ML tactics as
+ follows.\footnote{Unlike the corresponding Isar proof methods, the
+ ML tactics do not insist in changing the goal state.}
+
+ \begin{center}
+ \small
+ \begin{supertabular}{|l|l|p{0.3\textwidth}|}
+ \hline
+ Isar method & ML tactic & behavior \\\hline
+
+ @{text "(simp (no_asm))"} & @{ML simp_tac} & assumptions are ignored
+ completely \\\hline
+
+ @{text "(simp (no_asm_simp))"} & @{ML asm_simp_tac} & assumptions
+ are used in the simplification of the conclusion but are not
+ themselves simplified \\\hline
+
+ @{text "(simp (no_asm_use))"} & @{ML full_simp_tac} & assumptions
+ are simplified but are not used in the simplification of each other
+ or the conclusion \\\hline
+
+ @{text "(simp)"} & @{ML asm_full_simp_tac} & assumptions are used in
+ the simplification of the conclusion and to simplify other
+ assumptions \\\hline
+
+ @{text "(simp (asm_lr))"} & @{ML asm_lr_simp_tac} & compatibility
+ mode: an assumption is only used for simplifying assumptions which
+ are to the right of it \\\hline
+
+ \end{supertabular}
+ \end{center}
+*}
+
+
+subsubsection {* Examples *}
+
+text {* We consider basic algebraic simplifications in Isabelle/HOL.
+ The rather trivial goal @{prop "0 + (x + 0) = x + 0 + 0"} looks like
+ a good candidate to be solved by a single call of @{method simp}:
+*}
+
+lemma "0 + (x + 0) = x + 0 + 0" apply simp? oops
+
+text {* The above attempt \emph{fails}, because @{term "0"} and @{term
+ "op +"} in the HOL library are declared as generic type class
+ operations, without stating any algebraic laws yet. More specific
+ types are required to get access to certain standard simplifications
+ of the theory context, e.g.\ like this: *}
+
+lemma fixes x :: nat shows "0 + (x + 0) = x + 0 + 0" by simp
+lemma fixes x :: int shows "0 + (x + 0) = x + 0 + 0" by simp
+lemma fixes x :: "'a :: monoid_add" shows "0 + (x + 0) = x + 0 + 0" by simp
+
+text {*
+ \medskip In many cases, assumptions of a subgoal are also needed in
+ the simplification process. For example:
+*}
+
+lemma fixes x :: nat shows "x = 0 \<Longrightarrow> x + x = 0" by simp
+lemma fixes x :: nat assumes "x = 0" shows "x + x = 0" apply simp oops
+lemma fixes x :: nat assumes "x = 0" shows "x + x = 0" using assms by simp
+
+text {* As seen above, local assumptions that shall contribute to
+ simplification need to be part of the subgoal already, or indicated
+ explicitly for use by the subsequent method invocation. Both too
+ little or too much information can make simplification fail, for
+ different reasons.
+
+ In the next example the malicious assumption @{prop "\<And>x::nat. f x =
+ g (f (g x))"} does not contribute to solve the problem, but makes
+ the default @{method simp} method loop: the rewrite rule @{text "f
+ ?x \<equiv> g (f (g ?x))"} extracted from the assumption does not
+ terminate. The Simplifier notices certain simple forms of
+ nontermination, but not this one. The problem can be solved
+ nonetheless, by ignoring assumptions via special options as
+ explained before:
+*}
+
+lemma "(\<And>x::nat. f x = g (f (g x))) \<Longrightarrow> f 0 = f 0 + 0"
+ by (simp (no_asm))
+
+text {* The latter form is typical for long unstructured proof
+ scripts, where the control over the goal content is limited. In
+ structured proofs it is usually better to avoid pushing too many
+ facts into the goal state in the first place. Assumptions in the
+ Isar proof context do not intrude the reasoning if not used
+ explicitly. This is illustrated for a toplevel statement and a
+ local proof body as follows:
+*}
+
+lemma
+ assumes "\<And>x::nat. f x = g (f (g x))"
+ shows "f 0 = f 0 + 0" by simp
+
+notepad
+begin
+ assume "\<And>x::nat. f x = g (f (g x))"
+ have "f 0 = f 0 + 0" by simp
+end
+
+text {* \medskip Because assumptions may simplify each other, there
+ can be very subtle cases of nontermination. For example, the regular
+ @{method simp} method applied to @{prop "P (f x) \<Longrightarrow> y = x \<Longrightarrow> f x = f y
+ \<Longrightarrow> Q"} gives rise to the infinite reduction sequence
+ \[
+ @{text "P (f x)"} \stackrel{@{text "f x \<equiv> f y"}}{\longmapsto}
+ @{text "P (f y)"} \stackrel{@{text "y \<equiv> x"}}{\longmapsto}
+ @{text "P (f x)"} \stackrel{@{text "f x \<equiv> f y"}}{\longmapsto} \cdots
+ \]
+ whereas applying the same to @{prop "y = x \<Longrightarrow> f x = f y \<Longrightarrow> P (f x) \<Longrightarrow>
+ Q"} terminates (without solving the goal):
+*}
+
+lemma "y = x \<Longrightarrow> f x = f y \<Longrightarrow> P (f x) \<Longrightarrow> Q"
+ apply simp
+ oops
+
+text {* See also \secref{sec:simp-config} for options to enable
+ Simplifier trace mode, which often helps to diagnose problems with
+ rewrite systems.
+*}
+
+
+subsection {* Declaring rules \label{sec:simp-rules} *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{attribute_def simp} & : & @{text attribute} \\
+ @{attribute_def split} & : & @{text attribute} \\
+ @{attribute_def cong} & : & @{text attribute} \\
+ @{command_def "print_simpset"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
+ \end{matharray}
+
+ @{rail \<open>
+ (@@{attribute simp} | @@{attribute split} | @@{attribute cong})
+ (() | 'add' | 'del')
+ \<close>}
+
+ \begin{description}
+
+ \item @{attribute simp} declares rewrite rules, by adding or
+ deleting them from the simpset within the theory or proof context.
+ Rewrite rules are theorems expressing some form of equality, for
+ example:
+
+ @{text "Suc ?m + ?n = ?m + Suc ?n"} \\
+ @{text "?P \<and> ?P \<longleftrightarrow> ?P"} \\
+ @{text "?A \<union> ?B \<equiv> {x. x \<in> ?A \<or> x \<in> ?B}"}
+
+ \smallskip
+ Conditional rewrites such as @{text "?m < ?n \<Longrightarrow> ?m div ?n = 0"} are
+ also permitted; the conditions can be arbitrary formulas.
+
+ \medskip Internally, all rewrite rules are translated into Pure
+ equalities, theorems with conclusion @{text "lhs \<equiv> rhs"}. The
+ simpset contains a function for extracting equalities from arbitrary
+ theorems, which is usually installed when the object-logic is
+ configured initially. For example, @{text "\<not> ?x \<in> {}"} could be
+ turned into @{text "?x \<in> {} \<equiv> False"}. Theorems that are declared as
+ @{attribute simp} and local assumptions within a goal are treated
+ uniformly in this respect.
+
+ The Simplifier accepts the following formats for the @{text "lhs"}
+ term:
+
+ \begin{enumerate}
+
+ \item First-order patterns, considering the sublanguage of
+ application of constant operators to variable operands, without
+ @{text "\<lambda>"}-abstractions or functional variables.
+ For example:
+
+ @{text "(?x + ?y) + ?z \<equiv> ?x + (?y + ?z)"} \\
+ @{text "f (f ?x ?y) ?z \<equiv> f ?x (f ?y ?z)"}
+
+ \item Higher-order patterns in the sense of \cite{nipkow-patterns}.
+ These are terms in @{text "\<beta>"}-normal form (this will always be the
+ case unless you have done something strange) where each occurrence
+ of an unknown is of the form @{text "?F x\<^sub>1 \<dots> x\<^sub>n"}, where the
+ @{text "x\<^sub>i"} are distinct bound variables.
+
+ For example, @{text "(\<forall>x. ?P x \<and> ?Q x) \<equiv> (\<forall>x. ?P x) \<and> (\<forall>x. ?Q x)"}
+ or its symmetric form, since the @{text "rhs"} is also a
+ higher-order pattern.
+
+ \item Physical first-order patterns over raw @{text "\<lambda>"}-term
+ structure without @{text "\<alpha>\<beta>\<eta>"}-equality; abstractions and bound
+ variables are treated like quasi-constant term material.
+
+ For example, the rule @{text "?f ?x \<in> range ?f = True"} rewrites the
+ term @{text "g a \<in> range g"} to @{text "True"}, but will fail to
+ match @{text "g (h b) \<in> range (\<lambda>x. g (h x))"}. However, offending
+ subterms (in our case @{text "?f ?x"}, which is not a pattern) can
+ be replaced by adding new variables and conditions like this: @{text
+ "?y = ?f ?x \<Longrightarrow> ?y \<in> range ?f = True"} is acceptable as a conditional
+ rewrite rule of the second category since conditions can be
+ arbitrary terms.
+
+ \end{enumerate}
+
+ \item @{attribute split} declares case split rules.
+
+ \item @{attribute cong} declares congruence rules to the Simplifier
+ context.
+
+ Congruence rules are equalities of the form @{text [display]
+ "\<dots> \<Longrightarrow> f ?x\<^sub>1 \<dots> ?x\<^sub>n = f ?y\<^sub>1 \<dots> ?y\<^sub>n"}
+
+ This controls the simplification of the arguments of @{text f}. For
+ example, some arguments can be simplified under additional
+ assumptions: @{text [display] "?P\<^sub>1 \<longleftrightarrow> ?Q\<^sub>1 \<Longrightarrow> (?Q\<^sub>1 \<Longrightarrow> ?P\<^sub>2 \<longleftrightarrow> ?Q\<^sub>2) \<Longrightarrow>
+ (?P\<^sub>1 \<longrightarrow> ?P\<^sub>2) \<longleftrightarrow> (?Q\<^sub>1 \<longrightarrow> ?Q\<^sub>2)"}
+
+ Given this rule, the simplifier assumes @{text "?Q\<^sub>1"} and extracts
+ rewrite rules from it when simplifying @{text "?P\<^sub>2"}. Such local
+ assumptions are effective for rewriting formulae such as @{text "x =
+ 0 \<longrightarrow> y + x = y"}.
+
+ %FIXME
+ %The local assumptions are also provided as theorems to the solver;
+ %see \secref{sec:simp-solver} below.
+
+ \medskip The following congruence rule for bounded quantifiers also
+ supplies contextual information --- about the bound variable:
+ @{text [display] "(?A = ?B) \<Longrightarrow> (\<And>x. x \<in> ?B \<Longrightarrow> ?P x \<longleftrightarrow> ?Q x) \<Longrightarrow>
+ (\<forall>x \<in> ?A. ?P x) \<longleftrightarrow> (\<forall>x \<in> ?B. ?Q x)"}
+
+ \medskip This congruence rule for conditional expressions can
+ supply contextual information for simplifying the arms:
+ @{text [display] "?p = ?q \<Longrightarrow> (?q \<Longrightarrow> ?a = ?c) \<Longrightarrow> (\<not> ?q \<Longrightarrow> ?b = ?d) \<Longrightarrow>
+ (if ?p then ?a else ?b) = (if ?q then ?c else ?d)"}
+
+ A congruence rule can also \emph{prevent} simplification of some
+ arguments. Here is an alternative congruence rule for conditional
+ expressions that conforms to non-strict functional evaluation:
+ @{text [display] "?p = ?q \<Longrightarrow> (if ?p then ?a else ?b) = (if ?q then ?a else ?b)"}
+
+ Only the first argument is simplified; the others remain unchanged.
+ This can make simplification much faster, but may require an extra
+ case split over the condition @{text "?q"} to prove the goal.
+
+ \item @{command "print_simpset"} prints the collection of rules
+ declared to the Simplifier, which is also known as ``simpset''
+ internally.
+
+ For historical reasons, simpsets may occur independently from the
+ current context, but are conceptually dependent on it. When the
+ Simplifier is invoked via one of its main entry points in the Isar
+ source language (as proof method \secref{sec:simp-meth} or rule
+ attribute \secref{sec:simp-meth}), its simpset is derived from the
+ current proof context, and carries a back-reference to that for
+ other tools that might get invoked internally (e.g.\ simplification
+ procedures \secref{sec:simproc}). A mismatch of the context of the
+ simpset and the context of the problem being simplified may lead to
+ unexpected results.
+
+ \end{description}
+
+ The implicit simpset of the theory context is propagated
+ monotonically through the theory hierarchy: forming a new theory,
+ the union of the simpsets of its imports are taken as starting
+ point. Also note that definitional packages like @{command
+ "datatype"}, @{command "primrec"}, @{command "fun"} routinely
+ declare Simplifier rules to the target context, while plain
+ @{command "definition"} is an exception in \emph{not} declaring
+ anything.
+
+ \medskip It is up the user to manipulate the current simpset further
+ by explicitly adding or deleting theorems as simplification rules,
+ or installing other tools via simplification procedures
+ (\secref{sec:simproc}). Good simpsets are hard to design. Rules
+ that obviously simplify, like @{text "?n + 0 \<equiv> ?n"} are good
+ candidates for the implicit simpset, unless a special
+ non-normalizing behavior of certain operations is intended. More
+ specific rules (such as distributive laws, which duplicate subterms)
+ should be added only for specific proof steps. Conversely,
+ sometimes a rule needs to be deleted just for some part of a proof.
+ The need of frequent additions or deletions may indicate a poorly
+ designed simpset.
+
+ \begin{warn}
+ The union of simpsets from theory imports (as described above) is
+ not always a good starting point for the new theory. If some
+ ancestors have deleted simplification rules because they are no
+ longer wanted, while others have left those rules in, then the union
+ will contain the unwanted rules, and thus have to be deleted again
+ in the theory body.
+ \end{warn}
+*}
+
+
+subsection {* Ordered rewriting with permutative rules *}
+
+text {* A rewrite rule is \emph{permutative} if the left-hand side and
+ right-hand side are the equal up to renaming of variables. The most
+ common permutative rule is commutativity: @{text "?x + ?y = ?y +
+ ?x"}. Other examples include @{text "(?x - ?y) - ?z = (?x - ?z) -
+ ?y"} in arithmetic and @{text "insert ?x (insert ?y ?A) = insert ?y
+ (insert ?x ?A)"} for sets. Such rules are common enough to merit
+ special attention.
+
+ Because ordinary rewriting loops given such rules, the Simplifier
+ employs a special strategy, called \emph{ordered rewriting}.
+ Permutative rules are detected and only applied if the rewriting
+ step decreases the redex wrt.\ a given term ordering. For example,
+ commutativity rewrites @{text "b + a"} to @{text "a + b"}, but then
+ stops, because the redex cannot be decreased further in the sense of
+ the term ordering.
+
+ The default is lexicographic ordering of term structure, but this
+ could be also changed locally for special applications via
+ @{index_ML Simplifier.set_termless} in Isabelle/ML.
+
+ \medskip Permutative rewrite rules are declared to the Simplifier
+ just like other rewrite rules. Their special status is recognized
+ automatically, and their application is guarded by the term ordering
+ accordingly. *}
+
+
+subsubsection {* Rewriting with AC operators *}
+
+text {* Ordered rewriting is particularly effective in the case of
+ associative-commutative operators. (Associativity by itself is not
+ permutative.) When dealing with an AC-operator @{text "f"}, keep
+ the following points in mind:
+
+ \begin{itemize}
+
+ \item The associative law must always be oriented from left to
+ right, namely @{text "f (f x y) z = f x (f y z)"}. The opposite
+ orientation, if used with commutativity, leads to looping in
+ conjunction with the standard term order.
+
+ \item To complete your set of rewrite rules, you must add not just
+ associativity (A) and commutativity (C) but also a derived rule
+ \emph{left-commutativity} (LC): @{text "f x (f y z) = f y (f x z)"}.
+
+ \end{itemize}
+
+ Ordered rewriting with the combination of A, C, and LC sorts a term
+ lexicographically --- the rewriting engine imitates bubble-sort.
+*}
+
+locale AC_example =
+ fixes f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix "\<bullet>" 60)
+ assumes assoc: "(x \<bullet> y) \<bullet> z = x \<bullet> (y \<bullet> z)"
+ assumes commute: "x \<bullet> y = y \<bullet> x"
+begin
+
+lemma left_commute: "x \<bullet> (y \<bullet> z) = y \<bullet> (x \<bullet> z)"
+proof -
+ have "(x \<bullet> y) \<bullet> z = (y \<bullet> x) \<bullet> z" by (simp only: commute)
+ then show ?thesis by (simp only: assoc)
+qed
+
+lemmas AC_rules = assoc commute left_commute
+
+text {* Thus the Simplifier is able to establish equalities with
+ arbitrary permutations of subterms, by normalizing to a common
+ standard form. For example: *}
+
+lemma "(b \<bullet> c) \<bullet> a = xxx"
+ apply (simp only: AC_rules)
+ txt {* @{subgoals} *}
+ oops
+
+lemma "(b \<bullet> c) \<bullet> a = a \<bullet> (b \<bullet> c)" by (simp only: AC_rules)
+lemma "(b \<bullet> c) \<bullet> a = c \<bullet> (b \<bullet> a)" by (simp only: AC_rules)
+lemma "(b \<bullet> c) \<bullet> a = (c \<bullet> b) \<bullet> a" by (simp only: AC_rules)
+
+end
+
+text {* Martin and Nipkow \cite{martin-nipkow} discuss the theory and
+ give many examples; other algebraic structures are amenable to
+ ordered rewriting, such as boolean rings. The Boyer-Moore theorem
+ prover \cite{bm88book} also employs ordered rewriting.
+*}
+
+
+subsubsection {* Re-orienting equalities *}
+
+text {* Another application of ordered rewriting uses the derived rule
+ @{thm [source] eq_commute}: @{thm [source = false] eq_commute} to
+ reverse equations.
+
+ This is occasionally useful to re-orient local assumptions according
+ to the term ordering, when other built-in mechanisms of
+ reorientation and mutual simplification fail to apply. *}
+
+
+subsection {* Configuration options \label{sec:simp-config} *}
+
+text {*
+ \begin{tabular}{rcll}
+ @{attribute_def simp_depth_limit} & : & @{text attribute} & default @{text 100} \\
+ @{attribute_def simp_trace} & : & @{text attribute} & default @{text false} \\
+ @{attribute_def simp_trace_depth_limit} & : & @{text attribute} & default @{text 1} \\
+ @{attribute_def simp_debug} & : & @{text attribute} & default @{text false} \\
+ \end{tabular}
+ \medskip
+
+ These configurations options control further aspects of the Simplifier.
+ See also \secref{sec:config}.
+
+ \begin{description}
+
+ \item @{attribute simp_depth_limit} limits the number of recursive
+ invocations of the Simplifier during conditional rewriting.
+
+ \item @{attribute simp_trace} makes the Simplifier output internal
+ operations. This includes rewrite steps, but also bookkeeping like
+ modifications of the simpset.
+
+ \item @{attribute simp_trace_depth_limit} limits the effect of
+ @{attribute simp_trace} to the given depth of recursive Simplifier
+ invocations (when solving conditions of rewrite rules).
+
+ \item @{attribute simp_debug} makes the Simplifier output some extra
+ information about internal operations. This includes any attempted
+ invocation of simplification procedures.
+
+ \end{description}
+*}
+
+
+subsection {* Simplification procedures \label{sec:simproc} *}
+
+text {* Simplification procedures are ML functions that produce proven
+ rewrite rules on demand. They are associated with higher-order
+ patterns that approximate the left-hand sides of equations. The
+ Simplifier first matches the current redex against one of the LHS
+ patterns; if this succeeds, the corresponding ML function is
+ invoked, passing the Simplifier context and redex term. Thus rules
+ may be specifically fashioned for particular situations, resulting
+ in a more powerful mechanism than term rewriting by a fixed set of
+ rules.
+
+ Any successful result needs to be a (possibly conditional) rewrite
+ rule @{text "t \<equiv> u"} that is applicable to the current redex. The
+ rule will be applied just as any ordinary rewrite rule. It is
+ expected to be already in \emph{internal form}, bypassing the
+ automatic preprocessing of object-level equivalences.
+
+ \begin{matharray}{rcl}
+ @{command_def "simproc_setup"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+ simproc & : & @{text attribute} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{command simproc_setup} @{syntax name} '(' (@{syntax term} + '|') ')' '='
+ @{syntax text} \<newline> (@'identifier' (@{syntax nameref}+))?
+ ;
+
+ @@{attribute simproc} (('add' ':')? | 'del' ':') (@{syntax name}+)
+ \<close>}
+
+ \begin{description}
+
+ \item @{command "simproc_setup"} defines a named simplification
+ procedure that is invoked by the Simplifier whenever any of the
+ given term patterns match the current redex. The implementation,
+ which is provided as ML source text, needs to be of type @{ML_type
+ "morphism -> simpset -> cterm -> thm option"}, where the @{ML_type
+ cterm} represents the current redex @{text r} and the result is
+ supposed to be some proven rewrite rule @{text "r \<equiv> r'"} (or a
+ generalized version), or @{ML NONE} to indicate failure. The
+ @{ML_type simpset} argument holds the full context of the current
+ Simplifier invocation, including the actual Isar proof context. The
+ @{ML_type morphism} informs about the difference of the original
+ compilation context wrt.\ the one of the actual application later
+ on. The optional @{keyword "identifier"} specifies theorems that
+ represent the logical content of the abstract theory of this
+ simproc.
+
+ Morphisms and identifiers are only relevant for simprocs that are
+ defined within a local target context, e.g.\ in a locale.
+
+ \item @{text "simproc add: name"} and @{text "simproc del: name"}
+ add or delete named simprocs to the current Simplifier context. The
+ default is to add a simproc. Note that @{command "simproc_setup"}
+ already adds the new simproc to the subsequent context.
+
+ \end{description}
+*}
+
+
+subsubsection {* Example *}
+
+text {* The following simplification procedure for @{thm
+ [source=false, show_types] unit_eq} in HOL performs fine-grained
+ control over rule application, beyond higher-order pattern matching.
+ Declaring @{thm unit_eq} as @{attribute simp} directly would make
+ the simplifier loop! Note that a version of this simplification
+ procedure is already active in Isabelle/HOL. *}
+
+simproc_setup unit ("x::unit") = {*
+ fn _ => fn _ => fn ct =>
+ if HOLogic.is_unit (term_of ct) then NONE
+ else SOME (mk_meta_eq @{thm unit_eq})
+*}
+
+text {* Since the Simplifier applies simplification procedures
+ frequently, it is important to make the failure check in ML
+ reasonably fast. *}
+
+
+subsection {* Configurable Simplifier strategies \label{sec:simp-strategies} *}
+
+text {* The core term-rewriting engine of the Simplifier is normally
+ used in combination with some add-on components that modify the
+ strategy and allow to integrate other non-Simplifier proof tools.
+ These may be reconfigured in ML as explained below. Even if the
+ default strategies of object-logics like Isabelle/HOL are used
+ unchanged, it helps to understand how the standard Simplifier
+ strategies work. *}
+
+
+subsubsection {* The subgoaler *}
+
+text {*
+ \begin{mldecls}
+ @{index_ML Simplifier.set_subgoaler: "(Proof.context -> int -> tactic) ->
+ Proof.context -> Proof.context"} \\
+ @{index_ML Simplifier.prems_of: "Proof.context -> thm list"} \\
+ \end{mldecls}
+
+ The subgoaler is the tactic used to solve subgoals arising out of
+ conditional rewrite rules or congruence rules. The default should
+ be simplification itself. In rare situations, this strategy may
+ need to be changed. For example, if the premise of a conditional
+ rule is an instance of its conclusion, as in @{text "Suc ?m < ?n \<Longrightarrow>
+ ?m < ?n"}, the default strategy could loop. % FIXME !??
+
+ \begin{description}
+
+ \item @{ML Simplifier.set_subgoaler}~@{text "tac ctxt"} sets the
+ subgoaler of the context to @{text "tac"}. The tactic will
+ be applied to the context of the running Simplifier instance.
+
+ \item @{ML Simplifier.prems_of}~@{text "ctxt"} retrieves the current
+ set of premises from the context. This may be non-empty only if
+ the Simplifier has been told to utilize local assumptions in the
+ first place (cf.\ the options in \secref{sec:simp-meth}).
+
+ \end{description}
+
+ As an example, consider the following alternative subgoaler:
+*}
+
+ML {*
+ fun subgoaler_tac ctxt =
+ assume_tac ORELSE'
+ resolve_tac (Simplifier.prems_of ctxt) ORELSE'
+ asm_simp_tac ctxt
+*}
+
+text {* This tactic first tries to solve the subgoal by assumption or
+ by resolving with with one of the premises, calling simplification
+ only if that fails. *}
+
+
+subsubsection {* The solver *}
+
+text {*
+ \begin{mldecls}
+ @{index_ML_type solver} \\
+ @{index_ML Simplifier.mk_solver: "string ->
+ (Proof.context -> int -> tactic) -> solver"} \\
+ @{index_ML_op setSolver: "Proof.context * solver -> Proof.context"} \\
+ @{index_ML_op addSolver: "Proof.context * solver -> Proof.context"} \\
+ @{index_ML_op setSSolver: "Proof.context * solver -> Proof.context"} \\
+ @{index_ML_op addSSolver: "Proof.context * solver -> Proof.context"} \\
+ \end{mldecls}
+
+ A solver is a tactic that attempts to solve a subgoal after
+ simplification. Its core functionality is to prove trivial subgoals
+ such as @{prop "True"} and @{text "t = t"}, but object-logics might
+ be more ambitious. For example, Isabelle/HOL performs a restricted
+ version of linear arithmetic here.
+
+ Solvers are packaged up in abstract type @{ML_type solver}, with
+ @{ML Simplifier.mk_solver} as the only operation to create a solver.
+
+ \medskip Rewriting does not instantiate unknowns. For example,
+ rewriting alone cannot prove @{text "a \<in> ?A"} since this requires
+ instantiating @{text "?A"}. The solver, however, is an arbitrary
+ tactic and may instantiate unknowns as it pleases. This is the only
+ way the Simplifier can handle a conditional rewrite rule whose
+ condition contains extra variables. When a simplification tactic is
+ to be combined with other provers, especially with the Classical
+ Reasoner, it is important whether it can be considered safe or not.
+ For this reason a simpset contains two solvers: safe and unsafe.
+
+ The standard simplification strategy solely uses the unsafe solver,
+ which is appropriate in most cases. For special applications where
+ the simplification process is not allowed to instantiate unknowns
+ within the goal, simplification starts with the safe solver, but may
+ still apply the ordinary unsafe one in nested simplifications for
+ conditional rules or congruences. Note that in this way the overall
+ tactic is not totally safe: it may instantiate unknowns that appear
+ also in other subgoals.
+
+ \begin{description}
+
+ \item @{ML Simplifier.mk_solver}~@{text "name tac"} turns @{text
+ "tac"} into a solver; the @{text "name"} is only attached as a
+ comment and has no further significance.
+
+ \item @{text "ctxt setSSolver solver"} installs @{text "solver"} as
+ the safe solver of @{text "ctxt"}.
+
+ \item @{text "ctxt addSSolver solver"} adds @{text "solver"} as an
+ additional safe solver; it will be tried after the solvers which had
+ already been present in @{text "ctxt"}.
+
+ \item @{text "ctxt setSolver solver"} installs @{text "solver"} as the
+ unsafe solver of @{text "ctxt"}.
+
+ \item @{text "ctxt addSolver solver"} adds @{text "solver"} as an
+ additional unsafe solver; it will be tried after the solvers which
+ had already been present in @{text "ctxt"}.
+
+ \end{description}
+
+ \medskip The solver tactic is invoked with the context of the
+ running Simplifier. Further operations
+ may be used to retrieve relevant information, such as the list of
+ local Simplifier premises via @{ML Simplifier.prems_of} --- this
+ list may be non-empty only if the Simplifier runs in a mode that
+ utilizes local assumptions (see also \secref{sec:simp-meth}). The
+ solver is also presented the full goal including its assumptions in
+ any case. Thus it can use these (e.g.\ by calling @{ML
+ assume_tac}), even if the Simplifier proper happens to ignore local
+ premises at the moment.
+
+ \medskip As explained before, the subgoaler is also used to solve
+ the premises of congruence rules. These are usually of the form
+ @{text "s = ?x"}, where @{text "s"} needs to be simplified and
+ @{text "?x"} needs to be instantiated with the result. Typically,
+ the subgoaler will invoke the Simplifier at some point, which will
+ eventually call the solver. For this reason, solver tactics must be
+ prepared to solve goals of the form @{text "t = ?x"}, usually by
+ reflexivity. In particular, reflexivity should be tried before any
+ of the fancy automated proof tools.
+
+ It may even happen that due to simplification the subgoal is no
+ longer an equality. For example, @{text "False \<longleftrightarrow> ?Q"} could be
+ rewritten to @{text "\<not> ?Q"}. To cover this case, the solver could
+ try resolving with the theorem @{text "\<not> False"} of the
+ object-logic.
+
+ \medskip
+
+ \begin{warn}
+ If a premise of a congruence rule cannot be proved, then the
+ congruence is ignored. This should only happen if the rule is
+ \emph{conditional} --- that is, contains premises not of the form
+ @{text "t = ?x"}. Otherwise it indicates that some congruence rule,
+ or possibly the subgoaler or solver, is faulty.
+ \end{warn}
+*}
+
+
+subsubsection {* The looper *}
+
+text {*
+ \begin{mldecls}
+ @{index_ML_op setloop: "Proof.context *
+ (Proof.context -> int -> tactic) -> Proof.context"} \\
+ @{index_ML_op addloop: "Proof.context *
+ (string * (Proof.context -> int -> tactic))
+ -> Proof.context"} \\
+ @{index_ML_op delloop: "Proof.context * string -> Proof.context"} \\
+ @{index_ML Splitter.add_split: "thm -> Proof.context -> Proof.context"} \\
+ @{index_ML Splitter.del_split: "thm -> Proof.context -> Proof.context"} \\
+ \end{mldecls}
+
+ The looper is a list of tactics that are applied after
+ simplification, in case the solver failed to solve the simplified
+ goal. If the looper succeeds, the simplification process is started
+ all over again. Each of the subgoals generated by the looper is
+ attacked in turn, in reverse order.
+
+ A typical looper is \emph{case splitting}: the expansion of a
+ conditional. Another possibility is to apply an elimination rule on
+ the assumptions. More adventurous loopers could start an induction.
+
+ \begin{description}
+
+ \item @{text "ctxt setloop tac"} installs @{text "tac"} as the only
+ looper tactic of @{text "ctxt"}.
+
+ \item @{text "ctxt addloop (name, tac)"} adds @{text "tac"} as an
+ additional looper tactic with name @{text "name"}, which is
+ significant for managing the collection of loopers. The tactic will
+ be tried after the looper tactics that had already been present in
+ @{text "ctxt"}.
+
+ \item @{text "ctxt delloop name"} deletes the looper tactic that was
+ associated with @{text "name"} from @{text "ctxt"}.
+
+ \item @{ML Splitter.add_split}~@{text "thm ctxt"} adds split tactics
+ for @{text "thm"} as additional looper tactics of @{text "ctxt"}.
+
+ \item @{ML Splitter.del_split}~@{text "thm ctxt"} deletes the split
+ tactic corresponding to @{text thm} from the looper tactics of
+ @{text "ctxt"}.
+
+ \end{description}
+
+ The splitter replaces applications of a given function; the
+ right-hand side of the replacement can be anything. For example,
+ here is a splitting rule for conditional expressions:
+
+ @{text [display] "?P (if ?Q ?x ?y) \<longleftrightarrow> (?Q \<longrightarrow> ?P ?x) \<and> (\<not> ?Q \<longrightarrow> ?P ?y)"}
+
+ Another example is the elimination operator for Cartesian products
+ (which happens to be called @{text split} in Isabelle/HOL:
+
+ @{text [display] "?P (split ?f ?p) \<longleftrightarrow> (\<forall>a b. ?p = (a, b) \<longrightarrow> ?P (f a b))"}
+
+ For technical reasons, there is a distinction between case splitting
+ in the conclusion and in the premises of a subgoal. The former is
+ done by @{ML Splitter.split_tac} with rules like @{thm [source]
+ split_if} or @{thm [source] option.split}, which do not split the
+ subgoal, while the latter is done by @{ML Splitter.split_asm_tac}
+ with rules like @{thm [source] split_if_asm} or @{thm [source]
+ option.split_asm}, which split the subgoal. The function @{ML
+ Splitter.add_split} automatically takes care of which tactic to
+ call, analyzing the form of the rules given as argument; it is the
+ same operation behind @{text "split"} attribute or method modifier
+ syntax in the Isar source language.
+
+ Case splits should be allowed only when necessary; they are
+ expensive and hard to control. Case-splitting on if-expressions in
+ the conclusion is usually beneficial, so it is enabled by default in
+ Isabelle/HOL and Isabelle/FOL/ZF.
+
+ \begin{warn}
+ With @{ML Splitter.split_asm_tac} as looper component, the
+ Simplifier may split subgoals! This might cause unexpected problems
+ in tactic expressions that silently assume 0 or 1 subgoals after
+ simplification.
+ \end{warn}
+*}
+
+
+subsection {* Forward simplification \label{sec:simp-forward} *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{attribute_def simplified} & : & @{text attribute} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{attribute simplified} opt? @{syntax thmrefs}?
+ ;
+
+ opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use') ')'
+ \<close>}
+
+ \begin{description}
+
+ \item @{attribute simplified}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} causes a theorem to
+ be simplified, either by exactly the specified rules @{text "a\<^sub>1, \<dots>,
+ a\<^sub>n"}, or the implicit Simplifier context if no arguments are given.
+ The result is fully simplified by default, including assumptions and
+ conclusion; the options @{text no_asm} etc.\ tune the Simplifier in
+ the same way as the for the @{text simp} method.
+
+ Note that forward simplification restricts the simplifier to its
+ most basic operation of term rewriting; solver and looper tactics
+ (\secref{sec:simp-strategies}) are \emph{not} involved here. The
+ @{attribute simplified} attribute should be only rarely required
+ under normal circumstances.
+
+ \end{description}
+*}
+
+
+section {* The Classical Reasoner \label{sec:classical} *}
+
+subsection {* Basic concepts *}
+
+text {* Although Isabelle is generic, many users will be working in
+ some extension of classical first-order logic. Isabelle/ZF is built
+ upon theory FOL, while Isabelle/HOL conceptually contains
+ first-order logic as a fragment. Theorem-proving in predicate logic
+ is undecidable, but many automated strategies have been developed to
+ assist in this task.
+
+ Isabelle's classical reasoner is a generic package that accepts
+ certain information about a logic and delivers a suite of automatic
+ proof tools, based on rules that are classified and declared in the
+ context. These proof procedures are slow and simplistic compared
+ with high-end automated theorem provers, but they can save
+ considerable time and effort in practice. They can prove theorems
+ such as Pelletier's \cite{pelletier86} problems 40 and 41 in a few
+ milliseconds (including full proof reconstruction): *}
+
+lemma "(\<exists>y. \<forall>x. F x y \<longleftrightarrow> F x x) \<longrightarrow> \<not> (\<forall>x. \<exists>y. \<forall>z. F z y \<longleftrightarrow> \<not> F z x)"
+ by blast
+
+lemma "(\<forall>z. \<exists>y. \<forall>x. f x y \<longleftrightarrow> f x z \<and> \<not> f x x) \<longrightarrow> \<not> (\<exists>z. \<forall>x. f x z)"
+ by blast
+
+text {* The proof tools are generic. They are not restricted to
+ first-order logic, and have been heavily used in the development of
+ the Isabelle/HOL library and applications. The tactics can be
+ traced, and their components can be called directly; in this manner,
+ any proof can be viewed interactively. *}
+
+
+subsubsection {* The sequent calculus *}
+
+text {* Isabelle supports natural deduction, which is easy to use for
+ interactive proof. But natural deduction does not easily lend
+ itself to automation, and has a bias towards intuitionism. For
+ certain proofs in classical logic, it can not be called natural.
+ The \emph{sequent calculus}, a generalization of natural deduction,
+ is easier to automate.
+
+ A \textbf{sequent} has the form @{text "\<Gamma> \<turnstile> \<Delta>"}, where @{text "\<Gamma>"}
+ and @{text "\<Delta>"} are sets of formulae.\footnote{For first-order
+ logic, sequents can equivalently be made from lists or multisets of
+ formulae.} The sequent @{text "P\<^sub>1, \<dots>, P\<^sub>m \<turnstile> Q\<^sub>1, \<dots>, Q\<^sub>n"} is
+ \textbf{valid} if @{text "P\<^sub>1 \<and> \<dots> \<and> P\<^sub>m"} implies @{text "Q\<^sub>1 \<or> \<dots> \<or>
+ Q\<^sub>n"}. Thus @{text "P\<^sub>1, \<dots>, P\<^sub>m"} represent assumptions, each of which
+ is true, while @{text "Q\<^sub>1, \<dots>, Q\<^sub>n"} represent alternative goals. A
+ sequent is \textbf{basic} if its left and right sides have a common
+ formula, as in @{text "P, Q \<turnstile> Q, R"}; basic sequents are trivially
+ valid.
+
+ Sequent rules are classified as \textbf{right} or \textbf{left},
+ indicating which side of the @{text "\<turnstile>"} symbol they operate on.
+ Rules that operate on the right side are analogous to natural
+ deduction's introduction rules, and left rules are analogous to
+ elimination rules. The sequent calculus analogue of @{text "(\<longrightarrow>I)"}
+ is the rule
+ \[
+ \infer[@{text "(\<longrightarrow>R)"}]{@{text "\<Gamma> \<turnstile> \<Delta>, P \<longrightarrow> Q"}}{@{text "P, \<Gamma> \<turnstile> \<Delta>, Q"}}
+ \]
+ Applying the rule backwards, this breaks down some implication on
+ the right side of a sequent; @{text "\<Gamma>"} and @{text "\<Delta>"} stand for
+ the sets of formulae that are unaffected by the inference. The
+ analogue of the pair @{text "(\<or>I1)"} and @{text "(\<or>I2)"} is the
+ single rule
+ \[
+ \infer[@{text "(\<or>R)"}]{@{text "\<Gamma> \<turnstile> \<Delta>, P \<or> Q"}}{@{text "\<Gamma> \<turnstile> \<Delta>, P, Q"}}
+ \]
+ This breaks down some disjunction on the right side, replacing it by
+ both disjuncts. Thus, the sequent calculus is a kind of
+ multiple-conclusion logic.
+
+ To illustrate the use of multiple formulae on the right, let us
+ prove the classical theorem @{text "(P \<longrightarrow> Q) \<or> (Q \<longrightarrow> P)"}. Working
+ backwards, we reduce this formula to a basic sequent:
+ \[
+ \infer[@{text "(\<or>R)"}]{@{text "\<turnstile> (P \<longrightarrow> Q) \<or> (Q \<longrightarrow> P)"}}
+ {\infer[@{text "(\<longrightarrow>R)"}]{@{text "\<turnstile> (P \<longrightarrow> Q), (Q \<longrightarrow> P)"}}
+ {\infer[@{text "(\<longrightarrow>R)"}]{@{text "P \<turnstile> Q, (Q \<longrightarrow> P)"}}
+ {@{text "P, Q \<turnstile> Q, P"}}}}
+ \]
+
+ This example is typical of the sequent calculus: start with the
+ desired theorem and apply rules backwards in a fairly arbitrary
+ manner. This yields a surprisingly effective proof procedure.
+ Quantifiers add only few complications, since Isabelle handles
+ parameters and schematic variables. See \cite[Chapter
+ 10]{paulson-ml2} for further discussion. *}
+
+
+subsubsection {* Simulating sequents by natural deduction *}
+
+text {* Isabelle can represent sequents directly, as in the
+ object-logic LK. But natural deduction is easier to work with, and
+ most object-logics employ it. Fortunately, we can simulate the
+ sequent @{text "P\<^sub>1, \<dots>, P\<^sub>m \<turnstile> Q\<^sub>1, \<dots>, Q\<^sub>n"} by the Isabelle formula
+ @{text "P\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> P\<^sub>m \<Longrightarrow> \<not> Q\<^sub>2 \<Longrightarrow> ... \<Longrightarrow> \<not> Q\<^sub>n \<Longrightarrow> Q\<^sub>1"} where the order of
+ the assumptions and the choice of @{text "Q\<^sub>1"} are arbitrary.
+ Elim-resolution plays a key role in simulating sequent proofs.
+
+ We can easily handle reasoning on the left. Elim-resolution with
+ the rules @{text "(\<or>E)"}, @{text "(\<bottom>E)"} and @{text "(\<exists>E)"} achieves
+ a similar effect as the corresponding sequent rules. For the other
+ connectives, we use sequent-style elimination rules instead of
+ destruction rules such as @{text "(\<and>E1, 2)"} and @{text "(\<forall>E)"}.
+ But note that the rule @{text "(\<not>L)"} has no effect under our
+ representation of sequents!
+ \[
+ \infer[@{text "(\<not>L)"}]{@{text "\<not> P, \<Gamma> \<turnstile> \<Delta>"}}{@{text "\<Gamma> \<turnstile> \<Delta>, P"}}
+ \]
+
+ What about reasoning on the right? Introduction rules can only
+ affect the formula in the conclusion, namely @{text "Q\<^sub>1"}. The
+ other right-side formulae are represented as negated assumptions,
+ @{text "\<not> Q\<^sub>2, \<dots>, \<not> Q\<^sub>n"}. In order to operate on one of these, it
+ must first be exchanged with @{text "Q\<^sub>1"}. Elim-resolution with the
+ @{text swap} rule has this effect: @{text "\<not> P \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> R"}
+
+ To ensure that swaps occur only when necessary, each introduction
+ rule is converted into a swapped form: it is resolved with the
+ second premise of @{text "(swap)"}. The swapped form of @{text
+ "(\<and>I)"}, which might be called @{text "(\<not>\<and>E)"}, is
+ @{text [display] "\<not> (P \<and> Q) \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> (\<not> R \<Longrightarrow> Q) \<Longrightarrow> R"}
+
+ Similarly, the swapped form of @{text "(\<longrightarrow>I)"} is
+ @{text [display] "\<not> (P \<longrightarrow> Q) \<Longrightarrow> (\<not> R \<Longrightarrow> P \<Longrightarrow> Q) \<Longrightarrow> R"}
+
+ Swapped introduction rules are applied using elim-resolution, which
+ deletes the negated formula. Our representation of sequents also
+ requires the use of ordinary introduction rules. If we had no
+ regard for readability of intermediate goal states, we could treat
+ the right side more uniformly by representing sequents as @{text
+ [display] "P\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> P\<^sub>m \<Longrightarrow> \<not> Q\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> \<not> Q\<^sub>n \<Longrightarrow> \<bottom>"}
+*}
+
+
+subsubsection {* Extra rules for the sequent calculus *}
+
+text {* As mentioned, destruction rules such as @{text "(\<and>E1, 2)"} and
+ @{text "(\<forall>E)"} must be replaced by sequent-style elimination rules.
+ In addition, we need rules to embody the classical equivalence
+ between @{text "P \<longrightarrow> Q"} and @{text "\<not> P \<or> Q"}. The introduction
+ rules @{text "(\<or>I1, 2)"} are replaced by a rule that simulates
+ @{text "(\<or>R)"}: @{text [display] "(\<not> Q \<Longrightarrow> P) \<Longrightarrow> P \<or> Q"}
+
+ The destruction rule @{text "(\<longrightarrow>E)"} is replaced by @{text [display]
+ "(P \<longrightarrow> Q) \<Longrightarrow> (\<not> P \<Longrightarrow> R) \<Longrightarrow> (Q \<Longrightarrow> R) \<Longrightarrow> R"}
+
+ Quantifier replication also requires special rules. In classical
+ logic, @{text "\<exists>x. P x"} is equivalent to @{text "\<not> (\<forall>x. \<not> P x)"};
+ the rules @{text "(\<exists>R)"} and @{text "(\<forall>L)"} are dual:
+ \[
+ \infer[@{text "(\<exists>R)"}]{@{text "\<Gamma> \<turnstile> \<Delta>, \<exists>x. P x"}}{@{text "\<Gamma> \<turnstile> \<Delta>, \<exists>x. P x, P t"}}
+ \qquad
+ \infer[@{text "(\<forall>L)"}]{@{text "\<forall>x. P x, \<Gamma> \<turnstile> \<Delta>"}}{@{text "P t, \<forall>x. P x, \<Gamma> \<turnstile> \<Delta>"}}
+ \]
+ Thus both kinds of quantifier may be replicated. Theorems requiring
+ multiple uses of a universal formula are easy to invent; consider
+ @{text [display] "(\<forall>x. P x \<longrightarrow> P (f x)) \<and> P a \<longrightarrow> P (f\<^sup>n a)"} for any
+ @{text "n > 1"}. Natural examples of the multiple use of an
+ existential formula are rare; a standard one is @{text "\<exists>x. \<forall>y. P x
+ \<longrightarrow> P y"}.
+
+ Forgoing quantifier replication loses completeness, but gains
+ decidability, since the search space becomes finite. Many useful
+ theorems can be proved without replication, and the search generally
+ delivers its verdict in a reasonable time. To adopt this approach,
+ represent the sequent rules @{text "(\<exists>R)"}, @{text "(\<exists>L)"} and
+ @{text "(\<forall>R)"} by @{text "(\<exists>I)"}, @{text "(\<exists>E)"} and @{text "(\<forall>I)"},
+ respectively, and put @{text "(\<forall>E)"} into elimination form: @{text
+ [display] "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> Q) \<Longrightarrow> Q"}
+
+ Elim-resolution with this rule will delete the universal formula
+ after a single use. To replicate universal quantifiers, replace the
+ rule by @{text [display] "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> \<forall>x. P x \<Longrightarrow> Q) \<Longrightarrow> Q"}
+
+ To replicate existential quantifiers, replace @{text "(\<exists>I)"} by
+ @{text [display] "(\<not> (\<exists>x. P x) \<Longrightarrow> P t) \<Longrightarrow> \<exists>x. P x"}
+
+ All introduction rules mentioned above are also useful in swapped
+ form.
+
+ Replication makes the search space infinite; we must apply the rules
+ with care. The classical reasoner distinguishes between safe and
+ unsafe rules, applying the latter only when there is no alternative.
+ Depth-first search may well go down a blind alley; best-first search
+ is better behaved in an infinite search space. However, quantifier
+ replication is too expensive to prove any but the simplest theorems.
+*}
+
+
+subsection {* Rule declarations *}
+
+text {* The proof tools of the Classical Reasoner depend on
+ collections of rules declared in the context, which are classified
+ as introduction, elimination or destruction and as \emph{safe} or
+ \emph{unsafe}. In general, safe rules can be attempted blindly,
+ while unsafe rules must be used with care. A safe rule must never
+ reduce a provable goal to an unprovable set of subgoals.
+
+ The rule @{text "P \<Longrightarrow> P \<or> Q"} is unsafe because it reduces @{text "P
+ \<or> Q"} to @{text "P"}, which might turn out as premature choice of an
+ unprovable subgoal. Any rule is unsafe whose premises contain new
+ unknowns. The elimination rule @{text "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> Q) \<Longrightarrow> Q"} is
+ unsafe, since it is applied via elim-resolution, which discards the
+ assumption @{text "\<forall>x. P x"} and replaces it by the weaker
+ assumption @{text "P t"}. The rule @{text "P t \<Longrightarrow> \<exists>x. P x"} is
+ unsafe for similar reasons. The quantifier duplication rule @{text
+ "\<forall>x. P x \<Longrightarrow> (P t \<Longrightarrow> \<forall>x. P x \<Longrightarrow> Q) \<Longrightarrow> Q"} is unsafe in a different sense:
+ since it keeps the assumption @{text "\<forall>x. P x"}, it is prone to
+ looping. In classical first-order logic, all rules are safe except
+ those mentioned above.
+
+ The safe~/ unsafe distinction is vague, and may be regarded merely
+ as a way of giving some rules priority over others. One could argue
+ that @{text "(\<or>E)"} is unsafe, because repeated application of it
+ could generate exponentially many subgoals. Induction rules are
+ unsafe because inductive proofs are difficult to set up
+ automatically. Any inference is unsafe that instantiates an unknown
+ in the proof state --- thus matching must be used, rather than
+ unification. Even proof by assumption is unsafe if it instantiates
+ unknowns shared with other subgoals.
+
+ \begin{matharray}{rcl}
+ @{command_def "print_claset"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
+ @{attribute_def intro} & : & @{text attribute} \\
+ @{attribute_def elim} & : & @{text attribute} \\
+ @{attribute_def dest} & : & @{text attribute} \\
+ @{attribute_def rule} & : & @{text attribute} \\
+ @{attribute_def iff} & : & @{text attribute} \\
+ @{attribute_def swapped} & : & @{text attribute} \\
+ \end{matharray}
+
+ @{rail \<open>
+ (@@{attribute intro} | @@{attribute elim} | @@{attribute dest}) ('!' | () | '?') @{syntax nat}?
+ ;
+ @@{attribute rule} 'del'
+ ;
+ @@{attribute iff} (((() | 'add') '?'?) | 'del')
+ \<close>}
+
+ \begin{description}
+
+ \item @{command "print_claset"} prints the collection of rules
+ declared to the Classical Reasoner, i.e.\ the @{ML_type claset}
+ within the context.
+
+ \item @{attribute intro}, @{attribute elim}, and @{attribute dest}
+ declare introduction, elimination, and destruction rules,
+ respectively. By default, rules are considered as \emph{unsafe}
+ (i.e.\ not applied blindly without backtracking), while ``@{text
+ "!"}'' classifies as \emph{safe}. Rule declarations marked by
+ ``@{text "?"}'' coincide with those of Isabelle/Pure, cf.\
+ \secref{sec:pure-meth-att} (i.e.\ are only applied in single steps
+ of the @{method rule} method). The optional natural number
+ specifies an explicit weight argument, which is ignored by the
+ automated reasoning tools, but determines the search order of single
+ rule steps.
+
+ Introduction rules are those that can be applied using ordinary
+ resolution. Their swapped forms are generated internally, which
+ will be applied using elim-resolution. Elimination rules are
+ applied using elim-resolution. Rules are sorted by the number of
+ new subgoals they will yield; rules that generate the fewest
+ subgoals will be tried first. Otherwise, later declarations take
+ precedence over earlier ones.
+
+ Rules already present in the context with the same classification
+ are ignored. A warning is printed if the rule has already been
+ added with some other classification, but the rule is added anyway
+ as requested.
+
+ \item @{attribute rule}~@{text del} deletes all occurrences of a
+ rule from the classical context, regardless of its classification as
+ introduction~/ elimination~/ destruction and safe~/ unsafe.
+
+ \item @{attribute iff} declares logical equivalences to the
+ Simplifier and the Classical reasoner at the same time.
+ Non-conditional rules result in a safe introduction and elimination
+ pair; conditional ones are considered unsafe. Rules with negative
+ conclusion are automatically inverted (using @{text "\<not>"}-elimination
+ internally).
+
+ The ``@{text "?"}'' version of @{attribute iff} declares rules to
+ the Isabelle/Pure context only, and omits the Simplifier
+ declaration.
+
+ \item @{attribute swapped} turns an introduction rule into an
+ elimination, by resolving with the classical swap principle @{text
+ "\<not> P \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> R"} in the second position. This is mainly for
+ illustrative purposes: the Classical Reasoner already swaps rules
+ internally as explained above.
+
+ \end{description}
+*}
+
+
+subsection {* Structured methods *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{method_def rule} & : & @{text method} \\
+ @{method_def contradiction} & : & @{text method} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{method rule} @{syntax thmrefs}?
+ \<close>}
+
+ \begin{description}
+
+ \item @{method rule} as offered by the Classical Reasoner is a
+ refinement over the Pure one (see \secref{sec:pure-meth-att}). Both
+ versions work the same, but the classical version observes the
+ classical rule context in addition to that of Isabelle/Pure.
+
+ Common object logics (HOL, ZF, etc.) declare a rich collection of
+ classical rules (even if these would qualify as intuitionistic
+ ones), but only few declarations to the rule context of
+ Isabelle/Pure (\secref{sec:pure-meth-att}).
+
+ \item @{method contradiction} solves some goal by contradiction,
+ deriving any result from both @{text "\<not> A"} and @{text A}. Chained
+ facts, which are guaranteed to participate, may appear in either
+ order.
+
+ \end{description}
+*}
+
+
+subsection {* Fully automated methods *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{method_def blast} & : & @{text method} \\
+ @{method_def auto} & : & @{text method} \\
+ @{method_def force} & : & @{text method} \\
+ @{method_def fast} & : & @{text method} \\
+ @{method_def slow} & : & @{text method} \\
+ @{method_def best} & : & @{text method} \\
+ @{method_def fastforce} & : & @{text method} \\
+ @{method_def slowsimp} & : & @{text method} \\
+ @{method_def bestsimp} & : & @{text method} \\
+ @{method_def deepen} & : & @{text method} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{method blast} @{syntax nat}? (@{syntax clamod} * )
+ ;
+ @@{method auto} (@{syntax nat} @{syntax nat})? (@{syntax clasimpmod} * )
+ ;
+ @@{method force} (@{syntax clasimpmod} * )
+ ;
+ (@@{method fast} | @@{method slow} | @@{method best}) (@{syntax clamod} * )
+ ;
+ (@@{method fastforce} | @@{method slowsimp} | @@{method bestsimp})
+ (@{syntax clasimpmod} * )
+ ;
+ @@{method deepen} (@{syntax nat} ?) (@{syntax clamod} * )
+ ;
+ @{syntax_def clamod}:
+ (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del') ':' @{syntax thmrefs}
+ ;
+ @{syntax_def clasimpmod}: ('simp' (() | 'add' | 'del' | 'only') |
+ ('cong' | 'split') (() | 'add' | 'del') |
+ 'iff' (((() | 'add') '?'?) | 'del') |
+ (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del')) ':' @{syntax thmrefs}
+ \<close>}
+
+ \begin{description}
+
+ \item @{method blast} is a separate classical tableau prover that
+ uses the same classical rule declarations as explained before.
+
+ Proof search is coded directly in ML using special data structures.
+ A successful proof is then reconstructed using regular Isabelle
+ inferences. It is faster and more powerful than the other classical
+ reasoning tools, but has major limitations too.
+
+ \begin{itemize}
+
+ \item It does not use the classical wrapper tacticals, such as the
+ integration with the Simplifier of @{method fastforce}.
+
+ \item It does not perform higher-order unification, as needed by the
+ rule @{thm [source=false] rangeI} in HOL. There are often
+ alternatives to such rules, for example @{thm [source=false]
+ range_eqI}.
+
+ \item Function variables may only be applied to parameters of the
+ subgoal. (This restriction arises because the prover does not use
+ higher-order unification.) If other function variables are present
+ then the prover will fail with the message \texttt{Function Var's
+ argument not a bound variable}.
+
+ \item Its proof strategy is more general than @{method fast} but can
+ be slower. If @{method blast} fails or seems to be running forever,
+ try @{method fast} and the other proof tools described below.
+
+ \end{itemize}
+
+ The optional integer argument specifies a bound for the number of
+ unsafe steps used in a proof. By default, @{method blast} starts
+ with a bound of 0 and increases it successively to 20. In contrast,
+ @{text "(blast lim)"} tries to prove the goal using a search bound
+ of @{text "lim"}. Sometimes a slow proof using @{method blast} can
+ be made much faster by supplying the successful search bound to this
+ proof method instead.
+
+ \item @{method auto} combines classical reasoning with
+ simplification. It is intended for situations where there are a lot
+ of mostly trivial subgoals; it proves all the easy ones, leaving the
+ ones it cannot prove. Occasionally, attempting to prove the hard
+ ones may take a long time.
+
+ The optional depth arguments in @{text "(auto m n)"} refer to its
+ builtin classical reasoning procedures: @{text m} (default 4) is for
+ @{method blast}, which is tried first, and @{text n} (default 2) is
+ for a slower but more general alternative that also takes wrappers
+ into account.
+
+ \item @{method force} is intended to prove the first subgoal
+ completely, using many fancy proof tools and performing a rather
+ exhaustive search. As a result, proof attempts may take rather long
+ or diverge easily.
+
+ \item @{method fast}, @{method best}, @{method slow} attempt to
+ prove the first subgoal using sequent-style reasoning as explained
+ before. Unlike @{method blast}, they construct proofs directly in
+ Isabelle.
+
+ There is a difference in search strategy and back-tracking: @{method
+ fast} uses depth-first search and @{method best} uses best-first
+ search (guided by a heuristic function: normally the total size of
+ the proof state).
+
+ Method @{method slow} is like @{method fast}, but conducts a broader
+ search: it may, when backtracking from a failed proof attempt, undo
+ even the step of proving a subgoal by assumption.
+
+ \item @{method fastforce}, @{method slowsimp}, @{method bestsimp}
+ are like @{method fast}, @{method slow}, @{method best},
+ respectively, but use the Simplifier as additional wrapper. The name
+ @{method fastforce}, reflects the behaviour of this popular method
+ better without requiring an understanding of its implementation.
+
+ \item @{method deepen} works by exhaustive search up to a certain
+ depth. The start depth is 4 (unless specified explicitly), and the
+ depth is increased iteratively up to 10. Unsafe rules are modified
+ to preserve the formula they act on, so that it be used repeatedly.
+ This method can prove more goals than @{method fast}, but is much
+ slower, for example if the assumptions have many universal
+ quantifiers.
+
+ \end{description}
+
+ Any of the above methods support additional modifiers of the context
+ of classical (and simplifier) rules, but the ones related to the
+ Simplifier are explicitly prefixed by @{text simp} here. The
+ semantics of these ad-hoc rule declarations is analogous to the
+ attributes given before. Facts provided by forward chaining are
+ inserted into the goal before commencing proof search.
+*}
+
+
+subsection {* Partially automated methods *}
+
+text {* These proof methods may help in situations when the
+ fully-automated tools fail. The result is a simpler subgoal that
+ can be tackled by other means, such as by manual instantiation of
+ quantifiers.
+
+ \begin{matharray}{rcl}
+ @{method_def safe} & : & @{text method} \\
+ @{method_def clarify} & : & @{text method} \\
+ @{method_def clarsimp} & : & @{text method} \\
+ \end{matharray}
+
+ @{rail \<open>
+ (@@{method safe} | @@{method clarify}) (@{syntax clamod} * )
+ ;
+ @@{method clarsimp} (@{syntax clasimpmod} * )
+ \<close>}
+
+ \begin{description}
+
+ \item @{method safe} repeatedly performs safe steps on all subgoals.
+ It is deterministic, with at most one outcome.
+
+ \item @{method clarify} performs a series of safe steps without
+ splitting subgoals; see also @{method clarify_step}.
+
+ \item @{method clarsimp} acts like @{method clarify}, but also does
+ simplification. Note that if the Simplifier context includes a
+ splitter for the premises, the subgoal may still be split.
+
+ \end{description}
+*}
+
+
+subsection {* Single-step tactics *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{method_def safe_step} & : & @{text method} \\
+ @{method_def inst_step} & : & @{text method} \\
+ @{method_def step} & : & @{text method} \\
+ @{method_def slow_step} & : & @{text method} \\
+ @{method_def clarify_step} & : & @{text method} \\
+ \end{matharray}
+
+ These are the primitive tactics behind the automated proof methods
+ of the Classical Reasoner. By calling them yourself, you can
+ execute these procedures one step at a time.
+
+ \begin{description}
+
+ \item @{method safe_step} performs a safe step on the first subgoal.
+ The safe wrapper tacticals are applied to a tactic that may include
+ proof by assumption or Modus Ponens (taking care not to instantiate
+ unknowns), or substitution.
+
+ \item @{method inst_step} is like @{method safe_step}, but allows
+ unknowns to be instantiated.
+
+ \item @{method step} is the basic step of the proof procedure, it
+ operates on the first subgoal. The unsafe wrapper tacticals are
+ applied to a tactic that tries @{method safe}, @{method inst_step},
+ or applies an unsafe rule from the context.
+
+ \item @{method slow_step} resembles @{method step}, but allows
+ backtracking between using safe rules with instantiation (@{method
+ inst_step}) and using unsafe rules. The resulting search space is
+ larger.
+
+ \item @{method clarify_step} performs a safe step on the first
+ subgoal; no splitting step is applied. For example, the subgoal
+ @{text "A \<and> B"} is left as a conjunction. Proof by assumption,
+ Modus Ponens, etc., may be performed provided they do not
+ instantiate unknowns. Assumptions of the form @{text "x = t"} may
+ be eliminated. The safe wrapper tactical is applied.
+
+ \end{description}
+*}
+
+
+subsection {* Modifying the search step *}
+
+text {*
+ \begin{mldecls}
+ @{index_ML_type wrapper: "(int -> tactic) -> (int -> tactic)"} \\[0.5ex]
+ @{index_ML_op addSWrapper: "Proof.context *
+ (string * (Proof.context -> wrapper)) -> Proof.context"} \\
+ @{index_ML_op addSbefore: "Proof.context *
+ (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\
+ @{index_ML_op addSafter: "Proof.context *
+ (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\
+ @{index_ML_op delSWrapper: "Proof.context * string -> Proof.context"} \\[0.5ex]
+ @{index_ML_op addWrapper: "Proof.context *
+ (string * (Proof.context -> wrapper)) -> Proof.context"} \\
+ @{index_ML_op addbefore: "Proof.context *
+ (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\
+ @{index_ML_op addafter: "Proof.context *
+ (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\
+ @{index_ML_op delWrapper: "Proof.context * string -> Proof.context"} \\[0.5ex]
+ @{index_ML addSss: "Proof.context -> Proof.context"} \\
+ @{index_ML addss: "Proof.context -> Proof.context"} \\
+ \end{mldecls}
+
+ The proof strategy of the Classical Reasoner is simple. Perform as
+ many safe inferences as possible; or else, apply certain safe rules,
+ allowing instantiation of unknowns; or else, apply an unsafe rule.
+ The tactics also eliminate assumptions of the form @{text "x = t"}
+ by substitution if they have been set up to do so. They may perform
+ a form of Modus Ponens: if there are assumptions @{text "P \<longrightarrow> Q"} and
+ @{text "P"}, then replace @{text "P \<longrightarrow> Q"} by @{text "Q"}.
+
+ The classical reasoning tools --- except @{method blast} --- allow
+ to modify this basic proof strategy by applying two lists of
+ arbitrary \emph{wrapper tacticals} to it. The first wrapper list,
+ which is considered to contain safe wrappers only, affects @{method
+ safe_step} and all the tactics that call it. The second one, which
+ may contain unsafe wrappers, affects the unsafe parts of @{method
+ step}, @{method slow_step}, and the tactics that call them. A
+ wrapper transforms each step of the search, for example by
+ attempting other tactics before or after the original step tactic.
+ All members of a wrapper list are applied in turn to the respective
+ step tactic.
+
+ Initially the two wrapper lists are empty, which means no
+ modification of the step tactics. Safe and unsafe wrappers are added
+ to a claset with the functions given below, supplying them with
+ wrapper names. These names may be used to selectively delete
+ wrappers.
+
+ \begin{description}
+
+ \item @{text "ctxt addSWrapper (name, wrapper)"} adds a new wrapper,
+ which should yield a safe tactic, to modify the existing safe step
+ tactic.
+
+ \item @{text "ctxt addSbefore (name, tac)"} adds the given tactic as a
+ safe wrapper, such that it is tried \emph{before} each safe step of
+ the search.
+
+ \item @{text "ctxt addSafter (name, tac)"} adds the given tactic as a
+ safe wrapper, such that it is tried when a safe step of the search
+ would fail.
+
+ \item @{text "ctxt delSWrapper name"} deletes the safe wrapper with
+ the given name.
+
+ \item @{text "ctxt addWrapper (name, wrapper)"} adds a new wrapper to
+ modify the existing (unsafe) step tactic.
+
+ \item @{text "ctxt addbefore (name, tac)"} adds the given tactic as an
+ unsafe wrapper, such that it its result is concatenated
+ \emph{before} the result of each unsafe step.
+
+ \item @{text "ctxt addafter (name, tac)"} adds the given tactic as an
+ unsafe wrapper, such that it its result is concatenated \emph{after}
+ the result of each unsafe step.
+
+ \item @{text "ctxt delWrapper name"} deletes the unsafe wrapper with
+ the given name.
+
+ \item @{text "addSss"} adds the simpset of the context to its
+ classical set. The assumptions and goal will be simplified, in a
+ rather safe way, after each safe step of the search.
+
+ \item @{text "addss"} adds the simpset of the context to its
+ classical set. The assumptions and goal will be simplified, before
+ the each unsafe step of the search.
+
+ \end{description}
+*}
+
+
+section {* Object-logic setup \label{sec:object-logic} *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{command_def "judgment"} & : & @{text "theory \<rightarrow> theory"} \\
+ @{method_def atomize} & : & @{text method} \\
+ @{attribute_def atomize} & : & @{text attribute} \\
+ @{attribute_def rule_format} & : & @{text attribute} \\
+ @{attribute_def rulify} & : & @{text attribute} \\
+ \end{matharray}
+
+ The very starting point for any Isabelle object-logic is a ``truth
+ judgment'' that links object-level statements to the meta-logic
+ (with its minimal language of @{text prop} that covers universal
+ quantification @{text "\<And>"} and implication @{text "\<Longrightarrow>"}).
+
+ Common object-logics are sufficiently expressive to internalize rule
+ statements over @{text "\<And>"} and @{text "\<Longrightarrow>"} within their own
+ language. This is useful in certain situations where a rule needs
+ to be viewed as an atomic statement from the meta-level perspective,
+ e.g.\ @{text "\<And>x. x \<in> A \<Longrightarrow> P x"} versus @{text "\<forall>x \<in> A. P x"}.
+
+ From the following language elements, only the @{method atomize}
+ method and @{attribute rule_format} attribute are occasionally
+ required by end-users, the rest is for those who need to setup their
+ own object-logic. In the latter case existing formulations of
+ Isabelle/FOL or Isabelle/HOL may be taken as realistic examples.
+
+ Generic tools may refer to the information provided by object-logic
+ declarations internally.
+
+ @{rail \<open>
+ @@{command judgment} @{syntax name} '::' @{syntax type} @{syntax mixfix}?
+ ;
+ @@{attribute atomize} ('(' 'full' ')')?
+ ;
+ @@{attribute rule_format} ('(' 'noasm' ')')?
+ \<close>}
+
+ \begin{description}
+
+ \item @{command "judgment"}~@{text "c :: \<sigma> (mx)"} declares constant
+ @{text c} as the truth judgment of the current object-logic. Its
+ type @{text \<sigma>} should specify a coercion of the category of
+ object-level propositions to @{text prop} of the Pure meta-logic;
+ the mixfix annotation @{text "(mx)"} would typically just link the
+ object language (internally of syntactic category @{text logic})
+ with that of @{text prop}. Only one @{command "judgment"}
+ declaration may be given in any theory development.
+
+ \item @{method atomize} (as a method) rewrites any non-atomic
+ premises of a sub-goal, using the meta-level equations declared via
+ @{attribute atomize} (as an attribute) beforehand. As a result,
+ heavily nested goals become amenable to fundamental operations such
+ as resolution (cf.\ the @{method (Pure) rule} method). Giving the ``@{text
+ "(full)"}'' option here means to turn the whole subgoal into an
+ object-statement (if possible), including the outermost parameters
+ and assumptions as well.
+
+ A typical collection of @{attribute atomize} rules for a particular
+ object-logic would provide an internalization for each of the
+ connectives of @{text "\<And>"}, @{text "\<Longrightarrow>"}, and @{text "\<equiv>"}.
+ Meta-level conjunction should be covered as well (this is
+ particularly important for locales, see \secref{sec:locale}).
+
+ \item @{attribute rule_format} rewrites a theorem by the equalities
+ declared as @{attribute rulify} rules in the current object-logic.
+ By default, the result is fully normalized, including assumptions
+ and conclusions at any depth. The @{text "(no_asm)"} option
+ restricts the transformation to the conclusion of a rule.
+
+ In common object-logics (HOL, FOL, ZF), the effect of @{attribute
+ rule_format} is to replace (bounded) universal quantification
+ (@{text "\<forall>"}) and implication (@{text "\<longrightarrow>"}) by the corresponding
+ rule statements over @{text "\<And>"} and @{text "\<Longrightarrow>"}.
+
+ \end{description}
+*}
+
+
+section {* Tracing higher-order unification *}
+
+text {*
+ \begin{tabular}{rcll}
+ @{attribute_def unify_trace_simp} & : & @{text "attribute"} & default @{text "false"} \\
+ @{attribute_def unify_trace_types} & : & @{text "attribute"} & default @{text "false"} \\
+ @{attribute_def unify_trace_bound} & : & @{text "attribute"} & default @{text "50"} \\
+ @{attribute_def unify_search_bound} & : & @{text "attribute"} & default @{text "60"} \\
+ \end{tabular}
+ \medskip
+
+ Higher-order unification works well in most practical situations,
+ but sometimes needs extra care to identify problems. These tracing
+ options may help.
+
+ \begin{description}
+
+ \item @{attribute unify_trace_simp} controls tracing of the
+ simplification phase of higher-order unification.
+
+ \item @{attribute unify_trace_types} controls warnings of
+ incompleteness, when unification is not considering all possible
+ instantiations of schematic type variables.
+
+ \item @{attribute unify_trace_bound} determines the depth where
+ unification starts to print tracing information once it reaches
+ depth; 0 for full tracing. At the default value, tracing
+ information is almost never printed in practice.
+
+ \item @{attribute unify_search_bound} prevents unification from
+ searching past the given depth. Because of this bound, higher-order
+ unification cannot return an infinite sequence, though it can return
+ an exponentially long one. The search rarely approaches the default
+ value in practice. If the search is cut off, unification prints a
+ warning ``Unification bound exceeded''.
+
+ \end{description}
+
+ \begin{warn}
+ Options for unification cannot be modified in a local context. Only
+ the global theory content is taken into account.
+ \end{warn}
+*}
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Doc/Isar-Ref/HOL_Specific.thy Sat Apr 05 11:37:00 2014 +0200
@@ -0,0 +1,2677 @@
+theory HOL_Specific
+imports Base Main "~~/src/HOL/Library/Old_Recdef" "~~/src/Tools/Adhoc_Overloading"
+begin
+
+chapter {* Higher-Order Logic *}
+
+text {* 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{weaker} 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 @{text "\<lambda>"}-calculus and
+ functional programming. Function application is curried. To apply
+ the function @{text f} of type @{text "\<tau>\<^sub>1 \<Rightarrow> \<tau>\<^sub>2 \<Rightarrow> \<tau>\<^sub>3"} to the
+ arguments @{text a} and @{text b} in HOL, you simply write @{text "f
+ a b"} (as in ML or Haskell). There is no ``apply'' operator; the
+ existing application of the Pure @{text "\<lambda>"}-calculus is re-used.
+ Note that in HOL @{text "f (a, b)"} means ``@{text "f"} applied to
+ the pair @{text "(a, b)"} (which is notation for @{text "Pair a
+ b"}). 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. Beginners can set
+ @{attribute show_types} or even @{attribute show_sorts} to get more
+ explicit information about the result of type-inference. *}
+
+
+chapter {* Derived specification elements *}
+
+section {* Inductive and coinductive definitions \label{sec:hol-inductive} *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{command_def (HOL) "inductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+ @{command_def (HOL) "inductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+ @{command_def (HOL) "coinductive"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+ @{command_def (HOL) "coinductive_set"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+ @{command_def "print_inductives"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
+ @{attribute_def (HOL) mono} & : & @{text attribute} \\
+ \end{matharray}
+
+ An \emph{inductive definition} specifies the least predicate or set
+ @{text R} closed under given rules: applying a rule to elements of
+ @{text R} yields a result within @{text R}. For example, a
+ structural operational semantics is an inductive definition of an
+ evaluation relation.
+
+ Dually, a \emph{coinductive definition} specifies the greatest
+ predicate or set @{text R} that is consistent with given rules:
+ every element of @{text R} can be seen as arising by applying a rule
+ to elements of @{text R}. 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
+ @{text "A \<le> B \<Longrightarrow> \<M> A \<le> \<M> B"}, for each premise @{text "\<M> R t"} 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 target}? \<newline>
+ @{syntax "fixes"} (@'for' @{syntax "fixes"})? (@'where' clauses)? \<newline>
+ (@'monos' @{syntax thmrefs})?
+ ;
+ clauses: (@{syntax thmdecl}? @{syntax prop} + '|')
+ ;
+ @@{attribute (HOL) mono} (() | 'add' | 'del')
+ \<close>}
+
+ \begin{description}
+
+ \item @{command (HOL) "inductive"} and @{command (HOL)
+ "coinductive"} define (co)inductive predicates from the introduction
+ rules.
+
+ The propositions given as @{text "clauses"} in the @{keyword
+ "where"} part are either rules of the usual @{text "\<And>/\<Longrightarrow>"} format
+ (with arbitrary nesting), or equalities using @{text "\<equiv>"}. 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
+ definition, in contrast to arguments of the relation that may vary
+ in each occurrence within the given @{text "clauses"}.
+
+ The optional @{keyword "monos"} declaration contains additional
+ \emph{monotonicity theorems}, which are required for each operator
+ applied to a recursive set in the introduction rules.
+
+ \item @{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.
+
+ \item @{command "print_inductives"} prints (co)inductive definitions and
+ monotonicity rules.
+
+ \item @{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.
+
+ \end{description}
+*}
+
+
+subsection {* Derived rules *}
+
+text {* A (co)inductive definition of @{text R} provides the following
+ main theorems:
+
+ \begin{description}
+
+ \item @{text R.intros} 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;
+
+ \item @{text R.cases} is the case analysis (or elimination) rule;
+
+ \item @{text R.induct} or @{text R.coinduct} is the (co)induction
+ rule;
+
+ \item @{text R.simps} is the equation unrolling the fixpoint of the
+ predicate one step.
+
+ \end{description}
+
+ When several predicates @{text "R\<^sub>1, \<dots>, R\<^sub>n"} are
+ defined simultaneously, the list of introduction rules is called
+ @{text "R\<^sub>1_\<dots>_R\<^sub>n.intros"}, the case analysis rules are
+ called @{text "R\<^sub>1.cases, \<dots>, R\<^sub>n.cases"}, and the list
+ of mutual induction rules is called @{text
+ "R\<^sub>1_\<dots>_R\<^sub>n.inducts"}.
+*}
+
+
+subsection {* Monotonicity theorems *}
+
+text {* 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:
+
+ \begin{itemize}
+
+ \item Theorems of the form @{text "A \<le> B \<Longrightarrow> \<M> A \<le> \<M> B"}, for proving
+ monotonicity of inductive definitions whose introduction rules have
+ premises involving terms such as @{text "\<M> R t"}.
+
+ \item Monotonicity theorems for logical operators, which are of the
+ general form @{text "(\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> (\<dots> \<longrightarrow> \<dots>) \<Longrightarrow> \<dots> \<longrightarrow> \<dots>"}. For example, in
+ the case of the operator @{text "\<or>"}, the corresponding theorem is
+ \[
+ \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"}}
+ \]
+
+ \item De Morgan style equations for reasoning about the ``polarity''
+ of expressions, e.g.
+ \[
+ @{prop "\<not> \<not> P \<longleftrightarrow> P"} \qquad\qquad
+ @{prop "\<not> (P \<and> Q) \<longleftrightarrow> \<not> P \<or> \<not> Q"}
+ \]
+
+ \item Equations for reducing complex operators to more primitive
+ ones whose monotonicity can easily be proved, e.g.
+ \[
+ @{prop "(P \<longrightarrow> Q) \<longleftrightarrow> \<not> P \<or> Q"} \qquad\qquad
+ @{prop "Ball A P \<equiv> \<forall>x. x \<in> A \<longrightarrow> P x"}
+ \]
+
+ \end{itemize}
+*}
+
+subsubsection {* Examples *}
+
+text {* The finite powerset operator can be defined inductively like this: *}
+
+inductive_set Fin :: "'a set \<Rightarrow> 'a set set" for A :: "'a set"
+where
+ empty: "{} \<in> Fin A"
+| insert: "a \<in> A \<Longrightarrow> B \<in> Fin A \<Longrightarrow> insert a B \<in> Fin A"
+
+text {* The accessible part of a relation is defined as follows: *}
+
+inductive acc :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
+ for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<prec>" 50)
+where acc: "(\<And>y. y \<prec> x \<Longrightarrow> acc r y) \<Longrightarrow> acc r x"
+
+text {* Common logical connectives can be easily characterized as
+non-recursive inductive definitions with parameters, but without
+arguments. *}
+
+inductive AND for A B :: bool
+where "A \<Longrightarrow> B \<Longrightarrow> AND A B"
+
+inductive OR for A B :: bool
+where "A \<Longrightarrow> OR A B"
+ | "B \<Longrightarrow> OR A B"
+
+inductive EXISTS for B :: "'a \<Rightarrow> bool"
+where "B a \<Longrightarrow> EXISTS B"
+
+text {* Here the @{text "cases"} or @{text "induct"} 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. *}
+
+
+section {* Recursive functions \label{sec:recursion} *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{command_def (HOL) "primrec"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+ @{command_def (HOL) "fun"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+ @{command_def (HOL) "function"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
+ @{command_def (HOL) "termination"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
+ @{command_def (HOL) "fun_cases"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{command (HOL) primrec} @{syntax target}? @{syntax "fixes"} @'where' equations
+ ;
+ (@@{command (HOL) fun} | @@{command (HOL) function}) @{syntax target}? functionopts?
+ @{syntax "fixes"} \<newline> @'where' equations
+ ;
+
+ equations: (@{syntax thmdecl}? @{syntax prop} + '|')
+ ;
+ functionopts: '(' (('sequential' | 'domintros') + ',') ')'
+ ;
+ @@{command (HOL) termination} @{syntax term}?
+ ;
+ @@{command (HOL) fun_cases} (@{syntax thmdecl}? @{syntax prop} + @'and')
+ \<close>}
+
+ \begin{description}
+
+ \item @{command (HOL) "primrec"} defines primitive recursive
+ functions over datatypes (see also @{command_ref (HOL) datatype} and
+ @{command_ref (HOL) rep_datatype}). The given @{text equations}
+ 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 \<dots> x\<^sub>m (C y\<^sub>1 \<dots> y\<^sub>k) z\<^sub>1 \<dots> z\<^sub>n = rhs"}
+
+ such that @{text C} is a datatype constructor, @{text rhs} contains
+ only the free variables on the left-hand side (or from the context),
+ and all recursive occurrences of @{text "f"} in @{text "rhs"} are of
+ the form @{text "f \<dots> y\<^sub>i \<dots>"} for some @{text i}. 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 @{text "f"} applied to
+ datatype constructions, by simulating symbolic computation via
+ rewriting.
+
+ \item @{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 @{text "f.psimps"}) and induction rule
+ (named @{text "f.pinduct"}) are guarded by a generated domain
+ predicate @{text "f_dom"}. The @{command (HOL) "termination"}
+ command can then be used to establish that the function is total.
+
+ \item @{command (HOL) "fun"} is a shorthand notation for ``@{command
+ (HOL) "function"}~@{text "(sequential)"}, followed by automated
+ proof attempts regarding pattern matching and termination. See
+ \cite{isabelle-function} for further details.
+
+ \item @{command (HOL) "termination"}~@{text f} commences a
+ termination proof for the previously defined function @{text f}. 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.
+
+ \item @{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.
+ \end{description}
+
+ Recursive definitions introduced by the @{command (HOL) "function"}
+ command accommodate reasoning by induction (cf.\ @{method induct}):
+ rule @{text "f.induct"} 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 @{text "f.simps"}, where @{text f} 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.
+
+ \begin{description}
+
+ \item @{text sequential} 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 theorems. Also note that this automatic
+ transformation only works for ML-style datatype patterns.
+
+ \item @{text domintros} 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.
+
+ \end{description}
+*}
+
+subsubsection {* Example: evaluation of expressions *}
+
+text {* Subsequently, we define mutual datatypes for arithmetic and
+ boolean expressions, and use @{command primrec} for evaluation
+ functions that follow the same recursive structure. *}
+
+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 {* \medskip Evaluation of arithmetic and boolean expressions *}
+
+primrec evala :: "('a \<Rightarrow> nat) \<Rightarrow> 'a aexp \<Rightarrow> nat"
+ and evalb :: "('a \<Rightarrow> nat) \<Rightarrow> 'a bexp \<Rightarrow> 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 \<and> evalb env b2)"
+| "evalb env (Neg b) = (\<not> evalb env b)"
+
+text {* Since the value of an expression depends on the value of its
+ variables, the functions @{const evala} and @{const evalb} take an
+ additional parameter, an \emph{environment} that maps variables to
+ their values.
+
+ \medskip Substitution on expressions can be defined similarly. The
+ mapping @{text f} of type @{typ "'a \<Rightarrow> 'a aexp"} given as a
+ parameter is lifted canonically on the types @{typ "'a aexp"} and
+ @{typ "'a bexp"}, respectively.
+*}
+
+primrec substa :: "('a \<Rightarrow> 'b aexp) \<Rightarrow> 'a aexp \<Rightarrow> 'b aexp"
+ and substb :: "('a \<Rightarrow> 'b aexp) \<Rightarrow> 'a bexp \<Rightarrow> '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 {* In textbooks about semantics one often finds substitution
+ theorems, which express the relationship between substitution and
+ evaluation. For @{typ "'a aexp"} and @{typ "'a bexp"}, we can prove
+ such a theorem by mutual induction, followed by simplification.
+*}
+
+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 (\<lambda>x. evala env (s x)) a"
+ "evalb env (substb s b) = evalb (\<lambda>x. evala env (s x)) b"
+ by (induct a and b) simp_all
+
+
+subsubsection {* Example: a substitution function for terms *}
+
+text {* Functions on datatypes with nested recursion are also defined
+ by mutual primitive recursion. *}
+
+datatype ('a, 'b) "term" = Var 'a | App 'b "('a, 'b) term list"
+
+text {* A substitution function on type @{typ "('a, 'b) term"} can be
+ defined as follows, by working simultaneously on @{typ "('a, 'b)
+ term list"}: *}
+
+primrec subst_term :: "('a \<Rightarrow> ('a, 'b) term) \<Rightarrow> ('a, 'b) term \<Rightarrow> ('a, 'b) term" and
+ subst_term_list :: "('a \<Rightarrow> ('a, 'b) term) \<Rightarrow> ('a, 'b) term list \<Rightarrow> ('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 {* The recursion scheme follows the structure of the unfolded
+ definition of type @{typ "('a, 'b) term"}. To prove properties of this
+ substitution function, mutual induction is needed:
+*}
+
+lemma "subst_term (subst_term f1 \<circ> f2) t = subst_term f1 (subst_term f2 t)" and
+ "subst_term_list (subst_term f1 \<circ> f2) ts = subst_term_list f1 (subst_term_list f2 ts)"
+ by (induct t and ts) simp_all
+
+
+subsubsection {* Example: a map function for infinitely branching trees *}
+
+text {* Defining functions on infinitely branching datatypes by
+ primitive recursion is just as easy.
+*}
+
+datatype 'a tree = Atom 'a | Branch "nat \<Rightarrow> 'a tree"
+
+primrec map_tree :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a tree \<Rightarrow> 'b tree"
+where
+ "map_tree f (Atom a) = Atom (f a)"
+| "map_tree f (Branch ts) = Branch (\<lambda>x. map_tree f (ts x))"
+
+text {* Note that all occurrences of functions such as @{text ts}
+ above must be applied to an argument. In particular, @{term
+ "map_tree f \<circ> ts"} is not allowed here. *}
+
+text {* Here is a simple composition lemma for @{term map_tree}: *}
+
+lemma "map_tree g (map_tree f t) = map_tree (g \<circ> f) t"
+ by (induct t) simp_all
+
+
+subsection {* Proof methods related to recursive definitions *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{method_def (HOL) pat_completeness} & : & @{text method} \\
+ @{method_def (HOL) relation} & : & @{text method} \\
+ @{method_def (HOL) lexicographic_order} & : & @{text method} \\
+ @{method_def (HOL) size_change} & : & @{text method} \\
+ @{method_def (HOL) induction_schema} & : & @{text method} \\
+ \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>}
+
+ \begin{description}
+
+ \item @{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}).
+
+ \item @{method (HOL) relation}~@{text R} introduces a termination
+ proof using the relation @{text R}. The resulting proof state will
+ contain goals expressing that @{text R} is wellfounded, and that the
+ arguments of recursive calls decrease with respect to @{text R}.
+ Usually, this method is used as the initial proof step of manual
+ termination proofs.
+
+ \item @{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}).
+
+ \item @{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: @{text max}, @{text min},
+ and @{text ms} (multiset), which is only available when the theory
+ @{text Multiset} 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}).
+
+ \item @{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 "~~/src/HOL/ex/Induction_Schema.thy"} for examples.
+
+ \end{description}
+*}
+
+
+subsection {* Functions with explicit partiality *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{command_def (HOL) "partial_function"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+ @{attribute_def (HOL) "partial_function_mono"} & : & @{text attribute} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{command (HOL) partial_function} @{syntax target}?
+ '(' @{syntax nameref} ')' @{syntax "fixes"} \<newline>
+ @'where' @{syntax thmdecl}? @{syntax prop}
+ \<close>}
+
+ \begin{description}
+
+ \item @{command (HOL) "partial_function"}~@{text "(mode)"} 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 @{text mode} 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:
+
+ \begin{description}
+
+ \item @{text option} defines functions that map into the @{type
+ option} type. Here, the value @{term None} is used to model a
+ non-terminating computation. Monotonicity requires that if @{term
+ None} is returned by a recursive call, then the overall result must
+ also be @{term None}. This is best achieved through the use of the
+ monadic operator @{const "Option.bind"}.
+
+ \item @{text tailrec} defines functions with an arbitrary result
+ type and uses the slightly degenerated partial order where @{term
+ "undefined"} is the bottom element. Now, monotonicity requires that
+ if @{term undefined} is returned by a recursive call, then the
+ overall result must also be @{term undefined}. 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.
+
+ \end{description}
+
+ Experienced users may define new modes by instantiating the locale
+ @{const "partial_function_definitions"} appropriately.
+
+ \item @{attribute (HOL) partial_function_mono} declares rules for
+ use in the internal monotonicity proofs of partial function
+ definitions.
+
+ \end{description}
+
+*}
+
+
+subsection {* Old-style recursive function definitions (TFL) *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{command_def (HOL) "recdef"} & : & @{text "theory \<rightarrow> theory)"} \\
+ @{command_def (HOL) "recdef_tc"}@{text "\<^sup>*"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
+ \end{matharray}
+
+ The old TFL commands @{command (HOL) "recdef"} and @{command (HOL)
+ "recdef_tc"} for defining recursive are 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?
+ ;
+ recdeftc @{syntax thmdecl}? tc
+ ;
+ hints: '(' @'hints' ( recdefmod * ) ')'
+ ;
+ recdefmod: (('recdef_simp' | 'recdef_cong' | 'recdef_wf')
+ (() | 'add' | 'del') ':' @{syntax thmrefs}) | @{syntax clasimpmod}
+ ;
+ tc: @{syntax nameref} ('(' @{syntax nat} ')')?
+ \<close>}
+
+ \begin{description}
+
+ \item @{command (HOL) "recdef"} defines general well-founded
+ recursive functions (using the TFL package), see also
+ \cite{isabelle-HOL}. The ``@{text "(permissive)"}'' option tells
+ TFL to recover from failed proof attempts, returning unfinished
+ results. The @{text recdef_simp}, @{text recdef_cong}, and @{text
+ recdef_wf} 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}).
+
+ \item @{command (HOL) "recdef_tc"}~@{text "c (i)"} recommences the
+ proof for leftover termination condition number @{text i} (default
+ 1) as generated by a @{command (HOL) "recdef"} definition of
+ constant @{text c}.
+
+ Note that in most cases, @{command (HOL) "recdef"} is able to finish
+ its internal proofs without manual intervention.
+
+ \end{description}
+
+ \medskip Hints for @{command (HOL) "recdef"} may be also declared
+ globally, using the following attributes.
+
+ \begin{matharray}{rcl}
+ @{attribute_def (HOL) recdef_simp} & : & @{text attribute} \\
+ @{attribute_def (HOL) recdef_cong} & : & @{text attribute} \\
+ @{attribute_def (HOL) recdef_wf} & : & @{text attribute} \\
+ \end{matharray}
+
+ @{rail \<open>
+ (@@{attribute (HOL) recdef_simp} | @@{attribute (HOL) recdef_cong} |
+ @@{attribute (HOL) recdef_wf}) (() | 'add' | 'del')
+ \<close>}
+*}
+
+
+section {* Datatypes \label{sec:hol-datatype} *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{command_def (HOL) "datatype"} & : & @{text "theory \<rightarrow> theory"} \\
+ @{command_def (HOL) "rep_datatype"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{command (HOL) datatype} (spec + @'and')
+ ;
+ @@{command (HOL) rep_datatype} ('(' (@{syntax name} +) ')')? (@{syntax term} +)
+ ;
+
+ spec: @{syntax typespec_sorts} @{syntax mixfix}? '=' (cons + '|')
+ ;
+ cons: @{syntax name} (@{syntax type} * ) @{syntax mixfix}?
+ \<close>}
+
+ \begin{description}
+
+ \item @{command (HOL) "datatype"} defines inductive datatypes in
+ HOL.
+
+ \item @{command (HOL) "rep_datatype"} represents existing types as
+ datatypes.
+
+ For foundational reasons, some basic types such as @{typ nat}, @{typ
+ "'a \<times> 'b"}, @{typ "'a + 'b"}, @{typ bool} and @{typ unit} are
+ introduced by more primitive means using @{command_ref typedef}. To
+ recover the rich infrastructure of @{command datatype} (e.g.\ rules
+ for @{method cases} and @{method induct} and the primitive recursion
+ combinators), such types may be represented as actual datatypes
+ later. This is done by specifying the constructors of the desired
+ type, and giving a proof of the induction rule, distinctness and
+ injectivity of constructors.
+
+ For example, see @{file "~~/src/HOL/Sum_Type.thy"} for the
+ representation of the primitive sum type as fully-featured datatype.
+
+ \end{description}
+
+ The generated rules for @{method induct} and @{method cases} provide
+ case names according to the given constructors, while parameters are
+ named after the types (see also \secref{sec:cases-induct}).
+
+ See \cite{isabelle-HOL} for more details on datatypes, but beware of
+ the old-style theory syntax being used there! 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).
+*}
+
+
+subsubsection {* Examples *}
+
+text {* We define a type of finite sequences, with slightly different
+ names than the existing @{typ "'a list"} that is already in @{theory
+ Main}: *}
+
+datatype 'a seq = Empty | Seq 'a "'a seq"
+
+text {* We can now prove some simple lemma by structural induction: *}
+
+lemma "Seq x xs \<noteq> xs"
+proof (induct xs arbitrary: x)
+ case Empty
+ txt {* This case can be proved using the simplifier: the freeness
+ properties of the datatype are already declared as @{attribute
+ simp} rules. *}
+ show "Seq x Empty \<noteq> Empty"
+ by simp
+next
+ case (Seq y ys)
+ txt {* The step case is proved similarly. *}
+ show "Seq x (Seq y ys) \<noteq> Seq y ys"
+ using `Seq y ys \<noteq> ys` by simp
+qed
+
+text {* Here is a more succinct version of the same proof: *}
+
+lemma "Seq x xs \<noteq> xs"
+ by (induct xs arbitrary: x) simp_all
+
+
+section {* Records \label{sec:hol-record} *}
+
+text {*
+ 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.
+*}
+
+
+subsection {* Basic concepts *}
+
+text {*
+ Isabelle/HOL supports both \emph{fixed} and \emph{schematic} 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 & @{text "\<lparr>x = a, y = b\<rparr>"} & @{text "\<lparr>x :: A, y :: B\<rparr>"} \\
+ schematic & @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} &
+ @{text "\<lparr>x :: A, y :: B, \<dots> :: M\<rparr>"} \\
+ \end{tabular}
+ \end{center}
+
+ \noindent The ASCII representation of @{text "\<lparr>x = a\<rparr>"} is @{text
+ "(| x = a |)"}.
+
+ A fixed record @{text "\<lparr>x = a, y = b\<rparr>"} has field @{text x} of value
+ @{text a} and field @{text y} of value @{text b}. The corresponding
+ type is @{text "\<lparr>x :: A, y :: B\<rparr>"}, assuming that @{text "a :: A"}
+ and @{text "b :: B"}.
+
+ A record scheme like @{text "\<lparr>x = a, y = b, \<dots> = m\<rparr>"} contains fields
+ @{text x} and @{text y} as before, but also possibly further fields
+ as indicated by the ``@{text "\<dots>"}'' notation (which is actually part
+ of the syntax). The improper field ``@{text "\<dots>"}'' of a record
+ scheme is called the \emph{more part}. 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 @{text "\<lparr>x = a, y = b, z =
+ c, \<dots> = m'\<rparr>"}, where @{text m'} refers to a different more part.
+ Fixed records are special instances of record schemes, where
+ ``@{text "\<dots>"}'' is properly terminated by the @{text "() :: unit"}
+ element. In fact, @{text "\<lparr>x = a, y = b\<rparr>"} is just an abbreviation
+ for @{text "\<lparr>x = a, y = b, \<dots> = ()\<rparr>"}.
+
+ \medskip Two key observations make extensible records in a simply
+ typed language like HOL work out:
+
+ \begin{enumerate}
+
+ \item the more part is internalized, as a free term or type
+ variable,
+
+ \item field names are externalized, they cannot be accessed within
+ the logic as first-class values.
+
+ \end{enumerate}
+
+ \medskip In Isabelle/HOL record types have to be defined explicitly,
+ fixing their field names and types, and 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.
+*}
+
+
+subsection {* Record specifications *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{command_def (HOL) "record"} & : & @{text "theory \<rightarrow> theory"} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{command (HOL) record} @{syntax typespec_sorts} '=' \<newline>
+ (@{syntax type} '+')? (constdecl +)
+ ;
+ constdecl: @{syntax name} '::' @{syntax type} @{syntax mixfix}?
+ \<close>}
+
+ \begin{description}
+
+ \item @{command (HOL) "record"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t = \<tau> + c\<^sub>1 :: \<sigma>\<^sub>1
+ \<dots> c\<^sub>n :: \<sigma>\<^sub>n"} defines extensible record type @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"},
+ derived from the optional parent record @{text "\<tau>"} by adding new
+ field components @{text "c\<^sub>i :: \<sigma>\<^sub>i"} etc.
+
+ The type variables of @{text "\<tau>"} and @{text "\<sigma>\<^sub>i"} need to be
+ covered by the (distinct) parameters @{text "\<alpha>\<^sub>1, \<dots>,
+ \<alpha>\<^sub>m"}. Type constructor @{text t} has to be new, while @{text
+ \<tau>} needs to specify an instance of an existing record type. At
+ least one new field @{text "c\<^sub>i"} 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 @{text \<tau>} is optional; if omitted
+ @{text t} 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, @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} is made a
+ type abbreviation for the fixed record type @{text "\<lparr>c\<^sub>1 ::
+ \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n\<rparr>"}, likewise is @{text
+ "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m, \<zeta>) t_scheme"} made an abbreviation for
+ @{text "\<lparr>c\<^sub>1 :: \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n, \<dots> ::
+ \<zeta>\<rparr>"}.
+
+ \end{description}
+*}
+
+
+subsection {* Record operations *}
+
+text {*
+ Any record definition of the form presented above produces certain
+ standard operations. Selectors and updates are provided for any
+ field, including the improper one ``@{text more}''. There are also
+ cumulative record constructor functions. To simplify the
+ presentation below, we assume for now that @{text "(\<alpha>\<^sub>1, \<dots>,
+ \<alpha>\<^sub>m) t"} is a root record with fields @{text "c\<^sub>1 ::
+ \<sigma>\<^sub>1, \<dots>, c\<^sub>n :: \<sigma>\<^sub>n"}.
+
+ \medskip \textbf{Selectors} and \textbf{updates} are available for
+ any field (including ``@{text more}''):
+
+ \begin{matharray}{lll}
+ @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
+ @{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>"} \\
+ \end{matharray}
+
+ There is special syntax for application of updates: @{text "r\<lparr>x :=
+ a\<rparr>"} abbreviates term @{text "x_update a r"}. Further notation for
+ repeated updates is also available: @{text "r\<lparr>x := a\<rparr>\<lparr>y := b\<rparr>\<lparr>z :=
+ c\<rparr>"} may be written @{text "r\<lparr>x := a, y := b, z := c\<rparr>"}. 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 @{text "\<lparr>x
+ := a, y := b, z := c\<rparr>"}, 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 \textbf{make} operation provides a cumulative record
+ constructor function:
+
+ \begin{matharray}{lll}
+ @{text "t.make"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
+ \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{ancestor
+ records}. 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 @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>m) t"} has ancestor
+ fields @{text "b\<^sub>1 :: \<rho>\<^sub>1, \<dots>, b\<^sub>k :: \<rho>\<^sub>k"},
+ the above record operations will get the following types:
+
+ \medskip
+ \begin{tabular}{lll}
+ @{text "c\<^sub>i"} & @{text "::"} & @{text "\<lparr>\<^vec>b :: \<^vec>\<rho>, \<^vec>c :: \<^vec>\<sigma>, \<dots> :: \<zeta>\<rparr> \<Rightarrow> \<sigma>\<^sub>i"} \\
+ @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\<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>"} \\
+ @{text "t.make"} & @{text "::"} & @{text "\<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>"} \\
+ \end{tabular}
+ \medskip
+
+ \noindent Some further operations address the extension aspect of a
+ derived record scheme specifically: @{text "t.fields"} produces a
+ record fragment consisting of exactly the new fields introduced here
+ (the result may serve as a more part elsewhere); @{text "t.extend"}
+ takes a fixed record and adds a given more part; @{text
+ "t.truncate"} restricts a record scheme to a fixed record.
+
+ \medskip
+ \begin{tabular}{lll}
+ @{text "t.fields"} & @{text "::"} & @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>n \<Rightarrow> \<lparr>\<^vec>c :: \<^vec>\<sigma>\<rparr>"} \\
+ @{text "t.extend"} & @{text "::"} & @{text "\<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>"} \\
+ @{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>"} \\
+ \end{tabular}
+ \medskip
+
+ \noindent Note that @{text "t.make"} and @{text "t.fields"} coincide
+ for root records.
+*}
+
+
+subsection {* Derived rules and proof tools *}
+
+text {*
+ 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
+ @{text t} is a record type as specified above.
+
+ \begin{enumerate}
+
+ \item Standard conversions for selectors or updates applied to
+ record constructor terms are made part of the default Simplifier
+ context; thus proofs by reduction of basic operations merely require
+ the @{method simp} method without further arguments. These rules
+ are available as @{text "t.simps"}, too.
+
+ \item Selectors applied to updated records are automatically reduced
+ by an internal simplification procedure, which is also part of the
+ standard Simplifier setup.
+
+ \item Inject equations of a form analogous to @{prop "(x, y) = (x',
+ y') \<equiv> x = x' \<and> y = y'"} are declared to the Simplifier and Classical
+ Reasoner as @{attribute iff} rules. These rules are available as
+ @{text "t.iffs"}.
+
+ \item The introduction rule for record equality analogous to @{text
+ "x r = x r' \<Longrightarrow> y r = y r' \<dots> \<Longrightarrow> r = r'"} is declared to the Simplifier,
+ and as the basic rule context as ``@{attribute intro}@{text "?"}''.
+ The rule is called @{text "t.equality"}.
+
+ \item 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 ``@{text "(cases
+ r)"}'' to a certain proof problem.
+
+ \item The derived record operations @{text "t.make"}, @{text
+ "t.fields"}, @{text "t.extend"}, @{text "t.truncate"} are \emph{not}
+ treated automatically, but usually need to be expanded by hand,
+ using the collective fact @{text "t.defs"}.
+
+ \end{enumerate}
+*}
+
+
+subsubsection {* Examples *}
+
+text {* See @{file "~~/src/HOL/ex/Records.thy"}, for example. *}
+
+section {* Typedef axiomatization \label{sec:hol-typedef} *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{command_def (HOL) "typedef"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
+ \end{matharray}
+
+ A Gordon/HOL-style type definition is a certain axiom scheme that
+ identifies a new type with a subset of an existing type. More
+ precisely, the new type is defined by exhibiting an existing type
+ @{text \<tau>}, a set @{text "A :: \<tau> set"}, and a theorem that proves
+ @{prop "\<exists>x. x \<in> A"}. Thus @{text A} is a non-empty subset of @{text
+ \<tau>}, 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 @{text \<tau>} may involve type
+ variables @{text "\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n"} which means that the type definition
+ produces a type constructor @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t"} depending on
+ those type arguments.
+
+ The axiomatization can be considered a ``definition'' in the sense
+ of the particular set-theoretic interpretation of HOL
+ \cite{pitts93}, where the universe of types is required to be
+ downwards-closed wrt.\ arbitrary non-empty subsets. Thus genuinely
+ new types introduced by @{command "typedef"} stay within the range
+ of HOL models by construction. Note that @{command_ref
+ type_synonym} from Isabelle/Pure merely introduces syntactic
+ abbreviations, without any logical significance.
+
+ @{rail \<open>
+ @@{command (HOL) typedef} abs_type '=' rep_set
+ ;
+ abs_type: @{syntax typespec_sorts} @{syntax mixfix}?
+ ;
+ rep_set: @{syntax term} (@'morphisms' @{syntax name} @{syntax name})?
+ \<close>}
+
+ \begin{description}
+
+ \item @{command (HOL) "typedef"}~@{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t = A"}
+ axiomatizes a type definition in the background theory of the
+ current context, depending on a non-emptiness result of the set
+ @{text A} that needs to be proven here. The set @{text A} may
+ contain type variables @{text "\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n"} as specified on the LHS,
+ but no term variables.
+
+ Even though a local theory specification, the newly introduced type
+ constructor cannot depend on parameters or assumptions of the
+ context: this is structurally impossible in HOL. In contrast, the
+ non-emptiness proof may use local assumptions in unusual situations,
+ which could result in different interpretations in target contexts:
+ the meaning of the bijection between the representing set @{text A}
+ and the new type @{text t} may then change in different application
+ contexts.
+
+ For @{command (HOL) "typedef"}~@{text "t = A"} the newly introduced
+ type @{text t} 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 @{text Rep_t} and its inverse @{text
+ Abs_t}: An explicit @{keyword (HOL) "morphisms"} specification
+ allows to provide alternative names.
+
+ The core axiomatization uses the locale predicate @{const
+ type_definition} as 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 @{thm type_definition.Rep} is turned into the specific
+ @{text "Rep_t"}, 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.
+
+ \end{description}
+
+ \begin{warn}
+ If you introduce a new type axiomatically, i.e.\ via @{command_ref
+ typedecl} and @{command_ref axiomatization}, the minimum requirement
+ is that it has a non-empty model, to avoid immediate collapse of the
+ HOL logic. Moreover, one needs to demonstrate that the
+ interpretation of such free-form axiomatizations can coexist with
+ that of the regular @{command_def typedef} scheme, and any extension
+ that other people might have introduced elsewhere.
+ \end{warn}
+*}
+
+subsubsection {* Examples *}
+
+text {* Type definitions permit the introduction of abstract data
+ types in a safe way, namely by providing models based on already
+ existing types. Given some abstract axiomatic description @{text P}
+ of a type, this involves two steps:
+
+ \begin{enumerate}
+
+ \item Find an appropriate type @{text \<tau>} and subset @{text A} which
+ has the desired properties @{text P}, and make a type definition
+ based on this representation.
+
+ \item Prove that @{text P} holds for @{text \<tau>} by lifting @{text P}
+ from the representation.
+
+ \end{enumerate}
+
+ You can later forget about the representation and work solely in
+ terms of the abstract properties @{text P}.
+
+ \medskip The following trivial example pulls a three-element type
+ into existence within the formal logical environment of HOL. *}
+
+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 \<noteq> Two" "One \<noteq> Three" "Two \<noteq> 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)
+
+text {* Note that such trivial constructions are better done with
+ derived specification mechanisms such as @{command datatype}: *}
+
+datatype three' = One' | Two' | Three'
+
+text {* This avoids re-doing basic definitions and proofs from the
+ primitive @{command typedef} above. *}
+
+
+
+section {* Functorial structure of types *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{command_def (HOL) "functor"} & : & @{text "local_theory \<rightarrow> proof(prove)"}
+ \end{matharray}
+
+ @{rail \<open>
+ @@{command (HOL) functor} (@{syntax name} ':')? @{syntax term}
+ \<close>}
+
+ \begin{description}
+
+ \item @{command (HOL) "functor"}~@{text "prefix: m"} 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 theory. By
+ default, they are prefixed with the base name of the type
+ constructor, an explicit prefix can be given alternatively.
+
+ The given term @{text "m"} is considered as \emph{mapper} for the
+ corresponding type constructor and must conform to the following
+ type pattern:
+
+ \begin{matharray}{lll}
+ @{text "m"} & @{text "::"} &
+ @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>k \<Rightarrow> (\<^vec>\<alpha>\<^sub>n) t \<Rightarrow> (\<^vec>\<beta>\<^sub>n) t"} \\
+ \end{matharray}
+
+ \noindent where @{text t} is the type constructor, @{text
+ "\<^vec>\<alpha>\<^sub>n"} and @{text "\<^vec>\<beta>\<^sub>n"} are distinct
+ type variables free in the local theory and @{text "\<sigma>\<^sub>1"},
+ \ldots, @{text "\<sigma>\<^sub>k"} is a subsequence of @{text "\<alpha>\<^sub>1 \<Rightarrow>
+ \<beta>\<^sub>1"}, @{text "\<beta>\<^sub>1 \<Rightarrow> \<alpha>\<^sub>1"}, \ldots,
+ @{text "\<alpha>\<^sub>n \<Rightarrow> \<beta>\<^sub>n"}, @{text "\<beta>\<^sub>n \<Rightarrow>
+ \<alpha>\<^sub>n"}.
+
+ \end{description}
+*}
+
+
+section {* Quotient types *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{command_def (HOL) "quotient_type"} & : & @{text "local_theory \<rightarrow> proof(prove)"}\\
+ @{command_def (HOL) "quotient_definition"} & : & @{text "local_theory \<rightarrow> proof(prove)"}\\
+ @{command_def (HOL) "print_quotmapsQ3"} & : & @{text "context \<rightarrow>"}\\
+ @{command_def (HOL) "print_quotientsQ3"} & : & @{text "context \<rightarrow>"}\\
+ @{command_def (HOL) "print_quotconsts"} & : & @{text "context \<rightarrow>"}\\
+ @{method_def (HOL) "lifting"} & : & @{text method} \\
+ @{method_def (HOL) "lifting_setup"} & : & @{text method} \\
+ @{method_def (HOL) "descending"} & : & @{text method} \\
+ @{method_def (HOL) "descending_setup"} & : & @{text method} \\
+ @{method_def (HOL) "partiality_descending"} & : & @{text method} \\
+ @{method_def (HOL) "partiality_descending_setup"} & : & @{text method} \\
+ @{method_def (HOL) "regularize"} & : & @{text method} \\
+ @{method_def (HOL) "injection"} & : & @{text method} \\
+ @{method_def (HOL) "cleaning"} & : & @{text method} \\
+ @{attribute_def (HOL) "quot_thm"} & : & @{text attribute} \\
+ @{attribute_def (HOL) "quot_lifted"} & : & @{text attribute} \\
+ @{attribute_def (HOL) "quot_respect"} & : & @{text attribute} \\
+ @{attribute_def (HOL) "quot_preserve"} & : & @{text attribute} \\
+ \end{matharray}
+
+ The quotient package defines a new quotient type given a raw type
+ and a partial equivalence relation. The package also historically
+ includes automation for transporting definitions and theorems.
+ But most of this automation was superseded by the Lifting and Transfer
+ packages. The user should consider using these two new packages for
+ lifting definitions and transporting theorems.
+
+ @{rail \<open>
+ @@{command (HOL) quotient_type} (spec)
+ ;
+ spec: @{syntax typespec} @{syntax mixfix}? '=' \<newline>
+ @{syntax type} '/' ('partial' ':')? @{syntax term} \<newline>
+ (@'morphisms' @{syntax name} @{syntax name})? (@'parametric' @{syntax thmref})?
+ \<close>}
+
+ @{rail \<open>
+ @@{command (HOL) quotient_definition} constdecl? @{syntax thmdecl}? \<newline>
+ @{syntax term} 'is' @{syntax term}
+ ;
+ constdecl: @{syntax name} ('::' @{syntax type})? @{syntax mixfix}?
+ \<close>}
+
+ @{rail \<open>
+ @@{method (HOL) lifting} @{syntax thmrefs}?
+ ;
+ @@{method (HOL) lifting_setup} @{syntax thmrefs}?
+ \<close>}
+
+ \begin{description}
+
+ \item @{command (HOL) "quotient_type"} defines a new quotient type @{text \<tau>}. The
+ injection from a quotient type to a raw type is called @{text
+ rep_\<tau>}, its inverse @{text abs_\<tau>} 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 @{text equivp}), unless the
+ user specifies explicitly @{text partial} in which case the
+ obligation is @{text part_equivp}. A quotient defined with @{text
+ partial} 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.
+
+ \end{description}
+
+ The most of the rest of the package was superseded by the Lifting and Transfer
+ packages. The user should consider using these two new packages for
+ lifting definitions and transporting theorems.
+
+ \begin{description}
+
+ \item @{command (HOL) "quotient_definition"} defines a constant on
+ the quotient type.
+
+ \item @{command (HOL) "print_quotmapsQ3"} prints quotient map
+ functions.
+
+ \item @{command (HOL) "print_quotientsQ3"} prints quotients.
+
+ \item @{command (HOL) "print_quotconsts"} prints quotient constants.
+
+ \item @{method (HOL) "lifting"} and @{method (HOL) "lifting_setup"}
+ methods match the current goal with the given raw theorem to be
+ lifted producing three new subgoals: regularization, injection and
+ cleaning subgoals. @{method (HOL) "lifting"} tries to apply the
+ heuristics for automatically solving these three subgoals and
+ leaves only the subgoals unsolved by the heuristics to the user as
+ opposed to @{method (HOL) "lifting_setup"} which leaves the three
+ subgoals unsolved.
+
+ \item @{method (HOL) "descending"} and @{method (HOL)
+ "descending_setup"} try to guess a raw statement that would lift
+ to the current subgoal. Such statement is assumed as a new subgoal
+ and @{method (HOL) "descending"} continues in the same way as
+ @{method (HOL) "lifting"} does. @{method (HOL) "descending"} tries
+ to solve the arising regularization, injection and cleaning
+ subgoals with the analogous method @{method (HOL)
+ "descending_setup"} which leaves the four unsolved subgoals.
+
+ \item @{method (HOL) "partiality_descending"} finds the regularized
+ theorem that would lift to the current subgoal, lifts it and
+ leaves as a subgoal. This method can be used with partial
+ equivalence quotients where the non regularized statements would
+ not be true. @{method (HOL) "partiality_descending_setup"} leaves
+ the injection and cleaning subgoals unchanged.
+
+ \item @{method (HOL) "regularize"} applies the regularization
+ heuristics to the current subgoal.
+
+ \item @{method (HOL) "injection"} applies the injection heuristics
+ to the current goal using the stored quotient respectfulness
+ theorems.
+
+ \item @{method (HOL) "cleaning"} applies the injection cleaning
+ heuristics to the current subgoal using the stored quotient
+ preservation theorems.
+
+ \item @{attribute (HOL) quot_lifted} attribute tries to
+ automatically transport the theorem to the quotient type.
+ The attribute uses all the defined quotients types and quotient
+ constants often producing undesired results or theorems that
+ cannot be lifted.
+
+ \item @{attribute (HOL) quot_respect} and @{attribute (HOL)
+ quot_preserve} attributes declare a theorem as a respectfulness
+ and preservation theorem respectively. These are stored in the
+ local theory store and used by the @{method (HOL) "injection"}
+ and @{method (HOL) "cleaning"} methods respectively.
+
+ \item @{attribute (HOL) quot_thm} declares that a certain theorem
+ is a quotient extension theorem. Quotient extension theorems
+ allow for quotienting inside container types. Given a polymorphic
+ type that serves as a container, a map function defined for this
+ container using @{command (HOL) "functor"} and a relation
+ map defined for for the container type, the quotient extension
+ theorem should be @{term "Quotient3 R Abs Rep \<Longrightarrow> Quotient3
+ (rel_map R) (map Abs) (map Rep)"}. Quotient extension theorems
+ are stored in a database and are used all the steps of lifting
+ theorems.
+
+ \end{description}
+*}
+
+
+section {* Definition by specification \label{sec:hol-specification} *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{command_def (HOL) "specification"} & : & @{text "theory \<rightarrow> proof(prove)"} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{command (HOL) specification} '(' (decl +) ')' \<newline>
+ (@{syntax thmdecl}? @{syntax prop} +)
+ ;
+ decl: (@{syntax name} ':')? @{syntax term} ('(' @'overloaded' ')')?
+ \<close>}
+
+ \begin{description}
+
+ \item @{command (HOL) "specification"}~@{text "decls \<phi>"} sets up a
+ goal stating the existence of terms with the properties specified to
+ hold for the constants given in @{text decls}. 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.
+
+ @{text decl} declares a constant to be defined by the
+ specification given. The definition for the constant @{text c} is
+ bound to the name @{text c_def} unless a theorem name is given in
+ the declaration. Overloaded constants should be declared as such.
+
+ \end{description}
+*}
+
+
+section {* Adhoc overloading of constants *}
+
+text {*
+ \begin{tabular}{rcll}
+ @{command_def "adhoc_overloading"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+ @{command_def "no_adhoc_overloading"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+ @{attribute_def "show_variants"} & : & @{text "attribute"} & default @{text false} \\
+ \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 output) and the addition
+ of some variants (used internally). For examples see
+ @{file "~~/src/HOL/ex/Adhoc_Overloading_Examples.thy"} and
+ @{file "~~/src/HOL/Library/Monad_Syntax.thy"}.
+
+ @{rail \<open>
+ (@@{command adhoc_overloading} | @@{command no_adhoc_overloading})
+ (@{syntax nameref} (@{syntax term} + ) + @'and')
+ \<close>}
+
+ \begin{description}
+
+ \item @{command "adhoc_overloading"}~@{text "c v\<^sub>1 ... v\<^sub>n"}
+ associates variants with an existing constant.
+
+ \item @{command "no_adhoc_overloading"} is similar to
+ @{command "adhoc_overloading"}, but removes the specified variants
+ from the present context.
+
+ \item @{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.
+
+ \end{description}
+*}
+
+chapter {* Proof tools *}
+
+section {* Adhoc tuples *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{attribute_def (HOL) split_format}@{text "\<^sup>*"} & : & @{text attribute} \\
+ \end{matharray}
+
+ @{rail \<open>
+ @@{attribute (HOL) split_format} ('(' 'complete' ')')?
+ \<close>}
+
+ \begin{description}
+
+ \item @{attribute (HOL) split_format}\ @{text "(complete)"} causes
+ arguments in function applications to be represented canonically
+ according to their tuple type structure.
+
+ Note that this operation tends to invent funny names for new local
+ parameters introduced.
+
+ \end{description}
+*}
+
+
+section {* Transfer package *}
+
+text {*
+ \begin{matharray}{rcl}
+ @{method_def (HOL) "transfer"} & : & @{text method} \\
+ @{method_def (HOL) "transfer'"} & : & @{text method} \\
+ @{method_def (HOL) "transfer_prover"} & : & @{text method} \\
+ @{attribute_def (HOL) "Transfer.transferred"} & : & @{text attribute} \\
+ @{attribute_def (HOL) "untransferred"} & : & @{text attribute} \\
+ @{attribute_def (HOL) "transfer_rule"} & : & @{text attribute} \\
+ @{attribute_def (HOL) "transfer_domain_rule"} & : & @{text attribute} \\
+ @{attribute_def (HOL) "relator_eq"} & : & @{text attribute} \\
+ @{attribute_def (HOL) "relator_domain"} & : & @{text attribute} \\
+ \end{matharray}
+
+ \begin{description}
+
+ \item @{method (HOL) "transfer"} method replaces the current subgoal
+ with a logically equivalent one that uses different types and
+ constants. The replacement of types and constants is guided by the
+ database of transfer rules. Goals are generalized over all free
+ variables by default; this is necessary for variables whose types
+ change, but can be overridden for specific variables with e.g.
+ @{text "transfer fixing: x y z"}.
+
+ \item @{method (HOL) "transfer'"} is a variant of @{method (HOL)
+ transfer} that allows replacing a subgoal with one that is
+ logically stronger (rather than equivalent). For example, a
+ subgoal involving equality on a quotient type could be replaced
+ with a subgoal involving equality (instead of the corresponding
+ equivalence relation) on the underlying raw type.
+
+ \item @{method (HOL) "transfer_prover"} method assists with proving
+ a transfer rule for a new constant, provided the constant is
+ defined in terms of other constants that already have transfer
+ rules. It should be applied after unfolding the constant
+ definitions.
+
+ \item @{attribute (HOL) "untransferred"} proves the same equivalent theorem
+ as @{method (HOL) "transfer"} internally does.
+
+ \item @{attribute (HOL) Transfer.transferred} works in the opposite