--- a/NEWS Sun Oct 24 03:43:12 2010 -0700
+++ b/NEWS Sun Oct 24 15:11:24 2010 -0700
@@ -94,6 +94,8 @@
set_ext -> set_eqI
INCOMPATIBILITY.
+* Renamed lemma list: nat_number -> eval_nat_numeral
+
* Renamed class eq and constant eq (for code generation) to class equal
and constant equal, plus renaming of related facts and various tuning.
INCOMPATIBILITY.
@@ -276,18 +278,28 @@
meson_disj_FalseD2 ~> Meson.disj_FalseD2
INCOMPATIBILITY.
-* Sledgehammer: Renamed lemmas:
- COMBI_def ~> Meson.COMBI_def
- COMBK_def ~> Meson.COMBK_def
- COMBB_def ~> Meson.COMBB_def
- COMBC_def ~> Meson.COMBC_def
- COMBS_def ~> Meson.COMBS_def
- abs_I ~> Meson.abs_I
- abs_K ~> Meson.abs_K
- abs_B ~> Meson.abs_B
- abs_C ~> Meson.abs_C
- abs_S ~> Meson.abs_S
-INCOMPATIBILITY.
+* Sledgehammer:
+ - Renamed lemmas:
+ COMBI_def ~> Meson.COMBI_def
+ COMBK_def ~> Meson.COMBK_def
+ COMBB_def ~> Meson.COMBB_def
+ COMBC_def ~> Meson.COMBC_def
+ COMBS_def ~> Meson.COMBS_def
+ abs_I ~> Meson.abs_I
+ abs_K ~> Meson.abs_K
+ abs_B ~> Meson.abs_B
+ abs_C ~> Meson.abs_C
+ abs_S ~> Meson.abs_S
+ INCOMPATIBILITY.
+ - Renamed commands:
+ sledgehammer atp_info ~> sledgehammer running_provers
+ sledgehammer atp_kill ~> sledgehammer kill_provers
+ sledgehammer available_atps ~> sledgehammer available_provers
+ INCOMPATIBILITY.
+ - Renamed options:
+ sledgehammer [atps = ...] ~> sledgehammer [provers = ...]
+ sledgehammer [atp = ...] ~> sledgehammer [prover = ...]
+ INCOMPATIBILITY.
*** FOL ***
--- a/doc-src/IsarImplementation/IsaMakefile Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/IsaMakefile Sun Oct 24 15:11:24 2010 -0700
@@ -19,13 +19,13 @@
## sessions
-Thy: $(LOG)/Pure-Thy.gz
+Thy: $(LOG)/HOL-Thy.gz
-$(LOG)/Pure-Thy.gz: Thy/ROOT.ML Thy/Base.thy Thy/Integration.thy \
+$(LOG)/HOL-Thy.gz: Thy/ROOT.ML Thy/Base.thy Thy/Integration.thy \
Thy/Isar.thy Thy/Local_Theory.thy Thy/Logic.thy Thy/Prelim.thy \
Thy/Proof.thy Thy/Syntax.thy Thy/Tactic.thy Thy/ML.thy \
../antiquote_setup.ML
- @$(USEDIR) Pure Thy
+ @$(USEDIR) HOL Thy
@rm -f Thy/document/isabelle.sty Thy/document/isabellesym.sty \
Thy/document/pdfsetup.sty Thy/document/session.tex
@@ -33,4 +33,4 @@
## clean
clean:
- @rm -f $(LOG)/Pure-Thy.gz
+ @rm -f $(LOG)/HOL-Thy.gz
--- a/doc-src/IsarImplementation/Makefile Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Makefile Sun Oct 24 15:11:24 2010 -0700
@@ -11,8 +11,8 @@
NAME = implementation
FILES = ../extra.sty ../iman.sty ../isabelle.sty ../isabellesym.sty \
- ../isar.sty ../manual.bib ../pdfsetup.sty ../proof.sty \
- Thy/document/Integration.tex Thy/document/Isar.tex \
+ ../isar.sty ../manual.bib ../pdfsetup.sty ../proof.sty ../rail.sty \
+ ../railsetup.sty Thy/document/Integration.tex Thy/document/Isar.tex \
Thy/document/Local_Theory.tex Thy/document/Logic.tex \
Thy/document/Prelim.tex Thy/document/Proof.tex \
Thy/document/Syntax.tex Thy/document/Tactic.tex implementation.tex \
@@ -22,6 +22,7 @@
$(NAME).dvi: $(FILES) isabelle_isar.eps
$(LATEX) $(NAME)
+ $(RAIL) $(NAME)
$(BIBTEX) $(NAME)
$(LATEX) $(NAME)
$(LATEX) $(NAME)
@@ -33,6 +34,7 @@
$(NAME).pdf: $(FILES) isabelle_isar.pdf
$(PDFLATEX) $(NAME)
+ $(RAIL) $(NAME)
$(BIBTEX) $(NAME)
$(PDFLATEX) $(NAME)
$(PDFLATEX) $(NAME)
--- a/doc-src/IsarImplementation/Thy/Base.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Base.thy Sun Oct 24 15:11:24 2010 -0700
@@ -1,6 +1,12 @@
theory Base
-imports Pure
+imports Main
uses "../../antiquote_setup.ML"
begin
+(* FIXME move to src/Pure/ML/ml_antiquote.ML *)
+ML {*
+ ML_Antiquote.inline "assert"
+ (Scan.succeed "(fn b => if b then () else raise General.Fail \"Assertion failed\")")
+*}
+
end
--- a/doc-src/IsarImplementation/Thy/Integration.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Integration.thy Sun Oct 24 15:11:24 2010 -0700
@@ -8,13 +8,12 @@
text {* The Isar toplevel may be considered the centeral 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. We
- shall even incorporate the existing {\ML} toplevel of the compiler
- and run-time system (cf.\ \secref{sec:ML-toplevel}).
+ 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
+ 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
@@ -66,11 +65,11 @@
\begin{description}
- \item @{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 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,
@@ -94,13 +93,29 @@
information for each Isar command being executed.
\item @{ML Toplevel.profiling}~@{verbatim ":="}~@{text "n"} controls
- low-level profiling of the underlying {\ML} runtime system. For
+ 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} *}
@@ -182,123 +197,6 @@
*}
-subsection {* Toplevel control *}
-
-text {*
- There are a few special control commands that modify the behavior
- the toplevel itself, and only make sense in interactive mode. Under
- normal circumstances, the user encounters these only implicitly as
- part of the protocol between the Isabelle/Isar system and a
- user-interface such as Proof~General.
-
- \begin{description}
-
- \item \isacommand{undo} follows the three-level hierarchy of empty
- toplevel vs.\ theory vs.\ proof: undo within a proof reverts to the
- previous proof context, undo after a proof reverts to the theory
- before the initial goal statement, undo of a theory command reverts
- to the previous theory value, undo of a theory header discontinues
- the current theory development and removes it from the theory
- database (\secref{sec:theory-database}).
-
- \item \isacommand{kill} aborts the current level of development:
- kill in a proof context reverts to the theory before the initial
- goal statement, kill in a theory context aborts the current theory
- development, removing it from the database.
-
- \item \isacommand{exit} drops out of the Isar toplevel into the
- underlying {\ML} toplevel (\secref{sec:ML-toplevel}). The Isar
- toplevel state is preserved and may be continued later.
-
- \item \isacommand{quit} terminates the Isabelle/Isar process without
- saving.
-
- \end{description}
-*}
-
-
-section {* ML toplevel \label{sec:ML-toplevel} *}
-
-text {*
- The {\ML} toplevel provides a read-compile-eval-print loop for {\ML}
- values, types, structures, and functors. {\ML} declarations operate
- on the global system state, which consists of the compiler
- environment plus the values of {\ML} reference variables. There is
- no clean way to undo {\ML} declarations, except for reverting to a
- previously saved state of the whole Isabelle process. {\ML} input
- is either read interactively from a TTY, or from a string (usually
- within a theory text), or from a source file (usually loaded from a
- theory).
-
- Whenever the {\ML} toplevel is active, the current Isabelle theory
- context is passed as an internal reference variable. Thus {\ML}
- code may access the theory context during compilation, it may even
- change the value of a theory being under construction --- while
- observing the usual linearity restrictions
- (cf.~\secref{sec:context-theory}).
-*}
-
-text %mlref {*
- \begin{mldecls}
- @{index_ML ML_Context.the_generic_context: "unit -> Context.generic"} \\
- @{index_ML "Context.>> ": "(Context.generic -> Context.generic) -> 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 runtime. Moreover, persistent {\ML} toplevel bindings
- to an unfinished theory should be avoided: code should either
- project out the desired information immediately, or produce an
- explicit @{ML_type theory_ref} (cf.\ \secref{sec:context-theory}).
-
- \item @{ML "Context.>>"}~@{text f} applies context transformation
- @{text f} to the implicit context of the {\ML} toplevel.
-
- \end{description}
-
- It is very important to note that the above functions are really
- restricted to the compile time, even though the {\ML} compiler is
- invoked at runtime! The majority of {\ML} code uses explicit
- functional arguments of a theory or proof context instead. Thus it
- may be invoked for an arbitrary context later on, without having to
- worry about any operational details.
-
- \bigskip
-
- \begin{mldecls}
- @{index_ML Isar.main: "unit -> unit"} \\
- @{index_ML Isar.loop: "unit -> unit"} \\
- @{index_ML Isar.state: "unit -> Toplevel.state"} \\
- @{index_ML Isar.exn: "unit -> (exn * string) option"} \\
- @{index_ML Isar.goal: "unit ->
- {context: Proof.context, facts: thm list, goal: thm}"} \\
- \end{mldecls}
-
- \begin{description}
-
- \item @{ML "Isar.main ()"} invokes the Isar toplevel from {\ML},
- initializing an empty toplevel state.
-
- \item @{ML "Isar.loop ()"} continues the Isar toplevel with the
- current state, after having dropped out of the Isar toplevel loop.
-
- \item @{ML "Isar.state ()"} and @{ML "Isar.exn ()"} get current
- toplevel state and error condition, respectively. This only works
- after having dropped out of the Isar toplevel loop.
-
- \item @{ML "Isar.goal ()"} produces the full Isar goal state,
- consisting of proof context, facts that have been indicated for
- immediate use, and the tactical goal according to
- \secref{sec:tactical-goals}.
-
- \end{description}
-*}
-
-
section {* Theory database \label{sec:theory-database} *}
text {*
@@ -315,10 +213,10 @@
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
+ 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}
+ 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
--- a/doc-src/IsarImplementation/Thy/Isar.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Isar.thy Sun Oct 24 15:11:24 2010 -0700
@@ -5,16 +5,35 @@
chapter {* Isar language elements *}
text {*
- The primary Isar language consists of three main categories of
- language elements:
+ The Isar proof language (see also \cite[\S2]{isabelle-isar-ref})
+ consists of three main categories of language elements:
\begin{enumerate}
- \item Proof commands
+ \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 methods
+ \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}.
- \item Attributes
+ 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}
*}
@@ -22,16 +41,523 @@
section {* Proof commands *}
-text FIXME
+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: 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 stores the final
+ result in a suitable context data slot. 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 (see also \chref{ch:local-theory}).
+ This usually 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 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. *}
+
+example_proof
+ 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 FIXME
+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 is embedded
+ into Isar as arguments to certain commands, e.g.\ @{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. First some 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 here. 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-off any remaining
+ subgoals, but it does not see the facts of the initial step.
+
+ \medskip The second 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 together 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, Isar proof methods 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, instead of separate subgoals.
+
+ \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).
+ To make the special non-standard status clear, the naming convention
+ @{text "foo_tac"} needs to be observed.
+
+ 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 HEADGOAL: "(int -> tactic) -> tactic"} \\
+ @{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. Goal facts are
+ already inserted into the first subgoal before @{text "tactic"} is
+ applied to the same.
+
+ \item @{ML HEADGOAL}~@{text "tactic"} applies @{text "tactic"} to
+ the first subgoal. This is convenient to reproduce part the @{ML
+ SIMPLE_METHOD'} wrapping within regular @{ML METHOD}, for example.
+
+ \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 @{verbatim facts}). This allows immediate
+ experimentation without parsing of concrete syntax. *}
+
+example_proof
+ 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
+qed
+
+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
+ (HOL_basic_ss 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}). The resulting @{verbatim 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:
+*}
+
+example_proof
+ fix a b c
+ assume a: "a = b"
+ assume b: "b = c"
+ have "a = c" by (my_simp a b)
+qed
+
+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
+ (HOL_basic_ss addsimps thms)))))
+*} "rewrite all subgoals by given rules"
+
+example_proof
+ fix a b c
+ assume a: "a = b"
+ assume b: "b = c"
+ have "a = c" and "c = b" by (my_simp_all a b)
+
+qed
+
+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 = "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
+ (HOL_basic_ss 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:
+*}
+
+example_proof
+ fix a b c
+ assume [my_simp]: "a \<equiv> b"
+ assume [my_simp]: "b \<equiv> c"
+ have "a \<equiv> c" by my_simp'
+qed
+
+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 auto} little more can be
+ done here.
+
+ 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 requires it be maintained statically within the context
+ data, not dynamically on each tool invocation. *}
-section {* Attributes *}
+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 fully general 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 FIXME
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML_type attribute: "Context.generic * thm -> Context.generic * thm"} \\
+ @{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 %mlex {* See also @{command attribute_setup} in
+ \cite{isabelle-isar-ref} which includes some abstract examples. *}
end
--- a/doc-src/IsarImplementation/Thy/Local_Theory.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Local_Theory.thy Sun Oct 24 15:11:24 2010 -0700
@@ -31,8 +31,8 @@
Many definitional packages for local theories are available in
Isabelle. Although a few old packages only work for global
- theories, the local theory interface is already the standard way of
- implementing definitional packages in Isabelle.
+ theories, the standard way of implementing definitional packages in
+ Isabelle is via the local theory interface.
*}
@@ -86,13 +86,12 @@
\framebox{\quad auxiliary context \quad\framebox{\quad target context \quad\framebox{\quad background theory\quad}}}
\end{center}
- \noindent 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).
-*}
+ 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}
@@ -106,8 +105,8 @@
\begin{description}
- \item @{ML_type local_theory} represents local theories. Although
- this is merely an alias for @{ML_type Proof.context}, it is
+ \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
@@ -155,8 +154,11 @@
*}
-section {* Morphisms and declarations *}
+section {* Morphisms and declarations \label{sec:morphisms} *}
-text FIXME
+text {* FIXME
+
+ \medskip See also \cite{Chaieb-Wenzel:2007}.
+*}
end
--- a/doc-src/IsarImplementation/Thy/Logic.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Logic.thy Sun Oct 24 15:11:24 2010 -0700
@@ -121,8 +121,8 @@
@{index_ML_type sort: "class list"} \\
@{index_ML_type arity: "string * sort list * sort"} \\
@{index_ML_type typ} \\
- @{index_ML map_atyps: "(typ -> typ) -> typ -> typ"} \\
- @{index_ML fold_atyps: "(typ -> 'a -> 'a) -> typ -> 'a -> 'a"} \\
+ @{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"} \\
@@ -136,26 +136,27 @@
\begin{description}
- \item @{ML_type class} represents type classes.
+ \item Type @{ML_type class} represents type classes.
- \item @{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 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 @{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 arity} represents type arities. A triple
+ @{text "(\<kappa>, \<^vec>s, s) : arity"} represents @{text "\<kappa> ::
+ (\<^vec>s)s"} as described above.
- \item @{ML_type typ} represents types; this is a datatype with
+ \item Type @{ML_type typ} represents types; this is a datatype with
constructors @{ML TFree}, @{ML TVar}, @{ML Type}.
- \item @{ML 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.map_atyps}~@{text "f \<tau>"} applies the mapping @{text
+ "f"} to all atomic types (@{ML TFree}, @{ML TVar}) occurring in
+ @{text "\<tau>"}.
- \item @{ML 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 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\<^isub>1, s\<^isub>2)"}
tests the subsort relation @{text "s\<^isub>1 \<subseteq> s\<^isub>2"}.
@@ -185,6 +186,51 @@
\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}
+
+ \begin{rail}
+ 'class' nameref
+ ;
+ 'sort' sort
+ ;
+ ('type\_name' | 'type\_abbrev' | 'nonterminal') nameref
+ ;
+ 'typ' type
+ ;
+ \end{rail}
+
+ \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} *}
@@ -240,7 +286,7 @@
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.
+ \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\<^isub>\<tau> | ?x\<^isub>\<tau> | c\<^isub>\<tau> | \<lambda>\<^isub>\<tau>. t | t\<^isub>1 t\<^isub>2"}. Parsing and printing takes care of
@@ -309,10 +355,10 @@
\begin{mldecls}
@{index_ML_type term} \\
@{index_ML "op aconv": "term * term -> bool"} \\
- @{index_ML map_types: "(typ -> typ) -> term -> term"} \\
- @{index_ML fold_types: "(typ -> 'a -> 'a) -> term -> 'a -> 'a"} \\
- @{index_ML map_aterms: "(term -> term) -> term -> term"} \\
- @{index_ML fold_aterms: "(term -> 'a -> 'a) -> term -> 'a -> 'a"} \\
+ @{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"} \\
@@ -328,8 +374,8 @@
\begin{description}
- \item @{ML_type term} represents de-Bruijn terms, with comments in
- abstractions, and explicitly named free variables and constants;
+ \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 @{ML Bound}, @{ML Free}, @{ML
Var}, @{ML Const}, @{ML Abs}, @{ML "op $"}.
@@ -338,20 +384,20 @@
on type @{ML_type term}; raw datatype equality should only be used
for operations related to parsing or printing!
- \item @{ML map_types}~@{text "f t"} applies the mapping @{text
+ \item @{ML Term.map_types}~@{text "f t"} applies the mapping @{text
"f"} to all types occurring in @{text "t"}.
- \item @{ML fold_types}~@{text "f t"} iterates the operation @{text
- "f"} over all occurrences of types in @{text "t"}; the term
+ \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 map_aterms}~@{text "f t"} applies the mapping @{text "f"}
- to all atomic terms (@{ML Bound}, @{ML Free}, @{ML Var}, @{ML
+ \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 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
+ \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
@@ -381,6 +427,49 @@
\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}
+
+ \begin{rail}
+ ('const\_name' | 'const\_abbrev') nameref
+ ;
+ 'const' ('(' (type + ',') ')')?
+ ;
+ 'term' term
+ ;
+ 'prop' prop
+ ;
+ \end{rail}
+
+ \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} *}
@@ -531,7 +620,7 @@
"d(\<alpha>\<^isub>i)"} for some @{text "\<alpha>\<^isub>i"} projected from @{text
"\<^vec>\<alpha>"}. Thus overloaded definitions essentially work by
primitive recursion over the syntactic structure of a single type
- argument.
+ argument. See also \cite[\S4.3]{Haftmann-Wenzel:2006:classes}.
*}
text %mlref {*
@@ -552,18 +641,20 @@
@{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: "binding * term -> theory -> (string * thm) * theory"} \\
- @{index_ML Thm.add_oracle: "binding * ('a -> cterm) -> theory
- -> (string * ('a -> thm)) * theory"} \\
- @{index_ML Thm.add_def: "bool -> bool -> binding * term -> theory -> (string * thm) * theory"} \\
+ @{index_ML Thm.add_oracle: "binding * ('a -> cterm) -> theory ->
+ (string * ('a -> thm)) * theory"} \\
+ @{index_ML Thm.add_def: "bool -> bool -> binding * term -> theory ->
+ (string * thm) * theory"} \\
\end{mldecls}
\begin{mldecls}
- @{index_ML Theory.add_deps: "string -> string * typ -> (string * typ) list -> theory -> theory"} \\
+ @{index_ML Theory.add_deps: "string -> string * typ -> (string * typ) list ->
+ theory -> theory"} \\
\end{mldecls}
\begin{description}
- \item @{ML_type ctyp} and @{ML_type cterm} represent certified types
- and terms, respectively. These are abstract datatypes that
+ \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 theory.
@@ -577,8 +668,8 @@
reasoning loops. There are separate operations to decompose
certified entities (including actual theorems).
- \item @{ML_type thm} represents proven propositions. This is an
- abstract datatype that guarantees that its values have been
+ \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_struct Thm} module.
Every @{ML_type thm} value contains a sliding back-reference to the
enclosing theory, cf.\ \secref{sec:context-theory}.
@@ -631,6 +722,68 @@
*}
+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}
+
+ \begin{rail}
+ 'ctyp' typ
+ ;
+ 'cterm' term
+ ;
+ 'cprop' prop
+ ;
+ 'thm' thmref
+ ;
+ 'thms' thmrefs
+ ;
+ 'lemma' ('(open)')? ((prop +) + 'and') \\ 'by' method method?
+ \end{rail}
+
+ \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 definitions *}
text {*
@@ -776,9 +929,9 @@
\end{tabular}
\medskip
- \noindent 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
+ 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
--- a/doc-src/IsarImplementation/Thy/ML.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/ML.thy Sun Oct 24 15:11:24 2010 -0700
@@ -2,633 +2,1728 @@
imports Base
begin
-chapter {* Advanced ML programming *}
+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 are specific library
+ modules and infrastructure to address the needs for such difficult
+ tasks. 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.
+ One should always strife at least for local consistency of modules
+ and sub-systems, without deviating from some general principles how
+ to write Isabelle/ML.
+
+ In a sense, a common 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 source.
+*}
+
+
+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. E.g.\ see @{"file" "~~/src/Pure/thm.ML"}.
+
+ 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 some a layout via ML comments
+ as follows.
+
+\begin{verbatim}
+(*** section ***)
+
+(** subsection **)
+
+(* subsubsection *)
+
+(*short paragraph*)
+
+(*
+ long paragraph
+ 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 indicate 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 just for
+ certain ML categories as follows:
+
+ \medskip
+ \begin{tabular}{lll}
+ variant & example & ML categories \\\hline
+ lower-case & @{verbatim foo_bar} & values, types, record fields \\
+ capitalized & @{verbatim Foo_Bar} & datatype constructors, \\
+ & & structures, functors \\
+ upper-case & @{verbatim FOO_BAR} & special values (combinators), \\
+ & & exception constructors, signatures \\
+ \end{tabular}
+ \medskip
+
+ For historical reasons, many capitalized names omit underscores,
+ e.g.\ old-style @{verbatim FooBar} instead of @{verbatim Foo_Bar}.
+ Genuine mixed-case names are \emph{not} used --- 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 and keywords. Later it became a 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 are suffixed by extra markers like this:
+ @{verbatim foo_barT}.
+
+ Name variants are produced by adding 1--3 primes, e.g.\ @{verbatim
+ foo'}, @{verbatim foo''}, or @{verbatim foo'''}, but not @{verbatim
+ foo''''} or more. Decimal digits scale better to larger numbers,
+ e.g.\ @{verbatim foo0}, @{verbatim foo1}, @{verbatim 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 always seen in the sources and documentation.
+ For the same reasons, aliases of well-known library functions should
+ be avoided.
-section {* Style *}
+ Local names of function abstraction or case/lets bindings are
+ typically shorter, sometimes using only rudiments of ``words'',
+ while still avoiding cryptic shorthands. An auxiliary function
+ called @{verbatim helper}, @{verbatim aux}, or @{verbatim 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 important specific
+ name forms that occur frequently in the sources.
+
+ \begin{itemize}
+
+ \item A function that maps @{verbatim foo} to @{verbatim bar} is
+ called @{verbatim foo_to_bar} or @{verbatim bar_of_foo} (never
+ @{verbatim foo2bar}, @{verbatim bar_from_foo}, @{verbatim
+ bar_for_foo}, or @{verbatim bar4foo}).
+
+ \item The name component @{verbatim legacy} means that the operation
+ is about to be discontinued soon.
+
+ \item The name component @{verbatim old} means that this is historic
+ material that might disappear at some later stage.
+
+ \item The name component @{verbatim 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 to allow to visualize the their data flow
+ via plain regular expressions in the editor. In particular:
+
+ \begin{itemize}
+
+ \item theories are called @{verbatim thy}, rarely @{verbatim theory}
+ (never @{verbatim thry})
+
+ \item proof contexts are called @{verbatim ctxt}, rarely @{verbatim
+ context} (never @{verbatim ctx})
+
+ \item generic contexts are called @{verbatim context}, rarely
+ @{verbatim ctxt}
+
+ \item local theories are called @{verbatim lthy}, unless there is an
+ implicit conversion to a general proof context (semantic super-type)
+
+ \end{itemize}
+
+ Variations with primed or decimal numbers are always possible, as
+ well as ``sematic prefixes'' like @{verbatim foo_thy} or @{verbatim
+ bar_ctxt}, but the base conventions above need to be preserved.
+
+ \item The main logical entities (\secref{ch:logic}) also have
+ established naming convention:
+
+ \begin{itemize}
+
+ \item sorts are called @{verbatim S}
+
+ \item types are called @{verbatim T}, @{verbatim U}, or @{verbatim
+ ty} (never @{verbatim t})
+
+ \item terms are called @{verbatim t}, @{verbatim u}, or @{verbatim
+ tm} (never @{verbatim trm})
+
+ \item certified types are called @{verbatim cT}, rarely @{verbatim
+ T}, with variants as for types
+
+ \item certified terms are called @{verbatim ct}, rarely @{verbatim
+ t}, with variants as for terms
+
+ \item theorems are called @{verbatim th}, or @{verbatim 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
+ @{verbatim lhs} (not @{verbatim lhs_tm}). Or a term that is treated
+ specifically as a variable can be called @{verbatim v} or @{verbatim
+ x}.
+
+ \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. As a rule of thumb,
+ sources need to be printable on plain paper.} 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:
-text {*
- Like any style guide, also this one should not be interpreted dogmatically, but
- with care and discernment. It consists of a collection of
- recommendations which have been turned out useful over long years of
- Isabelle system development and are supposed to support writing of readable
- and managable code. Special cases encourage derivations,
- as far as you know what you are doing.
- \footnote{This style guide is loosely based on
- \url{http://caml.inria.fr/resources/doc/guides/guidelines.en.html}}
+ \begin{verbatim}
+ val x =
+ a +
+ b +
+ c;
+
+ val tuple =
+ (a,
+ b,
+ c);
+ \end{verbatim}
+
+ Some special infixes (e.g.\ @{verbatim "|>"}) 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: @{verbatim "f a b"} for a
+ curried function, or @{verbatim "g (a, b)"} for a tupled function.
+ Note that the space between @{verbatim g} and the pair @{verbatim
+ "(a, b)"} follows the important principle of
+ \emph{compositionality}: the layout of @{verbatim "g p"} does not
+ change when @{verbatim "p"} is refined to a concrete pair.
+
+ \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.
+
+ 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 stars etc. should be
+ avoided.
+
+ \paragraph{Complex expressions} consisting of multi-clausal function
+ definitions, @{verbatim case}, @{verbatim handle}, @{verbatim let},
+ or combinations require special attention. The syntax of Standard
+ ML is a bit too ambitious in admitting too much variance that can
+ distort the meaning of the text.
+
+ Clauses of @{verbatim fun}, @{verbatim fn}, @{verbatim case},
+ @{verbatim handle} get extra indentation to indicate the nesting
+ clearly. For 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 @{verbatim case} or @{verbatim let}
+ require care to maintain compositionality, to prevent loss of
+ logical indentation where it is particularly imporant to see the
+ structure of the text later on. 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 @{verbatim case} expressions are optional,
+ but help to analyse the nesting easily based on simple string
+ 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 @{verbatim "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 @{verbatim 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. For example:
+
+ \begin{verbatim}
+ (* RIGHT *)
+
+ fun foo x y = fold (fn z =>
+ expr)
+ \end{verbatim}
+
+ Here the visual appearance is that of three arguments @{verbatim x},
+ @{verbatim y}, @{verbatim z}.
+
+ \end{enumerate}
+
+ Such weakly structured layout should be use with great care. Here
+ is a counter-example involving @{verbatim let} expressions:
+
+ \begin{verbatim}
+ (* WRONG *)
+
+ fun foo x = let
+ val y = ...
+ in ... end
+
+ 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 environment
+ for arbitrary ML values (see also \secref{sec:context}). This
+ formal context holds logical entities as well as ML compiler
+ bindings, among many other things. Raw Standard ML is never
+ encountered again after the initial bootstrap of Isabelle/Pure.
+
+ Object-logics such as Isabelle/HOL are built within the
+ Isabelle/ML/Isar environment of Pure by introducing suitable
+ theories with associated ML text, 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 regular user-space repertoire of Isabelle.
+*}
+
+
+subsection {* Isar ML commands *}
+
+text {* The primary Isar source language provides various facilities
+ to open a ``window'' to the underlying ML compiler. Especially see
+ @{command_ref "use"} and @{command_ref "ML"}, which work exactly 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}.
+ An example of even more fine-grained embedding of ML into Isar is
+ 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 really 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 or similar command loops of Computer Algebra systems, 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.
+ Instead of @{command_ref "ML"} for theory mode, we use @{command_ref
+ "ML_prf"} for proof mode. As illustrated below, its effect on the
+ ML environment is local to the whole proof body, while ignoring its
+ internal block structure. *}
+
+example_proof
+ ML_prf %"ML" {* val a = 1 *}
+ { -- {* Isar block structure ignored by ML environment *}
+ ML_prf %"ML" {* val b = a + 1 *}
+ } -- {* Isar block structure ignored by ML environment *}
+ ML_prf %"ML" {* val c = b + 1 *}
+qed
+
+text {* By side-stepping the normal scoping rules for Isar proof
+ blocks, embedded ML code can refer to the different contexts
+ explicitly, 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 both
+ \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)) *}
+
+example_proof
+ ML_val {* factorial 100 *} (* FIXME check/fix indentation *)
+ ML_command {* writeln (string_of_int (factorial 100)) *}
+qed
+
+
+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 via some derived ML functions.
+*}
+
+text %mlref {*
+ \begin{mldecls}
+ @{index_ML ML_Context.the_generic_context: "unit -> Context.generic"} \\
+ @{index_ML "Context.>> ": "(Context.generic -> Context.generic) -> unit"} \\
+ @{index_ML bind_thms: "string * thm list -> unit"} \\
+ @{index_ML bind_thm: "string * thm -> unit"} \\
+ \end{mldecls}
\begin{description}
- \item[fundamental law of programming]
- Whenever writing code, keep in mind: A program is
- written once, modified ten times, and read
- hundred times. So simplify its writing,
- always keep future modifications in mind,
- and never jeopardize readability. Every second you hesitate
- to spend on making your code more clear you will
- have to spend ten times understanding what you have
- written later on.
+ \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 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 bind_thm} is similar to @{ML bind_thms} but refers to a
+ singleton theorem.
+
+ \end{description}
+
+ It is very 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}. First, the standard token
+ language of ML is augmented by special syntactic entities of the
+ following form:
+
+ \indexouternonterm{antiquote}
+ \begin{rail}
+ antiquote: atsign lbrace nameref args rbrace | lbracesym | rbracesym
+ ;
+ \end{rail}
+
+ Here @{syntax nameref} and @{syntax args} are regular outer syntax
+ categories. Note that 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.
+
+ Special antiquotations like @{text "@{let \<dots>}"} or @{text "@{note
+ \<dots>}"} augment the compilation context without generating code. The
+ non-ASCII braces @{text "\<lbrace>"} and @{text "\<rbrace>"} allow to delimit the
+ effect by introducing local blocks within the pre-compilation
+ environment.
+
+ \medskip See also \cite{Wenzel-Chaieb:2007b} for a slightly broader
+ perspective on Isabelle/ML antiquotations. *}
+
+text %mlantiq {*
+ \begin{matharray}{rcl}
+ @{ML_antiquotation_def "let"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "note"} & : & @{text ML_antiquotation} \\
+ \end{matharray}
- \item[white space matters]
- Treat white space in your code as if it determines
- the meaning of code.
+ \begin{rail}
+ 'let' ((term + 'and') '=' term + 'and')
+ ;
+
+ 'note' (thmdef? thmrefs + 'and')
+ ;
+ \end{rail}
+
+ \begin{description}
+
+ \item @{text "@{let p = t}"} binds schematic variables in the
+ pattern @{text "p"} by higher-order matching against the term @{text
+ "t"}. This is analogous to the regular @{command_ref let} command
+ in the Isar proof language. The pre-compilation environment is
+ augmented by auxiliary term bindings, without emitting ML source.
+
+ \item @{text "@{note a = b\<^sub>1 \<dots> b\<^sub>n}"} recalls existing facts @{text
+ "b\<^sub>1, \<dots>, b\<^sub>n"}, binding the result as @{text a}. This is analogous to
+ the regular @{command_ref note} command in the Isar proof language.
+ The pre-compilation environment is augmented by auxiliary fact
+ bindings, without emitting ML source.
+
+ \end{description}
+*}
+
+text %mlex {* The following artificial example gives a first
+ impression about using the antiquotation elements introduced so far,
+ together with the basic @{text "@{thm}"} antiquotation defined
+ later.
+*}
- \begin{itemize}
+ML {*
+ \<lbrace>
+ @{let ?t = my_term}
+ @{note my_refl = reflexive [of ?t]}
+ fun foo th = Thm.transitive th @{thm my_refl}
+ \<rbrace>
+*}
+
+text {*
+ The extra block delimiters do not affect the compiled code itself, i.e.\
+ function @{ML foo} is available in the present context.
+*}
+
+
+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.
+
+ First of all, functions are usually \emph{curried}: the idea of
+ @{text "f"} taking @{text "n"} arguments of type @{text "\<tau>\<^sub>i"} (for
+ @{text "i \<in> {1, \<dots> n}"}) with result @{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 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 additional 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>\<^bsub>2\<^esub> \<rightarrow> \<tau>"} can be applied to @{text "x: \<tau>\<^sub>1"}
+ and the remaining @{text "(f x): \<tau>\<^sub>2 \<rightarrow> \<tau>"} passed to other functions
+ 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.
- \item The space bar is the easiest key to find on the keyboard,
- press it as often as necessary. @{verbatim "2 + 2"} is better
- than @{verbatim "2+2"}, likewise @{verbatim "f (x, y)"} is
- better than @{verbatim "f(x,y)"}.
+ This can be avoided by \emph{canonical argument order}, which
+ observes certain standard patterns of function definitions and
+ minimizes adhoc permutations in their application. In Isabelle/ML,
+ large portions of non-trivial source code are written without ever
+ using the infamous function @{text "swap: \<alpha> \<times> \<beta> \<rightarrow> \<beta> \<times> \<alpha>"} or the
+ combinator @{text "C: (\<alpha> \<rightarrow> \<beta> \<rightarrow> \<gamma>) \<rightarrow> (\<beta> \<rightarrow> \<alpha> \<rightarrow> \<gamma>)"}, which is not even
+ defined in our 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.
+
+ \medskip
+ \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}
+ \medskip
+
+ 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 in the application
+ @{text "forall (member B)"} to get directly a predicate again.
+
+ The update operation naturally ``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
+ follows: @{text "insert c \<circ> insert b \<circ> insert a"}. This works well,
+ apart from some awkwardness due to conventional mathematical
+ function composition, which can be easily amended as follows.
+*}
+
+
+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 if \emph{linear transformation}
+ applies a cascade of functions as follows: @{text "f\<^sub>n (\<dots> (f\<^sub>1 x))"}.
+ This becomes hard to read and maintain if the functions are
+ themselves 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"} & @{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).
- \item Restrict your lines to 80 characters. This will allow
- you to keep the beginning of a line in view while watching
- its end.\footnote{To acknowledge the lax practice of
- text editing these days, we tolerate as much as 100
- characters per line, but anything beyond 120 is not
- considered proper source text.}
+ \medskip
+ \begin{tabular}{lll}
+ @{text "(x, y) |-> f"} & @{text "\<equiv>"} & @{text "f x y"} \\
+ @{text "f #-> g"} & @{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}
+
+ %FIXME description!?
+*}
+
+
+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>"} that is
+ 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\<^bsub>n\<^esub>\<^bsub>\<^esub>"}. 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.
- \item Ban tabulators; they are a context-sensitive formatting
- feature and likely to confuse anyone not using your favorite
- editor.\footnote{Some modern programming language even
- forbid tabulators altogether according to the formal syntax
- definition.}
+ 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 Get rid of trailing whitespace. Instead, do not
- suppress a trailing newline at the end of your files.
+ \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 still has the
+ slightly more convenient historic @{ML Library.foldl} and @{ML
+ Library.foldr}. To avoid further confusion, all of this 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.
+*}
- \item Choose a generally accepted style of indentation,
- then use it systematically throughout the whole
- application. An indentation of two spaces is appropriate.
- Avoid dangling indentation.
+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 requires some practice to read and
+ write it fluently.
+
+ \medskip The next example elaborates this idea and demonstrates 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 ^ ">");
- \end{itemize}
+ 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
+ afresh.
+
+ The incremental @{ML add_content} avoids this by operating on a
+ buffer that is passed-through in a linear fashion. Using @{verbatim
+ "#>"} and contraction over the actual @{ML_type "Buffer.T"} argument
+ saves some additional boiler-plate, but requires again some
+ practice. 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 @{ML_type Buffer.T} that can be continued in further
+ linear transformations, folding etc. Thus it is more compositional
+ than the naive @{ML slow_content}. As a realistic example, compare
+ the slightly 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 for completeness
+ of the example. In many practical situations, it is customary to
+ defined 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.
- \item[cut-and-paste succeeds over copy-and-paste]
- \emph{Never} copy-and-paste code when programming. If you
- need the same piece of code twice, introduce a
- reasonable auxiliary function (if there is no
- such function, very likely you got something wrong).
- Any copy-and-paste will turn out to be painful
- when something has to be changed later on.
+ 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[comments]
- are a device which requires careful thinking before using
- it. The best comment for your code should be the code itself.
- Prefer efforts to write clear, understandable code
- over efforts to explain nasty code.
+ \item @{ML warning}~@{text "text"} outputs @{text "text"} as
+ warning, which typically means some extra emphasis on the front-end
+ side (color highlighting, icon).
+
+ \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, icon).
+
+ 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_msg}, but
+ the 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 here.
+ \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.
+ \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
+ needs to be issued by a \emph{single} invocation of @{ML writeln}
+ etc. with the functional concatenation of all message constituents.
+ \end{warn}
+*}
- \item[functional programming is based on functions]
- Model things as abstract as appropriate. Avoid unnecessarrily
- concrete datatype representations. For example, consider a function
- taking some table data structure as argument and performing
- lookups on it. Instead of taking a table, it could likewise
- take just a lookup function. Accustom
- your way of coding to the level of expressiveness a functional
- programming language is giving onto you.
+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 no longer necessary, because the ML
+ runtime system prints a detailed source position of the
+ corresponding @{verbatim 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 end-users or ML
+ users of a module can be often minimized by using plain user errors.
+
+ \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.
- \item[tuples]
- are often in the way. When there is no striking argument
- to tuple function arguments, just write your function curried.
+ 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 an indispensable tool to implement managed
+ evaluation of Isar 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 ERROR: "string -> exn"} \\
+ @{index_ML Fail: "string -> exn"} \\
+ @{index_ML Exn.is_interrupt: "exn -> bool"} \\
+ @{index_ML reraise: "exn -> 'a"} \\
+ @{index_ML exception_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 @{verbatim "(SOME"}~@{text "f
+ x"}~@{verbatim "handle _ => NONE)"} that is occasionally seen in
+ books about SML.
+
+ \item @{ML can} is similar to @{ML try} with more abstract result.
- \item[telling names]
- Any name should tell its purpose as exactly as possible, while
- keeping its length to the absolutely necessary minimum. Always
- give the same name to function arguments which have the same
- meaning. Separate words by underscores (@{verbatim
- int_of_string}, not @{verbatim intOfString}).\footnote{Some
- recent tools for Emacs include special precautions to cope with
- bumpy names in @{text "camelCase"}, e.g.\ for improved on-screen
- readability. It is easier to abstain from using such names in the
- first place.}
+ \item @{ML ERROR}~@{text "msg"} represents user errors; this
+ exception is always raised 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 exception_trace}~@{verbatim "(fn _ =>"}~@{text
+ "e"}@{verbatim ")"} evaluates expression @{text "e"} while printing
+ a full trace of its stack of nested exceptions (if possible,
+ depending on the ML platform).\footnote{In various versions of
+ Poly/ML the trace will appear on raw stdout of the Isabelle
+ process.}
+
+ Inserting @{ML exception_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 @{text "bool \<Rightarrow> 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 {* Thread-safe programming *}
+section {* Basic data types *}
+
+text {* The basis library proposal of SML97 need 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 can cause serious
+ 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 a ``symbol'' (see
+ \secref{sec:symbols}).
+
+ \end{description}
+*}
+
+
+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 (e.g.\ @{ML
+ 123456789012345678901234567890}) are forced to be of this one true
+ integer type --- overloading of SML97 is disabled.
+
+ Structure @{ML_struct IntInf} of SML97 is obsolete, always use
+ @{ML_struct Int}. Structure @{ML_struct Integer} in @{"file"
+ "~~/src/Pure/General/integer.ML"} provides some additional
+ operations.
+
+ \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 operations defined in
+ structure @{ML_struct Option} are alien to Isabelle/ML. The
+ operations shown above are defined in @{"file"
+ "~~/src/Pure/General/basics.ML"}, among others. *}
+
+
+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 last 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: more recent declarations normally take precedence over
+ earlier ones.
+
+ \end{description}
+*}
+
+text %mlex {* Using canonical @{ML fold} together with canonical @{ML
+ cons}, or similar standard operations, alternates the orientation of
+ data. The is quite natural and should not altered forcible by
+ inserting extra applications @{ML rev}. The alternative @{ML
+ fold_rev} can be used in the relatively 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 {*
- Recent versions of Poly/ML (5.2.1 or later) support robust
- multithreaded execution, based on native operating system threads of
- the underlying platform. Thus threads will actually be executed in
- parallel on multi-core systems. A speedup-factor of approximately
- 1.5--3 can be expected on a regular 4-core machine.\footnote{There
- is some inherent limitation of the speedup factor due to garbage
- collection, which is still sequential. It helps to provide initial
- heap space generously, using the \texttt{-H} option of Poly/ML.}
- Threads also help to organize advanced operations of the system,
- with explicit communication between sub-components, real-time
- conditions, time-outs etc.
+ \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, with standard conventions for 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. *}
+
+
+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.
- 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 immediately available
- to other threads, without requiring untyped character streams,
- awkward serialization etc.
+ 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. A maximum speedup-factor
+ of 3.0 on 4 cores and 5.0 on 8 cores can be
+ expected.\footnote{Further scalability is limited due to garbage
+ collection, which is still sequential in Poly/ML 5.2/5.3/5.4. It
+ helps to provide initial heap space generously, using the
+ \texttt{-H} option. Initial heap size needs to be scaled-up
+ together with the number of CPU cores: approximately 1--2\,GB per
+ core..}
- On the other hand, some programming guidelines need to be observed
- in order to make unprotected parallelism work out smoothly. While
- the ML system implementation is responsible to maintain basic
- integrity of the representation of ML values in memory, the
- application programmer needs to ensure that multithreaded execution
- does not break the intended semantics.
+ \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
+ immediately available to other threads: abstract values can be
+ passed between threads without copying or awkward serialization that
+ is typically required for explicit message passing between 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.
- \medskip \paragraph{Critical shared resources.} Actually only those
- parts outside the purely functional world of ML are critical. In
- particular, this covers
+ \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.\ those 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 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 direct I/O on shared channels, notably @{text "stdin"}, @{text
- "stdout"}, @{text "stderr"}.
+ \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 other processes.
\end{itemize}
- The majority of tools implemented within the Isabelle/Isar framework
- will not require any of these critical elements: nothing special
- needs to be observed when staying in the purely functional fragment
- of ML. Note that output via the official Isabelle channels does not
- count as direct I/O, so the operations @{ML "writeln"}, @{ML
- "warning"}, @{ML "tracing"} etc.\ are safe.
-
- Moreover, ML bindings within the toplevel environment (@{verbatim
- "type"}, @{verbatim val}, @{verbatim "structure"} etc.) due to
- run-time invocation of the compiler are also safe, because
- Isabelle/Isar manages this as part of the theory or proof context.
-
- \paragraph{Multithreading in Isabelle/Isar.} The theory loader
- automatically exploits the overall parallelism of independent nodes
- in the development graph, as well as the inherent irrelevance of
- proofs for goals being fully specified in advance. This means,
- checking of individual Isar proofs is parallelized by default.
- Beyond that, very sophisticated proof tools may use local
- parallelism internally, via the general programming model of
- ``future values'' (see also @{"file"
- "~~/src/Pure/Concurrent/future.ML"}).
-
- Any ML code that works relatively to the present background theory
- is already safe. Contextual data may be easily stored within the
- theory or proof context, thanks to the generic data concept of
- Isabelle/Isar (see \secref{sec:context-data}). This greatly
- diminishes the demand for global state information in the first
- place.
-
- \medskip In rare situations where actual mutable content needs to be
- manipulated, Isabelle provides a single \emph{critical section} that
- may be entered while preventing any other thread from doing the
- same. Entering the critical section without contention is very
- fast, and several basic system operations do so frequently. This
- also means that each thread should leave the critical section
- quickly, otherwise parallel execution performance may degrade
- significantly.
-
- Despite this potential bottle-neck, centralized locking is
- convenient, because it prevents deadlocks without demanding special
- precautions. Explicit communication demands other means, though.
- The high-level abstraction of synchronized variables @{"file"
- "~~/src/Pure/Concurrent/synchronized.ML"} enables parallel
- components to communicate via shared state; see also @{"file"
- "~~/src/Pure/Concurrent/mailbox.ML"} as canonical example.
-
- \paragraph{Good conduct of impure programs.} The following
- guidelines enable non-functional programs to participate in
- multithreading.
+ Isabelle/ML provides various mechanisms to avoid critical shared
+ resources in most practical 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 Minimize global state information. Using proper theory and
- proof context data will actually return to functional update of
- values, without any special precautions for multithreading. Apart
- from the fully general functors for theory and proof data (see
- \secref{sec:context-data}) there are drop-in replacements that
- emulate primitive references for common cases of \emph{configuration
- options} for type @{ML_type "bool"}/@{ML_type "int"}/@{ML_type
- "string"} (see structure @{ML_struct Config} and @{ML
- Attrib.config_bool} etc.), and lists of theorems (see functor
- @{ML_functor Named_Thms}).
+ \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 formerly writable flags.
- \item Keep components with local state information
- \emph{re-entrant}. Instead of poking initial values into (private)
- global references, create a new state record on each invocation, and
- pass that 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 sees its
- own copy. Occasionally, one might even return to plain functional
- updates on non-mutable record values here.
+ \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 Isolate process configuration flags. The main legitimate
- application of global references is to configure the whole process
- in a certain way, essentially affecting all threads. A typical
- example is the @{ML show_types} flag, which tells the pretty printer
- to output explicit type information for terms. Such flags usually
- do not affect the functionality of the core system, but only the
- view being presented to the user.
+ \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.
- Occasionally, such global process flags are treated like implicit
- arguments to certain operations, by using the @{ML setmp_CRITICAL} combinator
- for safe temporary assignment. Its traditional purpose was to
- ensure proper recovery of the original value when exceptions are
- raised in the body, now the functionality is extended to enter the
- \emph{critical section} (with its usual potential of degrading
- parallelism).
+ 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).
- Note that recovery of plain value passing semantics via @{ML
- setmp_CRITICAL}~@{text "ref value"} assumes that this @{text "ref"} is
- exclusively manipulated within the critical section. In particular,
- any persistent global assignment of @{text "ref := value"} needs to
- be marked critical as well, to prevent intruding another threads
- local view, and a lost-update in the global scope, too.
+ \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}
+*}
- Recall that in an open ``LCF-style'' system like Isabelle/Isar, the
- user participates in constructing the overall environment. This
- means that state-based facilities offered by one component will
- require special caution later on. So minimizing critical elements,
- by staying within the plain value-oriented view relative to theory
- or proof contexts most of the time, will also reduce the chance of
- mishaps occurring to end-users.
+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"} \\
- @{index_ML setmp_CRITICAL: "'a Unsynchronized.ref -> 'a -> ('b -> 'c) -> 'b -> 'c"} \\
+ \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 f"} evaluates @{text "f ()"}
- while staying within the critical section of Isabelle/Isar. No
- other thread may do so at the same time, but non-critical parallel
- execution will continue. The @{text "name"} argument serves for
- diagnostic purposes and might help to spot sources of congestion.
+ \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, and
+ several basic system operations do so frequently. Each thread
+ should leave the critical section quickly, otherwise parallel
+ performance may degrade.
\item @{ML CRITICAL} is the same as @{ML NAMED_CRITICAL} with empty
name argument.
- \item @{ML setmp_CRITICAL}~@{text "ref value f x"} evaluates @{text "f x"}
- while staying within the critical section and having @{text "ref :=
- value"} assigned temporarily. This recovers a value-passing
- semantics involving global references, regardless of exceptions or
- concurrency.
+ \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}, we
+ continue 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')"} we
+ are satisfied and assign the new state value @{text "x'"}, broadcast
+ a signal to all waiting threads on the associated condition
+ variable, and return the result @{text "y"}.
\end{description}
-*}
-
-chapter {* Basic library functions *}
-
-text {*
- Beyond the proposal of the SML/NJ basis library, Isabelle comes
- with its own library, from which selected parts are given here.
- These should encourage a study of the Isabelle sources,
- in particular files \emph{Pure/library.ML} and \emph{Pure/General/*.ML}.
-*}
-
-section {* Linear transformations \label{sec:ML-linear-trans} *}
-
-text %mlref {*
- \begin{mldecls}
- @{index_ML "op |> ": "'a * ('a -> 'b) -> 'b"} \\
- \end{mldecls}
-*}
-
-(*<*)
-typedecl foo
-consts foo :: foo
-ML {*
-val thy = Theory.copy @{theory}
-*}
-(*>*)
-
-text {*
- \noindent Many problems in functional programming can be thought of
- as linear transformations, i.e.~a caluclation starts with a
- particular value @{ML_text "x : foo"} which is then transformed
- by application of a function @{ML_text "f : foo -> foo"},
- continued by an application of a function @{ML_text "g : foo -> bar"},
- and so on. As a canoncial example, take functions enriching
- a theory by constant declararion and primitive definitions:
-
- \smallskip\begin{mldecls}
- @{ML "Sign.declare_const: (binding * typ) * mixfix
- -> theory -> term * theory"} \\
- @{ML "Thm.add_def: bool -> bool -> binding * term -> theory -> (string * thm) * theory"}
- \end{mldecls}
-
- \noindent Written with naive application, an addition of constant
- @{term bar} with type @{typ "foo \<Rightarrow> foo"} and
- a corresponding definition @{term "bar \<equiv> \<lambda>x. x"} would look like:
-
- \smallskip\begin{mldecls}
- @{ML "(fn (t, thy) => Thm.add_def false false
- (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})) thy)
- (Sign.declare_const
- ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn) thy)"}
- \end{mldecls}
-
- \noindent With increasing numbers of applications, this code gets quite
- unreadable. Further, it is unintuitive that things are
- written down in the opposite order as they actually ``happen''.
-*}
-
-text {*
- \noindent At this stage, Isabelle offers some combinators which allow
- for more convenient notation, most notably reverse application:
-
- \smallskip\begin{mldecls}
-@{ML "thy
-|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
-|> (fn (t, thy) => thy
-|> Thm.add_def false false
- (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))"}
- \end{mldecls}
-*}
-
-text %mlref {*
- \begin{mldecls}
- @{index_ML "op |-> ": "('c * 'a) * ('c -> 'a -> 'b) -> 'b"} \\
- @{index_ML "op |>> ": "('a * 'c) * ('a -> 'b) -> 'b * 'c"} \\
- @{index_ML "op ||> ": "('c * 'a) * ('a -> 'b) -> 'c * 'b"} \\
- @{index_ML "op ||>> ": "('c * 'a) * ('a -> 'd * 'b) -> ('c * 'd) * 'b"} \\
- \end{mldecls}
-*}
-
-text {*
- \noindent Usually, functions involving theory updates also return
- side results; to use these conveniently, yet another
- set of combinators is at hand, most notably @{ML "op |->"}
- which allows curried access to side results:
-
- \smallskip\begin{mldecls}
-@{ML "thy
-|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
-|-> (fn t => Thm.add_def false false
- (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))
-"}
- \end{mldecls}
-
- \noindent @{ML "op |>>"} allows for processing side results on their own:
-
- \smallskip\begin{mldecls}
-@{ML "thy
-|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
-|>> (fn t => Logic.mk_equals (t, @{term \"%x. x\"}))
-|-> (fn def => Thm.add_def false false (Binding.name \"bar_def\", def))
-"}
- \end{mldecls}
-
- \noindent Orthogonally, @{ML "op ||>"} applies a transformation
- in the presence of side results which are left unchanged:
-
- \smallskip\begin{mldecls}
-@{ML "thy
-|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
-||> Sign.add_path \"foobar\"
-|-> (fn t => Thm.add_def false false
- (Binding.name \"bar_def\", Logic.mk_equals (t, @{term \"%x. x\"})))
-||> Sign.restore_naming thy
-"}
- \end{mldecls}
-
- \noindent @{ML "op ||>>"} accumulates side results:
-
- \smallskip\begin{mldecls}
-@{ML "thy
-|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
-||>> Sign.declare_const ((Binding.name \"foobar\", @{typ \"foo => foo\"}), NoSyn)
-|-> (fn (t1, t2) => Thm.add_def false false
- (Binding.name \"bar_def\", Logic.mk_equals (t1, t2)))
-"}
- \end{mldecls}
-*}
-
-text %mlref {*
- \begin{mldecls}
- @{index_ML fold: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\
- @{index_ML fold_map: "('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b"} \\
- \end{mldecls}
+ There are some further variants of the general @{ML
+ Synchronized.guarded_access} combinator, see @{"file"
+ "~~/src/Pure/Concurrent/synchronized.ML"} for details.
*}
-text {*
- \noindent This principles naturally lift to \emph{lists} using
- the @{ML fold} and @{ML fold_map} combinators.
- The first lifts a single function
- \begin{quote}\footnotesize
- @{ML_text "f : 'a -> 'b -> 'b"} to @{ML_text "'a list -> 'b -> 'b"}
- \end{quote}
- such that
- \begin{quote}\footnotesize
- @{ML_text "y |> fold f [x1, x2, ..., x_n]"} \\
- \hspace*{2ex}@{text "\<leadsto>"} @{ML_text "y |> f x1 |> f x2 |> ... |> f x_n"}
- \end{quote}
- \noindent The second accumulates side results in a list by lifting
- a single function
- \begin{quote}\footnotesize
- @{ML_text "f : 'a -> 'b -> 'c * 'b"} to @{ML_text "'a list -> 'b -> 'c list * 'b"}
- \end{quote}
- such that
- \begin{quote}\footnotesize
- @{ML_text "y |> fold_map f [x1, x2, ..., x_n]"} \\
- \hspace*{2ex}@{text "\<leadsto>"} @{ML_text "y |> f x1 ||>> f x2 ||>> ... ||>> f x_n"} \\
- \hspace*{6ex}@{ML_text "||> (fn ((z1, z2), ..., z_n) => [z1, z2, ..., z_n])"}
- \end{quote}
-
- \noindent Example:
-
- \smallskip\begin{mldecls}
-@{ML "let
- val consts = [\"foo\", \"bar\"];
-in
- thy
- |> fold_map (fn const => Sign.declare_const
- ((Binding.name const, @{typ \"foo => foo\"}), NoSyn)) consts
- |>> map (fn t => Logic.mk_equals (t, @{term \"%x. x\"}))
- |-> (fn defs => fold_map (fn def =>
- Thm.add_def false false (Binding.empty, def)) defs)
-end"}
- \end{mldecls}
-*}
-
-text %mlref {*
- \begin{mldecls}
- @{index_ML "op #> ": "('a -> 'b) * ('b -> 'c) -> 'a -> 'c"} \\
- @{index_ML "op #-> ": "('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd"} \\
- @{index_ML "op #>> ": "('a -> 'c * 'b) * ('c -> 'd) -> 'a -> 'd * 'b"} \\
- @{index_ML "op ##> ": "('a -> 'c * 'b) * ('b -> 'd) -> 'a -> 'c * 'd"} \\
- @{index_ML "op ##>> ": "('a -> 'c * 'b) * ('b -> 'e * 'd) -> 'a -> ('c * 'e) * 'd"} \\
- \end{mldecls}
-*}
+text %mlex {* See @{"file" "~~/src/Pure/Concurrent/mailbox.ML"} how to
+ implement a mailbox as synchronized variable over a purely
+ functional queue.
-text {*
- \noindent All those linear combinators also exist in higher-order
- variants which do not expect a value on the left hand side
- but a function.
-*}
-
-text %mlref {*
- \begin{mldecls}
- @{index_ML "op ` ": "('b -> 'a) -> 'b -> 'a * 'b"} \\
- @{index_ML tap: "('b -> 'a) -> 'b -> 'b"} \\
- \end{mldecls}
-*}
-
-text {*
- \noindent These operators allow to ``query'' a context
- in a series of context transformations:
-
- \smallskip\begin{mldecls}
-@{ML "thy
-|> tap (fn _ => writeln \"now adding constant\")
-|> Sign.declare_const ((Binding.name \"bar\", @{typ \"foo => foo\"}), NoSyn)
-||>> `(fn thy => Sign.declared_const thy
- (Sign.full_name thy (Binding.name \"foobar\")))
-|-> (fn (t, b) => if b then I
- else Sign.declare_const
- ((Binding.name \"foobar\", @{typ \"foo => foo\"}), NoSyn) #> snd)
-"}
- \end{mldecls}
-*}
-
-section {* Options and partiality *}
-
-text %mlref {*
- \begin{mldecls}
- @{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"} \\
- @{index_ML try: "('a -> 'b) -> 'a -> 'b option"} \\
- @{index_ML can: "('a -> 'b) -> 'a -> bool"} \\
- \end{mldecls}
-*}
-
-text {*
- Standard selector functions on @{text option}s are provided. The
- @{ML try} and @{ML can} functions provide a convenient interface for
- handling exceptions -- both take as arguments a function @{ML_text f}
- together with a parameter @{ML_text x} and handle any exception during
- the evaluation of the application of @{ML_text f} to @{ML_text x}, either
- return a lifted result (@{ML NONE} on failure) or a boolean value
- (@{ML false} on failure).
+ \medskip The following example implements a counter that produces
+ positive integers that are unique over the runtime of the Isabelle
+ process:
*}
-
-section {* Common data structures *}
-
-subsection {* Lists (as set-like data structures) *}
-
-text {*
- \begin{mldecls}
- @{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 merge: "('a * 'a -> bool) -> 'a list * 'a list -> 'a list"} \\
- \end{mldecls}
-*}
-
-text {*
- Lists are often used as set-like data structures -- set-like in
- the sense that they support a notion of @{ML member}-ship,
- @{ML insert}-ing and @{ML remove}-ing, but are order-sensitive.
- This is convenient when implementing a history-like mechanism:
- @{ML insert} adds an element \emph{to the front} of a list,
- if not yet present; @{ML remove} removes \emph{all} occurences
- of a particular element. Correspondingly @{ML merge} implements a
- a merge on two lists suitable for merges of context data
- (\secref{sec:context-theory}).
-
- Functions are parametrized by an explicit equality function
- to accomplish overloaded equality; in most cases of monomorphic
- equality, writing @{ML "op ="} should suffice.
-*}
-
-subsection {* Association lists *}
-
-text {*
- \begin{mldecls}
- @{index_ML_exn AList.DUP} \\
- @{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"} \\
- @{index_ML AList.default: "('a * 'a -> bool) -> ('a * 'b) -> ('a * 'b) list -> ('a * 'b) list"} \\
- @{index_ML AList.delete: "('a * 'b -> bool) -> 'a -> ('b * 'c) list -> ('b * 'c) list"} \\
- @{index_ML AList.map_entry: "('a * 'b -> bool) -> 'a
- -> ('c -> 'c) -> ('b * 'c) list -> ('b * 'c) list"} \\
- @{index_ML AList.map_default: "('a * 'a -> bool) -> 'a * 'b -> ('b -> 'b)
- -> ('a * 'b) list -> ('a * 'b) list"} \\
- @{index_ML AList.join: "('a * 'a -> bool) -> ('a -> 'b * 'b -> 'b) (*exception DUP*)
- -> ('a * 'b) list * ('a * 'b) list -> ('a * 'b) list (*exception AList.DUP*)"} \\
- @{index_ML AList.merge: "('a * 'a -> bool) -> ('b * 'b -> bool)
- -> ('a * 'b) list * ('a * 'b) list -> ('a * 'b) list (*exception AList.DUP*)"}
- \end{mldecls}
+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;
*}
-text {*
- Association lists can be seens as an extension of set-like lists:
- on the one hand, they may be used to implement finite mappings,
- on the other hand, they remain order-sensitive and allow for
- multiple key-value-pair with the same key: @{ML AList.lookup}
- returns the \emph{first} value corresponding to a particular
- key, if present. @{ML AList.update} updates
- the \emph{first} occurence of a particular key; if no such
- key exists yet, the key-value-pair is added \emph{to the front}.
- @{ML AList.delete} only deletes the \emph{first} occurence of a key.
- @{ML AList.merge} provides an operation suitable for merges of context data
- (\secref{sec:context-theory}), where an equality parameter on
- values determines whether a merge should be considered a conflict.
- A slightly generalized operation if implementend by the @{ML AList.join}
- function which allows for explicit conflict resolution.
-*}
-
-subsection {* Tables *}
-
-text {*
- \begin{mldecls}
- @{index_ML_type "'a Symtab.table"} \\
- @{index_ML_exn Symtab.DUP: string} \\
- @{index_ML_exn Symtab.SAME} \\
- @{index_ML_exn Symtab.UNDEF: string} \\
- @{index_ML Symtab.empty: "'a Symtab.table"} \\
- @{index_ML Symtab.lookup: "'a Symtab.table -> string -> 'a option"} \\
- @{index_ML Symtab.defined: "'a Symtab.table -> string -> bool"} \\
- @{index_ML Symtab.update: "(string * 'a) -> 'a Symtab.table -> 'a Symtab.table"} \\
- @{index_ML Symtab.default: "string * 'a -> 'a Symtab.table -> 'a Symtab.table"} \\
- @{index_ML Symtab.delete: "string
- -> 'a Symtab.table -> 'a Symtab.table (*exception Symtab.UNDEF*)"} \\
- @{index_ML Symtab.map_entry: "string -> ('a -> 'a)
- -> 'a Symtab.table -> 'a Symtab.table"} \\
- @{index_ML Symtab.map_default: "(string * 'a) -> ('a -> 'a)
- -> 'a Symtab.table -> 'a Symtab.table"} \\
- @{index_ML Symtab.join: "(string -> 'a * 'a -> 'a) (*exception Symtab.DUP/Symtab.SAME*)
- -> 'a Symtab.table * 'a Symtab.table
- -> 'a Symtab.table (*exception Symtab.DUP*)"} \\
- @{index_ML Symtab.merge: "('a * 'a -> bool)
- -> 'a Symtab.table * 'a Symtab.table
- -> 'a Symtab.table (*exception Symtab.DUP*)"}
- \end{mldecls}
-*}
-
-text {*
- Tables are an efficient representation of finite mappings without
- any notion of order; due to their efficiency they should be used
- whenever such pure finite mappings are neccessary.
-
- The key type of tables must be given explicitly by instantiating
- the @{ML_functor Table} functor which takes the key type
- together with its @{ML_type order}; for convience, we restrict
- here to the @{ML_struct Symtab} instance with @{ML_type string}
- as key type.
-
- Most table functions correspond to those of association lists.
+ML {*
+ val a = next ();
+ val b = next ();
+ @{assert} (a <> b);
*}
end
\ No newline at end of file
--- a/doc-src/IsarImplementation/Thy/Prelim.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Prelim.thy Sun Oct 24 15:11:24 2010 -0700
@@ -140,23 +140,26 @@
\medskip There is a separate notion of \emph{theory reference} for
maintaining a live link to an evolving theory context: updates on
- drafts are propagated automatically. Dynamic updating stops after
- an explicit @{text "end"} only.
+ drafts are propagated automatically. Dynamic updating stops when
+ the next @{text "checkpoint"} is reached.
Derived entities may store a theory reference in order to indicate
- the context they belong to. This implicitly assumes monotonic
- reasoning, because the referenced context may become larger without
- further notice.
+ the formal context from which they are derived. This implicitly
+ assumes monotonic reasoning, because the referenced context may
+ become larger without further notice.
*}
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.checkpoint: "theory -> theory"} \\
@{index_ML Theory.copy: "theory -> theory"} \\
@{index_ML Theory.merge: "theory * theory -> theory"} \\
@{index_ML Theory.begin_theory: "string -> theory list -> theory"} \\
+ @{index_ML Theory.parents_of: "theory -> theory list"} \\
+ @{index_ML Theory.ancestors_of: "theory -> theory list"} \\
\end{mldecls}
\begin{mldecls}
@{index_ML_type theory_ref} \\
@@ -166,10 +169,15 @@
\begin{description}
- \item @{ML_type theory} represents theory contexts. This is
- essentially a linear type, with explicit runtime checking! Most
- internal theory operations destroy the original version, which then
- becomes ``stale''.
+ \item Type @{ML_type theory} represents theory contexts. This is
+ essentially a linear type, with explicit runtime checking.
+ Primitive theory operations destroy the original version, which then
+ becomes ``stale''. This can be prevented by explicit checkpointing,
+ which the system does at least at the boundary of toplevel command
+ transactions \secref{sec:isar-toplevel}.
+
+ \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.
@@ -191,11 +199,17 @@
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
+ a new theory based on the given parents. This ML function is
normally not invoked directly.
- \item @{ML_type theory_ref} represents a sliding reference to an
- always valid theory; updates on the original are propagated
+ \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).
+
+ \item Type @{ML_type theory_ref} represents a sliding reference to
+ an always valid theory; updates on the original are propagated
automatically.
\item @{ML "Theory.deref"}~@{text "thy_ref"} turns a @{ML_type
@@ -209,12 +223,39 @@
\end{description}
*}
+text %mlantiq {*
+ \begin{matharray}{rcl}
+ @{ML_antiquotation_def "theory"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "theory_ref"} & : & @{text ML_antiquotation} \\
+ \end{matharray}
+
+ \begin{rail}
+ ('theory' | 'theory\_ref') nameref?
+ ;
+ \end{rail}
+
+ \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_ref}"} is similar to @{text "@{theory}"}, but
+ produces a @{ML_type theory_ref} via @{ML "Theory.check_thy"} as
+ explained above.
+
+ \end{description}
+*}
+
subsection {* Proof context \label{sec:context-proof} *}
text {* A proof context is a container for pure data with a
- back-reference to the theory it belongs to. The @{text "init"}
- operation creates a proof context from a given theory.
+ back-reference to the theory from which it is derived. The @{text
+ "init"} operation creates a proof context from a given theory.
Modifications to draft theories are propagated to the proof context
as usual, but there is also an explicit @{text "transfer"} operation
to force resynchronization with more substantial updates to the
@@ -250,9 +291,9 @@
\begin{description}
- \item @{ML_type Proof.context} represents proof contexts. Elements
- of this type are essentially pure values, with a sliding reference
- to the background theory.
+ \item Type @{ML_type Proof.context} represents proof contexts.
+ Elements of this type are essentially pure values, with a sliding
+ reference to the background theory.
\item @{ML ProofContext.init_global}~@{text "thy"} produces a proof context
derived from @{text "thy"}, initializing all data.
@@ -268,6 +309,23 @@
\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} *}
@@ -295,7 +353,7 @@
\begin{description}
- \item @{ML_type Context.generic} is the direct sum of @{ML_type
+ \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"}.
@@ -331,9 +389,9 @@
\end{tabular}
\medskip
- \noindent 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"}.
+ 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
@@ -355,19 +413,24 @@
\end{tabular}
\medskip
- \noindent The @{text "init"} operation is supposed to produce a pure
- value from the given background theory and should be somehow
+ 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.
+ 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 these data declaration over type @{text "T"} result
- in an ML structure with the following signature:
+ \bigskip Any of the above data declarations over type @{text "T"}
+ result in an ML structure with the following signature:
\medskip
\begin{tabular}{ll}
@@ -377,13 +440,12 @@
\end{tabular}
\medskip
- \noindent These other operations provide exclusive access for the
- particular kind of context (theory, proof, or generic context).
- This interface fully 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.
-*}
+ 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}
@@ -422,9 +484,9 @@
end;
*}
-text {* \noindent The implementation uses private theory data
- internally, and only exposes an operation that involves explicit
- argument checking wrt.\ the given theory. *}
+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 =
@@ -437,28 +499,31 @@
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;
+ 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 {* We use @{ML_type "term Ord_List.T"} for reasonably efficient
- representation of a set of terms: all operations are linear in the
- number of stored elements. Here we assume that our users do not
- care about the declaration order, since that data structure forces
- its own arrangement of elements.
+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 @{verbatim 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.)
+ exponential blowup. Plain list append etc.\ must never be used for
+ theory data merges!
\medskip Our intended invariant is achieved as follows:
\begin{enumerate}
@@ -479,7 +544,7 @@
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
- strengthened by later declarations, for example.
+ 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
@@ -488,6 +553,111 @@
*}
+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 *}
+
+example_proof
+ note [[show_types = true]]
+ -- {* declaration within proof (forward mode) *}
+ term x
+
+ have "x = x"
+ using [[show_types = false]]
+ -- {* declaration within proof (backward mode) *}
+ ..
+qed
+
+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.config_bool: "string -> (Context.generic -> bool) ->
+ bool Config.T * (theory -> theory)"} \\
+ @{index_ML Attrib.config_int: "string -> (Context.generic -> int) ->
+ int Config.T * (theory -> theory)"} \\
+ @{index_ML Attrib.config_string: "string -> (Context.generic -> string) ->
+ string Config.T * (theory -> theory)"} \\
+ \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, setup) ="}~@{ML Attrib.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. The @{text "setup"} function
+ needs to be applied to the theory initially, in order to make
+ concrete declaration syntax available to the user.
+
+ \item @{ML Attrib.config_int} 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, my_flag_setup) =
+ Attrib.config_bool "my_flag" (K false)
+*}
+setup my_flag_setup
+
+text {* Now the user can refer to @{attribute my_flag} in
+ declarations, while we 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) *}
+
+example_proof
+ {
+ note [[my_flag = false]]
+ ML_val {* @{assert} (Config.get @{context} my_flag = false) *}
+ }
+ ML_val {* @{assert} (Config.get @{context} my_flag = true) *}
+qed
+
+
section {* Names \label{sec:names} *}
text {* In principle, a name is just a string, but there are various
@@ -517,7 +687,7 @@
*}
-subsection {* Strings of symbols *}
+subsection {* Strings of symbols \label{sec:symbols} *}
text {* A \emph{symbol} constitutes the smallest textual unit in
Isabelle --- raw ML characters are normally not encountered at all!
@@ -549,13 +719,13 @@
\end{enumerate}
- \noindent 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 @{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
@@ -594,26 +764,25 @@
\begin{description}
- \item @{ML_type "Symbol.symbol"} represents individual Isabelle
+ \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 supercedes @{ML
+ 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 --- which is the most common case --- do not
- require extra memory in Poly/ML.}
+ 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 @{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 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.
@@ -692,8 +861,8 @@
\item @{ML Name.skolem}~@{text "name"} produces a Skolem name by
adding two underscores.
- \item @{ML_type Name.context} represents the context of already used
- names; the initial value is @{ML "Name.context"}.
+ \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.
@@ -716,6 +885,37 @@
\end{description}
*}
+text %mlex {* The following simple examples demonstrate how to produce
+ fresh names from the initial @{ML Name.context}. *}
+
+ML {*
+ val list1 = Name.invents Name.context "a" 5;
+ @{assert} (list1 = ["a", "b", "c", "d", "e"]);
+
+ val list2 =
+ #1 (Name.variants ["x", "x", "a", "a", "'a", "'a"] Name.context);
+ @{assert} (list2 = ["x", "xa", "a", "aa", "'a", "'aa"]);
+*}
+
+text {* \medskip The same works reletively to the formal context as
+ follows. *}
+
+locale ex = fixes a b c :: 'a
+begin
+
+ML {*
+ val names = Variable.names_of @{context};
+
+ val list1 = Name.invents names "a" 5;
+ @{assert} (list1 = ["d", "e", "f", "g", "h"]);
+
+ val list2 =
+ #1 (Name.variants ["x", "x", "a", "a", "'a", "'a"] names);
+ @{assert} (list2 = ["x", "xa", "aa", "ab", "'aa", "'ab"]);
+*}
+
+end
+
subsection {* Indexed names \label{sec:indexname} *}
@@ -756,15 +956,15 @@
text %mlref {*
\begin{mldecls}
- @{index_ML_type indexname} \\
+ @{index_ML_type indexname: "string * int"} \\
\end{mldecls}
\begin{description}
- \item @{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
+ \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}
@@ -847,13 +1047,12 @@
main equation of this ``chemical reaction'' when binding new
entities in a context is as follows:
- \smallskip
+ \medskip
\begin{tabular}{l}
@{text "binding + naming \<longrightarrow> long name + name space accesses"}
\end{tabular}
- \smallskip
- \medskip As a general principle, there is a separate name space for
+ \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
@@ -866,10 +1065,17 @@
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 are better suffixed in addition to the usual qualification,
- e.g.\ @{text "c_type.intro"} and @{text "c_class.intro"} for
- theorems related to type @{text "c"} and class @{text "c"},
- respectively.
+ 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 {*
@@ -901,13 +1107,15 @@
\begin{description}
- \item @{ML_type binding} represents the abstract concept of name
- bindings.
+ \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"}.
+ 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
@@ -932,8 +1140,8 @@
representation for human-readable output, together with some formal
markup that might get used in GUI front-ends, for example.
- \item @{ML_type Name_Space.naming} represents the abstract concept of
- a naming policy.
+ \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
@@ -946,7 +1154,7 @@
name binding (usually a basic name) into the fully qualified
internal name, according to the given naming policy.
- \item @{ML_type Name_Space.T} represents name spaces.
+ \item Type @{ML_type Name_Space.T} represents name spaces.
\item @{ML Name_Space.empty}~@{text "kind"} and @{ML Name_Space.merge}~@{text
"(space\<^isub>1, space\<^isub>2)"} are the canonical operations for
@@ -981,4 +1189,42 @@
\end{description}
*}
+text %mlantiq {*
+ \begin{matharray}{rcl}
+ @{ML_antiquotation_def "binding"} & : & @{text ML_antiquotation} \\
+ \end{matharray}
+
+ \begin{rail}
+ 'binding' name
+ ;
+ \end{rail}
+
+ \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.str_of (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 advanced feedback given to the user. *}
+
end
--- a/doc-src/IsarImplementation/Thy/Proof.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Proof.thy Sun Oct 24 15:11:24 2010 -0700
@@ -52,13 +52,26 @@
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\<^isub>\<alpha> = x\<^isub>\<alpha>"} produces in the first step
- @{text "x: term \<turnstile> x\<^isub>\<alpha> = x\<^isub>\<alpha>"} for fixed @{text "\<alpha>"},
- and only in the second step @{text "\<turnstile> ?x\<^isub>?\<^isub>\<alpha> =
- ?x\<^isub>?\<^isub>\<alpha>"} for schematic @{text "?x"} and @{text "?\<alpha>"}.
+ term, \<alpha>: type \<turnstile> x\<^isub>\<alpha> \<equiv> x\<^isub>\<alpha>"} produces in the first step @{text "x: term
+ \<turnstile> x\<^isub>\<alpha> \<equiv> x\<^isub>\<alpha>"} for fixed @{text "\<alpha>"}, and only in the second step
+ @{text "\<turnstile> ?x\<^isub>?\<^isub>\<alpha> \<equiv> ?x\<^isub>?\<^isub>\<alpha>"} for schematic @{text "?x"} and @{text "?\<alpha>"}.
+ The following Isar source text illustrates this scenario.
+*}
- \medskip The Isabelle/Isar proof context manages the gory details of
- term vs.\ type variables, with high-level principles for moving the
+example_proof
+ {
+ 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"} *}
+qed
+
+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
@@ -145,12 +158,8 @@
\end{description}
*}
-text %mlex {* The following example (in theory @{theory Pure}) shows
- how to work with fixed term and type parameters work with
- type-inference.
-*}
-
-typedecl foo -- {* some basic type for testing purposes *}
+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*)
@@ -164,11 +173,11 @@
val t1' = singleton (Variable.polymorphic ctxt1) t1;
(*term u enforces specific type assignment*)
- val u = Syntax.read_term ctxt1 "(x::foo) \<equiv> y";
+ 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::foo is enforced*)
+ val t2 = Syntax.read_term ctxt2 "x"; (*x::nat is enforced*)
*}
text {* In the above example, the starting context had been derived
@@ -185,15 +194,13 @@
ctxt0 |> Variable.variant_fixes ["x", "x", "x"];
*}
-text {* \noindent Subsequent ML code can now work with the invented
- names of @{verbatim x1}, @{verbatim x2}, @{verbatim 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:
-*}
+text {* The following ML code can now work with the invented names of
+ @{verbatim x1}, @{verbatim x2}, @{verbatim 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: *}
-lemma "PROP XXX"
-proof -
+example_proof
ML_prf %"ML" {*
val ctxt0 = @{context};
@@ -206,10 +213,10 @@
*}
oops
-text {* \noindent In this situation @{ML Variable.add_fixes} and @{ML
+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} *}
@@ -282,11 +289,11 @@
\begin{description}
- \item @{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 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 "A"} turns proposition @{text
"A"} into a primitive assumption @{text "A \<turnstile> A'"}, where the
@@ -330,11 +337,10 @@
val r = Assumption.export false ctxt1 ctxt0 eq';
*}
-text {* \noindent 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 "ProofContext.export"}).
-*}
+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 "ProofContext.export"}). *}
section {* Structured goals and results \label{sec:struct-goals} *}
@@ -373,18 +379,24 @@
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 @{text "\<OBTAIN>"} and
- @{text "\<GUESS>"} elements. Final results, which may not refer to
+ \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.
-*}
+ 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 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}
@@ -394,8 +406,8 @@
({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"} \\
+ @{index_ML Obtain.result: "(Proof.context -> tactic) -> thm list ->
+ Proof.context -> ((string * cterm) list * thm list) * Proof.context"} \\
\end{mldecls}
\begin{description}
@@ -412,6 +424,12 @@
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
@@ -431,4 +449,41 @@
\end{description}
*}
+text %mlex {* The following minimal example illustrates how to access
+ the focus information of a structured goal state. *}
+
+example_proof
+ 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}. *}
+
+example_proof
+ 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 (ProofContext.export ctxt1 ctxt0) @{thm refl};
+ *}
+ ML_prf %"ML" {*
+ ProofContext.export ctxt1 ctxt0 [Thm.reflexive x]
+ handle ERROR msg => (warning msg; []);
+ *}
+qed
+
end
--- a/doc-src/IsarImplementation/Thy/Syntax.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Syntax.thy Sun Oct 24 15:11:24 2010 -0700
@@ -6,10 +6,91 @@
text FIXME
-section {* Parsing and printing *}
+section {* Reading and pretty printing \label{sec:read-print} *}
+
+text FIXME
+
+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}
+
+ \begin{description}
+
+ \item FIXME
+
+ \end{description}
+*}
+
+
+section {* Parsing and unparsing \label{sec:parse-unparse} *}
text FIXME
+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}
+
+ \begin{description}
+
+ \item FIXME
+
+ \end{description}
+*}
+
+
section {* Checking and unchecking \label{sec:term-check} *}
+text FIXME
+
+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}
+
+ \begin{description}
+
+ \item FIXME
+
+ \end{description}
+*}
+
+
+section {* Syntax translations *}
+
+text FIXME
+
+text %mlantiq {*
+ \begin{matharray}{rcl}
+ @{ML_antiquotation_def "class_syntax"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "type_syntax"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "const_syntax"} & : & @{text ML_antiquotation} \\
+ @{ML_antiquotation_def "syntax_const"} & : & @{text ML_antiquotation} \\
+ \end{matharray}
+
+ \begin{rail}
+ ('class_syntax' | 'type_syntax' | 'const_syntax' | 'syntax_const') name
+ ;
+ \end{rail}
+
+ \begin{description}
+
+ \item FIXME
+
+ \end{description}
+*}
+
end
--- a/doc-src/IsarImplementation/Thy/Tactic.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Tactic.thy Sun Oct 24 15:11:24 2010 -0700
@@ -86,7 +86,7 @@
*}
-section {* Tactics *}
+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.\
@@ -180,14 +180,14 @@
\begin{description}
- \item @{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 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 @{ML_type "int -> tactic"} represents tactics with explicit
- subgoal addressing, with well-formedness conditions as described
- above.
+ \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.
@@ -406,6 +406,8 @@
Various search strategies may be expressed via tacticals.
\medskip FIXME
-*}
+
+ \medskip The chapter on tacticals in \cite{isabelle-ref} is still
+ applicable, despite a few outdated details. *}
end
--- a/doc-src/IsarImplementation/Thy/document/Base.tex Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/Base.tex Sun Oct 24 15:11:24 2010 -0700
@@ -9,10 +9,43 @@
\isatagtheory
\isacommand{theory}\isamarkupfalse%
\ Base\isanewline
-\isakeyword{imports}\ Pure\isanewline
+\isakeyword{imports}\ Main\isanewline
\isakeyword{uses}\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isacharslash}{\isachardot}{\isachardot}{\isacharslash}antiquote{\isacharunderscore}setup{\isachardot}ML{\isachardoublequoteclose}\isanewline
-\isakeyword{begin}\isanewline
+\isakeyword{begin}%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+\isanewline
+%
+\endisadelimtheory
+\isanewline
+%
+\isadelimML
\isanewline
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ ML{\isacharunderscore}Antiquote{\isachardot}inline\ {\isachardoublequote}assert{\isachardoublequote}\isanewline
+\ \ \ \ {\isacharparenleft}Scan{\isachardot}succeed\ {\isachardoublequote}{\isacharparenleft}fn\ b\ {\isacharequal}{\isachargreater}\ if\ b\ then\ {\isacharparenleft}{\isacharparenright}\ else\ raise\ General{\isachardot}Fail\ {\isacharbackslash}{\isachardoublequote}Assertion\ failed{\isacharbackslash}{\isachardoublequote}{\isacharparenright}{\isachardoublequote}{\isacharparenright}\isanewline
+{\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+\isanewline
+%
+\endisadelimML
+%
+\isadelimtheory
+\isanewline
+%
+\endisadelimtheory
+%
+\isatagtheory
\isacommand{end}\isamarkupfalse%
%
\endisatagtheory
--- a/doc-src/IsarImplementation/Thy/document/Integration.tex Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/Integration.tex Sun Oct 24 15:11:24 2010 -0700
@@ -29,13 +29,12 @@
\begin{isamarkuptext}%
The Isar toplevel may be considered the centeral 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. We
- shall even incorporate the existing {\ML} toplevel of the compiler
- and run-time system (cf.\ \secref{sec:ML-toplevel}).
+ 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
+ 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
@@ -93,11 +92,11 @@
\begin{description}
- \item \verb|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 Type \verb|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 \verb|Toplevel.UNDEF| is raised for undefined toplevel
operations. Many operations work only partially for certain cases,
@@ -121,7 +120,7 @@
information for each Isar command being executed.
\item \verb|Toplevel.profiling|~\verb|:=|~\isa{n} controls
- low-level profiling of the underlying {\ML} runtime system. For
+ low-level profiling of the underlying ML runtime system. For
Poly/ML, \isa{n\ {\isacharequal}\ {\isadigit{1}}} means time and \isa{n\ {\isacharequal}\ {\isadigit{2}}} space
profiling.
@@ -136,6 +135,35 @@
%
\endisadelimmlref
%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
+\isatagmlantiq
+%
+\begin{isamarkuptext}%
+\begin{matharray}{rcl}
+ \indexdef{}{ML antiquotation}{Isar.state}\hypertarget{ML antiquotation.Isar.state}{\hyperlink{ML antiquotation.Isar.state}{\mbox{\isa{Isar{\isachardot}state}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \end{matharray}
+
+ \begin{description}
+
+ \item \isa{{\isacharat}{\isacharbraceleft}Isar{\isachardot}state{\isacharbraceright}} refers to Isar toplevel state at that
+ point --- as abstract value.
+
+ This only works for diagnostic ML commands, such as \hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isacharunderscore}val}}}} or \hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isacharunderscore}command}}}}.
+
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlantiq
+{\isafoldmlantiq}%
+%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
\isamarkupsubsection{Toplevel transitions \label{sec:toplevel-transition}%
}
\isamarkuptrue%
@@ -232,141 +260,6 @@
%
\endisadelimmlref
%
-\isamarkupsubsection{Toplevel control%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-There are a few special control commands that modify the behavior
- the toplevel itself, and only make sense in interactive mode. Under
- normal circumstances, the user encounters these only implicitly as
- part of the protocol between the Isabelle/Isar system and a
- user-interface such as Proof~General.
-
- \begin{description}
-
- \item \isacommand{undo} follows the three-level hierarchy of empty
- toplevel vs.\ theory vs.\ proof: undo within a proof reverts to the
- previous proof context, undo after a proof reverts to the theory
- before the initial goal statement, undo of a theory command reverts
- to the previous theory value, undo of a theory header discontinues
- the current theory development and removes it from the theory
- database (\secref{sec:theory-database}).
-
- \item \isacommand{kill} aborts the current level of development:
- kill in a proof context reverts to the theory before the initial
- goal statement, kill in a theory context aborts the current theory
- development, removing it from the database.
-
- \item \isacommand{exit} drops out of the Isar toplevel into the
- underlying {\ML} toplevel (\secref{sec:ML-toplevel}). The Isar
- toplevel state is preserved and may be continued later.
-
- \item \isacommand{quit} terminates the Isabelle/Isar process without
- saving.
-
- \end{description}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsection{ML toplevel \label{sec:ML-toplevel}%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-The {\ML} toplevel provides a read-compile-eval-print loop for {\ML}
- values, types, structures, and functors. {\ML} declarations operate
- on the global system state, which consists of the compiler
- environment plus the values of {\ML} reference variables. There is
- no clean way to undo {\ML} declarations, except for reverting to a
- previously saved state of the whole Isabelle process. {\ML} input
- is either read interactively from a TTY, or from a string (usually
- within a theory text), or from a source file (usually loaded from a
- theory).
-
- Whenever the {\ML} toplevel is active, the current Isabelle theory
- context is passed as an internal reference variable. Thus {\ML}
- code may access the theory context during compilation, it may even
- change the value of a theory being under construction --- while
- observing the usual linearity restrictions
- (cf.~\secref{sec:context-theory}).%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimmlref
-%
-\endisadelimmlref
-%
-\isatagmlref
-%
-\begin{isamarkuptext}%
-\begin{mldecls}
- \indexdef{}{ML}{ML\_Context.the\_generic\_context}\verb|ML_Context.the_generic_context: unit -> Context.generic| \\
- \indexdef{}{ML}{Context.$>$$>$ }\verb|Context.>> : (Context.generic -> Context.generic) -> unit| \\
- \end{mldecls}
-
- \begin{description}
-
- \item \verb|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 \verb|ML_Context.the_generic_context ()|
- correctly. Recall that evaluation of a function body is delayed
- until actual runtime. Moreover, persistent {\ML} toplevel bindings
- to an unfinished theory should be avoided: code should either
- project out the desired information immediately, or produce an
- explicit \verb|theory_ref| (cf.\ \secref{sec:context-theory}).
-
- \item \verb|Context.>>|~\isa{f} applies context transformation
- \isa{f} to the implicit context of the {\ML} toplevel.
-
- \end{description}
-
- It is very important to note that the above functions are really
- restricted to the compile time, even though the {\ML} compiler is
- invoked at runtime! The majority of {\ML} code uses explicit
- functional arguments of a theory or proof context instead. Thus it
- may be invoked for an arbitrary context later on, without having to
- worry about any operational details.
-
- \bigskip
-
- \begin{mldecls}
- \indexdef{}{ML}{Isar.main}\verb|Isar.main: unit -> unit| \\
- \indexdef{}{ML}{Isar.loop}\verb|Isar.loop: unit -> unit| \\
- \indexdef{}{ML}{Isar.state}\verb|Isar.state: unit -> Toplevel.state| \\
- \indexdef{}{ML}{Isar.exn}\verb|Isar.exn: unit -> (exn * string) option| \\
- \indexdef{}{ML}{Isar.goal}\verb|Isar.goal: unit ->|\isasep\isanewline%
-\verb| {context: Proof.context, facts: thm list, goal: thm}| \\
- \end{mldecls}
-
- \begin{description}
-
- \item \verb|Isar.main ()| invokes the Isar toplevel from {\ML},
- initializing an empty toplevel state.
-
- \item \verb|Isar.loop ()| continues the Isar toplevel with the
- current state, after having dropped out of the Isar toplevel loop.
-
- \item \verb|Isar.state ()| and \verb|Isar.exn ()| get current
- toplevel state and error condition, respectively. This only works
- after having dropped out of the Isar toplevel loop.
-
- \item \verb|Isar.goal ()| produces the full Isar goal state,
- consisting of proof context, facts that have been indicated for
- immediate use, and the tactical goal according to
- \secref{sec:tactical-goals}.
-
- \end{description}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagmlref
-{\isafoldmlref}%
-%
-\isadelimmlref
-%
-\endisadelimmlref
-%
\isamarkupsection{Theory database \label{sec:theory-database}%
}
\isamarkuptrue%
@@ -384,10 +277,10 @@
name space is flat!
Theory \isa{A} is associated with the main theory file \isa{A}\verb,.thy,, which needs to be accessible through the theory
- loader path. Any number of additional {\ML} source files may be
+ loader path. Any number of additional ML source files may be
associated with each theory, by declaring these dependencies in the
theory header as \isa{{\isasymUSES}}, and loading them consecutively
- within the theory context. The system keeps track of incoming {\ML}
+ 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 \isa{update} and \isa{remove}:
--- a/doc-src/IsarImplementation/Thy/document/Isar.tex Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/Isar.tex Sun Oct 24 15:11:24 2010 -0700
@@ -23,16 +23,33 @@
\isamarkuptrue%
%
\begin{isamarkuptext}%
-The primary Isar language consists of three main categories of
- language elements:
+The Isar proof language (see also \cite[\S2]{isabelle-isar-ref})
+ consists of three main categories of language elements:
\begin{enumerate}
- \item Proof commands
+ \item Proof \emph{commands} define the primary language of
+ transactions of the underlying Isar/VM interpreter. Typical
+ examples are \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}, \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}, \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}, \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}, and \hyperlink{command.qed}{\mbox{\isa{\isacommand{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 methods
+ \item Proof \emph{methods} define a secondary language of mixed
+ forward-backward refinement steps involving facts and goals.
+ Typical examples are \hyperlink{method.rule}{\mbox{\isa{rule}}}, \hyperlink{method.unfold}{\mbox{\isa{unfold}}}, and \hyperlink{method.simp}{\mbox{\isa{simp}}}.
- \item Attributes
+ Methods can occur in certain well-defined parts of the Isar proof
+ language, say as arguments to \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}, \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}},
+ or \hyperlink{command.by}{\mbox{\isa{\isacommand{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 \hyperlink{attribute.intro}{\mbox{\isa{intro}}} (which affects the context),
+ and \hyperlink{attribute.symmetric}{\mbox{\isa{symmetric}}} (which affects the theorem).
\end{enumerate}%
\end{isamarkuptext}%
@@ -43,7 +60,227 @@
\isamarkuptrue%
%
\begin{isamarkuptext}%
-FIXME%
+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: 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 stores the final
+ result in a suitable context data slot. Thus a proof can be
+ incorporated into the context of some user-space tool, without
+ modifying the Isar proof language itself.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML type}{Proof.state}\verb|type Proof.state| \\
+ \indexdef{}{ML}{Proof.assert\_forward}\verb|Proof.assert_forward: Proof.state -> Proof.state| \\
+ \indexdef{}{ML}{Proof.assert\_chain}\verb|Proof.assert_chain: Proof.state -> Proof.state| \\
+ \indexdef{}{ML}{Proof.assert\_backward}\verb|Proof.assert_backward: Proof.state -> Proof.state| \\
+ \indexdef{}{ML}{Proof.simple\_goal}\verb|Proof.simple_goal: Proof.state -> {context: Proof.context, goal: thm}| \\
+ \indexdef{}{ML}{Proof.goal}\verb|Proof.goal: Proof.state ->|\isasep\isanewline%
+\verb| {context: Proof.context, facts: thm list, goal: thm}| \\
+ \indexdef{}{ML}{Proof.raw\_goal}\verb|Proof.raw_goal: Proof.state ->|\isasep\isanewline%
+\verb| {context: Proof.context, facts: thm list, goal: thm}| \\
+ \indexdef{}{ML}{Proof.theorem}\verb|Proof.theorem: Method.text option ->|\isasep\isanewline%
+\verb| (thm list list -> Proof.context -> Proof.context) ->|\isasep\isanewline%
+\verb| (term * term list) list list -> Proof.context -> Proof.state| \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type \verb|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 (``\isa{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 \verb|Proof.assert_forward|, \verb|Proof.assert_chain|, \verb|Proof.assert_backward| are partial identity functions that fail
+ unless a certain linguistic mode is active, namely ``\isa{proof{\isacharparenleft}state{\isacharparenright}}'', ``\isa{proof{\isacharparenleft}chain{\isacharparenright}}'', ``\isa{proof{\isacharparenleft}prove{\isacharparenright}}'', 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 \verb|Proof.simple_goal|~\isa{state} returns the structured
+ Isar goal (if available) in the form seen by ``simple'' methods
+ (like \hyperlink{method.simp}{\mbox{\isa{simp}}} or \hyperlink{method.blast}{\mbox{\isa{blast}}}). The Isar goal facts are
+ already inserted as premises into the subgoals, which are presented
+ individually as in \verb|Proof.goal|.
+
+ \item \verb|Proof.goal|~\isa{state} returns the structured Isar
+ goal (if available) in the form seen by regular methods (like
+ \hyperlink{method.rule}{\mbox{\isa{rule}}}). The auxiliary internal encoding of Pure
+ conjunctions is split into individual subgoals as usual.
+
+ \item \verb|Proof.raw_goal|~\isa{state} returns the structured
+ Isar goal (if available) in the raw internal form seen by ``raw''
+ methods (like \hyperlink{method.induct}{\mbox{\isa{induct}}}). This form is rarely appropriate
+ for dignostic tools; \verb|Proof.simple_goal| or \verb|Proof.goal|
+ should be used in most situations.
+
+ \item \verb|Proof.theorem|~\isa{before{\isacharunderscore}qed\ after{\isacharunderscore}qed\ statement\ ctxt}
+ initializes a toplevel Isar proof state within a given context.
+
+ The optional \isa{before{\isacharunderscore}qed} method is applied at the end of
+ the proof, just before extracting the result (this feature is rarely
+ used).
+
+ The \isa{after{\isacharunderscore}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 \verb|Proof.context|, but the usual
+ wrapping of toplevel proofs into command transactions will provide a
+ \verb|local_theory| here (see also \chref{ch:local-theory}).
+ This usually affects the way how results are stored.
+
+ The \isa{statement} is given as a nested list of terms, each
+ associated with optional \hyperlink{keyword.is}{\mbox{\isa{\isakeyword{is}}}} patterns as usual in the
+ Isar source language. The original list structure over terms is
+ turned into one over theorems when \isa{after{\isacharunderscore}qed} is invoked.
+
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
+\isatagmlantiq
+%
+\begin{isamarkuptext}%
+\begin{matharray}{rcl}
+ \indexdef{}{ML antiquotation}{Isar.goal}\hypertarget{ML antiquotation.Isar.goal}{\hyperlink{ML antiquotation.Isar.goal}{\mbox{\isa{Isar{\isachardot}goal}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \end{matharray}
+
+ \begin{description}
+
+ \item \isa{{\isacharat}{\isacharbraceleft}Isar{\isachardot}goal{\isacharbraceright}} 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 \hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isacharunderscore}val}}}} or \hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isacharunderscore}command}}}}.
+
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlantiq
+{\isafoldmlantiq}%
+%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isatagmlex
+%
+\begin{isamarkuptext}%
+The following example peeks at a certain goal configuration.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlex
+{\isafoldmlex}%
+%
+\isadelimmlex
+%
+\endisadelimmlex
+\isacommand{example{\isacharunderscore}proof}\isamarkupfalse%
+\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{have}\isamarkupfalse%
+\ A\ \isakeyword{and}\ B\ \isakeyword{and}\ C%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+%
+\isadelimML
+\ \ \ \ %
+\endisadelimML
+%
+\isatagML
+\isacommand{ML{\isacharunderscore}val}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ \ \ \ \ val\ n\ {\isacharequal}\ Thm{\isachardot}nprems{\isacharunderscore}of\ {\isacharparenleft}{\isacharhash}goal\ %
+\isaantiq
+Isar{\isachardot}goal%
+\endisaantiq
+{\isacharparenright}{\isacharsemicolon}\isanewline
+\ \ \ \ \ \ %
+\isaantiq
+assert%
+\endisaantiq
+\ {\isacharparenleft}n\ {\isacharequal}\ {\isadigit{3}}{\isacharparenright}{\isacharsemicolon}\isanewline
+\ \ \ \ {\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+\isanewline
+%
+\endisadelimML
+%
+\isadelimproof
+\ \ \ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{oops}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\begin{isamarkuptext}%
+Here we see 3 individual subgoals in the same way as regular
+ proof methods would do.%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -52,19 +289,627 @@
\isamarkuptrue%
%
\begin{isamarkuptext}%
-FIXME%
+A \isa{method} is a function \isa{context\ {\isasymrightarrow}\ thm\isactrlsup {\isacharasterisk}\ {\isasymrightarrow}\ goal\ {\isasymrightarrow}\ {\isacharparenleft}cases\ {\isasymtimes}\ goal{\isacharparenright}\isactrlsup {\isacharasterisk}\isactrlsup {\isacharasterisk}} 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.\ \hyperlink{method.rule-tac}{\mbox{\isa{rule{\isacharunderscore}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 \isa{meth\isactrlsup {\isacharplus}}
+ without looping, while the trivial do-nothing case can be recovered
+ via \isa{meth\isactrlsup {\isacharquery}}.}
+
+ Exception: trivial stuttering steps, such as ``\hyperlink{method.-}{\mbox{\isa{{\isacharminus}}}}'' or
+ \hyperlink{method.succeed}{\mbox{\isa{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 is embedded
+ into Isar as arguments to certain commands, e.g.\ \hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}} or
+ \hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}}. User-space additions are reasonably easy by
+ plugging suitable method-valued parser functions into the framework,
+ using the \hyperlink{command.method-setup}{\mbox{\isa{\isacommand{method{\isacharunderscore}setup}}}} command, for example.
+
+ To get a better idea about the range of possibilities, consider the
+ following Isar proof schemes. First some general form of structured
+ proof text:
+
+ \medskip
+ \begin{tabular}{l}
+ \hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{facts\isactrlsub {\isadigit{1}}}~\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{props}~\hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}~\isa{facts\isactrlsub {\isadigit{2}}} \\
+ \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\isa{{\isacharparenleft}initial{\isacharunderscore}method{\isacharparenright}} \\
+ \quad\isa{body} \\
+ \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}~\isa{{\isacharparenleft}terminal{\isacharunderscore}method{\isacharparenright}} \\
+ \end{tabular}
+ \medskip
+
+ The goal configuration consists of \isa{facts\isactrlsub {\isadigit{1}}} and \isa{facts\isactrlsub {\isadigit{2}}} appended in that order, and various \isa{props} being
+ claimed here. The \isa{initial{\isacharunderscore}method} is invoked with facts
+ and goals together and refines the problem to something that is
+ handled recursively in the proof \isa{body}. The \isa{terminal{\isacharunderscore}method} has another chance to finish-off any remaining
+ subgoals, but it does not see the facts of the initial step.
+
+ \medskip The second pattern illustrates unstructured proof scripts:
+
+ \medskip
+ \begin{tabular}{l}
+ \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{props} \\
+ \quad\hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}~\isa{facts\isactrlsub {\isadigit{1}}}~\hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}}~\isa{method\isactrlsub {\isadigit{1}}} \\
+ \quad\hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}}~\isa{method\isactrlsub {\isadigit{2}}} \\
+ \quad\hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}~\isa{facts\isactrlsub {\isadigit{3}}}~\hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}}~\isa{method\isactrlsub {\isadigit{3}}} \\
+ \quad\hyperlink{command.done}{\mbox{\isa{\isacommand{done}}}} \\
+ \end{tabular}
+ \medskip
+
+ The \isa{method\isactrlsub {\isadigit{1}}} operates on the original claim together while
+ using \isa{facts\isactrlsub {\isadigit{1}}}. Since the \hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}} command
+ structurally resets the facts, the \isa{method\isactrlsub {\isadigit{2}}} will operate on
+ the remaining goal state without facts. The \isa{method\isactrlsub {\isadigit{3}}} will
+ see again a collection of \isa{facts\isactrlsub {\isadigit{3}}} that has been inserted
+ into the script explicitly.
+
+ \medskip Empirically, Isar proof methods 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: \hyperlink{method.induct}{\mbox{\isa{induct}}}, which is also a ``raw'' method since it
+ operates on the internal representation of simultaneous claims as
+ Pure conjunction, instead of separate subgoals.
+
+ \item \emph{Structured method} with strong emphasis on facts outside
+ the goal state.
+
+ Example: \hyperlink{method.rule}{\mbox{\isa{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: \hyperlink{method.simp}{\mbox{\isa{simp}}}, \hyperlink{method.blast}{\mbox{\isa{blast}}}, \hyperlink{method.auto}{\mbox{\isa{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).
+ To make the special non-standard status clear, the naming convention
+ \isa{foo{\isacharunderscore}tac} needs to be observed.
+
+ Example: \hyperlink{method.rule-tac}{\mbox{\isa{rule{\isacharunderscore}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.}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML type}{Proof.method}\verb|type Proof.method| \\
+ \indexdef{}{ML}{METHOD\_CASES}\verb|METHOD_CASES: (thm list -> cases_tactic) -> Proof.method| \\
+ \indexdef{}{ML}{METHOD}\verb|METHOD: (thm list -> tactic) -> Proof.method| \\
+ \indexdef{}{ML}{SIMPLE\_METHOD}\verb|SIMPLE_METHOD: tactic -> Proof.method| \\
+ \indexdef{}{ML}{SIMPLE\_METHOD'}\verb|SIMPLE_METHOD': (int -> tactic) -> Proof.method| \\
+ \indexdef{}{ML}{HEADGOAL}\verb|HEADGOAL: (int -> tactic) -> tactic| \\
+ \indexdef{}{ML}{Method.insert\_tac}\verb|Method.insert_tac: thm list -> int -> tactic| \\
+ \indexdef{}{ML}{Method.setup}\verb|Method.setup: binding -> (Proof.context -> Proof.method) context_parser ->|\isasep\isanewline%
+\verb| string -> theory -> theory| \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type \verb|Proof.method| represents proof methods as
+ abstract type.
+
+ \item \verb|METHOD_CASES|~\isa{{\isacharparenleft}fn\ facts\ {\isacharequal}{\isachargreater}\ cases{\isacharunderscore}tactic{\isacharparenright}} wraps
+ \isa{cases{\isacharunderscore}tactic} depending on goal facts as proof method with
+ cases; the goal context is passed via method syntax.
+
+ \item \verb|METHOD|~\isa{{\isacharparenleft}fn\ facts\ {\isacharequal}{\isachargreater}\ tactic{\isacharparenright}} wraps \isa{tactic} depending on goal facts as regular proof method; the goal
+ context is passed via method syntax.
+
+ \item \verb|SIMPLE_METHOD|~\isa{tactic} wraps a tactic that
+ addresses all subgoals uniformly as simple proof method. Goal facts
+ are already inserted into all subgoals before \isa{tactic} is
+ applied.
+
+ \item \verb|SIMPLE_METHOD'|~\isa{tactic} wraps a tactic that
+ addresses a specific subgoal as simple proof method. Goal facts are
+ already inserted into the first subgoal before \isa{tactic} is
+ applied to the same.
+
+ \item \verb|HEADGOAL|~\isa{tactic} applies \isa{tactic} to
+ the first subgoal. This is convenient to reproduce part the \verb|SIMPLE_METHOD'| wrapping within regular \verb|METHOD|, for example.
+
+ \item \verb|Method.insert_tac|~\isa{facts\ i} inserts \isa{facts} into subgoal \isa{i}. This is convenient to reproduce
+ part of the \verb|SIMPLE_METHOD| or \verb|SIMPLE_METHOD'| wrapping
+ within regular \verb|METHOD|, for example.
+
+ \item \verb|Method.setup|~\isa{name\ parser\ description} provides
+ the functionality of the Isar command \hyperlink{command.method-setup}{\mbox{\isa{\isacommand{method{\isacharunderscore}setup}}}} as ML
+ function.
+
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isatagmlex
+%
+\begin{isamarkuptext}%
+See also \hyperlink{command.method-setup}{\mbox{\isa{\isacommand{method{\isacharunderscore}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 ``\hyperlink{method.tactic}{\mbox{\isa{tactic}}}'' wraps ML source of type \verb|tactic| (abstracted over \verb|facts|). This allows immediate
+ experimentation without parsing of concrete syntax.%
\end{isamarkuptext}%
\isamarkuptrue%
%
-\isamarkupsection{Attributes%
+\endisatagmlex
+{\isafoldmlex}%
+%
+\isadelimmlex
+%
+\endisadelimmlex
+\isacommand{example{\isacharunderscore}proof}\isamarkupfalse%
+\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{assume}\isamarkupfalse%
+\ a{\isacharcolon}\ A\ \isakeyword{and}\ b{\isacharcolon}\ B\isanewline
+\isanewline
+\ \ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}A\ {\isasymand}\ B{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}tactic\ {\isacharverbatimopen}\ rtac\ %
+\isaantiq
+thm\ conjI%
+\endisaantiq
+\ {\isadigit{1}}\ {\isacharverbatimclose}{\isacharparenright}\isanewline
+\ \ \ \ \isacommand{using}\isamarkupfalse%
+\ a\ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}tactic\ {\isacharverbatimopen}\ resolve{\isacharunderscore}tac\ facts\ {\isadigit{1}}\ {\isacharverbatimclose}{\isacharparenright}\isanewline
+\ \ \ \ \isacommand{using}\isamarkupfalse%
+\ b\ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}tactic\ {\isacharverbatimopen}\ resolve{\isacharunderscore}tac\ facts\ {\isadigit{1}}\ {\isacharverbatimclose}{\isacharparenright}\isanewline
+\ \ \ \ \isacommand{done}\isamarkupfalse%
+\isanewline
+\isanewline
+\ \ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}A\ {\isasymand}\ B{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{using}\isamarkupfalse%
+\ a\ \isakeyword{and}\ b%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+%
+\isadelimML
+\ \ \ \ %
+\endisadelimML
+%
+\isatagML
+\isacommand{ML{\isacharunderscore}val}\isamarkupfalse%
+\ {\isachardoublequoteopen}{\isacharat}{\isacharbraceleft}Isar{\isachardot}goal{\isacharbraceright}{\isachardoublequoteclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+\isanewline
+%
+\endisadelimML
+%
+\isadelimproof
+\ \ \ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}tactic\ {\isacharverbatimopen}\ Method{\isachardot}insert{\isacharunderscore}tac\ facts\ {\isadigit{1}}\ {\isacharverbatimclose}{\isacharparenright}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}tactic\ {\isacharverbatimopen}\ {\isacharparenleft}rtac\ %
+\isaantiq
+thm\ conjI%
+\endisaantiq
+\ THEN{\isacharunderscore}ALL{\isacharunderscore}NEW\ atac{\isacharparenright}\ {\isadigit{1}}\ {\isacharverbatimclose}{\isacharparenright}\isanewline
+\ \ \ \ \isacommand{done}\isamarkupfalse%
+\isanewline
+\isacommand{qed}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\begin{isamarkuptext}%
+\medskip The next example implements a method that simplifies
+ the first subgoal by rewrite rules given as arguments.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{method{\isacharunderscore}setup}\isamarkupfalse%
+\ my{\isacharunderscore}simp\ {\isacharequal}\ {\isacharverbatimopen}\isanewline
+\ \ Attrib{\isachardot}thms\ {\isachargreater}{\isachargreater}\ {\isacharparenleft}fn\ thms\ {\isacharequal}{\isachargreater}\ fn\ ctxt\ {\isacharequal}{\isachargreater}\isanewline
+\ \ \ \ SIMPLE{\isacharunderscore}METHOD{\isacharprime}\ {\isacharparenleft}fn\ i\ {\isacharequal}{\isachargreater}\isanewline
+\ \ \ \ \ \ CHANGED\ {\isacharparenleft}asm{\isacharunderscore}full{\isacharunderscore}simp{\isacharunderscore}tac\isanewline
+\ \ \ \ \ \ \ \ {\isacharparenleft}HOL{\isacharunderscore}basic{\isacharunderscore}ss\ addsimps\ thms{\isacharparenright}\ i{\isacharparenright}{\isacharparenright}{\isacharparenright}\isanewline
+{\isacharverbatimclose}\ {\isachardoublequoteopen}rewrite\ subgoal\ by\ given\ rules{\isachardoublequoteclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+The concrete syntax wrapping of \hyperlink{command.method-setup}{\mbox{\isa{\isacommand{method{\isacharunderscore}setup}}}} always
+ passes-through the proof context at the end of parsing, but it is
+ not used in this example.
+
+ The \verb|Attrib.thms| parser produces a list of theorems from the
+ usual Isar syntax involving attribute expressions etc.\ (syntax
+ category \hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}). The resulting \verb|thms| are
+ added to \verb|HOL_basic_ss| which already contains the basic
+ Simplifier setup for HOL.
+
+ The tactic \verb|asm_full_simp_tac| is the one that is also used in
+ method \hyperlink{method.simp}{\mbox{\isa{simp}}} by default. The extra wrapping by the \verb|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 \hyperlink{method.my-simp}{\mbox{\isa{my{\isacharunderscore}simp}}} can be used in Isar proofs like
+ this:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+\isacommand{example{\isacharunderscore}proof}\isamarkupfalse%
+\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{fix}\isamarkupfalse%
+\ a\ b\ c\isanewline
+\ \ \isacommand{assume}\isamarkupfalse%
+\ a{\isacharcolon}\ {\isachardoublequoteopen}a\ {\isacharequal}\ b{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{assume}\isamarkupfalse%
+\ b{\isacharcolon}\ {\isachardoublequoteopen}b\ {\isacharequal}\ c{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}a\ {\isacharequal}\ c{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}my{\isacharunderscore}simp\ a\ b{\isacharparenright}\isanewline
+\isacommand{qed}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\begin{isamarkuptext}%
+Here is a similar method that operates on all subgoals,
+ instead of just the first one.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{method{\isacharunderscore}setup}\isamarkupfalse%
+\ my{\isacharunderscore}simp{\isacharunderscore}all\ {\isacharequal}\ {\isacharverbatimopen}\isanewline
+\ \ Attrib{\isachardot}thms\ {\isachargreater}{\isachargreater}\ {\isacharparenleft}fn\ thms\ {\isacharequal}{\isachargreater}\ fn\ ctxt\ {\isacharequal}{\isachargreater}\isanewline
+\ \ \ \ SIMPLE{\isacharunderscore}METHOD\isanewline
+\ \ \ \ \ \ {\isacharparenleft}CHANGED\isanewline
+\ \ \ \ \ \ \ \ {\isacharparenleft}ALLGOALS\ {\isacharparenleft}asm{\isacharunderscore}full{\isacharunderscore}simp{\isacharunderscore}tac\isanewline
+\ \ \ \ \ \ \ \ \ \ {\isacharparenleft}HOL{\isacharunderscore}basic{\isacharunderscore}ss\ addsimps\ thms{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharparenright}\isanewline
+{\isacharverbatimclose}\ {\isachardoublequoteopen}rewrite\ all\ subgoals\ by\ given\ rules{\isachardoublequoteclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+\isanewline
+\isanewline
+\isacommand{example{\isacharunderscore}proof}\isamarkupfalse%
+\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{fix}\isamarkupfalse%
+\ a\ b\ c\isanewline
+\ \ \isacommand{assume}\isamarkupfalse%
+\ a{\isacharcolon}\ {\isachardoublequoteopen}a\ {\isacharequal}\ b{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{assume}\isamarkupfalse%
+\ b{\isacharcolon}\ {\isachardoublequoteopen}b\ {\isacharequal}\ c{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}a\ {\isacharequal}\ c{\isachardoublequoteclose}\ \isakeyword{and}\ {\isachardoublequoteopen}c\ {\isacharequal}\ b{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}my{\isacharunderscore}simp{\isacharunderscore}all\ a\ b{\isacharparenright}\isanewline
+\isanewline
+\isacommand{qed}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\begin{isamarkuptext}%
+\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
+ \verb|Named_Thms| (see also \hyperlink{file.~~/src/Pure/Tools/named-thms.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}Tools{\isacharslash}named{\isacharunderscore}thms{\isachardot}ML}}}}).%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ structure\ My{\isacharunderscore}Simps\ {\isacharequal}\isanewline
+\ \ \ \ Named{\isacharunderscore}Thms\isanewline
+\ \ \ \ \ \ {\isacharparenleft}val\ name\ {\isacharequal}\ {\isachardoublequote}my{\isacharunderscore}simp{\isachardoublequote}\ val\ description\ {\isacharequal}\ {\isachardoublequote}my{\isacharunderscore}simp\ rule{\isachardoublequote}{\isacharparenright}\isanewline
+{\isacharverbatimclose}\isanewline
+\isacommand{setup}\isamarkupfalse%
+\ My{\isacharunderscore}Simps{\isachardot}setup%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+This provides ML access to a list of theorems in canonical
+ declaration order via \verb|My_Simps.get|. The user can add or
+ delete rules via the attribute \hyperlink{attribute.my-simp}{\mbox{\isa{my{\isacharunderscore}simp}}}. The actual
+ proof method is now defined as before, but we append the explicit
+ arguments and the rules from the context.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{method{\isacharunderscore}setup}\isamarkupfalse%
+\ my{\isacharunderscore}simp{\isacharprime}\ {\isacharequal}\ {\isacharverbatimopen}\isanewline
+\ \ Attrib{\isachardot}thms\ {\isachargreater}{\isachargreater}\ {\isacharparenleft}fn\ thms\ {\isacharequal}{\isachargreater}\ fn\ ctxt\ {\isacharequal}{\isachargreater}\isanewline
+\ \ \ \ SIMPLE{\isacharunderscore}METHOD{\isacharprime}\ {\isacharparenleft}fn\ i\ {\isacharequal}{\isachargreater}\isanewline
+\ \ \ \ \ \ CHANGED\ {\isacharparenleft}asm{\isacharunderscore}full{\isacharunderscore}simp{\isacharunderscore}tac\isanewline
+\ \ \ \ \ \ \ \ {\isacharparenleft}HOL{\isacharunderscore}basic{\isacharunderscore}ss\ addsimps\ {\isacharparenleft}thms\ {\isacharat}\ My{\isacharunderscore}Simps{\isachardot}get\ ctxt{\isacharparenright}{\isacharparenright}\ i{\isacharparenright}{\isacharparenright}{\isacharparenright}\isanewline
+{\isacharverbatimclose}\ {\isachardoublequoteopen}rewrite\ subgoal\ by\ given\ rules\ and\ my{\isacharunderscore}simp\ rules\ from\ the\ context{\isachardoublequoteclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+\medskip Method \hyperlink{method.my-simp'}{\mbox{\isa{my{\isacharunderscore}simp{\isacharprime}}}} can be used in Isar proofs
+ like this:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+\isacommand{example{\isacharunderscore}proof}\isamarkupfalse%
+\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{fix}\isamarkupfalse%
+\ a\ b\ c\isanewline
+\ \ \isacommand{assume}\isamarkupfalse%
+\ {\isacharbrackleft}my{\isacharunderscore}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}a\ {\isasymequiv}\ b{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{assume}\isamarkupfalse%
+\ {\isacharbrackleft}my{\isacharunderscore}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}b\ {\isasymequiv}\ c{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}a\ {\isasymequiv}\ c{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
+\ my{\isacharunderscore}simp{\isacharprime}\isanewline
+\isacommand{qed}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\begin{isamarkuptext}%
+\medskip The \hyperlink{method.my-simp}{\mbox{\isa{my{\isacharunderscore}simp}}} variants defined above are
+ ``simple'' methods, i.e.\ the goal facts are merely inserted as goal
+ premises by the \verb|SIMPLE_METHOD'| or \verb|SIMPLE_METHOD| wrapper.
+ For proof methods that are similar to the standard collection of
+ \hyperlink{method.simp}{\mbox{\isa{simp}}}, \hyperlink{method.blast}{\mbox{\isa{blast}}}, \hyperlink{method.auto}{\mbox{\isa{auto}}} little more can be
+ done here.
+
+ 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 ``\hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}~\isa{foo}~\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{my{\isacharunderscore}simp}'' where \isa{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 \verb|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 \verb|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 \verb|simpset| is adequate, but
+ scalability requires it be maintained statically within the context
+ data, not dynamically on each tool invocation.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsection{Attributes \label{sec:attributes}%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
-FIXME%
+An \emph{attribute} is a function \isa{context\ {\isasymtimes}\ thm\ {\isasymrightarrow}\ context\ {\isasymtimes}\ thm}, which means both a (generic) context and a theorem
+ can be modified simultaneously. In practice this fully general form
+ is very rare, instead attributes are presented either as
+ \emph{declaration attribute:} \isa{thm\ {\isasymrightarrow}\ context\ {\isasymrightarrow}\ context} or
+ \emph{rule attribute:} \isa{context\ {\isasymrightarrow}\ thm\ {\isasymrightarrow}\ 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!%
\end{isamarkuptext}%
\isamarkuptrue%
%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML type}{attribute}\verb|type attribute = Context.generic * thm -> Context.generic * thm| \\
+ \indexdef{}{ML}{Thm.rule\_attribute}\verb|Thm.rule_attribute: (Context.generic -> thm -> thm) -> attribute| \\
+ \indexdef{}{ML}{Thm.declaration\_attribute}\verb|Thm.declaration_attribute: |\isasep\isanewline%
+\verb| (thm -> Context.generic -> Context.generic) -> attribute| \\
+ \indexdef{}{ML}{Attrib.setup}\verb|Attrib.setup: binding -> attribute context_parser ->|\isasep\isanewline%
+\verb| string -> theory -> theory| \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type \verb|attribute| represents attributes as concrete
+ type alias.
+
+ \item \verb|Thm.rule_attribute|~\isa{{\isacharparenleft}fn\ context\ {\isacharequal}{\isachargreater}\ rule{\isacharparenright}} wraps
+ a context-dependent rule (mapping on \verb|thm|) as attribute.
+
+ \item \verb|Thm.declaration_attribute|~\isa{{\isacharparenleft}fn\ thm\ {\isacharequal}{\isachargreater}\ decl{\isacharparenright}}
+ wraps a theorem-dependent declaration (mapping on \verb|Context.generic|) as attribute.
+
+ \item \verb|Attrib.setup|~\isa{name\ parser\ description} provides
+ the functionality of the Isar command \hyperlink{command.attribute-setup}{\mbox{\isa{\isacommand{attribute{\isacharunderscore}setup}}}} as
+ ML function.
+
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isatagmlex
+%
+\begin{isamarkuptext}%
+See also \hyperlink{command.attribute-setup}{\mbox{\isa{\isacommand{attribute{\isacharunderscore}setup}}}} in
+ \cite{isabelle-isar-ref} which includes some abstract examples.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlex
+{\isafoldmlex}%
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
\isadelimtheory
%
\endisadelimtheory
--- a/doc-src/IsarImplementation/Thy/document/Local_Theory.tex Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/Local_Theory.tex Sun Oct 24 15:11:24 2010 -0700
@@ -49,8 +49,8 @@
Many definitional packages for local theories are available in
Isabelle. Although a few old packages only work for global
- theories, the local theory interface is already the standard way of
- implementing definitional packages in Isabelle.%
+ theories, the standard way of implementing definitional packages in
+ Isabelle is via the local theory interface.%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -105,12 +105,11 @@
\framebox{\quad auxiliary context \quad\framebox{\quad target context \quad\framebox{\quad background theory\quad}}}
\end{center}
- \noindent 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
- \isa{{\isasymDEFINE}} and \isa{{\isasymNOTE}} elements, transformed by
- the particular target policy (see
- \cite[\S4--5]{Haftmann-Wenzel:2009} for details).%
+ 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 \isa{{\isasymDEFINE}} and \isa{{\isasymNOTE}} elements, transformed by the
+ particular target policy (see \cite[\S4--5]{Haftmann-Wenzel:2009}
+ for details).%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -132,8 +131,8 @@
\begin{description}
- \item \verb|local_theory| represents local theories. Although
- this is merely an alias for \verb|Proof.context|, it is
+ \item Type \verb|local_theory| represents local theories.
+ Although this is merely an alias for \verb|Proof.context|, it is
semantically a subtype of the same: a \verb|local_theory| holds
target information as special context data. Subtyping means that
any value \isa{lthy{\isacharcolon}}~\verb|local_theory| can be also used
@@ -185,12 +184,14 @@
%
\endisadelimmlref
%
-\isamarkupsection{Morphisms and declarations%
+\isamarkupsection{Morphisms and declarations \label{sec:morphisms}%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
-FIXME%
+FIXME
+
+ \medskip See also \cite{Chaieb-Wenzel:2007}.%
\end{isamarkuptext}%
\isamarkuptrue%
%
--- a/doc-src/IsarImplementation/Thy/document/Logic.tex Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/Logic.tex Sun Oct 24 15:11:24 2010 -0700
@@ -132,8 +132,8 @@
\indexdef{}{ML type}{sort}\verb|type sort = class list| \\
\indexdef{}{ML type}{arity}\verb|type arity = string * sort list * sort| \\
\indexdef{}{ML type}{typ}\verb|type typ| \\
- \indexdef{}{ML}{map\_atyps}\verb|map_atyps: (typ -> typ) -> typ -> typ| \\
- \indexdef{}{ML}{fold\_atyps}\verb|fold_atyps: (typ -> 'a -> 'a) -> typ -> 'a -> 'a| \\
+ \indexdef{}{ML}{Term.map\_atyps}\verb|Term.map_atyps: (typ -> typ) -> typ -> typ| \\
+ \indexdef{}{ML}{Term.fold\_atyps}\verb|Term.fold_atyps: (typ -> 'a -> 'a) -> typ -> 'a -> 'a| \\
\end{mldecls}
\begin{mldecls}
\indexdef{}{ML}{Sign.subsort}\verb|Sign.subsort: theory -> sort * sort -> bool| \\
@@ -147,23 +147,24 @@
\begin{description}
- \item \verb|class| represents type classes.
+ \item Type \verb|class| represents type classes.
- \item \verb|sort| represents sorts, i.e.\ finite intersections
- of classes. The empty list \verb|[]: sort| refers to the empty
- class intersection, i.e.\ the ``full sort''.
+ \item Type \verb|sort| represents sorts, i.e.\ finite
+ intersections of classes. The empty list \verb|[]: sort| refers to
+ the empty class intersection, i.e.\ the ``full sort''.
- \item \verb|arity| represents type arities. A triple \isa{{\isacharparenleft}{\isasymkappa}{\isacharcomma}\ \isactrlvec s{\isacharcomma}\ s{\isacharparenright}\ {\isacharcolon}\ arity} represents \isa{{\isasymkappa}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}\isactrlvec s{\isacharparenright}s} as
- described above.
+ \item Type \verb|arity| represents type arities. A triple
+ \isa{{\isacharparenleft}{\isasymkappa}{\isacharcomma}\ \isactrlvec s{\isacharcomma}\ s{\isacharparenright}\ {\isacharcolon}\ arity} represents \isa{{\isasymkappa}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}\isactrlvec s{\isacharparenright}s} as described above.
- \item \verb|typ| represents types; this is a datatype with
+ \item Type \verb|typ| represents types; this is a datatype with
constructors \verb|TFree|, \verb|TVar|, \verb|Type|.
- \item \verb|map_atyps|~\isa{f\ {\isasymtau}} applies the mapping \isa{f}
- to all atomic types (\verb|TFree|, \verb|TVar|) occurring in \isa{{\isasymtau}}.
+ \item \verb|Term.map_atyps|~\isa{f\ {\isasymtau}} applies the mapping \isa{f} to all atomic types (\verb|TFree|, \verb|TVar|) occurring in
+ \isa{{\isasymtau}}.
- \item \verb|fold_atyps|~\isa{f\ {\isasymtau}} iterates the operation \isa{f} over all occurrences of atomic types (\verb|TFree|, \verb|TVar|)
- in \isa{{\isasymtau}}; the type structure is traversed from left to right.
+ \item \verb|Term.fold_atyps|~\isa{f\ {\isasymtau}} iterates the operation
+ \isa{f} over all occurrences of atomic types (\verb|TFree|, \verb|TVar|) in \isa{{\isasymtau}}; the type structure is traversed from left to
+ right.
\item \verb|Sign.subsort|~\isa{thy\ {\isacharparenleft}s\isactrlisub {\isadigit{1}}{\isacharcomma}\ s\isactrlisub {\isadigit{2}}{\isacharparenright}}
tests the subsort relation \isa{s\isactrlisub {\isadigit{1}}\ {\isasymsubseteq}\ s\isactrlisub {\isadigit{2}}}.
@@ -196,6 +197,64 @@
%
\endisadelimmlref
%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
+\isatagmlantiq
+%
+\begin{isamarkuptext}%
+\begin{matharray}{rcl}
+ \indexdef{}{ML antiquotation}{class}\hypertarget{ML antiquotation.class}{\hyperlink{ML antiquotation.class}{\mbox{\isa{class}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{sort}\hypertarget{ML antiquotation.sort}{\hyperlink{ML antiquotation.sort}{\mbox{\isa{sort}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{type\_name}\hypertarget{ML antiquotation.type-name}{\hyperlink{ML antiquotation.type-name}{\mbox{\isa{type{\isacharunderscore}name}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{type\_abbrev}\hypertarget{ML antiquotation.type-abbrev}{\hyperlink{ML antiquotation.type-abbrev}{\mbox{\isa{type{\isacharunderscore}abbrev}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{nonterminal}\hypertarget{ML antiquotation.nonterminal}{\hyperlink{ML antiquotation.nonterminal}{\mbox{\isa{nonterminal}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{typ}\hypertarget{ML antiquotation.typ}{\hyperlink{ML antiquotation.typ}{\mbox{\isa{typ}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \end{matharray}
+
+ \begin{rail}
+ 'class' nameref
+ ;
+ 'sort' sort
+ ;
+ ('type\_name' | 'type\_abbrev' | 'nonterminal') nameref
+ ;
+ 'typ' type
+ ;
+ \end{rail}
+
+ \begin{description}
+
+ \item \isa{{\isacharat}{\isacharbraceleft}class\ c{\isacharbraceright}} inlines the internalized class \isa{c} --- as \verb|string| literal.
+
+ \item \isa{{\isacharat}{\isacharbraceleft}sort\ s{\isacharbraceright}} inlines the internalized sort \isa{s}
+ --- as \verb|string list| literal.
+
+ \item \isa{{\isacharat}{\isacharbraceleft}type{\isacharunderscore}name\ c{\isacharbraceright}} inlines the internalized type
+ constructor \isa{c} --- as \verb|string| literal.
+
+ \item \isa{{\isacharat}{\isacharbraceleft}type{\isacharunderscore}abbrev\ c{\isacharbraceright}} inlines the internalized type
+ abbreviation \isa{c} --- as \verb|string| literal.
+
+ \item \isa{{\isacharat}{\isacharbraceleft}nonterminal\ c{\isacharbraceright}} inlines the internalized syntactic
+ type~/ grammar nonterminal \isa{c} --- as \verb|string|
+ literal.
+
+ \item \isa{{\isacharat}{\isacharbraceleft}typ\ {\isasymtau}{\isacharbraceright}} inlines the internalized type \isa{{\isasymtau}}
+ --- as constructor term for datatype \verb|typ|.
+
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlantiq
+{\isafoldmlantiq}%
+%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
\isamarkupsection{Terms \label{sec:terms}%
}
\isamarkuptrue%
@@ -248,7 +307,7 @@
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.
+ \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: \isa{t\ {\isacharequal}\ b\ {\isacharbar}\ x\isactrlisub {\isasymtau}\ {\isacharbar}\ {\isacharquery}x\isactrlisub {\isasymtau}\ {\isacharbar}\ c\isactrlisub {\isasymtau}\ {\isacharbar}\ {\isasymlambda}\isactrlisub {\isasymtau}{\isachardot}\ t\ {\isacharbar}\ t\isactrlisub {\isadigit{1}}\ t\isactrlisub {\isadigit{2}}}. Parsing and printing takes care of
converting between an external representation with named bound
@@ -317,11 +376,11 @@
\begin{isamarkuptext}%
\begin{mldecls}
\indexdef{}{ML type}{term}\verb|type term| \\
- \indexdef{}{ML}{op aconv}\verb|op aconv: term * term -> bool| \\
- \indexdef{}{ML}{map\_types}\verb|map_types: (typ -> typ) -> term -> term| \\
- \indexdef{}{ML}{fold\_types}\verb|fold_types: (typ -> 'a -> 'a) -> term -> 'a -> 'a| \\
- \indexdef{}{ML}{map\_aterms}\verb|map_aterms: (term -> term) -> term -> term| \\
- \indexdef{}{ML}{fold\_aterms}\verb|fold_aterms: (term -> 'a -> 'a) -> term -> 'a -> 'a| \\
+ \indexdef{}{ML}{aconv}\verb|op aconv: term * term -> bool| \\
+ \indexdef{}{ML}{Term.map\_types}\verb|Term.map_types: (typ -> typ) -> term -> term| \\
+ \indexdef{}{ML}{Term.fold\_types}\verb|Term.fold_types: (typ -> 'a -> 'a) -> term -> 'a -> 'a| \\
+ \indexdef{}{ML}{Term.map\_aterms}\verb|Term.map_aterms: (term -> term) -> term -> term| \\
+ \indexdef{}{ML}{Term.fold\_aterms}\verb|Term.fold_aterms: (term -> 'a -> 'a) -> term -> 'a -> 'a| \\
\end{mldecls}
\begin{mldecls}
\indexdef{}{ML}{fastype\_of}\verb|fastype_of: term -> typ| \\
@@ -337,24 +396,24 @@
\begin{description}
- \item \verb|term| represents de-Bruijn terms, with comments in
- abstractions, and explicitly named free variables and constants;
+ \item Type \verb|term| represents de-Bruijn terms, with comments
+ in abstractions, and explicitly named free variables and constants;
this is a datatype with constructors \verb|Bound|, \verb|Free|, \verb|Var|, \verb|Const|, \verb|Abs|, \verb|op $|.
\item \isa{t}~\verb|aconv|~\isa{u} checks \isa{{\isasymalpha}}-equivalence of two terms. This is the basic equality relation
on type \verb|term|; raw datatype equality should only be used
for operations related to parsing or printing!
- \item \verb|map_types|~\isa{f\ t} applies the mapping \isa{f} to all types occurring in \isa{t}.
+ \item \verb|Term.map_types|~\isa{f\ t} applies the mapping \isa{f} to all types occurring in \isa{t}.
- \item \verb|fold_types|~\isa{f\ t} iterates the operation \isa{f} over all occurrences of types in \isa{t}; the term
+ \item \verb|Term.fold_types|~\isa{f\ t} iterates the operation
+ \isa{f} over all occurrences of types in \isa{t}; the term
structure is traversed from left to right.
- \item \verb|map_aterms|~\isa{f\ t} applies the mapping \isa{f}
- to all atomic terms (\verb|Bound|, \verb|Free|, \verb|Var|, \verb|Const|) occurring in \isa{t}.
+ \item \verb|Term.map_aterms|~\isa{f\ t} applies the mapping \isa{f} to all atomic terms (\verb|Bound|, \verb|Free|, \verb|Var|, \verb|Const|) occurring in \isa{t}.
- \item \verb|fold_aterms|~\isa{f\ t} iterates the operation \isa{f} over all occurrences of atomic terms (\verb|Bound|, \verb|Free|,
- \verb|Var|, \verb|Const|) in \isa{t}; the term structure is
+ \item \verb|Term.fold_aterms|~\isa{f\ t} iterates the operation
+ \isa{f} over all occurrences of atomic terms (\verb|Bound|, \verb|Free|, \verb|Var|, \verb|Const|) in \isa{t}; the term structure is
traversed from left to right.
\item \verb|fastype_of|~\isa{t} determines the type of a
@@ -389,6 +448,63 @@
%
\endisadelimmlref
%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
+\isatagmlantiq
+%
+\begin{isamarkuptext}%
+\begin{matharray}{rcl}
+ \indexdef{}{ML antiquotation}{const\_name}\hypertarget{ML antiquotation.const-name}{\hyperlink{ML antiquotation.const-name}{\mbox{\isa{const{\isacharunderscore}name}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{const\_abbrev}\hypertarget{ML antiquotation.const-abbrev}{\hyperlink{ML antiquotation.const-abbrev}{\mbox{\isa{const{\isacharunderscore}abbrev}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{const}\hypertarget{ML antiquotation.const}{\hyperlink{ML antiquotation.const}{\mbox{\isa{const}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{term}\hypertarget{ML antiquotation.term}{\hyperlink{ML antiquotation.term}{\mbox{\isa{term}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{prop}\hypertarget{ML antiquotation.prop}{\hyperlink{ML antiquotation.prop}{\mbox{\isa{prop}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \end{matharray}
+
+ \begin{rail}
+ ('const\_name' | 'const\_abbrev') nameref
+ ;
+ 'const' ('(' (type + ',') ')')?
+ ;
+ 'term' term
+ ;
+ 'prop' prop
+ ;
+ \end{rail}
+
+ \begin{description}
+
+ \item \isa{{\isacharat}{\isacharbraceleft}const{\isacharunderscore}name\ c{\isacharbraceright}} inlines the internalized logical
+ constant name \isa{c} --- as \verb|string| literal.
+
+ \item \isa{{\isacharat}{\isacharbraceleft}const{\isacharunderscore}abbrev\ c{\isacharbraceright}} inlines the internalized
+ abbreviated constant name \isa{c} --- as \verb|string|
+ literal.
+
+ \item \isa{{\isacharat}{\isacharbraceleft}const\ c{\isacharparenleft}\isactrlvec {\isasymtau}{\isacharparenright}{\isacharbraceright}} inlines the internalized
+ constant \isa{c} with precise type instantiation in the sense of
+ \verb|Sign.const_instance| --- as \verb|Const| constructor term for
+ datatype \verb|term|.
+
+ \item \isa{{\isacharat}{\isacharbraceleft}term\ t{\isacharbraceright}} inlines the internalized term \isa{t}
+ --- as constructor term for datatype \verb|term|.
+
+ \item \isa{{\isacharat}{\isacharbraceleft}prop\ {\isasymphi}{\isacharbraceright}} inlines the internalized proposition
+ \isa{{\isasymphi}} --- as constructor term for datatype \verb|term|.
+
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlantiq
+{\isafoldmlantiq}%
+%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
\isamarkupsection{Theorems \label{sec:thms}%
}
\isamarkuptrue%
@@ -530,7 +646,7 @@
distinct variables \isa{\isactrlvec {\isasymalpha}}). The RHS may mention
previously defined constants as above, or arbitrary constants \isa{d{\isacharparenleft}{\isasymalpha}\isactrlisub i{\isacharparenright}} for some \isa{{\isasymalpha}\isactrlisub i} projected from \isa{\isactrlvec {\isasymalpha}}. Thus overloaded definitions essentially work by
primitive recursion over the syntactic structure of a single type
- argument.%
+ argument. See also \cite[\S4.3]{Haftmann-Wenzel:2006:classes}.%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -558,18 +674,20 @@
\indexdef{}{ML}{Thm.generalize}\verb|Thm.generalize: string list * string list -> int -> thm -> thm| \\
\indexdef{}{ML}{Thm.instantiate}\verb|Thm.instantiate: (ctyp * ctyp) list * (cterm * cterm) list -> thm -> thm| \\
\indexdef{}{ML}{Thm.add\_axiom}\verb|Thm.add_axiom: binding * term -> theory -> (string * thm) * theory| \\
- \indexdef{}{ML}{Thm.add\_oracle}\verb|Thm.add_oracle: binding * ('a -> cterm) -> theory|\isasep\isanewline%
-\verb| -> (string * ('a -> thm)) * theory| \\
- \indexdef{}{ML}{Thm.add\_def}\verb|Thm.add_def: bool -> bool -> binding * term -> theory -> (string * thm) * theory| \\
+ \indexdef{}{ML}{Thm.add\_oracle}\verb|Thm.add_oracle: binding * ('a -> cterm) -> theory ->|\isasep\isanewline%
+\verb| (string * ('a -> thm)) * theory| \\
+ \indexdef{}{ML}{Thm.add\_def}\verb|Thm.add_def: bool -> bool -> binding * term -> theory ->|\isasep\isanewline%
+\verb| (string * thm) * theory| \\
\end{mldecls}
\begin{mldecls}
- \indexdef{}{ML}{Theory.add\_deps}\verb|Theory.add_deps: string -> string * typ -> (string * typ) list -> theory -> theory| \\
+ \indexdef{}{ML}{Theory.add\_deps}\verb|Theory.add_deps: string -> string * typ -> (string * typ) list ->|\isasep\isanewline%
+\verb| theory -> theory| \\
\end{mldecls}
\begin{description}
- \item \verb|ctyp| and \verb|cterm| represent certified types
- and terms, respectively. These are abstract datatypes that
+ \item Types \verb|ctyp| and \verb|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 theory.
@@ -582,8 +700,8 @@
reasoning loops. There are separate operations to decompose
certified entities (including actual theorems).
- \item \verb|thm| represents proven propositions. This is an
- abstract datatype that guarantees that its values have been
+ \item Type \verb|thm| represents proven propositions. This is
+ an abstract datatype that guarantees that its values have been
constructed by basic principles of the \verb|Thm| module.
Every \verb|thm| value contains a sliding back-reference to the
enclosing theory, cf.\ \secref{sec:context-theory}.
@@ -639,6 +757,78 @@
%
\endisadelimmlref
%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
+\isatagmlantiq
+%
+\begin{isamarkuptext}%
+\begin{matharray}{rcl}
+ \indexdef{}{ML antiquotation}{ctyp}\hypertarget{ML antiquotation.ctyp}{\hyperlink{ML antiquotation.ctyp}{\mbox{\isa{ctyp}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{cterm}\hypertarget{ML antiquotation.cterm}{\hyperlink{ML antiquotation.cterm}{\mbox{\isa{cterm}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{cprop}\hypertarget{ML antiquotation.cprop}{\hyperlink{ML antiquotation.cprop}{\mbox{\isa{cprop}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{thm}\hypertarget{ML antiquotation.thm}{\hyperlink{ML antiquotation.thm}{\mbox{\isa{thm}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{thms}\hypertarget{ML antiquotation.thms}{\hyperlink{ML antiquotation.thms}{\mbox{\isa{thms}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{lemma}\hypertarget{ML antiquotation.lemma}{\hyperlink{ML antiquotation.lemma}{\mbox{\isa{lemma}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \end{matharray}
+
+ \begin{rail}
+ 'ctyp' typ
+ ;
+ 'cterm' term
+ ;
+ 'cprop' prop
+ ;
+ 'thm' thmref
+ ;
+ 'thms' thmrefs
+ ;
+ 'lemma' ('(open)')? ((prop +) + 'and') \\ 'by' method method?
+ \end{rail}
+
+ \begin{description}
+
+ \item \isa{{\isacharat}{\isacharbraceleft}ctyp\ {\isasymtau}{\isacharbraceright}} produces a certified type wrt.\ the
+ current background theory --- as abstract value of type \verb|ctyp|.
+
+ \item \isa{{\isacharat}{\isacharbraceleft}cterm\ t{\isacharbraceright}} and \isa{{\isacharat}{\isacharbraceleft}cprop\ {\isasymphi}{\isacharbraceright}} produce a
+ certified term wrt.\ the current background theory --- as abstract
+ value of type \verb|cterm|.
+
+ \item \isa{{\isacharat}{\isacharbraceleft}thm\ a{\isacharbraceright}} produces a singleton fact --- as abstract
+ value of type \verb|thm|.
+
+ \item \isa{{\isacharat}{\isacharbraceleft}thms\ a{\isacharbraceright}} produces a general fact --- as abstract
+ value of type \verb|thm list|.
+
+ \item \isa{{\isacharat}{\isacharbraceleft}lemma\ {\isasymphi}\ by\ meth{\isacharbraceright}} 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 \verb|thm|
+ or \verb|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 \isa{{\isacharparenleft}open{\isacharparenright}} 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 \hyperlink{command.by}{\mbox{\isa{\isacommand{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}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlantiq
+{\isafoldmlantiq}%
+%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
\isamarkupsubsection{Auxiliary definitions%
}
\isamarkuptrue%
@@ -795,9 +985,9 @@
\end{tabular}
\medskip
- \noindent Thus we essentially impose nesting levels on propositions
- formed from \isa{{\isasymAnd}} and \isa{{\isasymLongrightarrow}}. At each level there is a
- prefix of parameters and compound premises, concluding an atomic
+ Thus we essentially impose nesting levels on propositions formed
+ from \isa{{\isasymAnd}} and \isa{{\isasymLongrightarrow}}. At each level there is a prefix
+ of parameters and compound premises, concluding an atomic
proposition. Typical examples are \isa{{\isasymlongrightarrow}}-introduction \isa{{\isacharparenleft}A\ {\isasymLongrightarrow}\ B{\isacharparenright}\ {\isasymLongrightarrow}\ A\ {\isasymlongrightarrow}\ B} or mathematical induction \isa{P\ {\isadigit{0}}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\isasymAnd}n{\isachardot}\ P\ n\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}Suc\ n{\isacharparenright}{\isacharparenright}\ {\isasymLongrightarrow}\ P\ n}. Even deeper nesting occurs in well-founded
induction \isa{{\isacharparenleft}{\isasymAnd}x{\isachardot}\ {\isacharparenleft}{\isasymAnd}y{\isachardot}\ y\ {\isasymprec}\ x\ {\isasymLongrightarrow}\ P\ y{\isacharparenright}\ {\isasymLongrightarrow}\ P\ x{\isacharparenright}\ {\isasymLongrightarrow}\ P\ x}, but this
already marks the limit of rule complexity that is usually seen in
@@ -920,8 +1110,8 @@
%
\begin{isamarkuptext}%
\begin{mldecls}
- \indexdef{}{ML}{op RS}\verb|op RS: thm * thm -> thm| \\
- \indexdef{}{ML}{op OF}\verb|op OF: thm * thm list -> thm| \\
+ \indexdef{}{ML}{RS}\verb|op RS: thm * thm -> thm| \\
+ \indexdef{}{ML}{OF}\verb|op OF: thm * thm list -> thm| \\
\end{mldecls}
\begin{description}
--- a/doc-src/IsarImplementation/Thy/document/ML.tex Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/ML.tex Sun Oct 24 15:11:24 2010 -0700
@@ -18,252 +18,2165 @@
%
\endisadelimtheory
%
-\isamarkupchapter{Advanced ML programming%
+\isamarkupchapter{Isabelle/ML%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+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 are specific library
+ modules and infrastructure to address the needs for such difficult
+ tasks. 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.}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsection{Style and orthography%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+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.
+ One should always strife at least for local consistency of modules
+ and sub-systems, without deviating from some general principles how
+ to write Isabelle/ML.
+
+ In a sense, a common 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 source.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsubsection{Header and sectioning%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+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. E.g.\ see \hyperlink{file.~~/src/Pure/thm.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}thm{\isachardot}ML}}}}.
+
+ The header includes at least \verb|Title| and \verb|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 some a layout via ML comments
+ as follows.
+
+\begin{verbatim}
+(*** section ***)
+
+(** subsection **)
+
+(* subsubsection *)
+
+(*short paragraph*)
+
+(*
+ long paragraph
+ 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 indicate the main theme of the
+ subsequent formal ML text.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsubsection{Naming conventions%
}
\isamarkuptrue%
%
-\isamarkupsection{Style%
+\begin{isamarkuptext}%
+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 just for
+ certain ML categories as follows:
+
+ \medskip
+ \begin{tabular}{lll}
+ variant & example & ML categories \\\hline
+ lower-case & \verb|foo_bar| & values, types, record fields \\
+ capitalized & \verb|Foo_Bar| & datatype constructors, \\
+ & & structures, functors \\
+ upper-case & \verb|FOO_BAR| & special values (combinators), \\
+ & & exception constructors, signatures \\
+ \end{tabular}
+ \medskip
+
+ For historical reasons, many capitalized names omit underscores,
+ e.g.\ old-style \verb|FooBar| instead of \verb|Foo_Bar|.
+ Genuine mixed-case names are \emph{not} used --- 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 and keywords. Later it became a 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 are suffixed by extra markers like this:
+ \verb|foo_barT|.
+
+ Name variants are produced by adding 1--3 primes, e.g.\ \verb|foo'|, \verb|foo''|, or \verb|foo'''|, but not \verb|foo''''| or more. Decimal digits scale better to larger numbers,
+ e.g.\ \verb|foo0|, \verb|foo1|, \verb|foo42|.
+
+ \paragraph{Scopes.} Apart from very basic library modules, ML
+ structures are not ``opened'', but names are referenced with
+ explicit qualification, as in \verb|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 always 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/lets bindings are
+ typically shorter, sometimes using only rudiments of ``words'',
+ while still avoiding cryptic shorthands. An auxiliary function
+ called \verb|helper|, \verb|aux|, or \verb|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 important specific
+ name forms that occur frequently in the sources.
+
+ \begin{itemize}
+
+ \item A function that maps \verb|foo| to \verb|bar| is
+ called \verb|foo_to_bar| or \verb|bar_of_foo| (never
+ \verb|foo2bar|, \verb|bar_from_foo|, \verb|bar_for_foo|, or \verb|bar4foo|).
+
+ \item The name component \verb|legacy| means that the operation
+ is about to be discontinued soon.
+
+ \item The name component \verb|old| means that this is historic
+ material that might disappear at some later stage.
+
+ \item The name component \verb|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 to allow to visualize the their data flow
+ via plain regular expressions in the editor. In particular:
+
+ \begin{itemize}
+
+ \item theories are called \verb|thy|, rarely \verb|theory|
+ (never \verb|thry|)
+
+ \item proof contexts are called \verb|ctxt|, rarely \verb|context| (never \verb|ctx|)
+
+ \item generic contexts are called \verb|context|, rarely
+ \verb|ctxt|
+
+ \item local theories are called \verb|lthy|, unless there is an
+ implicit conversion to a general proof context (semantic super-type)
+
+ \end{itemize}
+
+ Variations with primed or decimal numbers are always possible, as
+ well as ``sematic prefixes'' like \verb|foo_thy| or \verb|bar_ctxt|, but the base conventions above need to be preserved.
+
+ \item The main logical entities (\secref{ch:logic}) also have
+ established naming convention:
+
+ \begin{itemize}
+
+ \item sorts are called \verb|S|
+
+ \item types are called \verb|T|, \verb|U|, or \verb|ty| (never \verb|t|)
+
+ \item terms are called \verb|t|, \verb|u|, or \verb|tm| (never \verb|trm|)
+
+ \item certified types are called \verb|cT|, rarely \verb|T|, with variants as for types
+
+ \item certified terms are called \verb|ct|, rarely \verb|t|, with variants as for terms
+
+ \item theorems are called \verb|th|, or \verb|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
+ \verb|lhs| (not \verb|lhs_tm|). Or a term that is treated
+ specifically as a variable can be called \verb|v| or \verb|x|.
+
+ \end{itemize}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsubsection{General source layout%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
-Like any style guide, also this one should not be interpreted dogmatically, but
- with care and discernment. It consists of a collection of
- recommendations which have been turned out useful over long years of
- Isabelle system development and are supposed to support writing of readable
- and managable code. Special cases encourage derivations,
- as far as you know what you are doing.
- \footnote{This style guide is loosely based on
- \url{http://caml.inria.fr/resources/doc/guides/guidelines.en.html}}
+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. As a rule of thumb,
+ sources need to be printable on plain paper.} 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.\ \verb||\verb,|,\verb|>|) work better at the
+ start of the line, but punctuation is always at the end.
+
+ Function application follows the tradition of \isa{{\isasymlambda}}-calculus,
+ not informal mathematics. For example: \verb|f a b| for a
+ curried function, or \verb|g (a, b)| for a tupled function.
+ Note that the space between \verb|g| and the pair \verb|(a, b)| follows the important principle of
+ \emph{compositionality}: the layout of \verb|g p| does not
+ change when \verb|p| is refined to a concrete pair.
+
+ \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.
+
+ 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 stars etc. should be
+ avoided.
+
+ \paragraph{Complex expressions} consisting of multi-clausal function
+ definitions, \verb|case|, \verb|handle|, \verb|let|,
+ or combinations require special attention. The syntax of Standard
+ ML is a bit too ambitious in admitting too much variance that can
+ distort the meaning of the text.
+
+ Clauses of \verb|fun|, \verb|fn|, \verb|case|,
+ \verb|handle| get extra indentation to indicate the nesting
+ clearly. For 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 \verb|case| or \verb|let|
+ require care to maintain compositionality, to prevent loss of
+ logical indentation where it is particularly imporant to see the
+ structure of the text later on. 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 \verb|case| expressions are optional,
+ but help to analyse the nesting easily based on simple string
+ 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 \verb|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 \verb|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. For example:
+
+ \begin{verbatim}
+ (* RIGHT *)
+
+ fun foo x y = fold (fn z =>
+ expr)
+ \end{verbatim}
+
+ Here the visual appearance is that of three arguments \verb|x|,
+ \verb|y|, \verb|z|.
+
+ \end{enumerate}
+
+ Such weakly structured layout should be use with great care. Here
+ is a counter-example involving \verb|let| expressions:
+
+ \begin{verbatim}
+ (* WRONG *)
+
+ fun foo x = let
+ val y = ...
+ in ... end
+
+ 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).%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsection{SML embedded into Isabelle/Isar%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+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 environment
+ for arbitrary ML values (see also \secref{sec:context}). This
+ formal context holds logical entities as well as ML compiler
+ bindings, among many other things. Raw Standard ML is never
+ encountered again after the initial bootstrap of Isabelle/Pure.
+
+ Object-logics such as Isabelle/HOL are built within the
+ Isabelle/ML/Isar environment of Pure by introducing suitable
+ theories with associated ML text, 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 regular user-space repertoire of Isabelle.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsubsection{Isar ML commands%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+The primary Isar source language provides various facilities
+ to open a ``window'' to the underlying ML compiler. Especially see
+ \indexref{}{command}{use}\hyperlink{command.use}{\mbox{\isa{\isacommand{use}}}} and \indexref{}{command}{ML}\hyperlink{command.ML}{\mbox{\isa{\isacommand{ML}}}}, which work exactly 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 \indexref{}{command}{setup}\hyperlink{command.setup}{\mbox{\isa{\isacommand{setup}}}} or \indexref{}{command}{declaration}\hyperlink{command.declaration}{\mbox{\isa{\isacommand{declaration}}}}.
+ An example of even more fine-grained embedding of ML into Isar is
+ the proof method \indexref{}{method}{tactic}\hyperlink{method.tactic}{\mbox{\isa{tactic}}}, which refines the pending goal state
+ via a given expression of type \verb|tactic|.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isatagmlex
+%
+\begin{isamarkuptext}%
+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.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlex
+{\isafoldmlex}%
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ fun\ factorial\ {\isadigit{0}}\ {\isacharequal}\ {\isadigit{1}}\isanewline
+\ \ \ \ {\isacharbar}\ factorial\ n\ {\isacharequal}\ n\ {\isacharasterisk}\ factorial\ {\isacharparenleft}n\ {\isacharminus}\ {\isadigit{1}}{\isacharparenright}\isanewline
+{\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+Here the ML environment is really managed by Isabelle, i.e.\
+ the \verb|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 or similar command loops of Computer Algebra systems, 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.
+ Instead of \indexref{}{command}{ML}\hyperlink{command.ML}{\mbox{\isa{\isacommand{ML}}}} for theory mode, we use \indexref{}{command}{ML\_prf}\hyperlink{command.ML-prf}{\mbox{\isa{\isacommand{ML{\isacharunderscore}prf}}}} for proof mode. As illustrated below, its effect on the
+ ML environment is local to the whole proof body, while ignoring its
+ internal block structure.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+\isacommand{example{\isacharunderscore}proof}\isamarkupfalse%
+\isanewline
+%
+\isadelimML
+\ \ %
+\endisadelimML
+%
+\isatagML
+\isacommand{ML{\isacharunderscore}prf}\isamarkupfalse%
+\ {\isacharverbatimopen}\ val\ a\ {\isacharequal}\ {\isadigit{1}}\ {\isacharverbatimclose}\isanewline
+\ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse%
+\ %
+\isamarkupcmt{Isar block structure ignored by ML environment%
+}
+\isanewline
+\ \ \ \ \isacommand{ML{\isacharunderscore}prf}\isamarkupfalse%
+\ {\isacharverbatimopen}\ val\ b\ {\isacharequal}\ a\ {\isacharplus}\ {\isadigit{1}}\ {\isacharverbatimclose}\isanewline
+\ \ \isacommand{{\isacharbraceright}}\isamarkupfalse%
+\ %
+\isamarkupcmt{Isar block structure ignored by ML environment%
+}
+\isanewline
+\ \ \isacommand{ML{\isacharunderscore}prf}\isamarkupfalse%
+\ {\isacharverbatimopen}\ val\ c\ {\isacharequal}\ b\ {\isacharplus}\ {\isadigit{1}}\ {\isacharverbatimclose}\isanewline
+\isacommand{qed}\isamarkupfalse%
+%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+By side-stepping the normal scoping rules for Isar proof
+ blocks, embedded ML code can refer to the different contexts
+ explicitly, and manipulate corresponding entities, e.g.\ export a
+ fact from a block context.
+
+ \medskip Two further ML commands are useful in certain situations:
+ \indexref{}{command}{ML\_val}\hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isacharunderscore}val}}}} and \indexref{}{command}{ML\_command}\hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isacharunderscore}command}}}} are both
+ \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 \verb|factorial|: \hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isacharunderscore}val}}}} already takes care of
+ printing the ML toplevel result, but \hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isacharunderscore}command}}}} is silent
+ so we produce an explicit output message.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML{\isacharunderscore}val}\isamarkupfalse%
+\ {\isacharverbatimopen}\ factorial\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}\ {\isacharverbatimclose}\isanewline
+\isacommand{ML{\isacharunderscore}command}\isamarkupfalse%
+\ {\isacharverbatimopen}\ writeln\ {\isacharparenleft}string{\isacharunderscore}of{\isacharunderscore}int\ {\isacharparenleft}factorial\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharparenright}{\isacharparenright}\ {\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+\isanewline
+\isanewline
+\isacommand{example{\isacharunderscore}proof}\isamarkupfalse%
+\isanewline
+%
+\isadelimML
+\ \ %
+\endisadelimML
+%
+\isatagML
+\isacommand{ML{\isacharunderscore}val}\isamarkupfalse%
+\ {\isacharverbatimopen}\ factorial\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}\ {\isacharverbatimclose}\ \ \isanewline
+\ \ \isacommand{ML{\isacharunderscore}command}\isamarkupfalse%
+\ {\isacharverbatimopen}\ writeln\ {\isacharparenleft}string{\isacharunderscore}of{\isacharunderscore}int\ {\isacharparenleft}factorial\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharparenright}{\isacharparenright}\ {\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+\isanewline
+%
+\endisadelimML
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{qed}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isamarkupsubsection{Compile-time context%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+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 via some derived ML functions.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML}{ML\_Context.the\_generic\_context}\verb|ML_Context.the_generic_context: unit -> Context.generic| \\
+ \indexdef{}{ML}{Context.$>$$>$}\verb|Context.>> : (Context.generic -> Context.generic) -> unit| \\
+ \indexdef{}{ML}{bind\_thms}\verb|bind_thms: string * thm list -> unit| \\
+ \indexdef{}{ML}{bind\_thm}\verb|bind_thm: string * thm -> unit| \\
+ \end{mldecls}
\begin{description}
- \item[fundamental law of programming]
- Whenever writing code, keep in mind: A program is
- written once, modified ten times, and read
- hundred times. So simplify its writing,
- always keep future modifications in mind,
- and never jeopardize readability. Every second you hesitate
- to spend on making your code more clear you will
- have to spend ten times understanding what you have
- written later on.
+ \item \verb|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 \verb|ML_Context.the_generic_context ()|
+ correctly. Recall that evaluation of a function body is delayed
+ until actual run-time.
+
+ \item \verb|Context.>>|~\isa{f} applies context transformation
+ \isa{f} to the implicit context of the ML toplevel.
- \item[white space matters]
- Treat white space in your code as if it determines
- the meaning of code.
+ \item \verb|bind_thms|~\isa{{\isacharparenleft}name{\isacharcomma}\ thms{\isacharparenright}} 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.
- \begin{itemize}
+ \item \verb|bind_thm| is similar to \verb|bind_thms| but refers to a
+ singleton theorem.
+
+ \end{description}
- \item The space bar is the easiest key to find on the keyboard,
- press it as often as necessary. \verb|2 + 2| is better
- than \verb|2+2|, likewise \verb|f (x, y)| is
- better than \verb|f(x,y)|.
+ It is very 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.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isamarkupsubsection{Antiquotations \label{sec:ML-antiq}%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+A very important consequence of embedding SML into Isar is the
+ concept of \emph{ML antiquotation}. First, the standard token
+ language of ML is augmented by special syntactic entities of the
+ following form:
- \item Restrict your lines to 80 characters. This will allow
- you to keep the beginning of a line in view while watching
- its end.\footnote{To acknowledge the lax practice of
- text editing these days, we tolerate as much as 100
- characters per line, but anything beyond 120 is not
- considered proper source text.}
-
- \item Ban tabulators; they are a context-sensitive formatting
- feature and likely to confuse anyone not using your favorite
- editor.\footnote{Some modern programming language even
- forbid tabulators altogether according to the formal syntax
- definition.}
-
- \item Get rid of trailing whitespace. Instead, do not
- suppress a trailing newline at the end of your files.
+ \indexouternonterm{antiquote}
+ \begin{rail}
+ antiquote: atsign lbrace nameref args rbrace | lbracesym | rbracesym
+ ;
+ \end{rail}
- \item Choose a generally accepted style of indentation,
- then use it systematically throughout the whole
- application. An indentation of two spaces is appropriate.
- Avoid dangling indentation.
-
- \end{itemize}
+ Here \hyperlink{syntax.nameref}{\mbox{\isa{nameref}}} and \hyperlink{syntax.args}{\mbox{\isa{args}}} are regular outer syntax
+ categories. Note that attributes and proof methods use similar
+ syntax.
- \item[cut-and-paste succeeds over copy-and-paste]
- \emph{Never} copy-and-paste code when programming. If you
- need the same piece of code twice, introduce a
- reasonable auxiliary function (if there is no
- such function, very likely you got something wrong).
- Any copy-and-paste will turn out to be painful
- when something has to be changed later on.
+ \medskip A regular antiquotation \isa{{\isacharat}{\isacharbraceleft}name\ args{\isacharbraceright}} 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. \isa{{\isacharat}{\isacharbraceleft}term\ t{\isacharbraceright}}) or abstract
+ \emph{value} (e.g. \isa{{\isacharat}{\isacharbraceleft}thm\ th{\isacharbraceright}}). 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.
- \item[comments]
- are a device which requires careful thinking before using
- it. The best comment for your code should be the code itself.
- Prefer efforts to write clear, understandable code
- over efforts to explain nasty code.
+ Special antiquotations like \isa{{\isacharat}{\isacharbraceleft}let\ {\isasymdots}{\isacharbraceright}} or \isa{{\isacharat}{\isacharbraceleft}note\ {\isasymdots}{\isacharbraceright}} augment the compilation context without generating code. The
+ non-ASCII braces \isa{{\isasymlbrace}} and \isa{{\isasymrbrace}} allow to delimit the
+ effect by introducing local blocks within the pre-compilation
+ environment.
- \item[functional programming is based on functions]
- Model things as abstract as appropriate. Avoid unnecessarrily
- concrete datatype representations. For example, consider a function
- taking some table data structure as argument and performing
- lookups on it. Instead of taking a table, it could likewise
- take just a lookup function. Accustom
- your way of coding to the level of expressiveness a functional
- programming language is giving onto you.
+ \medskip See also \cite{Wenzel-Chaieb:2007b} for a slightly broader
+ perspective on Isabelle/ML antiquotations.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
+\isatagmlantiq
+%
+\begin{isamarkuptext}%
+\begin{matharray}{rcl}
+ \indexdef{}{ML antiquotation}{let}\hypertarget{ML antiquotation.let}{\hyperlink{ML antiquotation.let}{\mbox{\isa{let}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{note}\hypertarget{ML antiquotation.note}{\hyperlink{ML antiquotation.note}{\mbox{\isa{note}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \end{matharray}
- \item[tuples]
- are often in the way. When there is no striking argument
- to tuple function arguments, just write your function curried.
+ \begin{rail}
+ 'let' ((term + 'and') '=' term + 'and')
+ ;
+
+ 'note' (thmdef? thmrefs + 'and')
+ ;
+ \end{rail}
+
+ \begin{description}
- \item[telling names]
- Any name should tell its purpose as exactly as possible, while
- keeping its length to the absolutely necessary minimum. Always
- give the same name to function arguments which have the same
- meaning. Separate words by underscores (\verb|int_of_string|, not \verb|intOfString|).\footnote{Some
- recent tools for Emacs include special precautions to cope with
- bumpy names in \isa{camelCase}, e.g.\ for improved on-screen
- readability. It is easier to abstain from using such names in the
- first place.}
+ \item \isa{{\isacharat}{\isacharbraceleft}let\ p\ {\isacharequal}\ t{\isacharbraceright}} binds schematic variables in the
+ pattern \isa{p} by higher-order matching against the term \isa{t}. This is analogous to the regular \indexref{}{command}{let}\hyperlink{command.let}{\mbox{\isa{\isacommand{let}}}} command
+ in the Isar proof language. The pre-compilation environment is
+ augmented by auxiliary term bindings, without emitting ML source.
+
+ \item \isa{{\isacharat}{\isacharbraceleft}note\ a\ {\isacharequal}\ b\isactrlsub {\isadigit{1}}\ {\isasymdots}\ b\isactrlsub n{\isacharbraceright}} recalls existing facts \isa{b\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ b\isactrlsub n}, binding the result as \isa{a}. This is analogous to
+ the regular \indexref{}{command}{note}\hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}} command in the Isar proof language.
+ The pre-compilation environment is augmented by auxiliary fact
+ bindings, without emitting ML source.
\end{description}%
\end{isamarkuptext}%
\isamarkuptrue%
%
-\isamarkupsection{Thread-safe programming%
+\endisatagmlantiq
+{\isafoldmlantiq}%
+%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isatagmlex
+%
+\begin{isamarkuptext}%
+The following artificial example gives a first
+ impression about using the antiquotation elements introduced so far,
+ together with the basic \isa{{\isacharat}{\isacharbraceleft}thm{\isacharbraceright}} antiquotation defined
+ later.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlex
+{\isafoldmlex}%
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ {\isaantiqopen}\isanewline
+\ \ \ \ %
+\isaantiq
+let\ {\isacharquery}t\ {\isacharequal}\ my{\isacharunderscore}term%
+\endisaantiq
+\isanewline
+\ \ \ \ %
+\isaantiq
+note\ my{\isacharunderscore}refl\ {\isacharequal}\ reflexive\ {\isacharbrackleft}of\ {\isacharquery}t{\isacharbrackright}%
+\endisaantiq
+\isanewline
+\ \ \ \ fun\ foo\ th\ {\isacharequal}\ Thm{\isachardot}transitive\ th\ %
+\isaantiq
+thm\ my{\isacharunderscore}refl%
+\endisaantiq
+\isanewline
+\ \ {\isaantiqclose}\isanewline
+{\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+The extra block delimiters do not affect the compiled code itself, i.e.\
+ function \verb|foo| is available in the present context.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsection{Canonical argument order \label{sec:canonical-argument-order}%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+Standard ML is a language in the tradition of \isa{{\isasymlambda}}-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.
+
+ First of all, functions are usually \emph{curried}: the idea of
+ \isa{f} taking \isa{n} arguments of type \isa{{\isasymtau}\isactrlsub i} (for
+ \isa{i\ {\isasymin}\ {\isacharbraceleft}{\isadigit{1}}{\isacharcomma}\ {\isasymdots}\ n{\isacharbraceright}}) with result \isa{{\isasymtau}} is represented by
+ the iterated function space \isa{{\isasymtau}\isactrlsub {\isadigit{1}}\ {\isasymrightarrow}\ {\isasymdots}\ {\isasymrightarrow}\ {\isasymtau}\isactrlsub n\ {\isasymrightarrow}\ {\isasymtau}}. This is
+ isomorphic to the encoding via tuples \isa{{\isasymtau}\isactrlsub {\isadigit{1}}\ {\isasymtimes}\ {\isasymdots}\ {\isasymtimes}\ {\isasymtau}\isactrlsub n\ {\isasymrightarrow}\ {\isasymtau}}, but
+ the curried version fits more smoothly into the basic
+ calculus.\footnote{The difference is even more significant in
+ higher-order logic, because additional the redundant tuple structure
+ needs to be accommodated by formal reasoning.}
+
+ Currying gives some flexiblity due to \emph{partial application}. A
+ function \isa{f{\isacharcolon}\ {\isasymtau}\isactrlsub {\isadigit{1}}\ {\isasymrightarrow}\ {\isasymtau}\isactrlbsub {\isadigit{2}}\isactrlesub \ {\isasymrightarrow}\ {\isasymtau}} can be applied to \isa{x{\isacharcolon}\ {\isasymtau}\isactrlsub {\isadigit{1}}}
+ and the remaining \isa{{\isacharparenleft}f\ x{\isacharparenright}{\isacharcolon}\ {\isasymtau}\isactrlsub {\isadigit{2}}\ {\isasymrightarrow}\ {\isasymtau}} passed to other functions
+ 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 of function definitions and
+ minimizes adhoc permutations in their application. In Isabelle/ML,
+ large portions of non-trivial source code are written without ever
+ using the infamous function \isa{swap{\isacharcolon}\ {\isasymalpha}\ {\isasymtimes}\ {\isasymbeta}\ {\isasymrightarrow}\ {\isasymbeta}\ {\isasymtimes}\ {\isasymalpha}} or the
+ combinator \isa{C{\isacharcolon}\ {\isacharparenleft}{\isasymalpha}\ {\isasymrightarrow}\ {\isasymbeta}\ {\isasymrightarrow}\ {\isasymgamma}{\isacharparenright}\ {\isasymrightarrow}\ {\isacharparenleft}{\isasymbeta}\ {\isasymrightarrow}\ {\isasymalpha}\ {\isasymrightarrow}\ {\isasymgamma}{\isacharparenright}}, which is not even
+ defined in our 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 \isa{{\isasymbeta}} that manages elements of type \isa{{\isasymalpha}}. Both
+ the names and types of the associated operations are canonical for
+ Isabelle/ML.
+
+ \medskip
+ \begin{tabular}{ll}
+ kind & canonical name and type \\\hline
+ selector & \isa{member{\isacharcolon}\ {\isasymbeta}\ {\isasymrightarrow}\ {\isasymalpha}\ {\isasymrightarrow}\ bool} \\
+ update & \isa{insert{\isacharcolon}\ {\isasymalpha}\ {\isasymrightarrow}\ {\isasymbeta}\ {\isasymrightarrow}\ {\isasymbeta}} \\
+ \end{tabular}
+ \medskip
+
+ Given a container \isa{B{\isacharcolon}\ {\isasymbeta}}, the partially applied \isa{member\ B} is a predicate over elements \isa{{\isasymalpha}\ {\isasymrightarrow}\ 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: \isa{forall\ {\isacharparenleft}member\ B{\isacharparenright}\ list} will check a list of
+ elements for membership in \isa{B} etc. Often the explicit
+ \isa{list} is pointless and can be contracted in the application
+ \isa{forall\ {\isacharparenleft}member\ B{\isacharparenright}} to get directly a predicate again.
+
+ The update operation naturally ``varies'' the container, so it moves
+ to the right: \isa{insert\ a} is a function \isa{{\isasymbeta}\ {\isasymrightarrow}\ {\isasymbeta}} to
+ insert a value \isa{a}. These can be composed naturally as
+ follows: \isa{insert\ c\ {\isasymcirc}\ insert\ b\ {\isasymcirc}\ insert\ a}. This works well,
+ apart from some awkwardness due to conventional mathematical
+ function composition, which can be easily amended as follows.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsubsection{Forward application and composition%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+Regular function application and infix notation works best for
+ relatively deeply structured expressions, e.g.\ \isa{h\ {\isacharparenleft}f\ x\ y\ {\isacharplus}\ g\ z{\isacharparenright}}. The important special case if \emph{linear transformation}
+ applies a cascade of functions as follows: \isa{f\isactrlsub n\ {\isacharparenleft}{\isasymdots}\ {\isacharparenleft}f\isactrlsub {\isadigit{1}}\ x{\isacharparenright}{\isacharparenright}}.
+ This becomes hard to read and maintain if the functions are
+ themselves complex expressions. The notation can be significantly
+ improved by introducing \emph{forward} versions of application and
+ composition as follows:
+
+ \medskip
+ \begin{tabular}{lll}
+ \isa{x\ {\isacharbar}{\isachargreater}\ f} & \isa{{\isasymequiv}} & \isa{f\ x} \\
+ \isa{f\ {\isacharhash}{\isachargreater}\ g} & \isa{{\isasymequiv}} & \isa{x\ {\isacharbar}{\isachargreater}\ f\ {\isacharbar}{\isachargreater}\ g} \\
+ \end{tabular}
+ \medskip
+
+ This enables to write conveniently \isa{x\ {\isacharbar}{\isachargreater}\ f\isactrlsub {\isadigit{1}}\ {\isacharbar}{\isachargreater}\ {\isasymdots}\ {\isacharbar}{\isachargreater}\ f\isactrlsub n} or
+ \isa{f\isactrlsub {\isadigit{1}}\ {\isacharhash}{\isachargreater}\ {\isasymdots}\ {\isacharhash}{\isachargreater}\ f\isactrlsub n} for its functional abstraction over \isa{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}
+ \isa{{\isacharparenleft}x{\isacharcomma}\ y{\isacharparenright}\ {\isacharbar}{\isacharminus}{\isachargreater}\ f} & \isa{{\isasymequiv}} & \isa{f\ x\ y} \\
+ \isa{f\ {\isacharhash}{\isacharminus}{\isachargreater}\ g} & \isa{{\isasymequiv}} & \isa{x\ {\isacharbar}{\isachargreater}\ f\ {\isacharbar}{\isacharminus}{\isachargreater}\ g} \\
+ \end{tabular}
+ \medskip%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML}{$\mid$$>$}\verb|op |\verb,|,\verb|> : 'a * ('a -> 'b) -> 'b| \\
+ \indexdef{}{ML}{$\mid$-$>$}\verb|op |\verb,|,\verb|-> : ('c * 'a) * ('c -> 'a -> 'b) -> 'b| \\
+ \indexdef{}{ML}{\#$>$}\verb|op #> : ('a -> 'b) * ('b -> 'c) -> 'a -> 'c| \\
+ \indexdef{}{ML}{\#-$>$}\verb|op #-> : ('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd| \\
+ \end{mldecls}
+
+ %FIXME description!?%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isamarkupsubsection{Canonical iteration%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+As explained above, a function \isa{f{\isacharcolon}\ {\isasymalpha}\ {\isasymrightarrow}\ {\isasymbeta}\ {\isasymrightarrow}\ {\isasymbeta}} can be
+ understood as update on a configuration of type \isa{{\isasymbeta}} that is
+ parametrized by arguments of type \isa{{\isasymalpha}}. Given \isa{a{\isacharcolon}\ {\isasymalpha}}
+ the partial application \isa{{\isacharparenleft}f\ a{\isacharparenright}{\isacharcolon}\ {\isasymbeta}\ {\isasymrightarrow}\ {\isasymbeta}} operates
+ homogeneously on \isa{{\isasymbeta}}. This can be iterated naturally over a
+ list of parameters \isa{{\isacharbrackleft}a\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ a\isactrlsub n{\isacharbrackright}} as \isa{f\ a\isactrlsub {\isadigit{1}}\ {\isacharhash}{\isachargreater}\ {\isasymdots}\ {\isacharhash}{\isachargreater}\ f\ a\isactrlbsub n\isactrlesub \isactrlbsub \isactrlesub }. The latter expression is again a function \isa{{\isasymbeta}\ {\isasymrightarrow}\ {\isasymbeta}}.
+ It can be applied to an initial configuration \isa{b{\isacharcolon}\ {\isasymbeta}} to
+ start the iteration over the given list of arguments: each \isa{a} in \isa{a\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ a\isactrlsub n} is applied consecutively by updating a
+ cumulative configuration.
+
+ The \isa{fold} combinator in Isabelle/ML lifts a function \isa{f} as above to its iterated version over a list of arguments.
+ Lifting can be repeated, e.g.\ \isa{{\isacharparenleft}fold\ {\isasymcirc}\ fold{\isacharparenright}\ f} iterates
+ over a list of lists as expected.
+
+ The variant \isa{fold{\isacharunderscore}rev} works inside-out over the list of
+ arguments, such that \isa{fold{\isacharunderscore}rev\ f\ {\isasymequiv}\ fold\ f\ {\isasymcirc}\ rev} holds.
+
+ The \isa{fold{\isacharunderscore}map} combinator essentially performs \isa{fold} and \isa{map} simultaneously: each application of \isa{f} produces an updated configuration together with a side-result;
+ the iteration collects all such side-results as a separate list.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML}{fold}\verb|fold: ('a -> 'b -> 'b) -> 'a list -> 'b -> 'b| \\
+ \indexdef{}{ML}{fold\_rev}\verb|fold_rev: ('a -> 'b -> 'b) -> 'a list -> 'b -> 'b| \\
+ \indexdef{}{ML}{fold\_map}\verb|fold_map: ('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b| \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item \verb|fold|~\isa{f} lifts the parametrized update function
+ \isa{f} to a list of parameters.
+
+ \item \verb|fold_rev|~\isa{f} is similar to \verb|fold|~\isa{f}, but works inside-out.
+
+ \item \verb|fold_map|~\isa{f} lifts the parametrized update
+ function \isa{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 \isa{foldl}, \isa{foldr} etc. SML97
+ provides its own variations as \verb|List.foldl| and \verb|List.foldr|, while the classic Isabelle library still has the
+ slightly more convenient historic \verb|Library.foldl| and \verb|Library.foldr|. To avoid further confusion, all of this should be
+ ignored, and \verb|fold| (or \verb|fold_rev|) used exclusively.
+ \end{warn}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isatagmlex
+%
+\begin{isamarkuptext}%
+The following example shows how to fill a text buffer
+ incrementally by adding strings, either individually or from a given
+ list.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlex
+{\isafoldmlex}%
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ val\ s\ {\isacharequal}\isanewline
+\ \ \ \ Buffer{\isachardot}empty\isanewline
+\ \ \ \ {\isacharbar}{\isachargreater}\ Buffer{\isachardot}add\ {\isachardoublequote}digits{\isacharcolon}\ {\isachardoublequote}\isanewline
+\ \ \ \ {\isacharbar}{\isachargreater}\ fold\ {\isacharparenleft}Buffer{\isachardot}add\ o\ string{\isacharunderscore}of{\isacharunderscore}int{\isacharparenright}\ {\isacharparenleft}{\isadigit{0}}\ upto\ {\isadigit{9}}{\isacharparenright}\isanewline
+\ \ \ \ {\isacharbar}{\isachargreater}\ Buffer{\isachardot}content{\isacharsemicolon}\isanewline
+\isanewline
+\ \ %
+\isaantiq
+assert%
+\endisaantiq
+\ {\isacharparenleft}s\ {\isacharequal}\ {\isachardoublequote}digits{\isacharcolon}\ {\isadigit{0}}{\isadigit{1}}{\isadigit{2}}{\isadigit{3}}{\isadigit{4}}{\isadigit{5}}{\isadigit{6}}{\isadigit{7}}{\isadigit{8}}{\isadigit{9}}{\isachardoublequote}{\isacharparenright}{\isacharsemicolon}\isanewline
+{\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+Note how \verb|fold (Buffer.add o string_of_int)| above saves
+ an extra \verb|map| over the given list. This kind of peephole
+ optimization reduces both the code size and the tree structures in
+ memory (``deforestation''), but requires some practice to read and
+ write it fluently.
+
+ \medskip The next example elaborates this idea and demonstrates fast
+ accumulation of tree content using a text buffer.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ datatype\ tree\ {\isacharequal}\ Text\ of\ string\ {\isacharbar}\ Elem\ of\ string\ {\isacharasterisk}\ tree\ list{\isacharsemicolon}\isanewline
+\isanewline
+\ \ fun\ slow{\isacharunderscore}content\ {\isacharparenleft}Text\ txt{\isacharparenright}\ {\isacharequal}\ txt\isanewline
+\ \ \ \ {\isacharbar}\ slow{\isacharunderscore}content\ {\isacharparenleft}Elem\ {\isacharparenleft}name{\isacharcomma}\ ts{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline
+\ \ \ \ \ \ \ \ {\isachardoublequote}{\isacharless}{\isachardoublequote}\ {\isacharcircum}\ name\ {\isacharcircum}\ {\isachardoublequote}{\isachargreater}{\isachardoublequote}\ {\isacharcircum}\isanewline
+\ \ \ \ \ \ \ \ implode\ {\isacharparenleft}map\ slow{\isacharunderscore}content\ ts{\isacharparenright}\ {\isacharcircum}\isanewline
+\ \ \ \ \ \ \ \ {\isachardoublequote}{\isacharless}{\isacharslash}{\isachardoublequote}\ {\isacharcircum}\ name\ {\isacharcircum}\ {\isachardoublequote}{\isachargreater}{\isachardoublequote}\isanewline
+\isanewline
+\ \ fun\ add{\isacharunderscore}content\ {\isacharparenleft}Text\ txt{\isacharparenright}\ {\isacharequal}\ Buffer{\isachardot}add\ txt\isanewline
+\ \ \ \ {\isacharbar}\ add{\isacharunderscore}content\ {\isacharparenleft}Elem\ {\isacharparenleft}name{\isacharcomma}\ ts{\isacharparenright}{\isacharparenright}\ {\isacharequal}\isanewline
+\ \ \ \ \ \ \ \ Buffer{\isachardot}add\ {\isacharparenleft}{\isachardoublequote}{\isacharless}{\isachardoublequote}\ {\isacharcircum}\ name\ {\isacharcircum}\ {\isachardoublequote}{\isachargreater}{\isachardoublequote}{\isacharparenright}\ {\isacharhash}{\isachargreater}\isanewline
+\ \ \ \ \ \ \ \ fold\ add{\isacharunderscore}content\ ts\ {\isacharhash}{\isachargreater}\isanewline
+\ \ \ \ \ \ \ \ Buffer{\isachardot}add\ {\isacharparenleft}{\isachardoublequote}{\isacharless}{\isacharslash}{\isachardoublequote}\ {\isacharcircum}\ name\ {\isacharcircum}\ {\isachardoublequote}{\isachargreater}{\isachardoublequote}{\isacharparenright}{\isacharsemicolon}\isanewline
+\isanewline
+\ \ fun\ fast{\isacharunderscore}content\ tree\ {\isacharequal}\isanewline
+\ \ \ \ Buffer{\isachardot}empty\ {\isacharbar}{\isachargreater}\ add{\isacharunderscore}content\ tree\ {\isacharbar}{\isachargreater}\ Buffer{\isachardot}content{\isacharsemicolon}\isanewline
+{\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+The slow part of \verb|slow_content| is the \verb|implode| of
+ the recursive results, because it copies previously produced strings
+ afresh.
+
+ The incremental \verb|add_content| avoids this by operating on a
+ buffer that is passed-through in a linear fashion. Using \verb|#>| and contraction over the actual \verb|Buffer.T| argument
+ saves some additional boiler-plate, but requires again some
+ practice. Of course, the two \verb|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 \verb|add_content| is its ``open'' form as a
+ function on \verb|Buffer.T| that can be continued in further
+ linear transformations, folding etc. Thus it is more compositional
+ than the naive \verb|slow_content|. As a realistic example, compare
+ the slightly old-style \verb|Term.maxidx_of_term: term -> int| with
+ the newer \verb|Term.maxidx_term: term -> int -> int| in
+ Isabelle/Pure.
+
+ Note that \verb|fast_content| above is only defined for completeness
+ of the example. In many practical situations, it is customary to
+ defined the incremental \verb|add_content| only and leave the
+ initialization and termination to the concrete application by the
+ user.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsection{Message output channels \label{sec:message-channels}%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+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 \verb|###| and errors by \verb|***|, 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.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML}{writeln}\verb|writeln: string -> unit| \\
+ \indexdef{}{ML}{tracing}\verb|tracing: string -> unit| \\
+ \indexdef{}{ML}{warning}\verb|warning: string -> unit| \\
+ \indexdef{}{ML}{error}\verb|error: string -> 'a| \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item \verb|writeln|~\isa{text} outputs \isa{text} as regular
+ message. This is the primary message output operation of Isabelle
+ and should be used by default.
+
+ \item \verb|tracing|~\isa{text} outputs \isa{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 \verb|warning|~\isa{text} outputs \isa{text} as
+ warning, which typically means some extra emphasis on the front-end
+ side (color highlighting, icon).
+
+ \item \verb|error|~\isa{text} raises exception \verb|ERROR|~\isa{text} and thus lets the Isar toplevel print \isa{text} on the
+ error channel, which typically means some extra emphasis on the
+ front-end side (color highlighting, icon).
+
+ This assumes that the exception is not handled before the command
+ terminates. Handling exception \verb|ERROR|~\isa{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 \verb|Output.error_msg|, but
+ the 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 here.
+ \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 \verb|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.
+ \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
+ needs to be issued by a \emph{single} invocation of \verb|writeln|
+ etc. with the functional concatenation of all message constituents.
+ \end{warn}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isatagmlex
+%
+\begin{isamarkuptext}%
+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.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlex
+{\isafoldmlex}%
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML{\isacharunderscore}command}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ warning\ {\isacharparenleft}cat{\isacharunderscore}lines\isanewline
+\ \ \ {\isacharbrackleft}{\isachardoublequote}Beware\ the\ Jabberwock{\isacharcomma}\ my\ son{\isacharbang}{\isachardoublequote}{\isacharcomma}\isanewline
+\ \ \ \ {\isachardoublequote}The\ jaws\ that\ bite{\isacharcomma}\ the\ claws\ that\ catch{\isacharbang}{\isachardoublequote}{\isacharcomma}\isanewline
+\ \ \ \ {\isachardoublequote}Beware\ the\ Jubjub\ Bird{\isacharcomma}\ and\ shun{\isachardoublequote}{\isacharcomma}\isanewline
+\ \ \ \ {\isachardoublequote}The\ frumious\ Bandersnatch{\isacharbang}{\isachardoublequote}{\isacharbrackright}{\isacharparenright}{\isacharsemicolon}\isanewline
+{\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isamarkupsection{Exceptions \label{sec:exceptions}%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
-Recent versions of Poly/ML (5.2.1 or later) support robust
- multithreaded execution, based on native operating system threads of
- the underlying platform. Thus threads will actually be executed in
- parallel on multi-core systems. A speedup-factor of approximately
- 1.5--3 can be expected on a regular 4-core machine.\footnote{There
- is some inherent limitation of the speedup factor due to garbage
- collection, which is still sequential. It helps to provide initial
- heap space generously, using the \texttt{-H} option of Poly/ML.}
- Threads also help to organize advanced operations of the system,
- with explicit communication between sub-components, real-time
- conditions, time-outs etc.
+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 no longer necessary, because the ML
+ runtime system prints a detailed source position of the
+ corresponding \verb|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 end-users or ML
+ users of a module can be often minimized by using plain user errors.
+
+ \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 an indispensable tool to implement managed
+ evaluation of Isar 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}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML}{try}\verb|try: ('a -> 'b) -> 'a -> 'b option| \\
+ \indexdef{}{ML}{can}\verb|can: ('a -> 'b) -> 'a -> bool| \\
+ \indexdef{}{ML}{ERROR}\verb|ERROR: string -> exn| \\
+ \indexdef{}{ML}{Fail}\verb|Fail: string -> exn| \\
+ \indexdef{}{ML}{Exn.is\_interrupt}\verb|Exn.is_interrupt: exn -> bool| \\
+ \indexdef{}{ML}{reraise}\verb|reraise: exn -> 'a| \\
+ \indexdef{}{ML}{exception\_trace}\verb|exception_trace: (unit -> 'a) -> 'a| \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item \verb|try|~\isa{f\ x} makes the partiality of evaluating
+ \isa{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 \verb|(SOME|~\isa{f\ x}~\verb|handle _ => NONE)| that is occasionally seen in
+ books about SML.
+
+ \item \verb|can| is similar to \verb|try| with more abstract result.
+
+ \item \verb|ERROR|~\isa{msg} represents user errors; this
+ exception is always raised via the \verb|error| function (see
+ \secref{sec:message-channels}).
+
+ \item \verb|Fail|~\isa{msg} represents general program failures.
+
+ \item \verb|Exn.is_interrupt| identifies interrupts robustly, without
+ mentioning concrete exception constructors in user code. Handled
+ interrupts need to be re-raised promptly!
+
+ \item \verb|reraise|~\isa{exn} raises exception \isa{exn}
+ while preserving its implicit position information (if possible,
+ depending on the ML platform).
+
+ \item \verb|exception_trace|~\verb|(fn _ =>|~\isa{e}\verb|)| evaluates expression \isa{e} while printing
+ a full trace of its stack of nested exceptions (if possible,
+ depending on the ML platform).\footnote{In various versions of
+ Poly/ML the trace will appear on raw stdout of the Isabelle
+ process.}
+
+ Inserting \verb|exception_trace| into ML code temporarily is useful
+ for debugging, but not suitable for production code.
+
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
+\isatagmlantiq
+%
+\begin{isamarkuptext}%
+\begin{matharray}{rcl}
+ \indexdef{}{ML antiquotation}{assert}\hypertarget{ML antiquotation.assert}{\hyperlink{ML antiquotation.assert}{\mbox{\isa{assert}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \end{matharray}
+
+ \begin{description}
+
+ \item \isa{{\isacharat}{\isacharbraceleft}assert{\isacharbraceright}} inlines a function \isa{bool\ {\isasymRightarrow}\ unit}
+ that raises \verb|Fail| if the argument is \verb|false|. Due to
+ inlining the source position of failed assertions is included in the
+ error output.
+
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlantiq
+{\isafoldmlantiq}%
+%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
+\isamarkupsection{Basic data types%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+The basis library proposal of SML97 need 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 can cause serious
+ problems with the parallel evaluation model of Isabelle/ML (such as
+ \verb|TextIO.print| or \verb|OS.Process.system|).
+
+ Subsequently we give a brief overview of important operations on
+ basic ML data types.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsubsection{Characters%
+}
+\isamarkuptrue%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML type}{char}\verb|type char| \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type \verb|char| is \emph{not} used. The smallest textual
+ unit in Isabelle is represented a ``symbol'' (see
+ \secref{sec:symbols}).
+
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isamarkupsubsection{Integers%
+}
+\isamarkuptrue%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML type}{int}\verb|type int| \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item Type \verb|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 (e.g.\ \verb|123456789012345678901234567890|) are forced to be of this one true
+ integer type --- overloading of SML97 is disabled.
+
+ Structure \verb|IntInf| of SML97 is obsolete, always use
+ \verb|Int|. Structure \verb|Integer| in \hyperlink{file.~~/src/Pure/General/integer.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}General{\isacharslash}integer{\isachardot}ML}}}} provides some additional
+ operations.
- 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 immediately available
- to other threads, without requiring untyped character streams,
- awkward serialization etc.
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isamarkupsubsection{Options%
+}
+\isamarkuptrue%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML}{Option.map}\verb|Option.map: ('a -> 'b) -> 'a option -> 'b option| \\
+ \indexdef{}{ML}{is\_some}\verb|is_some: 'a option -> bool| \\
+ \indexdef{}{ML}{is\_none}\verb|is_none: 'a option -> bool| \\
+ \indexdef{}{ML}{the}\verb|the: 'a option -> 'a| \\
+ \indexdef{}{ML}{these}\verb|these: 'a list option -> 'a list| \\
+ \indexdef{}{ML}{the\_list}\verb|the_list: 'a option -> 'a list| \\
+ \indexdef{}{ML}{the\_default}\verb|the_default: 'a -> 'a option -> 'a| \\
+ \end{mldecls}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\begin{isamarkuptext}%
+Apart from \verb|Option.map| most operations defined in
+ structure \verb|Option| are alien to Isabelle/ML. The
+ operations shown above are defined in \hyperlink{file.~~/src/Pure/General/basics.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}General{\isacharslash}basics{\isachardot}ML}}}}, among others.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsubsection{Lists%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+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.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML}{cons}\verb|cons: 'a -> 'a list -> 'a list| \\
+ \indexdef{}{ML}{member}\verb|member: ('b * 'a -> bool) -> 'a list -> 'b -> bool| \\
+ \indexdef{}{ML}{insert}\verb|insert: ('a * 'a -> bool) -> 'a -> 'a list -> 'a list| \\
+ \indexdef{}{ML}{remove}\verb|remove: ('b * 'a -> bool) -> 'b -> 'a list -> 'a list| \\
+ \indexdef{}{ML}{update}\verb|update: ('a * 'a -> bool) -> 'a -> 'a list -> 'a list| \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item \verb|cons|~\isa{x\ xs} evaluates to \isa{x\ {\isacharcolon}{\isacharcolon}\ xs}.
+
+ Tupled infix operators are a historical accident in Standard ML.
+ The curried \verb|cons| amends this, but it should be only used when
+ partial application is required.
+
+ \item \verb|member|, \verb|insert|, \verb|remove|, \verb|update| treat
+ lists as a set-like container that maintains the order of elements.
+ See \hyperlink{file.~~/src/Pure/library.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}library{\isachardot}ML}}}} for the full specifications
+ (written in ML). There are some further derived operations like
+ \verb|union| or \verb|inter|.
+
+ Note that \verb|insert| is conservative about elements that are
+ already a \verb|member| of the list, while \verb|update| ensures that
+ the last 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: more recent declarations normally take precedence over
+ earlier ones.
- On the other hand, some programming guidelines need to be observed
- in order to make unprotected parallelism work out smoothly. While
- the ML system implementation is responsible to maintain basic
- integrity of the representation of ML values in memory, the
- application programmer needs to ensure that multithreaded execution
- does not break the intended semantics.
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isatagmlex
+%
+\begin{isamarkuptext}%
+Using canonical \verb|fold| together with canonical \verb|cons|, or similar standard operations, alternates the orientation of
+ data. The is quite natural and should not altered forcible by
+ inserting extra applications \verb|rev|. The alternative \verb|fold_rev| can be used in the relatively few situations, where
+ alternation should be prevented.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlex
+{\isafoldmlex}%
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ val\ items\ {\isacharequal}\ {\isacharbrackleft}{\isadigit{1}}{\isacharcomma}\ {\isadigit{2}}{\isacharcomma}\ {\isadigit{3}}{\isacharcomma}\ {\isadigit{4}}{\isacharcomma}\ {\isadigit{5}}{\isacharcomma}\ {\isadigit{6}}{\isacharcomma}\ {\isadigit{7}}{\isacharcomma}\ {\isadigit{8}}{\isacharcomma}\ {\isadigit{9}}{\isacharcomma}\ {\isadigit{1}}{\isadigit{0}}{\isacharbrackright}{\isacharsemicolon}\isanewline
+\isanewline
+\ \ val\ list{\isadigit{1}}\ {\isacharequal}\ fold\ cons\ items\ {\isacharbrackleft}{\isacharbrackright}{\isacharsemicolon}\isanewline
+\ \ %
+\isaantiq
+assert%
+\endisaantiq
+\ {\isacharparenleft}list{\isadigit{1}}\ {\isacharequal}\ rev\ items{\isacharparenright}{\isacharsemicolon}\isanewline
+\isanewline
+\ \ val\ list{\isadigit{2}}\ {\isacharequal}\ fold{\isacharunderscore}rev\ cons\ items\ {\isacharbrackleft}{\isacharbrackright}{\isacharsemicolon}\isanewline
+\ \ %
+\isaantiq
+assert%
+\endisaantiq
+\ {\isacharparenleft}list{\isadigit{2}}\ {\isacharequal}\ items{\isacharparenright}{\isacharsemicolon}\isanewline
+{\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+The subsequent example demonstrates how to \emph{merge} two
+ lists in a natural way.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ fun\ merge{\isacharunderscore}lists\ eq\ {\isacharparenleft}xs{\isacharcomma}\ ys{\isacharparenright}\ {\isacharequal}\ fold{\isacharunderscore}rev\ {\isacharparenleft}insert\ eq{\isacharparenright}\ ys\ xs{\isacharsemicolon}\isanewline
+{\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+Here the first list is treated conservatively: only the new
+ elements from the second list are inserted. The inside-out order of
+ insertion via \verb|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 \verb|merge| as defined in
+ \hyperlink{file.~~/src/Pure/library.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}library{\isachardot}ML}}}}.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsubsection{Association lists%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+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.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML}{AList.lookup}\verb|AList.lookup: ('a * 'b -> bool) -> ('b * 'c) list -> 'a -> 'c option| \\
+ \indexdef{}{ML}{AList.defined}\verb|AList.defined: ('a * 'b -> bool) -> ('b * 'c) list -> 'a -> bool| \\
+ \indexdef{}{ML}{AList.update}\verb|AList.update: ('a * 'a -> bool) -> 'a * 'b -> ('a * 'b) list -> ('a * 'b) list| \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item \verb|AList.lookup|, \verb|AList.defined|, \verb|AList.update|
+ implement the main ``framework operations'' for mappings in
+ Isabelle/ML, with standard conventions for names and types.
+
+ Note that a function called \isa{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
+ \isa{the{\isacharunderscore}element} or \isa{get}.
+
+ The \isa{defined} operation is essentially a contraction of \verb|is_some| and \isa{lookup}, but this is sufficiently frequent to
+ justify its independent existence. This also gives the
+ implementation some opportunity for peep-hole optimization.
+
+ \end{description}
- \medskip \paragraph{Critical shared resources.} Actually only those
- parts outside the purely functional world of ML are critical. In
- particular, this covers
+ 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 \hyperlink{file.~~/src/Pure/General/table.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}General{\isacharslash}table{\isachardot}ML}}}}; that version scales easily to
+ thousands or millions of elements.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsubsection{Unsynchronized references%
+}
+\isamarkuptrue%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML type}{Unsynchronized.ref}\verb|type 'a Unsynchronized.ref| \\
+ \indexdef{}{ML}{Unsynchronized.ref}\verb|Unsynchronized.ref: 'a -> 'a Unsynchronized.ref| \\
+ \indexdef{}{ML}{!}\verb|! : 'a Unsynchronized.ref -> 'a| \\
+ \indexdef{}{ML}{:=}\verb|op := : 'a Unsynchronized.ref * 'a -> unit| \\
+ \end{mldecls}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\begin{isamarkuptext}%
+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 \verb|Unsynchronized.ref| for the constructor
+ for references in Isabelle/ML emphasizes the inconveniences caused by
+ mutability. Existing operations \verb|!| and \verb|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.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsection{Thread-safe programming \label{sec:multi-threading}%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+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.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsubsection{Multi-threading with shared memory%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+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. A maximum speedup-factor
+ of 3.0 on 4 cores and 5.0 on 8 cores can be
+ expected.\footnote{Further scalability is limited due to garbage
+ collection, which is still sequential in Poly/ML 5.2/5.3/5.4. It
+ helps to provide initial heap space generously, using the
+ \texttt{-H} option. Initial heap size needs to be scaled-up
+ together with the number of CPU cores: approximately 1--2\,GB per
+ core..}
+
+ \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
+ immediately available to other threads: abstract values can be
+ passed between threads without copying or awkward serialization that
+ is typically required for explicit message passing between 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}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isamarkupsubsection{Critical shared resources%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+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.\ those 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 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 direct I/O on shared channels, notably \isa{stdin}, \isa{stdout}, \isa{stderr}.
+ \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 other processes.
\end{itemize}
- The majority of tools implemented within the Isabelle/Isar framework
- will not require any of these critical elements: nothing special
- needs to be observed when staying in the purely functional fragment
- of ML. Note that output via the official Isabelle channels does not
- count as direct I/O, so the operations \verb|writeln|, \verb|warning|, \verb|tracing| etc.\ are safe.
-
- Moreover, ML bindings within the toplevel environment (\verb|type|, \verb|val|, \verb|structure| etc.) due to
- run-time invocation of the compiler are also safe, because
- Isabelle/Isar manages this as part of the theory or proof context.
-
- \paragraph{Multithreading in Isabelle/Isar.} The theory loader
- automatically exploits the overall parallelism of independent nodes
- in the development graph, as well as the inherent irrelevance of
- proofs for goals being fully specified in advance. This means,
- checking of individual Isar proofs is parallelized by default.
- Beyond that, very sophisticated proof tools may use local
- parallelism internally, via the general programming model of
- ``future values'' (see also \hyperlink{file.~~/src/Pure/Concurrent/future.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}Concurrent{\isacharslash}future{\isachardot}ML}}}}).
-
- Any ML code that works relatively to the present background theory
- is already safe. Contextual data may be easily stored within the
- theory or proof context, thanks to the generic data concept of
- Isabelle/Isar (see \secref{sec:context-data}). This greatly
- diminishes the demand for global state information in the first
- place.
-
- \medskip In rare situations where actual mutable content needs to be
- manipulated, Isabelle provides a single \emph{critical section} that
- may be entered while preventing any other thread from doing the
- same. Entering the critical section without contention is very
- fast, and several basic system operations do so frequently. This
- also means that each thread should leave the critical section
- quickly, otherwise parallel execution performance may degrade
- significantly.
-
- Despite this potential bottle-neck, centralized locking is
- convenient, because it prevents deadlocks without demanding special
- precautions. Explicit communication demands other means, though.
- The high-level abstraction of synchronized variables \hyperlink{file.~~/src/Pure/Concurrent/synchronized.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}Concurrent{\isacharslash}synchronized{\isachardot}ML}}}} enables parallel
- components to communicate via shared state; see also \hyperlink{file.~~/src/Pure/Concurrent/mailbox.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}Concurrent{\isacharslash}mailbox{\isachardot}ML}}}} as canonical example.
-
- \paragraph{Good conduct of impure programs.} The following
- guidelines enable non-functional programs to participate in
- multithreading.
+ Isabelle/ML provides various mechanisms to avoid critical shared
+ resources in most practical 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 Minimize global state information. Using proper theory and
- proof context data will actually return to functional update of
- values, without any special precautions for multithreading. Apart
- from the fully general functors for theory and proof data (see
- \secref{sec:context-data}) there are drop-in replacements that
- emulate primitive references for common cases of \emph{configuration
- options} for type \verb|bool|/\verb|int|/\verb|string| (see structure \verb|Config| and \verb|Attrib.config_bool| etc.), and lists of theorems (see functor
- \verb|Named_Thms|).
+ \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 formerly writable flags.
+
+ \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 \isa{stdout} or \isa{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 Keep components with local state information
- \emph{re-entrant}. Instead of poking initial values into (private)
- global references, create a new state record on each invocation, and
- pass that 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 sees its
- own copy. Occasionally, one might even return to plain functional
- updates on non-mutable record values here.
+ \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}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML}{File.tmp\_path}\verb|File.tmp_path: Path.T -> Path.T| \\
+ \indexdef{}{ML}{serial\_string}\verb|serial_string: unit -> string| \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item \verb|File.tmp_path|~\isa{path} relocates the base
+ component of \isa{path} into the unique temporary directory of
+ the running Isabelle/ML process.
+
+ \item \verb|serial_string|~\isa{{\isacharparenleft}{\isacharparenright}} creates a new serial number
+ that is unique over the runtime of the Isabelle/ML process.
- \item Isolate process configuration flags. The main legitimate
- application of global references is to configure the whole process
- in a certain way, essentially affecting all threads. A typical
- example is the \verb|show_types| flag, which tells the pretty printer
- to output explicit type information for terms. Such flags usually
- do not affect the functionality of the core system, but only the
- view being presented to the user.
-
- Occasionally, such global process flags are treated like implicit
- arguments to certain operations, by using the \verb|setmp_CRITICAL| combinator
- for safe temporary assignment. Its traditional purpose was to
- ensure proper recovery of the original value when exceptions are
- raised in the body, now the functionality is extended to enter the
- \emph{critical section} (with its usual potential of degrading
- parallelism).
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isatagmlex
+%
+\begin{isamarkuptext}%
+The following example shows how to create unique
+ temporary file names.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlex
+{\isafoldmlex}%
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ val\ tmp{\isadigit{1}}\ {\isacharequal}\ File{\isachardot}tmp{\isacharunderscore}path\ {\isacharparenleft}Path{\isachardot}basic\ {\isacharparenleft}{\isachardoublequote}foo{\isachardoublequote}\ {\isacharcircum}\ serial{\isacharunderscore}string\ {\isacharparenleft}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharsemicolon}\isanewline
+\ \ val\ tmp{\isadigit{2}}\ {\isacharequal}\ File{\isachardot}tmp{\isacharunderscore}path\ {\isacharparenleft}Path{\isachardot}basic\ {\isacharparenleft}{\isachardoublequote}foo{\isachardoublequote}\ {\isacharcircum}\ serial{\isacharunderscore}string\ {\isacharparenleft}{\isacharparenright}{\isacharparenright}{\isacharparenright}{\isacharsemicolon}\isanewline
+\ \ %
+\isaantiq
+assert%
+\endisaantiq
+\ {\isacharparenleft}tmp{\isadigit{1}}\ {\isacharless}{\isachargreater}\ tmp{\isadigit{2}}{\isacharparenright}{\isacharsemicolon}\isanewline
+{\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isamarkupsubsection{Explicit synchronization%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+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!
- Note that recovery of plain value passing semantics via \verb|setmp_CRITICAL|~\isa{ref\ value} assumes that this \isa{ref} is
- exclusively manipulated within the critical section. In particular,
- any persistent global assignment of \isa{ref\ {\isacharcolon}{\isacharequal}\ value} needs to
- be marked critical as well, to prevent intruding another threads
- local view, and a lost-update in the global scope, too.
+ \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.
- \end{itemize}
+ Such centralized locking is convenient, because it prevents
+ deadlocks by construction.
- Recall that in an open ``LCF-style'' system like Isabelle/Isar, the
- user participates in constructing the overall environment. This
- means that state-based facilities offered by one component will
- require special caution later on. So minimizing critical elements,
- by staying within the plain value-oriented view relative to theory
- or proof contexts most of the time, will also reduce the chance of
- mishaps occurring to end-users.%
+ \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!%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -277,215 +2190,49 @@
\begin{mldecls}
\indexdef{}{ML}{NAMED\_CRITICAL}\verb|NAMED_CRITICAL: string -> (unit -> 'a) -> 'a| \\
\indexdef{}{ML}{CRITICAL}\verb|CRITICAL: (unit -> 'a) -> 'a| \\
- \indexdef{}{ML}{setmp\_CRITICAL}\verb|setmp_CRITICAL: 'a Unsynchronized.ref -> 'a -> ('b -> 'c) -> 'b -> 'c| \\
+ \end{mldecls}
+ \begin{mldecls}
+ \indexdef{}{ML type}{Synchronized.var}\verb|type 'a Synchronized.var| \\
+ \indexdef{}{ML}{Synchronized.var}\verb|Synchronized.var: string -> 'a -> 'a Synchronized.var| \\
+ \indexdef{}{ML}{Synchronized.guarded\_access}\verb|Synchronized.guarded_access: 'a Synchronized.var ->|\isasep\isanewline%
+\verb| ('a -> ('b * 'a) option) -> 'b| \\
\end{mldecls}
\begin{description}
- \item \verb|NAMED_CRITICAL|~\isa{name\ f} evaluates \isa{f\ {\isacharparenleft}{\isacharparenright}}
- while staying within the critical section of Isabelle/Isar. No
- other thread may do so at the same time, but non-critical parallel
- execution will continue. The \isa{name} argument serves for
- diagnostic purposes and might help to spot sources of congestion.
+ \item \verb|NAMED_CRITICAL|~\isa{name\ e} evaluates \isa{e\ {\isacharparenleft}{\isacharparenright}}
+ 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 \isa{name} argument is used for tracing and might
+ help to spot sources of congestion.
+
+ Entering the critical section without contention is very fast, and
+ several basic system operations do so frequently. Each thread
+ should leave the critical section quickly, otherwise parallel
+ performance may degrade.
\item \verb|CRITICAL| is the same as \verb|NAMED_CRITICAL| with empty
name argument.
- \item \verb|setmp_CRITICAL|~\isa{ref\ value\ f\ x} evaluates \isa{f\ x}
- while staying within the critical section and having \isa{ref\ {\isacharcolon}{\isacharequal}\ value} assigned temporarily. This recovers a value-passing
- semantics involving global references, regardless of exceptions or
- concurrency.
+ \item Type \verb|'a Synchronized.var| represents synchronized
+ variables with state of type \verb|'a|.
- \end{description}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagmlref
-{\isafoldmlref}%
-%
-\isadelimmlref
-%
-\endisadelimmlref
-%
-\isamarkupchapter{Basic library functions%
-}
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Beyond the proposal of the SML/NJ basis library, Isabelle comes
- with its own library, from which selected parts are given here.
- These should encourage a study of the Isabelle sources,
- in particular files \emph{Pure/library.ML} and \emph{Pure/General/*.ML}.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsection{Linear transformations \label{sec:ML-linear-trans}%
-}
-\isamarkuptrue%
-%
-\isadelimmlref
-%
-\endisadelimmlref
-%
-\isatagmlref
-%
-\begin{isamarkuptext}%
-\begin{mldecls}
- \indexdef{}{ML}{op $\mid$$>$ }\verb|op |\verb,|,\verb|> : 'a * ('a -> 'b) -> 'b| \\
- \end{mldecls}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagmlref
-{\isafoldmlref}%
-%
-\isadelimmlref
-%
-\endisadelimmlref
-%
-\isadelimML
-%
-\endisadelimML
-%
-\isatagML
-%
-\endisatagML
-{\isafoldML}%
-%
-\isadelimML
-%
-\endisadelimML
-%
-\begin{isamarkuptext}%
-\noindent Many problems in functional programming can be thought of
- as linear transformations, i.e.~a caluclation starts with a
- particular value \verb|x : foo| which is then transformed
- by application of a function \verb|f : foo -> foo|,
- continued by an application of a function \verb|g : foo -> bar|,
- and so on. As a canoncial example, take functions enriching
- a theory by constant declararion and primitive definitions:
-
- \smallskip\begin{mldecls}
- \verb|Sign.declare_const: (binding * typ) * mixfix|\isasep\isanewline%
-\verb| -> theory -> term * theory| \\
- \verb|Thm.add_def: bool -> bool -> binding * term -> theory -> (string * thm) * theory|
- \end{mldecls}
-
- \noindent Written with naive application, an addition of constant
- \isa{bar} with type \isa{foo\ {\isasymRightarrow}\ foo} and
- a corresponding definition \isa{bar\ {\isasymequiv}\ {\isasymlambda}x{\isachardot}\ x} would look like:
-
- \smallskip\begin{mldecls}
- \verb|(fn (t, thy) => Thm.add_def false false|\isasep\isanewline%
-\verb| (Binding.name "bar_def", Logic.mk_equals (t, @{term "%x. x"})) thy)|\isasep\isanewline%
-\verb| (Sign.declare_const|\isasep\isanewline%
-\verb| ((Binding.name "bar", @{typ "foo => foo"}), NoSyn) thy)|
- \end{mldecls}
+ \item \verb|Synchronized.var|~\isa{name\ x} creates a synchronized
+ variable that is initialized with value \isa{x}. The \isa{name} is used for tracing.
- \noindent With increasing numbers of applications, this code gets quite
- unreadable. Further, it is unintuitive that things are
- written down in the opposite order as they actually ``happen''.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-\noindent At this stage, Isabelle offers some combinators which allow
- for more convenient notation, most notably reverse application:
-
- \smallskip\begin{mldecls}
-\verb|thy|\isasep\isanewline%
-\verb||\verb,|,\verb|> Sign.declare_const ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
-\verb||\verb,|,\verb|> (fn (t, thy) => thy|\isasep\isanewline%
-\verb||\verb,|,\verb|> Thm.add_def false false|\isasep\isanewline%
-\verb| (Binding.name "bar_def", Logic.mk_equals (t, @{term "%x. x"})))|
- \end{mldecls}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimmlref
-%
-\endisadelimmlref
-%
-\isatagmlref
-%
-\begin{isamarkuptext}%
-\begin{mldecls}
- \indexdef{}{ML}{op $\mid$-$>$ }\verb|op |\verb,|,\verb|-> : ('c * 'a) * ('c -> 'a -> 'b) -> 'b| \\
- \indexdef{}{ML}{op $\mid$$>$$>$ }\verb|op |\verb,|,\verb|>> : ('a * 'c) * ('a -> 'b) -> 'b * 'c| \\
- \indexdef{}{ML}{op $\mid$$\mid$$>$ }\verb|op |\verb,|,\verb||\verb,|,\verb|> : ('c * 'a) * ('a -> 'b) -> 'c * 'b| \\
- \indexdef{}{ML}{op $\mid$$\mid$$>$$>$ }\verb|op |\verb,|,\verb||\verb,|,\verb|>> : ('c * 'a) * ('a -> 'd * 'b) -> ('c * 'd) * 'b| \\
- \end{mldecls}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagmlref
-{\isafoldmlref}%
-%
-\isadelimmlref
-%
-\endisadelimmlref
-%
-\begin{isamarkuptext}%
-\noindent Usually, functions involving theory updates also return
- side results; to use these conveniently, yet another
- set of combinators is at hand, most notably \verb|op |\verb,|,\verb|->|
- which allows curried access to side results:
+ \item \verb|Synchronized.guarded_access|~\isa{var\ f} lets the
+ function \isa{f} operate within a critical section on the state
+ \isa{x} as follows: if \isa{f\ x} produces \verb|NONE|, we
+ continue to wait on the internal condition variable, expecting that
+ some other thread will eventually change the content in a suitable
+ manner; if \isa{f\ x} produces \verb|SOME|~\isa{{\isacharparenleft}y{\isacharcomma}\ x{\isacharprime}{\isacharparenright}} we
+ are satisfied and assign the new state value \isa{x{\isacharprime}}, broadcast
+ a signal to all waiting threads on the associated condition
+ variable, and return the result \isa{y}.
- \smallskip\begin{mldecls}
-\verb|thy|\isasep\isanewline%
-\verb||\verb,|,\verb|> Sign.declare_const ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
-\verb||\verb,|,\verb|-> (fn t => Thm.add_def false false|\isasep\isanewline%
-\verb| (Binding.name "bar_def", Logic.mk_equals (t, @{term "%x. x"})))|\isasep\isanewline%
-
- \end{mldecls}
-
- \noindent \verb|op |\verb,|,\verb|>>| allows for processing side results on their own:
-
- \smallskip\begin{mldecls}
-\verb|thy|\isasep\isanewline%
-\verb||\verb,|,\verb|> Sign.declare_const ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
-\verb||\verb,|,\verb|>> (fn t => Logic.mk_equals (t, @{term "%x. x"}))|\isasep\isanewline%
-\verb||\verb,|,\verb|-> (fn def => Thm.add_def false false (Binding.name "bar_def", def))|\isasep\isanewline%
-
- \end{mldecls}
-
- \noindent Orthogonally, \verb|op |\verb,|,\verb||\verb,|,\verb|>| applies a transformation
- in the presence of side results which are left unchanged:
+ \end{description}
- \smallskip\begin{mldecls}
-\verb|thy|\isasep\isanewline%
-\verb||\verb,|,\verb|> Sign.declare_const ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
-\verb||\verb,|,\verb||\verb,|,\verb|> Sign.add_path "foobar"|\isasep\isanewline%
-\verb||\verb,|,\verb|-> (fn t => Thm.add_def false false|\isasep\isanewline%
-\verb| (Binding.name "bar_def", Logic.mk_equals (t, @{term "%x. x"})))|\isasep\isanewline%
-\verb||\verb,|,\verb||\verb,|,\verb|> Sign.restore_naming thy|\isasep\isanewline%
-
- \end{mldecls}
-
- \noindent \verb|op |\verb,|,\verb||\verb,|,\verb|>>| accumulates side results:
-
- \smallskip\begin{mldecls}
-\verb|thy|\isasep\isanewline%
-\verb||\verb,|,\verb|> Sign.declare_const ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
-\verb||\verb,|,\verb||\verb,|,\verb|>> Sign.declare_const ((Binding.name "foobar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
-\verb||\verb,|,\verb|-> (fn (t1, t2) => Thm.add_def false false|\isasep\isanewline%
-\verb| (Binding.name "bar_def", Logic.mk_equals (t1, t2)))|\isasep\isanewline%
-
- \end{mldecls}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimmlref
-%
-\endisadelimmlref
-%
-\isatagmlref
-%
-\begin{isamarkuptext}%
-\begin{mldecls}
- \indexdef{}{ML}{fold}\verb|fold: ('a -> 'b -> 'b) -> 'a list -> 'b -> 'b| \\
- \indexdef{}{ML}{fold\_map}\verb|fold_map: ('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b| \\
- \end{mldecls}%
+ There are some further variants of the general \verb|Synchronized.guarded_access| combinator, see \hyperlink{file.~~/src/Pure/Concurrent/synchronized.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}Concurrent{\isacharslash}synchronized{\isachardot}ML}}}} for details.%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -496,283 +2243,68 @@
%
\endisadelimmlref
%
-\begin{isamarkuptext}%
-\noindent This principles naturally lift to \emph{lists} using
- the \verb|fold| and \verb|fold_map| combinators.
- The first lifts a single function
- \begin{quote}\footnotesize
- \verb|f : 'a -> 'b -> 'b| to \verb|'a list -> 'b -> 'b|
- \end{quote}
- such that
- \begin{quote}\footnotesize
- \verb|y |\verb,|,\verb|> fold f [x1, x2, ..., x_n]| \\
- \hspace*{2ex}\isa{{\isasymleadsto}} \verb|y |\verb,|,\verb|> f x1 |\verb,|,\verb|> f x2 |\verb,|,\verb|> ... |\verb,|,\verb|> f x_n|
- \end{quote}
- \noindent The second accumulates side results in a list by lifting
- a single function
- \begin{quote}\footnotesize
- \verb|f : 'a -> 'b -> 'c * 'b| to \verb|'a list -> 'b -> 'c list * 'b|
- \end{quote}
- such that
- \begin{quote}\footnotesize
- \verb|y |\verb,|,\verb|> fold_map f [x1, x2, ..., x_n]| \\
- \hspace*{2ex}\isa{{\isasymleadsto}} \verb|y |\verb,|,\verb|> f x1 |\verb,|,\verb||\verb,|,\verb|>> f x2 |\verb,|,\verb||\verb,|,\verb|>> ... |\verb,|,\verb||\verb,|,\verb|>> f x_n| \\
- \hspace*{6ex}\verb||\verb,|,\verb||\verb,|,\verb|> (fn ((z1, z2), ..., z_n) => [z1, z2, ..., z_n])|
- \end{quote}
-
- \noindent Example:
-
- \smallskip\begin{mldecls}
-\verb|let|\isasep\isanewline%
-\verb| val consts = ["foo", "bar"];|\isasep\isanewline%
-\verb|in|\isasep\isanewline%
-\verb| thy|\isasep\isanewline%
-\verb| |\verb,|,\verb|> fold_map (fn const => Sign.declare_const|\isasep\isanewline%
-\verb| ((Binding.name const, @{typ "foo => foo"}), NoSyn)) consts|\isasep\isanewline%
-\verb| |\verb,|,\verb|>> map (fn t => Logic.mk_equals (t, @{term "%x. x"}))|\isasep\isanewline%
-\verb| |\verb,|,\verb|-> (fn defs => fold_map (fn def =>|\isasep\isanewline%
-\verb| Thm.add_def false false (Binding.empty, def)) defs)|\isasep\isanewline%
-\verb|end|
- \end{mldecls}%
-\end{isamarkuptext}%
-\isamarkuptrue%
+\isadelimmlex
%
-\isadelimmlref
-%
-\endisadelimmlref
-%
-\isatagmlref
+\endisadelimmlex
%
-\begin{isamarkuptext}%
-\begin{mldecls}
- \indexdef{}{ML}{op \#$>$ }\verb|op #> : ('a -> 'b) * ('b -> 'c) -> 'a -> 'c| \\
- \indexdef{}{ML}{op \#-$>$ }\verb|op #-> : ('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd| \\
- \indexdef{}{ML}{op \#$>$$>$ }\verb|op #>> : ('a -> 'c * 'b) * ('c -> 'd) -> 'a -> 'd * 'b| \\
- \indexdef{}{ML}{op \#\#$>$ }\verb|op ##> : ('a -> 'c * 'b) * ('b -> 'd) -> 'a -> 'c * 'd| \\
- \indexdef{}{ML}{op \#\#$>$$>$ }\verb|op ##>> : ('a -> 'c * 'b) * ('b -> 'e * 'd) -> 'a -> ('c * 'e) * 'd| \\
- \end{mldecls}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagmlref
-{\isafoldmlref}%
-%
-\isadelimmlref
-%
-\endisadelimmlref
+\isatagmlex
%
\begin{isamarkuptext}%
-\noindent All those linear combinators also exist in higher-order
- variants which do not expect a value on the left hand side
- but a function.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isadelimmlref
-%
-\endisadelimmlref
-%
-\isatagmlref
-%
-\begin{isamarkuptext}%
-\begin{mldecls}
- \indexdef{}{ML}{op ` }\verb|op ` : ('b -> 'a) -> 'b -> 'a * 'b| \\
- \indexdef{}{ML}{tap}\verb|tap: ('b -> 'a) -> 'b -> 'b| \\
- \end{mldecls}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\endisatagmlref
-{\isafoldmlref}%
-%
-\isadelimmlref
-%
-\endisadelimmlref
-%
-\begin{isamarkuptext}%
-\noindent These operators allow to ``query'' a context
- in a series of context transformations:
+See \hyperlink{file.~~/src/Pure/Concurrent/mailbox.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}Concurrent{\isacharslash}mailbox{\isachardot}ML}}}} how to
+ implement a mailbox as synchronized variable over a purely
+ functional queue.
- \smallskip\begin{mldecls}
-\verb|thy|\isasep\isanewline%
-\verb||\verb,|,\verb|> tap (fn _ => writeln "now adding constant")|\isasep\isanewline%
-\verb||\verb,|,\verb|> Sign.declare_const ((Binding.name "bar", @{typ "foo => foo"}), NoSyn)|\isasep\isanewline%
-\verb||\verb,|,\verb||\verb,|,\verb|>> `(fn thy => Sign.declared_const thy|\isasep\isanewline%
-\verb| (Sign.full_name thy (Binding.name "foobar")))|\isasep\isanewline%
-\verb||\verb,|,\verb|-> (fn (t, b) => if b then I|\isasep\isanewline%
-\verb| else Sign.declare_const|\isasep\isanewline%
-\verb| ((Binding.name "foobar", @{typ "foo => foo"}), NoSyn) #> snd)|\isasep\isanewline%
-
- \end{mldecls}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsection{Options and partiality%
-}
-\isamarkuptrue%
-%
-\isadelimmlref
-%
-\endisadelimmlref
-%
-\isatagmlref
-%
-\begin{isamarkuptext}%
-\begin{mldecls}
- \indexdef{}{ML}{is\_some}\verb|is_some: 'a option -> bool| \\
- \indexdef{}{ML}{is\_none}\verb|is_none: 'a option -> bool| \\
- \indexdef{}{ML}{the}\verb|the: 'a option -> 'a| \\
- \indexdef{}{ML}{these}\verb|these: 'a list option -> 'a list| \\
- \indexdef{}{ML}{the\_list}\verb|the_list: 'a option -> 'a list| \\
- \indexdef{}{ML}{the\_default}\verb|the_default: 'a -> 'a option -> 'a| \\
- \indexdef{}{ML}{try}\verb|try: ('a -> 'b) -> 'a -> 'b option| \\
- \indexdef{}{ML}{can}\verb|can: ('a -> 'b) -> 'a -> bool| \\
- \end{mldecls}%
+ \medskip The following example implements a counter that produces
+ positive integers that are unique over the runtime of the Isabelle
+ process:%
\end{isamarkuptext}%
\isamarkuptrue%
%
-\endisatagmlref
-{\isafoldmlref}%
-%
-\isadelimmlref
-%
-\endisadelimmlref
+\endisatagmlex
+{\isafoldmlex}%
%
-\begin{isamarkuptext}%
-Standard selector functions on \isa{option}s are provided. The
- \verb|try| and \verb|can| functions provide a convenient interface for
- handling exceptions -- both take as arguments a function \verb|f|
- together with a parameter \verb|x| and handle any exception during
- the evaluation of the application of \verb|f| to \verb|x|, either
- return a lifted result (\verb|NONE| on failure) or a boolean value
- (\verb|false| on failure).%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsection{Common data structures%
-}
-\isamarkuptrue%
-%
-\isamarkupsubsection{Lists (as set-like data structures)%
-}
-\isamarkuptrue%
+\isadelimmlex
%
-\begin{isamarkuptext}%
-\begin{mldecls}
- \indexdef{}{ML}{member}\verb|member: ('b * 'a -> bool) -> 'a list -> 'b -> bool| \\
- \indexdef{}{ML}{insert}\verb|insert: ('a * 'a -> bool) -> 'a -> 'a list -> 'a list| \\
- \indexdef{}{ML}{remove}\verb|remove: ('b * 'a -> bool) -> 'b -> 'a list -> 'a list| \\
- \indexdef{}{ML}{merge}\verb|merge: ('a * 'a -> bool) -> 'a list * 'a list -> 'a list| \\
- \end{mldecls}%
-\end{isamarkuptext}%
-\isamarkuptrue%
+\endisadelimmlex
%
-\begin{isamarkuptext}%
-Lists are often used as set-like data structures -- set-like in
- the sense that they support a notion of \verb|member|-ship,
- \verb|insert|-ing and \verb|remove|-ing, but are order-sensitive.
- This is convenient when implementing a history-like mechanism:
- \verb|insert| adds an element \emph{to the front} of a list,
- if not yet present; \verb|remove| removes \emph{all} occurences
- of a particular element. Correspondingly \verb|merge| implements a
- a merge on two lists suitable for merges of context data
- (\secref{sec:context-theory}).
-
- Functions are parametrized by an explicit equality function
- to accomplish overloaded equality; in most cases of monomorphic
- equality, writing \verb|op =| should suffice.%
-\end{isamarkuptext}%
-\isamarkuptrue%
+\isadelimML
%
-\isamarkupsubsection{Association lists%
-}
-\isamarkuptrue%
+\endisadelimML
%
-\begin{isamarkuptext}%
-\begin{mldecls}
- \indexdef{}{ML exception}{AList.DUP}\verb|exception AList.DUP| \\
- \indexdef{}{ML}{AList.lookup}\verb|AList.lookup: ('a * 'b -> bool) -> ('b * 'c) list -> 'a -> 'c option| \\
- \indexdef{}{ML}{AList.defined}\verb|AList.defined: ('a * 'b -> bool) -> ('b * 'c) list -> 'a -> bool| \\
- \indexdef{}{ML}{AList.update}\verb|AList.update: ('a * 'a -> bool) -> ('a * 'b) -> ('a * 'b) list -> ('a * 'b) list| \\
- \indexdef{}{ML}{AList.default}\verb|AList.default: ('a * 'a -> bool) -> ('a * 'b) -> ('a * 'b) list -> ('a * 'b) list| \\
- \indexdef{}{ML}{AList.delete}\verb|AList.delete: ('a * 'b -> bool) -> 'a -> ('b * 'c) list -> ('b * 'c) list| \\
- \indexdef{}{ML}{AList.map\_entry}\verb|AList.map_entry: ('a * 'b -> bool) -> 'a|\isasep\isanewline%
-\verb| -> ('c -> 'c) -> ('b * 'c) list -> ('b * 'c) list| \\
- \indexdef{}{ML}{AList.map\_default}\verb|AList.map_default: ('a * 'a -> bool) -> 'a * 'b -> ('b -> 'b)|\isasep\isanewline%
-\verb| -> ('a * 'b) list -> ('a * 'b) list| \\
- \indexdef{}{ML}{AList.join}\verb|AList.join: ('a * 'a -> bool) -> ('a -> 'b * 'b -> 'b) (*exception DUP*)|\isasep\isanewline%
-\verb| -> ('a * 'b) list * ('a * 'b) list -> ('a * 'b) list (*exception AList.DUP*)| \\
- \indexdef{}{ML}{AList.merge}\verb|AList.merge: ('a * 'a -> bool) -> ('b * 'b -> bool)|\isasep\isanewline%
-\verb| -> ('a * 'b) list * ('a * 'b) list -> ('a * 'b) list (*exception AList.DUP*)|
- \end{mldecls}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\begin{isamarkuptext}%
-Association lists can be seens as an extension of set-like lists:
- on the one hand, they may be used to implement finite mappings,
- on the other hand, they remain order-sensitive and allow for
- multiple key-value-pair with the same key: \verb|AList.lookup|
- returns the \emph{first} value corresponding to a particular
- key, if present. \verb|AList.update| updates
- the \emph{first} occurence of a particular key; if no such
- key exists yet, the key-value-pair is added \emph{to the front}.
- \verb|AList.delete| only deletes the \emph{first} occurence of a key.
- \verb|AList.merge| provides an operation suitable for merges of context data
- (\secref{sec:context-theory}), where an equality parameter on
- values determines whether a merge should be considered a conflict.
- A slightly generalized operation if implementend by the \verb|AList.join|
- function which allows for explicit conflict resolution.%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
-\isamarkupsubsection{Tables%
-}
-\isamarkuptrue%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ local\isanewline
+\ \ \ \ val\ counter\ {\isacharequal}\ Synchronized{\isachardot}var\ {\isachardoublequote}counter{\isachardoublequote}\ {\isadigit{0}}{\isacharsemicolon}\isanewline
+\ \ in\isanewline
+\ \ \ \ fun\ next\ {\isacharparenleft}{\isacharparenright}\ {\isacharequal}\isanewline
+\ \ \ \ \ \ Synchronized{\isachardot}guarded{\isacharunderscore}access\ counter\isanewline
+\ \ \ \ \ \ \ \ {\isacharparenleft}fn\ i\ {\isacharequal}{\isachargreater}\isanewline
+\ \ \ \ \ \ \ \ \ \ let\ val\ j\ {\isacharequal}\ i\ {\isacharplus}\ {\isadigit{1}}\isanewline
+\ \ \ \ \ \ \ \ \ \ in\ SOME\ {\isacharparenleft}j{\isacharcomma}\ j{\isacharparenright}\ end{\isacharparenright}{\isacharsemicolon}\isanewline
+\ \ end{\isacharsemicolon}\isanewline
+{\isacharverbatimclose}\isanewline
+\isanewline
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ val\ a\ {\isacharequal}\ next\ {\isacharparenleft}{\isacharparenright}{\isacharsemicolon}\isanewline
+\ \ val\ b\ {\isacharequal}\ next\ {\isacharparenleft}{\isacharparenright}{\isacharsemicolon}\isanewline
+\ \ %
+\isaantiq
+assert%
+\endisaantiq
+\ {\isacharparenleft}a\ {\isacharless}{\isachargreater}\ b{\isacharparenright}{\isacharsemicolon}\isanewline
+{\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
%
-\begin{isamarkuptext}%
-\begin{mldecls}
- \indexdef{}{ML type}{'a Symtab.table}\verb|type 'a Symtab.table| \\
- \indexdef{}{ML exception}{Symtab.DUP}\verb|exception Symtab.DUP of string| \\
- \indexdef{}{ML exception}{Symtab.SAME}\verb|exception Symtab.SAME| \\
- \indexdef{}{ML exception}{Symtab.UNDEF}\verb|exception Symtab.UNDEF of string| \\
- \indexdef{}{ML}{Symtab.empty}\verb|Symtab.empty: 'a Symtab.table| \\
- \indexdef{}{ML}{Symtab.lookup}\verb|Symtab.lookup: 'a Symtab.table -> string -> 'a option| \\
- \indexdef{}{ML}{Symtab.defined}\verb|Symtab.defined: 'a Symtab.table -> string -> bool| \\
- \indexdef{}{ML}{Symtab.update}\verb|Symtab.update: (string * 'a) -> 'a Symtab.table -> 'a Symtab.table| \\
- \indexdef{}{ML}{Symtab.default}\verb|Symtab.default: string * 'a -> 'a Symtab.table -> 'a Symtab.table| \\
- \indexdef{}{ML}{Symtab.delete}\verb|Symtab.delete: string|\isasep\isanewline%
-\verb| -> 'a Symtab.table -> 'a Symtab.table (*exception Symtab.UNDEF*)| \\
- \indexdef{}{ML}{Symtab.map\_entry}\verb|Symtab.map_entry: string -> ('a -> 'a)|\isasep\isanewline%
-\verb| -> 'a Symtab.table -> 'a Symtab.table| \\
- \indexdef{}{ML}{Symtab.map\_default}\verb|Symtab.map_default: (string * 'a) -> ('a -> 'a)|\isasep\isanewline%
-\verb| -> 'a Symtab.table -> 'a Symtab.table| \\
- \indexdef{}{ML}{Symtab.join}\verb|Symtab.join: (string -> 'a * 'a -> 'a) (*exception Symtab.DUP/Symtab.SAME*)|\isasep\isanewline%
-\verb| -> 'a Symtab.table * 'a Symtab.table|\isasep\isanewline%
-\verb| -> 'a Symtab.table (*exception Symtab.DUP*)| \\
- \indexdef{}{ML}{Symtab.merge}\verb|Symtab.merge: ('a * 'a -> bool)|\isasep\isanewline%
-\verb| -> 'a Symtab.table * 'a Symtab.table|\isasep\isanewline%
-\verb| -> 'a Symtab.table (*exception Symtab.DUP*)|
- \end{mldecls}%
-\end{isamarkuptext}%
-\isamarkuptrue%
+\isadelimML
%
-\begin{isamarkuptext}%
-Tables are an efficient representation of finite mappings without
- any notion of order; due to their efficiency they should be used
- whenever such pure finite mappings are neccessary.
-
- The key type of tables must be given explicitly by instantiating
- the \verb|Table| functor which takes the key type
- together with its \verb|order|; for convience, we restrict
- here to the \verb|Symtab| instance with \verb|string|
- as key type.
-
- Most table functions correspond to those of association lists.%
-\end{isamarkuptext}%
-\isamarkuptrue%
+\endisadelimML
+\isanewline
%
\isadelimtheory
+\isanewline
%
\endisadelimtheory
%
--- a/doc-src/IsarImplementation/Thy/document/Prelim.tex Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/Prelim.tex Sun Oct 24 15:11:24 2010 -0700
@@ -159,13 +159,13 @@
\medskip There is a separate notion of \emph{theory reference} for
maintaining a live link to an evolving theory context: updates on
- drafts are propagated automatically. Dynamic updating stops after
- an explicit \isa{end} only.
+ drafts are propagated automatically. Dynamic updating stops when
+ the next \isa{checkpoint} is reached.
Derived entities may store a theory reference in order to indicate
- the context they belong to. This implicitly assumes monotonic
- reasoning, because the referenced context may become larger without
- further notice.%
+ the formal context from which they are derived. This implicitly
+ assumes monotonic reasoning, because the referenced context may
+ become larger without further notice.%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -178,11 +178,14 @@
\begin{isamarkuptext}%
\begin{mldecls}
\indexdef{}{ML type}{theory}\verb|type theory| \\
+ \indexdef{}{ML}{Theory.eq\_thy}\verb|Theory.eq_thy: theory * theory -> bool| \\
\indexdef{}{ML}{Theory.subthy}\verb|Theory.subthy: theory * theory -> bool| \\
\indexdef{}{ML}{Theory.checkpoint}\verb|Theory.checkpoint: theory -> theory| \\
\indexdef{}{ML}{Theory.copy}\verb|Theory.copy: theory -> theory| \\
\indexdef{}{ML}{Theory.merge}\verb|Theory.merge: theory * theory -> theory| \\
\indexdef{}{ML}{Theory.begin\_theory}\verb|Theory.begin_theory: string -> theory list -> theory| \\
+ \indexdef{}{ML}{Theory.parents\_of}\verb|Theory.parents_of: theory -> theory list| \\
+ \indexdef{}{ML}{Theory.ancestors\_of}\verb|Theory.ancestors_of: theory -> theory list| \\
\end{mldecls}
\begin{mldecls}
\indexdef{}{ML type}{theory\_ref}\verb|type theory_ref| \\
@@ -192,10 +195,15 @@
\begin{description}
- \item \verb|theory| represents theory contexts. This is
- essentially a linear type, with explicit runtime checking! Most
- internal theory operations destroy the original version, which then
- becomes ``stale''.
+ \item Type \verb|theory| represents theory contexts. This is
+ essentially a linear type, with explicit runtime checking.
+ Primitive theory operations destroy the original version, which then
+ becomes ``stale''. This can be prevented by explicit checkpointing,
+ which the system does at least at the boundary of toplevel command
+ transactions \secref{sec:isar-toplevel}.
+
+ \item \verb|Theory.eq_thy|~\isa{{\isacharparenleft}thy\isactrlsub {\isadigit{1}}{\isacharcomma}\ thy\isactrlsub {\isadigit{2}}{\isacharparenright}} check strict
+ identity of two theories.
\item \verb|Theory.subthy|~\isa{{\isacharparenleft}thy\isactrlsub {\isadigit{1}}{\isacharcomma}\ thy\isactrlsub {\isadigit{2}}{\isacharparenright}} compares theories
according to the intrinsic graph structure of the construction.
@@ -216,11 +224,17 @@
This version of ad-hoc theory merge fails for unrelated theories!
\item \verb|Theory.begin_theory|~\isa{name\ parents} constructs
- a new theory based on the given parents. This {\ML} function is
+ a new theory based on the given parents. This ML function is
normally not invoked directly.
- \item \verb|theory_ref| represents a sliding reference to an
- always valid theory; updates on the original are propagated
+ \item \verb|Theory.parents_of|~\isa{thy} returns the direct
+ ancestors of \isa{thy}.
+
+ \item \verb|Theory.ancestors_of|~\isa{thy} returns all
+ ancestors of \isa{thy} (not including \isa{thy} itself).
+
+ \item Type \verb|theory_ref| represents a sliding reference to
+ an always valid theory; updates on the original are propagated
automatically.
\item \verb|Theory.deref|~\isa{thy{\isacharunderscore}ref} turns a \verb|theory_ref| into an \verb|theory| value. As the referenced
@@ -239,14 +253,54 @@
%
\endisadelimmlref
%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
+\isatagmlantiq
+%
+\begin{isamarkuptext}%
+\begin{matharray}{rcl}
+ \indexdef{}{ML antiquotation}{theory}\hypertarget{ML antiquotation.theory}{\hyperlink{ML antiquotation.theory}{\mbox{\isa{theory}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{theory\_ref}\hypertarget{ML antiquotation.theory-ref}{\hyperlink{ML antiquotation.theory-ref}{\mbox{\isa{theory{\isacharunderscore}ref}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \end{matharray}
+
+ \begin{rail}
+ ('theory' | 'theory\_ref') nameref?
+ ;
+ \end{rail}
+
+ \begin{description}
+
+ \item \isa{{\isacharat}{\isacharbraceleft}theory{\isacharbraceright}} refers to the background theory of the
+ current context --- as abstract value.
+
+ \item \isa{{\isacharat}{\isacharbraceleft}theory\ A{\isacharbraceright}} refers to an explicitly named ancestor
+ theory \isa{A} of the background theory of the current context
+ --- as abstract value.
+
+ \item \isa{{\isacharat}{\isacharbraceleft}theory{\isacharunderscore}ref{\isacharbraceright}} is similar to \isa{{\isacharat}{\isacharbraceleft}theory{\isacharbraceright}}, but
+ produces a \verb|theory_ref| via \verb|Theory.check_thy| as
+ explained above.
+
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlantiq
+{\isafoldmlantiq}%
+%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
\isamarkupsubsection{Proof context \label{sec:context-proof}%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
A proof context is a container for pure data with a
- back-reference to the theory it belongs to. The \isa{init}
- operation creates a proof context from a given theory.
+ back-reference to the theory from which it is derived. The \isa{init} operation creates a proof context from a given theory.
Modifications to draft theories are propagated to the proof context
as usual, but there is also an explicit \isa{transfer} operation
to force resynchronization with more substantial updates to the
@@ -289,9 +343,9 @@
\begin{description}
- \item \verb|Proof.context| represents proof contexts. Elements
- of this type are essentially pure values, with a sliding reference
- to the background theory.
+ \item Type \verb|Proof.context| represents proof contexts.
+ Elements of this type are essentially pure values, with a sliding
+ reference to the background theory.
\item \verb|ProofContext.init_global|~\isa{thy} produces a proof context
derived from \isa{thy}, initializing all data.
@@ -314,6 +368,37 @@
%
\endisadelimmlref
%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
+\isatagmlantiq
+%
+\begin{isamarkuptext}%
+\begin{matharray}{rcl}
+ \indexdef{}{ML antiquotation}{context}\hypertarget{ML antiquotation.context}{\hyperlink{ML antiquotation.context}{\mbox{\isa{context}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \end{matharray}
+
+ \begin{description}
+
+ \item \isa{{\isacharat}{\isacharbraceleft}context{\isacharbraceright}} 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}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlantiq
+{\isafoldmlantiq}%
+%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
\isamarkupsubsection{Generic contexts \label{sec:generic-context}%
}
\isamarkuptrue%
@@ -348,7 +433,7 @@
\begin{description}
- \item \verb|Context.generic| is the direct sum of \verb|theory| and \verb|Proof.context|, with the datatype
+ \item Type \verb|Context.generic| is the direct sum of \verb|theory| and \verb|Proof.context|, with the datatype
constructors \verb|Context.Theory| and \verb|Context.Proof|.
\item \verb|Context.theory_of|~\isa{context} always produces a
@@ -391,8 +476,9 @@
\end{tabular}
\medskip
- \noindent The \isa{empty} value acts as initial default for
- \emph{any} theory that does not declare actual data content; \isa{extend} is acts like a unitary version of \isa{merge}.
+ The \isa{empty} value acts as initial default for \emph{any}
+ theory that does not declare actual data content; \isa{extend}
+ is acts like a unitary version of \isa{merge}.
Implementing \isa{merge} can be tricky. The general idea is
that \isa{merge\ {\isacharparenleft}data\isactrlsub {\isadigit{1}}{\isacharcomma}\ data\isactrlsub {\isadigit{2}}{\isacharparenright}} inserts those parts of \isa{data\isactrlsub {\isadigit{2}}} into \isa{data\isactrlsub {\isadigit{1}}} that are not yet present, while
@@ -413,19 +499,24 @@
\end{tabular}
\medskip
- \noindent The \isa{init} operation is supposed to produce a pure
- value from the given background theory and should be somehow
+ The \isa{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 \isa{init}
- operation of \emph{all} theory data slots ever declared.
+ 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 \isa{init} operation for proof contexts is
predefined to select the current data value from the background
theory.
- \bigskip Any of these data declaration over type \isa{T} result
- in an ML structure with the following signature:
+ \bigskip Any of the above data declarations over type \isa{T}
+ result in an ML structure with the following signature:
\medskip
\begin{tabular}{ll}
@@ -435,12 +526,12 @@
\end{tabular}
\medskip
- \noindent These other operations provide exclusive access for the
- particular kind of context (theory, proof, or generic context).
- This interface fully 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.%
+ 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.%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -522,9 +613,9 @@
\endisadelimML
%
\begin{isamarkuptext}%
-\noindent The implementation uses private theory data
- internally, and only exposes an operation that involves explicit
- argument checking wrt.\ the given theory.%
+The implementation uses private theory data internally, and
+ only exposes an operation that involves explicit argument checking
+ wrt.\ the given theory.%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -545,13 +636,16 @@
\ \ \ \ val\ extend\ {\isacharequal}\ I{\isacharsemicolon}\isanewline
\ \ \ \ fun\ merge\ {\isacharparenleft}ts{\isadigit{1}}{\isacharcomma}\ ts{\isadigit{2}}{\isacharparenright}\ {\isacharequal}\isanewline
\ \ \ \ \ \ Ord{\isacharunderscore}List{\isachardot}union\ Term{\isacharunderscore}Ord{\isachardot}fast{\isacharunderscore}term{\isacharunderscore}ord\ ts{\isadigit{1}}\ ts{\isadigit{2}}{\isacharsemicolon}\isanewline
-\ \ {\isacharparenright}\isanewline
+\ \ {\isacharparenright}{\isacharsemicolon}\isanewline
\isanewline
\ \ val\ get\ {\isacharequal}\ Terms{\isachardot}get{\isacharsemicolon}\isanewline
\isanewline
\ \ fun\ add\ raw{\isacharunderscore}t\ thy\ {\isacharequal}\isanewline
-\ \ \ \ let\ val\ t\ {\isacharequal}\ Sign{\isachardot}cert{\isacharunderscore}term\ thy\ raw{\isacharunderscore}t\isanewline
-\ \ \ \ in\ Terms{\isachardot}map\ {\isacharparenleft}Ord{\isacharunderscore}List{\isachardot}insert\ Term{\isacharunderscore}Ord{\isachardot}fast{\isacharunderscore}term{\isacharunderscore}ord\ t{\isacharparenright}\ thy\ end{\isacharsemicolon}\isanewline
+\ \ \ \ let\isanewline
+\ \ \ \ \ \ val\ t\ {\isacharequal}\ Sign{\isachardot}cert{\isacharunderscore}term\ thy\ raw{\isacharunderscore}t{\isacharsemicolon}\isanewline
+\ \ \ \ in\isanewline
+\ \ \ \ \ \ Terms{\isachardot}map\ {\isacharparenleft}Ord{\isacharunderscore}List{\isachardot}insert\ Term{\isacharunderscore}Ord{\isachardot}fast{\isacharunderscore}term{\isacharunderscore}ord\ t{\isacharparenright}\ thy\isanewline
+\ \ \ \ end{\isacharsemicolon}\isanewline
\isanewline
\ \ end{\isacharsemicolon}\isanewline
{\isacharverbatimclose}%
@@ -563,17 +657,17 @@
\endisadelimML
%
\begin{isamarkuptext}%
-We use \verb|term Ord_List.T| for reasonably efficient
- representation of a set of terms: all operations are linear in the
- number of stored elements. Here we assume that our users do not
- care about the declaration order, since that data structure forces
- its own arrangement of elements.
+Type \verb|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 \verb|merge| operation joins the data slots of
the two constituents: \verb|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.)
+ exponential blowup. Plain list append etc.\ must never be used for
+ theory data merges!
\medskip Our intended invariant is achieved as follows:
\begin{enumerate}
@@ -594,7 +688,7 @@
type-inference via \verb|Syntax.check_term| (cf.\
\secref{sec:term-check}) is not necessarily monotonic wrt.\ the
background theory, since constraints of term constants can be
- strengthened by later declarations, for example.
+ 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
@@ -603,6 +697,333 @@
\end{isamarkuptext}%
\isamarkuptrue%
%
+\isamarkupsubsection{Configuration options \label{sec:config-options}%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+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 \hyperlink{attribute.show-types}{\mbox{\isa{show{\isacharunderscore}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:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+\isacommand{declare}\isamarkupfalse%
+\ {\isacharbrackleft}{\isacharbrackleft}show{\isacharunderscore}types\ {\isacharequal}\ false{\isacharbrackright}{\isacharbrackright}\isanewline
+\ \ %
+\isamarkupcmt{declaration within (local) theory context%
+}
+\isanewline
+\isanewline
+\isacommand{example{\isacharunderscore}proof}\isamarkupfalse%
+\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{note}\isamarkupfalse%
+\ {\isacharbrackleft}{\isacharbrackleft}show{\isacharunderscore}types\ {\isacharequal}\ true{\isacharbrackright}{\isacharbrackright}\isanewline
+\ \ \ \ %
+\isamarkupcmt{declaration within proof (forward mode)%
+}
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\ \ \isacommand{term}\isamarkupfalse%
+\ x\isanewline
+%
+\isadelimproof
+\isanewline
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}x\ {\isacharequal}\ x{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{using}\isamarkupfalse%
+\ {\isacharbrackleft}{\isacharbrackleft}show{\isacharunderscore}types\ {\isacharequal}\ false{\isacharbrackright}{\isacharbrackright}\isanewline
+\ \ \ \ \ \ %
+\isamarkupcmt{declaration within proof (backward mode)%
+}
+\isanewline
+\ \ \ \ \isacommand{{\isachardot}{\isachardot}}\isamarkupfalse%
+\isanewline
+\isacommand{qed}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\begin{isamarkuptext}%
+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.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML}{Config.get}\verb|Config.get: Proof.context -> 'a Config.T -> 'a| \\
+ \indexdef{}{ML}{Config.map}\verb|Config.map: 'a Config.T -> ('a -> 'a) -> Proof.context -> Proof.context| \\
+ \indexdef{}{ML}{Attrib.config\_bool}\verb|Attrib.config_bool: string -> (Context.generic -> bool) ->|\isasep\isanewline%
+\verb| bool Config.T * (theory -> theory)| \\
+ \indexdef{}{ML}{Attrib.config\_int}\verb|Attrib.config_int: string -> (Context.generic -> int) ->|\isasep\isanewline%
+\verb| int Config.T * (theory -> theory)| \\
+ \indexdef{}{ML}{Attrib.config\_string}\verb|Attrib.config_string: string -> (Context.generic -> string) ->|\isasep\isanewline%
+\verb| string Config.T * (theory -> theory)| \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item \verb|Config.get|~\isa{ctxt\ config} gets the value of
+ \isa{config} in the given context.
+
+ \item \verb|Config.map|~\isa{config\ f\ ctxt} updates the context
+ by updating the value of \isa{config}.
+
+ \item \isa{{\isacharparenleft}config{\isacharcomma}\ setup{\isacharparenright}\ {\isacharequal}}~\verb|Attrib.config_bool|~\isa{name\ default} creates a named configuration option of type
+ \verb|bool|, with the given \isa{default} depending on the
+ application context. The resulting \isa{config} can be used to
+ get/map its value in a given context. The \isa{setup} function
+ needs to be applied to the theory initially, in order to make
+ concrete declaration syntax available to the user.
+
+ \item \verb|Attrib.config_int| and \verb|Attrib.config_string| work
+ like \verb|Attrib.config_bool|, but for types \verb|int| and
+ \verb|string|, respectively.
+
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isatagmlex
+%
+\begin{isamarkuptext}%
+The following example shows how to declare and use a
+ Boolean configuration option called \isa{my{\isacharunderscore}flag} with constant
+ default value \verb|false|.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlex
+{\isafoldmlex}%
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ val\ {\isacharparenleft}my{\isacharunderscore}flag{\isacharcomma}\ my{\isacharunderscore}flag{\isacharunderscore}setup{\isacharparenright}\ {\isacharequal}\isanewline
+\ \ \ \ Attrib{\isachardot}config{\isacharunderscore}bool\ {\isachardoublequote}my{\isacharunderscore}flag{\isachardoublequote}\ {\isacharparenleft}K\ false{\isacharparenright}\isanewline
+{\isacharverbatimclose}\isanewline
+\isacommand{setup}\isamarkupfalse%
+\ my{\isacharunderscore}flag{\isacharunderscore}setup%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+Now the user can refer to \hyperlink{attribute.my-flag}{\mbox{\isa{my{\isacharunderscore}flag}}} in
+ declarations, while we can retrieve the current value from the
+ context via \verb|Config.get|.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML{\isacharunderscore}val}\isamarkupfalse%
+\ {\isacharverbatimopen}\ %
+\isaantiq
+assert%
+\endisaantiq
+\ {\isacharparenleft}Config{\isachardot}get\ %
+\isaantiq
+context%
+\endisaantiq
+\ my{\isacharunderscore}flag\ {\isacharequal}\ false{\isacharparenright}\ {\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+\isanewline
+\isanewline
+\isacommand{declare}\isamarkupfalse%
+\ {\isacharbrackleft}{\isacharbrackleft}my{\isacharunderscore}flag\ {\isacharequal}\ true{\isacharbrackright}{\isacharbrackright}\isanewline
+%
+\isadelimML
+\isanewline
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML{\isacharunderscore}val}\isamarkupfalse%
+\ {\isacharverbatimopen}\ %
+\isaantiq
+assert%
+\endisaantiq
+\ {\isacharparenleft}Config{\isachardot}get\ %
+\isaantiq
+context%
+\endisaantiq
+\ my{\isacharunderscore}flag\ {\isacharequal}\ true{\isacharparenright}\ {\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+\isanewline
+%
+\endisadelimML
+\isanewline
+\isacommand{example{\isacharunderscore}proof}\isamarkupfalse%
+\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{{\isacharbraceleft}}\isamarkupfalse%
+\isanewline
+\ \ \ \ \isacommand{note}\isamarkupfalse%
+\ {\isacharbrackleft}{\isacharbrackleft}my{\isacharunderscore}flag\ {\isacharequal}\ false{\isacharbrackright}{\isacharbrackright}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+%
+\isadelimML
+\ \ \ \ %
+\endisadelimML
+%
+\isatagML
+\isacommand{ML{\isacharunderscore}val}\isamarkupfalse%
+\ {\isacharverbatimopen}\ %
+\isaantiq
+assert%
+\endisaantiq
+\ {\isacharparenleft}Config{\isachardot}get\ %
+\isaantiq
+context%
+\endisaantiq
+\ my{\isacharunderscore}flag\ {\isacharequal}\ false{\isacharparenright}\ {\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+\isanewline
+%
+\endisadelimML
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{{\isacharbraceright}}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+%
+\isadelimML
+\ \ %
+\endisadelimML
+%
+\isatagML
+\isacommand{ML{\isacharunderscore}val}\isamarkupfalse%
+\ {\isacharverbatimopen}\ %
+\isaantiq
+assert%
+\endisaantiq
+\ {\isacharparenleft}Config{\isachardot}get\ %
+\isaantiq
+context%
+\endisaantiq
+\ my{\isacharunderscore}flag\ {\isacharequal}\ true{\isacharparenright}\ {\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+\isanewline
+%
+\endisadelimML
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{qed}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
\isamarkupsection{Names \label{sec:names}%
}
\isamarkuptrue%
@@ -635,7 +1056,7 @@
\end{isamarkuptext}%
\isamarkuptrue%
%
-\isamarkupsubsection{Strings of symbols%
+\isamarkupsubsection{Strings of symbols \label{sec:symbols}%
}
\isamarkuptrue%
%
@@ -669,11 +1090,11 @@
\end{enumerate}
- \noindent The \isa{ident} syntax for symbol names is \isa{letter\ {\isacharparenleft}letter\ {\isacharbar}\ digit{\isacharparenright}\isactrlsup {\isacharasterisk}}, where \isa{letter\ {\isacharequal}\ A{\isachardot}{\isachardot}Za{\isachardot}{\isachardot}z} and \isa{digit\ {\isacharequal}\ {\isadigit{0}}{\isachardot}{\isachardot}{\isadigit{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 \isa{ident} syntax for symbol names is \isa{letter\ {\isacharparenleft}letter\ {\isacharbar}\ digit{\isacharparenright}\isactrlsup {\isacharasterisk}}, where \isa{letter\ {\isacharequal}\ A{\isachardot}{\isachardot}Za{\isachardot}{\isachardot}z} and \isa{digit\ {\isacharequal}\ {\isadigit{0}}{\isachardot}{\isachardot}{\isadigit{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
@@ -718,22 +1139,22 @@
\begin{description}
- \item \verb|Symbol.symbol| represents individual Isabelle
+ \item Type \verb|Symbol.symbol| represents individual Isabelle
symbols.
\item \verb|Symbol.explode|~\isa{str} produces a symbol list
- from the packed form. This function supercedes \verb|String.explode| for virtually all purposes of manipulating text in
+ from the packed form. This function supersedes \verb|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 --- which is the most common case --- do not
- require extra memory in Poly/ML.}
+ be a singleton string do not require extra memory in Poly/ML.}
\item \verb|Symbol.is_letter|, \verb|Symbol.is_digit|, \verb|Symbol.is_quasi|, \verb|Symbol.is_blank| classify standard
symbols according to fixed syntactic conventions of Isabelle, cf.\
\cite{isabelle-isar-ref}.
- \item \verb|Symbol.sym| is a concrete datatype that represents
- the different kinds of symbols explicitly, with constructors \verb|Symbol.Char|, \verb|Symbol.Sym|, \verb|Symbol.UTF8|, \verb|Symbol.Ctrl|, \verb|Symbol.Raw|.
+ \item Type \verb|Symbol.sym| is a concrete datatype that
+ represents the different kinds of symbols explicitly, with
+ constructors \verb|Symbol.Char|, \verb|Symbol.Sym|, \verb|Symbol.UTF8|, \verb|Symbol.Ctrl|, \verb|Symbol.Raw|.
\item \verb|Symbol.decode| converts the string representation of a
symbol into the datatype version.
@@ -823,8 +1244,8 @@
\item \verb|Name.skolem|~\isa{name} produces a Skolem name by
adding two underscores.
- \item \verb|Name.context| represents the context of already used
- names; the initial value is \verb|Name.context|.
+ \item Type \verb|Name.context| represents the context of already
+ used names; the initial value is \verb|Name.context|.
\item \verb|Name.declare|~\isa{name} enters a used name into the
context.
@@ -853,6 +1274,102 @@
%
\endisadelimmlref
%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isatagmlex
+%
+\begin{isamarkuptext}%
+The following simple examples demonstrate how to produce
+ fresh names from the initial \verb|Name.context|.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlex
+{\isafoldmlex}%
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ val\ list{\isadigit{1}}\ {\isacharequal}\ Name{\isachardot}invents\ Name{\isachardot}context\ {\isachardoublequote}a{\isachardoublequote}\ {\isadigit{5}}{\isacharsemicolon}\isanewline
+\ \ %
+\isaantiq
+assert%
+\endisaantiq
+\ {\isacharparenleft}list{\isadigit{1}}\ {\isacharequal}\ {\isacharbrackleft}{\isachardoublequote}a{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}b{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}c{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}d{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}e{\isachardoublequote}{\isacharbrackright}{\isacharparenright}{\isacharsemicolon}\isanewline
+\isanewline
+\ \ val\ list{\isadigit{2}}\ {\isacharequal}\isanewline
+\ \ \ \ {\isacharhash}{\isadigit{1}}\ {\isacharparenleft}Name{\isachardot}variants\ {\isacharbrackleft}{\isachardoublequote}x{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}x{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}a{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}a{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}{\isacharprime}a{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}{\isacharprime}a{\isachardoublequote}{\isacharbrackright}\ Name{\isachardot}context{\isacharparenright}{\isacharsemicolon}\isanewline
+\ \ %
+\isaantiq
+assert%
+\endisaantiq
+\ {\isacharparenleft}list{\isadigit{2}}\ {\isacharequal}\ {\isacharbrackleft}{\isachardoublequote}x{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}xa{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}a{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}aa{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}{\isacharprime}a{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}{\isacharprime}aa{\isachardoublequote}{\isacharbrackright}{\isacharparenright}{\isacharsemicolon}\isanewline
+{\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+\medskip The same works reletively to the formal context as
+ follows.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+\isacommand{locale}\isamarkupfalse%
+\ ex\ {\isacharequal}\ \isakeyword{fixes}\ a\ b\ c\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\isanewline
+\isakeyword{begin}\isanewline
+%
+\isadelimML
+\isanewline
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ val\ names\ {\isacharequal}\ Variable{\isachardot}names{\isacharunderscore}of\ %
+\isaantiq
+context%
+\endisaantiq
+{\isacharsemicolon}\isanewline
+\isanewline
+\ \ val\ list{\isadigit{1}}\ {\isacharequal}\ Name{\isachardot}invents\ names\ {\isachardoublequote}a{\isachardoublequote}\ {\isadigit{5}}{\isacharsemicolon}\isanewline
+\ \ %
+\isaantiq
+assert%
+\endisaantiq
+\ {\isacharparenleft}list{\isadigit{1}}\ {\isacharequal}\ {\isacharbrackleft}{\isachardoublequote}d{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}e{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}f{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}g{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}h{\isachardoublequote}{\isacharbrackright}{\isacharparenright}{\isacharsemicolon}\isanewline
+\isanewline
+\ \ val\ list{\isadigit{2}}\ {\isacharequal}\isanewline
+\ \ \ \ {\isacharhash}{\isadigit{1}}\ {\isacharparenleft}Name{\isachardot}variants\ {\isacharbrackleft}{\isachardoublequote}x{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}x{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}a{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}a{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}{\isacharprime}a{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}{\isacharprime}a{\isachardoublequote}{\isacharbrackright}\ names{\isacharparenright}{\isacharsemicolon}\isanewline
+\ \ %
+\isaantiq
+assert%
+\endisaantiq
+\ {\isacharparenleft}list{\isadigit{2}}\ {\isacharequal}\ {\isacharbrackleft}{\isachardoublequote}x{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}xa{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}aa{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}ab{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}{\isacharprime}aa{\isachardoublequote}{\isacharcomma}\ {\isachardoublequote}{\isacharprime}ab{\isachardoublequote}{\isacharbrackright}{\isacharparenright}{\isacharsemicolon}\isanewline
+{\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+\isanewline
+%
+\endisadelimML
+\isanewline
+\isacommand{end}\isamarkupfalse%
+%
\isamarkupsubsection{Indexed names \label{sec:indexname}%
}
\isamarkuptrue%
@@ -901,15 +1418,14 @@
%
\begin{isamarkuptext}%
\begin{mldecls}
- \indexdef{}{ML type}{indexname}\verb|type indexname| \\
+ \indexdef{}{ML type}{indexname}\verb|type indexname = string * int| \\
\end{mldecls}
\begin{description}
- \item \verb|indexname| represents indexed names. This is an
- abbreviation for \verb|string * int|. The second component is
- usually non-negative, except for situations where \isa{{\isacharparenleft}x{\isacharcomma}\ {\isacharminus}{\isadigit{1}}{\isacharparenright}}
- is used to inject basic names into this type. Other negative
+ \item Type \verb|indexname| represents indexed names. This is
+ an abbreviation for \verb|string * int|. The second component
+ is usually non-negative, except for situations where \isa{{\isacharparenleft}x{\isacharcomma}\ {\isacharminus}{\isadigit{1}}{\isacharparenright}} is used to inject basic names into this type. Other negative
indexes should not be used.
\end{description}%
@@ -1017,13 +1533,12 @@
main equation of this ``chemical reaction'' when binding new
entities in a context is as follows:
- \smallskip
+ \medskip
\begin{tabular}{l}
\isa{binding\ {\isacharplus}\ naming\ {\isasymlongrightarrow}\ long\ name\ {\isacharplus}\ name\ space\ accesses}
\end{tabular}
- \smallskip
- \medskip As a general principle, there is a separate name space for
+ \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
@@ -1036,10 +1551,17 @@
constant \isa{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 are better suffixed in addition to the usual qualification,
- e.g.\ \isa{c{\isacharunderscore}type{\isachardot}intro} and \isa{c{\isacharunderscore}class{\isachardot}intro} for
- theorems related to type \isa{c} and class \isa{c},
- respectively.%
+ 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 \isa{c} & \isa{c{\isachardot}intro} \\
+ type \isa{c} & \isa{c{\isacharunderscore}type{\isachardot}intro} \\
+ class \isa{c} & \isa{c{\isacharunderscore}class{\isachardot}intro} \\
+ \end{tabular}%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -1078,13 +1600,14 @@
\begin{description}
- \item \verb|binding| represents the abstract concept of name
- bindings.
+ \item Type \verb|binding| represents the abstract concept of
+ name bindings.
\item \verb|Binding.empty| is the empty binding.
\item \verb|Binding.name|~\isa{name} produces a binding with base
- name \isa{name}.
+ name \isa{name}. Note that this lacks proper source position
+ information; see also the ML antiquotation \hyperlink{ML antiquotation.binding}{\mbox{\isa{binding}}}.
\item \verb|Binding.qualify|~\isa{mandatory\ name\ binding}
prefixes qualifier \isa{name} to \isa{binding}. The \isa{mandatory} flag tells if this name component always needs to be
@@ -1107,8 +1630,8 @@
representation for human-readable output, together with some formal
markup that might get used in GUI front-ends, for example.
- \item \verb|Name_Space.naming| represents the abstract concept of
- a naming policy.
+ \item Type \verb|Name_Space.naming| represents the abstract
+ concept of a naming policy.
\item \verb|Name_Space.default_naming| is the default naming policy.
In a theory context, this is usually augmented by a path prefix
@@ -1121,7 +1644,7 @@
name binding (usually a basic name) into the fully qualified
internal name, according to the given naming policy.
- \item \verb|Name_Space.T| represents name spaces.
+ \item Type \verb|Name_Space.T| represents name spaces.
\item \verb|Name_Space.empty|~\isa{kind} and \verb|Name_Space.merge|~\isa{{\isacharparenleft}space\isactrlisub {\isadigit{1}}{\isacharcomma}\ space\isactrlisub {\isadigit{2}}{\isacharparenright}} are the canonical operations for
maintaining name spaces according to theory data management
@@ -1161,6 +1684,111 @@
%
\endisadelimmlref
%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
+\isatagmlantiq
+%
+\begin{isamarkuptext}%
+\begin{matharray}{rcl}
+ \indexdef{}{ML antiquotation}{binding}\hypertarget{ML antiquotation.binding}{\hyperlink{ML antiquotation.binding}{\mbox{\isa{binding}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \end{matharray}
+
+ \begin{rail}
+ 'binding' name
+ ;
+ \end{rail}
+
+ \begin{description}
+
+ \item \isa{{\isacharat}{\isacharbraceleft}binding\ name{\isacharbraceright}} produces a binding with base name
+ \isa{name} and the source position taken from the concrete
+ syntax of this antiquotation. In many situations this is more
+ appropriate than the more basic \verb|Binding.name| function.
+
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlantiq
+{\isafoldmlantiq}%
+%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isatagmlex
+%
+\begin{isamarkuptext}%
+The following example yields the source position of some
+ concrete binding inlined into the text.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlex
+{\isafoldmlex}%
+%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\ Binding{\isachardot}pos{\isacharunderscore}of\ %
+\isaantiq
+binding\ here%
+\endisaantiq
+\ {\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+\medskip That position can be also printed in a message as
+ follows.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML{\isacharunderscore}command}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ writeln\isanewline
+\ \ \ \ {\isacharparenleft}{\isachardoublequote}Look\ here{\isachardoublequote}\ {\isacharcircum}\ Position{\isachardot}str{\isacharunderscore}of\ {\isacharparenleft}Binding{\isachardot}pos{\isacharunderscore}of\ %
+\isaantiq
+binding\ here%
+\endisaantiq
+{\isacharparenright}{\isacharparenright}\isanewline
+{\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\begin{isamarkuptext}%
+This illustrates a key virtue of formalized bindings as
+ opposed to raw specifications of base names: the system can use this
+ additional information for advanced feedback given to the user.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
\isadelimtheory
%
\endisadelimtheory
--- a/doc-src/IsarImplementation/Thy/document/Proof.tex Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/Proof.tex Sun Oct 24 15:11:24 2010 -0700
@@ -67,12 +67,87 @@
This twist of dependencies is also accommodated by the reverse
operation of exporting results from a context: a type variable
\isa{{\isasymalpha}} is considered fixed as long as it occurs in some fixed
- term variable of the context. For example, exporting \isa{x{\isacharcolon}\ term{\isacharcomma}\ {\isasymalpha}{\isacharcolon}\ type\ {\isasymturnstile}\ x\isactrlisub {\isasymalpha}\ {\isacharequal}\ x\isactrlisub {\isasymalpha}} produces in the first step
- \isa{x{\isacharcolon}\ term\ {\isasymturnstile}\ x\isactrlisub {\isasymalpha}\ {\isacharequal}\ x\isactrlisub {\isasymalpha}} for fixed \isa{{\isasymalpha}},
- and only in the second step \isa{{\isasymturnstile}\ {\isacharquery}x\isactrlisub {\isacharquery}\isactrlisub {\isasymalpha}\ {\isacharequal}\ {\isacharquery}x\isactrlisub {\isacharquery}\isactrlisub {\isasymalpha}} for schematic \isa{{\isacharquery}x} and \isa{{\isacharquery}{\isasymalpha}}.
-
- \medskip The Isabelle/Isar proof context manages the gory details of
- term vs.\ type variables, with high-level principles for moving the
+ term variable of the context. For example, exporting \isa{x{\isacharcolon}\ term{\isacharcomma}\ {\isasymalpha}{\isacharcolon}\ type\ {\isasymturnstile}\ x\isactrlisub {\isasymalpha}\ {\isasymequiv}\ x\isactrlisub {\isasymalpha}} produces in the first step \isa{x{\isacharcolon}\ term\ {\isasymturnstile}\ x\isactrlisub {\isasymalpha}\ {\isasymequiv}\ x\isactrlisub {\isasymalpha}} for fixed \isa{{\isasymalpha}}, and only in the second step
+ \isa{{\isasymturnstile}\ {\isacharquery}x\isactrlisub {\isacharquery}\isactrlisub {\isasymalpha}\ {\isasymequiv}\ {\isacharquery}x\isactrlisub {\isacharquery}\isactrlisub {\isasymalpha}} for schematic \isa{{\isacharquery}x} and \isa{{\isacharquery}{\isasymalpha}}.
+ The following Isar source text illustrates this scenario.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+\isacommand{example{\isacharunderscore}proof}\isamarkupfalse%
+\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{{\isacharbraceleft}}\isamarkupfalse%
+\isanewline
+\ \ \ \ \isacommand{fix}\isamarkupfalse%
+\ x\ \ %
+\isamarkupcmt{all potential occurrences of some \isa{x{\isacharcolon}{\isacharcolon}{\isasymtau}} are fixed%
+}
+\isanewline
+\ \ \ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse%
+\isanewline
+\ \ \ \ \ \ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}x{\isacharcolon}{\isacharcolon}{\isacharprime}a\ {\isasymequiv}\ x{\isachardoublequoteclose}\ \ %
+\isamarkupcmt{implicit type assigment by concrete occurrence%
+}
+\isanewline
+\ \ \ \ \ \ \ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}rule\ reflexive{\isacharparenright}\isanewline
+\ \ \ \ \isacommand{{\isacharbraceright}}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\ \ \ \ \isacommand{thm}\isamarkupfalse%
+\ this\ \ %
+\isamarkupcmt{result still with fixed type \isa{{\isacharprime}a}%
+}
+\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{{\isacharbraceright}}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\ \ \isacommand{thm}\isamarkupfalse%
+\ this\ \ %
+\isamarkupcmt{fully general result for arbitrary \isa{{\isacharquery}x{\isacharcolon}{\isacharcolon}{\isacharquery}{\isacharprime}a}%
+}
+\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{qed}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\begin{isamarkuptext}%
+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 \isa{add{\isacharunderscore}fixes} operation explictly declares fixed
@@ -176,9 +251,8 @@
\isatagmlex
%
\begin{isamarkuptext}%
-The following example (in theory \hyperlink{theory.Pure}{\mbox{\isa{Pure}}}) shows
- how to work with fixed term and type parameters work with
- type-inference.%
+The following example shows how to work with fixed term
+ and type parameters and with type-inference.%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -188,14 +262,8 @@
\isadelimmlex
%
\endisadelimmlex
-\isacommand{typedecl}\isamarkupfalse%
-\ foo\ \ %
-\isamarkupcmt{some basic type for testing purposes%
-}
-\isanewline
%
\isadelimML
-\isanewline
%
\endisadelimML
%
@@ -217,11 +285,11 @@
\ \ val\ t{\isadigit{1}}{\isacharprime}\ {\isacharequal}\ singleton\ {\isacharparenleft}Variable{\isachardot}polymorphic\ ctxt{\isadigit{1}}{\isacharparenright}\ t{\isadigit{1}}{\isacharsemicolon}\isanewline
\isanewline
\ \ {\isacharparenleft}{\isacharasterisk}term\ u\ enforces\ specific\ type\ assignment{\isacharasterisk}{\isacharparenright}\isanewline
-\ \ val\ u\ {\isacharequal}\ Syntax{\isachardot}read{\isacharunderscore}term\ ctxt{\isadigit{1}}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcolon}{\isacharcolon}foo{\isacharparenright}\ {\isasymequiv}\ y{\isachardoublequote}{\isacharsemicolon}\isanewline
+\ \ val\ u\ {\isacharequal}\ Syntax{\isachardot}read{\isacharunderscore}term\ ctxt{\isadigit{1}}\ {\isachardoublequote}{\isacharparenleft}x{\isacharcolon}{\isacharcolon}nat{\isacharparenright}\ {\isasymequiv}\ y{\isachardoublequote}{\isacharsemicolon}\isanewline
\isanewline
\ \ {\isacharparenleft}{\isacharasterisk}official\ declaration\ of\ u\ {\isacharminus}{\isacharminus}\ propagates\ constraints\ etc{\isachardot}{\isacharasterisk}{\isacharparenright}\isanewline
\ \ val\ ctxt{\isadigit{2}}\ {\isacharequal}\ ctxt{\isadigit{1}}\ {\isacharbar}{\isachargreater}\ Variable{\isachardot}declare{\isacharunderscore}term\ u{\isacharsemicolon}\isanewline
-\ \ val\ t{\isadigit{2}}\ {\isacharequal}\ Syntax{\isachardot}read{\isacharunderscore}term\ ctxt{\isadigit{2}}\ {\isachardoublequote}x{\isachardoublequote}{\isacharsemicolon}\ \ {\isacharparenleft}{\isacharasterisk}x{\isacharcolon}{\isacharcolon}foo\ is\ enforced{\isacharasterisk}{\isacharparenright}\isanewline
+\ \ val\ t{\isadigit{2}}\ {\isacharequal}\ Syntax{\isachardot}read{\isacharunderscore}term\ ctxt{\isadigit{2}}\ {\isachardoublequote}x{\isachardoublequote}{\isacharsemicolon}\ \ {\isacharparenleft}{\isacharasterisk}x{\isacharcolon}{\isacharcolon}nat\ is\ enforced{\isacharasterisk}{\isacharparenright}\isanewline
{\isacharverbatimclose}%
\endisatagML
{\isafoldML}%
@@ -263,31 +331,16 @@
\endisadelimML
%
\begin{isamarkuptext}%
-\noindent Subsequent ML code can now work with the invented
- names of \verb|x1|, \verb|x2|, \verb|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:%
+The following ML code can now work with the invented names of
+ \verb|x1|, \verb|x2|, \verb|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:%
\end{isamarkuptext}%
\isamarkuptrue%
-\isacommand{lemma}\isamarkupfalse%
-\ {\isachardoublequoteopen}PROP\ XXX{\isachardoublequoteclose}\isanewline
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-\isacommand{proof}\isamarkupfalse%
-\ {\isacharminus}%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
+\isacommand{example{\isacharunderscore}proof}\isamarkupfalse%
\isanewline
%
-\endisadelimproof
-%
\isadelimML
\ \ %
\endisadelimML
@@ -318,7 +371,7 @@
\endisadelimML
%
\begin{isamarkuptext}%
-\noindent In this situation \verb|Variable.add_fixes| and \verb|Variable.variant_fixes| are very similar, but identical name
+In this situation \verb|Variable.add_fixes| and \verb|Variable.variant_fixes| are very similar, but identical name
proposals given in a row are only accepted by the second version.%
\end{isamarkuptext}%
\isamarkuptrue%
@@ -398,10 +451,11 @@
\begin{description}
- \item \verb|Assumption.export| represents arbitrary export
- rules, which is any function of type \verb|bool -> cterm list -> thm -> thm|,
- where the \verb|bool| indicates goal mode, and the \verb|cterm list| the collection of assumptions to be discharged
- simultaneously.
+ \item Type \verb|Assumption.export| represents arbitrary export
+ rules, which is any function of type \verb|bool -> cterm list|\isasep\isanewline%
+\verb| -> thm -> thm|, where the \verb|bool| indicates goal mode,
+ and the \verb|cterm list| the collection of assumptions to be
+ discharged simultaneously.
\item \verb|Assumption.assume|~\isa{A} turns proposition \isa{A} into a primitive assumption \isa{A\ {\isasymturnstile}\ A{\isacharprime}}, where the
conclusion \isa{A{\isacharprime}} is in HHF normal form.
@@ -487,10 +541,9 @@
\endisadelimML
%
\begin{isamarkuptext}%
-\noindent 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.\
- \verb|Variable.export| or the combined \verb|ProofContext.export|).%
+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.\ \verb|Variable.export| or the combined \verb|ProofContext.export|).%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -531,8 +584,8 @@
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 \isa{{\isasymOBTAIN}} and
- \isa{{\isasymGUESS}} elements. Final results, which may not refer to
+ \cite{isabelle-isar-ref} for the user-level \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}} and
+ \hyperlink{command.guess}{\mbox{\isa{\isacommand{guess}}}} elements. Final results, which may not refer to
the parameters in the conclusion, need to exported explicitly into
the original context.%
\end{isamarkuptext}%
@@ -546,10 +599,17 @@
%
\begin{isamarkuptext}%
\begin{mldecls}
- \indexdef{}{ML}{SUBPROOF}\verb|SUBPROOF: (Subgoal.focus -> tactic) -> Proof.context -> int -> tactic| \\
- \indexdef{}{ML}{Subgoal.FOCUS}\verb|Subgoal.FOCUS: (Subgoal.focus -> tactic) -> Proof.context -> int -> tactic| \\
- \indexdef{}{ML}{Subgoal.FOCUS\_PREMS}\verb|Subgoal.FOCUS_PREMS: (Subgoal.focus -> tactic) -> Proof.context -> int -> tactic| \\
- \indexdef{}{ML}{Subgoal.FOCUS\_PARAMS}\verb|Subgoal.FOCUS_PARAMS: (Subgoal.focus -> tactic) -> Proof.context -> int -> tactic| \\
+ \indexdef{}{ML}{SUBPROOF}\verb|SUBPROOF: (Subgoal.focus -> tactic) ->|\isasep\isanewline%
+\verb| Proof.context -> int -> tactic| \\
+ \indexdef{}{ML}{Subgoal.FOCUS}\verb|Subgoal.FOCUS: (Subgoal.focus -> tactic) ->|\isasep\isanewline%
+\verb| Proof.context -> int -> tactic| \\
+ \indexdef{}{ML}{Subgoal.FOCUS\_PREMS}\verb|Subgoal.FOCUS_PREMS: (Subgoal.focus -> tactic) ->|\isasep\isanewline%
+\verb| Proof.context -> int -> tactic| \\
+ \indexdef{}{ML}{Subgoal.FOCUS\_PARAMS}\verb|Subgoal.FOCUS_PARAMS: (Subgoal.focus -> tactic) ->|\isasep\isanewline%
+\verb| Proof.context -> int -> tactic| \\
+ \indexdef{}{ML}{Subgoal.focus}\verb|Subgoal.focus: Proof.context -> int -> thm -> Subgoal.focus * thm| \\
+ \indexdef{}{ML}{Subgoal.focus\_prems}\verb|Subgoal.focus_prems: Proof.context -> int -> thm -> Subgoal.focus * thm| \\
+ \indexdef{}{ML}{Subgoal.focus\_params}\verb|Subgoal.focus_params: Proof.context -> int -> thm -> Subgoal.focus * thm| \\
\end{mldecls}
\begin{mldecls}
@@ -559,8 +619,8 @@
\verb| ({prems: thm list, context: Proof.context} -> tactic) -> thm list| \\
\end{mldecls}
\begin{mldecls}
- \indexdef{}{ML}{Obtain.result}\verb|Obtain.result: (Proof.context -> tactic) ->|\isasep\isanewline%
-\verb| thm list -> Proof.context -> ((string * cterm) list * thm list) * Proof.context| \\
+ \indexdef{}{ML}{Obtain.result}\verb|Obtain.result: (Proof.context -> tactic) -> thm list ->|\isasep\isanewline%
+\verb| Proof.context -> ((string * cterm) list * thm list) * Proof.context| \\
\end{mldecls}
\begin{description}
@@ -576,6 +636,11 @@
imported into the context, and the body tactic may introduce new
subgoals and schematic variables.
+ \item \verb|Subgoal.focus|, \verb|Subgoal.focus_prems|, \verb|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 \verb|Goal.prove|~\isa{ctxt\ xs\ As\ C\ tac} states goal \isa{C} in the context augmented by fixed variables \isa{xs} and
assumptions \isa{As}, and applies tactic \isa{tac} to solve
it. The latter may depend on the local assumptions being presented
@@ -602,7 +667,157 @@
%
\endisadelimmlref
%
+\isadelimmlex
+%
+\endisadelimmlex
+%
+\isatagmlex
+%
+\begin{isamarkuptext}%
+The following minimal example illustrates how to access
+ the focus information of a structured goal state.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlex
+{\isafoldmlex}%
+%
+\isadelimmlex
+%
+\endisadelimmlex
+\isacommand{example{\isacharunderscore}proof}\isamarkupfalse%
+\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{fix}\isamarkupfalse%
+\ A\ B\ C\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
+\isanewline
+\ \ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}{\isasymAnd}x{\isachardot}\ A\ x\ {\isasymLongrightarrow}\ B\ x\ {\isasymLongrightarrow}\ C\ x{\isachardoublequoteclose}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+%
+\isadelimML
+\ \ \ \ %
+\endisadelimML
+%
+\isatagML
+\isacommand{ML{\isacharunderscore}val}\isamarkupfalse%
+\isanewline
+\ \ \ \ {\isacharverbatimopen}\isanewline
+\ \ \ \ \ \ val\ {\isacharbraceleft}goal{\isacharcomma}\ context\ {\isacharequal}\ goal{\isacharunderscore}ctxt{\isacharcomma}\ {\isachardot}{\isachardot}{\isachardot}{\isacharbraceright}\ {\isacharequal}\ %
+\isaantiq
+Isar{\isachardot}goal%
+\endisaantiq
+{\isacharsemicolon}\isanewline
+\ \ \ \ \ \ val\ {\isacharparenleft}focus\ as\ {\isacharbraceleft}params{\isacharcomma}\ asms{\isacharcomma}\ concl{\isacharcomma}\ {\isachardot}{\isachardot}{\isachardot}{\isacharbraceright}{\isacharcomma}\ goal{\isacharprime}{\isacharparenright}\ {\isacharequal}\isanewline
+\ \ \ \ \ \ \ \ Subgoal{\isachardot}focus\ goal{\isacharunderscore}ctxt\ {\isadigit{1}}\ goal{\isacharsemicolon}\isanewline
+\ \ \ \ \ \ val\ {\isacharbrackleft}A{\isacharcomma}\ B{\isacharbrackright}\ {\isacharequal}\ {\isacharhash}prems\ focus{\isacharsemicolon}\isanewline
+\ \ \ \ \ \ val\ {\isacharbrackleft}{\isacharparenleft}{\isacharunderscore}{\isacharcomma}\ x{\isacharparenright}{\isacharbrackright}\ {\isacharequal}\ {\isacharhash}params\ focus{\isacharsemicolon}\isanewline
+\ \ \ \ {\isacharverbatimclose}%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+\isanewline
+%
+\endisadelimML
+%
+\isadelimproof
+\ \ \ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{oops}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\begin{isamarkuptext}%
+\medskip The next example demonstrates forward-elimination in
+ a local context, using \verb|Obtain.result|.%
+\end{isamarkuptext}%
+\isamarkuptrue%
+\isacommand{example{\isacharunderscore}proof}\isamarkupfalse%
+\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{assume}\isamarkupfalse%
+\ ex{\isacharcolon}\ {\isachardoublequoteopen}{\isasymexists}x{\isachardot}\ B\ x{\isachardoublequoteclose}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+%
+\isadelimML
+\isanewline
+\ \ %
+\endisadelimML
+%
+\isatagML
+\isacommand{ML{\isacharunderscore}prf}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ \ \ val\ ctxt{\isadigit{0}}\ {\isacharequal}\ %
+\isaantiq
+context%
+\endisaantiq
+{\isacharsemicolon}\isanewline
+\ \ \ \ val\ {\isacharparenleft}{\isacharparenleft}{\isacharbrackleft}{\isacharparenleft}{\isacharunderscore}{\isacharcomma}\ x{\isacharparenright}{\isacharbrackright}{\isacharcomma}\ {\isacharbrackleft}B{\isacharbrackright}{\isacharparenright}{\isacharcomma}\ ctxt{\isadigit{1}}{\isacharparenright}\ {\isacharequal}\ ctxt{\isadigit{0}}\isanewline
+\ \ \ \ \ \ {\isacharbar}{\isachargreater}\ Obtain{\isachardot}result\ {\isacharparenleft}fn\ {\isacharunderscore}\ {\isacharequal}{\isachargreater}\ etac\ %
+\isaantiq
+thm\ exE%
+\endisaantiq
+\ {\isadigit{1}}{\isacharparenright}\ {\isacharbrackleft}%
+\isaantiq
+thm\ ex%
+\endisaantiq
+{\isacharbrackright}{\isacharsemicolon}\isanewline
+\ \ {\isacharverbatimclose}\isanewline
+\ \ \isacommand{ML{\isacharunderscore}prf}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ \ \ singleton\ {\isacharparenleft}ProofContext{\isachardot}export\ ctxt{\isadigit{1}}\ ctxt{\isadigit{0}}{\isacharparenright}\ %
+\isaantiq
+thm\ refl%
+\endisaantiq
+{\isacharsemicolon}\isanewline
+\ \ {\isacharverbatimclose}\isanewline
+\ \ \isacommand{ML{\isacharunderscore}prf}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ \ \ ProofContext{\isachardot}export\ ctxt{\isadigit{1}}\ ctxt{\isadigit{0}}\ {\isacharbrackleft}Thm{\isachardot}reflexive\ x{\isacharbrackright}\isanewline
+\ \ \ \ \ \ handle\ ERROR\ msg\ {\isacharequal}{\isachargreater}\ {\isacharparenleft}warning\ msg{\isacharsemicolon}\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}{\isacharsemicolon}\isanewline
+\ \ {\isacharverbatimclose}\isanewline
+\isacommand{qed}\isamarkupfalse%
+%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+\isanewline
+%
+\endisadelimML
+%
\isadelimtheory
+\isanewline
%
\endisadelimtheory
%
@@ -613,9 +828,9 @@
{\isafoldtheory}%
%
\isadelimtheory
+\isanewline
%
\endisadelimtheory
-\isanewline
\end{isabellebody}%
%%% Local Variables:
%%% mode: latex
--- a/doc-src/IsarImplementation/Thy/document/Syntax.tex Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/Syntax.tex Sun Oct 24 15:11:24 2010 -0700
@@ -27,7 +27,46 @@
\end{isamarkuptext}%
\isamarkuptrue%
%
-\isamarkupsection{Parsing and printing%
+\isamarkupsection{Reading and pretty printing \label{sec:read-print}%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+FIXME%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML}{Syntax.read\_typ}\verb|Syntax.read_typ: Proof.context -> string -> typ| \\
+ \indexdef{}{ML}{Syntax.read\_term}\verb|Syntax.read_term: Proof.context -> string -> term| \\
+ \indexdef{}{ML}{Syntax.read\_prop}\verb|Syntax.read_prop: Proof.context -> string -> term| \\
+ \indexdef{}{ML}{Syntax.pretty\_typ}\verb|Syntax.pretty_typ: Proof.context -> typ -> Pretty.T| \\
+ \indexdef{}{ML}{Syntax.pretty\_term}\verb|Syntax.pretty_term: Proof.context -> term -> Pretty.T| \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item FIXME
+
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isamarkupsection{Parsing and unparsing \label{sec:parse-unparse}%
}
\isamarkuptrue%
%
@@ -36,10 +75,118 @@
\end{isamarkuptext}%
\isamarkuptrue%
%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML}{Syntax.parse\_typ}\verb|Syntax.parse_typ: Proof.context -> string -> typ| \\
+ \indexdef{}{ML}{Syntax.parse\_term}\verb|Syntax.parse_term: Proof.context -> string -> term| \\
+ \indexdef{}{ML}{Syntax.parse\_prop}\verb|Syntax.parse_prop: Proof.context -> string -> term| \\
+ \indexdef{}{ML}{Syntax.unparse\_typ}\verb|Syntax.unparse_typ: Proof.context -> typ -> Pretty.T| \\
+ \indexdef{}{ML}{Syntax.unparse\_term}\verb|Syntax.unparse_term: Proof.context -> term -> Pretty.T| \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item FIXME
+
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
\isamarkupsection{Checking and unchecking \label{sec:term-check}%
}
\isamarkuptrue%
%
+\begin{isamarkuptext}%
+FIXME%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isatagmlref
+%
+\begin{isamarkuptext}%
+\begin{mldecls}
+ \indexdef{}{ML}{Syntax.check\_typs}\verb|Syntax.check_typs: Proof.context -> typ list -> typ list| \\
+ \indexdef{}{ML}{Syntax.check\_terms}\verb|Syntax.check_terms: Proof.context -> term list -> term list| \\
+ \indexdef{}{ML}{Syntax.check\_props}\verb|Syntax.check_props: Proof.context -> term list -> term list| \\
+ \indexdef{}{ML}{Syntax.uncheck\_typs}\verb|Syntax.uncheck_typs: Proof.context -> typ list -> typ list| \\
+ \indexdef{}{ML}{Syntax.uncheck\_terms}\verb|Syntax.uncheck_terms: Proof.context -> term list -> term list| \\
+ \end{mldecls}
+
+ \begin{description}
+
+ \item FIXME
+
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlref
+{\isafoldmlref}%
+%
+\isadelimmlref
+%
+\endisadelimmlref
+%
+\isamarkupsection{Syntax translations%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+FIXME%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
+\isatagmlantiq
+%
+\begin{isamarkuptext}%
+\begin{matharray}{rcl}
+ \indexdef{}{ML antiquotation}{class\_syntax}\hypertarget{ML antiquotation.class-syntax}{\hyperlink{ML antiquotation.class-syntax}{\mbox{\isa{class{\isacharunderscore}syntax}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{type\_syntax}\hypertarget{ML antiquotation.type-syntax}{\hyperlink{ML antiquotation.type-syntax}{\mbox{\isa{type{\isacharunderscore}syntax}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{const\_syntax}\hypertarget{ML antiquotation.const-syntax}{\hyperlink{ML antiquotation.const-syntax}{\mbox{\isa{const{\isacharunderscore}syntax}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \indexdef{}{ML antiquotation}{syntax\_const}\hypertarget{ML antiquotation.syntax-const}{\hyperlink{ML antiquotation.syntax-const}{\mbox{\isa{syntax{\isacharunderscore}const}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
+ \end{matharray}
+
+ \begin{rail}
+ ('class_syntax' | 'type_syntax' | 'const_syntax' | 'syntax_const') name
+ ;
+ \end{rail}
+
+ \begin{description}
+
+ \item FIXME
+
+ \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagmlantiq
+{\isafoldmlantiq}%
+%
+\isadelimmlantiq
+%
+\endisadelimmlantiq
+%
\isadelimtheory
%
\endisadelimtheory
--- a/doc-src/IsarImplementation/Thy/document/Tactic.tex Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/Tactic.tex Sun Oct 24 15:11:24 2010 -0700
@@ -116,7 +116,7 @@
%
\endisadelimmlref
%
-\isamarkupsection{Tactics%
+\isamarkupsection{Tactics\label{sec:tactics}%
}
\isamarkuptrue%
%
@@ -220,13 +220,14 @@
\begin{description}
- \item \verb|tactic| represents tactics. The well-formedness
- conditions described above need to be observed. See also \hyperlink{file.~~/src/Pure/General/seq.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}General{\isacharslash}seq{\isachardot}ML}}}} for the underlying implementation of
- lazy sequences.
+ \item Type \verb|tactic| represents tactics. The
+ well-formedness conditions described above need to be observed. See
+ also \hyperlink{file.~~/src/Pure/General/seq.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}General{\isacharslash}seq{\isachardot}ML}}}} for the underlying
+ implementation of lazy sequences.
- \item \verb|int -> tactic| represents tactics with explicit
- subgoal addressing, with well-formedness conditions as described
- above.
+ \item Type \verb|int -> tactic| represents tactics with
+ explicit subgoal addressing, with well-formedness conditions as
+ described above.
\item \verb|no_tac| is a tactic that always fails, returning the
empty sequence.
@@ -480,7 +481,10 @@
composition, disjunction (choice), iteration, or goal addressing.
Various search strategies may be expressed via tacticals.
- \medskip FIXME%
+ \medskip FIXME
+
+ \medskip The chapter on tacticals in \cite{isabelle-ref} is still
+ applicable, despite a few outdated details.%
\end{isamarkuptext}%
\isamarkuptrue%
%
--- a/doc-src/IsarImplementation/implementation.tex Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/implementation.tex Sun Oct 24 15:11:24 2010 -0700
@@ -4,6 +4,7 @@
\usepackage{../iman,../extra,../isar,../proof}
\usepackage[nohyphen,strings]{../underscore}
\usepackage{../isabelle,../isabellesym}
+\usepackage{../ttbox,../rail,../railsetup}
\usepackage{style}
\usepackage{../pdfsetup}
@@ -22,6 +23,11 @@
\makeindex
+\railterm{lbrace,rbrace,atsign}
+\railalias{lbracesym}{\isasymlbrace}\railterm{lbracesym}
+\railalias{rbracesym}{\isasymrbrace}\railterm{rbracesym}
+\railalias{dots}{\isasymdots}\railterm{dots}
+
\begin{document}
@@ -76,18 +82,18 @@
\listoffigures
\clearfirst
+\setcounter{chapter}{-1}
+
+\input{Thy/document/ML.tex}
\input{Thy/document/Prelim.tex}
\input{Thy/document/Logic.tex}
+\input{Thy/document/Syntax.tex}
\input{Thy/document/Tactic.tex}
\input{Thy/document/Proof.tex}
-\input{Thy/document/Syntax.tex}
\input{Thy/document/Isar.tex}
\input{Thy/document/Local_Theory.tex}
\input{Thy/document/Integration.tex}
-\appendix
-\input{Thy/document/ML.tex}
-
\begingroup
\tocentry{\bibname}
\bibliographystyle{abbrv} \small\raggedright\frenchspacing
--- a/doc-src/IsarImplementation/style.sty Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarImplementation/style.sty Sun Oct 24 15:11:24 2010 -0700
@@ -20,17 +20,24 @@
\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}\begingroup\def\isastyletext{\rm}\small}
-\renewcommand{\endisatagmlref}{\endgroup}
+\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{\isatagmlex}{\subsection*{\makebox[0pt][r]{\fbox{ML}~~}Examples}}
\renewcommand{\endisatagmlex}{}
\renewcommand{\isatagML}{\begingroup\isabellestyle{default}\isastyle\def\isadigit##1{##1}}
--- a/doc-src/IsarRef/Thy/Misc.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarRef/Thy/Misc.thy Sun Oct 24 15:11:24 2010 -0700
@@ -143,7 +143,7 @@
\end{matharray}
\begin{rail}
- ('cd' | 'use\_thy' | 'update\_thy') name
+ ('cd' | 'use\_thy') name
;
\end{rail}
@@ -159,6 +159,13 @@
since loading of theories is done automatically as required.
\end{description}
+
+ %FIXME proper place (!?)
+ Isabelle file specification may contain path variables (e.g.\
+ @{verbatim "$ISABELLE_HOME"}) that are expanded accordingly. Note
+ that @{verbatim "~"} abbreviates @{verbatim "$HOME"}, and @{verbatim
+ "~~"} abbreviates @{verbatim "$ISABELLE_HOME"}. The general syntax
+ for path specifications follows POSIX conventions.
*}
end
--- a/doc-src/IsarRef/Thy/Spec.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarRef/Thy/Spec.thy Sun Oct 24 15:11:24 2010 -0700
@@ -837,11 +837,6 @@
@{command_def "attribute_setup"} & : & @{text "theory \<rightarrow> theory"} \\
\end{matharray}
- \begin{mldecls}
- @{index_ML bind_thms: "string * thm list -> unit"} \\
- @{index_ML bind_thm: "string * thm -> unit"} \\
- \end{mldecls}
-
\begin{rail}
'use' name
;
@@ -918,20 +913,6 @@
let val context' = context
in context' end)) *} "my declaration"
-text {*
- \begin{description}
-
- \item @{ML bind_thms}~@{text "(name, thms)"} stores a list of
- theorems produced in ML both in the 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 bind_thm} is similar to @{ML bind_thms} but refers to a
- singleton theorem.
-
- \end{description}
-*}
-
section {* Primitive specification elements *}
@@ -973,7 +954,7 @@
Usually the default sort is only changed when defining a new
object-logic. For example, the default sort in Isabelle/HOL is
- @{text type}, the class of all HOL types. %FIXME sort antiq?
+ @{class type}, the class of all HOL types.
When merging theories, the default sorts of the parents are
logically intersected, i.e.\ the representations as lists of classes
--- a/doc-src/IsarRef/Thy/document/Misc.tex Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarRef/Thy/document/Misc.tex Sun Oct 24 15:11:24 2010 -0700
@@ -164,7 +164,7 @@
\end{matharray}
\begin{rail}
- ('cd' | 'use\_thy' | 'update\_thy') name
+ ('cd' | 'use\_thy') name
;
\end{rail}
@@ -179,7 +179,13 @@
These system commands are scarcely used when working interactively,
since loading of theories is done automatically as required.
- \end{description}%
+ \end{description}
+
+ %FIXME proper place (!?)
+ Isabelle file specification may contain path variables (e.g.\
+ \verb|$ISABELLE_HOME|) that are expanded accordingly. Note
+ that \verb|~| abbreviates \verb|$HOME|, and \verb|~~| abbreviates \verb|$ISABELLE_HOME|. The general syntax
+ for path specifications follows POSIX conventions.%
\end{isamarkuptext}%
\isamarkuptrue%
%
--- a/doc-src/IsarRef/Thy/document/Spec.tex Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/IsarRef/Thy/document/Spec.tex Sun Oct 24 15:11:24 2010 -0700
@@ -859,11 +859,6 @@
\indexdef{}{command}{attribute\_setup}\hypertarget{command.attribute-setup}{\hyperlink{command.attribute-setup}{\mbox{\isa{\isacommand{attribute{\isacharunderscore}setup}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
\end{matharray}
- \begin{mldecls}
- \indexdef{}{ML}{bind\_thms}\verb|bind_thms: string * thm list -> unit| \\
- \indexdef{}{ML}{bind\_thm}\verb|bind_thm: string * thm -> unit| \\
- \end{mldecls}
-
\begin{rail}
'use' name
;
@@ -951,21 +946,6 @@
%
\endisadelimML
%
-\begin{isamarkuptext}%
-\begin{description}
-
- \item \verb|bind_thms|~\isa{{\isachardoublequote}{\isacharparenleft}name{\isacharcomma}\ thms{\isacharparenright}{\isachardoublequote}} stores a list of
- theorems produced in ML both in the theory context and the ML
- toplevel, associating it with the provided name. Theorems are put
- into a global ``standard'' format before being stored.
-
- \item \verb|bind_thm| is similar to \verb|bind_thms| but refers to a
- singleton theorem.
-
- \end{description}%
-\end{isamarkuptext}%
-\isamarkuptrue%
-%
\isamarkupsection{Primitive specification elements%
}
\isamarkuptrue%
@@ -1010,7 +990,7 @@
Usually the default sort is only changed when defining a new
object-logic. For example, the default sort in Isabelle/HOL is
- \isa{type}, the class of all HOL types. %FIXME sort antiq?
+ \isa{type}, the class of all HOL types.
When merging theories, the default sorts of the parents are
logically intersected, i.e.\ the representations as lists of classes
--- a/doc-src/Ref/introduction.tex Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/Ref/introduction.tex Sun Oct 24 15:11:24 2010 -0700
@@ -1,82 +1,5 @@
-\chapter{Basic Use of Isabelle}\index{sessions|(}
-
-\section{Basic interaction with Isabelle}
-\index{starting up|bold}\nobreak
-%
-We assume that your local Isabelle administrator (this might be you!) has
-already installed the Isabelle system together with appropriate object-logics.
-
-\medskip Let $\langle isabellehome \rangle$ denote the location where
-the distribution has been installed. To run Isabelle from a the shell
-prompt within an ordinary text terminal session, simply type
-\begin{ttbox}
-\({\langle}isabellehome{\rangle}\)/bin/isabelle
-\end{ttbox}
-This should start an interactive \ML{} session with the default object-logic
-(usually HOL) already pre-loaded.
-
-Subsequently, we assume that the \texttt{isabelle} executable is determined
-automatically by the shell, e.g.\ by adding {\tt \(\langle isabellehome
- \rangle\)/bin} to your search path.\footnote{Depending on your installation,
- there may be stand-alone binaries located in some global directory such as
- \texttt{/usr/bin}. Do not attempt to copy {\tt \(\langle isabellehome
- \rangle\)/bin/isabelle}, though! See \texttt{isabelle install} in
- \emph{The Isabelle System Manual} of how to do this properly.}
-
-\medskip
-
-The object-logic image to load may be also specified explicitly as an argument
-to the {\tt isabelle} command, e.g.
-\begin{ttbox}
-isabelle FOL
-\end{ttbox}
-This should put you into the world of polymorphic first-order logic (assuming
-that an image of FOL has been pre-built).
-
-\index{saving your session|bold} Isabelle provides no means of storing
-theorems or internal proof objects on files. Theorems are simply part of the
-\ML{} state. To save your work between sessions, you may dump the \ML{}
-system state to a file. This is done automatically when ending the session
-normally (e.g.\ by typing control-D), provided that the image has been opened
-\emph{writable} in the first place. The standard object-logic images are
-usually read-only, so you have to create a private working copy first. For
-example, the following shell command puts you into a writable Isabelle session
-of name \texttt{Foo} that initially contains just plain HOL:
-\begin{ttbox}
-isabelle HOL Foo
-\end{ttbox}
-Ending the \texttt{Foo} session with control-D will cause the complete
-\ML-world to be saved somewhere in your home directory\footnote{The default
- location is in \texttt{\~\relax/isabelle/heaps}, but this depends on your
- local configuration.}. Make sure there is enough space available! Then one
-may later continue at exactly the same point by running
-\begin{ttbox}
-isabelle Foo
-\end{ttbox}
-
-\medskip Saving the {\ML} state is not enough. Record, on a file, the
-top-level commands that generate your theories and proofs. Such a record
-allows you to replay the proofs whenever required, for instance after making
-minor changes to the axioms. Ideally, these sources will be somewhat
-intelligible to others as a formal description of your work.
-
-It is good practice to put all source files that constitute a separate
-Isabelle session into an individual directory, together with an {\ML} file
-called \texttt{ROOT.ML} that contains appropriate commands to load all other
-files required. Running \texttt{isabelle} with option \texttt{-u}
-automatically loads \texttt{ROOT.ML} on entering the session. The
-\texttt{isabelle usedir} utility provides some more options to manage Isabelle
-sessions, such as automatic generation of theory browsing information.
-
-\medskip More details about the \texttt{isabelle} and \texttt{isabelle}
-commands may be found in \emph{The Isabelle System Manual}.
-
-\medskip There are more comfortable user interfaces than the bare-bones \ML{}
-top-level run from a text terminal. The \texttt{Isabelle} executable (note
-the capital I) runs one such interface, depending on your local configuration.
-Again, see \emph{The Isabelle System Manual} for more information.
-
+\chapter{Basic Use of Isabelle}
\section{Ending a session}
\begin{ttbox}
@@ -99,149 +22,6 @@
Typing control-D also finishes the session in essentially the same way
as the sequence {\tt commit(); quit();} would.
-
-\section{Reading ML files}
-\index{files!reading}
-\begin{ttbox}
-cd : string -> unit
-pwd : unit -> string
-use : string -> unit
-time_use : string -> unit
-\end{ttbox}
-\begin{ttdescription}
-\item[\ttindexbold{cd} "{\it dir}";] changes the current directory to
- {\it dir}. This is the default directory for reading files.
-
-\item[\ttindexbold{pwd}();] returns the full path of the current
- directory.
-
-\item[\ttindexbold{use} "$file$";]
-reads the given {\it file} as input to the \ML{} session. Reading a file
-of Isabelle commands is the usual way of replaying a proof.
-
-\item[\ttindexbold{time_use} "$file$";]
-performs {\tt use~"$file$"} and prints the total execution time.
-\end{ttdescription}
-
-The $dir$ and $file$ specifications of the \texttt{cd} and \texttt{use}
-commands may contain path variables (e.g.\ \texttt{\$ISABELLE_HOME}) that are
-expanded appropriately. Note that \texttt{\~\relax} abbreviates
-\texttt{\$HOME}, and \texttt{\~\relax\~\relax} abbreviates
-\texttt{\$ISABELLE_HOME}\index{*\$ISABELLE_HOME}. The syntax for path
-specifications follows Unix conventions.
-
-
-\section{Reading theories}\label{sec:intro-theories}
-\index{theories!reading}
-
-In Isabelle, any kind of declarations, definitions, etc.\ are organized around
-named \emph{theory} objects. Logical reasoning always takes place within a
-certain theory context, which may be switched at any time. Theory $name$ is
-defined by a theory file $name$\texttt{.thy}, containing declarations of
-\texttt{consts}, \texttt{types}, \texttt{defs}, etc.\ (see
-\S\ref{sec:ref-defining-theories} for more details on concrete syntax).
-Furthermore, there may be an associated {\ML} file $name$\texttt{.ML} with
-proof scripts that are to be run in the context of the theory.
-
-\begin{ttbox}
-context : theory -> unit
-the_context : unit -> theory
-theory : string -> theory
-use_thy : string -> unit
-time_use_thy : string -> unit
-update_thy : string -> unit
-\end{ttbox}
-
-\begin{ttdescription}
-
-\item[\ttindexbold{context} $thy$;] switches the current theory context. Any
- subsequent command with ``implicit theory argument'' (e.g.\ \texttt{Goal})
- will refer to $thy$ as its theory.
-
-\item[\ttindexbold{the_context}();] obtains the current theory context, or
- raises an error if absent.
-
-\item[\ttindexbold{theory} "$name$";] retrieves the theory called $name$ from
- the internal data\-base of loaded theories, raising an error if absent.
-
-\item[\ttindexbold{use_thy} "$name$";] reads theory $name$ from the file
- system, looking for $name$\texttt{.thy} and $name$\texttt{.ML} (the latter
- being optional). It also ensures that all parent theories are loaded as
- well. In case some older versions have already been present,
- \texttt{use_thy} only tries to reload $name$ itself, but is content with any
- version of its ancestors.
-
-\item[\ttindexbold{time_use_thy} "$name$";] same as \texttt{use_thy}, but
- reports the time taken to process the actual theory parts and {\ML} files
- separately.
-
-\item[\ttindexbold{update_thy} "$name$";] is similar to \texttt{use_thy}, but
- ensures that theory $name$ is fully up-to-date with respect to the file
- system --- apart from theory $name$ itself, any of its ancestors may be
- reloaded as well.
-
-\end{ttdescription}
-
-Note that theories of pre-built logic images (e.g.\ HOL) are marked as
-\emph{finished} and cannot be updated any more. See \S\ref{sec:more-theories}
-for further information on Isabelle's theory loader.
-
-
-\section{Setting flags}
-\begin{ttbox}
-set : bool ref -> bool
-reset : bool ref -> bool
-toggle : bool ref -> bool
-\end{ttbox}\index{*set}\index{*reset}\index{*toggle}
-These are some shorthands for manipulating boolean references. The new
-value is returned.
-
-
-\section{Diagnostic messages}
-\index{error messages}
-\index{warnings}
-
-Isabelle conceptually provides three output channels for different kinds of
-messages: ordinary text, warnings, errors. Depending on the user interface
-involved, these messages may appear in different text styles or colours.
-
-The default setup of an \texttt{isabelle} terminal session is as
-follows: plain output of ordinary text, warnings prefixed by
-\texttt{\#\#\#}'s, errors prefixed by \texttt{***}'s. For example, a
-typical warning would look like this:
-\begin{ttbox}
-\#\#\# Beware the Jabberwock, my son!
-\#\#\# The jaws that bite, the claws that catch!
-\#\#\# Beware the Jubjub Bird, and shun
-\#\#\# The frumious Bandersnatch!
-\end{ttbox}
-
-\texttt{ML} programs may output diagnostic messages using the
-following functions:
-\begin{ttbox}
-writeln : string -> unit
-warning : string -> unit
-error : string -> 'a
-\end{ttbox}
-Note that \ttindex{error} fails by raising exception \ttindex{ERROR}
-after having output the text, while \ttindex{writeln} and
-\ttindex{warning} resume normal program execution.
-
-
-\section{Timing}
-\index{timing statistics}\index{proofs!timing}
-\begin{ttbox}
-timing: bool ref \hfill{\bf initially false}
-\end{ttbox}
-
-\begin{ttdescription}
-\item[set \ttindexbold{timing};] enables global timing in Isabelle.
- This information is compiler-dependent.
-\end{ttdescription}
-
-\index{sessions|)}
-
-
%%% Local Variables:
%%% mode: latex
%%% TeX-master: "ref"
--- a/doc-src/Ref/theories.tex Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/Ref/theories.tex Sun Oct 24 15:11:24 2010 -0700
@@ -7,37 +7,6 @@
Isabelle's theory loader manages dependencies of the internal graph of theory
nodes (the \emph{theory database}) and the external view of the file system.
-See \S\ref{sec:intro-theories} for its most basic commands, such as
-\texttt{use_thy}. There are a few more operations available.
-
-\begin{ttbox}
-use_thy_only : string -> unit
-update_thy_only : string -> unit
-touch_thy : string -> unit
-remove_thy : string -> unit
-delete_tmpfiles : bool ref \hfill\textbf{initially true}
-\end{ttbox}
-
-\begin{ttdescription}
-\item[\ttindexbold{use_thy_only} "$name$";] is similar to \texttt{use_thy},
- but processes the actual theory file $name$\texttt{.thy} only, ignoring
- $name$\texttt{.ML}. This might be useful in replaying proof scripts by hand
- from the very beginning, starting with the fresh theory.
-
-\item[\ttindexbold{update_thy_only} "$name$";] is similar to
- \texttt{update_thy}, but processes the actual theory file
- $name$\texttt{.thy} only, ignoring $name$\texttt{.ML}.
-
-\item[\ttindexbold{touch_thy} "$name$";] marks theory node $name$ of the
- internal graph as outdated. While the theory remains usable, subsequent
- operations such as \texttt{use_thy} may cause a reload.
-
-\item[\ttindexbold{remove_thy} "$name$";] deletes theory node $name$,
- including \emph{all} of its descendants. Beware! This is a quick way to
- dispose a large number of theories at once. Note that {\ML} bindings to
- theorems etc.\ of removed theories may still persist.
-
-\end{ttdescription}
\medskip Theory and {\ML} files are located by skimming through the
directories listed in Isabelle's internal load path, which merely contains the
@@ -83,133 +52,18 @@
\subsection{*Theory inclusion}
\begin{ttbox}
-subthy : theory * theory -> bool
-eq_thy : theory * theory -> bool
transfer : theory -> thm -> thm
-transfer_sg : Sign.sg -> thm -> thm
\end{ttbox}
-Inclusion and equality of theories is determined by unique
-identification stamps that are created when declaring new components.
-Theorems contain a reference to the theory (actually to its signature)
-they have been derived in. Transferring theorems to super theories
-has no logical significance, but may affect some operations in subtle
-ways (e.g.\ implicit merges of signatures when applying rules, or
-pretty printing of theorems).
+Transferring theorems to super theories has no logical significance,
+but may affect some operations in subtle ways (e.g.\ implicit merges
+of signatures when applying rules, or pretty printing of theorems).
\begin{ttdescription}
-\item[\ttindexbold{subthy} ($thy@1$, $thy@2$)] determines if $thy@1$
- is included in $thy@2$ wrt.\ identification stamps.
-
-\item[\ttindexbold{eq_thy} ($thy@1$, $thy@2$)] determines if $thy@1$
- is exactly the same as $thy@2$.
-
\item[\ttindexbold{transfer} $thy$ $thm$] transfers theorem $thm$ to
theory $thy$, provided the latter includes the theory of $thm$.
-\item[\ttindexbold{transfer_sg} $sign$ $thm$] is similar to
- \texttt{transfer}, but identifies the super theory via its
- signature.
-
-\end{ttdescription}
-
-
-\subsection{*Primitive theories}
-\begin{ttbox}
-ProtoPure.thy : theory
-Pure.thy : theory
-CPure.thy : theory
-\end{ttbox}
-\begin{description}
-\item[\ttindexbold{ProtoPure.thy}, \ttindexbold{Pure.thy},
- \ttindexbold{CPure.thy}] contain the syntax and signature of the
- meta-logic. There are basically no axioms: meta-level inferences
- are carried out by \ML\ functions. \texttt{Pure} and \texttt{CPure}
- just differ in their concrete syntax of prefix function application:
- $t(u@1, \ldots, u@n)$ in \texttt{Pure} vs.\ $t\,u@1,\ldots\,u@n$ in
- \texttt{CPure}. \texttt{ProtoPure} is their common parent,
- containing no syntax for printing prefix applications at all!
-
-%% FIXME
-%\item [\ttindexbold{extend_theory} $thy$ {\tt"}$T${\tt"} $\cdots$] extends
-% the theory $thy$ with new types, constants, etc. $T$ identifies the theory
-% internally. When a theory is redeclared, say to change an incorrect axiom,
-% bindings to the old axiom may persist. Isabelle ensures that the old and
-% new theories are not involved in the same proof. Attempting to combine
-% different theories having the same name $T$ yields the fatal error
-%extend_theory : theory -> string -> \(\cdots\) -> theory
-%\begin{ttbox}
-%Attempt to merge different versions of theory: \(T\)
-%\end{ttbox}
-\end{description}
-
-%% FIXME
-%\item [\ttindexbold{extend_theory} $thy$ {\tt"}$T${\tt"}
-% ($classes$, $default$, $types$, $arities$, $consts$, $sextopt$) $rules$]
-%\hfill\break %%% include if line is just too short
-%is the \ML{} equivalent of the following theory definition:
-%\begin{ttbox}
-%\(T\) = \(thy\) +
-%classes \(c\) < \(c@1\),\(\dots\),\(c@m\)
-% \dots
-%default {\(d@1,\dots,d@r\)}
-%types \(tycon@1\),\dots,\(tycon@i\) \(n\)
-% \dots
-%arities \(tycon@1'\),\dots,\(tycon@j'\) :: (\(s@1\),\dots,\(s@n\))\(c\)
-% \dots
-%consts \(b@1\),\dots,\(b@k\) :: \(\tau\)
-% \dots
-%rules \(name\) \(rule\)
-% \dots
-%end
-%\end{ttbox}
-%where
-%\begin{tabular}[t]{l@{~=~}l}
-%$classes$ & \tt[("$c$",["$c@1$",\dots,"$c@m$"]),\dots] \\
-%$default$ & \tt["$d@1$",\dots,"$d@r$"]\\
-%$types$ & \tt[([$tycon@1$,\dots,$tycon@i$], $n$),\dots] \\
-%$arities$ & \tt[([$tycon'@1$,\dots,$tycon'@j$], ([$s@1$,\dots,$s@n$],$c$)),\dots]
-%\\
-%$consts$ & \tt[([$b@1$,\dots,$b@k$],$\tau$),\dots] \\
-%$rules$ & \tt[("$name$",$rule$),\dots]
-%\end{tabular}
-
-
-\subsection{Inspecting a theory}\label{sec:inspct-thy}
-\index{theories!inspecting|bold}
-\begin{ttbox}
-print_syntax : theory -> unit
-print_theory : theory -> unit
-parents_of : theory -> theory list
-ancestors_of : theory -> theory list
-sign_of : theory -> Sign.sg
-Sign.stamp_names_of : Sign.sg -> string list
-\end{ttbox}
-These provide means of viewing a theory's components.
-\begin{ttdescription}
-\item[\ttindexbold{print_syntax} $thy$] prints the syntax of $thy$
- (grammar, macros, translation functions etc., see
- page~\pageref{pg:print_syn} for more details).
-
-\item[\ttindexbold{print_theory} $thy$] prints the logical parts of
- $thy$, excluding the syntax.
-
-\item[\ttindexbold{parents_of} $thy$] returns the direct ancestors
- of~$thy$.
-
-\item[\ttindexbold{ancestors_of} $thy$] returns all ancestors of~$thy$
- (not including $thy$ itself).
-
-\item[\ttindexbold{sign_of} $thy$] returns the signature associated
- with~$thy$. It is useful with functions like {\tt
- read_instantiate_sg}, which take a signature as an argument.
-
-\item[\ttindexbold{Sign.stamp_names_of} $sg$]\index{signatures}
- returns the names of the identification \rmindex{stamps} of ax
- signature. These coincide with the names of its full ancestry
- including that of $sg$ itself.
-
\end{ttdescription}
--- a/doc-src/Sledgehammer/sledgehammer.tex Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/Sledgehammer/sledgehammer.tex Sun Oct 24 15:11:24 2010 -0700
@@ -78,24 +78,26 @@
\label{introduction}
Sledgehammer is a tool that applies first-order automatic theorem provers (ATPs)
-on the current goal. The supported ATPs are E \cite{schulz-2002}, SPASS
-\cite{weidenbach-et-al-2009}, and Vampire \cite{riazanov-voronkov-2002}, which
-can be run locally or remotely via the SystemOnTPTP web service
-\cite{sutcliffe-2000}.
+and satisfiability-modulo-theory (SMT) solvers on the current goal. The
+supported ATPs are E \cite{schulz-2002}, SPASS \cite{weidenbach-et-al-2009},
+Vampire \cite{riazanov-voronkov-2002}, SInE-E \cite{sine}, and SNARK
+\cite{snark}. The ATPs are run either locally or remotely via the
+System\-On\-TPTP web service \cite{sutcliffe-2000}. In addition to the ATPs, the
+\textit{smt} proof method (which typically relies on the Z3 SMT solver
+\cite{z3}) is tried as well.
-The problem passed to ATPs consists of your current goal together with a
-heuristic selection of hundreds of facts (theorems) from the current theory
-context, filtered by relevance. Because jobs are run in the background, you can
-continue to work on your proof by other means. Provers can be run in parallel.
-Any reply (which may arrive minutes later) will appear in the Proof General
-response buffer.
+The problem passed to the automatic provers consists of your current goal
+together with a heuristic selection of hundreds of facts (theorems) from the
+current theory context, filtered by relevance. Because jobs are run in the
+background, you can continue to work on your proof by other means. Provers can
+be run in parallel. Any reply (which may arrive half a minute later) will appear
+in the Proof General response buffer.
-The result of a successful ATP proof search is some source text that usually
-(but not always) reconstructs the proof within Isabelle, without requiring the
-ATPs again. The reconstructed proof relies on the general-purpose Metis prover
-\cite{metis}, which is fully integrated into Isabelle/HOL, with explicit
-inferences going through the kernel. Thus its results are correct by
-construction.
+The result of a successful proof search is some source text that usually (but
+not always) reconstructs the proof within Isabelle. For ATPs, the reconstructed
+proof relies on the general-purpose Metis prover \cite{metis}, which is fully
+integrated into Isabelle/HOL, with explicit inferences going through the kernel.
+Thus its results are correct by construction.
In this manual, we will explicitly invoke the \textbf{sledgehammer} command.
Sledgehammer also provides an automatic mode that can be enabled via the
@@ -123,10 +125,10 @@
\label{installation}
Sledgehammer is part of Isabelle, so you don't need to install it. However, it
-relies on third-party automatic theorem provers (ATPs). Currently, E, SPASS, and
-Vampire are supported. All of these are available remotely via SystemOnTPTP
-\cite{sutcliffe-2000}, but if you want better performance you will need to
-install at least E and SPASS locally.
+relies on third-party automatic theorem provers (ATPs) and SAT solvers.
+Currently, E, SPASS, and Vampire can be run locally; in addition, E, Vampire,
+SInE-E, and SNARK are available remotely via SystemOnTPTP \cite{sutcliffe-2000}.
+If you want better performance, you should install E and SPASS locally.
There are three main ways to install ATPs on your machine:
@@ -199,38 +201,47 @@
\prew
\slshape
-Sledgehammer: ATP ``\textit{e}'' for subgoal 1: \\
+Sledgehammer: ``\textit{e}'' for subgoal 1: \\
$([a] = [b]) = (a = b)$ \\
Try this command: \textbf{by} (\textit{metis hd.simps}). \\
To minimize the number of lemmas, try this: \\
-\textbf{sledgehammer} \textit{minimize} [\textit{atp} = \textit{e}] (\textit{hd.simps}). \\[3\smallskipamount]
+\textbf{sledgehammer} \textit{minimize} [\textit{prover} = \textit{e}] (\textit{hd.simps}). \\[3\smallskipamount]
%
-Sledgehammer: ATP ``\textit{spass}'' for subgoal 1: \\
+Sledgehammer: ``\textit{spass}'' for subgoal 1: \\
$([a] = [b]) = (a = b)$ \\
Try this command: \textbf{by} (\textit{metis insert\_Nil last\_ConsL}). \\
To minimize the number of lemmas, try this: \\
-\textbf{sledgehammer} \textit{minimize} [\textit{atp} = \textit{spass}] (\textit{insert\_Nil last\_ConsL}). \\[3\smallskipamount]
+\textbf{sledgehammer} \textit{minimize} [\textit{prover} = \textit{spass}] (\textit{insert\_Nil last\_ConsL}). \\[3\smallskipamount]
%
-Sledgehammer: ATP ``\textit{remote\_vampire}'' for subgoal 1: \\
+Sledgehammer: ``\textit{vampire}'' for subgoal 1: \\
+$([a] = [b]) = (a = b)$ \\
+Try this command: \textbf{by} (\textit{metis eq\_commute last\_snoc}) \\
+To minimize the number of lemmas, try this: \\
+\textbf{sledgehammer} \textit{minimize} [\textit{prover} = \textit{vampire}]~(\textit{eq\_commute last\_snoc}). \\[3\smallskipamount]
+%
+Sledgehammer: ``\textit{remote\_sine\_e}'' for subgoal 1: \\
$([a] = [b]) = (a = b)$ \\
-Try this command: \textbf{by} (\textit{metis One\_nat\_def\_raw empty\_replicate} \\
-\phantom{Try this command: \textbf{by} (\textit{metis~}}\textit{insert\_Nil last\_ConsL replicate\_Suc}). \\
+Try this command: \textbf{by} (\textit{metis hd.simps}) \\
To minimize the number of lemmas, try this: \\
-\textbf{sledgehammer} \textit{minimize} [\textit{atp} = \textit{remote\_vampire}] \\
-\phantom{\textbf{sledgehammer}~}(\textit{One\_nat\_def\_raw empty\_replicate insert\_Nil} \\
-\phantom{\textbf{sledgehammer}~(}\textit{last\_ConsL replicate\_Suc}).
+\textbf{sledgehammer} \textit{minimize} [\textit{prover} = \textit{remote\_sine\_e}]~(\textit{hd.simps}). \\[3\smallskipamount]
+%
+Sledgehammer: ``\textit{smt}'' for subgoal 1: \\
+$([a] = [b]) = (a = b)$ \\
+Try this command: \textbf{by} (\textit{smt hd.simps}) \\
+To minimize the number of lemmas, try this: \\
+\textbf{sledgehammer} \textit{minimize} [\textit{prover} = \textit{smt}]~(\textit{hd.simps}).
\postw
-Sledgehammer ran E, SPASS, and the remote version of Vampire in parallel. If E
-is not installed (\S\ref{installation}), you will see references to
-its remote American cousin \textit{remote\_e} instead of
-\textit{e}; and if SPASS is not installed, it will not appear in the output.
+Sledgehammer ran E, SPASS, Vampire, SInE-E, and the \textit{smt} proof method in
+parallel. Depending on which provers are installed and how many processor cores
+are available, some of the provers might be missing or present with a
+\textit{remote\_} prefix.
-Based on each ATP proof, Sledgehammer gives a one-liner proof that uses the
-\textit{metis} method. You can click them and insert them into the theory text.
-You can click the ``\textbf{sledgehammer} \textit{minimize}'' command if you
-want to look for a shorter (and faster) proof. But here the proof found by E
-looks perfect, so click it to finish the proof.
+For each successful prover, Sledgehammer gives a one-liner proof that uses the
+\textit{metis} or \textit{smt} method. You can click the proof to insert it into
+the theory text. You can click the ``\textbf{sledgehammer} \textit{minimize}''
+command if you want to look for a shorter (and probably faster) proof. But here
+the proof found by E looks perfect, so click it to finish the proof.
You can ask Sledgehammer for an Isar text proof by passing the
\textit{isar\_proof} option:
@@ -241,7 +252,7 @@
When Isar proof construction is successful, it can yield proofs that are more
readable and also faster than the \textit{metis} one-liners. This feature is
-experimental.
+experimental and is only available for ATPs.
\section{Hints}
\label{hints}
@@ -291,14 +302,16 @@
to Sledgehammer's asynchronous nature. The \textit{num} argument specifies a
limit on the number of messages to display (5 by default).
-\item[$\bullet$] \textbf{\textit{available\_atps}:} Prints the list of installed ATPs.
+\item[$\bullet$] \textbf{\textit{available\_provers}:} Prints the list of installed provers.
See \S\ref{installation} and \S\ref{mode-of-operation} for more information on
-how to install ATPs.
+how to install automatic provers.
-\item[$\bullet$] \textbf{\textit{running\_atps}:} Prints information about currently
-running ATPs, including elapsed runtime and remaining time until timeout.
+\item[$\bullet$] \textbf{\textit{running\_provers}:} Prints information about
+currently running automatic provers, including elapsed runtime and remaining
+time until timeout.
-\item[$\bullet$] \textbf{\textit{kill\_atps}:} Terminates all running ATPs.
+\item[$\bullet$] \textbf{\textit{kill\_provers}:} Terminates all running
+automatic provers.
\item[$\bullet$] \textbf{\textit{refresh\_tptp}:} Refreshes the list of remote
ATPs available at System\-On\-TPTP \cite{sutcliffe-2000}.
@@ -334,12 +347,12 @@
You can instruct Sledgehammer to run automatically on newly entered theorems by
enabling the ``Auto Sledgehammer'' option from the ``Isabelle'' menu in Proof
-General. For automatic runs, only the first ATP set using \textit{atps}
+General. For automatic runs, only the first prover set using \textit{provers}
(\S\ref{mode-of-operation}) is considered, \textit{verbose}
(\S\ref{output-format}) and \textit{debug} (\S\ref{output-format}) are disabled,
-fewer facts are passed to the ATP, and \textit{timeout} (\S\ref{timeouts}) is
-superseded by the ``Auto Tools Time Limit'' in Proof General's ``Isabelle''
-menu. Sledgehammer's output is also more concise.
+fewer facts are passed to the prover, and \textit{timeout}
+(\S\ref{mode-of-operation}) is superseded by the ``Auto Tools Time Limit'' in
+Proof General's ``Isabelle'' menu. Sledgehammer's output is also more concise.
\section{Option Reference}
\label{option-reference}
@@ -382,9 +395,9 @@
\label{mode-of-operation}
\begin{enum}
-\opnodefault{atps}{string}
-Specifies the ATPs (automated theorem provers) to use as a space-separated list
-(e.g., ``\textit{e}~\textit{spass}''). The following ATPs are supported:
+\opnodefault{provers}{string}
+Specifies the automatic provers to use as a space-separated list (e.g.,
+``\textit{e}~\textit{spass}''). The following provers are supported:
\begin{enum}
\item[$\bullet$] \textbf{\textit{e}:} E is an ATP developed by Stephan Schulz
@@ -405,6 +418,9 @@
Vampire, set the environment variable \texttt{VAMPIRE\_HOME} to the directory
that contains the \texttt{vampire} executable.
+\item[$\bullet$] \textbf{\textit{smt}:} The \textit{smt} proof method invokes an
+external SMT prover (usually Z3 \cite{z3}).
+
\item[$\bullet$] \textbf{\textit{remote\_e}:} The remote version of E runs
on Geoff Sutcliffe's Miami servers \cite{sutcliffe-2000}.
@@ -418,27 +434,38 @@
\item[$\bullet$] \textbf{\textit{remote\_snark}:} SNARK is a prover
developed by Stickel et al.\ \cite{snark}. The remote version of
SNARK runs on Geoff Sutcliffe's Miami servers.
+
+\item[$\bullet$] \textbf{\textit{remote\_smt}:} The remote version of the
+\textit{smt} proof method runs the SMT solver on servers at the TU M\"unchen (or
+wherever \texttt{REMOTE\_SMT\_URL} is set to point).
+
\end{enum}
-By default, Sledgehammer will run E, SPASS, Vampire, and SInE-E in parallel.
-For at most two of E, SPASS, and Vampire, it will use any locally installed
-version if available. For historical reasons, the default value of this option
-can be overridden using the option ``Sledgehammer: ATPs'' from the ``Isabelle''
-menu in Proof General.
+By default, Sledgehammer will run E, SPASS, Vampire, SInE-E, and \textit{smt} in
+parallel---either locally or remotely, depending on the number of processor
+cores available. For historical reasons, the default value of this option can be
+overridden using the option ``Sledgehammer: Provers'' from the ``Isabelle'' menu
+in Proof General.
-It is a good idea to run several ATPs in parallel, although it could slow down
-your machine. Running E, SPASS, and Vampire together for 5 seconds yields about
-the same success rate as running the most effective of these (Vampire) for 120
-seconds \cite{boehme-nipkow-2010}.
+It is a good idea to run several provers in parallel, although it could slow
+down your machine. Running E, SPASS, Vampire, and SInE-E together for 5 seconds
+yields a better success rate than running the most effective of these (Vampire)
+for 120 seconds \cite{boehme-nipkow-2010}.
+
+\opnodefault{prover}{string}
+Alias for \textit{provers}.
+
+\opnodefault{atps}{string}
+Legacy alias for \textit{provers}.
\opnodefault{atp}{string}
-Alias for \textit{atps}.
+Legacy alias for \textit{provers}.
\opdefault{timeout}{time}{$\mathbf{30}$ s}
-Specifies the maximum amount of time that the ATPs should spend searching for a
-proof. For historical reasons, the default value of this option can be
-overridden using the option ``Sledgehammer: Time Limit'' from the ``Isabelle''
-menu in Proof General.
+Specifies the maximum amount of time that the automatic provers should spend
+searching for a proof. For historical reasons, the default value of this option
+can be overridden using the option ``Sledgehammer: Time Limit'' from the
+``Isabelle'' menu in Proof General.
\opfalse{blocking}{non\_blocking}
Specifies whether the \textbf{sledgehammer} command should operate
@@ -463,16 +490,17 @@
\begin{enum}
\opfalse{explicit\_apply}{implicit\_apply}
Specifies whether function application should be encoded as an explicit
-``apply'' operator. If the option is set to \textit{false}, each function will
-be directly applied to as many arguments as possible. Enabling this option can
-sometimes help discover higher-order proofs that otherwise would not be found.
+``apply'' operator in ATP problems. If the option is set to \textit{false}, each
+function will be directly applied to as many arguments as possible. Enabling
+this option can sometimes help discover higher-order proofs that otherwise would
+not be found.
\opfalse{full\_types}{partial\_types}
-Specifies whether full-type information is exported. Enabling this option can
-prevent the discovery of type-incorrect proofs, but it also tends to slow down
-the ATPs significantly. For historical reasons, the default value of this option
-can be overridden using the option ``Sledgehammer: ATPs'' from the ``Isabelle''
-menu in Proof General.
+Specifies whether full-type information is encoded in ATP problems. Enabling
+this option can prevent the discovery of type-incorrect proofs, but it also
+tends to slow down the ATPs significantly. For historical reasons, the default
+value of this option can be overridden using the option ``Sledgehammer: Full
+Types'' from the ``Isabelle'' menu in Proof General.
\end{enum}
\subsection{Relevance Filter}
@@ -490,18 +518,8 @@
\opdefault{max\_relevant}{int\_or\_smart}{\textit{smart}}
Specifies the maximum number of facts that may be returned by the relevance
filter. If the option is set to \textit{smart}, it is set to a value that was
-empirically found to be appropriate for the ATP. A typical value would be 300.
-
-%\opsmartx{theory\_relevant}{theory\_irrelevant}
-%Specifies whether the theory from which a fact comes should be taken into
-%consideration by the relevance filter. If the option is set to \textit{smart},
-%it is taken to be \textit{true} for SPASS and \textit{false} for the other ATPs;
-%empirical results suggest that these are the best settings.
-
-%\opfalse{defs\_relevant}{defs\_irrelevant}
-%Specifies whether the definition of constants occurring in the formula to prove
-%should be considered particularly relevant. Enabling this option tends to lead
-%to larger problems and typically slows down the ATPs.
+empirically found to be appropriate for the prover. A typical value would be
+300.
\end{enum}
\subsection{Output Format}
--- a/doc-src/antiquote_setup.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/antiquote_setup.ML Sun Oct 24 15:11:24 2010 -0700
@@ -44,18 +44,25 @@
local
fun ml_val (txt1, "") = "fn _ => (" ^ txt1 ^ ");"
- | ml_val (txt1, txt2) = "fn _ => (" ^ txt1 ^ ": " ^ txt2 ^ ");";
+ | ml_val (txt1, txt2) = "fn _ => (" ^ txt1 ^ " : " ^ txt2 ^ ");";
fun ml_type (txt1, "") = "val _ = NONE : (" ^ txt1 ^ ") option;"
| ml_type (txt1, txt2) = "val _ = [NONE : (" ^ txt1 ^ ") option, NONE : (" ^ txt2 ^ ") option];";
-fun ml_exn (txt1, "") = "fn _ => (" ^ txt1 ^ ": exn);"
- | ml_exn (txt1, txt2) = "fn _ => (" ^ txt1 ^ ": " ^ txt2 ^ " -> exn);";
+fun ml_exn (txt1, "") = "fn _ => (" ^ txt1 ^ " : exn);"
+ | ml_exn (txt1, txt2) = "fn _ => (" ^ txt1 ^ " : " ^ txt2 ^ " -> exn);";
fun ml_structure (txt, _) = "functor XXX() = struct structure XX = " ^ txt ^ " end;";
fun ml_functor (txt, _) = "ML_Env.check_functor " ^ ML_Syntax.print_string txt;
+val is_name = ML_Lex.kind_of #> (fn kind => kind = ML_Lex.Ident orelse kind = ML_Lex.LongIdent);
+
+fun ml_name txt =
+ (case filter is_name (ML_Lex.tokenize txt) of
+ toks as [_] => ML_Lex.flatten toks
+ | _ => error ("Single ML name expected in input: " ^ quote txt));
+
fun index_ml name kind ml = Thy_Output.antiquotation name
(Scan.lift (Args.name -- Scan.optional (Args.colon |-- Args.name) ""))
(fn {context = ctxt, ...} => fn (txt1, txt2) =>
@@ -64,12 +71,13 @@
if txt2 = "" then txt1
else if kind = "type" then txt1 ^ " = " ^ txt2
else if kind = "exception" then txt1 ^ " of " ^ txt2
- else txt1 ^ ": " ^ txt2;
+ else if Syntax.is_identifier (Long_Name.base_name (ml_name txt1)) then txt1 ^ ": " ^ txt2
+ else txt1 ^ " : " ^ txt2;
val txt' = if kind = "" then txt else kind ^ " " ^ txt;
val _ = ML_Context.eval_text_in (SOME ctxt) false Position.none (ml (txt1, txt2)); (* ML_Lex.read (!?) *)
val kind' = if kind = "" then "ML" else "ML " ^ kind;
in
- "\\indexdef{}{" ^ kind' ^ "}{" ^ clean_string txt1 ^ "}" ^
+ "\\indexdef{}{" ^ kind' ^ "}{" ^ clean_string (ml_name txt1) ^ "}" ^
(txt'
|> (if Config.get ctxt Thy_Output.quotes then quote else I)
|> (if Config.get ctxt Thy_Output.display then enclose "\\begin{verbatim}\n" "\n\\end{verbatim}"
@@ -184,6 +192,7 @@
val _ = entity_antiqs (K check_tool) "isatt" "tool";
val _ = entity_antiqs (K (File.exists o Path.explode)) "isatt" "file";
val _ = entity_antiqs (K (can Thy_Info.get_theory)) "" "theory";
+val _ = entity_antiqs no_check "" "ML antiquotation"; (* FIXME proper check *)
end;
--- a/doc-src/iman.sty Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/iman.sty Sun Oct 24 15:11:24 2010 -0700
@@ -82,7 +82,8 @@
\def\dbend{\vtop to 0pt{\vss\hbox{\Huge\bf!}\vss}}
\newenvironment{warn}{\medskip\medbreak\begingroup \clubpenalty=10000
\small %%WAS\baselineskip=0.9\baselineskip
- \noindent \hangindent\parindent \hangafter=-2
+ \noindent \ifdim\parindent > 0pt\hangindent\parindent\else\hangindent1.5em\fi
+ \hangafter=-2
\hbox to0pt{\hskip-\hangindent\dbend\hfill}\ignorespaces}%
{\par\endgroup\medbreak}
--- a/doc-src/manual.bib Sun Oct 24 03:43:12 2010 -0700
+++ b/doc-src/manual.bib Sun Oct 24 15:11:24 2010 -0700
@@ -319,6 +319,18 @@
%C
+@InProceedings{Chaieb-Wenzel:2007,
+ author = {Amine Chaieb and Makarius Wenzel},
+ title = {Context aware Calculation and Deduction ---
+ Ring Equalities via {Gr\"obner Bases} in {Isabelle}},
+ booktitle = {Towards Mechanized Mathematical Assistants (CALCULEMUS 2007)},
+ editor = {Manuel Kauers and Manfred Kerber and Robert Miner and Wolfgang Windsteiger},
+ series = LNAI,
+ volume = 4573,
+ year = 2007,
+ publisher = Springer
+}
+
@TechReport{camilleri92,
author = {J. Camilleri and T. F. Melham},
title = {Reasoning with Inductively Defined Relations in the
@@ -1053,7 +1065,7 @@
@manual{isabelle-intro,
author = {Lawrence C. Paulson},
- title = {Introduction to {Isabelle}},
+ title = {Old Introduction to {Isabelle}},
institution = CUCL,
note = {\url{http://isabelle.in.tum.de/doc/intro.pdf}}}
@@ -1065,7 +1077,7 @@
@manual{isabelle-ref,
author = {Lawrence C. Paulson},
- title = {The {Isabelle} Reference Manual},
+ title = {The Old {Isabelle} Reference Manual},
institution = CUCL,
note = {\url{http://isabelle.in.tum.de/doc/ref.pdf}}}
@@ -1393,6 +1405,14 @@
year = 2000,
publisher = Springer}
+@Article{Sutter:2005,
+ author = {H. Sutter},
+ title = {The Free Lunch Is Over --- A Fundamental Turn Toward Concurrency in Software},
+ journal = {Dr. Dobb's Journal},
+ year = 2005,
+ volume = 30,
+ number = 3}
+
@InCollection{szasz93,
author = {Nora Szasz},
title = {A Machine Checked Proof that {Ackermann's} Function is not
@@ -1569,6 +1589,24 @@
note = {\url{http://www.in.tum.de/~wenzelm/papers/isar-framework.pdf}}
}
+@InProceedings{Wenzel-Chaieb:2007b,
+ author = {Makarius Wenzel and Amine Chaieb},
+ title = {{SML} with antiquotations embedded into {Isabelle/Isar}},
+ booktitle = {Workshop on Programming Languages for Mechanized Mathematics
+ (satellite of CALCULEMUS 2007). Hagenberg, Austria},
+ editor = {Jacques Carette and Freek Wiedijk},
+ month = {June},
+ year = {2007}
+}
+
+@InProceedings{Wenzel:2009,
+ author = {M. Wenzel},
+ title = {Parallel Proof Checking in {Isabelle/Isar}},
+ booktitle = {ACM SIGSAM Workshop on Programming Languages for Mechanized Mathematics Systems (PLMMS 2009)},
+ year = 2009,
+ editor = {Dos Reis, G. and L. Th\'ery},
+ publisher = {ACM Digital Library}}
+
@book{principia,
author = {A. N. Whitehead and B. Russell},
title = {Principia Mathematica},
@@ -1593,6 +1631,11 @@
note = {\url{http://www.cs.kun.nl/~freek/notes/mv.ps.gz}}
}
+@misc{wikipedia-2009-aa-trees,
+ key = "Wikipedia",
+ title = "Wikipedia: {AA} Tree",
+ note = "\url{http://en.wikipedia.org/wiki/AA_tree}"}
+
@book{winskel93,
author = {Glynn Winskel},
title = {The Formal Semantics of Programming Languages},
@@ -1611,6 +1654,11 @@
%Z
+@misc{z3,
+ key = "Z3",
+ title = "Z3: An Efficient {SMT} Solver",
+ note = "\url{http://research.microsoft.com/en-us/um/redmond/projects/z3/}"}
+
% CROSS REFERENCES
@@ -1855,8 +1903,3 @@
title = {{ML} Modules and {Haskell} Type Classes: A Constructive Comparison},
author = {Stefan Wehr et. al.}
}
-
-@misc{wikipedia-2009-aa-trees,
- key = "Wikipedia",
- title = "Wikipedia: {AA} Tree",
- note = "\url{http://en.wikipedia.org/wiki/AA_tree}"}
--- a/src/HOL/Decision_Procs/Commutative_Ring.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Decision_Procs/Commutative_Ring.thy Sun Oct 24 15:11:24 2010 -0700
@@ -266,7 +266,7 @@
proof cases
assume "even l"
then have "Suc l div 2 = l div 2"
- by (simp add: nat_number even_nat_plus_one_div_two)
+ by (simp add: eval_nat_numeral even_nat_plus_one_div_two)
moreover
from Suc have "l < k" by simp
with 1 have "\<And>P. Ipol ls (pow l P) = Ipol ls P ^ l" by simp
--- a/src/HOL/IsaMakefile Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/IsaMakefile Sun Oct 24 15:11:24 2010 -0700
@@ -327,8 +327,8 @@
Tools/Sledgehammer/sledgehammer_filter.ML \
Tools/Sledgehammer/sledgehammer_minimize.ML \
Tools/Sledgehammer/sledgehammer_isar.ML \
- Tools/Sledgehammer/sledgehammer_reconstruct.ML \
- Tools/Sledgehammer/sledgehammer_translate.ML \
+ Tools/Sledgehammer/sledgehammer_atp_reconstruct.ML \
+ Tools/Sledgehammer/sledgehammer_atp_translate.ML \
Tools/Sledgehammer/sledgehammer_util.ML \
Tools/SMT/cvc3_solver.ML \
Tools/SMT/smtlib_interface.ML \
--- a/src/HOL/Lazy_Sequence.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Lazy_Sequence.thy Sun Oct 24 15:11:24 2010 -0700
@@ -140,6 +140,71 @@
datatypes lazy_sequence = Lazy_Sequence
functions map yield yieldn
+subsection {* Generator Sequences *}
+
+
+subsubsection {* General lazy sequence operation *}
+
+definition product :: "'a Lazy_Sequence.lazy_sequence \<Rightarrow> 'b Lazy_Sequence.lazy_sequence \<Rightarrow> ('a * 'b) Lazy_Sequence.lazy_sequence"
+where
+ "product s1 s2 = Lazy_Sequence.bind s1 (%a. Lazy_Sequence.bind s2 (%b. Lazy_Sequence.single (a, b)))"
+
+
+subsubsection {* Small lazy typeclasses *}
+
+class small_lazy =
+ fixes small_lazy :: "int \<Rightarrow> 'a Lazy_Sequence.lazy_sequence"
+
+instantiation unit :: small_lazy
+begin
+
+definition "small_lazy d = Lazy_Sequence.single ()"
+
+instance ..
+
+end
+
+instantiation int :: small_lazy
+begin
+
+text {* maybe optimise this expression -> append (single x) xs == cons x xs
+Performance difference? *}
+
+function small_lazy' :: "int => int => int Lazy_Sequence.lazy_sequence"
+where "small_lazy' d i = (if d < i then Lazy_Sequence.empty else
+ Lazy_Sequence.append (Lazy_Sequence.single i) (small_lazy' d (i + 1)))"
+by pat_completeness auto
+
+termination
+ by (relation "measure (%(d, i). nat (d + 1 - i))") auto
+
+definition "small_lazy d = small_lazy' d (- d)"
+
+instance ..
+
+end
+
+instantiation prod :: (small_lazy, small_lazy) small_lazy
+begin
+
+definition
+ "small_lazy d = product (small_lazy d) (small_lazy d)"
+
+instance ..
+
+end
+
+instantiation list :: (small_lazy) small_lazy
+begin
+
+fun small_lazy_list :: "int => 'a list Lazy_Sequence.lazy_sequence"
+where
+ "small_lazy_list d = Lazy_Sequence.append (Lazy_Sequence.single []) (if d > 0 then Lazy_Sequence.bind (product (small_lazy (d - 1)) (small_lazy (d - 1))) (%(x, xs). Lazy_Sequence.single (x # xs)) else Lazy_Sequence.empty)"
+
+instance ..
+
+end
+
subsection {* With Hit Bound Value *}
text {* assuming in negative context *}
@@ -149,7 +214,6 @@
where
[code]: "hit_bound = Lazy_Sequence (%u. Some (None, empty))"
-
definition hb_single :: "'a => 'a hit_bound_lazy_sequence"
where
[code]: "hb_single x = Lazy_Sequence (%u. Some (Some x, empty))"
@@ -194,7 +258,10 @@
"hb_not_seq xq = (case yield xq of None => single () | Some (x, xq) => empty)"
hide_type (open) lazy_sequence
-hide_const (open) Empty Insert Lazy_Sequence yield empty single append flat map bind if_seq iterate_upto not_seq
-hide_fact yield.simps empty_def single_def append.simps flat.simps map.simps bind_def iterate_upto.simps if_seq_def not_seq_def
+hide_const (open) Empty Insert Lazy_Sequence yield empty single append flat map bind if_seq
+ iterate_upto not_seq product
+
+hide_fact yield.simps empty_def single_def append.simps flat.simps map.simps bind_def
+ iterate_upto.simps product_def if_seq_def not_seq_def
end
--- a/src/HOL/Library/Char_nat.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Library/Char_nat.thy Sun Oct 24 15:11:24 2010 -0700
@@ -171,7 +171,7 @@
proof -
fix n m
have nat_of_nibble_less_eq_15: "\<And>n. nat_of_nibble n \<le> 15"
- using Suc_leI [OF nat_of_nibble_less_16] by (auto simp add: nat_number)
+ using Suc_leI [OF nat_of_nibble_less_16] by (auto simp add: eval_nat_numeral)
have less_eq_240: "nat_of_nibble n * 16 \<le> 240"
using nat_of_nibble_less_eq_15 by auto
have "nat_of_nibble n * 16 + nat_of_nibble m \<le> 240 + 15"
--- a/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy Sun Oct 24 15:11:24 2010 -0700
@@ -96,8 +96,8 @@
val ioi = Fun (Input, Fun (Output, Fun (Input, Bool)))
val oii = Fun (Output, Fun (Input, Fun (Input, Bool)))
val ooi = Fun (Output, Fun (Output, Fun (Input, Bool)))
- val plus_nat = Predicate_Compile_Core.functional_compilation @{const_name plus} iio
- val minus_nat = Predicate_Compile_Core.functional_compilation @{const_name "minus"} iio
+ val plus_nat = Core_Data.functional_compilation @{const_name plus} iio
+ val minus_nat = Core_Data.functional_compilation @{const_name "minus"} iio
fun subtract_nat compfuns (_ : typ) =
let
val T = Predicate_Compile_Aux.mk_predT compfuns @{typ nat}
@@ -128,21 +128,21 @@
(single_const $ (@{term "op + :: nat => nat => nat"} $ Bound 1 $ Bound 0))))
end
in
- Predicate_Compile_Core.force_modes_and_compilations @{const_name plus_eq_nat}
+ Core_Data.force_modes_and_compilations @{const_name plus_eq_nat}
[(iio, (plus_nat, false)), (oii, (subtract_nat, false)), (ioi, (subtract_nat, false)),
(ooi, (enumerate_addups_nat, false))]
#> Predicate_Compile_Fun.add_function_predicate_translation
(@{term "plus :: nat => nat => nat"}, @{term "plus_eq_nat"})
- #> Predicate_Compile_Core.force_modes_and_compilations @{const_name minus_eq_nat}
+ #> Core_Data.force_modes_and_compilations @{const_name minus_eq_nat}
[(iio, (minus_nat, false)), (oii, (enumerate_nats, false))]
#> Predicate_Compile_Fun.add_function_predicate_translation
(@{term "minus :: nat => nat => nat"}, @{term "minus_eq_nat"})
- #> Predicate_Compile_Core.force_modes_and_functions @{const_name plus_eq_int}
+ #> Core_Data.force_modes_and_functions @{const_name plus_eq_int}
[(iio, (@{const_name plus}, false)), (ioi, (@{const_name subtract}, false)),
(oii, (@{const_name subtract}, false))]
#> Predicate_Compile_Fun.add_function_predicate_translation
(@{term "plus :: int => int => int"}, @{term "plus_eq_int"})
- #> Predicate_Compile_Core.force_modes_and_functions @{const_name minus_eq_int}
+ #> Core_Data.force_modes_and_functions @{const_name minus_eq_int}
[(iio, (@{const_name minus}, false)), (oii, (@{const_name plus}, false)),
(ioi, (@{const_name minus}, false))]
#> Predicate_Compile_Fun.add_function_predicate_translation
--- a/src/HOL/Library/Ramsey.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Library/Ramsey.thy Sun Oct 24 15:11:24 2010 -0700
@@ -218,7 +218,7 @@
"\<exists>Y t. Y \<subseteq> Z & infinite Y & t < s & (\<forall>x\<in>Y. \<forall>y\<in>Y. x\<noteq>y --> f{x,y} = t)"
proof -
have part2: "\<forall>X. X \<subseteq> Z & finite X & card X = 2 --> f X < s"
- using part by (fastsimp simp add: nat_number card_Suc_eq)
+ using part by (fastsimp simp add: eval_nat_numeral card_Suc_eq)
obtain Y t
where "Y \<subseteq> Z" "infinite Y" "t < s"
"(\<forall>X. X \<subseteq> Y & finite X & card X = 2 --> f X = t)"
--- a/src/HOL/List.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/List.thy Sun Oct 24 15:11:24 2010 -0700
@@ -262,9 +262,9 @@
@{lemma "[a,b,c,d][2 := x] = [a,b,x,d]" by simp}\\
@{lemma "sublist [a,b,c,d,e] {0,2,3} = [a,c,d]" by (simp add:sublist_def)}\\
@{lemma "rotate1 [a,b,c,d] = [b,c,d,a]" by (simp add:rotate1_def)}\\
-@{lemma "rotate 3 [a,b,c,d] = [d,a,b,c]" by (simp add:rotate1_def rotate_def nat_number')}\\
-@{lemma "replicate 4 a = [a,a,a,a]" by (simp add:nat_number')}\\
-@{lemma "[2..<5] = [2,3,4]" by (simp add:nat_number')}\\
+@{lemma "rotate 3 [a,b,c,d] = [d,a,b,c]" by (simp add:rotate1_def rotate_def eval_nat_numeral)}\\
+@{lemma "replicate 4 a = [a,a,a,a]" by (simp add:eval_nat_numeral)}\\
+@{lemma "[2..<5] = [2,3,4]" by (simp add:eval_nat_numeral)}\\
@{lemma "listsum [1,2,3::nat] = 6" by (simp add: listsum_def)}
\end{tabular}}
\caption{Characteristic examples}
--- a/src/HOL/MicroJava/BV/BVExample.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/MicroJava/BV/BVExample.thy Sun Oct 24 15:11:24 2010 -0700
@@ -258,7 +258,7 @@
lemma bounded_append [simp]:
"check_bounded append_ins [(Suc 0, 2, 8, Xcpt NullPointer)]"
apply (simp add: check_bounded_def)
- apply (simp add: nat_number append_ins_def)
+ apply (simp add: eval_nat_numeral append_ins_def)
apply (rule allI, rule impI)
apply (elim pc_end pc_next pc_0)
apply auto
@@ -327,7 +327,7 @@
lemma bounded_makelist [simp]: "check_bounded make_list_ins []"
apply (simp add: check_bounded_def)
- apply (simp add: nat_number make_list_ins_def)
+ apply (simp add: eval_nat_numeral make_list_ins_def)
apply (rule allI, rule impI)
apply (elim pc_end pc_next pc_0)
apply auto
@@ -343,7 +343,7 @@
"wt_method E test_name [] (PrimT Void) 3 2 make_list_ins [] \<phi>\<^sub>m"
apply (simp add: wt_method_def)
apply (simp add: make_list_ins_def phi_makelist_def)
- apply (simp add: wt_start_def nat_number)
+ apply (simp add: wt_start_def eval_nat_numeral)
apply (simp add: wt_instr_def)
apply clarify
apply (elim pc_end pc_next pc_0)
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Sun Oct 24 15:11:24 2010 -0700
@@ -29,8 +29,8 @@
lemmas: int,
max_lems: int,
time_isa: int,
- time_atp: int,
- time_atp_fail: int}
+ time_prover: int,
+ time_prover_fail: int}
datatype me_data = MeData of {
calls: int,
@@ -51,10 +51,11 @@
fun make_sh_data
(calls,success,nontriv_calls,nontriv_success,lemmas,max_lems,time_isa,
- time_atp,time_atp_fail) =
+ time_prover,time_prover_fail) =
ShData{calls=calls, success=success, nontriv_calls=nontriv_calls,
nontriv_success=nontriv_success, lemmas=lemmas, max_lems=max_lems,
- time_isa=time_isa, time_atp=time_atp, time_atp_fail=time_atp_fail}
+ time_isa=time_isa, time_prover=time_prover,
+ time_prover_fail=time_prover_fail}
fun make_min_data (succs, ab_ratios) =
MinData{succs=succs, ab_ratios=ab_ratios}
@@ -71,8 +72,8 @@
fun tuple_of_sh_data (ShData {calls, success, nontriv_calls, nontriv_success,
lemmas, max_lems, time_isa,
- time_atp, time_atp_fail}) = (calls, success, nontriv_calls, nontriv_success,
- lemmas, max_lems, time_isa, time_atp, time_atp_fail)
+ time_prover, time_prover_fail}) = (calls, success, nontriv_calls,
+ nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail)
fun tuple_of_min_data (MinData {succs, ab_ratios}) = (succs, ab_ratios)
@@ -127,40 +128,40 @@
fun inc_max (n:int) (s,sos,m) = (s+n, sos + n*n, Int.max(m,n));
val inc_sh_calls = map_sh_data
- (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_atp, time_atp_fail)
- => (calls + 1, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_atp, time_atp_fail))
+ (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)
+ => (calls + 1, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail))
val inc_sh_success = map_sh_data
- (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_atp, time_atp_fail)
- => (calls, success + 1, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_atp, time_atp_fail))
+ (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)
+ => (calls, success + 1, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail))
val inc_sh_nontriv_calls = map_sh_data
- (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_atp, time_atp_fail)
- => (calls, success, nontriv_calls + 1, nontriv_success, lemmas, max_lems, time_isa, time_atp, time_atp_fail))
+ (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)
+ => (calls, success, nontriv_calls + 1, nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail))
val inc_sh_nontriv_success = map_sh_data
- (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_atp, time_atp_fail)
- => (calls, success, nontriv_calls, nontriv_success + 1, lemmas,max_lems, time_isa, time_atp, time_atp_fail))
+ (fn (calls, success, nontriv_calls, nontriv_success, lemmas,max_lems, time_isa, time_prover, time_prover_fail)
+ => (calls, success, nontriv_calls, nontriv_success + 1, lemmas,max_lems, time_isa, time_prover, time_prover_fail))
fun inc_sh_lemmas n = map_sh_data
- (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_atp,time_atp_fail)
- => (calls,success,nontriv_calls, nontriv_success, lemmas+n,max_lems,time_isa,time_atp,time_atp_fail))
+ (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
+ => (calls,success,nontriv_calls, nontriv_success, lemmas+n,max_lems,time_isa,time_prover,time_prover_fail))
fun inc_sh_max_lems n = map_sh_data
- (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_atp,time_atp_fail)
- => (calls,success,nontriv_calls, nontriv_success, lemmas,Int.max(max_lems,n),time_isa,time_atp,time_atp_fail))
+ (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
+ => (calls,success,nontriv_calls, nontriv_success, lemmas,Int.max(max_lems,n),time_isa,time_prover,time_prover_fail))
fun inc_sh_time_isa t = map_sh_data
- (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_atp,time_atp_fail)
- => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa + t,time_atp,time_atp_fail))
+ (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
+ => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa + t,time_prover,time_prover_fail))
-fun inc_sh_time_atp t = map_sh_data
- (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_atp,time_atp_fail)
- => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_atp + t,time_atp_fail))
+fun inc_sh_time_prover t = map_sh_data
+ (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
+ => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover + t,time_prover_fail))
-fun inc_sh_time_atp_fail t = map_sh_data
- (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_atp,time_atp_fail)
- => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_atp,time_atp_fail + t))
+fun inc_sh_time_prover_fail t = map_sh_data
+ (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
+ => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail + t))
val inc_min_succs = map_min_data
(fn (succs,ab_ratios) => (succs+1, ab_ratios))
@@ -214,7 +215,7 @@
if n > 0 then (Real.fromInt t / 1000.0) / Real.fromInt n else 0.0
fun log_sh_data log
- (calls, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_atp, time_atp_fail) =
+ (calls, success, nontriv_calls, nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail) =
(log ("Total number of sledgehammer calls: " ^ str calls);
log ("Number of successful sledgehammer calls: " ^ str success);
log ("Number of sledgehammer lemmas: " ^ str lemmas);
@@ -223,14 +224,14 @@
log ("Total number of nontrivial sledgehammer calls: " ^ str nontriv_calls);
log ("Number of successful nontrivial sledgehammer calls: " ^ str nontriv_success);
log ("Total time for sledgehammer calls (Isabelle): " ^ str3 (time time_isa));
- log ("Total time for successful sledgehammer calls (ATP): " ^ str3 (time time_atp));
- log ("Total time for failed sledgehammer calls (ATP): " ^ str3 (time time_atp_fail));
+ log ("Total time for successful sledgehammer calls (ATP): " ^ str3 (time time_prover));
+ log ("Total time for failed sledgehammer calls (ATP): " ^ str3 (time time_prover_fail));
log ("Average time for sledgehammer calls (Isabelle): " ^
str3 (avg_time time_isa calls));
log ("Average time for successful sledgehammer calls (ATP): " ^
- str3 (avg_time time_atp success));
+ str3 (avg_time time_prover success));
log ("Average time for failed sledgehammer calls (ATP): " ^
- str3 (avg_time time_atp_fail (calls - success)))
+ str3 (avg_time time_prover_fail (calls - success)))
)
@@ -313,16 +314,17 @@
fun change_data id f = (Unsynchronized.change data (AList.map_entry (op =) id f); ())
-fun get_atp thy args =
+fun get_prover ctxt args =
let
- fun default_atp_name () =
- hd (#atps (Sledgehammer_Isar.default_params thy []))
+ val thy = ProofContext.theory_of ctxt
+ fun default_prover_name () =
+ hd (#provers (Sledgehammer_Isar.default_params ctxt []))
handle Empty => error "No ATP available."
- fun get_prover name = (name, Sledgehammer.get_prover_fun thy name)
+ fun get_prover name = (name, Sledgehammer.get_prover thy false name)
in
(case AList.lookup (op =) args proverK of
SOME name => get_prover name
- | NONE => get_prover (default_atp_name ()))
+ | NONE => get_prover (default_prover_name ()))
end
type locality = Sledgehammer_Filter.locality
@@ -346,31 +348,35 @@
(change_dir dir
#> Config.put Sledgehammer.measure_run_time true)
val params as {full_types, relevance_thresholds, max_relevant, ...} =
- Sledgehammer_Isar.default_params thy
+ Sledgehammer_Isar.default_params ctxt
[("timeout", Int.toString timeout ^ " s")]
+ val default_max_relevant =
+ Sledgehammer.default_max_relevant_for_prover thy prover_name
+ val irrelevant_consts =
+ Sledgehammer.irrelevant_consts_for_prover prover_name
+ val relevance_fudge = Sledgehammer.relevance_fudge_for_prover prover_name
val relevance_override = {add = [], del = [], only = false}
- val {default_max_relevant, ...} = ATP_Systems.get_prover thy prover_name
val (_, hyp_ts, concl_t) = Sledgehammer_Util.strip_subgoal goal i
val axioms =
Sledgehammer_Filter.relevant_facts ctxt full_types relevance_thresholds
- (the_default default_max_relevant max_relevant)
- relevance_override chained_ths hyp_ts concl_t
+ (the_default default_max_relevant max_relevant) irrelevant_consts
+ relevance_fudge relevance_override chained_ths hyp_ts concl_t
val problem =
{state = st', goal = goal, subgoal = i,
- axioms = map (Sledgehammer_Translate.prepare_axiom ctxt) axioms,
- only = true}
+ subgoal_count = Sledgehammer_Util.subgoal_count st,
+ axioms = axioms |> map Sledgehammer.Unprepared, only = true}
val time_limit =
(case hard_timeout of
NONE => I
| SOME secs => TimeLimit.timeLimit (Time.fromSeconds secs))
- val ({outcome, message, used_thm_names,
- atp_run_time_in_msecs = time_atp, ...} : Sledgehammer.prover_result,
+ val ({outcome, message, used_axioms, run_time_in_msecs = SOME time_prover, ...}
+ : Sledgehammer.prover_result,
time_isa) = time_limit (Mirabelle.cpu_time
(prover params (K ""))) problem
in
case outcome of
- NONE => (message, SH_OK (time_isa, time_atp, used_thm_names))
- | SOME _ => (message, SH_FAIL (time_isa, time_atp))
+ NONE => (message, SH_OK (time_isa, time_prover, used_axioms))
+ | SOME _ => (message, SH_FAIL (time_isa, time_prover))
end
handle ERROR msg => ("error: " ^ msg, SH_ERROR)
| TimeLimit.TimeOut => ("timeout", SH_ERROR)
@@ -395,7 +401,7 @@
val triv_str = if trivial then "[T] " else ""
val _ = change_data id inc_sh_calls
val _ = if trivial then () else change_data id inc_sh_nontriv_calls
- val (prover_name, prover) = get_atp (Proof.theory_of st) args
+ val (prover_name, prover) = get_prover (Proof.context_of st) args
val dir = AList.lookup (op =) args keepK
val timeout = Mirabelle.get_int_setting args (prover_timeoutK, 30)
val hard_timeout = AList.lookup (op =) args prover_hard_timeoutK
@@ -403,7 +409,7 @@
val (msg, result) = run_sh prover_name prover hard_timeout timeout dir st
in
case result of
- SH_OK (time_isa, time_atp, names) =>
+ SH_OK (time_isa, time_prover, names) =>
let
fun get_thms (_, Sledgehammer_Filter.Chained) = NONE
| get_thms (name, loc) =
@@ -414,38 +420,36 @@
change_data id (inc_sh_lemmas (length names));
change_data id (inc_sh_max_lems (length names));
change_data id (inc_sh_time_isa time_isa);
- change_data id (inc_sh_time_atp time_atp);
+ change_data id (inc_sh_time_prover time_prover);
named_thms := SOME (map_filter get_thms names);
log (sh_tag id ^ triv_str ^ "succeeded (" ^ string_of_int time_isa ^ "+" ^
- string_of_int time_atp ^ ") [" ^ prover_name ^ "]:\n" ^ msg)
+ string_of_int time_prover ^ ") [" ^ prover_name ^ "]:\n" ^ msg)
end
- | SH_FAIL (time_isa, time_atp) =>
+ | SH_FAIL (time_isa, time_prover) =>
let
val _ = change_data id (inc_sh_time_isa time_isa)
- val _ = change_data id (inc_sh_time_atp_fail time_atp)
+ val _ = change_data id (inc_sh_time_prover_fail time_prover)
in log (sh_tag id ^ triv_str ^ "failed: " ^ msg) end
| SH_ERROR => log (sh_tag id ^ "failed: " ^ msg)
end
end
-
-val subgoal_count = Logic.count_prems o prop_of o #goal o Proof.goal
-
fun run_minimize args named_thms id ({pre=st, log, ...}: Mirabelle.run_args) =
let
open Metis_Translate
- val thy = Proof.theory_of st
+ val ctxt = Proof.context_of st
val n0 = length (these (!named_thms))
- val (prover_name, _) = get_atp thy args
+ val (prover_name, _) = get_prover ctxt args
val timeout =
AList.lookup (op =) args minimize_timeoutK
|> Option.map (fst o read_int o explode)
|> the_default 5
- val params = Sledgehammer_Isar.default_params thy
- [("atps", prover_name), ("timeout", Int.toString timeout ^ " s")]
+ val params = Sledgehammer_Isar.default_params ctxt
+ [("provers", prover_name), ("timeout", Int.toString timeout ^ " s")]
val minimize =
- Sledgehammer_Minimize.minimize_theorems params 1 (subgoal_count st)
+ Sledgehammer_Minimize.minimize_facts params 1
+ (Sledgehammer_Util.subgoal_count st)
val _ = log separator
in
case minimize st (these (!named_thms)) of
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML Sun Oct 24 15:11:24 2010 -0700
@@ -5,26 +5,37 @@
structure Mirabelle_Sledgehammer_Filter : MIRABELLE_ACTION =
struct
-structure SF = Sledgehammer_Filter
+fun get args name default_value =
+ case AList.lookup (op =) args name of
+ SOME value => the (Real.fromString value)
+ | NONE => default_value
-val relevance_filter_args =
- [("worse_irrel_freq", SF.worse_irrel_freq),
- ("higher_order_irrel_weight", SF.higher_order_irrel_weight),
- ("abs_rel_weight", SF.abs_rel_weight),
- ("abs_irrel_weight", SF.abs_irrel_weight),
- ("skolem_irrel_weight", SF.skolem_irrel_weight),
- ("theory_const_rel_weight", SF.theory_const_rel_weight),
- ("theory_const_irrel_weight", SF.theory_const_irrel_weight),
- ("intro_bonus", SF.intro_bonus),
- ("elim_bonus", SF.elim_bonus),
- ("simp_bonus", SF.simp_bonus),
- ("local_bonus", SF.local_bonus),
- ("assum_bonus", SF.assum_bonus),
- ("chained_bonus", SF.chained_bonus),
- ("max_imperfect", SF.max_imperfect),
- ("max_imperfect_exp", SF.max_imperfect_exp),
- ("threshold_divisor", SF.threshold_divisor),
- ("ridiculous_threshold", SF.ridiculous_threshold)]
+fun extract_relevance_fudge args
+ {worse_irrel_freq, higher_order_irrel_weight, abs_rel_weight,
+ abs_irrel_weight, skolem_irrel_weight, theory_const_rel_weight,
+ theory_const_irrel_weight, intro_bonus, elim_bonus, simp_bonus,
+ local_bonus, assum_bonus, chained_bonus, max_imperfect,
+ max_imperfect_exp, threshold_divisor, ridiculous_threshold} =
+ {worse_irrel_freq = get args "worse_irrel_freq" worse_irrel_freq,
+ higher_order_irrel_weight =
+ get args "higher_order_irrel_weight" higher_order_irrel_weight,
+ abs_rel_weight = get args "abs_rel_weight" abs_rel_weight,
+ abs_irrel_weight = get args "abs_irrel_weight" abs_irrel_weight,
+ skolem_irrel_weight = get args "skolem_irrel_weight" skolem_irrel_weight,
+ theory_const_rel_weight =
+ get args "theory_const_rel_weight" theory_const_rel_weight,
+ theory_const_irrel_weight =
+ get args "theory_const_irrel_weight" theory_const_irrel_weight,
+ intro_bonus = get args "intro_bonus" intro_bonus,
+ elim_bonus = get args "elim_bonus" elim_bonus,
+ simp_bonus = get args "simp_bonus" simp_bonus,
+ local_bonus = get args "local_bonus" local_bonus,
+ assum_bonus = get args "assum_bonus" assum_bonus,
+ chained_bonus = get args "chained_bonus" chained_bonus,
+ max_imperfect = get args "max_imperfect" max_imperfect,
+ max_imperfect_exp = get args "max_imperfect_exp" max_imperfect_exp,
+ threshold_divisor = get args "threshold_divisor" threshold_divisor,
+ ridiculous_threshold = get args "ridiculous_threshold" ridiculous_threshold}
structure Prooftab =
Table(type key = int * int val ord = prod_ord int_ord int_ord);
@@ -85,6 +96,7 @@
()
val default_max_relevant = 300
+val prover_name = ATP_Systems.eN (* arbitrary ATP *)
fun with_index (i, s) = s ^ "@" ^ string_of_int i
@@ -95,23 +107,22 @@
SOME proofs =>
let
val {context = ctxt, facts, goal} = Proof.goal pre
- val thy = ProofContext.theory_of ctxt
- val args =
- args
- |> filter (fn (key, value) =>
- case AList.lookup (op =) relevance_filter_args key of
- SOME rf => (rf := the (Real.fromString value); false)
- | NONE => true)
-
+ val irrelevant_consts =
+ Sledgehammer.irrelevant_consts_for_prover prover_name
+ val relevance_fudge =
+ extract_relevance_fudge args
+ (Sledgehammer.relevance_fudge_for_prover prover_name)
+ val relevance_override = {add = [], del = [], only = false}
val {relevance_thresholds, full_types, max_relevant, ...} =
- Sledgehammer_Isar.default_params thy args
+ Sledgehammer_Isar.default_params ctxt args
val subgoal = 1
val (_, hyp_ts, concl_t) = Sledgehammer_Util.strip_subgoal goal subgoal
val facts =
- SF.relevant_facts ctxt full_types
+ Sledgehammer_Filter.relevant_facts ctxt full_types
relevance_thresholds
(the_default default_max_relevant max_relevant)
- {add = [], del = [], only = false} facts hyp_ts concl_t
+ irrelevant_consts relevance_fudge relevance_override facts hyp_ts
+ concl_t
|> map (fst o fst)
val (found_facts, lost_facts) =
List.concat proofs |> sort_distinct string_ord
--- a/src/HOL/Multivariate_Analysis/Determinants.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Multivariate_Analysis/Determinants.thy Sun Oct 24 15:11:24 2010 -0700
@@ -1042,9 +1042,9 @@
lemma setprod_1: "setprod f {(1::nat)..1} = f 1" by simp
lemma setprod_2: "setprod f {(1::nat)..2} = f 1 * f 2"
- by (simp add: nat_number setprod_numseg mult_commute)
+ by (simp add: eval_nat_numeral setprod_numseg mult_commute)
lemma setprod_3: "setprod f {(1::nat)..3} = f 1 * f 2 * f 3"
- by (simp add: nat_number setprod_numseg mult_commute)
+ by (simp add: eval_nat_numeral setprod_numseg mult_commute)
lemma det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1"
by (simp add: det_def sign_id UNIV_1)
--- a/src/HOL/Nat_Numeral.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Nat_Numeral.thy Sun Oct 24 15:11:24 2010 -0700
@@ -693,11 +693,7 @@
apply (simp only: nat_add_distrib, simp)
done
-lemmas nat_number =
- nat_number_of_Pls nat_number_of_Min
- nat_number_of_Bit0 nat_number_of_Bit1
-
-lemmas nat_number' =
+lemmas eval_nat_numeral =
nat_number_of_Bit0 nat_number_of_Bit1
lemmas nat_arith =
--- a/src/HOL/New_DSequence.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/New_DSequence.thy Sun Oct 24 15:11:24 2010 -0700
@@ -21,7 +21,11 @@
definition pos_bind :: "'a pos_dseq => ('a => 'b pos_dseq) => 'b pos_dseq"
where
- "pos_bind x f = (%i.
+ "pos_bind x f = (%i. Lazy_Sequence.bind (x i) (%a. f a i))"
+
+definition pos_decr_bind :: "'a pos_dseq => ('a => 'b pos_dseq) => 'b pos_dseq"
+where
+ "pos_decr_bind x f = (%i.
if i = 0 then
Lazy_Sequence.empty
else
@@ -57,7 +61,11 @@
definition neg_bind :: "'a neg_dseq => ('a => 'b neg_dseq) => 'b neg_dseq"
where
- "neg_bind x f = (%i.
+ "neg_bind x f = (%i. hb_bind (x i) (%a. f a i))"
+
+definition neg_decr_bind :: "'a neg_dseq => ('a => 'b neg_dseq) => 'b neg_dseq"
+where
+ "neg_decr_bind x f = (%i.
if i = 0 then
Lazy_Sequence.hit_bound
else
@@ -94,8 +102,8 @@
hide_type (open) pos_dseq neg_dseq
hide_const (open)
- pos_empty pos_single pos_bind pos_union pos_if_seq pos_iterate_upto pos_not_seq pos_map
- neg_empty neg_single neg_bind neg_union neg_if_seq neg_iterate_upto neg_not_seq neg_map
+ pos_empty pos_single pos_bind pos_decr_bind pos_union pos_if_seq pos_iterate_upto pos_not_seq pos_map
+ neg_empty neg_single neg_bind neg_decr_bind neg_union neg_if_seq neg_iterate_upto neg_not_seq neg_map
hide_fact (open)
pos_empty_def pos_single_def pos_bind_def pos_union_def pos_if_seq_def pos_iterate_upto_def pos_not_seq_def pos_map_def
neg_empty_def neg_single_def neg_bind_def neg_union_def neg_if_seq_def neg_iterate_upto_def neg_not_seq_def neg_map_def
--- a/src/HOL/New_Random_Sequence.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/New_Random_Sequence.thy Sun Oct 24 15:11:24 2010 -0700
@@ -20,6 +20,10 @@
where
"pos_bind R f = (\<lambda>nrandom size seed. New_DSequence.pos_bind (R nrandom size seed) (%a. f a nrandom size seed))"
+definition pos_decr_bind :: "'a pos_random_dseq => ('a \<Rightarrow> 'b pos_random_dseq) \<Rightarrow> 'b pos_random_dseq"
+where
+ "pos_decr_bind R f = (\<lambda>nrandom size seed. New_DSequence.pos_decr_bind (R nrandom size seed) (%a. f a nrandom size seed))"
+
definition pos_union :: "'a pos_random_dseq => 'a pos_random_dseq => 'a pos_random_dseq"
where
"pos_union R1 R2 = (\<lambda>nrandom size seed. New_DSequence.pos_union (R1 nrandom size seed) (R2 nrandom size seed))"
@@ -62,6 +66,10 @@
where
"neg_bind R f = (\<lambda>nrandom size seed. New_DSequence.neg_bind (R nrandom size seed) (%a. f a nrandom size seed))"
+definition neg_decr_bind :: "'a neg_random_dseq => ('a \<Rightarrow> 'b neg_random_dseq) \<Rightarrow> 'b neg_random_dseq"
+where
+ "neg_decr_bind R f = (\<lambda>nrandom size seed. New_DSequence.neg_decr_bind (R nrandom size seed) (%a. f a nrandom size seed))"
+
definition neg_union :: "'a neg_random_dseq => 'a neg_random_dseq => 'a neg_random_dseq"
where
"neg_union R1 R2 = (\<lambda>nrandom size seed. New_DSequence.neg_union (R1 nrandom size seed) (R2 nrandom size seed))"
@@ -97,8 +105,8 @@
*)
hide_const (open)
- pos_empty pos_single pos_bind pos_union pos_if_random_dseq pos_iterate_upto pos_not_random_dseq iter Random pos_map
- neg_empty neg_single neg_bind neg_union neg_if_random_dseq neg_iterate_upto neg_not_random_dseq neg_map
+ pos_empty pos_single pos_bind pos_decr_bind pos_union pos_if_random_dseq pos_iterate_upto pos_not_random_dseq iter Random pos_map
+ neg_empty neg_single neg_bind neg_decr_bind neg_union neg_if_random_dseq neg_iterate_upto neg_not_random_dseq neg_map
hide_type New_DSequence.pos_dseq New_DSequence.neg_dseq pos_random_dseq neg_random_dseq
(* FIXME: hide facts *)
(*hide_fact (open) empty_def single_def bind_def union_def if_random_dseq_def not_random_dseq_def Random.simps map_def*)
--- a/src/HOL/Predicate_Compile.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Predicate_Compile.thy Sun Oct 24 15:11:24 2010 -0700
@@ -9,6 +9,9 @@
uses
"Tools/Predicate_Compile/predicate_compile_aux.ML"
"Tools/Predicate_Compile/predicate_compile_compilations.ML"
+ "Tools/Predicate_Compile/core_data.ML"
+ "Tools/Predicate_Compile/mode_inference.ML"
+ "Tools/Predicate_Compile/predicate_compile_proof.ML"
"Tools/Predicate_Compile/predicate_compile_core.ML"
"Tools/Predicate_Compile/predicate_compile_data.ML"
"Tools/Predicate_Compile/predicate_compile_fun.ML"
--- a/src/HOL/Predicate_Compile_Examples/ROOT.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Predicate_Compile_Examples/ROOT.ML Sun Oct 24 15:11:24 2010 -0700
@@ -1,7 +1,7 @@
use_thys [
"Examples",
"Predicate_Compile_Tests",
- "Predicate_Compile_Quickcheck_Examples",
+(* "Predicate_Compile_Quickcheck_Examples", -- should be added again soon *)
"Specialisation_Examples",
"IMP_1",
"IMP_2",
--- a/src/HOL/Predicate_Compile_Examples/Specialisation_Examples.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Predicate_Compile_Examples/Specialisation_Examples.thy Sun Oct 24 15:11:24 2010 -0700
@@ -13,7 +13,7 @@
"greater_than_index xs = (\<forall>i x. nth_el' xs i = Some x --> x > i)"
code_pred (expected_modes: i => bool) [inductify, skip_proof, specialise] greater_than_index .
-ML {* Predicate_Compile_Core.intros_of @{context} @{const_name specialised_nth_el'P} *}
+ML {* Core_Data.intros_of @{context} @{const_name specialised_nth_el'P} *}
thm greater_than_index.equation
@@ -42,7 +42,7 @@
thm max_of_my_SucP.equation
-ML {* Predicate_Compile_Core.intros_of @{context} @{const_name specialised_max_natP} *}
+ML {* Core_Data.intros_of @{context} @{const_name specialised_max_natP} *}
values "{x. max_of_my_SucP x 6}"
--- a/src/HOL/Sledgehammer.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Sledgehammer.thy Sun Oct 24 15:11:24 2010 -0700
@@ -10,8 +10,8 @@
imports ATP
uses "Tools/Sledgehammer/sledgehammer_util.ML"
"Tools/Sledgehammer/sledgehammer_filter.ML"
- "Tools/Sledgehammer/sledgehammer_translate.ML"
- "Tools/Sledgehammer/sledgehammer_reconstruct.ML"
+ "Tools/Sledgehammer/sledgehammer_atp_translate.ML"
+ "Tools/Sledgehammer/sledgehammer_atp_reconstruct.ML"
"Tools/Sledgehammer/sledgehammer.ML"
"Tools/Sledgehammer/sledgehammer_minimize.ML"
"Tools/Sledgehammer/sledgehammer_isar.ML"
--- a/src/HOL/Tools/ATP/atp_systems.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/ATP/atp_systems.ML Sun Oct 24 15:11:24 2010 -0700
@@ -9,7 +9,7 @@
sig
type failure = ATP_Proof.failure
- type prover_config =
+ type atp_config =
{exec: string * string,
required_execs: (string * string) list,
arguments: bool -> Time.time -> string,
@@ -20,11 +20,17 @@
explicit_forall: bool,
use_conjecture_for_hypotheses: bool}
- val add_prover: string * prover_config -> theory -> theory
- val get_prover: theory -> string -> prover_config
- val available_atps: theory -> unit
+ val eN : string
+ val spassN : string
+ val vampireN : string
+ val sine_eN : string
+ val snarkN : string
+ val remote_prefix : string
+ val add_atp : string * atp_config -> theory -> theory
+ val get_atp : theory -> string -> atp_config
+ val available_atps : theory -> string list
+ val is_atp_installed : theory -> string -> bool
val refresh_systems_on_tptp : unit -> unit
- val default_atps_param_value : unit -> string
val setup : theory -> theory
end;
@@ -33,9 +39,9 @@
open ATP_Proof
-(* prover configuration *)
+(* ATP configuration *)
-type prover_config =
+type atp_config =
{exec: string * string,
required_execs: (string * string) list,
arguments: bool -> Time.time -> string,
@@ -51,33 +57,28 @@
(NoPerl, "env: perl"),
(NoLibwwwPerl, "Can't locate HTTP")]
-(* named provers *)
+(* named ATPs *)
+
+val eN = "e"
+val spassN = "spass"
+val vampireN = "vampire"
+val sine_eN = "sine_e"
+val snarkN = "snark"
+val remote_prefix = "remote_"
structure Data = Theory_Data
(
- type T = (prover_config * stamp) Symtab.table
+ type T = (atp_config * stamp) Symtab.table
val empty = Symtab.empty
val extend = I
fun merge data : T = Symtab.merge (eq_snd op =) data
handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".")
)
-fun add_prover (name, config) thy =
- Data.map (Symtab.update_new (name, (config, stamp ()))) thy
- handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".")
-
-fun get_prover thy name =
- the (Symtab.lookup (Data.get thy) name) |> fst
- handle Option.Option => error ("Unknown ATP: " ^ name ^ ".")
-
-fun available_atps thy =
- priority ("Available ATPs: " ^
- commas (sort_strings (Symtab.keys (Data.get thy))) ^ ".")
-
fun to_secs bonus time = (Time.toMilliseconds time + bonus + 999) div 1000
-(* E prover *)
+(* E *)
(* Give older versions of E an extra second, because the "eproof" script wrongly
subtracted an entire second to account for the overhead of the script
@@ -92,7 +93,7 @@
val tstp_proof_delims =
("# SZS output start CNFRefutation.", "# SZS output end CNFRefutation")
-val e_config : prover_config =
+val e_config : atp_config =
{exec = ("E_HOME", "eproof"),
required_execs = [],
arguments = fn _ => fn timeout =>
@@ -113,14 +114,14 @@
explicit_forall = false,
use_conjecture_for_hypotheses = true}
-val e = ("e", e_config)
+val e = (eN, e_config)
(* SPASS *)
(* The "-VarWeight=3" option helps the higher-order problems, probably by
counteracting the presence of "hAPP". *)
-val spass_config : prover_config =
+val spass_config : atp_config =
{exec = ("ISABELLE_ATP", "scripts/spass"),
required_execs = [("SPASS_HOME", "SPASS"), ("SPASS_HOME", "tptp2dfg")],
arguments = fn complete => fn timeout =>
@@ -142,12 +143,12 @@
explicit_forall = true,
use_conjecture_for_hypotheses = true}
-val spass = ("spass", spass_config)
+val spass = (spassN, spass_config)
(* Vampire *)
-val vampire_config : prover_config =
+val vampire_config : atp_config =
{exec = ("VAMPIRE_HOME", "vampire"),
required_execs = [],
arguments = fn complete => fn timeout =>
@@ -172,10 +173,10 @@
explicit_forall = false,
use_conjecture_for_hypotheses = true}
-val vampire = ("vampire", vampire_config)
+val vampire = (vampireN, vampire_config)
-(* Remote prover invocation via SystemOnTPTP *)
+(* Remote ATP invocation via SystemOnTPTP *)
val systems = Synchronized.var "atp_systems" ([] : string list)
@@ -187,9 +188,6 @@
SOME failure => string_for_failure failure
| NONE => perhaps (try (unsuffix "\n")) output ^ ".")
-fun refresh_systems_on_tptp () =
- Synchronized.change systems (fn _ => get_systems ())
-
fun find_system name [] systems = find_first (String.isPrefix name) systems
| find_system name (version :: versions) systems =
case find_first (String.isPrefix (name ^ "---" ^ version)) systems of
@@ -208,7 +206,7 @@
fun remote_config system_name system_versions proof_delims known_failures
default_max_relevant use_conjecture_for_hypotheses
- : prover_config =
+ : atp_config =
{exec = ("ISABELLE_ATP", "scripts/remote_atp"),
required_execs = [],
arguments = fn _ => fn timeout =>
@@ -225,48 +223,49 @@
fun remotify_config system_name system_versions
({proof_delims, known_failures, default_max_relevant,
- use_conjecture_for_hypotheses, ...} : prover_config) : prover_config =
+ use_conjecture_for_hypotheses, ...} : atp_config) : atp_config =
remote_config system_name system_versions proof_delims known_failures
default_max_relevant use_conjecture_for_hypotheses
-val remotify_name = prefix "remote_"
-fun remote_prover name system_name system_versions proof_delims known_failures
- default_max_relevant use_conjecture_for_hypotheses =
- (remotify_name name,
+fun remote_atp name system_name system_versions proof_delims known_failures
+ default_max_relevant use_conjecture_for_hypotheses =
+ (remote_prefix ^ name,
remote_config system_name system_versions proof_delims known_failures
default_max_relevant use_conjecture_for_hypotheses)
-fun remotify_prover (name, config) system_name system_versions =
- (remotify_name name, remotify_config system_name system_versions config)
+fun remotify_atp (name, config) system_name system_versions =
+ (remote_prefix ^ name, remotify_config system_name system_versions config)
-val remote_e = remotify_prover e "EP" ["1.0", "1.1", "1.2"]
-val remote_vampire = remotify_prover vampire "Vampire" ["0.6", "9.0", "1.0"]
+val remote_e = remotify_atp e "EP" ["1.0", "1.1", "1.2"]
+val remote_vampire = remotify_atp vampire "Vampire" ["0.6", "9.0", "1.0"]
val remote_sine_e =
- remote_prover "sine_e" "SInE" [] [] [(IncompleteUnprovable, "says Unknown")]
+ remote_atp sine_eN "SInE" [] [] [(IncompleteUnprovable, "says Unknown")]
800 (* FUDGE *) true
val remote_snark =
- remote_prover "snark" "SNARK---" [] [("refutation.", "end_refutation.")] []
- 250 (* FUDGE *) true
+ remote_atp snarkN "SNARK---" [] [("refutation.", "end_refutation.")] []
+ 250 (* FUDGE *) true
(* Setup *)
-fun is_installed ({exec, required_execs, ...} : prover_config) =
- forall (curry (op <>) "" o getenv o fst) (exec :: required_execs)
-fun maybe_remote (name, config) =
- name |> not (is_installed config) ? remotify_name
+fun add_atp (name, config) thy =
+ Data.map (Symtab.update_new (name, (config, stamp ()))) thy
+ handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".")
+
+fun get_atp thy name =
+ the (Symtab.lookup (Data.get thy) name) |> fst
+ handle Option.Option => error ("Unknown ATP: " ^ name ^ ".")
+
+val available_atps = Symtab.keys o Data.get
-(* The first prover of the list is used by Auto Sledgehammer. Because of the low
- timeout, it makes sense to put SPASS first. *)
-fun default_atps_param_value () =
- space_implode " " ((if is_installed (snd spass) then [fst spass] else []) @
- [maybe_remote e] @
- [if forall (is_installed o snd) [e, spass] then
- remotify_name (fst vampire)
- else
- maybe_remote vampire,
- fst remote_sine_e])
+fun is_atp_installed thy name =
+ let val {exec, required_execs, ...} = get_atp thy name in
+ forall (curry (op <>) "" o getenv o fst) (exec :: required_execs)
+ end
-val provers = [e, spass, vampire, remote_e, remote_vampire, remote_sine_e,
- remote_snark]
-val setup = fold add_prover provers
+fun refresh_systems_on_tptp () =
+ Synchronized.change systems (fn _ => get_systems ())
+
+val atps = [e, spass, vampire, remote_e, remote_vampire, remote_sine_e,
+ remote_snark]
+val setup = fold add_atp atps
end;
--- a/src/HOL/Tools/Function/function.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/Function/function.ML Sun Oct 24 15:11:24 2010 -0700
@@ -108,8 +108,9 @@
lthy
|> addsmps (conceal_partial o Binding.qualify false "partial")
"psimps" conceal_partial psimp_attribs psimps
- ||> fold_option (snd oo addsmps I "simps" I simp_attribs) trsimps
- ||> fold_option (Spec_Rules.add Spec_Rules.Equational o pair fs) trsimps
+ ||> (case trsimps of NONE => I | SOME thms =>
+ addsmps I "simps" I simp_attribs thms #> snd
+ #> Spec_Rules.add Spec_Rules.Equational (fs, thms))
||>> Local_Theory.note ((conceal_partial (qualify "pinduct"),
[Attrib.internal (K (Rule_Cases.case_names cnames)),
Attrib.internal (K (Rule_Cases.consumes 1)),
@@ -117,7 +118,8 @@
||>> Local_Theory.note ((Binding.conceal (qualify "termination"), []), [termination])
||> (snd o Local_Theory.note ((qualify "cases",
[Attrib.internal (K (Rule_Cases.case_names cnames))]), [cases]))
- ||> fold_option (snd oo curry Local_Theory.note (qualify "domintros", [])) domintros
+ ||> (case domintros of NONE => I | SOME thms =>
+ Local_Theory.note ((qualify "domintros", []), thms) #> snd)
val info = { add_simps=addsmps, case_names=cnames, psimps=psimps',
pinducts=snd pinducts', simps=NONE, inducts=NONE, termination=termination',
--- a/src/HOL/Tools/Function/function_lib.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/Function/function_lib.ML Sun Oct 24 15:11:24 2010 -0700
@@ -8,29 +8,15 @@
structure Function_Lib = (* FIXME proper signature *)
struct
-fun map_option f NONE = NONE
- | map_option f (SOME x) = SOME (f x);
-
-fun fold_option f NONE y = y
- | fold_option f (SOME x) y = f x y;
-
-(* Ex: "The variable" ^ plural " is" "s are" vs *)
+(* "The variable" ^ plural " is" "s are" vs *)
fun plural sg pl [x] = sg
| plural sg pl _ = pl
-fun dest_all (Const ("all", _) $ Abs (a as (_,T,_))) =
- let
- val (n, body) = Term.dest_abs a
- in
- (Free (n, T), body)
- end
- | dest_all _ = raise Match
-
(* Removes all quantifiers from a term, replacing bound variables by frees. *)
fun dest_all_all (t as (Const ("all",_) $ _)) =
let
- val (v,b) = dest_all t
+ val (v,b) = Logic.dest_all t |> apfst Free
val (vs, b') = dest_all_all b
in
(v :: vs, b')
@@ -56,18 +42,10 @@
(ctx, [], t)
-fun map3 _ [] [] [] = []
- | map3 f (x :: xs) (y :: ys) (z :: zs) = f x y z :: map3 f xs ys zs
- | map3 _ _ _ _ = raise Library.UnequalLengths;
-
fun map4 _ [] [] [] [] = []
| map4 f (x :: xs) (y :: ys) (z :: zs) (u :: us) = f x y z u :: map4 f xs ys zs us
| map4 _ _ _ _ _ = raise Library.UnequalLengths;
-fun map6 _ [] [] [] [] [] [] = []
- | map6 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) = f x y z u v w :: map6 f xs ys zs us vs ws
- | map6 _ _ _ _ _ _ _ = raise Library.UnequalLengths;
-
fun map7 _ [] [] [] [] [] [] [] = []
| map7 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) (b :: bs) = f x y z u v w b :: map7 f xs ys zs us vs ws bs
| map7 _ _ _ _ _ _ _ _ = raise Library.UnequalLengths;
--- a/src/HOL/Tools/Function/mutual.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/Function/mutual.ML Sun Oct 24 15:11:24 2010 -0700
@@ -267,10 +267,10 @@
val rew_ss = HOL_basic_ss addsimps all_f_defs
val mpsimps = map2 mk_mpsimp fqgars psimps
- val mtrsimps = map_option (map2 mk_mpsimp fqgars) trsimps
+ val mtrsimps = Option.map (map2 mk_mpsimp fqgars) trsimps
val minducts = mutual_induct_rules lthy simple_pinduct all_f_defs m
val mtermination = full_simplify rew_ss termination
- val mdomintros = map_option (map (full_simplify rew_ss)) domintros
+ val mdomintros = Option.map (map (full_simplify rew_ss)) domintros
in
FunctionResult { fs=fs, G=G, R=R,
psimps=mpsimps, simple_pinducts=minducts,
--- a/src/HOL/Tools/Function/relation.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/Function/relation.ML Sun Oct 24 15:11:24 2010 -0700
@@ -14,24 +14,32 @@
structure Function_Relation : FUNCTION_RELATION =
struct
-fun inst_state_tac ctxt rel st =
- let
- val cert = Thm.cterm_of (ProofContext.theory_of ctxt)
- val rel' = cert (singleton (Variable.polymorphic ctxt) rel)
- val st' = Thm.incr_indexes (#maxidx (Thm.rep_cterm rel') + 1) st
- in case Term.add_vars (prop_of st') [] of
- [v] =>
- PRIMITIVE (Drule.cterm_instantiate [(cert (Var v), rel')]) st'
- | _ => Seq.empty
- end
+fun gen_inst_state_tac prep ctxt rel st =
+ case Term.add_vars (prop_of st) [] of
+ [v as (_, T)] =>
+ let
+ val cert = Thm.cterm_of (ProofContext.theory_of ctxt)
+ val rel' = prep T (singleton (Variable.polymorphic ctxt) rel)
+ |> cert
+ val st' = Thm.incr_indexes (#maxidx (Thm.rep_cterm rel') + 1) st (*FIXME??*)
+ in
+ PRIMITIVE (Drule.cterm_instantiate [(cert (Var v), rel')]) st'
+ end
+ | _ => Seq.empty;
-fun relation_tac ctxt rel i =
+fun gen_relation_tac prep ctxt rel i =
TRY (Function_Common.apply_termination_rule ctxt i)
- THEN inst_state_tac ctxt rel
+ THEN gen_inst_state_tac prep ctxt rel
+
+val relation_tac = gen_relation_tac (K I)
+
+fun relation_meth rel ctxt = SIMPLE_METHOD'
+ (gen_relation_tac (fn T => map_types Type_Infer.paramify_vars
+ #> Type.constraint T #> Syntax.check_term ctxt)
+ ctxt rel)
val setup =
- Method.setup @{binding relation}
- (Args.term >> (fn rel => fn ctxt => SIMPLE_METHOD' (relation_tac ctxt rel)))
+ Method.setup @{binding relation} (Args.term >> relation_meth)
"proves termination using a user-specified wellfounded relation"
end
--- a/src/HOL/Tools/Nitpick/nitpick_mono.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/Nitpick/nitpick_mono.ML Sun Oct 24 15:11:24 2010 -0700
@@ -448,8 +448,7 @@
fun add_notin_mtype_fv sn M ((lits, comps, sexps) : constraint_set) =
(trace_msg (fn () => "*** Add " ^ string_for_mtype M ^ " is " ^
- (case sn of Minus => "concrete" | Plus => "complete") ^
- ".");
+ (case sn of Minus => "concrete" | Plus => "complete"));
case do_notin_mtype_fv sn [] M (SOME (lits, sexps)) of
NONE => (trace_msg (K "**** Unsolvable"); raise UNSOLVABLE ())
| SOME (lits, sexps) => (lits, comps, sexps))
--- a/src/HOL/Tools/Predicate_Compile/code_prolog.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/Predicate_Compile/code_prolog.ML Sun Oct 24 15:11:24 2010 -0700
@@ -386,9 +386,9 @@
(fn s => member (op =) (the (AList.lookup (op =) random s))))
val (preds, all_vs, param_vs, all_modes, clauses) =
Predicate_Compile_Core.prepare_intrs options ctxt prednames
- (maps (Predicate_Compile_Core.intros_of ctxt) prednames)
+ (maps (Core_Data.intros_of ctxt) prednames)
val ((moded_clauses, random'), _) =
- Predicate_Compile_Core.infer_modes mode_analysis_options options
+ Mode_Inference.infer_modes mode_analysis_options options
(lookup_modes, lookup_neg_modes, needs_random) ctxt preds all_modes param_vs clauses
val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
val pos_modes' = map (apsnd (map_filter (fn (true, m) => SOME m | _ => NONE))) modes
@@ -474,7 +474,7 @@
let
fun strong_conn_of gr keys =
Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
- val gr = Predicate_Compile_Core.intros_graph_of ctxt
+ val gr = Core_Data.intros_graph_of ctxt
val gr' = add_edges depending_preds_of const gr
val scc = strong_conn_of gr' [const]
val initial_constant_table =
@@ -917,6 +917,7 @@
no_higher_order_predicate = [],
inductify = false,
detect_switches = true,
+ smart_depth_limiting = true,
compilation = Predicate_Compile_Aux.Pred
}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/core_data.ML Sun Oct 24 15:11:24 2010 -0700
@@ -0,0 +1,593 @@
+(* Title: HOL/Tools/Predicate_Compile/core_data.ML
+ Author: Lukas Bulwahn, TU Muenchen
+
+Data of the predicate compiler core
+
+*)
+signature CORE_DATA =
+sig
+ type mode = Predicate_Compile_Aux.mode
+ type compilation = Predicate_Compile_Aux.compilation
+ type compilation_funs = Predicate_Compile_Aux.compilation_funs
+
+ datatype predfun_data = PredfunData of {
+ definition : thm,
+ intro : thm,
+ elim : thm,
+ neg_intro : thm option
+ };
+
+ datatype pred_data = PredData of {
+ intros : (string option * thm) list,
+ elim : thm option,
+ preprocessed : bool,
+ function_names : (compilation * (mode * string) list) list,
+ predfun_data : (mode * predfun_data) list,
+ needs_random : mode list
+ };
+
+ (* general operations *)
+ val unify_consts : theory -> term list -> term list -> (term list * term list)
+ val mk_casesrule : Proof.context -> term -> thm list -> term
+ val preprocess_intro : theory -> thm -> thm
+
+ structure PredData : THEORY_DATA
+
+ (* queries *)
+ val defined_functions : compilation -> Proof.context -> string -> bool
+ val is_registered : Proof.context -> string -> bool
+ val function_name_of : compilation -> Proof.context -> string -> mode -> string
+ val the_elim_of : Proof.context -> string -> thm
+ val has_elim : Proof.context -> string -> bool
+
+ val needs_random : Proof.context -> string -> mode -> bool
+
+ val predfun_intro_of : Proof.context -> string -> mode -> thm
+ val predfun_neg_intro_of : Proof.context -> string -> mode -> thm option
+ val predfun_elim_of : Proof.context -> string -> mode -> thm
+ val predfun_definition_of : Proof.context -> string -> mode -> thm
+
+ val all_preds_of : Proof.context -> string list
+ val modes_of: compilation -> Proof.context -> string -> mode list
+ val all_modes_of : compilation -> Proof.context -> (string * mode list) list
+ val all_random_modes_of : Proof.context -> (string * mode list) list
+ val intros_of : Proof.context -> string -> thm list
+ val names_of : Proof.context -> string -> string option list
+
+ val intros_graph_of : Proof.context -> thm list Graph.T
+
+ (* updaters *)
+
+ val register_predicate : (string * thm list * thm) -> theory -> theory
+ val register_intros : string * thm list -> theory -> theory
+
+ (* FIXME: naming of function is strange *)
+ val defined_function_of : compilation -> string -> theory -> theory
+ val add_intro : string option * thm -> theory -> theory
+ val set_elim : thm -> theory -> theory
+ val set_function_name : compilation -> string -> mode -> string -> theory -> theory
+ val add_predfun_data : string -> mode -> thm * ((thm * thm) * thm option) -> theory -> theory
+ val set_needs_random : string -> mode list -> theory -> theory
+ (* sophisticated updaters *)
+ val extend_intro_graph : string list -> theory -> theory
+ val preprocess_intros : string -> theory -> theory
+
+ (* alternative function definitions *)
+ val register_alternative_function : string -> mode -> string -> theory -> theory
+ val alternative_compilation_of_global : theory -> string -> mode ->
+ (compilation_funs -> typ -> term) option
+ val alternative_compilation_of : Proof.context -> string -> mode ->
+ (compilation_funs -> typ -> term) option
+ val functional_compilation : string -> mode -> compilation_funs -> typ -> term
+ val force_modes_and_functions : string -> (mode * (string * bool)) list -> theory -> theory
+ val force_modes_and_compilations : string ->
+ (mode * ((compilation_funs -> typ -> term) * bool)) list -> theory -> theory
+
+end;
+
+structure Core_Data : CORE_DATA =
+struct
+
+open Predicate_Compile_Aux;
+
+(* FIXME: should be moved to Predicate_Compile_Aux *)
+val strip_intro_concl = (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl o prop_of)
+
+
+(* book-keeping *)
+
+datatype predfun_data = PredfunData of {
+ definition : thm,
+ intro : thm,
+ elim : thm,
+ neg_intro : thm option
+};
+
+fun rep_predfun_data (PredfunData data) = data;
+
+fun mk_predfun_data (definition, ((intro, elim), neg_intro)) =
+ PredfunData {definition = definition, intro = intro, elim = elim, neg_intro = neg_intro}
+
+datatype pred_data = PredData of {
+ intros : (string option * thm) list,
+ elim : thm option,
+ preprocessed : bool,
+ function_names : (compilation * (mode * string) list) list,
+ predfun_data : (mode * predfun_data) list,
+ needs_random : mode list
+};
+
+fun rep_pred_data (PredData data) = data;
+
+fun mk_pred_data (((intros, elim), preprocessed), (function_names, (predfun_data, needs_random))) =
+ PredData {intros = intros, elim = elim, preprocessed = preprocessed,
+ function_names = function_names, predfun_data = predfun_data, needs_random = needs_random}
+
+fun map_pred_data f (PredData {intros, elim, preprocessed, function_names, predfun_data, needs_random}) =
+ mk_pred_data (f (((intros, elim), preprocessed), (function_names, (predfun_data, needs_random))))
+
+fun eq_option eq (NONE, NONE) = true
+ | eq_option eq (SOME x, SOME y) = eq (x, y)
+ | eq_option eq _ = false
+
+fun eq_pred_data (PredData d1, PredData d2) =
+ eq_list (eq_pair (op =) Thm.eq_thm) (#intros d1, #intros d2) andalso
+ eq_option Thm.eq_thm (#elim d1, #elim d2)
+
+structure PredData = Theory_Data
+(
+ type T = pred_data Graph.T;
+ val empty = Graph.empty;
+ val extend = I;
+ val merge = Graph.merge eq_pred_data;
+);
+
+(* queries *)
+
+fun lookup_pred_data ctxt name =
+ Option.map rep_pred_data (try (Graph.get_node (PredData.get (ProofContext.theory_of ctxt))) name)
+
+fun the_pred_data ctxt name = case lookup_pred_data ctxt name
+ of NONE => error ("No such predicate " ^ quote name)
+ | SOME data => data;
+
+val is_registered = is_some oo lookup_pred_data
+
+val all_preds_of = Graph.keys o PredData.get o ProofContext.theory_of
+
+val intros_of = map snd o #intros oo the_pred_data
+
+val names_of = map fst o #intros oo the_pred_data
+
+fun the_elim_of ctxt name = case #elim (the_pred_data ctxt name)
+ of NONE => error ("No elimination rule for predicate " ^ quote name)
+ | SOME thm => thm
+
+val has_elim = is_some o #elim oo the_pred_data
+
+fun function_names_of compilation ctxt name =
+ case AList.lookup (op =) (#function_names (the_pred_data ctxt name)) compilation of
+ NONE => error ("No " ^ string_of_compilation compilation
+ ^ " functions defined for predicate " ^ quote name)
+ | SOME fun_names => fun_names
+
+fun function_name_of compilation ctxt name mode =
+ case AList.lookup eq_mode
+ (function_names_of compilation ctxt name) mode of
+ NONE => error ("No " ^ string_of_compilation compilation
+ ^ " function defined for mode " ^ string_of_mode mode ^ " of predicate " ^ quote name)
+ | SOME function_name => function_name
+
+fun modes_of compilation ctxt name = map fst (function_names_of compilation ctxt name)
+
+fun all_modes_of compilation ctxt =
+ map_filter (fn name => Option.map (pair name) (try (modes_of compilation ctxt) name))
+ (all_preds_of ctxt)
+
+val all_random_modes_of = all_modes_of Random
+
+fun defined_functions compilation ctxt name = case lookup_pred_data ctxt name of
+ NONE => false
+ | SOME data => AList.defined (op =) (#function_names data) compilation
+
+fun needs_random ctxt s m =
+ member (op =) (#needs_random (the_pred_data ctxt s)) m
+
+fun lookup_predfun_data ctxt name mode =
+ Option.map rep_predfun_data
+ (AList.lookup (op =) (#predfun_data (the_pred_data ctxt name)) mode)
+
+fun the_predfun_data ctxt name mode =
+ case lookup_predfun_data ctxt name mode of
+ NONE => error ("No function defined for mode " ^ string_of_mode mode ^
+ " of predicate " ^ name)
+ | SOME data => data;
+
+val predfun_definition_of = #definition ooo the_predfun_data
+
+val predfun_intro_of = #intro ooo the_predfun_data
+
+val predfun_elim_of = #elim ooo the_predfun_data
+
+val predfun_neg_intro_of = #neg_intro ooo the_predfun_data
+
+val intros_graph_of =
+ Graph.map (K (map snd o #intros o rep_pred_data)) o PredData.get o ProofContext.theory_of
+
+(** preprocessing rules **)
+
+fun Trueprop_conv cv ct =
+ case Thm.term_of ct of
+ Const (@{const_name Trueprop}, _) $ _ => Conv.arg_conv cv ct
+ | _ => raise Fail "Trueprop_conv"
+
+fun preprocess_equality thy rule =
+ Conv.fconv_rule
+ (imp_prems_conv
+ (Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq})))))
+ (Thm.transfer thy rule)
+
+fun preprocess_intro thy = expand_tuples thy #> preprocess_equality thy
+
+(* importing introduction rules *)
+
+fun unify_consts thy cs intr_ts =
+ (let
+ val add_term_consts_2 = fold_aterms (fn Const c => insert (op =) c | _ => I);
+ fun varify (t, (i, ts)) =
+ let val t' = map_types (Logic.incr_tvar (i + 1)) (#2 (Type.varify_global [] t))
+ in (maxidx_of_term t', t'::ts) end;
+ val (i, cs') = List.foldr varify (~1, []) cs;
+ val (i', intr_ts') = List.foldr varify (i, []) intr_ts;
+ val rec_consts = fold add_term_consts_2 cs' [];
+ val intr_consts = fold add_term_consts_2 intr_ts' [];
+ fun unify (cname, cT) =
+ let val consts = map snd (filter (fn c => fst c = cname) intr_consts)
+ in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end;
+ val (env, _) = fold unify rec_consts (Vartab.empty, i');
+ val subst = map_types (Envir.norm_type env)
+ in (map subst cs', map subst intr_ts')
+ end) handle Type.TUNIFY =>
+ (warning "Occurrences of recursive constant have non-unifiable types"; (cs, intr_ts));
+
+fun import_intros inp_pred [] ctxt =
+ let
+ val ([outp_pred], ctxt') = Variable.import_terms true [inp_pred] ctxt
+ val T = fastype_of outp_pred
+ val paramTs = ho_argsT_of_typ (binder_types T)
+ val (param_names, ctxt'') = Variable.variant_fixes
+ (map (fn i => "p" ^ (string_of_int i)) (1 upto (length paramTs))) ctxt'
+ val params = map2 (curry Free) param_names paramTs
+ in
+ (((outp_pred, params), []), ctxt')
+ end
+ | import_intros inp_pred (th :: ths) ctxt =
+ let
+ val ((_, [th']), ctxt') = Variable.import true [th] ctxt
+ val thy = ProofContext.theory_of ctxt'
+ val (pred, args) = strip_intro_concl th'
+ val T = fastype_of pred
+ val ho_args = ho_args_of_typ T args
+ fun subst_of (pred', pred) =
+ let
+ val subst = Sign.typ_match thy (fastype_of pred', fastype_of pred) Vartab.empty
+ handle Type.TYPE_MATCH => error ("Type mismatch of predicate " ^ fst (dest_Const pred)
+ ^ " (trying to match " ^ Syntax.string_of_typ ctxt (fastype_of pred')
+ ^ " and " ^ Syntax.string_of_typ ctxt (fastype_of pred) ^ ")"
+ ^ " in " ^ Display.string_of_thm ctxt th)
+ in map (fn (indexname, (s, T)) => ((indexname, s), T)) (Vartab.dest subst) end
+ fun instantiate_typ th =
+ let
+ val (pred', _) = strip_intro_concl th
+ val _ = if not (fst (dest_Const pred) = fst (dest_Const pred')) then
+ raise Fail "Trying to instantiate another predicate" else ()
+ in Thm.certify_instantiate (subst_of (pred', pred), []) th end;
+ fun instantiate_ho_args th =
+ let
+ val (_, args') = (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl o prop_of) th
+ val ho_args' = map dest_Var (ho_args_of_typ T args')
+ in Thm.certify_instantiate ([], ho_args' ~~ ho_args) th end
+ val outp_pred =
+ Term_Subst.instantiate (subst_of (inp_pred, pred), []) inp_pred
+ val ((_, ths'), ctxt1) =
+ Variable.import false (map (instantiate_typ #> instantiate_ho_args) ths) ctxt'
+ in
+ (((outp_pred, ho_args), th' :: ths'), ctxt1)
+ end
+
+(* generation of case rules from user-given introduction rules *)
+
+fun mk_args2 (Type (@{type_name Product_Type.prod}, [T1, T2])) st =
+ let
+ val (t1, st') = mk_args2 T1 st
+ val (t2, st'') = mk_args2 T2 st'
+ in
+ (HOLogic.mk_prod (t1, t2), st'')
+ end
+ (*| mk_args2 (T as Type ("fun", _)) (params, ctxt) =
+ let
+ val (S, U) = strip_type T
+ in
+ if U = HOLogic.boolT then
+ (hd params, (tl params, ctxt))
+ else
+ let
+ val ([x], ctxt') = Variable.variant_fixes ["x"] ctxt
+ in
+ (Free (x, T), (params, ctxt'))
+ end
+ end*)
+ | mk_args2 T (params, ctxt) =
+ let
+ val ([x], ctxt') = Variable.variant_fixes ["x"] ctxt
+ in
+ (Free (x, T), (params, ctxt'))
+ end
+
+fun mk_casesrule ctxt pred introrules =
+ let
+ (* TODO: can be simplified if parameters are not treated specially ? *)
+ val (((pred, params), intros_th), ctxt1) = import_intros pred introrules ctxt
+ (* TODO: distinct required ? -- test case with more than one parameter! *)
+ val params = distinct (op aconv) params
+ val intros = map prop_of intros_th
+ val ([propname], ctxt2) = Variable.variant_fixes ["thesis"] ctxt1
+ val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT))
+ val argsT = binder_types (fastype_of pred)
+ (* TODO: can be simplified if parameters are not treated specially ? <-- see uncommented code! *)
+ val (argvs, _) = fold_map mk_args2 argsT (params, ctxt2)
+ fun mk_case intro =
+ let
+ val (_, args) = (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl) intro
+ val prems = Logic.strip_imp_prems intro
+ val eqprems =
+ map2 (HOLogic.mk_Trueprop oo (curry HOLogic.mk_eq)) argvs args
+ val frees = map Free (fold Term.add_frees (args @ prems) [])
+ in fold Logic.all frees (Logic.list_implies (eqprems @ prems, prop)) end
+ val assm = HOLogic.mk_Trueprop (list_comb (pred, argvs))
+ val cases = map mk_case intros
+ in Logic.list_implies (assm :: cases, prop) end;
+
+fun prove_casesrule ctxt (pred, (pre_cases_rule, nparams)) cases_rule =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val nargs = length (binder_types (fastype_of pred))
+ fun PEEK f dependent_tactic st = dependent_tactic (f st) st
+ fun meta_eq_of th = th RS @{thm eq_reflection}
+ val tuple_rew_rules = map meta_eq_of [@{thm fst_conv}, @{thm snd_conv}, @{thm Pair_eq}]
+ fun instantiate i n {context = ctxt, params = p, prems = prems,
+ asms = a, concl = cl, schematics = s} =
+ let
+ fun term_pair_of (ix, (ty,t)) = (Var (ix,ty), t)
+ fun inst_of_matches tts = fold (Pattern.match thy) tts (Vartab.empty, Vartab.empty)
+ |> snd |> Vartab.dest |> map (pairself (cterm_of thy) o term_pair_of)
+ val (cases, (eqs, prems)) = apsnd (chop (nargs - nparams)) (chop n prems)
+ val case_th = MetaSimplifier.simplify true
+ (@{thm Predicate.eq_is_eq} :: map meta_eq_of eqs) (nth cases (i - 1))
+ val prems' = maps (dest_conjunct_prem o MetaSimplifier.simplify true tuple_rew_rules) prems
+ val pats = map (swap o HOLogic.dest_eq o HOLogic.dest_Trueprop) (take nargs (prems_of case_th))
+ val case_th' = Thm.instantiate ([], inst_of_matches pats) case_th
+ OF (replicate nargs @{thm refl})
+ val thesis =
+ Thm.instantiate ([], inst_of_matches (prems_of case_th' ~~ map prop_of prems')) case_th'
+ OF prems'
+ in
+ (rtac thesis 1)
+ end
+ val tac =
+ etac pre_cases_rule 1
+ THEN
+ (PEEK nprems_of
+ (fn n =>
+ ALLGOALS (fn i =>
+ MetaSimplifier.rewrite_goal_tac [@{thm split_paired_all}] i
+ THEN (SUBPROOF (instantiate i n) ctxt i))))
+ in
+ Goal.prove ctxt (Term.add_free_names cases_rule []) [] cases_rule (fn _ => tac)
+ end
+
+(* updaters *)
+
+(* fetching introduction rules or registering introduction rules *)
+
+val no_compilation = ([], ([], []))
+
+fun fetch_pred_data ctxt name =
+ case try (Inductive.the_inductive ctxt) name of
+ SOME (info as (_, result)) =>
+ let
+ fun is_intro_of intro =
+ let
+ val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro))
+ in (fst (dest_Const const) = name) end;
+ val thy = ProofContext.theory_of ctxt
+ val intros = map (preprocess_intro thy) (filter is_intro_of (#intrs result))
+ val index = find_index (fn s => s = name) (#names (fst info))
+ val pre_elim = nth (#elims result) index
+ val pred = nth (#preds result) index
+ val elim_t = mk_casesrule ctxt pred intros
+ val nparams = length (Inductive.params_of (#raw_induct result))
+ val elim = prove_casesrule ctxt (pred, (pre_elim, nparams)) elim_t
+ in
+ mk_pred_data (((map (pair NONE) intros, SOME elim), true), no_compilation)
+ end
+ | NONE => error ("No such predicate: " ^ quote name)
+
+fun add_predfun_data name mode data =
+ let
+ val add = (apsnd o apsnd o apfst) (cons (mode, mk_predfun_data data))
+ in PredData.map (Graph.map_node name (map_pred_data add)) end
+
+fun is_inductive_predicate ctxt name =
+ is_some (try (Inductive.the_inductive ctxt) name)
+
+fun depending_preds_of ctxt (key, value) =
+ let
+ val intros = map (Thm.prop_of o snd) ((#intros o rep_pred_data) value)
+ in
+ fold Term.add_const_names intros []
+ |> (fn cs =>
+ if member (op =) cs @{const_name "HOL.eq"} then
+ insert (op =) @{const_name Predicate.eq} cs
+ else cs)
+ |> filter (fn c => (not (c = key)) andalso
+ (is_inductive_predicate ctxt c orelse is_registered ctxt c))
+ end;
+
+fun add_intro (opt_case_name, thm) thy =
+ let
+ val (name, T) = dest_Const (fst (strip_intro_concl thm))
+ fun cons_intro gr =
+ case try (Graph.get_node gr) name of
+ SOME pred_data => Graph.map_node name (map_pred_data
+ (apfst (apfst (apfst (fn intros => intros @ [(opt_case_name, thm)]))))) gr
+ | NONE => Graph.new_node (name, mk_pred_data ((([(opt_case_name, thm)], NONE), false), no_compilation)) gr
+ in PredData.map cons_intro thy end
+
+fun set_elim thm =
+ let
+ val (name, _) = dest_Const (fst
+ (strip_comb (HOLogic.dest_Trueprop (hd (prems_of thm)))))
+ in PredData.map (Graph.map_node name (map_pred_data (apfst (apfst (apsnd (K (SOME thm))))))) end
+
+fun register_predicate (constname, intros, elim) thy =
+ let
+ val named_intros = map (pair NONE) intros
+ in
+ if not (member (op =) (Graph.keys (PredData.get thy)) constname) then
+ PredData.map
+ (Graph.new_node (constname,
+ mk_pred_data (((named_intros, SOME elim), false), no_compilation))) thy
+ else thy
+ end
+
+fun register_intros (constname, pre_intros) thy =
+ let
+ val T = Sign.the_const_type thy constname
+ fun constname_of_intro intr = fst (dest_Const (fst (strip_intro_concl intr)))
+ val _ = if not (forall (fn intr => constname_of_intro intr = constname) pre_intros) then
+ error ("register_intros: Introduction rules of different constants are used\n" ^
+ "expected rules for " ^ constname ^ ", but received rules for " ^
+ commas (map constname_of_intro pre_intros))
+ else ()
+ val pred = Const (constname, T)
+ val pre_elim =
+ (Drule.export_without_context o Skip_Proof.make_thm thy)
+ (mk_casesrule (ProofContext.init_global thy) pred pre_intros)
+ in register_predicate (constname, pre_intros, pre_elim) thy end
+
+fun defined_function_of compilation pred =
+ let
+ val set = (apsnd o apfst) (cons (compilation, []))
+ in
+ PredData.map (Graph.map_node pred (map_pred_data set))
+ end
+
+fun set_function_name compilation pred mode name =
+ let
+ val set = (apsnd o apfst)
+ (AList.map_default (op =) (compilation, [(mode, name)]) (cons (mode, name)))
+ in
+ PredData.map (Graph.map_node pred (map_pred_data set))
+ end
+
+fun set_needs_random name modes =
+ let
+ val set = (apsnd o apsnd o apsnd) (K modes)
+ in
+ PredData.map (Graph.map_node name (map_pred_data set))
+ end
+
+fun extend' value_of edges_of key (G, visited) =
+ let
+ val (G', v) = case try (Graph.get_node G) key of
+ SOME v => (G, v)
+ | NONE => (Graph.new_node (key, value_of key) G, value_of key)
+ val (G'', visited') = fold (extend' value_of edges_of)
+ (subtract (op =) visited (edges_of (key, v)))
+ (G', key :: visited)
+ in
+ (fold (Graph.add_edge o (pair key)) (edges_of (key, v)) G'', visited')
+ end;
+
+fun extend value_of edges_of key G = fst (extend' value_of edges_of key (G, []))
+
+fun extend_intro_graph names thy =
+ let
+ val ctxt = ProofContext.init_global thy
+ in
+ PredData.map (fold (extend (fetch_pred_data ctxt) (depending_preds_of ctxt)) names) thy
+ end
+
+fun preprocess_intros name thy =
+ PredData.map (Graph.map_node name (map_pred_data (apfst (fn (rules, preprocessed) =>
+ if preprocessed then (rules, preprocessed)
+ else
+ let
+ val (named_intros, SOME elim) = rules
+ val named_intros' = map (apsnd (preprocess_intro thy)) named_intros
+ val pred = Const (name, Sign.the_const_type thy name)
+ val ctxt = ProofContext.init_global thy
+ val elim_t = mk_casesrule ctxt pred (map snd named_intros')
+ val nparams = (case try (Inductive.the_inductive ctxt) name of
+ SOME (_, result) => length (Inductive.params_of (#raw_induct result))
+ | NONE => 0)
+ val elim' = prove_casesrule ctxt (pred, (elim, 0)) elim_t
+ in
+ ((named_intros', SOME elim'), true)
+ end))))
+ thy
+
+(* registration of alternative function names *)
+
+structure Alt_Compilations_Data = Theory_Data
+(
+ type T = (mode * (compilation_funs -> typ -> term)) list Symtab.table;
+ val empty = Symtab.empty;
+ val extend = I;
+ fun merge data : T = Symtab.merge (K true) data;
+);
+
+fun alternative_compilation_of_global thy pred_name mode =
+ AList.lookup eq_mode (Symtab.lookup_list (Alt_Compilations_Data.get thy) pred_name) mode
+
+fun alternative_compilation_of ctxt pred_name mode =
+ AList.lookup eq_mode
+ (Symtab.lookup_list (Alt_Compilations_Data.get (ProofContext.theory_of ctxt)) pred_name) mode
+
+fun force_modes_and_compilations pred_name compilations =
+ let
+ (* thm refl is a dummy thm *)
+ val modes = map fst compilations
+ val (needs_random, non_random_modes) = pairself (map fst)
+ (List.partition (fn (m, (fun_name, random)) => random) compilations)
+ val non_random_dummys = map (rpair "dummy") non_random_modes
+ val all_dummys = map (rpair "dummy") modes
+ val dummy_function_names = map (rpair all_dummys) Predicate_Compile_Aux.random_compilations
+ @ map (rpair non_random_dummys) Predicate_Compile_Aux.non_random_compilations
+ val alt_compilations = map (apsnd fst) compilations
+ in
+ PredData.map (Graph.new_node
+ (pred_name, mk_pred_data ((([], SOME @{thm refl}), true), (dummy_function_names, ([], needs_random)))))
+ #> Alt_Compilations_Data.map (Symtab.insert (K false) (pred_name, alt_compilations))
+ end
+
+fun functional_compilation fun_name mode compfuns T =
+ let
+ val (inpTs, outpTs) = split_map_modeT (fn _ => fn T => (SOME T, NONE))
+ mode (binder_types T)
+ val bs = map (pair "x") inpTs
+ val bounds = map Bound (rev (0 upto (length bs) - 1))
+ val f = Const (fun_name, inpTs ---> HOLogic.mk_tupleT outpTs)
+ in list_abs (bs, mk_single compfuns (list_comb (f, bounds))) end
+
+fun register_alternative_function pred_name mode fun_name =
+ Alt_Compilations_Data.map (Symtab.insert_list (eq_pair eq_mode (K false))
+ (pred_name, (mode, functional_compilation fun_name mode)))
+
+fun force_modes_and_functions pred_name fun_names =
+ force_modes_and_compilations pred_name
+ (map (fn (mode, (fun_name, random)) => (mode, (functional_compilation fun_name mode, random)))
+ fun_names)
+
+end;
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/mode_inference.ML Sun Oct 24 15:11:24 2010 -0700
@@ -0,0 +1,643 @@
+(* Title: HOL/Tools/Predicate_Compile/mode_inference.ML
+ Author: Lukas Bulwahn, TU Muenchen
+
+Mode inference for the predicate compiler
+
+*)
+
+signature MODE_INFERENCE =
+sig
+ type mode = Predicate_Compile_Aux.mode
+
+ (* options *)
+ type mode_analysis_options =
+ {use_generators : bool,
+ reorder_premises : bool,
+ infer_pos_and_neg_modes : bool}
+
+ (* mode operation *)
+ val all_input_of : typ -> mode
+ (* mode derivation and operations *)
+ datatype mode_derivation = Mode_App of mode_derivation * mode_derivation | Context of mode
+ | Mode_Pair of mode_derivation * mode_derivation | Term of mode
+
+ val head_mode_of : mode_derivation -> mode
+ val param_derivations_of : mode_derivation -> mode_derivation list
+ val collect_context_modes : mode_derivation -> mode list
+
+ type moded_clause = term list * (Predicate_Compile_Aux.indprem * mode_derivation) list
+ type 'a pred_mode_table = (string * ((bool * mode) * 'a) list) list
+
+ (* mode inference operations *)
+ val all_derivations_of :
+ Proof.context -> (string * mode list) list -> string list -> term
+ -> (mode_derivation * string list) list
+ (* TODO: move all_modes creation to infer_modes *)
+ val infer_modes :
+ mode_analysis_options -> Predicate_Compile_Aux.options ->
+ (string -> mode list) * (string -> mode list)
+ * (string -> mode -> bool) -> Proof.context -> (string * typ) list ->
+ (string * mode list) list ->
+ string list -> (string * (Term.term list * Predicate_Compile_Aux.indprem list) list) list ->
+ ((moded_clause list pred_mode_table * (string * mode list) list) * string list)
+
+ (* mode and term operations -- to be moved to Predicate_Compile_Aux *)
+ val collect_non_invertible_subterms :
+ Proof.context -> term -> string list * term list -> (term * (string list * term list))
+ val is_all_input : mode -> bool
+ val term_vs : term -> string list
+ val terms_vs : term list -> string list
+
+end;
+
+structure Mode_Inference : MODE_INFERENCE =
+struct
+
+open Predicate_Compile_Aux;
+open Core_Data;
+
+(* derivation trees for modes of premises *)
+
+datatype mode_derivation = Mode_App of mode_derivation * mode_derivation | Context of mode
+ | Mode_Pair of mode_derivation * mode_derivation | Term of mode
+
+fun string_of_derivation (Mode_App (m1, m2)) =
+ "App (" ^ string_of_derivation m1 ^ ", " ^ string_of_derivation m2 ^ ")"
+ | string_of_derivation (Mode_Pair (m1, m2)) =
+ "Pair (" ^ string_of_derivation m1 ^ ", " ^ string_of_derivation m2 ^ ")"
+ | string_of_derivation (Term m) = "Term (" ^ string_of_mode m ^ ")"
+ | string_of_derivation (Context m) = "Context (" ^ string_of_mode m ^ ")"
+
+fun strip_mode_derivation deriv =
+ let
+ fun strip (Mode_App (deriv1, deriv2)) ds = strip deriv1 (deriv2 :: ds)
+ | strip deriv ds = (deriv, ds)
+ in
+ strip deriv []
+ end
+
+fun mode_of (Context m) = m
+ | mode_of (Term m) = m
+ | mode_of (Mode_App (d1, d2)) =
+ (case mode_of d1 of Fun (m, m') =>
+ (if eq_mode (m, mode_of d2) then m' else raise Fail "mode_of: derivation has mismatching modes")
+ | _ => raise Fail "mode_of: derivation has a non-functional mode")
+ | mode_of (Mode_Pair (d1, d2)) =
+ Pair (mode_of d1, mode_of d2)
+
+fun head_mode_of deriv = mode_of (fst (strip_mode_derivation deriv))
+
+fun param_derivations_of deriv =
+ let
+ val (_, argument_derivs) = strip_mode_derivation deriv
+ fun param_derivation (Mode_Pair (m1, m2)) =
+ param_derivation m1 @ param_derivation m2
+ | param_derivation (Term _) = []
+ | param_derivation m = [m]
+ in
+ maps param_derivation argument_derivs
+ end
+
+fun collect_context_modes (Mode_App (m1, m2)) =
+ collect_context_modes m1 @ collect_context_modes m2
+ | collect_context_modes (Mode_Pair (m1, m2)) =
+ collect_context_modes m1 @ collect_context_modes m2
+ | collect_context_modes (Context m) = [m]
+ | collect_context_modes (Term _) = []
+
+type moded_clause = term list * (Predicate_Compile_Aux.indprem * mode_derivation) list
+type 'a pred_mode_table = (string * ((bool * mode) * 'a) list) list
+
+
+(** string_of functions **)
+
+fun string_of_prem ctxt (Prem t) =
+ (Syntax.string_of_term ctxt t) ^ "(premise)"
+ | string_of_prem ctxt (Negprem t) =
+ (Syntax.string_of_term ctxt (HOLogic.mk_not t)) ^ "(negative premise)"
+ | string_of_prem ctxt (Sidecond t) =
+ (Syntax.string_of_term ctxt t) ^ "(sidecondition)"
+ | string_of_prem ctxt _ = raise Fail "string_of_prem: unexpected input"
+
+fun string_of_clause ctxt pred (ts, prems) =
+ (space_implode " --> "
+ (map (string_of_prem ctxt) prems)) ^ " --> " ^ pred ^ " "
+ ^ (space_implode " " (map (Syntax.string_of_term ctxt) ts))
+
+type mode_analysis_options =
+ {use_generators : bool,
+ reorder_premises : bool,
+ infer_pos_and_neg_modes : bool}
+
+fun is_constrt thy =
+ let
+ val cnstrs = flat (maps
+ (map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd)
+ (Symtab.dest (Datatype.get_all thy)));
+ fun check t = (case strip_comb t of
+ (Free _, []) => true
+ | (Const (s, T), ts) => (case (AList.lookup (op =) cnstrs s, body_type T) of
+ (SOME (i, Tname), Type (Tname', _)) =>
+ length ts = i andalso Tname = Tname' andalso forall check ts
+ | _ => false)
+ | _ => false)
+ in check end;
+
+(*** check if a type is an equality type (i.e. doesn't contain fun)
+ FIXME this is only an approximation ***)
+fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts
+ | is_eqT _ = true;
+
+fun term_vs tm = fold_aterms (fn Free (x, T) => cons x | _ => I) tm [];
+val terms_vs = distinct (op =) o maps term_vs;
+
+(** collect all Frees in a term (with duplicates!) **)
+fun term_vTs tm =
+ fold_aterms (fn Free xT => cons xT | _ => I) tm [];
+
+fun subsets i j =
+ if i <= j then
+ let
+ fun merge xs [] = xs
+ | merge [] ys = ys
+ | merge (x::xs) (y::ys) = if length x >= length y then x::merge xs (y::ys)
+ else y::merge (x::xs) ys;
+ val is = subsets (i+1) j
+ in merge (map (fn ks => i::ks) is) is end
+ else [[]];
+
+fun print_failed_mode options thy modes p (pol, m) rs is =
+ if show_mode_inference options then
+ let
+ val _ = tracing ("Clauses " ^ commas (map (fn i => string_of_int (i + 1)) is) ^ " of " ^
+ p ^ " violates mode " ^ string_of_mode m)
+ in () end
+ else ()
+
+fun error_of p (pol, m) is =
+ " Clauses " ^ commas (map (fn i => string_of_int (i + 1)) is) ^ " of " ^
+ p ^ " violates mode " ^ string_of_mode m
+
+fun is_all_input mode =
+ let
+ fun is_all_input' (Fun _) = true
+ | is_all_input' (Pair (m1, m2)) = is_all_input' m1 andalso is_all_input' m2
+ | is_all_input' Input = true
+ | is_all_input' Output = false
+ in
+ forall is_all_input' (strip_fun_mode mode)
+ end
+
+fun all_input_of T =
+ let
+ val (Ts, U) = strip_type T
+ fun input_of (Type (@{type_name Product_Type.prod}, [T1, T2])) = Pair (input_of T1, input_of T2)
+ | input_of _ = Input
+ in
+ if U = HOLogic.boolT then
+ fold_rev (curry Fun) (map input_of Ts) Bool
+ else
+ raise Fail "all_input_of: not a predicate"
+ end
+
+fun find_least ord xs =
+ let
+ fun find' x y = (case y of NONE => SOME x | SOME y' => if ord (x, y') = LESS then SOME x else y)
+ in
+ fold find' xs NONE
+ end
+
+fun term_vs tm = fold_aterms (fn Free (x, T) => cons x | _ => I) tm [];
+val terms_vs = distinct (op =) o maps term_vs;
+
+fun input_mode T =
+ let
+ val (Ts, U) = strip_type T
+ in
+ fold_rev (curry Fun) (map (K Input) Ts) Input
+ end
+
+fun output_mode T =
+ let
+ val (Ts, U) = strip_type T
+ in
+ fold_rev (curry Fun) (map (K Output) Ts) Output
+ end
+
+fun is_invertible_function ctxt (Const (f, _)) = is_constr ctxt f
+ | is_invertible_function ctxt _ = false
+
+fun non_invertible_subterms ctxt (t as Free _) = []
+ | non_invertible_subterms ctxt t =
+ let
+ val (f, args) = strip_comb t
+ in
+ if is_invertible_function ctxt f then
+ maps (non_invertible_subterms ctxt) args
+ else
+ [t]
+ end
+
+fun collect_non_invertible_subterms ctxt (f as Free _) (names, eqs) = (f, (names, eqs))
+ | collect_non_invertible_subterms ctxt t (names, eqs) =
+ case (strip_comb t) of (f, args) =>
+ if is_invertible_function ctxt f then
+ let
+ val (args', (names', eqs')) =
+ fold_map (collect_non_invertible_subterms ctxt) args (names, eqs)
+ in
+ (list_comb (f, args'), (names', eqs'))
+ end
+ else
+ let
+ val s = Name.variant names "x"
+ val v = Free (s, fastype_of t)
+ in
+ (v, (s :: names, HOLogic.mk_eq (v, t) :: eqs))
+ end
+(*
+ if is_constrt thy t then (t, (names, eqs)) else
+ let
+ val s = Name.variant names "x"
+ val v = Free (s, fastype_of t)
+ in (v, (s::names, HOLogic.mk_eq (v, t)::eqs)) end;
+*)
+
+fun is_possible_output ctxt vs t =
+ forall
+ (fn t => is_eqT (fastype_of t) andalso forall (member (op =) vs) (term_vs t))
+ (non_invertible_subterms ctxt t)
+ andalso
+ (forall (is_eqT o snd)
+ (inter (fn ((f', _), f) => f = f') vs (Term.add_frees t [])))
+
+fun vars_of_destructable_term ctxt (Free (x, _)) = [x]
+ | vars_of_destructable_term ctxt t =
+ let
+ val (f, args) = strip_comb t
+ in
+ if is_invertible_function ctxt f then
+ maps (vars_of_destructable_term ctxt) args
+ else
+ []
+ end
+
+fun is_constructable vs t = forall (member (op =) vs) (term_vs t)
+
+fun missing_vars vs t = subtract (op =) vs (term_vs t)
+
+fun output_terms (Const (@{const_name Pair}, _) $ t1 $ t2, Mode_Pair (d1, d2)) =
+ output_terms (t1, d1) @ output_terms (t2, d2)
+ | output_terms (t1 $ t2, Mode_App (d1, d2)) =
+ output_terms (t1, d1) @ output_terms (t2, d2)
+ | output_terms (t, Term Output) = [t]
+ | output_terms _ = []
+
+fun lookup_mode modes (Const (s, T)) =
+ (case (AList.lookup (op =) modes s) of
+ SOME ms => SOME (map (fn m => (Context m, [])) ms)
+ | NONE => NONE)
+ | lookup_mode modes (Free (x, _)) =
+ (case (AList.lookup (op =) modes x) of
+ SOME ms => SOME (map (fn m => (Context m , [])) ms)
+ | NONE => NONE)
+
+fun derivations_of (ctxt : Proof.context) modes vs (Const (@{const_name Pair}, _) $ t1 $ t2) (Pair (m1, m2)) =
+ map_product
+ (fn (m1, mvars1) => fn (m2, mvars2) => (Mode_Pair (m1, m2), union (op =) mvars1 mvars2))
+ (derivations_of ctxt modes vs t1 m1) (derivations_of ctxt modes vs t2 m2)
+ | derivations_of ctxt modes vs t (m as Fun _) =
+ (*let
+ val (p, args) = strip_comb t
+ in
+ (case lookup_mode modes p of
+ SOME ms => map_filter (fn (Context m, []) => let
+ val ms = strip_fun_mode m
+ val (argms, restms) = chop (length args) ms
+ val m' = fold_rev (curry Fun) restms Bool
+ in
+ if forall (fn m => eq_mode (Input, m)) argms andalso eq_mode (m', mode) then
+ SOME (fold (curry Mode_App) (map Term argms) (Context m), missing_vars vs t)
+ else NONE
+ end) ms
+ | NONE => (if is_all_input mode then [(Context mode, [])] else []))
+ end*)
+ (case try (all_derivations_of ctxt modes vs) t of
+ SOME derivs =>
+ filter (fn (d, mvars) => eq_mode (mode_of d, m) andalso null (output_terms (t, d))) derivs
+ | NONE => (if is_all_input m then [(Context m, [])] else []))
+ | derivations_of ctxt modes vs t m =
+ if eq_mode (m, Input) then
+ [(Term Input, missing_vars vs t)]
+ else if eq_mode (m, Output) then
+ (if is_possible_output ctxt vs t then [(Term Output, [])] else [])
+ else []
+and all_derivations_of ctxt modes vs (Const (@{const_name Pair}, _) $ t1 $ t2) =
+ let
+ val derivs1 = all_derivations_of ctxt modes vs t1
+ val derivs2 = all_derivations_of ctxt modes vs t2
+ in
+ map_product
+ (fn (m1, mvars1) => fn (m2, mvars2) => (Mode_Pair (m1, m2), union (op =) mvars1 mvars2))
+ derivs1 derivs2
+ end
+ | all_derivations_of ctxt modes vs (t1 $ t2) =
+ let
+ val derivs1 = all_derivations_of ctxt modes vs t1
+ in
+ maps (fn (d1, mvars1) =>
+ case mode_of d1 of
+ Fun (m', _) => map (fn (d2, mvars2) =>
+ (Mode_App (d1, d2), union (op =) mvars1 mvars2)) (derivations_of ctxt modes vs t2 m')
+ | _ => raise Fail "all_derivations_of: derivation has an unexpected non-functional mode") derivs1
+ end
+ | all_derivations_of _ modes vs (Const (s, T)) = the (lookup_mode modes (Const (s, T)))
+ | all_derivations_of _ modes vs (Free (x, T)) = the (lookup_mode modes (Free (x, T)))
+ | all_derivations_of _ modes vs _ = raise Fail "all_derivations_of: unexpected term"
+
+fun rev_option_ord ord (NONE, NONE) = EQUAL
+ | rev_option_ord ord (NONE, SOME _) = GREATER
+ | rev_option_ord ord (SOME _, NONE) = LESS
+ | rev_option_ord ord (SOME x, SOME y) = ord (x, y)
+
+fun random_mode_in_deriv modes t deriv =
+ case try dest_Const (fst (strip_comb t)) of
+ SOME (s, _) =>
+ (case AList.lookup (op =) modes s of
+ SOME ms =>
+ (case AList.lookup (op =) (map (fn ((p, m), r) => (m, r)) ms) (head_mode_of deriv) of
+ SOME r => r
+ | NONE => false)
+ | NONE => false)
+ | NONE => false
+
+fun number_of_output_positions mode =
+ let
+ val args = strip_fun_mode mode
+ fun contains_output (Fun _) = false
+ | contains_output Input = false
+ | contains_output Output = true
+ | contains_output (Pair (m1, m2)) = contains_output m1 orelse contains_output m2
+ in
+ length (filter contains_output args)
+ end
+
+fun lex_ord ord1 ord2 (x, x') =
+ case ord1 (x, x') of
+ EQUAL => ord2 (x, x')
+ | ord => ord
+
+fun lexl_ord [] (x, x') = EQUAL
+ | lexl_ord (ord :: ords') (x, x') =
+ case ord (x, x') of
+ EQUAL => lexl_ord ords' (x, x')
+ | ord => ord
+
+fun deriv_ord' ctxt pol pred modes t1 t2 ((deriv1, mvars1), (deriv2, mvars2)) =
+ let
+ (* prefer functional modes if it is a function *)
+ fun fun_ord ((t1, deriv1, mvars1), (t2, deriv2, mvars2)) =
+ let
+ fun is_functional t mode =
+ case try (fst o dest_Const o fst o strip_comb) t of
+ NONE => false
+ | SOME c => is_some (alternative_compilation_of ctxt c mode)
+ in
+ case (is_functional t1 (head_mode_of deriv1), is_functional t2 (head_mode_of deriv2)) of
+ (true, true) => EQUAL
+ | (true, false) => LESS
+ | (false, true) => GREATER
+ | (false, false) => EQUAL
+ end
+ (* prefer modes without requirement for generating random values *)
+ fun mvars_ord ((t1, deriv1, mvars1), (t2, deriv2, mvars2)) =
+ int_ord (length mvars1, length mvars2)
+ (* prefer non-random modes *)
+ fun random_mode_ord ((t1, deriv1, mvars1), (t2, deriv2, mvars2)) =
+ int_ord (if random_mode_in_deriv modes t1 deriv1 then 1 else 0,
+ if random_mode_in_deriv modes t2 deriv2 then 1 else 0)
+ (* prefer modes with more input and less output *)
+ fun output_mode_ord ((t1, deriv1, mvars1), (t2, deriv2, mvars2)) =
+ int_ord (number_of_output_positions (head_mode_of deriv1),
+ number_of_output_positions (head_mode_of deriv2))
+ (* prefer recursive calls *)
+ fun is_rec_premise t =
+ case fst (strip_comb t) of Const (c, T) => c = pred | _ => false
+ fun recursive_ord ((t1, deriv1, mvars1), (t2, deriv2, mvars2)) =
+ int_ord (if is_rec_premise t1 then 0 else 1,
+ if is_rec_premise t2 then 0 else 1)
+ val ord = lexl_ord [mvars_ord, fun_ord, random_mode_ord, output_mode_ord, recursive_ord]
+ in
+ ord ((t1, deriv1, mvars1), (t2, deriv2, mvars2))
+ end
+
+fun deriv_ord ctxt pol pred modes t = deriv_ord' ctxt pol pred modes t t
+
+fun premise_ord thy pol pred modes ((prem1, a1), (prem2, a2)) =
+ rev_option_ord (deriv_ord' thy pol pred modes (dest_indprem prem1) (dest_indprem prem2)) (a1, a2)
+
+fun print_mode_list modes =
+ tracing ("modes: " ^ (commas (map (fn (s, ms) => s ^ ": " ^
+ commas (map (fn (m, r) => string_of_mode m ^ (if r then " random " else " not ")) ms)) modes)))
+
+fun select_mode_prem (mode_analysis_options : mode_analysis_options) (ctxt : Proof.context) pred
+ pol (modes, (pos_modes, neg_modes)) vs ps =
+ let
+ fun choose_mode_of_prem (Prem t) =
+ find_least (deriv_ord ctxt pol pred modes t) (all_derivations_of ctxt pos_modes vs t)
+ | choose_mode_of_prem (Sidecond t) = SOME (Context Bool, missing_vars vs t)
+ | choose_mode_of_prem (Negprem t) = find_least (deriv_ord ctxt (not pol) pred modes t)
+ (filter (fn (d, missing_vars) => is_all_input (head_mode_of d))
+ (all_derivations_of ctxt neg_modes vs t))
+ | choose_mode_of_prem p = raise Fail ("choose_mode_of_prem: unexpected premise " ^ string_of_prem ctxt p)
+ in
+ if #reorder_premises mode_analysis_options then
+ find_least (premise_ord ctxt pol pred modes) (ps ~~ map choose_mode_of_prem ps)
+ else
+ SOME (hd ps, choose_mode_of_prem (hd ps))
+ end
+
+fun check_mode_clause' (mode_analysis_options : mode_analysis_options) ctxt pred param_vs (modes :
+ (string * ((bool * mode) * bool) list) list) ((pol, mode) : bool * mode) (ts, ps) =
+ let
+ val vTs = distinct (op =) (fold Term.add_frees (map dest_indprem ps) (fold Term.add_frees ts []))
+ val modes' = modes @ (param_vs ~~ map (fn x => [((true, x), false), ((false, x), false)]) (ho_arg_modes_of mode))
+ fun retrieve_modes_of_pol pol = map (fn (s, ms) =>
+ (s, map_filter (fn ((p, m), r) => if p = pol then SOME m else NONE | _ => NONE) ms))
+ val (pos_modes', neg_modes') =
+ if #infer_pos_and_neg_modes mode_analysis_options then
+ (retrieve_modes_of_pol pol modes', retrieve_modes_of_pol (not pol) modes')
+ else
+ let
+ val modes = map (fn (s, ms) => (s, map (fn ((p, m), r) => m) ms)) modes'
+ in (modes, modes) end
+ val (in_ts, out_ts) = split_mode mode ts
+ val in_vs = union (op =) param_vs (maps (vars_of_destructable_term ctxt) in_ts)
+ val out_vs = terms_vs out_ts
+ fun known_vs_after p vs = (case p of
+ Prem t => union (op =) vs (term_vs t)
+ | Sidecond t => union (op =) vs (term_vs t)
+ | Negprem t => union (op =) vs (term_vs t)
+ | _ => raise Fail "unexpected premise")
+ fun check_mode_prems acc_ps rnd vs [] = SOME (acc_ps, vs, rnd)
+ | check_mode_prems acc_ps rnd vs ps =
+ (case
+ (select_mode_prem mode_analysis_options ctxt pred pol (modes', (pos_modes', neg_modes')) vs ps) of
+ SOME (p, SOME (deriv, [])) => check_mode_prems ((p, deriv) :: acc_ps) rnd
+ (known_vs_after p vs) (filter_out (equal p) ps)
+ | SOME (p, SOME (deriv, missing_vars)) =>
+ if #use_generators mode_analysis_options andalso pol then
+ check_mode_prems ((p, deriv) :: (map
+ (fn v => (Generator (v, the (AList.lookup (op =) vTs v)), Term Output))
+ (distinct (op =) missing_vars))
+ @ acc_ps) true (known_vs_after p vs) (filter_out (equal p) ps)
+ else NONE
+ | SOME (p, NONE) => NONE
+ | NONE => NONE)
+ in
+ case check_mode_prems [] false in_vs ps of
+ NONE => NONE
+ | SOME (acc_ps, vs, rnd) =>
+ if forall (is_constructable vs) (in_ts @ out_ts) then
+ SOME (ts, rev acc_ps, rnd)
+ else
+ if #use_generators mode_analysis_options andalso pol then
+ let
+ val generators = map
+ (fn v => (Generator (v, the (AList.lookup (op =) vTs v)), Term Output))
+ (subtract (op =) vs (terms_vs (in_ts @ out_ts)))
+ in
+ SOME (ts, rev (generators @ acc_ps), true)
+ end
+ else
+ NONE
+ end
+
+datatype result = Success of bool | Error of string
+
+fun check_modes_pred' mode_analysis_options options thy param_vs clauses modes (p, (ms : ((bool * mode) * bool) list)) =
+ let
+ fun split xs =
+ let
+ fun split' [] (ys, zs) = (rev ys, rev zs)
+ | split' ((m, Error z) :: xs) (ys, zs) = split' xs (ys, z :: zs)
+ | split' (((m : bool * mode), Success rnd) :: xs) (ys, zs) = split' xs ((m, rnd) :: ys, zs)
+ in
+ split' xs ([], [])
+ end
+ val rs = these (AList.lookup (op =) clauses p)
+ fun check_mode m =
+ let
+ val res = Output.cond_timeit false "work part of check_mode for one mode" (fn _ =>
+ map (check_mode_clause' mode_analysis_options thy p param_vs modes m) rs)
+ in
+ Output.cond_timeit false "aux part of check_mode for one mode" (fn _ =>
+ case find_indices is_none res of
+ [] => Success (exists (fn SOME (_, _, true) => true | _ => false) res)
+ | is => (print_failed_mode options thy modes p m rs is; Error (error_of p m is)))
+ end
+ val _ = if show_mode_inference options then
+ tracing ("checking " ^ string_of_int (length ms) ^ " modes ...")
+ else ()
+ val res = Output.cond_timeit false "check_mode" (fn _ => map (fn (m, _) => (m, check_mode m)) ms)
+ val (ms', errors) = split res
+ in
+ ((p, (ms' : ((bool * mode) * bool) list)), errors)
+ end;
+
+fun get_modes_pred' mode_analysis_options thy param_vs clauses modes (p, ms) =
+ let
+ val rs = these (AList.lookup (op =) clauses p)
+ in
+ (p, map (fn (m, rnd) =>
+ (m, map
+ ((fn (ts, ps, rnd) => (ts, ps)) o the o
+ check_mode_clause' mode_analysis_options thy p param_vs modes m) rs)) ms)
+ end;
+
+fun fixp f (x : (string * ((bool * mode) * bool) list) list) =
+ let val y = f x
+ in if x = y then x else fixp f y end;
+
+fun fixp_with_state f (x : (string * ((bool * mode) * bool) list) list, state) =
+ let
+ val (y, state') = f (x, state)
+ in
+ if x = y then (y, state') else fixp_with_state f (y, state')
+ end
+
+fun string_of_ext_mode ((pol, mode), rnd) =
+ string_of_mode mode ^ "(" ^ (if pol then "pos" else "neg") ^ ", "
+ ^ (if rnd then "rnd" else "nornd") ^ ")"
+
+fun print_extra_modes options modes =
+ if show_mode_inference options then
+ tracing ("Modes of inferred predicates: " ^
+ cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map string_of_ext_mode ms)) modes))
+ else ()
+
+fun infer_modes mode_analysis_options options (lookup_mode, lookup_neg_mode, needs_random) ctxt
+ preds all_modes param_vs clauses =
+ let
+ fun appair f (x1, x2) (y1, y2) = (f x1 y1, f x2 y2)
+ fun add_needs_random s (false, m) = ((false, m), false)
+ | add_needs_random s (true, m) = ((true, m), needs_random s m)
+ fun add_polarity_and_random_bit s b ms = map (fn m => add_needs_random s (b, m)) ms
+ val prednames = map fst preds
+ (* extramodes contains all modes of all constants, should we only use the necessary ones
+ - what is the impact on performance? *)
+ fun predname_of (Prem t) =
+ (case try dest_Const (fst (strip_comb t)) of SOME (c, _) => insert (op =) c | NONE => I)
+ | predname_of (Negprem t) =
+ (case try dest_Const (fst (strip_comb t)) of SOME (c, _) => insert (op =) c | NONE => I)
+ | predname_of _ = I
+ val relevant_prednames = fold (fn (_, clauses') =>
+ fold (fn (_, ps) => fold Term.add_const_names (map dest_indprem ps)) clauses') clauses []
+ |> filter_out (fn name => member (op =) prednames name)
+ val extra_modes =
+ if #infer_pos_and_neg_modes mode_analysis_options then
+ let
+ val pos_extra_modes =
+ map_filter (fn name => Option.map (pair name) (try lookup_mode name))
+ relevant_prednames
+ val neg_extra_modes =
+ map_filter (fn name => Option.map (pair name) (try lookup_neg_mode name))
+ relevant_prednames
+ in
+ map (fn (s, ms) => (s, (add_polarity_and_random_bit s true ms)
+ @ add_polarity_and_random_bit s false (the (AList.lookup (op =) neg_extra_modes s))))
+ pos_extra_modes
+ end
+ else
+ map (fn (s, ms) => (s, (add_polarity_and_random_bit s true ms)))
+ (map_filter (fn name => Option.map (pair name) (try lookup_mode name))
+ relevant_prednames)
+ val _ = print_extra_modes options extra_modes
+ val start_modes =
+ if #infer_pos_and_neg_modes mode_analysis_options then
+ map (fn (s, ms) => (s, map (fn m => ((true, m), false)) ms @
+ (map (fn m => ((false, m), false)) ms))) all_modes
+ else
+ map (fn (s, ms) => (s, map (fn m => ((true, m), false)) ms)) all_modes
+ fun iteration modes = map
+ (check_modes_pred' mode_analysis_options options ctxt param_vs clauses
+ (modes @ extra_modes)) modes
+ val ((modes : (string * ((bool * mode) * bool) list) list), errors) =
+ Output.cond_timeit false "Fixpount computation of mode analysis" (fn () =>
+ if show_invalid_clauses options then
+ fixp_with_state (fn (modes, errors) =>
+ let
+ val (modes', new_errors) = split_list (iteration modes)
+ in (modes', errors @ flat new_errors) end) (start_modes, [])
+ else
+ (fixp (fn modes => map fst (iteration modes)) start_modes, []))
+ val moded_clauses = map (get_modes_pred' mode_analysis_options ctxt param_vs clauses
+ (modes @ extra_modes)) modes
+ val need_random = fold (fn (s, ms) => if member (op =) (map fst preds) s then
+ cons (s, (map_filter (fn ((true, m), true) => SOME m | _ => NONE) ms)) else I)
+ modes []
+ in
+ ((moded_clauses, need_random), errors)
+ end;
+
+end;
\ No newline at end of file
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile.ML Sun Oct 24 15:11:24 2010 -0700
@@ -75,7 +75,7 @@
fun preprocess_strong_conn_constnames options gr ts thy =
if forall (fn (Const (c, _)) =>
- Predicate_Compile_Core.is_registered (ProofContext.init_global thy) c) ts then
+ Core_Data.is_registered (ProofContext.init_global thy) c) ts then
thy
else
let
@@ -121,7 +121,7 @@
val intross10 = map_specs (map_filter (peephole_optimisation thy3)) intross9
val _ = print_intross options thy3 "introduction rules before registering: " intross10
val _ = print_step options "Registering introduction rules..."
- val thy4 = fold Predicate_Compile_Core.register_intros intross10 thy3
+ val thy4 = fold Core_Data.register_intros intross10 thy3
in
thy4
end;
@@ -173,6 +173,7 @@
no_higher_order_predicate = [],
inductify = chk "inductify",
detect_switches = chk "detect_switches",
+ smart_depth_limiting = chk "smart_depth_limiting",
compilation = compilation
}
end
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Sun Oct 24 15:11:24 2010 -0700
@@ -90,8 +90,10 @@
(* Different compilations *)
datatype compilation = Pred | Depth_Limited | Random | Depth_Limited_Random | DSeq | Annotated
| Pos_Random_DSeq | Neg_Random_DSeq | New_Pos_Random_DSeq | New_Neg_Random_DSeq
+ | Pos_Generator_DSeq | Neg_Generator_DSeq
val negative_compilation_of : compilation -> compilation
val compilation_for_polarity : bool -> compilation -> compilation
+ val is_depth_limited_compilation : compilation -> bool
val string_of_compilation : compilation -> string
val compilation_names : (string * compilation) list
val non_random_compilations : compilation list
@@ -117,6 +119,7 @@
no_higher_order_predicate : string list,
inductify : bool,
detect_switches : bool,
+ smart_depth_limiting : bool,
compilation : compilation
};
val expected_modes : options -> (string * mode list) option
@@ -138,6 +141,7 @@
val no_higher_order_predicate : options -> string list
val is_inductify : options -> bool
val detect_switches : options -> bool
+ val smart_depth_limiting : options -> bool
val compilation : options -> compilation
val default_options : options
val bool_options : string list
@@ -146,6 +150,7 @@
val imp_prems_conv : conv -> conv
(* simple transformations *)
val split_conjuncts_in_assms : Proof.context -> thm -> thm
+ val dest_conjunct_prem : thm -> thm list
val expand_tuples : theory -> thm -> thm
val case_betapply : theory -> term -> term
val eta_contract_ho_arguments : theory -> thm -> thm
@@ -153,7 +158,7 @@
val remove_pointless_clauses : thm -> thm list
val peephole_optimisation : theory -> thm -> thm option
val define_quickcheck_predicate :
- term -> theory -> (((string * typ) * (string * typ) list) * thm) * theory
+ term -> theory -> (((string * typ) * (string * typ) list) * thm) * theory
end;
structure Predicate_Compile_Aux : PREDICATE_COMPILE_AUX =
@@ -632,18 +637,25 @@
(* Different compilations *)
datatype compilation = Pred | Depth_Limited | Random | Depth_Limited_Random | DSeq | Annotated
- | Pos_Random_DSeq | Neg_Random_DSeq | New_Pos_Random_DSeq | New_Neg_Random_DSeq
+ | Pos_Random_DSeq | Neg_Random_DSeq | New_Pos_Random_DSeq | New_Neg_Random_DSeq |
+ Pos_Generator_DSeq | Neg_Generator_DSeq
fun negative_compilation_of Pos_Random_DSeq = Neg_Random_DSeq
| negative_compilation_of Neg_Random_DSeq = Pos_Random_DSeq
| negative_compilation_of New_Pos_Random_DSeq = New_Neg_Random_DSeq
| negative_compilation_of New_Neg_Random_DSeq = New_Pos_Random_DSeq
+ | negative_compilation_of Pos_Generator_DSeq = Neg_Generator_DSeq
+ | negative_compilation_of Pos_Generator_DSeq = Pos_Generator_DSeq
| negative_compilation_of c = c
fun compilation_for_polarity false Pos_Random_DSeq = Neg_Random_DSeq
| compilation_for_polarity false New_Pos_Random_DSeq = New_Neg_Random_DSeq
| compilation_for_polarity _ c = c
+fun is_depth_limited_compilation c =
+ (c = New_Pos_Random_DSeq) orelse (c = New_Neg_Random_DSeq) orelse
+ (c = Pos_Generator_DSeq) orelse (c = Pos_Generator_DSeq)
+
fun string_of_compilation c =
case c of
Pred => ""
@@ -656,14 +668,18 @@
| Neg_Random_DSeq => "neg_random_dseq"
| New_Pos_Random_DSeq => "new_pos_random dseq"
| New_Neg_Random_DSeq => "new_neg_random_dseq"
+ | Pos_Generator_DSeq => "pos_generator_dseq"
+ | Neg_Generator_DSeq => "neg_generator_dseq"
val compilation_names = [("pred", Pred),
("random", Random),
("depth_limited", Depth_Limited),
("depth_limited_random", Depth_Limited_Random),
(*("annotated", Annotated),*)
- ("dseq", DSeq), ("random_dseq", Pos_Random_DSeq),
- ("new_random_dseq", New_Pos_Random_DSeq)]
+ ("dseq", DSeq),
+ ("random_dseq", Pos_Random_DSeq),
+ ("new_random_dseq", New_Pos_Random_DSeq),
+ ("generator_dseq", Pos_Generator_DSeq)]
val non_random_compilations = [Pred, Depth_Limited, DSeq, Annotated]
@@ -729,6 +745,7 @@
no_higher_order_predicate : string list,
inductify : bool,
detect_switches : bool,
+ smart_depth_limiting : bool,
compilation : compilation
};
@@ -759,6 +776,8 @@
fun detect_switches (Options opt) = #detect_switches opt
+fun smart_depth_limiting (Options opt) = #smart_depth_limiting opt
+
val default_options = Options {
expected_modes = NONE,
proposed_modes = [],
@@ -779,12 +798,14 @@
no_higher_order_predicate = [],
inductify = false,
detect_switches = true,
+ smart_depth_limiting = false,
compilation = Pred
}
val bool_options = ["show_steps", "show_intermediate_results", "show_proof_trace", "show_modes",
"show_mode_inference", "show_compilation", "show_invalid_clauses", "skip_proof", "inductify",
- "no_function_flattening", "detect_switches", "specialise", "no_topmost_reordering"]
+ "no_function_flattening", "detect_switches", "specialise", "no_topmost_reordering",
+ "smart_depth_limiting"]
fun print_step options s =
if show_steps options then tracing s else ()
@@ -835,7 +856,14 @@
in
singleton (Variable.export ctxt' ctxt) (split_conjs 1 (Thm.nprems_of fixed_th) fixed_th)
end
-
+
+fun dest_conjunct_prem th =
+ case HOLogic.dest_Trueprop (prop_of th) of
+ (Const (@{const_name HOL.conj}, _) $ t $ t') =>
+ dest_conjunct_prem (th RS @{thm conjunct1})
+ @ dest_conjunct_prem (th RS @{thm conjunct2})
+ | _ => [th]
+
fun expand_tuples thy intro =
let
val ctxt = ProofContext.init_global thy
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_compilations.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_compilations.ML Sun Oct 24 15:11:24 2010 -0700
@@ -157,6 +157,170 @@
end;
+structure DSequence_CompFuns =
+struct
+
+fun mk_dseqT T = Type ("fun", [@{typ code_numeral}, Type ("fun", [@{typ bool},
+ Type (@{type_name Option.option}, [Type (@{type_name Lazy_Sequence.lazy_sequence}, [T])])])])
+
+fun dest_dseqT (Type ("fun", [@{typ code_numeral}, Type ("fun", [@{typ bool},
+ Type (@{type_name Option.option}, [Type (@{type_name Lazy_Sequence.lazy_sequence}, [T])])])])) = T
+ | dest_dseqT T = raise TYPE ("dest_dseqT", [T], []);
+
+fun mk_bot T = Const (@{const_name DSequence.empty}, mk_dseqT T);
+
+fun mk_single t =
+ let val T = fastype_of t
+ in Const(@{const_name DSequence.single}, T --> mk_dseqT T) $ t end;
+
+fun mk_bind (x, f) =
+ let val T as Type ("fun", [_, U]) = fastype_of f
+ in
+ Const (@{const_name DSequence.bind}, fastype_of x --> T --> U) $ x $ f
+ end;
+
+val mk_sup = HOLogic.mk_binop @{const_name DSequence.union};
+
+fun mk_if cond = Const (@{const_name DSequence.if_seq},
+ HOLogic.boolT --> mk_dseqT HOLogic.unitT) $ cond;
+
+fun mk_iterate_upto T (f, from, to) = raise Fail "No iterate_upto compilation"
+
+fun mk_not t = let val T = mk_dseqT HOLogic.unitT
+ in Const (@{const_name DSequence.not_seq}, T --> T) $ t end
+
+fun mk_map T1 T2 tf tp = Const (@{const_name DSequence.map},
+ (T1 --> T2) --> mk_dseqT T1 --> mk_dseqT T2) $ tf $ tp
+
+val compfuns = Predicate_Compile_Aux.CompilationFuns
+ {mk_predT = mk_dseqT, dest_predT = dest_dseqT,
+ mk_bot = mk_bot, mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if,
+ mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}
+
+end;
+
+structure New_Pos_DSequence_CompFuns =
+struct
+
+fun mk_pos_dseqT T =
+ @{typ code_numeral} --> Type (@{type_name Lazy_Sequence.lazy_sequence}, [T])
+
+fun dest_pos_dseqT (Type ("fun", [@{typ code_numeral},
+ Type (@{type_name Lazy_Sequence.lazy_sequence}, [T])])) = T
+ | dest_pos_dseqT T = raise TYPE ("dest_pos_dseqT", [T], []);
+
+fun mk_bot T = Const (@{const_name New_DSequence.pos_empty}, mk_pos_dseqT T);
+
+fun mk_single t =
+ let
+ val T = fastype_of t
+ in Const(@{const_name New_DSequence.pos_single}, T --> mk_pos_dseqT T) $ t end;
+
+fun mk_bind (x, f) =
+ let
+ val T as Type ("fun", [_, U]) = fastype_of f
+ in
+ Const (@{const_name New_DSequence.pos_bind}, fastype_of x --> T --> U) $ x $ f
+ end;
+
+fun mk_decr_bind (x, f) =
+ let
+ val T as Type ("fun", [_, U]) = fastype_of f
+ in
+ Const (@{const_name New_DSequence.pos_decr_bind}, fastype_of x --> T --> U) $ x $ f
+ end;
+
+val mk_sup = HOLogic.mk_binop @{const_name New_DSequence.pos_union};
+
+fun mk_if cond = Const (@{const_name New_DSequence.pos_if_seq},
+ HOLogic.boolT --> mk_pos_dseqT HOLogic.unitT) $ cond;
+
+fun mk_iterate_upto T (f, from, to) = raise Fail "No iterate_upto compilation"
+
+fun mk_not t =
+ let
+ val pT = mk_pos_dseqT HOLogic.unitT
+ val nT =
+ @{typ code_numeral} --> Type (@{type_name Lazy_Sequence.lazy_sequence},
+ [Type (@{type_name Option.option}, [@{typ unit}])])
+ in Const (@{const_name New_DSequence.pos_not_seq}, nT --> pT) $ t end
+
+fun mk_map T1 T2 tf tp = Const (@{const_name New_DSequence.pos_map},
+ (T1 --> T2) --> mk_pos_dseqT T1 --> mk_pos_dseqT T2) $ tf $ tp
+
+val depth_limited_compfuns = Predicate_Compile_Aux.CompilationFuns
+ {mk_predT = mk_pos_dseqT, dest_predT = dest_pos_dseqT,
+ mk_bot = mk_bot, mk_single = mk_single, mk_bind = mk_decr_bind, mk_sup = mk_sup, mk_if = mk_if,
+ mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}
+
+val depth_unlimited_compfuns = Predicate_Compile_Aux.CompilationFuns
+ {mk_predT = mk_pos_dseqT, dest_predT = dest_pos_dseqT,
+ mk_bot = mk_bot, mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if,
+ mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}
+
+end;
+
+structure New_Neg_DSequence_CompFuns =
+struct
+
+fun mk_neg_dseqT T = @{typ code_numeral} -->
+ Type (@{type_name Lazy_Sequence.lazy_sequence}, [Type (@{type_name Option.option}, [T])])
+
+fun dest_neg_dseqT (Type ("fun", [@{typ code_numeral},
+ Type (@{type_name Lazy_Sequence.lazy_sequence}, [Type (@{type_name Option.option}, [T])])])) = T
+ | dest_neg_dseqT T = raise TYPE ("dest_neg_dseqT", [T], []);
+
+fun mk_bot T = Const (@{const_name New_DSequence.neg_empty}, mk_neg_dseqT T);
+
+fun mk_single t =
+ let
+ val T = fastype_of t
+ in Const(@{const_name New_DSequence.neg_single}, T --> mk_neg_dseqT T) $ t end;
+
+fun mk_bind (x, f) =
+ let
+ val T as Type ("fun", [_, U]) = fastype_of f
+ in
+ Const (@{const_name New_DSequence.neg_bind}, fastype_of x --> T --> U) $ x $ f
+ end;
+
+fun mk_decr_bind (x, f) =
+ let
+ val T as Type ("fun", [_, U]) = fastype_of f
+ in
+ Const (@{const_name New_DSequence.neg_decr_bind}, fastype_of x --> T --> U) $ x $ f
+ end;
+
+val mk_sup = HOLogic.mk_binop @{const_name New_DSequence.neg_union};
+
+fun mk_if cond = Const (@{const_name New_DSequence.neg_if_seq},
+ HOLogic.boolT --> mk_neg_dseqT HOLogic.unitT) $ cond;
+
+fun mk_iterate_upto T (f, from, to) = raise Fail "No iterate_upto compilation"
+
+fun mk_not t =
+ let
+ val nT = mk_neg_dseqT HOLogic.unitT
+ val pT =
+ @{typ code_numeral} --> Type (@{type_name Lazy_Sequence.lazy_sequence},
+ [@{typ unit}])
+ in Const (@{const_name New_DSequence.neg_not_seq}, pT --> nT) $ t end
+
+fun mk_map T1 T2 tf tp = Const (@{const_name New_DSequence.neg_map},
+ (T1 --> T2) --> mk_neg_dseqT T1 --> mk_neg_dseqT T2) $ tf $ tp
+
+val depth_limited_compfuns = Predicate_Compile_Aux.CompilationFuns
+ {mk_predT = mk_neg_dseqT, dest_predT = dest_neg_dseqT,
+ mk_bot = mk_bot, mk_single = mk_single, mk_bind = mk_decr_bind, mk_sup = mk_sup, mk_if = mk_if,
+ mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}
+
+val depth_unlimited_compfuns = Predicate_Compile_Aux.CompilationFuns
+ {mk_predT = mk_neg_dseqT, dest_predT = dest_neg_dseqT,
+ mk_bot = mk_bot, mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if,
+ mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}
+
+end;
+
structure New_Pos_Random_Sequence_CompFuns =
struct
@@ -183,6 +347,13 @@
Const (@{const_name New_Random_Sequence.pos_bind}, fastype_of x --> T --> U) $ x $ f
end;
+fun mk_decr_bind (x, f) =
+ let
+ val T as Type ("fun", [_, U]) = fastype_of f
+ in
+ Const (@{const_name New_Random_Sequence.pos_decr_bind}, fastype_of x --> T --> U) $ x $ f
+ end;
+
val mk_sup = HOLogic.mk_binop @{const_name New_Random_Sequence.pos_union};
fun mk_if cond = Const (@{const_name New_Random_Sequence.pos_if_random_dseq},
@@ -206,11 +377,15 @@
fun mk_map T1 T2 tf tp = Const (@{const_name New_Random_Sequence.pos_map},
(T1 --> T2) --> mk_pos_random_dseqT T1 --> mk_pos_random_dseqT T2) $ tf $ tp
-val compfuns = Predicate_Compile_Aux.CompilationFuns
+val depth_limited_compfuns = Predicate_Compile_Aux.CompilationFuns
+ {mk_predT = mk_pos_random_dseqT, dest_predT = dest_pos_random_dseqT,
+ mk_bot = mk_bot, mk_single = mk_single, mk_bind = mk_decr_bind, mk_sup = mk_sup, mk_if = mk_if,
+ mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}
+
+val depth_unlimited_compfuns = Predicate_Compile_Aux.CompilationFuns
{mk_predT = mk_pos_random_dseqT, dest_predT = dest_pos_random_dseqT,
mk_bot = mk_bot, mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if,
mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}
-
end;
structure New_Neg_Random_Sequence_CompFuns =
@@ -241,6 +416,13 @@
Const (@{const_name New_Random_Sequence.neg_bind}, fastype_of x --> T --> U) $ x $ f
end;
+fun mk_decr_bind (x, f) =
+ let
+ val T as Type ("fun", [_, U]) = fastype_of f
+ in
+ Const (@{const_name New_Random_Sequence.neg_decr_bind}, fastype_of x --> T --> U) $ x $ f
+ end;
+
val mk_sup = HOLogic.mk_binop @{const_name New_Random_Sequence.neg_union};
fun mk_if cond = Const (@{const_name New_Random_Sequence.neg_if_random_dseq},
@@ -262,7 +444,12 @@
fun mk_map T1 T2 tf tp = Const (@{const_name New_Random_Sequence.neg_map},
(T1 --> T2) --> mk_neg_random_dseqT T1 --> mk_neg_random_dseqT T2) $ tf $ tp
-val compfuns = Predicate_Compile_Aux.CompilationFuns
+val depth_limited_compfuns = Predicate_Compile_Aux.CompilationFuns
+ {mk_predT = mk_neg_random_dseqT, dest_predT = dest_neg_random_dseqT,
+ mk_bot = mk_bot, mk_single = mk_single, mk_bind = mk_decr_bind, mk_sup = mk_sup, mk_if = mk_if,
+ mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}
+
+val depth_unlimited_compfuns = Predicate_Compile_Aux.CompilationFuns
{mk_predT = mk_neg_random_dseqT, dest_predT = dest_neg_random_dseqT,
mk_bot = mk_bot, mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if,
mk_iterate_upto = mk_iterate_upto, mk_not = mk_not, mk_map = mk_map}
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Sun Oct 24 15:11:24 2010 -0700
@@ -16,33 +16,9 @@
val code_pred_cmd : options -> string -> Proof.context -> Proof.state
val values_cmd : string list -> mode option list option
-> ((string option * bool) * (compilation * int list)) -> int -> string -> Toplevel.state -> unit
- val register_predicate : (string * thm list * thm) -> theory -> theory
- val register_intros : string * thm list -> theory -> theory
- val is_registered : Proof.context -> string -> bool
- val function_name_of : compilation -> Proof.context -> string -> mode -> string
- val predfun_intro_of: Proof.context -> string -> mode -> thm
- val predfun_elim_of: Proof.context -> string -> mode -> thm
- val all_preds_of : Proof.context -> string list
- val modes_of: compilation -> Proof.context -> string -> mode list
- val all_modes_of : compilation -> Proof.context -> (string * mode list) list
- val all_random_modes_of : Proof.context -> (string * mode list) list
- val intros_of : Proof.context -> string -> thm list
- val intros_graph_of : Proof.context -> thm list Graph.T
- val add_intro : string option * thm -> theory -> theory
- val set_elim : thm -> theory -> theory
- val register_alternative_function : string -> mode -> string -> theory -> theory
- val alternative_compilation_of_global : theory -> string -> mode ->
- (compilation_funs -> typ -> term) option
- val alternative_compilation_of : Proof.context -> string -> mode ->
- (compilation_funs -> typ -> term) option
- val functional_compilation : string -> mode -> compilation_funs -> typ -> term
- val force_modes_and_functions : string -> (mode * (string * bool)) list -> theory -> theory
- val force_modes_and_compilations : string ->
- (mode * ((compilation_funs -> typ -> term) * bool)) list -> theory -> theory
- val preprocess_intro : theory -> thm -> thm
+
val print_stored_rules : Proof.context -> unit
val print_all_modes : compilation -> Proof.context -> unit
- val mk_casesrule : Proof.context -> term -> thm list -> term
val put_pred_result : (unit -> term Predicate.pred) -> Proof.context -> Proof.context
val put_pred_random_result : (unit -> int * int -> term Predicate.pred * (int * int)) ->
@@ -78,35 +54,16 @@
type moded_clause = term list * (Predicate_Compile_Aux.indprem * mode_derivation) list
type 'a pred_mode_table = (string * ((bool * mode) * 'a) list) list
- val infer_modes :
- mode_analysis_options -> options ->
- (string -> Predicate_Compile_Aux.mode list) * (string -> Predicate_Compile_Aux.mode list)
- * (string -> Predicate_Compile_Aux.mode -> bool) -> Proof.context -> (string * typ) list ->
- (string * mode list) list ->
- string list -> (string * (Term.term list * Predicate_Compile_Aux.indprem list) list) list ->
- ((moded_clause list pred_mode_table * (string * mode list) list) * string list)
+
end;
structure Predicate_Compile_Core : PREDICATE_COMPILE_CORE =
struct
open Predicate_Compile_Aux;
-
-(** auxiliary **)
-
-(* debug stuff *)
-
-fun print_tac options s =
- if show_proof_trace options then Tactical.print_tac s else Seq.single;
-
-fun assert b = if not b then raise Fail "Assertion failed" else warning "Assertion holds"
-
-datatype assertion = Max_number_of_subgoals of int
-fun assert_tac (Max_number_of_subgoals i) st =
- if (nprems_of st <= i) then Seq.single st
- else raise Fail ("assert_tac: Numbers of subgoals mismatch at goal state :"
- ^ "\n" ^ Pretty.string_of (Pretty.chunks
- (Goal_Display.pretty_goals_without_context st)));
+open Core_Data;
+open Mode_Inference;
+open Predicate_Compile_Proof;
(** fundamentals **)
@@ -147,183 +104,12 @@
Const(@{const_name Code_Evaluation.tracing},
@{typ String.literal} --> (fastype_of t) --> (fastype_of t)) $ (HOLogic.mk_literal s) $ t
-val strip_intro_concl = (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl o prop_of)
-
-(* derivation trees for modes of premises *)
-
-datatype mode_derivation = Mode_App of mode_derivation * mode_derivation | Context of mode
- | Mode_Pair of mode_derivation * mode_derivation | Term of mode
-
-fun string_of_derivation (Mode_App (m1, m2)) =
- "App (" ^ string_of_derivation m1 ^ ", " ^ string_of_derivation m2 ^ ")"
- | string_of_derivation (Mode_Pair (m1, m2)) =
- "Pair (" ^ string_of_derivation m1 ^ ", " ^ string_of_derivation m2 ^ ")"
- | string_of_derivation (Term m) = "Term (" ^ string_of_mode m ^ ")"
- | string_of_derivation (Context m) = "Context (" ^ string_of_mode m ^ ")"
-
-fun strip_mode_derivation deriv =
- let
- fun strip (Mode_App (deriv1, deriv2)) ds = strip deriv1 (deriv2 :: ds)
- | strip deriv ds = (deriv, ds)
- in
- strip deriv []
- end
-
-fun mode_of (Context m) = m
- | mode_of (Term m) = m
- | mode_of (Mode_App (d1, d2)) =
- (case mode_of d1 of Fun (m, m') =>
- (if eq_mode (m, mode_of d2) then m' else raise Fail "mode_of: derivation has mismatching modes")
- | _ => raise Fail "mode_of: derivation has a non-functional mode")
- | mode_of (Mode_Pair (d1, d2)) =
- Pair (mode_of d1, mode_of d2)
-
-fun head_mode_of deriv = mode_of (fst (strip_mode_derivation deriv))
-
-fun param_derivations_of deriv =
- let
- val (_, argument_derivs) = strip_mode_derivation deriv
- fun param_derivation (Mode_Pair (m1, m2)) =
- param_derivation m1 @ param_derivation m2
- | param_derivation (Term _) = []
- | param_derivation m = [m]
- in
- maps param_derivation argument_derivs
- end
-
-fun collect_context_modes (Mode_App (m1, m2)) =
- collect_context_modes m1 @ collect_context_modes m2
- | collect_context_modes (Mode_Pair (m1, m2)) =
- collect_context_modes m1 @ collect_context_modes m2
- | collect_context_modes (Context m) = [m]
- | collect_context_modes (Term _) = []
-
(* representation of inferred clauses with modes *)
type moded_clause = term list * (indprem * mode_derivation) list
type 'a pred_mode_table = (string * ((bool * mode) * 'a) list) list
-(* book-keeping *)
-
-datatype predfun_data = PredfunData of {
- definition : thm,
- intro : thm,
- elim : thm,
- neg_intro : thm option
-};
-
-fun rep_predfun_data (PredfunData data) = data;
-
-fun mk_predfun_data (definition, ((intro, elim), neg_intro)) =
- PredfunData {definition = definition, intro = intro, elim = elim, neg_intro = neg_intro}
-
-datatype pred_data = PredData of {
- intros : (string option * thm) list,
- elim : thm option,
- preprocessed : bool,
- function_names : (compilation * (mode * string) list) list,
- predfun_data : (mode * predfun_data) list,
- needs_random : mode list
-};
-
-fun rep_pred_data (PredData data) = data;
-
-fun mk_pred_data (((intros, elim), preprocessed), (function_names, (predfun_data, needs_random))) =
- PredData {intros = intros, elim = elim, preprocessed = preprocessed,
- function_names = function_names, predfun_data = predfun_data, needs_random = needs_random}
-
-fun map_pred_data f (PredData {intros, elim, preprocessed, function_names, predfun_data, needs_random}) =
- mk_pred_data (f (((intros, elim), preprocessed), (function_names, (predfun_data, needs_random))))
-
-fun eq_option eq (NONE, NONE) = true
- | eq_option eq (SOME x, SOME y) = eq (x, y)
- | eq_option eq _ = false
-
-fun eq_pred_data (PredData d1, PredData d2) =
- eq_list (eq_pair (op =) Thm.eq_thm) (#intros d1, #intros d2) andalso
- eq_option Thm.eq_thm (#elim d1, #elim d2)
-
-structure PredData = Theory_Data
-(
- type T = pred_data Graph.T;
- val empty = Graph.empty;
- val extend = I;
- val merge = Graph.merge eq_pred_data;
-);
-
-(* queries *)
-
-fun lookup_pred_data ctxt name =
- Option.map rep_pred_data (try (Graph.get_node (PredData.get (ProofContext.theory_of ctxt))) name)
-
-fun the_pred_data ctxt name = case lookup_pred_data ctxt name
- of NONE => error ("No such predicate " ^ quote name)
- | SOME data => data;
-
-val is_registered = is_some oo lookup_pred_data
-
-val all_preds_of = Graph.keys o PredData.get o ProofContext.theory_of
-
-val intros_of = map snd o #intros oo the_pred_data
-
-val names_of = map fst o #intros oo the_pred_data
-
-fun the_elim_of ctxt name = case #elim (the_pred_data ctxt name)
- of NONE => error ("No elimination rule for predicate " ^ quote name)
- | SOME thm => thm
-
-val has_elim = is_some o #elim oo the_pred_data
-
-fun function_names_of compilation ctxt name =
- case AList.lookup (op =) (#function_names (the_pred_data ctxt name)) compilation of
- NONE => error ("No " ^ string_of_compilation compilation
- ^ "functions defined for predicate " ^ quote name)
- | SOME fun_names => fun_names
-
-fun function_name_of compilation ctxt name mode =
- case AList.lookup eq_mode
- (function_names_of compilation ctxt name) mode of
- NONE => error ("No " ^ string_of_compilation compilation
- ^ " function defined for mode " ^ string_of_mode mode ^ " of predicate " ^ quote name)
- | SOME function_name => function_name
-
-fun modes_of compilation ctxt name = map fst (function_names_of compilation ctxt name)
-
-fun all_modes_of compilation ctxt =
- map_filter (fn name => Option.map (pair name) (try (modes_of compilation ctxt) name))
- (all_preds_of ctxt)
-
-val all_random_modes_of = all_modes_of Random
-
-fun defined_functions compilation ctxt name = case lookup_pred_data ctxt name of
- NONE => false
- | SOME data => AList.defined (op =) (#function_names data) compilation
-
-fun needs_random ctxt s m =
- member (op =) (#needs_random (the_pred_data ctxt s)) m
-
-fun lookup_predfun_data ctxt name mode =
- Option.map rep_predfun_data
- (AList.lookup (op =) (#predfun_data (the_pred_data ctxt name)) mode)
-
-fun the_predfun_data ctxt name mode =
- case lookup_predfun_data ctxt name mode of
- NONE => error ("No function defined for mode " ^ string_of_mode mode ^
- " of predicate " ^ name)
- | SOME data => data;
-
-val predfun_definition_of = #definition ooo the_predfun_data
-
-val predfun_intro_of = #intro ooo the_predfun_data
-
-val predfun_elim_of = #elim ooo the_predfun_data
-
-val predfun_neg_intro_of = #neg_intro ooo the_predfun_data
-
-val intros_graph_of =
- Graph.map (K (map snd o #intros o rep_pred_data)) o PredData.get o ProofContext.theory_of
-
(* diagnostic display functions *)
fun print_modes options modes =
@@ -342,19 +128,6 @@
val _ = tracing (cat_lines (map print_pred pred_mode_table))
in () end;
-fun string_of_prem ctxt (Prem t) =
- (Syntax.string_of_term ctxt t) ^ "(premise)"
- | string_of_prem ctxt (Negprem t) =
- (Syntax.string_of_term ctxt (HOLogic.mk_not t)) ^ "(negative premise)"
- | string_of_prem ctxt (Sidecond t) =
- (Syntax.string_of_term ctxt t) ^ "(sidecondition)"
- | string_of_prem ctxt _ = raise Fail "string_of_prem: unexpected input"
-
-fun string_of_clause ctxt pred (ts, prems) =
- (space_implode " --> "
- (map (string_of_prem ctxt) prems)) ^ " --> " ^ pred ^ " "
- ^ (space_implode " " (map (Syntax.string_of_term ctxt) ts))
-
fun print_compiled_terms options ctxt =
if show_compilation options then
print_pred_mode_table (fn _ => fn _ => Syntax.string_of_term ctxt)
@@ -465,348 +238,6 @@
ms
end
-
-(* importing introduction rules *)
-
-fun unify_consts thy cs intr_ts =
- (let
- val add_term_consts_2 = fold_aterms (fn Const c => insert (op =) c | _ => I);
- fun varify (t, (i, ts)) =
- let val t' = map_types (Logic.incr_tvar (i + 1)) (#2 (Type.varify_global [] t))
- in (maxidx_of_term t', t'::ts) end;
- val (i, cs') = List.foldr varify (~1, []) cs;
- val (i', intr_ts') = List.foldr varify (i, []) intr_ts;
- val rec_consts = fold add_term_consts_2 cs' [];
- val intr_consts = fold add_term_consts_2 intr_ts' [];
- fun unify (cname, cT) =
- let val consts = map snd (filter (fn c => fst c = cname) intr_consts)
- in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end;
- val (env, _) = fold unify rec_consts (Vartab.empty, i');
- val subst = map_types (Envir.norm_type env)
- in (map subst cs', map subst intr_ts')
- end) handle Type.TUNIFY =>
- (warning "Occurrences of recursive constant have non-unifiable types"; (cs, intr_ts));
-
-fun import_intros inp_pred [] ctxt =
- let
- val ([outp_pred], ctxt') = Variable.import_terms true [inp_pred] ctxt
- val T = fastype_of outp_pred
- val paramTs = ho_argsT_of_typ (binder_types T)
- val (param_names, ctxt'') = Variable.variant_fixes
- (map (fn i => "p" ^ (string_of_int i)) (1 upto (length paramTs))) ctxt'
- val params = map2 (curry Free) param_names paramTs
- in
- (((outp_pred, params), []), ctxt')
- end
- | import_intros inp_pred (th :: ths) ctxt =
- let
- val ((_, [th']), ctxt') = Variable.import true [th] ctxt
- val thy = ProofContext.theory_of ctxt'
- val (pred, args) = strip_intro_concl th'
- val T = fastype_of pred
- val ho_args = ho_args_of_typ T args
- fun subst_of (pred', pred) =
- let
- val subst = Sign.typ_match thy (fastype_of pred', fastype_of pred) Vartab.empty
- handle Type.TYPE_MATCH => error ("Type mismatch of predicate " ^ fst (dest_Const pred)
- ^ " (trying to match " ^ Syntax.string_of_typ ctxt (fastype_of pred')
- ^ " and " ^ Syntax.string_of_typ ctxt (fastype_of pred) ^ ")"
- ^ " in " ^ Display.string_of_thm ctxt th)
- in map (fn (indexname, (s, T)) => ((indexname, s), T)) (Vartab.dest subst) end
- fun instantiate_typ th =
- let
- val (pred', _) = strip_intro_concl th
- val _ = if not (fst (dest_Const pred) = fst (dest_Const pred')) then
- raise Fail "Trying to instantiate another predicate" else ()
- in Thm.certify_instantiate (subst_of (pred', pred), []) th end;
- fun instantiate_ho_args th =
- let
- val (_, args') = (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl o prop_of) th
- val ho_args' = map dest_Var (ho_args_of_typ T args')
- in Thm.certify_instantiate ([], ho_args' ~~ ho_args) th end
- val outp_pred =
- Term_Subst.instantiate (subst_of (inp_pred, pred), []) inp_pred
- val ((_, ths'), ctxt1) =
- Variable.import false (map (instantiate_typ #> instantiate_ho_args) ths) ctxt'
- in
- (((outp_pred, ho_args), th' :: ths'), ctxt1)
- end
-
-(* generation of case rules from user-given introduction rules *)
-
-fun mk_args2 (Type (@{type_name Product_Type.prod}, [T1, T2])) st =
- let
- val (t1, st') = mk_args2 T1 st
- val (t2, st'') = mk_args2 T2 st'
- in
- (HOLogic.mk_prod (t1, t2), st'')
- end
- (*| mk_args2 (T as Type ("fun", _)) (params, ctxt) =
- let
- val (S, U) = strip_type T
- in
- if U = HOLogic.boolT then
- (hd params, (tl params, ctxt))
- else
- let
- val ([x], ctxt') = Variable.variant_fixes ["x"] ctxt
- in
- (Free (x, T), (params, ctxt'))
- end
- end*)
- | mk_args2 T (params, ctxt) =
- let
- val ([x], ctxt') = Variable.variant_fixes ["x"] ctxt
- in
- (Free (x, T), (params, ctxt'))
- end
-
-fun mk_casesrule ctxt pred introrules =
- let
- (* TODO: can be simplified if parameters are not treated specially ? *)
- val (((pred, params), intros_th), ctxt1) = import_intros pred introrules ctxt
- (* TODO: distinct required ? -- test case with more than one parameter! *)
- val params = distinct (op aconv) params
- val intros = map prop_of intros_th
- val ([propname], ctxt2) = Variable.variant_fixes ["thesis"] ctxt1
- val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT))
- val argsT = binder_types (fastype_of pred)
- (* TODO: can be simplified if parameters are not treated specially ? <-- see uncommented code! *)
- val (argvs, _) = fold_map mk_args2 argsT (params, ctxt2)
- fun mk_case intro =
- let
- val (_, args) = (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl) intro
- val prems = Logic.strip_imp_prems intro
- val eqprems =
- map2 (HOLogic.mk_Trueprop oo (curry HOLogic.mk_eq)) argvs args
- val frees = map Free (fold Term.add_frees (args @ prems) [])
- in fold Logic.all frees (Logic.list_implies (eqprems @ prems, prop)) end
- val assm = HOLogic.mk_Trueprop (list_comb (pred, argvs))
- val cases = map mk_case intros
- in Logic.list_implies (assm :: cases, prop) end;
-
-fun dest_conjunct_prem th =
- case HOLogic.dest_Trueprop (prop_of th) of
- (Const (@{const_name HOL.conj}, _) $ t $ t') =>
- dest_conjunct_prem (th RS @{thm conjunct1})
- @ dest_conjunct_prem (th RS @{thm conjunct2})
- | _ => [th]
-
-fun prove_casesrule ctxt (pred, (pre_cases_rule, nparams)) cases_rule =
- let
- val thy = ProofContext.theory_of ctxt
- val nargs = length (binder_types (fastype_of pred))
- fun PEEK f dependent_tactic st = dependent_tactic (f st) st
- fun meta_eq_of th = th RS @{thm eq_reflection}
- val tuple_rew_rules = map meta_eq_of [@{thm fst_conv}, @{thm snd_conv}, @{thm Pair_eq}]
- fun instantiate i n {context = ctxt, params = p, prems = prems,
- asms = a, concl = cl, schematics = s} =
- let
- fun term_pair_of (ix, (ty,t)) = (Var (ix,ty), t)
- fun inst_of_matches tts = fold (Pattern.match thy) tts (Vartab.empty, Vartab.empty)
- |> snd |> Vartab.dest |> map (pairself (cterm_of thy) o term_pair_of)
- val (cases, (eqs, prems)) = apsnd (chop (nargs - nparams)) (chop n prems)
- val case_th = MetaSimplifier.simplify true
- (@{thm Predicate.eq_is_eq} :: map meta_eq_of eqs) (nth cases (i - 1))
- val prems' = maps (dest_conjunct_prem o MetaSimplifier.simplify true tuple_rew_rules) prems
- val pats = map (swap o HOLogic.dest_eq o HOLogic.dest_Trueprop) (take nargs (prems_of case_th))
- val case_th' = Thm.instantiate ([], inst_of_matches pats) case_th
- OF (replicate nargs @{thm refl})
- val thesis =
- Thm.instantiate ([], inst_of_matches (prems_of case_th' ~~ map prop_of prems')) case_th'
- OF prems'
- in
- (rtac thesis 1)
- end
- val tac =
- etac pre_cases_rule 1
- THEN
- (PEEK nprems_of
- (fn n =>
- ALLGOALS (fn i =>
- MetaSimplifier.rewrite_goal_tac [@{thm split_paired_all}] i
- THEN (SUBPROOF (instantiate i n) ctxt i))))
- in
- Goal.prove ctxt (Term.add_free_names cases_rule []) [] cases_rule (fn _ => tac)
- end
-
-(** preprocessing rules **)
-
-fun Trueprop_conv cv ct =
- case Thm.term_of ct of
- Const (@{const_name Trueprop}, _) $ _ => Conv.arg_conv cv ct
- | _ => raise Fail "Trueprop_conv"
-
-fun preprocess_equality thy rule =
- Conv.fconv_rule
- (imp_prems_conv
- (Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq})))))
- (Thm.transfer thy rule)
-
-fun preprocess_intro thy = expand_tuples thy #> preprocess_equality thy
-
-(* fetching introduction rules or registering introduction rules *)
-
-val no_compilation = ([], ([], []))
-
-fun fetch_pred_data ctxt name =
- case try (Inductive.the_inductive ctxt) name of
- SOME (info as (_, result)) =>
- let
- fun is_intro_of intro =
- let
- val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro))
- in (fst (dest_Const const) = name) end;
- val thy = ProofContext.theory_of ctxt
- val intros = map (preprocess_intro thy) (filter is_intro_of (#intrs result))
- val index = find_index (fn s => s = name) (#names (fst info))
- val pre_elim = nth (#elims result) index
- val pred = nth (#preds result) index
- val elim_t = mk_casesrule ctxt pred intros
- val nparams = length (Inductive.params_of (#raw_induct result))
- val elim = prove_casesrule ctxt (pred, (pre_elim, nparams)) elim_t
- in
- mk_pred_data (((map (pair NONE) intros, SOME elim), true), no_compilation)
- end
- | NONE => error ("No such predicate: " ^ quote name)
-
-fun add_predfun_data name mode data =
- let
- val add = (apsnd o apsnd o apfst) (cons (mode, mk_predfun_data data))
- in PredData.map (Graph.map_node name (map_pred_data add)) end
-
-fun is_inductive_predicate ctxt name =
- is_some (try (Inductive.the_inductive ctxt) name)
-
-fun depending_preds_of ctxt (key, value) =
- let
- val intros = map (Thm.prop_of o snd) ((#intros o rep_pred_data) value)
- in
- fold Term.add_const_names intros []
- |> (fn cs =>
- if member (op =) cs @{const_name "HOL.eq"} then
- insert (op =) @{const_name Predicate.eq} cs
- else cs)
- |> filter (fn c => (not (c = key)) andalso
- (is_inductive_predicate ctxt c orelse is_registered ctxt c))
- end;
-
-fun add_intro (opt_case_name, thm) thy =
- let
- val (name, T) = dest_Const (fst (strip_intro_concl thm))
- fun cons_intro gr =
- case try (Graph.get_node gr) name of
- SOME pred_data => Graph.map_node name (map_pred_data
- (apfst (apfst (apfst (fn intros => intros @ [(opt_case_name, thm)]))))) gr
- | NONE => Graph.new_node (name, mk_pred_data ((([(opt_case_name, thm)], NONE), false), no_compilation)) gr
- in PredData.map cons_intro thy end
-
-fun set_elim thm =
- let
- val (name, _) = dest_Const (fst
- (strip_comb (HOLogic.dest_Trueprop (hd (prems_of thm)))))
- in PredData.map (Graph.map_node name (map_pred_data (apfst (apfst (apsnd (K (SOME thm))))))) end
-
-fun register_predicate (constname, intros, elim) thy =
- let
- val named_intros = map (pair NONE) intros
- in
- if not (member (op =) (Graph.keys (PredData.get thy)) constname) then
- PredData.map
- (Graph.new_node (constname,
- mk_pred_data (((named_intros, SOME elim), false), no_compilation))) thy
- else thy
- end
-
-fun register_intros (constname, pre_intros) thy =
- let
- val T = Sign.the_const_type thy constname
- fun constname_of_intro intr = fst (dest_Const (fst (strip_intro_concl intr)))
- val _ = if not (forall (fn intr => constname_of_intro intr = constname) pre_intros) then
- error ("register_intros: Introduction rules of different constants are used\n" ^
- "expected rules for " ^ constname ^ ", but received rules for " ^
- commas (map constname_of_intro pre_intros))
- else ()
- val pred = Const (constname, T)
- val pre_elim =
- (Drule.export_without_context o Skip_Proof.make_thm thy)
- (mk_casesrule (ProofContext.init_global thy) pred pre_intros)
- in register_predicate (constname, pre_intros, pre_elim) thy end
-
-fun defined_function_of compilation pred =
- let
- val set = (apsnd o apfst) (cons (compilation, []))
- in
- PredData.map (Graph.map_node pred (map_pred_data set))
- end
-
-fun set_function_name compilation pred mode name =
- let
- val set = (apsnd o apfst)
- (AList.map_default (op =) (compilation, [(mode, name)]) (cons (mode, name)))
- in
- PredData.map (Graph.map_node pred (map_pred_data set))
- end
-
-fun set_needs_random name modes =
- let
- val set = (apsnd o apsnd o apsnd) (K modes)
- in
- PredData.map (Graph.map_node name (map_pred_data set))
- end
-
-(* registration of alternative function names *)
-
-structure Alt_Compilations_Data = Theory_Data
-(
- type T = (mode * (compilation_funs -> typ -> term)) list Symtab.table;
- val empty = Symtab.empty;
- val extend = I;
- fun merge data : T = Symtab.merge (K true) data;
-);
-
-fun alternative_compilation_of_global thy pred_name mode =
- AList.lookup eq_mode (Symtab.lookup_list (Alt_Compilations_Data.get thy) pred_name) mode
-
-fun alternative_compilation_of ctxt pred_name mode =
- AList.lookup eq_mode
- (Symtab.lookup_list (Alt_Compilations_Data.get (ProofContext.theory_of ctxt)) pred_name) mode
-
-fun force_modes_and_compilations pred_name compilations =
- let
- (* thm refl is a dummy thm *)
- val modes = map fst compilations
- val (needs_random, non_random_modes) = pairself (map fst)
- (List.partition (fn (m, (fun_name, random)) => random) compilations)
- val non_random_dummys = map (rpair "dummy") non_random_modes
- val all_dummys = map (rpair "dummy") modes
- val dummy_function_names = map (rpair all_dummys) Predicate_Compile_Aux.random_compilations
- @ map (rpair non_random_dummys) Predicate_Compile_Aux.non_random_compilations
- val alt_compilations = map (apsnd fst) compilations
- in
- PredData.map (Graph.new_node
- (pred_name, mk_pred_data ((([], SOME @{thm refl}), true), (dummy_function_names, ([], needs_random)))))
- #> Alt_Compilations_Data.map (Symtab.insert (K false) (pred_name, alt_compilations))
- end
-
-fun functional_compilation fun_name mode compfuns T =
- let
- val (inpTs, outpTs) = split_map_modeT (fn _ => fn T => (SOME T, NONE))
- mode (binder_types T)
- val bs = map (pair "x") inpTs
- val bounds = map Bound (rev (0 upto (length bs) - 1))
- val f = Const (fun_name, inpTs ---> HOLogic.mk_tupleT outpTs)
- in list_abs (bs, mk_single compfuns (list_comb (f, bounds))) end
-
-fun register_alternative_function pred_name mode fun_name =
- Alt_Compilations_Data.map (Symtab.insert_list (eq_pair eq_mode (K false))
- (pred_name, (mode, functional_compilation fun_name mode)))
-
-fun force_modes_and_functions pred_name fun_names =
- force_modes_and_compilations pred_name
- (map (fn (mode, (fun_name, random)) => (mode, (functional_compilation fun_name mode, random)))
- fun_names)
-
(* compilation modifiers *)
structure Comp_Mod =
@@ -839,8 +270,37 @@
val wrap_compilation = #wrap_compilation o dest_comp_modifiers
val transform_additional_arguments = #transform_additional_arguments o dest_comp_modifiers
+fun set_compfuns compfuns' (Comp_Modifiers {compilation, function_name_prefix, compfuns, mk_random,
+ modify_funT, additional_arguments, wrap_compilation, transform_additional_arguments}) =
+ (Comp_Modifiers {compilation = compilation, function_name_prefix = function_name_prefix,
+ compfuns = compfuns', mk_random = mk_random, modify_funT = modify_funT,
+ additional_arguments = additional_arguments, wrap_compilation = wrap_compilation,
+ transform_additional_arguments = transform_additional_arguments})
+
end;
+fun unlimited_compfuns_of true New_Pos_Random_DSeq =
+ New_Pos_Random_Sequence_CompFuns.depth_unlimited_compfuns
+ | unlimited_compfuns_of false New_Pos_Random_DSeq =
+ New_Neg_Random_Sequence_CompFuns.depth_unlimited_compfuns
+ | unlimited_compfuns_of true Pos_Generator_DSeq =
+ New_Pos_DSequence_CompFuns.depth_unlimited_compfuns
+ | unlimited_compfuns_of false Pos_Generator_DSeq =
+ New_Neg_DSequence_CompFuns.depth_unlimited_compfuns
+ | unlimited_compfuns_of _ c =
+ raise Fail ("No unlimited compfuns for compilation " ^ string_of_compilation c)
+
+fun limited_compfuns_of true Predicate_Compile_Aux.New_Pos_Random_DSeq =
+ New_Pos_Random_Sequence_CompFuns.depth_limited_compfuns
+ | limited_compfuns_of false Predicate_Compile_Aux.New_Pos_Random_DSeq =
+ New_Neg_Random_Sequence_CompFuns.depth_limited_compfuns
+ | limited_compfuns_of true Pos_Generator_DSeq =
+ New_Pos_DSequence_CompFuns.depth_limited_compfuns
+ | limited_compfuns_of false Pos_Generator_DSeq =
+ New_Neg_DSequence_CompFuns.depth_limited_compfuns
+ | limited_compfuns_of _ c =
+ raise Fail ("No limited compfuns for compilation " ^ string_of_compilation c)
+
val depth_limited_comp_modifiers = Comp_Mod.Comp_Modifiers
{
compilation = Depth_Limited,
@@ -979,7 +439,7 @@
compilation = DSeq,
function_name_prefix = "dseq_",
compfuns = DSequence_CompFuns.compfuns,
- mk_random = (fn _ => error "no random generation"),
+ mk_random = (fn _ => error "no random generation in dseq"),
modify_funT = I,
additional_arguments = K [],
wrap_compilation = K (K (K (K (K I))))
@@ -1028,7 +488,7 @@
{
compilation = New_Pos_Random_DSeq,
function_name_prefix = "new_random_dseq_",
- compfuns = New_Pos_Random_Sequence_CompFuns.compfuns,
+ compfuns = New_Pos_Random_Sequence_CompFuns.depth_limited_compfuns,
mk_random = (fn T => fn additional_arguments =>
let
val random = Const ("Quickcheck.random_class.random",
@@ -1050,7 +510,42 @@
{
compilation = New_Neg_Random_DSeq,
function_name_prefix = "new_random_dseq_neg_",
- compfuns = New_Neg_Random_Sequence_CompFuns.compfuns,
+ compfuns = New_Neg_Random_Sequence_CompFuns.depth_limited_compfuns,
+ mk_random = (fn _ => error "no random generation"),
+ modify_funT = I,
+ additional_arguments = K [],
+ wrap_compilation = K (K (K (K (K I))))
+ : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
+ transform_additional_arguments = K I : (indprem -> term list -> term list)
+ }
+
+val pos_generator_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
+ {
+ compilation = Pos_Generator_DSeq,
+ function_name_prefix = "generator_dseq_",
+ compfuns = New_Pos_DSequence_CompFuns.depth_limited_compfuns,
+ mk_random =
+ (fn T => fn additional_arguments =>
+ let
+ val small_lazy =
+ Const (@{const_name "Lazy_Sequence.small_lazy_class.small_lazy"},
+ @{typ "Int.int"} --> Type (@{type_name "Lazy_Sequence.lazy_sequence"}, [T]))
+ in
+ absdummy (@{typ code_numeral}, small_lazy $ HOLogic.mk_number @{typ int} 3)
+ end
+ ),
+ modify_funT = I,
+ additional_arguments = K [],
+ wrap_compilation = K (K (K (K (K I))))
+ : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
+ transform_additional_arguments = K I : (indprem -> term list -> term list)
+ }
+
+val neg_generator_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
+ {
+ compilation = Neg_Generator_DSeq,
+ function_name_prefix = "generator_dseq_neg_",
+ compfuns = New_Neg_DSequence_CompFuns.depth_limited_compfuns,
mk_random = (fn _ => error "no random generation"),
modify_funT = I,
additional_arguments = K [],
@@ -1064,527 +559,11 @@
Pos_Random_DSeq => neg_random_dseq_comp_modifiers
| Neg_Random_DSeq => pos_random_dseq_comp_modifiers
| New_Pos_Random_DSeq => new_neg_random_dseq_comp_modifiers
- | New_Neg_Random_DSeq => new_pos_random_dseq_comp_modifiers
+ | New_Neg_Random_DSeq => new_pos_random_dseq_comp_modifiers
+ | Pos_Generator_DSeq => neg_generator_dseq_comp_modifiers
+ | Neg_Generator_DSeq => pos_generator_dseq_comp_modifiers
| c => comp_modifiers)
-(** mode analysis **)
-
-type mode_analysis_options =
- {use_generators : bool,
- reorder_premises : bool,
- infer_pos_and_neg_modes : bool}
-
-fun is_constrt thy =
- let
- val cnstrs = flat (maps
- (map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd)
- (Symtab.dest (Datatype.get_all thy)));
- fun check t = (case strip_comb t of
- (Free _, []) => true
- | (Const (s, T), ts) => (case (AList.lookup (op =) cnstrs s, body_type T) of
- (SOME (i, Tname), Type (Tname', _)) =>
- length ts = i andalso Tname = Tname' andalso forall check ts
- | _ => false)
- | _ => false)
- in check end;
-
-(*** check if a type is an equality type (i.e. doesn't contain fun)
- FIXME this is only an approximation ***)
-fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts
- | is_eqT _ = true;
-
-fun term_vs tm = fold_aterms (fn Free (x, T) => cons x | _ => I) tm [];
-val terms_vs = distinct (op =) o maps term_vs;
-
-(** collect all Frees in a term (with duplicates!) **)
-fun term_vTs tm =
- fold_aterms (fn Free xT => cons xT | _ => I) tm [];
-
-fun subsets i j =
- if i <= j then
- let
- fun merge xs [] = xs
- | merge [] ys = ys
- | merge (x::xs) (y::ys) = if length x >= length y then x::merge xs (y::ys)
- else y::merge (x::xs) ys;
- val is = subsets (i+1) j
- in merge (map (fn ks => i::ks) is) is end
- else [[]];
-
-fun print_failed_mode options thy modes p (pol, m) rs is =
- if show_mode_inference options then
- let
- val _ = tracing ("Clauses " ^ commas (map (fn i => string_of_int (i + 1)) is) ^ " of " ^
- p ^ " violates mode " ^ string_of_mode m)
- in () end
- else ()
-
-fun error_of p (pol, m) is =
- " Clauses " ^ commas (map (fn i => string_of_int (i + 1)) is) ^ " of " ^
- p ^ " violates mode " ^ string_of_mode m
-
-fun is_all_input mode =
- let
- fun is_all_input' (Fun _) = true
- | is_all_input' (Pair (m1, m2)) = is_all_input' m1 andalso is_all_input' m2
- | is_all_input' Input = true
- | is_all_input' Output = false
- in
- forall is_all_input' (strip_fun_mode mode)
- end
-
-fun all_input_of T =
- let
- val (Ts, U) = strip_type T
- fun input_of (Type (@{type_name Product_Type.prod}, [T1, T2])) = Pair (input_of T1, input_of T2)
- | input_of _ = Input
- in
- if U = HOLogic.boolT then
- fold_rev (curry Fun) (map input_of Ts) Bool
- else
- raise Fail "all_input_of: not a predicate"
- end
-
-fun find_least ord xs =
- let
- fun find' x y = (case y of NONE => SOME x | SOME y' => if ord (x, y') = LESS then SOME x else y)
- in
- fold find' xs NONE
- end
-
-fun term_vs tm = fold_aterms (fn Free (x, T) => cons x | _ => I) tm [];
-val terms_vs = distinct (op =) o maps term_vs;
-
-fun input_mode T =
- let
- val (Ts, U) = strip_type T
- in
- fold_rev (curry Fun) (map (K Input) Ts) Input
- end
-
-fun output_mode T =
- let
- val (Ts, U) = strip_type T
- in
- fold_rev (curry Fun) (map (K Output) Ts) Output
- end
-
-fun is_invertible_function ctxt (Const (f, _)) = is_constr ctxt f
- | is_invertible_function ctxt _ = false
-
-fun non_invertible_subterms ctxt (t as Free _) = []
- | non_invertible_subterms ctxt t =
- let
- val (f, args) = strip_comb t
- in
- if is_invertible_function ctxt f then
- maps (non_invertible_subterms ctxt) args
- else
- [t]
- end
-
-fun collect_non_invertible_subterms ctxt (f as Free _) (names, eqs) = (f, (names, eqs))
- | collect_non_invertible_subterms ctxt t (names, eqs) =
- case (strip_comb t) of (f, args) =>
- if is_invertible_function ctxt f then
- let
- val (args', (names', eqs')) =
- fold_map (collect_non_invertible_subterms ctxt) args (names, eqs)
- in
- (list_comb (f, args'), (names', eqs'))
- end
- else
- let
- val s = Name.variant names "x"
- val v = Free (s, fastype_of t)
- in
- (v, (s :: names, HOLogic.mk_eq (v, t) :: eqs))
- end
-(*
- if is_constrt thy t then (t, (names, eqs)) else
- let
- val s = Name.variant names "x"
- val v = Free (s, fastype_of t)
- in (v, (s::names, HOLogic.mk_eq (v, t)::eqs)) end;
-*)
-
-fun is_possible_output ctxt vs t =
- forall
- (fn t => is_eqT (fastype_of t) andalso forall (member (op =) vs) (term_vs t))
- (non_invertible_subterms ctxt t)
- andalso
- (forall (is_eqT o snd)
- (inter (fn ((f', _), f) => f = f') vs (Term.add_frees t [])))
-
-fun vars_of_destructable_term ctxt (Free (x, _)) = [x]
- | vars_of_destructable_term ctxt t =
- let
- val (f, args) = strip_comb t
- in
- if is_invertible_function ctxt f then
- maps (vars_of_destructable_term ctxt) args
- else
- []
- end
-
-fun is_constructable vs t = forall (member (op =) vs) (term_vs t)
-
-fun missing_vars vs t = subtract (op =) vs (term_vs t)
-
-fun output_terms (Const (@{const_name Pair}, _) $ t1 $ t2, Mode_Pair (d1, d2)) =
- output_terms (t1, d1) @ output_terms (t2, d2)
- | output_terms (t1 $ t2, Mode_App (d1, d2)) =
- output_terms (t1, d1) @ output_terms (t2, d2)
- | output_terms (t, Term Output) = [t]
- | output_terms _ = []
-
-fun lookup_mode modes (Const (s, T)) =
- (case (AList.lookup (op =) modes s) of
- SOME ms => SOME (map (fn m => (Context m, [])) ms)
- | NONE => NONE)
- | lookup_mode modes (Free (x, _)) =
- (case (AList.lookup (op =) modes x) of
- SOME ms => SOME (map (fn m => (Context m , [])) ms)
- | NONE => NONE)
-
-fun derivations_of (ctxt : Proof.context) modes vs (Const (@{const_name Pair}, _) $ t1 $ t2) (Pair (m1, m2)) =
- map_product
- (fn (m1, mvars1) => fn (m2, mvars2) => (Mode_Pair (m1, m2), union (op =) mvars1 mvars2))
- (derivations_of ctxt modes vs t1 m1) (derivations_of ctxt modes vs t2 m2)
- | derivations_of ctxt modes vs t (m as Fun _) =
- (*let
- val (p, args) = strip_comb t
- in
- (case lookup_mode modes p of
- SOME ms => map_filter (fn (Context m, []) => let
- val ms = strip_fun_mode m
- val (argms, restms) = chop (length args) ms
- val m' = fold_rev (curry Fun) restms Bool
- in
- if forall (fn m => eq_mode (Input, m)) argms andalso eq_mode (m', mode) then
- SOME (fold (curry Mode_App) (map Term argms) (Context m), missing_vars vs t)
- else NONE
- end) ms
- | NONE => (if is_all_input mode then [(Context mode, [])] else []))
- end*)
- (case try (all_derivations_of ctxt modes vs) t of
- SOME derivs =>
- filter (fn (d, mvars) => eq_mode (mode_of d, m) andalso null (output_terms (t, d))) derivs
- | NONE => (if is_all_input m then [(Context m, [])] else []))
- | derivations_of ctxt modes vs t m =
- if eq_mode (m, Input) then
- [(Term Input, missing_vars vs t)]
- else if eq_mode (m, Output) then
- (if is_possible_output ctxt vs t then [(Term Output, [])] else [])
- else []
-and all_derivations_of ctxt modes vs (Const (@{const_name Pair}, _) $ t1 $ t2) =
- let
- val derivs1 = all_derivations_of ctxt modes vs t1
- val derivs2 = all_derivations_of ctxt modes vs t2
- in
- map_product
- (fn (m1, mvars1) => fn (m2, mvars2) => (Mode_Pair (m1, m2), union (op =) mvars1 mvars2))
- derivs1 derivs2
- end
- | all_derivations_of ctxt modes vs (t1 $ t2) =
- let
- val derivs1 = all_derivations_of ctxt modes vs t1
- in
- maps (fn (d1, mvars1) =>
- case mode_of d1 of
- Fun (m', _) => map (fn (d2, mvars2) =>
- (Mode_App (d1, d2), union (op =) mvars1 mvars2)) (derivations_of ctxt modes vs t2 m')
- | _ => raise Fail "all_derivations_of: derivation has an unexpected non-functional mode") derivs1
- end
- | all_derivations_of _ modes vs (Const (s, T)) = the (lookup_mode modes (Const (s, T)))
- | all_derivations_of _ modes vs (Free (x, T)) = the (lookup_mode modes (Free (x, T)))
- | all_derivations_of _ modes vs _ = raise Fail "all_derivations_of: unexpected term"
-
-fun rev_option_ord ord (NONE, NONE) = EQUAL
- | rev_option_ord ord (NONE, SOME _) = GREATER
- | rev_option_ord ord (SOME _, NONE) = LESS
- | rev_option_ord ord (SOME x, SOME y) = ord (x, y)
-
-fun random_mode_in_deriv modes t deriv =
- case try dest_Const (fst (strip_comb t)) of
- SOME (s, _) =>
- (case AList.lookup (op =) modes s of
- SOME ms =>
- (case AList.lookup (op =) (map (fn ((p, m), r) => (m, r)) ms) (head_mode_of deriv) of
- SOME r => r
- | NONE => false)
- | NONE => false)
- | NONE => false
-
-fun number_of_output_positions mode =
- let
- val args = strip_fun_mode mode
- fun contains_output (Fun _) = false
- | contains_output Input = false
- | contains_output Output = true
- | contains_output (Pair (m1, m2)) = contains_output m1 orelse contains_output m2
- in
- length (filter contains_output args)
- end
-
-fun lex_ord ord1 ord2 (x, x') =
- case ord1 (x, x') of
- EQUAL => ord2 (x, x')
- | ord => ord
-
-fun lexl_ord [] (x, x') = EQUAL
- | lexl_ord (ord :: ords') (x, x') =
- case ord (x, x') of
- EQUAL => lexl_ord ords' (x, x')
- | ord => ord
-
-fun deriv_ord' ctxt pol pred modes t1 t2 ((deriv1, mvars1), (deriv2, mvars2)) =
- let
- (* prefer functional modes if it is a function *)
- fun fun_ord ((t1, deriv1, mvars1), (t2, deriv2, mvars2)) =
- let
- fun is_functional t mode =
- case try (fst o dest_Const o fst o strip_comb) t of
- NONE => false
- | SOME c => is_some (alternative_compilation_of ctxt c mode)
- in
- case (is_functional t1 (head_mode_of deriv1), is_functional t2 (head_mode_of deriv2)) of
- (true, true) => EQUAL
- | (true, false) => LESS
- | (false, true) => GREATER
- | (false, false) => EQUAL
- end
- (* prefer modes without requirement for generating random values *)
- fun mvars_ord ((t1, deriv1, mvars1), (t2, deriv2, mvars2)) =
- int_ord (length mvars1, length mvars2)
- (* prefer non-random modes *)
- fun random_mode_ord ((t1, deriv1, mvars1), (t2, deriv2, mvars2)) =
- int_ord (if random_mode_in_deriv modes t1 deriv1 then 1 else 0,
- if random_mode_in_deriv modes t2 deriv2 then 1 else 0)
- (* prefer modes with more input and less output *)
- fun output_mode_ord ((t1, deriv1, mvars1), (t2, deriv2, mvars2)) =
- int_ord (number_of_output_positions (head_mode_of deriv1),
- number_of_output_positions (head_mode_of deriv2))
- (* prefer recursive calls *)
- fun is_rec_premise t =
- case fst (strip_comb t) of Const (c, T) => c = pred | _ => false
- fun recursive_ord ((t1, deriv1, mvars1), (t2, deriv2, mvars2)) =
- int_ord (if is_rec_premise t1 then 0 else 1,
- if is_rec_premise t2 then 0 else 1)
- val ord = lexl_ord [mvars_ord, fun_ord, random_mode_ord, output_mode_ord, recursive_ord]
- in
- ord ((t1, deriv1, mvars1), (t2, deriv2, mvars2))
- end
-
-fun deriv_ord ctxt pol pred modes t = deriv_ord' ctxt pol pred modes t t
-
-fun premise_ord thy pol pred modes ((prem1, a1), (prem2, a2)) =
- rev_option_ord (deriv_ord' thy pol pred modes (dest_indprem prem1) (dest_indprem prem2)) (a1, a2)
-
-fun print_mode_list modes =
- tracing ("modes: " ^ (commas (map (fn (s, ms) => s ^ ": " ^
- commas (map (fn (m, r) => string_of_mode m ^ (if r then " random " else " not ")) ms)) modes)))
-
-fun select_mode_prem (mode_analysis_options : mode_analysis_options) (ctxt : Proof.context) pred
- pol (modes, (pos_modes, neg_modes)) vs ps =
- let
- fun choose_mode_of_prem (Prem t) =
- find_least (deriv_ord ctxt pol pred modes t) (all_derivations_of ctxt pos_modes vs t)
- | choose_mode_of_prem (Sidecond t) = SOME (Context Bool, missing_vars vs t)
- | choose_mode_of_prem (Negprem t) = find_least (deriv_ord ctxt (not pol) pred modes t)
- (filter (fn (d, missing_vars) => is_all_input (head_mode_of d))
- (all_derivations_of ctxt neg_modes vs t))
- | choose_mode_of_prem p = raise Fail ("choose_mode_of_prem: unexpected premise " ^ string_of_prem ctxt p)
- in
- if #reorder_premises mode_analysis_options then
- find_least (premise_ord ctxt pol pred modes) (ps ~~ map choose_mode_of_prem ps)
- else
- SOME (hd ps, choose_mode_of_prem (hd ps))
- end
-
-fun check_mode_clause' (mode_analysis_options : mode_analysis_options) ctxt pred param_vs (modes :
- (string * ((bool * mode) * bool) list) list) ((pol, mode) : bool * mode) (ts, ps) =
- let
- val vTs = distinct (op =) (fold Term.add_frees (map dest_indprem ps) (fold Term.add_frees ts []))
- val modes' = modes @ (param_vs ~~ map (fn x => [((true, x), false), ((false, x), false)]) (ho_arg_modes_of mode))
- fun retrieve_modes_of_pol pol = map (fn (s, ms) =>
- (s, map_filter (fn ((p, m), r) => if p = pol then SOME m else NONE | _ => NONE) ms))
- val (pos_modes', neg_modes') =
- if #infer_pos_and_neg_modes mode_analysis_options then
- (retrieve_modes_of_pol pol modes', retrieve_modes_of_pol (not pol) modes')
- else
- let
- val modes = map (fn (s, ms) => (s, map (fn ((p, m), r) => m) ms)) modes'
- in (modes, modes) end
- val (in_ts, out_ts) = split_mode mode ts
- val in_vs = union (op =) param_vs (maps (vars_of_destructable_term ctxt) in_ts)
- val out_vs = terms_vs out_ts
- fun known_vs_after p vs = (case p of
- Prem t => union (op =) vs (term_vs t)
- | Sidecond t => union (op =) vs (term_vs t)
- | Negprem t => union (op =) vs (term_vs t)
- | _ => raise Fail "unexpected premise")
- fun check_mode_prems acc_ps rnd vs [] = SOME (acc_ps, vs, rnd)
- | check_mode_prems acc_ps rnd vs ps =
- (case
- (select_mode_prem mode_analysis_options ctxt pred pol (modes', (pos_modes', neg_modes')) vs ps) of
- SOME (p, SOME (deriv, [])) => check_mode_prems ((p, deriv) :: acc_ps) rnd
- (known_vs_after p vs) (filter_out (equal p) ps)
- | SOME (p, SOME (deriv, missing_vars)) =>
- if #use_generators mode_analysis_options andalso pol then
- check_mode_prems ((p, deriv) :: (map
- (fn v => (Generator (v, the (AList.lookup (op =) vTs v)), Term Output))
- (distinct (op =) missing_vars))
- @ acc_ps) true (known_vs_after p vs) (filter_out (equal p) ps)
- else NONE
- | SOME (p, NONE) => NONE
- | NONE => NONE)
- in
- case check_mode_prems [] false in_vs ps of
- NONE => NONE
- | SOME (acc_ps, vs, rnd) =>
- if forall (is_constructable vs) (in_ts @ out_ts) then
- SOME (ts, rev acc_ps, rnd)
- else
- if #use_generators mode_analysis_options andalso pol then
- let
- val generators = map
- (fn v => (Generator (v, the (AList.lookup (op =) vTs v)), Term Output))
- (subtract (op =) vs (terms_vs (in_ts @ out_ts)))
- in
- SOME (ts, rev (generators @ acc_ps), true)
- end
- else
- NONE
- end
-
-datatype result = Success of bool | Error of string
-
-fun check_modes_pred' mode_analysis_options options thy param_vs clauses modes (p, (ms : ((bool * mode) * bool) list)) =
- let
- fun split xs =
- let
- fun split' [] (ys, zs) = (rev ys, rev zs)
- | split' ((m, Error z) :: xs) (ys, zs) = split' xs (ys, z :: zs)
- | split' (((m : bool * mode), Success rnd) :: xs) (ys, zs) = split' xs ((m, rnd) :: ys, zs)
- in
- split' xs ([], [])
- end
- val rs = these (AList.lookup (op =) clauses p)
- fun check_mode m =
- let
- val res = Output.cond_timeit false "work part of check_mode for one mode" (fn _ =>
- map (check_mode_clause' mode_analysis_options thy p param_vs modes m) rs)
- in
- Output.cond_timeit false "aux part of check_mode for one mode" (fn _ =>
- case find_indices is_none res of
- [] => Success (exists (fn SOME (_, _, true) => true | _ => false) res)
- | is => (print_failed_mode options thy modes p m rs is; Error (error_of p m is)))
- end
- val _ = if show_mode_inference options then
- tracing ("checking " ^ string_of_int (length ms) ^ " modes ...")
- else ()
- val res = Output.cond_timeit false "check_mode" (fn _ => map (fn (m, _) => (m, check_mode m)) ms)
- val (ms', errors) = split res
- in
- ((p, (ms' : ((bool * mode) * bool) list)), errors)
- end;
-
-fun get_modes_pred' mode_analysis_options thy param_vs clauses modes (p, ms) =
- let
- val rs = these (AList.lookup (op =) clauses p)
- in
- (p, map (fn (m, rnd) =>
- (m, map
- ((fn (ts, ps, rnd) => (ts, ps)) o the o
- check_mode_clause' mode_analysis_options thy p param_vs modes m) rs)) ms)
- end;
-
-fun fixp f (x : (string * ((bool * mode) * bool) list) list) =
- let val y = f x
- in if x = y then x else fixp f y end;
-
-fun fixp_with_state f (x : (string * ((bool * mode) * bool) list) list, state) =
- let
- val (y, state') = f (x, state)
- in
- if x = y then (y, state') else fixp_with_state f (y, state')
- end
-
-fun string_of_ext_mode ((pol, mode), rnd) =
- string_of_mode mode ^ "(" ^ (if pol then "pos" else "neg") ^ ", "
- ^ (if rnd then "rnd" else "nornd") ^ ")"
-
-fun print_extra_modes options modes =
- if show_mode_inference options then
- tracing ("Modes of inferred predicates: " ^
- cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map string_of_ext_mode ms)) modes))
- else ()
-
-fun infer_modes mode_analysis_options options (lookup_mode, lookup_neg_mode, needs_random) ctxt
- preds all_modes param_vs clauses =
- let
- fun appair f (x1, x2) (y1, y2) = (f x1 y1, f x2 y2)
- fun add_needs_random s (false, m) = ((false, m), false)
- | add_needs_random s (true, m) = ((true, m), needs_random s m)
- fun add_polarity_and_random_bit s b ms = map (fn m => add_needs_random s (b, m)) ms
- val prednames = map fst preds
- (* extramodes contains all modes of all constants, should we only use the necessary ones
- - what is the impact on performance? *)
- fun predname_of (Prem t) =
- (case try dest_Const (fst (strip_comb t)) of SOME (c, _) => insert (op =) c | NONE => I)
- | predname_of (Negprem t) =
- (case try dest_Const (fst (strip_comb t)) of SOME (c, _) => insert (op =) c | NONE => I)
- | predname_of _ = I
- val relevant_prednames = fold (fn (_, clauses') =>
- fold (fn (_, ps) => fold Term.add_const_names (map dest_indprem ps)) clauses') clauses []
- |> filter_out (fn name => member (op =) prednames name)
- val extra_modes =
- if #infer_pos_and_neg_modes mode_analysis_options then
- let
- val pos_extra_modes =
- map_filter (fn name => Option.map (pair name) (try lookup_mode name))
- relevant_prednames
- val neg_extra_modes =
- map_filter (fn name => Option.map (pair name) (try lookup_neg_mode name))
- relevant_prednames
- in
- map (fn (s, ms) => (s, (add_polarity_and_random_bit s true ms)
- @ add_polarity_and_random_bit s false (the (AList.lookup (op =) neg_extra_modes s))))
- pos_extra_modes
- end
- else
- map (fn (s, ms) => (s, (add_polarity_and_random_bit s true ms)))
- (map_filter (fn name => Option.map (pair name) (try lookup_mode name))
- relevant_prednames)
- val _ = print_extra_modes options extra_modes
- val start_modes =
- if #infer_pos_and_neg_modes mode_analysis_options then
- map (fn (s, ms) => (s, map (fn m => ((true, m), false)) ms @
- (map (fn m => ((false, m), false)) ms))) all_modes
- else
- map (fn (s, ms) => (s, map (fn m => ((true, m), false)) ms)) all_modes
- fun iteration modes = map
- (check_modes_pred' mode_analysis_options options ctxt param_vs clauses
- (modes @ extra_modes)) modes
- val ((modes : (string * ((bool * mode) * bool) list) list), errors) =
- Output.cond_timeit false "Fixpount computation of mode analysis" (fn () =>
- if show_invalid_clauses options then
- fixp_with_state (fn (modes, errors) =>
- let
- val (modes', new_errors) = split_list (iteration modes)
- in (modes', errors @ flat new_errors) end) (start_modes, [])
- else
- (fixp (fn modes => map fst (iteration modes)) start_modes, []))
- val moded_clauses = map (get_modes_pred' mode_analysis_options ctxt param_vs clauses
- (modes @ extra_modes)) modes
- val need_random = fold (fn (s, ms) => if member (op =) (map fst preds) s then
- cons (s, (map_filter (fn ((true, m), true) => SOME m | _ => NONE) ms)) else I)
- modes []
- in
- ((moded_clauses, need_random), errors)
- end;
-
(* term construction *)
fun mk_v (names, vs) s T = (case AList.lookup (op =) vs s of
@@ -1964,8 +943,16 @@
fun compile_pred options compilation_modifiers ctxt all_vs param_vs s T (pol, mode) moded_cls =
let
- val compilation_modifiers = if pol then compilation_modifiers else
- negative_comp_modifiers_of compilation_modifiers
+ val is_terminating = false (* FIXME: requires an termination analysis *)
+ val compilation_modifiers =
+ (if pol then compilation_modifiers else
+ negative_comp_modifiers_of compilation_modifiers)
+ |> (if is_depth_limited_compilation (Comp_Mod.compilation compilation_modifiers) then
+ (if is_terminating then
+ (Comp_Mod.set_compfuns (unlimited_compfuns_of pol (Comp_Mod.compilation compilation_modifiers)))
+ else
+ (Comp_Mod.set_compfuns (limited_compfuns_of pol (Comp_Mod.compilation compilation_modifiers))))
+ else I)
val additional_arguments = Comp_Mod.additional_arguments compilation_modifiers
(all_vs @ param_vs)
val compfuns = Comp_Mod.compfuns compilation_modifiers
@@ -2014,11 +1001,6 @@
(HOLogic.mk_eq (list_comb (fun_const, in_ts @ additional_arguments), compilation))
end;
-(** special setup for simpset **)
-val HOL_basic_ss' = HOL_basic_ss addsimps (@{thms HOL.simp_thms} @ [@{thm Pair_eq}])
- setSolver (mk_solver "all_tac_solver" (fn _ => fn _ => all_tac))
- setSolver (mk_solver "True_solver" (fn _ => rtac @{thm TrueI}))
-
(* Definition of executable functions and their intro and elim rules *)
fun print_arities arities = tracing ("Arities:\n" ^
@@ -2182,483 +1164,6 @@
|> fold create_definition modes
end;
-(* Proving equivalence of term *)
-
-fun is_Type (Type _) = true
- | is_Type _ = false
-
-(* returns true if t is an application of an datatype constructor *)
-(* which then consequently would be splitted *)
-(* else false *)
-fun is_constructor thy t =
- if (is_Type (fastype_of t)) then
- (case Datatype.get_info thy ((fst o dest_Type o fastype_of) t) of
- NONE => false
- | SOME info => (let
- val constr_consts = maps (fn (_, (_, _, constrs)) => map fst constrs) (#descr info)
- val (c, _) = strip_comb t
- in (case c of
- Const (name, _) => member (op =) constr_consts name
- | _ => false) end))
- else false
-
-(* MAJOR FIXME: prove_params should be simple
- - different form of introrule for parameters ? *)
-
-fun prove_param options ctxt nargs t deriv =
- let
- val (f, args) = strip_comb (Envir.eta_contract t)
- val mode = head_mode_of deriv
- val param_derivations = param_derivations_of deriv
- val ho_args = ho_args_of mode args
- val f_tac = case f of
- Const (name, T) => simp_tac (HOL_basic_ss addsimps
- [@{thm eval_pred}, predfun_definition_of ctxt name mode,
- @{thm split_eta}, @{thm split_beta}, @{thm fst_conv},
- @{thm snd_conv}, @{thm pair_collapse}, @{thm Product_Type.split_conv}]) 1
- | Free _ =>
- Subgoal.FOCUS_PREMS (fn {context = ctxt, params = params, prems, asms, concl, schematics} =>
- let
- val prems' = maps dest_conjunct_prem (take nargs prems)
- in
- MetaSimplifier.rewrite_goal_tac
- (map (fn th => th RS @{thm sym} RS @{thm eq_reflection}) prems') 1
- end) ctxt 1
- | Abs _ => raise Fail "prove_param: No valid parameter term"
- in
- REPEAT_DETERM (rtac @{thm ext} 1)
- THEN print_tac options "prove_param"
- THEN f_tac
- THEN print_tac options "after prove_param"
- THEN (REPEAT_DETERM (atac 1))
- THEN (EVERY (map2 (prove_param options ctxt nargs) ho_args param_derivations))
- THEN REPEAT_DETERM (rtac @{thm refl} 1)
- end
-
-fun prove_expr options ctxt nargs (premposition : int) (t, deriv) =
- case strip_comb t of
- (Const (name, T), args) =>
- let
- val mode = head_mode_of deriv
- val introrule = predfun_intro_of ctxt name mode
- val param_derivations = param_derivations_of deriv
- val ho_args = ho_args_of mode args
- in
- print_tac options "before intro rule:"
- THEN rtac introrule 1
- THEN print_tac options "after intro rule"
- (* for the right assumption in first position *)
- THEN rotate_tac premposition 1
- THEN atac 1
- THEN print_tac options "parameter goal"
- (* work with parameter arguments *)
- THEN (EVERY (map2 (prove_param options ctxt nargs) ho_args param_derivations))
- THEN (REPEAT_DETERM (atac 1))
- end
- | (Free _, _) =>
- print_tac options "proving parameter call.."
- THEN Subgoal.FOCUS_PREMS (fn {context = ctxt, params, prems, asms, concl, schematics} =>
- let
- val param_prem = nth prems premposition
- val (param, _) = strip_comb (HOLogic.dest_Trueprop (prop_of param_prem))
- val prems' = maps dest_conjunct_prem (take nargs prems)
- fun param_rewrite prem =
- param = snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of prem)))
- val SOME rew_eq = find_first param_rewrite prems'
- val param_prem' = MetaSimplifier.rewrite_rule
- (map (fn th => th RS @{thm eq_reflection})
- [rew_eq RS @{thm sym}, @{thm split_beta}, @{thm fst_conv}, @{thm snd_conv}])
- param_prem
- in
- rtac param_prem' 1
- end) ctxt 1
- THEN print_tac options "after prove parameter call"
-
-fun SOLVED tac st = FILTER (fn st' => nprems_of st' = nprems_of st - 1) tac st;
-
-fun SOLVEDALL tac st = FILTER (fn st' => nprems_of st' = 0) tac st
-
-fun check_format ctxt st =
- let
- val concl' = Logic.strip_assums_concl (hd (prems_of st))
- val concl = HOLogic.dest_Trueprop concl'
- val expr = fst (strip_comb (fst (PredicateCompFuns.dest_Eval concl)))
- fun valid_expr (Const (@{const_name Predicate.bind}, _)) = true
- | valid_expr (Const (@{const_name Predicate.single}, _)) = true
- | valid_expr _ = false
- in
- if valid_expr expr then
- ((*tracing "expression is valid";*) Seq.single st)
- else
- ((*tracing "expression is not valid";*) Seq.empty) (* error "check_format: wrong format" *)
- end
-
-fun prove_match options ctxt nargs out_ts =
- let
- val thy = ProofContext.theory_of ctxt
- val eval_if_P =
- @{lemma "P ==> Predicate.eval x z ==> Predicate.eval (if P then x else y) z" by simp}
- fun get_case_rewrite t =
- if (is_constructor thy t) then let
- val case_rewrites = (#case_rewrites (Datatype.the_info thy
- ((fst o dest_Type o fastype_of) t)))
- in fold (union Thm.eq_thm) (case_rewrites :: map get_case_rewrite (snd (strip_comb t))) [] end
- else []
- val simprules = insert Thm.eq_thm @{thm "unit.cases"} (insert Thm.eq_thm @{thm "prod.cases"}
- (fold (union Thm.eq_thm) (map get_case_rewrite out_ts) []))
- (* replace TRY by determining if it necessary - are there equations when calling compile match? *)
- in
- (* make this simpset better! *)
- asm_full_simp_tac (HOL_basic_ss' addsimps simprules) 1
- THEN print_tac options "after prove_match:"
- THEN (DETERM (TRY
- (rtac eval_if_P 1
- THEN (SUBPROOF (fn {context = ctxt, params, prems, asms, concl, schematics} =>
- (REPEAT_DETERM (rtac @{thm conjI} 1
- THEN (SOLVED (asm_simp_tac HOL_basic_ss' 1))))
- THEN print_tac options "if condition to be solved:"
- THEN asm_simp_tac HOL_basic_ss' 1
- THEN TRY (
- let
- val prems' = maps dest_conjunct_prem (take nargs prems)
- in
- MetaSimplifier.rewrite_goal_tac
- (map (fn th => th RS @{thm sym} RS @{thm eq_reflection}) prems') 1
- end
- THEN REPEAT_DETERM (rtac @{thm refl} 1))
- THEN print_tac options "after if simp; in SUBPROOF") ctxt 1))))
- THEN print_tac options "after if simplification"
- end;
-
-(* corresponds to compile_fun -- maybe call that also compile_sidecond? *)
-
-fun prove_sidecond ctxt t =
- let
- fun preds_of t nameTs = case strip_comb t of
- (f as Const (name, T), args) =>
- if is_registered ctxt name then (name, T) :: nameTs
- else fold preds_of args nameTs
- | _ => nameTs
- val preds = preds_of t []
- val defs = map
- (fn (pred, T) => predfun_definition_of ctxt pred
- (all_input_of T))
- preds
- in
- simp_tac (HOL_basic_ss addsimps
- (@{thms HOL.simp_thms} @ (@{thm eval_pred} :: defs))) 1
- (* need better control here! *)
- end
-
-fun prove_clause options ctxt nargs mode (_, clauses) (ts, moded_ps) =
- let
- val (in_ts, clause_out_ts) = split_mode mode ts;
- fun prove_prems out_ts [] =
- (prove_match options ctxt nargs out_ts)
- THEN print_tac options "before simplifying assumptions"
- THEN asm_full_simp_tac HOL_basic_ss' 1
- THEN print_tac options "before single intro rule"
- THEN Subgoal.FOCUS_PREMS
- (fn {context = ctxt, params = params, prems, asms, concl, schematics} =>
- let
- val prems' = maps dest_conjunct_prem (take nargs prems)
- in
- MetaSimplifier.rewrite_goal_tac
- (map (fn th => th RS @{thm sym} RS @{thm eq_reflection}) prems') 1
- end) ctxt 1
- THEN (rtac (if null clause_out_ts then @{thm singleI_unit} else @{thm singleI}) 1)
- | prove_prems out_ts ((p, deriv) :: ps) =
- let
- val premposition = (find_index (equal p) clauses) + nargs
- val mode = head_mode_of deriv
- val rest_tac =
- rtac @{thm bindI} 1
- THEN (case p of Prem t =>
- let
- val (_, us) = strip_comb t
- val (_, out_ts''') = split_mode mode us
- val rec_tac = prove_prems out_ts''' ps
- in
- print_tac options "before clause:"
- (*THEN asm_simp_tac HOL_basic_ss 1*)
- THEN print_tac options "before prove_expr:"
- THEN prove_expr options ctxt nargs premposition (t, deriv)
- THEN print_tac options "after prove_expr:"
- THEN rec_tac
- end
- | Negprem t =>
- let
- val (t, args) = strip_comb t
- val (_, out_ts''') = split_mode mode args
- val rec_tac = prove_prems out_ts''' ps
- val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
- val neg_intro_rule =
- Option.map (fn name =>
- the (predfun_neg_intro_of ctxt name mode)) name
- val param_derivations = param_derivations_of deriv
- val params = ho_args_of mode args
- in
- print_tac options "before prove_neg_expr:"
- THEN full_simp_tac (HOL_basic_ss addsimps
- [@{thm split_eta}, @{thm split_beta}, @{thm fst_conv},
- @{thm snd_conv}, @{thm pair_collapse}, @{thm Product_Type.split_conv}]) 1
- THEN (if (is_some name) then
- print_tac options "before applying not introduction rule"
- THEN Subgoal.FOCUS_PREMS
- (fn {context = ctxt, params = params, prems, asms, concl, schematics} =>
- rtac (the neg_intro_rule) 1
- THEN rtac (nth prems premposition) 1) ctxt 1
- THEN print_tac options "after applying not introduction rule"
- THEN (EVERY (map2 (prove_param options ctxt nargs) params param_derivations))
- THEN (REPEAT_DETERM (atac 1))
- else
- rtac @{thm not_predI'} 1
- (* test: *)
- THEN dtac @{thm sym} 1
- THEN asm_full_simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1)
- THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
- THEN rec_tac
- end
- | Sidecond t =>
- rtac @{thm if_predI} 1
- THEN print_tac options "before sidecond:"
- THEN prove_sidecond ctxt t
- THEN print_tac options "after sidecond:"
- THEN prove_prems [] ps)
- in (prove_match options ctxt nargs out_ts)
- THEN rest_tac
- end;
- val prems_tac = prove_prems in_ts moded_ps
- in
- print_tac options "Proving clause..."
- THEN rtac @{thm bindI} 1
- THEN rtac @{thm singleI} 1
- THEN prems_tac
- end;
-
-fun select_sup 1 1 = []
- | select_sup _ 1 = [rtac @{thm supI1}]
- | select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1));
-
-fun prove_one_direction options ctxt clauses preds pred mode moded_clauses =
- let
- val T = the (AList.lookup (op =) preds pred)
- val nargs = length (binder_types T)
- val pred_case_rule = the_elim_of ctxt pred
- in
- REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"}))
- THEN print_tac options "before applying elim rule"
- THEN etac (predfun_elim_of ctxt pred mode) 1
- THEN etac pred_case_rule 1
- THEN print_tac options "after applying elim rule"
- THEN (EVERY (map
- (fn i => EVERY' (select_sup (length moded_clauses) i) i)
- (1 upto (length moded_clauses))))
- THEN (EVERY (map2 (prove_clause options ctxt nargs mode) clauses moded_clauses))
- THEN print_tac options "proved one direction"
- end;
-
-(** Proof in the other direction **)
-
-fun prove_match2 options ctxt out_ts =
- let
- val thy = ProofContext.theory_of ctxt
- fun split_term_tac (Free _) = all_tac
- | split_term_tac t =
- if (is_constructor thy t) then
- let
- val info = Datatype.the_info thy ((fst o dest_Type o fastype_of) t)
- val num_of_constrs = length (#case_rewrites info)
- val (_, ts) = strip_comb t
- in
- print_tac options ("Term " ^ (Syntax.string_of_term ctxt t) ^
- "splitting with rules \n" ^ Display.string_of_thm ctxt (#split_asm info))
- THEN TRY ((Splitter.split_asm_tac [#split_asm info] 1)
- THEN (print_tac options "after splitting with split_asm rules")
- (* THEN (Simplifier.asm_full_simp_tac HOL_basic_ss 1)
- THEN (DETERM (TRY (etac @{thm Pair_inject} 1)))*)
- THEN (REPEAT_DETERM_N (num_of_constrs - 1)
- (etac @{thm botE} 1 ORELSE etac @{thm botE} 2)))
- THEN (assert_tac (Max_number_of_subgoals 2))
- THEN (EVERY (map split_term_tac ts))
- end
- else all_tac
- in
- split_term_tac (HOLogic.mk_tuple out_ts)
- THEN (DETERM (TRY ((Splitter.split_asm_tac [@{thm "split_if_asm"}] 1)
- THEN (etac @{thm botE} 2))))
- end
-
-(* VERY LARGE SIMILIRATIY to function prove_param
--- join both functions
-*)
-(* TODO: remove function *)
-
-fun prove_param2 options ctxt t deriv =
- let
- val (f, args) = strip_comb (Envir.eta_contract t)
- val mode = head_mode_of deriv
- val param_derivations = param_derivations_of deriv
- val ho_args = ho_args_of mode args
- val f_tac = case f of
- Const (name, T) => full_simp_tac (HOL_basic_ss addsimps
- (@{thm eval_pred}::(predfun_definition_of ctxt name mode)
- :: @{thm "Product_Type.split_conv"}::[])) 1
- | Free _ => all_tac
- | _ => error "prove_param2: illegal parameter term"
- in
- print_tac options "before simplification in prove_args:"
- THEN f_tac
- THEN print_tac options "after simplification in prove_args"
- THEN EVERY (map2 (prove_param2 options ctxt) ho_args param_derivations)
- end
-
-fun prove_expr2 options ctxt (t, deriv) =
- (case strip_comb t of
- (Const (name, T), args) =>
- let
- val mode = head_mode_of deriv
- val param_derivations = param_derivations_of deriv
- val ho_args = ho_args_of mode args
- in
- etac @{thm bindE} 1
- THEN (REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"})))
- THEN print_tac options "prove_expr2-before"
- THEN etac (predfun_elim_of ctxt name mode) 1
- THEN print_tac options "prove_expr2"
- THEN (EVERY (map2 (prove_param2 options ctxt) ho_args param_derivations))
- THEN print_tac options "finished prove_expr2"
- end
- | _ => etac @{thm bindE} 1)
-
-fun prove_sidecond2 options ctxt t = let
- fun preds_of t nameTs = case strip_comb t of
- (f as Const (name, T), args) =>
- if is_registered ctxt name then (name, T) :: nameTs
- else fold preds_of args nameTs
- | _ => nameTs
- val preds = preds_of t []
- val defs = map
- (fn (pred, T) => predfun_definition_of ctxt pred
- (all_input_of T))
- preds
- in
- (* only simplify the one assumption *)
- full_simp_tac (HOL_basic_ss' addsimps @{thm eval_pred} :: defs) 1
- (* need better control here! *)
- THEN print_tac options "after sidecond2 simplification"
- end
-
-fun prove_clause2 options ctxt pred mode (ts, ps) i =
- let
- val pred_intro_rule = nth (intros_of ctxt pred) (i - 1)
- val (in_ts, clause_out_ts) = split_mode mode ts;
- val split_ss = HOL_basic_ss' addsimps [@{thm split_eta}, @{thm split_beta},
- @{thm fst_conv}, @{thm snd_conv}, @{thm pair_collapse}]
- fun prove_prems2 out_ts [] =
- print_tac options "before prove_match2 - last call:"
- THEN prove_match2 options ctxt out_ts
- THEN print_tac options "after prove_match2 - last call:"
- THEN (etac @{thm singleE} 1)
- THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
- THEN (asm_full_simp_tac HOL_basic_ss' 1)
- THEN TRY (
- (REPEAT_DETERM (etac @{thm Pair_inject} 1))
- THEN (asm_full_simp_tac HOL_basic_ss' 1)
-
- THEN SOLVED (print_tac options "state before applying intro rule:"
- THEN (rtac pred_intro_rule
- (* How to handle equality correctly? *)
- THEN_ALL_NEW (K (print_tac options "state before assumption matching")
- THEN' (atac ORELSE' ((CHANGED o asm_full_simp_tac split_ss) THEN' (TRY o atac)))
- THEN' (K (print_tac options "state after pre-simplification:"))
- THEN' (K (print_tac options "state after assumption matching:")))) 1))
- | prove_prems2 out_ts ((p, deriv) :: ps) =
- let
- val mode = head_mode_of deriv
- val rest_tac = (case p of
- Prem t =>
- let
- val (_, us) = strip_comb t
- val (_, out_ts''') = split_mode mode us
- val rec_tac = prove_prems2 out_ts''' ps
- in
- (prove_expr2 options ctxt (t, deriv)) THEN rec_tac
- end
- | Negprem t =>
- let
- val (_, args) = strip_comb t
- val (_, out_ts''') = split_mode mode args
- val rec_tac = prove_prems2 out_ts''' ps
- val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
- val param_derivations = param_derivations_of deriv
- val ho_args = ho_args_of mode args
- in
- print_tac options "before neg prem 2"
- THEN etac @{thm bindE} 1
- THEN (if is_some name then
- full_simp_tac (HOL_basic_ss addsimps
- [predfun_definition_of ctxt (the name) mode]) 1
- THEN etac @{thm not_predE} 1
- THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
- THEN (EVERY (map2 (prove_param2 options ctxt) ho_args param_derivations))
- else
- etac @{thm not_predE'} 1)
- THEN rec_tac
- end
- | Sidecond t =>
- etac @{thm bindE} 1
- THEN etac @{thm if_predE} 1
- THEN prove_sidecond2 options ctxt t
- THEN prove_prems2 [] ps)
- in print_tac options "before prove_match2:"
- THEN prove_match2 options ctxt out_ts
- THEN print_tac options "after prove_match2:"
- THEN rest_tac
- end;
- val prems_tac = prove_prems2 in_ts ps
- in
- print_tac options "starting prove_clause2"
- THEN etac @{thm bindE} 1
- THEN (etac @{thm singleE'} 1)
- THEN (TRY (etac @{thm Pair_inject} 1))
- THEN print_tac options "after singleE':"
- THEN prems_tac
- end;
-
-fun prove_other_direction options ctxt pred mode moded_clauses =
- let
- fun prove_clause clause i =
- (if i < length moded_clauses then etac @{thm supE} 1 else all_tac)
- THEN (prove_clause2 options ctxt pred mode clause i)
- in
- (DETERM (TRY (rtac @{thm unit.induct} 1)))
- THEN (REPEAT_DETERM (CHANGED (rewtac @{thm split_paired_all})))
- THEN (rtac (predfun_intro_of ctxt pred mode) 1)
- THEN (REPEAT_DETERM (rtac @{thm refl} 2))
- THEN (if null moded_clauses then
- etac @{thm botE} 1
- else EVERY (map2 prove_clause moded_clauses (1 upto (length moded_clauses))))
- end;
-
-(** proof procedure **)
-
-fun prove_pred options thy clauses preds pred (pol, mode) (moded_clauses, compiled_term) =
- let
- val ctxt = ProofContext.init_global thy
- val clauses = case AList.lookup (op =) clauses pred of SOME rs => rs | NONE => []
- in
- Goal.prove ctxt (Term.add_free_names compiled_term []) [] compiled_term
- (if not (skip_proof options) then
- (fn _ =>
- rtac @{thm pred_iffI} 1
- THEN print_tac options "after pred_iffI"
- THEN prove_one_direction options ctxt clauses preds pred mode moded_clauses
- THEN print_tac options "proved one direction"
- THEN prove_other_direction options ctxt pred mode moded_clauses
- THEN print_tac options "proved other direction")
- else (fn _ => Skip_Proof.cheat_tac thy))
- end;
(* composition of mode inference, definition, compilation and proof *)
@@ -2905,51 +1410,16 @@
in
thy'''
end
-
-fun extend' value_of edges_of key (G, visited) =
- let
- val (G', v) = case try (Graph.get_node G) key of
- SOME v => (G, v)
- | NONE => (Graph.new_node (key, value_of key) G, value_of key)
- val (G'', visited') = fold (extend' value_of edges_of)
- (subtract (op =) visited (edges_of (key, v)))
- (G', key :: visited)
- in
- (fold (Graph.add_edge o (pair key)) (edges_of (key, v)) G'', visited')
- end;
-
-fun extend value_of edges_of key G = fst (extend' value_of edges_of key (G, []))
fun gen_add_equations steps options names thy =
let
fun dest_steps (Steps s) = s
val defined = defined_functions (Comp_Mod.compilation (#comp_modifiers (dest_steps steps)))
- val ctxt = ProofContext.init_global thy
- val thy' = thy
- |> PredData.map (fold (extend (fetch_pred_data ctxt) (depending_preds_of ctxt)) names)
- |> Theory.checkpoint;
+ val thy' = extend_intro_graph names thy |> Theory.checkpoint;
fun strong_conn_of gr keys =
Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
val scc = strong_conn_of (PredData.get thy') names
- fun preprocess name thy =
- PredData.map (Graph.map_node name (map_pred_data (apfst (fn (rules, preprocessed) =>
- if preprocessed then (rules, preprocessed)
- else
- let
- val (named_intros, SOME elim) = rules
- val named_intros' = map (apsnd (preprocess_intro thy)) named_intros
- val pred = Const (name, Sign.the_const_type thy name)
- val ctxt = ProofContext.init_global thy
- val elim_t = mk_casesrule ctxt pred (map snd named_intros')
- val nparams = (case try (Inductive.the_inductive ctxt) name of
- SOME (_, result) => length (Inductive.params_of (#raw_induct result))
- | NONE => 0)
- val elim' = prove_casesrule ctxt (pred, (elim, 0)) elim_t
- in
- ((named_intros', SOME elim'), true)
- end))))
- thy
- val thy'' = fold preprocess (flat scc) thy'
+ val thy'' = fold preprocess_intros (flat scc) thy'
val thy''' = fold_rev
(fn preds => fn thy =>
if not (forall (defined (ProofContext.init_global thy)) preds) then
@@ -3062,9 +1532,9 @@
let
val pos_modes = map_filter (fn (true, m) => SOME m | _ => NONE) modes
val neg_modes = map_filter (fn (false, m) => SOME m | _ => NONE) modes
- in define_functions new_pos_random_dseq_comp_modifiers New_Pos_Random_Sequence_CompFuns.compfuns
+ in define_functions new_pos_random_dseq_comp_modifiers New_Pos_Random_Sequence_CompFuns.depth_limited_compfuns
options preds (s, pos_modes)
- #> define_functions new_neg_random_dseq_comp_modifiers New_Neg_Random_Sequence_CompFuns.compfuns
+ #> define_functions new_neg_random_dseq_comp_modifiers New_Neg_Random_Sequence_CompFuns.depth_limited_compfuns
options preds (s, neg_modes)
end,
prove = prove_by_skip,
@@ -3073,6 +1543,25 @@
use_generators = true,
qname = "new_random_dseq_equation"})
+val add_generator_dseq_equations = gen_add_equations
+ (Steps {
+ define_functions =
+ fn options => fn preds => fn (s, modes) =>
+ let
+ val pos_modes = map_filter (fn (true, m) => SOME m | _ => NONE) modes
+ val neg_modes = map_filter (fn (false, m) => SOME m | _ => NONE) modes
+ in
+ define_functions pos_generator_dseq_comp_modifiers New_Pos_DSequence_CompFuns.depth_limited_compfuns
+ options preds (s, pos_modes)
+ #> define_functions neg_generator_dseq_comp_modifiers New_Neg_DSequence_CompFuns.depth_limited_compfuns
+ options preds (s, neg_modes)
+ end,
+ prove = prove_by_skip,
+ add_code_equations = K (K I),
+ comp_modifiers = pos_generator_dseq_comp_modifiers,
+ use_generators = true,
+ qname = "generator_dseq_equation"})
+
(** user interface **)
(* code_pred_intro attribute *)
@@ -3097,9 +1586,7 @@
let
val thy = ProofContext.theory_of lthy
val const = prep_const thy raw_const
- val ctxt = ProofContext.init_global thy
- val lthy' = Local_Theory.background_theory (PredData.map
- (extend (fetch_pred_data ctxt) (depending_preds_of ctxt) const)) lthy
+ val lthy' = Local_Theory.background_theory (extend_intro_graph [const]) lthy
val thy' = ProofContext.theory_of lthy'
val ctxt' = ProofContext.init_global thy'
val preds = Graph.all_succs (PredData.get thy') [const] |> filter_out (has_elim ctxt')
@@ -3134,6 +1621,7 @@
| Random => add_random_equations
| Depth_Limited_Random => add_depth_limited_random_equations
| New_Pos_Random_DSeq => add_new_random_dseq_equations
+ | Pos_Generator_DSeq => add_generator_dseq_equations
| compilation => error ("Compilation not supported")
) options [const]))
end
@@ -3292,7 +1780,7 @@
Random => PredicateCompFuns.compfuns
| DSeq => DSequence_CompFuns.compfuns
| Pos_Random_DSeq => Random_Sequence_CompFuns.compfuns
- | New_Pos_Random_DSeq => New_Pos_Random_Sequence_CompFuns.compfuns
+ | New_Pos_Random_DSeq => New_Pos_Random_Sequence_CompFuns.depth_unlimited_compfuns
| _ => PredicateCompFuns.compfuns
val t = analyze_compr ctxt compfuns param_user_modes options t_compr;
val T = dest_predT compfuns (fastype_of t);
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Sun Oct 24 15:11:24 2010 -0700
@@ -255,7 +255,7 @@
let
val ctxt = ProofContext.init_global thy
fun is_nondefining_const (c, T) = member (op =) logic_operator_names c
- fun has_code_pred_intros (c, T) = can (Predicate_Compile_Core.intros_of ctxt) c
+ fun has_code_pred_intros (c, T) = can (Core_Data.intros_of ctxt) c
fun case_consts (c, T) = is_some (Datatype.info_of_case thy c)
fun is_datatype_constructor (c, T) = is_some (Datatype.info_of_constr thy (c, T))
fun defiants_of specs =
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML Sun Oct 24 15:11:24 2010 -0700
@@ -312,7 +312,7 @@
fun functional_mode_of T =
list_fun_mode (replicate (length (binder_types T)) Input @ [Output])
val thy''' = fold
- (fn (predname, Const (name, T)) => Predicate_Compile_Core.register_alternative_function
+ (fn (predname, Const (name, T)) => Core_Data.register_alternative_function
predname (functional_mode_of T) name)
(dst_prednames ~~ dst_funs) thy''
in
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_proof.ML Sun Oct 24 15:11:24 2010 -0700
@@ -0,0 +1,525 @@
+(* Title: HOL/Tools/Predicate_Compile/predicate_compile_proof.ML
+ Author: Lukas Bulwahn, TU Muenchen
+
+Proof procedure for the compiler from predicates specified by intro/elim rules to equations.
+*)
+
+signature PREDICATE_COMPILE_PROOF =
+sig
+ type indprem = Predicate_Compile_Aux.indprem;
+ type mode = Predicate_Compile_Aux.mode
+ val prove_pred : Predicate_Compile_Aux.options -> theory
+ -> (string * (term list * indprem list) list) list
+ -> (string * typ) list -> string -> bool * mode
+ -> (term list * (indprem * Mode_Inference.mode_derivation) list) list * term
+ -> Thm.thm
+end;
+
+structure Predicate_Compile_Proof : PREDICATE_COMPILE_PROOF =
+struct
+
+open Predicate_Compile_Aux;
+open Core_Data;
+open Mode_Inference;
+
+(* debug stuff *)
+
+fun print_tac options s =
+ if show_proof_trace options then Tactical.print_tac s else Seq.single;
+
+(** auxiliary **)
+
+fun assert b = if not b then raise Fail "Assertion failed" else warning "Assertion holds"
+
+datatype assertion = Max_number_of_subgoals of int
+fun assert_tac (Max_number_of_subgoals i) st =
+ if (nprems_of st <= i) then Seq.single st
+ else raise Fail ("assert_tac: Numbers of subgoals mismatch at goal state :"
+ ^ "\n" ^ Pretty.string_of (Pretty.chunks
+ (Goal_Display.pretty_goals_without_context st)));
+
+
+(** special setup for simpset **)
+val HOL_basic_ss' = HOL_basic_ss addsimps (@{thms HOL.simp_thms} @ [@{thm Pair_eq}])
+ setSolver (mk_solver "all_tac_solver" (fn _ => fn _ => all_tac))
+ setSolver (mk_solver "True_solver" (fn _ => rtac @{thm TrueI}))
+
+(* auxillary functions *)
+
+fun is_Type (Type _) = true
+ | is_Type _ = false
+
+(* returns true if t is an application of an datatype constructor *)
+(* which then consequently would be splitted *)
+(* else false *)
+fun is_constructor thy t =
+ if (is_Type (fastype_of t)) then
+ (case Datatype.get_info thy ((fst o dest_Type o fastype_of) t) of
+ NONE => false
+ | SOME info => (let
+ val constr_consts = maps (fn (_, (_, _, constrs)) => map fst constrs) (#descr info)
+ val (c, _) = strip_comb t
+ in (case c of
+ Const (name, _) => member (op =) constr_consts name
+ | _ => false) end))
+ else false
+
+(* MAJOR FIXME: prove_params should be simple
+ - different form of introrule for parameters ? *)
+
+fun prove_param options ctxt nargs t deriv =
+ let
+ val (f, args) = strip_comb (Envir.eta_contract t)
+ val mode = head_mode_of deriv
+ val param_derivations = param_derivations_of deriv
+ val ho_args = ho_args_of mode args
+ val f_tac = case f of
+ Const (name, T) => simp_tac (HOL_basic_ss addsimps
+ [@{thm eval_pred}, predfun_definition_of ctxt name mode,
+ @{thm split_eta}, @{thm split_beta}, @{thm fst_conv},
+ @{thm snd_conv}, @{thm pair_collapse}, @{thm Product_Type.split_conv}]) 1
+ | Free _ =>
+ Subgoal.FOCUS_PREMS (fn {context = ctxt, params = params, prems, asms, concl, schematics} =>
+ let
+ val prems' = maps dest_conjunct_prem (take nargs prems)
+ in
+ MetaSimplifier.rewrite_goal_tac
+ (map (fn th => th RS @{thm sym} RS @{thm eq_reflection}) prems') 1
+ end) ctxt 1
+ | Abs _ => raise Fail "prove_param: No valid parameter term"
+ in
+ REPEAT_DETERM (rtac @{thm ext} 1)
+ THEN print_tac options "prove_param"
+ THEN f_tac
+ THEN print_tac options "after prove_param"
+ THEN (REPEAT_DETERM (atac 1))
+ THEN (EVERY (map2 (prove_param options ctxt nargs) ho_args param_derivations))
+ THEN REPEAT_DETERM (rtac @{thm refl} 1)
+ end
+
+fun prove_expr options ctxt nargs (premposition : int) (t, deriv) =
+ case strip_comb t of
+ (Const (name, T), args) =>
+ let
+ val mode = head_mode_of deriv
+ val introrule = predfun_intro_of ctxt name mode
+ val param_derivations = param_derivations_of deriv
+ val ho_args = ho_args_of mode args
+ in
+ print_tac options "before intro rule:"
+ THEN rtac introrule 1
+ THEN print_tac options "after intro rule"
+ (* for the right assumption in first position *)
+ THEN rotate_tac premposition 1
+ THEN atac 1
+ THEN print_tac options "parameter goal"
+ (* work with parameter arguments *)
+ THEN (EVERY (map2 (prove_param options ctxt nargs) ho_args param_derivations))
+ THEN (REPEAT_DETERM (atac 1))
+ end
+ | (Free _, _) =>
+ print_tac options "proving parameter call.."
+ THEN Subgoal.FOCUS_PREMS (fn {context = ctxt, params, prems, asms, concl, schematics} =>
+ let
+ val param_prem = nth prems premposition
+ val (param, _) = strip_comb (HOLogic.dest_Trueprop (prop_of param_prem))
+ val prems' = maps dest_conjunct_prem (take nargs prems)
+ fun param_rewrite prem =
+ param = snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of prem)))
+ val SOME rew_eq = find_first param_rewrite prems'
+ val param_prem' = MetaSimplifier.rewrite_rule
+ (map (fn th => th RS @{thm eq_reflection})
+ [rew_eq RS @{thm sym}, @{thm split_beta}, @{thm fst_conv}, @{thm snd_conv}])
+ param_prem
+ in
+ rtac param_prem' 1
+ end) ctxt 1
+ THEN print_tac options "after prove parameter call"
+
+fun SOLVED tac st = FILTER (fn st' => nprems_of st' = nprems_of st - 1) tac st;
+
+fun SOLVEDALL tac st = FILTER (fn st' => nprems_of st' = 0) tac st
+
+fun check_format ctxt st =
+ let
+ val concl' = Logic.strip_assums_concl (hd (prems_of st))
+ val concl = HOLogic.dest_Trueprop concl'
+ val expr = fst (strip_comb (fst (PredicateCompFuns.dest_Eval concl)))
+ fun valid_expr (Const (@{const_name Predicate.bind}, _)) = true
+ | valid_expr (Const (@{const_name Predicate.single}, _)) = true
+ | valid_expr _ = false
+ in
+ if valid_expr expr then
+ ((*tracing "expression is valid";*) Seq.single st)
+ else
+ ((*tracing "expression is not valid";*) Seq.empty) (* error "check_format: wrong format" *)
+ end
+
+fun prove_match options ctxt nargs out_ts =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val eval_if_P =
+ @{lemma "P ==> Predicate.eval x z ==> Predicate.eval (if P then x else y) z" by simp}
+ fun get_case_rewrite t =
+ if (is_constructor thy t) then let
+ val case_rewrites = (#case_rewrites (Datatype.the_info thy
+ ((fst o dest_Type o fastype_of) t)))
+ in fold (union Thm.eq_thm) (case_rewrites :: map get_case_rewrite (snd (strip_comb t))) [] end
+ else []
+ val simprules = insert Thm.eq_thm @{thm "unit.cases"} (insert Thm.eq_thm @{thm "prod.cases"}
+ (fold (union Thm.eq_thm) (map get_case_rewrite out_ts) []))
+ (* replace TRY by determining if it necessary - are there equations when calling compile match? *)
+ in
+ (* make this simpset better! *)
+ asm_full_simp_tac (HOL_basic_ss' addsimps simprules) 1
+ THEN print_tac options "after prove_match:"
+ THEN (DETERM (TRY
+ (rtac eval_if_P 1
+ THEN (SUBPROOF (fn {context = ctxt, params, prems, asms, concl, schematics} =>
+ (REPEAT_DETERM (rtac @{thm conjI} 1
+ THEN (SOLVED (asm_simp_tac HOL_basic_ss' 1))))
+ THEN print_tac options "if condition to be solved:"
+ THEN asm_simp_tac HOL_basic_ss' 1
+ THEN TRY (
+ let
+ val prems' = maps dest_conjunct_prem (take nargs prems)
+ in
+ MetaSimplifier.rewrite_goal_tac
+ (map (fn th => th RS @{thm sym} RS @{thm eq_reflection}) prems') 1
+ end
+ THEN REPEAT_DETERM (rtac @{thm refl} 1))
+ THEN print_tac options "after if simp; in SUBPROOF") ctxt 1))))
+ THEN print_tac options "after if simplification"
+ end;
+
+(* corresponds to compile_fun -- maybe call that also compile_sidecond? *)
+
+fun prove_sidecond ctxt t =
+ let
+ fun preds_of t nameTs = case strip_comb t of
+ (f as Const (name, T), args) =>
+ if is_registered ctxt name then (name, T) :: nameTs
+ else fold preds_of args nameTs
+ | _ => nameTs
+ val preds = preds_of t []
+ val defs = map
+ (fn (pred, T) => predfun_definition_of ctxt pred
+ (all_input_of T))
+ preds
+ in
+ simp_tac (HOL_basic_ss addsimps
+ (@{thms HOL.simp_thms} @ (@{thm eval_pred} :: defs))) 1
+ (* need better control here! *)
+ end
+
+fun prove_clause options ctxt nargs mode (_, clauses) (ts, moded_ps) =
+ let
+ val (in_ts, clause_out_ts) = split_mode mode ts;
+ fun prove_prems out_ts [] =
+ (prove_match options ctxt nargs out_ts)
+ THEN print_tac options "before simplifying assumptions"
+ THEN asm_full_simp_tac HOL_basic_ss' 1
+ THEN print_tac options "before single intro rule"
+ THEN Subgoal.FOCUS_PREMS
+ (fn {context = ctxt, params = params, prems, asms, concl, schematics} =>
+ let
+ val prems' = maps dest_conjunct_prem (take nargs prems)
+ in
+ MetaSimplifier.rewrite_goal_tac
+ (map (fn th => th RS @{thm sym} RS @{thm eq_reflection}) prems') 1
+ end) ctxt 1
+ THEN (rtac (if null clause_out_ts then @{thm singleI_unit} else @{thm singleI}) 1)
+ | prove_prems out_ts ((p, deriv) :: ps) =
+ let
+ val premposition = (find_index (equal p) clauses) + nargs
+ val mode = head_mode_of deriv
+ val rest_tac =
+ rtac @{thm bindI} 1
+ THEN (case p of Prem t =>
+ let
+ val (_, us) = strip_comb t
+ val (_, out_ts''') = split_mode mode us
+ val rec_tac = prove_prems out_ts''' ps
+ in
+ print_tac options "before clause:"
+ (*THEN asm_simp_tac HOL_basic_ss 1*)
+ THEN print_tac options "before prove_expr:"
+ THEN prove_expr options ctxt nargs premposition (t, deriv)
+ THEN print_tac options "after prove_expr:"
+ THEN rec_tac
+ end
+ | Negprem t =>
+ let
+ val (t, args) = strip_comb t
+ val (_, out_ts''') = split_mode mode args
+ val rec_tac = prove_prems out_ts''' ps
+ val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
+ val neg_intro_rule =
+ Option.map (fn name =>
+ the (predfun_neg_intro_of ctxt name mode)) name
+ val param_derivations = param_derivations_of deriv
+ val params = ho_args_of mode args
+ in
+ print_tac options "before prove_neg_expr:"
+ THEN full_simp_tac (HOL_basic_ss addsimps
+ [@{thm split_eta}, @{thm split_beta}, @{thm fst_conv},
+ @{thm snd_conv}, @{thm pair_collapse}, @{thm Product_Type.split_conv}]) 1
+ THEN (if (is_some name) then
+ print_tac options "before applying not introduction rule"
+ THEN Subgoal.FOCUS_PREMS
+ (fn {context = ctxt, params = params, prems, asms, concl, schematics} =>
+ rtac (the neg_intro_rule) 1
+ THEN rtac (nth prems premposition) 1) ctxt 1
+ THEN print_tac options "after applying not introduction rule"
+ THEN (EVERY (map2 (prove_param options ctxt nargs) params param_derivations))
+ THEN (REPEAT_DETERM (atac 1))
+ else
+ rtac @{thm not_predI'} 1
+ (* test: *)
+ THEN dtac @{thm sym} 1
+ THEN asm_full_simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1)
+ THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
+ THEN rec_tac
+ end
+ | Sidecond t =>
+ rtac @{thm if_predI} 1
+ THEN print_tac options "before sidecond:"
+ THEN prove_sidecond ctxt t
+ THEN print_tac options "after sidecond:"
+ THEN prove_prems [] ps)
+ in (prove_match options ctxt nargs out_ts)
+ THEN rest_tac
+ end;
+ val prems_tac = prove_prems in_ts moded_ps
+ in
+ print_tac options "Proving clause..."
+ THEN rtac @{thm bindI} 1
+ THEN rtac @{thm singleI} 1
+ THEN prems_tac
+ end;
+
+fun select_sup 1 1 = []
+ | select_sup _ 1 = [rtac @{thm supI1}]
+ | select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1));
+
+fun prove_one_direction options ctxt clauses preds pred mode moded_clauses =
+ let
+ val T = the (AList.lookup (op =) preds pred)
+ val nargs = length (binder_types T)
+ val pred_case_rule = the_elim_of ctxt pred
+ in
+ REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"}))
+ THEN print_tac options "before applying elim rule"
+ THEN etac (predfun_elim_of ctxt pred mode) 1
+ THEN etac pred_case_rule 1
+ THEN print_tac options "after applying elim rule"
+ THEN (EVERY (map
+ (fn i => EVERY' (select_sup (length moded_clauses) i) i)
+ (1 upto (length moded_clauses))))
+ THEN (EVERY (map2 (prove_clause options ctxt nargs mode) clauses moded_clauses))
+ THEN print_tac options "proved one direction"
+ end;
+
+(** Proof in the other direction **)
+
+fun prove_match2 options ctxt out_ts =
+ let
+ val thy = ProofContext.theory_of ctxt
+ fun split_term_tac (Free _) = all_tac
+ | split_term_tac t =
+ if (is_constructor thy t) then
+ let
+ val info = Datatype.the_info thy ((fst o dest_Type o fastype_of) t)
+ val num_of_constrs = length (#case_rewrites info)
+ val (_, ts) = strip_comb t
+ in
+ print_tac options ("Term " ^ (Syntax.string_of_term ctxt t) ^
+ "splitting with rules \n" ^ Display.string_of_thm ctxt (#split_asm info))
+ THEN TRY ((Splitter.split_asm_tac [#split_asm info] 1)
+ THEN (print_tac options "after splitting with split_asm rules")
+ (* THEN (Simplifier.asm_full_simp_tac HOL_basic_ss 1)
+ THEN (DETERM (TRY (etac @{thm Pair_inject} 1)))*)
+ THEN (REPEAT_DETERM_N (num_of_constrs - 1)
+ (etac @{thm botE} 1 ORELSE etac @{thm botE} 2)))
+ THEN (assert_tac (Max_number_of_subgoals 2))
+ THEN (EVERY (map split_term_tac ts))
+ end
+ else all_tac
+ in
+ split_term_tac (HOLogic.mk_tuple out_ts)
+ THEN (DETERM (TRY ((Splitter.split_asm_tac [@{thm "split_if_asm"}] 1)
+ THEN (etac @{thm botE} 2))))
+ end
+
+(* VERY LARGE SIMILIRATIY to function prove_param
+-- join both functions
+*)
+(* TODO: remove function *)
+
+fun prove_param2 options ctxt t deriv =
+ let
+ val (f, args) = strip_comb (Envir.eta_contract t)
+ val mode = head_mode_of deriv
+ val param_derivations = param_derivations_of deriv
+ val ho_args = ho_args_of mode args
+ val f_tac = case f of
+ Const (name, T) => full_simp_tac (HOL_basic_ss addsimps
+ (@{thm eval_pred}::(predfun_definition_of ctxt name mode)
+ :: @{thm "Product_Type.split_conv"}::[])) 1
+ | Free _ => all_tac
+ | _ => error "prove_param2: illegal parameter term"
+ in
+ print_tac options "before simplification in prove_args:"
+ THEN f_tac
+ THEN print_tac options "after simplification in prove_args"
+ THEN EVERY (map2 (prove_param2 options ctxt) ho_args param_derivations)
+ end
+
+fun prove_expr2 options ctxt (t, deriv) =
+ (case strip_comb t of
+ (Const (name, T), args) =>
+ let
+ val mode = head_mode_of deriv
+ val param_derivations = param_derivations_of deriv
+ val ho_args = ho_args_of mode args
+ in
+ etac @{thm bindE} 1
+ THEN (REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"})))
+ THEN print_tac options "prove_expr2-before"
+ THEN etac (predfun_elim_of ctxt name mode) 1
+ THEN print_tac options "prove_expr2"
+ THEN (EVERY (map2 (prove_param2 options ctxt) ho_args param_derivations))
+ THEN print_tac options "finished prove_expr2"
+ end
+ | _ => etac @{thm bindE} 1)
+
+fun prove_sidecond2 options ctxt t = let
+ fun preds_of t nameTs = case strip_comb t of
+ (f as Const (name, T), args) =>
+ if is_registered ctxt name then (name, T) :: nameTs
+ else fold preds_of args nameTs
+ | _ => nameTs
+ val preds = preds_of t []
+ val defs = map
+ (fn (pred, T) => predfun_definition_of ctxt pred
+ (all_input_of T))
+ preds
+ in
+ (* only simplify the one assumption *)
+ full_simp_tac (HOL_basic_ss' addsimps @{thm eval_pred} :: defs) 1
+ (* need better control here! *)
+ THEN print_tac options "after sidecond2 simplification"
+ end
+
+fun prove_clause2 options ctxt pred mode (ts, ps) i =
+ let
+ val pred_intro_rule = nth (intros_of ctxt pred) (i - 1)
+ val (in_ts, clause_out_ts) = split_mode mode ts;
+ val split_ss = HOL_basic_ss' addsimps [@{thm split_eta}, @{thm split_beta},
+ @{thm fst_conv}, @{thm snd_conv}, @{thm pair_collapse}]
+ fun prove_prems2 out_ts [] =
+ print_tac options "before prove_match2 - last call:"
+ THEN prove_match2 options ctxt out_ts
+ THEN print_tac options "after prove_match2 - last call:"
+ THEN (etac @{thm singleE} 1)
+ THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
+ THEN (asm_full_simp_tac HOL_basic_ss' 1)
+ THEN TRY (
+ (REPEAT_DETERM (etac @{thm Pair_inject} 1))
+ THEN (asm_full_simp_tac HOL_basic_ss' 1)
+
+ THEN SOLVED (print_tac options "state before applying intro rule:"
+ THEN (rtac pred_intro_rule
+ (* How to handle equality correctly? *)
+ THEN_ALL_NEW (K (print_tac options "state before assumption matching")
+ THEN' (atac ORELSE' ((CHANGED o asm_full_simp_tac split_ss) THEN' (TRY o atac)))
+ THEN' (K (print_tac options "state after pre-simplification:"))
+ THEN' (K (print_tac options "state after assumption matching:")))) 1))
+ | prove_prems2 out_ts ((p, deriv) :: ps) =
+ let
+ val mode = head_mode_of deriv
+ val rest_tac = (case p of
+ Prem t =>
+ let
+ val (_, us) = strip_comb t
+ val (_, out_ts''') = split_mode mode us
+ val rec_tac = prove_prems2 out_ts''' ps
+ in
+ (prove_expr2 options ctxt (t, deriv)) THEN rec_tac
+ end
+ | Negprem t =>
+ let
+ val (_, args) = strip_comb t
+ val (_, out_ts''') = split_mode mode args
+ val rec_tac = prove_prems2 out_ts''' ps
+ val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
+ val param_derivations = param_derivations_of deriv
+ val ho_args = ho_args_of mode args
+ in
+ print_tac options "before neg prem 2"
+ THEN etac @{thm bindE} 1
+ THEN (if is_some name then
+ full_simp_tac (HOL_basic_ss addsimps
+ [predfun_definition_of ctxt (the name) mode]) 1
+ THEN etac @{thm not_predE} 1
+ THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
+ THEN (EVERY (map2 (prove_param2 options ctxt) ho_args param_derivations))
+ else
+ etac @{thm not_predE'} 1)
+ THEN rec_tac
+ end
+ | Sidecond t =>
+ etac @{thm bindE} 1
+ THEN etac @{thm if_predE} 1
+ THEN prove_sidecond2 options ctxt t
+ THEN prove_prems2 [] ps)
+ in print_tac options "before prove_match2:"
+ THEN prove_match2 options ctxt out_ts
+ THEN print_tac options "after prove_match2:"
+ THEN rest_tac
+ end;
+ val prems_tac = prove_prems2 in_ts ps
+ in
+ print_tac options "starting prove_clause2"
+ THEN etac @{thm bindE} 1
+ THEN (etac @{thm singleE'} 1)
+ THEN (TRY (etac @{thm Pair_inject} 1))
+ THEN print_tac options "after singleE':"
+ THEN prems_tac
+ end;
+
+fun prove_other_direction options ctxt pred mode moded_clauses =
+ let
+ fun prove_clause clause i =
+ (if i < length moded_clauses then etac @{thm supE} 1 else all_tac)
+ THEN (prove_clause2 options ctxt pred mode clause i)
+ in
+ (DETERM (TRY (rtac @{thm unit.induct} 1)))
+ THEN (REPEAT_DETERM (CHANGED (rewtac @{thm split_paired_all})))
+ THEN (rtac (predfun_intro_of ctxt pred mode) 1)
+ THEN (REPEAT_DETERM (rtac @{thm refl} 2))
+ THEN (if null moded_clauses then
+ etac @{thm botE} 1
+ else EVERY (map2 prove_clause moded_clauses (1 upto (length moded_clauses))))
+ end;
+
+(** proof procedure **)
+
+fun prove_pred options thy clauses preds pred (pol, mode) (moded_clauses, compiled_term) =
+ let
+ val ctxt = ProofContext.init_global thy
+ val clauses = case AList.lookup (op =) clauses pred of SOME rs => rs | NONE => []
+ in
+ Goal.prove ctxt (Term.add_free_names compiled_term []) [] compiled_term
+ (if not (skip_proof options) then
+ (fn _ =>
+ rtac @{thm pred_iffI} 1
+ THEN print_tac options "after pred_iffI"
+ THEN prove_one_direction options ctxt clauses preds pred mode moded_clauses
+ THEN print_tac options "proved one direction"
+ THEN prove_other_direction options ctxt pred mode moded_clauses
+ THEN print_tac options "proved other direction")
+ else (fn _ => Skip_Proof.cheat_tac thy))
+ end;
+
+end;
\ No newline at end of file
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML Sun Oct 24 15:11:24 2010 -0700
@@ -85,6 +85,7 @@
function_flattening = true,
fail_safe_function_flattening = false,
no_higher_order_predicate = [],
+ smart_depth_limiting = true,
no_topmost_reordering = false
}
@@ -108,6 +109,7 @@
function_flattening = true,
fail_safe_function_flattening = false,
no_higher_order_predicate = [],
+ smart_depth_limiting = true,
no_topmost_reordering = true
}
@@ -119,14 +121,14 @@
show_invalid_clauses = s_ic, skip_proof = s_p,
compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = f_f,
fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho,
- no_topmost_reordering = re}) =
+ smart_depth_limiting = sm_dl, no_topmost_reordering = re}) =
(Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf,
show_invalid_clauses = s_ic, skip_proof = s_p,
compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = b,
fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho,
- no_topmost_reordering = re})
+ smart_depth_limiting = sm_dl, no_topmost_reordering = re})
fun set_fail_safe_function_flattening b
(Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
@@ -135,14 +137,14 @@
show_invalid_clauses = s_ic, skip_proof = s_p,
compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = f_f,
fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho,
- no_topmost_reordering = re}) =
+ smart_depth_limiting = sm_dl, no_topmost_reordering = re}) =
(Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf,
show_invalid_clauses = s_ic, skip_proof = s_p,
compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = f_f,
fail_safe_function_flattening = b, no_higher_order_predicate = no_ho,
- no_topmost_reordering = re})
+ smart_depth_limiting = sm_dl, no_topmost_reordering = re})
fun set_no_higher_order_predicate ss
(Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
@@ -151,13 +153,14 @@
show_invalid_clauses = s_ic, skip_proof = s_p,
compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = f_f,
fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho,
- no_topmost_reordering = re}) =
+ smart_depth_limiting = sm_dl, no_topmost_reordering = re}) =
(Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf,
show_invalid_clauses = s_ic, skip_proof = s_p,
compilation = c, inductify = i, specialise = sp, detect_switches = ds, function_flattening = f_f,
- fail_safe_function_flattening = fs_ff, no_higher_order_predicate = ss, no_topmost_reordering = re})
+ fail_safe_function_flattening = fs_ff, no_higher_order_predicate = ss,
+ smart_depth_limiting = sm_dl, no_topmost_reordering = re})
fun get_options () =
@@ -173,9 +176,12 @@
val mk_return = Predicate_Compile_Aux.mk_single RandomPredCompFuns.compfuns
val mk_bind = Predicate_Compile_Aux.mk_bind RandomPredCompFuns.compfuns
-val mk_new_randompredT = Predicate_Compile_Aux.mk_predT New_Pos_Random_Sequence_CompFuns.compfuns
-val mk_new_return = Predicate_Compile_Aux.mk_single New_Pos_Random_Sequence_CompFuns.compfuns
-val mk_new_bind = Predicate_Compile_Aux.mk_bind New_Pos_Random_Sequence_CompFuns.compfuns
+val mk_new_randompredT =
+ Predicate_Compile_Aux.mk_predT New_Pos_Random_Sequence_CompFuns.depth_limited_compfuns
+val mk_new_return =
+ Predicate_Compile_Aux.mk_single New_Pos_Random_Sequence_CompFuns.depth_limited_compfuns
+val mk_new_bind =
+ Predicate_Compile_Aux.mk_bind New_Pos_Random_Sequence_CompFuns.depth_limited_compfuns
val mk_split_lambda = HOLogic.tupled_lambda o HOLogic.mk_tuple
@@ -207,13 +213,12 @@
val _ = Output.tracing ("Preprocessing time: " ^ string_of_int (snd preproc_time))
val _ = Output.tracing ("Core compilation time: " ^ string_of_int (snd core_comp_time))
val ctxt4 = ProofContext.init_global thy4
- val modes = Predicate_Compile_Core.modes_of compilation ctxt4 full_constname
+ val modes = Core_Data.modes_of compilation ctxt4 full_constname
val output_mode = fold_rev (curry Fun) (map (K Output) (binder_types constT)) Bool
val prog =
if member eq_mode modes output_mode then
let
- val name = Predicate_Compile_Core.function_name_of compilation ctxt4
- full_constname output_mode
+ val name = Core_Data.function_name_of compilation ctxt4 full_constname output_mode
val T =
case compilation of
Pos_Random_DSeq => mk_randompredT (HOLogic.mk_tupleT (map snd vs'))
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_specialisation.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_specialisation.ML Sun Oct 24 15:11:24 2010 -0700
@@ -109,7 +109,7 @@
val optimised_intros =
map_filter (Predicate_Compile_Aux.peephole_optimisation thy'') exported_intros
val ([spec], thy''') = find_specialisations black_list [(constname, optimised_intros)] thy''
- val thy'''' = Predicate_Compile_Core.register_intros spec thy'''
+ val thy'''' = Core_Data.register_intros spec thy'''
in
thy''''
end
@@ -180,7 +180,7 @@
if member (op =) ((map fst specs) @ black_list) pred_name then
thy
else
- (case try (Predicate_Compile_Core.intros_of (ProofContext.init_global thy)) pred_name of
+ (case try (Core_Data.intros_of (ProofContext.init_global thy)) pred_name of
NONE => thy
| SOME [] => thy
| SOME intros =>
--- a/src/HOL/Tools/Sledgehammer/sledgehammer.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer.ML Sun Oct 24 15:11:24 2010 -0700
@@ -10,16 +10,17 @@
sig
type failure = ATP_Systems.failure
type locality = Sledgehammer_Filter.locality
+ type relevance_fudge = Sledgehammer_Filter.relevance_fudge
type relevance_override = Sledgehammer_Filter.relevance_override
- type fol_formula = Sledgehammer_Translate.fol_formula
- type minimize_command = Sledgehammer_Reconstruct.minimize_command
+ type prepared_formula = Sledgehammer_ATP_Translate.prepared_formula
+ type minimize_command = Sledgehammer_ATP_Reconstruct.minimize_command
type params =
{blocking: bool,
debug: bool,
verbose: bool,
overlord: bool,
- atps: string list,
+ provers: string list,
full_types: bool,
explicit_apply: bool,
relevance_thresholds: real * real,
@@ -29,33 +30,40 @@
timeout: Time.time,
expect: string}
- type problem =
+ datatype axiom =
+ Unprepared of (string * locality) * thm |
+ Prepared of term * ((string * locality) * prepared_formula) option
+
+ type prover_problem =
{state: Proof.state,
goal: thm,
subgoal: int,
- axioms: (term * ((string * locality) * fol_formula) option) list,
+ subgoal_count: int,
+ axioms: axiom list,
only: bool}
type prover_result =
{outcome: failure option,
- message: string,
- pool: string Symtab.table,
- used_thm_names: (string * locality) list,
- atp_run_time_in_msecs: int,
- output: string,
- tstplike_proof: string,
- axiom_names: (string * locality) list vector,
- conjecture_shape: int list list}
+ used_axioms: (string * locality) list,
+ run_time_in_msecs: int option,
+ message: string}
+
+ type prover = params -> minimize_command -> prover_problem -> prover_result
- type prover = params -> minimize_command -> problem -> prover_result
-
+ val smtN : string
+ val is_prover_available : theory -> string -> bool
+ val is_prover_installed : Proof.context -> string -> bool
+ val default_max_relevant_for_prover : theory -> string -> int
+ val irrelevant_consts_for_prover : string -> string list
+ val relevance_fudge_for_prover : string -> relevance_fudge
val dest_dir : string Config.T
val problem_prefix : string Config.T
val measure_run_time : bool Config.T
- val kill_atps: unit -> unit
- val running_atps: unit -> unit
- val messages: int option -> unit
- val get_prover_fun : theory -> string -> prover
+ val available_provers : theory -> unit
+ val kill_provers : unit -> unit
+ val running_provers : unit -> unit
+ val messages : int option -> unit
+ val get_prover : theory -> bool -> string -> prover
val run_sledgehammer :
params -> bool -> int -> relevance_override -> (string -> minimize_command)
-> Proof.state -> bool * Proof.state
@@ -71,8 +79,8 @@
open Metis_Translate
open Sledgehammer_Util
open Sledgehammer_Filter
-open Sledgehammer_Translate
-open Sledgehammer_Reconstruct
+open Sledgehammer_ATP_Translate
+open Sledgehammer_ATP_Reconstruct
(** The Sledgehammer **)
@@ -81,18 +89,102 @@
"Async_Manager". *)
val das_Tool = "Sledgehammer"
-fun kill_atps () = Async_Manager.kill_threads das_Tool "ATPs"
-fun running_atps () = Async_Manager.running_threads das_Tool "ATPs"
-val messages = Async_Manager.thread_messages das_Tool "ATP"
+val smtN = "smt"
+val smt_prover_names = [smtN, remote_prefix ^ smtN]
+
+val is_smt_prover = member (op =) smt_prover_names
+
+fun is_prover_available thy name =
+ is_smt_prover name orelse member (op =) (available_atps thy) name
+
+fun is_prover_installed ctxt name =
+ let val thy = ProofContext.theory_of ctxt in
+ if is_smt_prover name then true (* FIXME *)
+ else is_atp_installed thy name
+ end
+
+val smt_default_max_relevant = 300 (* FUDGE (FIXME) *)
+val auto_max_relevant_divisor = 2 (* FUDGE *)
+
+fun default_max_relevant_for_prover thy name =
+ if is_smt_prover name then smt_default_max_relevant
+ else #default_max_relevant (get_atp thy name)
+
+(* These are typically simplified away by "Meson.presimplify". Equality is
+ handled specially via "fequal". *)
+val atp_irrelevant_consts =
+ [@{const_name False}, @{const_name True}, @{const_name If}, @{const_name Let},
+ @{const_name HOL.eq}]
+val smt_irrelevant_consts = atp_irrelevant_consts (* FIXME *)
+
+fun irrelevant_consts_for_prover name =
+ if is_smt_prover name then smt_irrelevant_consts else atp_irrelevant_consts
-(** problems, results, provers, etc. **)
+(* FUDGE *)
+val atp_relevance_fudge =
+ {worse_irrel_freq = 100.0,
+ higher_order_irrel_weight = 1.05,
+ abs_rel_weight = 0.5,
+ abs_irrel_weight = 2.0,
+ skolem_irrel_weight = 0.75,
+ theory_const_rel_weight = 0.5,
+ theory_const_irrel_weight = 0.25,
+ intro_bonus = 0.15,
+ elim_bonus = 0.15,
+ simp_bonus = 0.15,
+ local_bonus = 0.55,
+ assum_bonus = 1.05,
+ chained_bonus = 1.5,
+ max_imperfect = 11.5,
+ max_imperfect_exp = 1.0,
+ threshold_divisor = 2.0,
+ ridiculous_threshold = 0.1}
+
+(* FUDGE (FIXME) *)
+val smt_relevance_fudge =
+ {worse_irrel_freq = #worse_irrel_freq atp_relevance_fudge,
+ higher_order_irrel_weight = #higher_order_irrel_weight atp_relevance_fudge,
+ abs_rel_weight = #abs_rel_weight atp_relevance_fudge,
+ abs_irrel_weight = #abs_irrel_weight atp_relevance_fudge,
+ skolem_irrel_weight = #skolem_irrel_weight atp_relevance_fudge,
+ theory_const_rel_weight = #theory_const_rel_weight atp_relevance_fudge,
+ theory_const_irrel_weight = #theory_const_irrel_weight atp_relevance_fudge,
+ intro_bonus = #intro_bonus atp_relevance_fudge,
+ elim_bonus = #elim_bonus atp_relevance_fudge,
+ simp_bonus = #simp_bonus atp_relevance_fudge,
+ local_bonus = #local_bonus atp_relevance_fudge,
+ assum_bonus = #assum_bonus atp_relevance_fudge,
+ chained_bonus = #chained_bonus atp_relevance_fudge,
+ max_imperfect = #max_imperfect atp_relevance_fudge,
+ max_imperfect_exp = #max_imperfect_exp atp_relevance_fudge,
+ threshold_divisor = #threshold_divisor atp_relevance_fudge,
+ ridiculous_threshold = #ridiculous_threshold atp_relevance_fudge}
+
+fun relevance_fudge_for_prover name =
+ if is_smt_prover name then smt_relevance_fudge else atp_relevance_fudge
+
+fun available_provers thy =
+ let
+ val (remote_provers, local_provers) =
+ sort_strings (available_atps thy) @ smt_prover_names
+ |> List.partition (String.isPrefix remote_prefix)
+ in
+ priority ("Available provers: " ^ commas (local_provers @ remote_provers) ^
+ ".")
+ end
+
+fun kill_provers () = Async_Manager.kill_threads das_Tool "provers"
+fun running_provers () = Async_Manager.running_threads das_Tool "provers"
+val messages = Async_Manager.thread_messages das_Tool "prover"
+
+(** problems, results, ATPs, etc. **)
type params =
{blocking: bool,
debug: bool,
verbose: bool,
overlord: bool,
- atps: string list,
+ provers: string list,
full_types: bool,
explicit_apply: bool,
relevance_thresholds: real * real,
@@ -102,25 +194,25 @@
timeout: Time.time,
expect: string}
-type problem =
+datatype axiom =
+ Unprepared of (string * locality) * thm |
+ Prepared of term * ((string * locality) * prepared_formula) option
+
+type prover_problem =
{state: Proof.state,
goal: thm,
subgoal: int,
- axioms: (term * ((string * locality) * fol_formula) option) list,
+ subgoal_count: int,
+ axioms: axiom list,
only: bool}
type prover_result =
{outcome: failure option,
message: string,
- pool: string Symtab.table,
- used_thm_names: (string * locality) list,
- atp_run_time_in_msecs: int,
- output: string,
- tstplike_proof: string,
- axiom_names: (string * locality) list vector,
- conjecture_shape: int list list}
+ used_axioms: (string * locality) list,
+ run_time_in_msecs: int option}
-type prover = params -> minimize_command -> problem -> prover_result
+type prover = params -> minimize_command -> prover_problem -> prover_result
(* configuration attributes *)
@@ -140,41 +232,67 @@
|> Exn.release
|> tap (after path)
-(* generic TPTP-based provers *)
+fun prover_description ctxt ({blocking, verbose, ...} : params) name num_axioms
+ i n goal =
+ quote name ^
+ (if verbose then
+ " with " ^ string_of_int num_axioms ^ " fact" ^ plural_s num_axioms
+ else
+ "") ^
+ " on " ^ (if n = 1 then "goal" else "subgoal " ^ string_of_int i) ^ ":" ^
+ (if blocking then
+ ""
+ else
+ "\n" ^ Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i)))
+
+fun proof_banner auto =
+ if auto then "Sledgehammer found a proof" else "Try this command"
+
+(* generic TPTP-based ATPs *)
+
+fun dest_Unprepared (Unprepared p) = p
+ | dest_Unprepared (Prepared _) = raise Fail "dest_Unprepared"
+fun prepared_axiom ctxt (Unprepared p) = prepare_axiom ctxt p
+ | prepared_axiom _ (Prepared p) = p
+
+fun int_option_add (SOME m) (SOME n) = SOME (m + n)
+ | int_option_add _ _ = NONE
(* Important messages are important but not so important that users want to see
them each time. *)
val important_message_keep_factor = 0.1
-fun prover_fun auto atp_name
+fun run_atp auto atp_name
{exec, required_execs, arguments, has_incomplete_mode, proof_delims,
known_failures, default_max_relevant, explicit_forall,
use_conjecture_for_hypotheses}
({debug, verbose, overlord, full_types, explicit_apply,
max_relevant, isar_proof, isar_shrink_factor, timeout, ...} : params)
- minimize_command ({state, goal, subgoal, axioms, only} : problem) =
+ minimize_command
+ ({state, goal, subgoal, axioms, only, ...} : prover_problem) =
let
val ctxt = Proof.context_of state
val (_, hyp_ts, concl_t) = strip_subgoal goal subgoal
- val axioms = axioms |> not only
- ? take (the_default default_max_relevant max_relevant)
- val the_dest_dir = if overlord then getenv "ISABELLE_HOME_USER"
- else Config.get ctxt dest_dir
- val the_problem_prefix = Config.get ctxt problem_prefix
+ val axioms =
+ axioms |> not only ? take (the_default default_max_relevant max_relevant)
+ |> map (prepared_axiom ctxt)
+ val dest_dir = if overlord then getenv "ISABELLE_HOME_USER"
+ else Config.get ctxt dest_dir
+ val problem_prefix = Config.get ctxt problem_prefix
val problem_file_name =
Path.basic ((if overlord then "prob_" ^ atp_name
- else the_problem_prefix ^ serial_string ())
+ else problem_prefix ^ serial_string ())
^ "_" ^ string_of_int subgoal)
val problem_path_name =
- if the_dest_dir = "" then
+ if dest_dir = "" then
File.tmp_path problem_file_name
- else if File.exists (Path.explode the_dest_dir) then
- Path.append (Path.explode the_dest_dir) problem_file_name
+ else if File.exists (Path.explode dest_dir) then
+ Path.append (Path.explode dest_dir) problem_file_name
else
- error ("No such directory: " ^ quote the_dest_dir ^ ".")
+ error ("No such directory: " ^ quote dest_dir ^ ".")
val measure_run_time = verbose orelse Config.get ctxt measure_run_time
val command = Path.explode (getenv (fst exec) ^ "/" ^ snd exec)
- (* write out problem file and call prover *)
+ (* write out problem file and call ATP *)
fun command_line complete timeout probfile =
let
val core = File.shell_path command ^ " " ^ arguments complete timeout ^
@@ -192,7 +310,7 @@
val digit = Scan.one Symbol.is_ascii_digit;
val num3 = as_num (digit ::: digit ::: (digit >> single));
val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b);
- val as_time = the_default 0 o Scan.read Symbol.stopper time o explode;
+ val as_time = Scan.read Symbol.stopper time o explode
in (output, as_time t) end;
fun run_on probfile =
case filter (curry (op =) "" o getenv o fst) (exec :: required_execs) of
@@ -210,17 +328,17 @@
prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n")
else
I)
- |>> (if measure_run_time then split_time else rpair 0)
+ |>> (if measure_run_time then split_time else rpair NONE)
val (tstplike_proof, outcome) =
extract_tstplike_proof_and_outcome complete res_code
proof_delims known_failures output
in (output, msecs, tstplike_proof, outcome) end
val readable_names = debug andalso overlord
- val (problem, pool, conjecture_offset, axiom_names) =
- prepare_problem ctxt readable_names explicit_forall full_types
- explicit_apply hyp_ts concl_t axioms
+ val (atp_problem, pool, conjecture_offset, axiom_names) =
+ prepare_atp_problem ctxt readable_names explicit_forall full_types
+ explicit_apply hyp_ts concl_t axioms
val ss = tptp_strings_for_atp_problem use_conjecture_for_hypotheses
- problem
+ atp_problem
val _ = File.write_list probfile ss
val conjecture_shape =
conjecture_offset + 1 upto conjecture_offset + length hyp_ts + 1
@@ -237,7 +355,8 @@
? (fn (_, msecs0, _, SOME _) =>
run true (Time.- (timeout, Timer.checkRealTimer timer))
|> (fn (output, msecs, tstplike_proof, outcome) =>
- (output, msecs0 + msecs, tstplike_proof, outcome))
+ (output, int_option_add msecs0 msecs,
+ tstplike_proof, outcome))
| result => result)
in ((pool, conjecture_shape, axiom_names), result) end
else
@@ -246,9 +365,9 @@
(* If the problem file has not been exported, remove it; otherwise, export
the proof file too. *)
fun cleanup probfile =
- if the_dest_dir = "" then try File.rm probfile else NONE
+ if dest_dir = "" then try File.rm probfile else NONE
fun export probfile (_, (output, _, _, _)) =
- if the_dest_dir = "" then
+ if dest_dir = "" then
()
else
File.write (Path.explode (Path.implode probfile ^ "_proof")) output
@@ -263,19 +382,17 @@
extract_important_message output
else
""
- val banner = if auto then "Sledgehammer found a proof"
- else "Try this command"
- val (message, used_thm_names) =
+ val (message, used_axioms) =
case outcome of
NONE =>
proof_text isar_proof
(pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
- (banner, full_types, minimize_command, tstplike_proof, axiom_names,
- goal, subgoal)
+ (proof_banner auto, full_types, minimize_command, tstplike_proof,
+ axiom_names, goal, subgoal)
|>> (fn message =>
message ^ (if verbose then
- "\nATP real CPU time: " ^ string_of_int msecs ^
- " ms."
+ "\nATP real CPU time: " ^
+ string_of_int (the msecs) ^ " ms."
else
"") ^
(if important_message <> "" then
@@ -285,40 +402,58 @@
""))
| SOME failure => (string_for_failure failure, [])
in
- {outcome = outcome, message = message, pool = pool,
- used_thm_names = used_thm_names, atp_run_time_in_msecs = msecs,
- output = output, tstplike_proof = tstplike_proof,
- axiom_names = axiom_names, conjecture_shape = conjecture_shape}
+ {outcome = outcome, message = message, used_axioms = used_axioms,
+ run_time_in_msecs = msecs}
end
-fun get_prover_fun thy name = prover_fun false name (get_prover thy name)
+(* FIXME: dummy *)
+fun saschas_run_smt_solver remote timeout state axioms i =
+ (OS.Process.sleep (Time.fromMilliseconds 1500);
+ {outcome = NONE, used_axioms = axioms |> take 5 |> map fst,
+ run_time_in_msecs = NONE})
-fun run_prover (params as {blocking, debug, verbose, max_relevant, timeout,
- expect, ...})
- auto i n minimize_command (problem as {state, goal, axioms, ...})
- (prover as {default_max_relevant, ...}, atp_name) =
+fun run_smt_solver remote ({timeout, ...} : params) minimize_command
+ ({state, subgoal, subgoal_count, axioms, ...}
+ : prover_problem) =
let
+ val {outcome, used_axioms, run_time_in_msecs} =
+ saschas_run_smt_solver remote timeout state
+ (map_filter (try dest_Unprepared) axioms) subgoal
+ val message =
+ if outcome = NONE then
+ try_command_line (proof_banner false)
+ (apply_on_subgoal subgoal subgoal_count ^
+ command_call smtN (map fst used_axioms)) ^
+ minimize_line minimize_command (map fst used_axioms)
+ else
+ ""
+ in
+ {outcome = outcome, used_axioms = used_axioms,
+ run_time_in_msecs = run_time_in_msecs, message = message}
+ end
+
+fun get_prover thy auto name =
+ if is_smt_prover name then run_smt_solver (String.isPrefix remote_prefix)
+ else run_atp auto name (get_atp thy name)
+
+fun run_prover (params as {blocking, debug, max_relevant, timeout, expect, ...})
+ auto minimize_command
+ (problem as {state, goal, subgoal, subgoal_count, axioms, ...})
+ name =
+ let
+ val thy = Proof.theory_of state
val ctxt = Proof.context_of state
val birth_time = Time.now ()
val death_time = Time.+ (birth_time, timeout)
- val max_relevant = the_default default_max_relevant max_relevant
+ val max_relevant =
+ the_default (default_max_relevant_for_prover thy name) max_relevant
val num_axioms = Int.min (length axioms, max_relevant)
val desc =
- "ATP " ^ quote atp_name ^
- (if verbose then
- " with " ^ string_of_int num_axioms ^ " fact" ^ plural_s num_axioms
- else
- "") ^
- " on " ^ (if n = 1 then "goal" else "subgoal " ^ string_of_int i) ^ ":" ^
- (if blocking then
- ""
- else
- "\n" ^ Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i)))
+ prover_description ctxt params name num_axioms subgoal subgoal_count goal
fun go () =
let
fun really_go () =
- prover_fun auto atp_name prover params (minimize_command atp_name)
- problem
+ get_prover thy auto name params (minimize_command name) problem
|> (fn {outcome, message, ...} =>
(if is_some outcome then "none" else "some", message))
val (outcome_code, message) =
@@ -355,51 +490,68 @@
(false, state))
end
-val auto_max_relevant_divisor = 2
-
-fun run_sledgehammer (params as {blocking, atps, full_types,
+fun run_sledgehammer (params as {blocking, provers, full_types,
relevance_thresholds, max_relevant, ...})
auto i (relevance_override as {only, ...}) minimize_command
state =
- if null atps then
- error "No ATP is set."
+ if null provers then
+ error "No prover is set."
else case subgoal_count state of
0 => (priority "No subgoal!"; (false, state))
| n =>
let
val _ = Proof.assert_backward state
val thy = Proof.theory_of state
- val _ = () |> not blocking ? kill_atps
+ val {context = ctxt, facts = chained_ths, goal} = Proof.goal state
+ val (_, hyp_ts, concl_t) = strip_subgoal goal i
+ val _ = () |> not blocking ? kill_provers
val _ = if auto then () else priority "Sledgehammering..."
- val provers = map (`(get_prover thy)) atps
- fun go () =
- let
- val {context = ctxt, facts = chained_ths, goal} = Proof.goal state
- val (_, hyp_ts, concl_t) = strip_subgoal goal i
- val max_max_relevant =
- case max_relevant of
- SOME n => n
- | NONE =>
- 0 |> fold (Integer.max o #default_max_relevant o fst) provers
- |> auto ? (fn n => n div auto_max_relevant_divisor)
- val axioms =
- relevant_facts ctxt full_types relevance_thresholds
- max_max_relevant relevance_override chained_ths
- hyp_ts concl_t
- val problem =
- {state = state, goal = goal, subgoal = i,
- axioms = map (prepare_axiom ctxt) axioms, only = only}
- val run_prover = run_prover params auto i n minimize_command problem
- in
- if auto then
- fold (fn prover => fn (true, state) => (true, state)
- | (false, _) => run_prover prover)
- provers (false, state)
- else
- (if blocking then Par_List.map else map) run_prover provers
- |> exists fst |> rpair state
- end
- in if blocking then go () else Future.fork (tap go) |> K (false, state) end
+ val (smts, atps) = provers |> List.partition is_smt_prover
+ fun run_provers full_types irrelevant_consts relevance_fudge maybe_prepare
+ provers (res as (success, state)) =
+ if success orelse null provers then
+ res
+ else
+ let
+ val max_max_relevant =
+ case max_relevant of
+ SOME n => n
+ | NONE =>
+ 0 |> fold (Integer.max o default_max_relevant_for_prover thy)
+ provers
+ |> auto ? (fn n => n div auto_max_relevant_divisor)
+ val axioms =
+ relevant_facts ctxt full_types relevance_thresholds
+ max_max_relevant irrelevant_consts relevance_fudge
+ relevance_override chained_ths hyp_ts concl_t
+ |> map maybe_prepare
+ val problem =
+ {state = state, goal = goal, subgoal = i, subgoal_count = n,
+ axioms = axioms, only = only}
+ val run_prover = run_prover params auto minimize_command
+ in
+ if auto then
+ fold (fn prover => fn (true, state) => (true, state)
+ | (false, _) => run_prover problem prover)
+ provers (false, state)
+ else
+ provers |> (if blocking then Par_List.map else map)
+ (run_prover problem)
+ |> exists fst |> rpair state
+ end
+ val run_atps =
+ run_provers full_types atp_irrelevant_consts atp_relevance_fudge
+ (Prepared o prepare_axiom ctxt) atps
+ val run_smts =
+ run_provers true smt_irrelevant_consts smt_relevance_fudge Unprepared
+ smts
+ fun run_atps_and_smt_solvers () =
+ [run_atps, run_smts] |> Par_List.map (fn f => f (false, state) |> K ())
+ in
+ (false, state)
+ |> (if blocking then run_atps #> not auto ? run_smts
+ else (fn p => Future.fork (tap run_atps_and_smt_solvers) |> K p))
+ end
val setup =
dest_dir_setup
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_atp_reconstruct.ML Sun Oct 24 15:11:24 2010 -0700
@@ -0,0 +1,949 @@
+(* Title: HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML
+ Author: Lawrence C. Paulson, Cambridge University Computer Laboratory
+ Author: Claire Quigley, Cambridge University Computer Laboratory
+ Author: Jasmin Blanchette, TU Muenchen
+
+Proof reconstruction for Sledgehammer.
+*)
+
+signature SLEDGEHAMMER_ATP_RECONSTRUCT =
+sig
+ type locality = Sledgehammer_Filter.locality
+ type minimize_command = string list -> string
+ type metis_params =
+ string * bool * minimize_command * string * (string * locality) list vector
+ * thm * int
+ type isar_params =
+ string Symtab.table * bool * int * Proof.context * int list list
+ type text_result = string * (string * locality) list
+
+ val repair_conjecture_shape_and_axiom_names :
+ string -> int list list -> (string * locality) list vector
+ -> int list list * (string * locality) list vector
+ val apply_on_subgoal : int -> int -> string
+ val command_call : string -> string list -> string
+ val try_command_line : string -> string -> string
+ val minimize_line : ('a list -> string) -> 'a list -> string
+ val metis_proof_text : metis_params -> text_result
+ val isar_proof_text : isar_params -> metis_params -> text_result
+ val proof_text : bool -> isar_params -> metis_params -> text_result
+end;
+
+structure Sledgehammer_ATP_Reconstruct : SLEDGEHAMMER_ATP_RECONSTRUCT =
+struct
+
+open ATP_Problem
+open ATP_Proof
+open Metis_Translate
+open Sledgehammer_Util
+open Sledgehammer_Filter
+open Sledgehammer_ATP_Translate
+
+type minimize_command = string list -> string
+type metis_params =
+ string * bool * minimize_command * string * (string * locality) list vector
+ * thm * int
+type isar_params =
+ string Symtab.table * bool * int * Proof.context * int list list
+type text_result = string * (string * locality) list
+
+fun is_head_digit s = Char.isDigit (String.sub (s, 0))
+val scan_integer = Scan.many1 is_head_digit >> (the o Int.fromString o implode)
+
+fun find_first_in_list_vector vec key =
+ Vector.foldl (fn (ps, NONE) => AList.lookup (op =) ps key
+ | (_, value) => value) NONE vec
+
+
+(** SPASS's Flotter hack **)
+
+(* This is a hack required for keeping track of axioms after they have been
+ clausified by SPASS's Flotter tool. The "ATP/scripts/spass" script is also
+ part of this hack. *)
+
+val set_ClauseFormulaRelationN = "set_ClauseFormulaRelation"
+
+fun extract_clause_sequence output =
+ let
+ val tokens_of = String.tokens (not o Char.isAlphaNum)
+ fun extract_num ("clause" :: (ss as _ :: _)) =
+ Int.fromString (List.last ss)
+ | extract_num _ = NONE
+ in output |> split_lines |> map_filter (extract_num o tokens_of) end
+
+val parse_clause_formula_pair =
+ $$ "(" |-- scan_integer --| $$ ","
+ -- (Symbol.scan_id ::: Scan.repeat ($$ "," |-- Symbol.scan_id)) --| $$ ")"
+ --| Scan.option ($$ ",")
+val parse_clause_formula_relation =
+ Scan.this_string set_ClauseFormulaRelationN |-- $$ "("
+ |-- Scan.repeat parse_clause_formula_pair
+val extract_clause_formula_relation =
+ Substring.full #> Substring.position set_ClauseFormulaRelationN
+ #> snd #> Substring.position "." #> fst #> Substring.string
+ #> explode #> filter_out Symbol.is_blank #> parse_clause_formula_relation
+ #> fst
+
+fun repair_conjecture_shape_and_axiom_names output conjecture_shape
+ axiom_names =
+ if String.isSubstring set_ClauseFormulaRelationN output then
+ let
+ val j0 = hd (hd conjecture_shape)
+ val seq = extract_clause_sequence output
+ val name_map = extract_clause_formula_relation output
+ fun renumber_conjecture j =
+ conjecture_prefix ^ string_of_int (j - j0)
+ |> AList.find (fn (s, ss) => member (op =) ss s) name_map
+ |> map (fn s => find_index (curry (op =) s) seq + 1)
+ fun names_for_number j =
+ j |> AList.lookup (op =) name_map |> these
+ |> map_filter (try (unprefix axiom_prefix)) |> map unascii_of
+ |> map (fn name =>
+ (name, name |> find_first_in_list_vector axiom_names
+ |> the)
+ handle Option.Option =>
+ error ("No such fact: " ^ quote name ^ "."))
+ in
+ (conjecture_shape |> map (maps renumber_conjecture),
+ seq |> map names_for_number |> Vector.fromList)
+ end
+ else
+ (conjecture_shape, axiom_names)
+
+
+(** Soft-core proof reconstruction: Metis one-liner **)
+
+fun string_for_label (s, num) = s ^ string_of_int num
+
+fun apply_on_subgoal _ 1 = "by "
+ | apply_on_subgoal 1 _ = "apply "
+ | apply_on_subgoal i _ = "prefer " ^ string_of_int i ^ " apply "
+fun command_call name [] = name
+ | command_call name args = "(" ^ name ^ " " ^ space_implode " " args ^ ")"
+fun try_command_line banner command =
+ banner ^ ": " ^ Markup.markup Markup.sendback command ^ "."
+fun using_labels [] = ""
+ | using_labels ls =
+ "using " ^ space_implode " " (map string_for_label ls) ^ " "
+fun metis_name full_types = if full_types then "metisFT" else "metis"
+fun metis_call full_types ss = command_call (metis_name full_types) ss
+fun metis_command full_types i n (ls, ss) =
+ using_labels ls ^ apply_on_subgoal i n ^ metis_call full_types ss
+fun metis_line banner full_types i n ss =
+ try_command_line banner (metis_command full_types i n ([], ss))
+fun minimize_line _ [] = ""
+ | minimize_line minimize_command ss =
+ case minimize_command ss of
+ "" => ""
+ | command =>
+ "\nTo minimize the number of lemmas, try this: " ^
+ Markup.markup Markup.sendback command ^ "."
+
+fun resolve_axiom axiom_names ((_, SOME s)) =
+ (case strip_prefix_and_unascii axiom_prefix s of
+ SOME s' => (case find_first_in_list_vector axiom_names s' of
+ SOME x => [(s', x)]
+ | NONE => [])
+ | NONE => [])
+ | resolve_axiom axiom_names (num, NONE) =
+ case Int.fromString num of
+ SOME j =>
+ if j > 0 andalso j <= Vector.length axiom_names then
+ Vector.sub (axiom_names, j - 1)
+ else
+ []
+ | NONE => []
+
+fun add_fact axiom_names (Inference (name, _, [])) =
+ append (resolve_axiom axiom_names name)
+ | add_fact _ _ = I
+
+fun used_facts_in_tstplike_proof axiom_names =
+ atp_proof_from_tstplike_string #> rpair [] #-> fold (add_fact axiom_names)
+
+fun used_facts axiom_names =
+ used_facts_in_tstplike_proof axiom_names
+ #> List.partition (curry (op =) Chained o snd)
+ #> pairself (sort_distinct (string_ord o pairself fst))
+
+fun metis_proof_text (banner, full_types, minimize_command,
+ tstplike_proof, axiom_names, goal, i) =
+ let
+ val (chained_lemmas, other_lemmas) =
+ used_facts axiom_names tstplike_proof
+ val n = Logic.count_prems (prop_of goal)
+ in
+ (metis_line banner full_types i n (map fst other_lemmas) ^
+ minimize_line minimize_command (map fst (other_lemmas @ chained_lemmas)),
+ other_lemmas @ chained_lemmas)
+ end
+
+
+(** Hard-core proof reconstruction: structured Isar proofs **)
+
+(* Simple simplifications to ensure that sort annotations don't leave a trail of
+ spurious "True"s. *)
+fun s_not @{const False} = @{const True}
+ | s_not @{const True} = @{const False}
+ | s_not (@{const Not} $ t) = t
+ | s_not t = @{const Not} $ t
+fun s_conj (@{const True}, t2) = t2
+ | s_conj (t1, @{const True}) = t1
+ | s_conj p = HOLogic.mk_conj p
+fun s_disj (@{const False}, t2) = t2
+ | s_disj (t1, @{const False}) = t1
+ | s_disj p = HOLogic.mk_disj p
+fun s_imp (@{const True}, t2) = t2
+ | s_imp (t1, @{const False}) = s_not t1
+ | s_imp p = HOLogic.mk_imp p
+fun s_iff (@{const True}, t2) = t2
+ | s_iff (t1, @{const True}) = t1
+ | s_iff (t1, t2) = HOLogic.eq_const HOLogic.boolT $ t1 $ t2
+
+fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda v t
+fun exists_of v t = HOLogic.exists_const (fastype_of v) $ lambda v t
+
+fun negate_term (Const (@{const_name All}, T) $ Abs (s, T', t')) =
+ Const (@{const_name Ex}, T) $ Abs (s, T', negate_term t')
+ | negate_term (Const (@{const_name Ex}, T) $ Abs (s, T', t')) =
+ Const (@{const_name All}, T) $ Abs (s, T', negate_term t')
+ | negate_term (@{const HOL.implies} $ t1 $ t2) =
+ @{const HOL.conj} $ t1 $ negate_term t2
+ | negate_term (@{const HOL.conj} $ t1 $ t2) =
+ @{const HOL.disj} $ negate_term t1 $ negate_term t2
+ | negate_term (@{const HOL.disj} $ t1 $ t2) =
+ @{const HOL.conj} $ negate_term t1 $ negate_term t2
+ | negate_term (@{const Not} $ t) = t
+ | negate_term t = @{const Not} $ t
+
+val indent_size = 2
+val no_label = ("", ~1)
+
+val raw_prefix = "X"
+val assum_prefix = "A"
+val fact_prefix = "F"
+
+fun resolve_conjecture conjecture_shape (num, s_opt) =
+ let
+ val k = case try (unprefix conjecture_prefix) (the_default "" s_opt) of
+ SOME s => Int.fromString s |> the_default ~1
+ | NONE => case Int.fromString num of
+ SOME j => find_index (exists (curry (op =) j))
+ conjecture_shape
+ | NONE => ~1
+ in if k >= 0 then [k] else [] end
+
+fun is_axiom conjecture_shape = not o null o resolve_axiom conjecture_shape
+fun is_conjecture conjecture_shape = not o null o resolve_conjecture conjecture_shape
+
+fun raw_label_for_name conjecture_shape name =
+ case resolve_conjecture conjecture_shape name of
+ [j] => (conjecture_prefix, j)
+ | _ => case Int.fromString (fst name) of
+ SOME j => (raw_prefix, j)
+ | NONE => (raw_prefix ^ fst name, 0)
+
+(**** INTERPRETATION OF TSTP SYNTAX TREES ****)
+
+exception FO_TERM of string fo_term list
+exception FORMULA of (string, string fo_term) formula list
+exception SAME of unit
+
+(* Type variables are given the basic sort "HOL.type". Some will later be
+ constrained by information from type literals, or by type inference. *)
+fun type_from_fo_term tfrees (u as ATerm (a, us)) =
+ let val Ts = map (type_from_fo_term tfrees) us in
+ case strip_prefix_and_unascii type_const_prefix a of
+ SOME b => Type (invert_const b, Ts)
+ | NONE =>
+ if not (null us) then
+ raise FO_TERM [u] (* only "tconst"s have type arguments *)
+ else case strip_prefix_and_unascii tfree_prefix a of
+ SOME b =>
+ let val s = "'" ^ b in
+ TFree (s, AList.lookup (op =) tfrees s |> the_default HOLogic.typeS)
+ end
+ | NONE =>
+ case strip_prefix_and_unascii tvar_prefix a of
+ SOME b => TVar (("'" ^ b, 0), HOLogic.typeS)
+ | NONE =>
+ (* Variable from the ATP, say "X1" *)
+ Type_Infer.param 0 (a, HOLogic.typeS)
+ end
+
+(* Type class literal applied to a type. Returns triple of polarity, class,
+ type. *)
+fun type_constraint_from_term pos tfrees (u as ATerm (a, us)) =
+ case (strip_prefix_and_unascii class_prefix a,
+ map (type_from_fo_term tfrees) us) of
+ (SOME b, [T]) => (pos, b, T)
+ | _ => raise FO_TERM [u]
+
+(** Accumulate type constraints in a formula: negative type literals **)
+fun add_var (key, z) = Vartab.map_default (key, []) (cons z)
+fun add_type_constraint (false, cl, TFree (a ,_)) = add_var ((a, ~1), cl)
+ | add_type_constraint (false, cl, TVar (ix, _)) = add_var (ix, cl)
+ | add_type_constraint _ = I
+
+fun repair_atp_variable_name f s =
+ let
+ fun subscript_name s n = s ^ nat_subscript n
+ val s = String.map f s
+ in
+ case space_explode "_" s of
+ [_] => (case take_suffix Char.isDigit (String.explode s) of
+ (cs1 as _ :: _, cs2 as _ :: _) =>
+ subscript_name (String.implode cs1)
+ (the (Int.fromString (String.implode cs2)))
+ | (_, _) => s)
+ | [s1, s2] => (case Int.fromString s2 of
+ SOME n => subscript_name s1 n
+ | NONE => s)
+ | _ => s
+ end
+
+(* First-order translation. No types are known for variables. "HOLogic.typeT"
+ should allow them to be inferred. *)
+fun raw_term_from_pred thy full_types tfrees =
+ let
+ fun aux opt_T extra_us u =
+ case u of
+ ATerm ("hBOOL", [u1]) => aux (SOME @{typ bool}) [] u1
+ | ATerm ("hAPP", [u1, u2]) => aux opt_T (u2 :: extra_us) u1
+ | ATerm (a, us) =>
+ if a = type_wrapper_name then
+ case us of
+ [typ_u, term_u] =>
+ aux (SOME (type_from_fo_term tfrees typ_u)) extra_us term_u
+ | _ => raise FO_TERM us
+ else case strip_prefix_and_unascii const_prefix a of
+ SOME "equal" =>
+ let val ts = map (aux NONE []) us in
+ if length ts = 2 andalso hd ts aconv List.last ts then
+ (* Vampire is keen on producing these. *)
+ @{const True}
+ else
+ list_comb (Const (@{const_name HOL.eq}, HOLogic.typeT), ts)
+ end
+ | SOME b =>
+ let
+ val c = invert_const b
+ val num_type_args = num_type_args thy c
+ val (type_us, term_us) =
+ chop (if full_types then 0 else num_type_args) us
+ (* Extra args from "hAPP" come after any arguments given directly to
+ the constant. *)
+ val term_ts = map (aux NONE []) term_us
+ val extra_ts = map (aux NONE []) extra_us
+ val t =
+ Const (c, if full_types then
+ case opt_T of
+ SOME T => map fastype_of term_ts ---> T
+ | NONE =>
+ if num_type_args = 0 then
+ Sign.const_instance thy (c, [])
+ else
+ raise Fail ("no type information for " ^ quote c)
+ else
+ Sign.const_instance thy (c,
+ map (type_from_fo_term tfrees) type_us))
+ in list_comb (t, term_ts @ extra_ts) end
+ | NONE => (* a free or schematic variable *)
+ let
+ val ts = map (aux NONE []) (us @ extra_us)
+ val T = map fastype_of ts ---> HOLogic.typeT
+ val t =
+ case strip_prefix_and_unascii fixed_var_prefix a of
+ SOME b => Free (b, T)
+ | NONE =>
+ case strip_prefix_and_unascii schematic_var_prefix a of
+ SOME b => Var ((b, 0), T)
+ | NONE =>
+ if is_atp_variable a then
+ Var ((repair_atp_variable_name Char.toLower a, 0), T)
+ else
+ (* Skolem constants? *)
+ Var ((repair_atp_variable_name Char.toUpper a, 0), T)
+ in list_comb (t, ts) end
+ in aux (SOME HOLogic.boolT) [] end
+
+fun term_from_pred thy full_types tfrees pos (u as ATerm (s, _)) =
+ if String.isPrefix class_prefix s then
+ add_type_constraint (type_constraint_from_term pos tfrees u)
+ #> pair @{const True}
+ else
+ pair (raw_term_from_pred thy full_types tfrees u)
+
+val combinator_table =
+ [(@{const_name Meson.COMBI}, @{thm Meson.COMBI_def_raw}),
+ (@{const_name Meson.COMBK}, @{thm Meson.COMBK_def_raw}),
+ (@{const_name Meson.COMBB}, @{thm Meson.COMBB_def_raw}),
+ (@{const_name Meson.COMBC}, @{thm Meson.COMBC_def_raw}),
+ (@{const_name Meson.COMBS}, @{thm Meson.COMBS_def_raw})]
+
+fun uncombine_term (t1 $ t2) = betapply (pairself uncombine_term (t1, t2))
+ | uncombine_term (Abs (s, T, t')) = Abs (s, T, uncombine_term t')
+ | uncombine_term (t as Const (x as (s, _))) =
+ (case AList.lookup (op =) combinator_table s of
+ SOME thm => thm |> prop_of |> specialize_type @{theory} x |> Logic.dest_equals |> snd
+ | NONE => t)
+ | uncombine_term t = t
+
+(* Update schematic type variables with detected sort constraints. It's not
+ totally clear when this code is necessary. *)
+fun repair_tvar_sorts (t, tvar_tab) =
+ let
+ fun do_type (Type (a, Ts)) = Type (a, map do_type Ts)
+ | do_type (TVar (xi, s)) =
+ TVar (xi, the_default s (Vartab.lookup tvar_tab xi))
+ | do_type (TFree z) = TFree z
+ fun do_term (Const (a, T)) = Const (a, do_type T)
+ | do_term (Free (a, T)) = Free (a, do_type T)
+ | do_term (Var (xi, T)) = Var (xi, do_type T)
+ | do_term (t as Bound _) = t
+ | do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t)
+ | do_term (t1 $ t2) = do_term t1 $ do_term t2
+ in t |> not (Vartab.is_empty tvar_tab) ? do_term end
+
+fun quantify_over_var quant_of var_s t =
+ let
+ val vars = [] |> Term.add_vars t |> filter (fn ((s, _), _) => s = var_s)
+ |> map Var
+ in fold_rev quant_of vars t end
+
+(* Interpret an ATP formula as a HOL term, extracting sort constraints as they
+ appear in the formula. *)
+fun prop_from_formula thy full_types tfrees phi =
+ let
+ fun do_formula pos phi =
+ case phi of
+ AQuant (_, [], phi) => do_formula pos phi
+ | AQuant (q, x :: xs, phi') =>
+ do_formula pos (AQuant (q, xs, phi'))
+ #>> quantify_over_var (case q of
+ AForall => forall_of
+ | AExists => exists_of)
+ (repair_atp_variable_name Char.toLower x)
+ | AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not
+ | AConn (c, [phi1, phi2]) =>
+ do_formula (pos |> c = AImplies ? not) phi1
+ ##>> do_formula pos phi2
+ #>> (case c of
+ AAnd => s_conj
+ | AOr => s_disj
+ | AImplies => s_imp
+ | AIf => s_imp o swap
+ | AIff => s_iff
+ | ANotIff => s_not o s_iff
+ | _ => raise Fail "unexpected connective")
+ | AAtom tm => term_from_pred thy full_types tfrees pos tm
+ | _ => raise FORMULA [phi]
+ in repair_tvar_sorts (do_formula true phi Vartab.empty) end
+
+fun check_formula ctxt =
+ Type.constraint HOLogic.boolT
+ #> Syntax.check_term (ProofContext.set_mode ProofContext.mode_schematic ctxt)
+
+
+(**** Translation of TSTP files to Isar Proofs ****)
+
+fun unvarify_term (Var ((s, 0), T)) = Free (s, T)
+ | unvarify_term t = raise TERM ("unvarify_term: non-Var", [t])
+
+fun decode_line full_types tfrees (Definition (name, phi1, phi2)) ctxt =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val t1 = prop_from_formula thy full_types tfrees phi1
+ val vars = snd (strip_comb t1)
+ val frees = map unvarify_term vars
+ val unvarify_args = subst_atomic (vars ~~ frees)
+ val t2 = prop_from_formula thy full_types tfrees phi2
+ val (t1, t2) =
+ HOLogic.eq_const HOLogic.typeT $ t1 $ t2
+ |> unvarify_args |> uncombine_term |> check_formula ctxt
+ |> HOLogic.dest_eq
+ in
+ (Definition (name, t1, t2),
+ fold Variable.declare_term (maps OldTerm.term_frees [t1, t2]) ctxt)
+ end
+ | decode_line full_types tfrees (Inference (name, u, deps)) ctxt =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val t = u |> prop_from_formula thy full_types tfrees
+ |> uncombine_term |> check_formula ctxt
+ in
+ (Inference (name, t, deps),
+ fold Variable.declare_term (OldTerm.term_frees t) ctxt)
+ end
+fun decode_lines ctxt full_types tfrees lines =
+ fst (fold_map (decode_line full_types tfrees) lines ctxt)
+
+fun is_same_inference _ (Definition _) = false
+ | is_same_inference t (Inference (_, t', _)) = t aconv t'
+
+(* No "real" literals means only type information (tfree_tcs, clsrel, or
+ clsarity). *)
+val is_only_type_information = curry (op aconv) HOLogic.true_const
+
+fun replace_one_dependency (old, new) dep =
+ if is_same_step (dep, old) then new else [dep]
+fun replace_dependencies_in_line _ (line as Definition _) = line
+ | replace_dependencies_in_line p (Inference (name, t, deps)) =
+ Inference (name, t, fold (union (op =) o replace_one_dependency p) deps [])
+
+(* Discard axioms; consolidate adjacent lines that prove the same formula, since
+ they differ only in type information.*)
+fun add_line _ _ (line as Definition _) lines = line :: lines
+ | add_line conjecture_shape axiom_names (Inference (name, t, [])) lines =
+ (* No dependencies: axiom, conjecture, or (for Vampire) internal axioms or
+ definitions. *)
+ if is_axiom axiom_names name then
+ (* Axioms are not proof lines. *)
+ if is_only_type_information t then
+ map (replace_dependencies_in_line (name, [])) lines
+ (* Is there a repetition? If so, replace later line by earlier one. *)
+ else case take_prefix (not o is_same_inference t) lines of
+ (_, []) => lines (* no repetition of proof line *)
+ | (pre, Inference (name', _, _) :: post) =>
+ pre @ map (replace_dependencies_in_line (name', [name])) post
+ | _ => raise Fail "unexpected inference"
+ else if is_conjecture conjecture_shape name then
+ Inference (name, negate_term t, []) :: lines
+ else
+ map (replace_dependencies_in_line (name, [])) lines
+ | add_line _ _ (Inference (name, t, deps)) lines =
+ (* Type information will be deleted later; skip repetition test. *)
+ if is_only_type_information t then
+ Inference (name, t, deps) :: lines
+ (* Is there a repetition? If so, replace later line by earlier one. *)
+ else case take_prefix (not o is_same_inference t) lines of
+ (* FIXME: Doesn't this code risk conflating proofs involving different
+ types? *)
+ (_, []) => Inference (name, t, deps) :: lines
+ | (pre, Inference (name', t', _) :: post) =>
+ Inference (name, t', deps) ::
+ pre @ map (replace_dependencies_in_line (name', [name])) post
+ | _ => raise Fail "unexpected inference"
+
+(* Recursively delete empty lines (type information) from the proof. *)
+fun add_nontrivial_line (Inference (name, t, [])) lines =
+ if is_only_type_information t then delete_dependency name lines
+ else Inference (name, t, []) :: lines
+ | add_nontrivial_line line lines = line :: lines
+and delete_dependency name lines =
+ fold_rev add_nontrivial_line
+ (map (replace_dependencies_in_line (name, [])) lines) []
+
+(* ATPs sometimes reuse free variable names in the strangest ways. Removing
+ offending lines often does the trick. *)
+fun is_bad_free frees (Free x) = not (member (op =) frees x)
+ | is_bad_free _ _ = false
+
+fun add_desired_line _ _ _ _ (line as Definition (name, _, _)) (j, lines) =
+ (j, line :: map (replace_dependencies_in_line (name, [])) lines)
+ | add_desired_line isar_shrink_factor conjecture_shape axiom_names frees
+ (Inference (name, t, deps)) (j, lines) =
+ (j + 1,
+ if is_axiom axiom_names name orelse
+ is_conjecture conjecture_shape name orelse
+ (* the last line must be kept *)
+ j = 0 orelse
+ (not (is_only_type_information t) andalso
+ null (Term.add_tvars t []) andalso
+ not (exists_subterm (is_bad_free frees) t) andalso
+ length deps >= 2 andalso j mod isar_shrink_factor = 0 andalso
+ (* kill next to last line, which usually results in a trivial step *)
+ j <> 1) then
+ Inference (name, t, deps) :: lines (* keep line *)
+ else
+ map (replace_dependencies_in_line (name, deps)) lines) (* drop line *)
+
+(** Isar proof construction and manipulation **)
+
+fun merge_fact_sets (ls1, ss1) (ls2, ss2) =
+ (union (op =) ls1 ls2, union (op =) ss1 ss2)
+
+type label = string * int
+type facts = label list * string list
+
+datatype isar_qualifier = Show | Then | Moreover | Ultimately
+
+datatype isar_step =
+ Fix of (string * typ) list |
+ Let of term * term |
+ Assume of label * term |
+ Have of isar_qualifier list * label * term * byline
+and byline =
+ ByMetis of facts |
+ CaseSplit of isar_step list list * facts
+
+fun smart_case_split [] facts = ByMetis facts
+ | smart_case_split proofs facts = CaseSplit (proofs, facts)
+
+fun add_fact_from_dependency conjecture_shape axiom_names name =
+ if is_axiom axiom_names name then
+ apsnd (union (op =) (map fst (resolve_axiom axiom_names name)))
+ else
+ apfst (insert (op =) (raw_label_for_name conjecture_shape name))
+
+fun step_for_line _ _ _ (Definition (_, t1, t2)) = Let (t1, t2)
+ | step_for_line conjecture_shape _ _ (Inference (name, t, [])) =
+ Assume (raw_label_for_name conjecture_shape name, t)
+ | step_for_line conjecture_shape axiom_names j (Inference (name, t, deps)) =
+ Have (if j = 1 then [Show] else [],
+ raw_label_for_name conjecture_shape name,
+ fold_rev forall_of (map Var (Term.add_vars t [])) t,
+ ByMetis (fold (add_fact_from_dependency conjecture_shape axiom_names)
+ deps ([], [])))
+
+fun repair_name "$true" = "c_True"
+ | repair_name "$false" = "c_False"
+ | repair_name "$$e" = "c_equal" (* seen in Vampire proofs *)
+ | repair_name "equal" = "c_equal" (* needed by SPASS? *)
+ | repair_name s =
+ if String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s then
+ "c_equal" (* seen in Vampire proofs *)
+ else
+ s
+
+fun isar_proof_from_tstplike_proof pool ctxt full_types tfrees isar_shrink_factor
+ tstplike_proof conjecture_shape axiom_names params frees =
+ let
+ val lines =
+ tstplike_proof
+ |> atp_proof_from_tstplike_string
+ |> nasty_atp_proof pool
+ |> map_term_names_in_atp_proof repair_name
+ |> decode_lines ctxt full_types tfrees
+ |> rpair [] |-> fold_rev (add_line conjecture_shape axiom_names)
+ |> rpair [] |-> fold_rev add_nontrivial_line
+ |> rpair (0, []) |-> fold_rev (add_desired_line isar_shrink_factor
+ conjecture_shape axiom_names frees)
+ |> snd
+ in
+ (if null params then [] else [Fix params]) @
+ map2 (step_for_line conjecture_shape axiom_names) (length lines downto 1)
+ lines
+ end
+
+(* When redirecting proofs, we keep information about the labels seen so far in
+ the "backpatches" data structure. The first component indicates which facts
+ should be associated with forthcoming proof steps. The second component is a
+ pair ("assum_ls", "drop_ls"), where "assum_ls" are the labels that should
+ become assumptions and "drop_ls" are the labels that should be dropped in a
+ case split. *)
+type backpatches = (label * facts) list * (label list * label list)
+
+fun used_labels_of_step (Have (_, _, _, by)) =
+ (case by of
+ ByMetis (ls, _) => ls
+ | CaseSplit (proofs, (ls, _)) =>
+ fold (union (op =) o used_labels_of) proofs ls)
+ | used_labels_of_step _ = []
+and used_labels_of proof = fold (union (op =) o used_labels_of_step) proof []
+
+fun new_labels_of_step (Fix _) = []
+ | new_labels_of_step (Let _) = []
+ | new_labels_of_step (Assume (l, _)) = [l]
+ | new_labels_of_step (Have (_, l, _, _)) = [l]
+val new_labels_of = maps new_labels_of_step
+
+val join_proofs =
+ let
+ fun aux _ [] = NONE
+ | aux proof_tail (proofs as (proof1 :: _)) =
+ if exists null proofs then
+ NONE
+ else if forall (curry (op =) (hd proof1) o hd) (tl proofs) then
+ aux (hd proof1 :: proof_tail) (map tl proofs)
+ else case hd proof1 of
+ Have ([], l, t, _) => (* FIXME: should we really ignore the "by"? *)
+ if forall (fn Have ([], l', t', _) :: _ => (l, t) = (l', t')
+ | _ => false) (tl proofs) andalso
+ not (exists (member (op =) (maps new_labels_of proofs))
+ (used_labels_of proof_tail)) then
+ SOME (l, t, map rev proofs, proof_tail)
+ else
+ NONE
+ | _ => NONE
+ in aux [] o map rev end
+
+fun case_split_qualifiers proofs =
+ case length proofs of
+ 0 => []
+ | 1 => [Then]
+ | _ => [Ultimately]
+
+fun redirect_proof hyp_ts concl_t proof =
+ let
+ (* The first pass outputs those steps that are independent of the negated
+ conjecture. The second pass flips the proof by contradiction to obtain a
+ direct proof, introducing case splits when an inference depends on
+ several facts that depend on the negated conjecture. *)
+ val concl_l = (conjecture_prefix, length hyp_ts)
+ fun first_pass ([], contra) = ([], contra)
+ | first_pass ((step as Fix _) :: proof, contra) =
+ first_pass (proof, contra) |>> cons step
+ | first_pass ((step as Let _) :: proof, contra) =
+ first_pass (proof, contra) |>> cons step
+ | first_pass ((step as Assume (l as (_, j), _)) :: proof, contra) =
+ if l = concl_l then first_pass (proof, contra ||> cons step)
+ else first_pass (proof, contra) |>> cons (Assume (l, nth hyp_ts j))
+ | first_pass (Have (qs, l, t, ByMetis (ls, ss)) :: proof, contra) =
+ let val step = Have (qs, l, t, ByMetis (ls, ss)) in
+ if exists (member (op =) (fst contra)) ls then
+ first_pass (proof, contra |>> cons l ||> cons step)
+ else
+ first_pass (proof, contra) |>> cons step
+ end
+ | first_pass _ = raise Fail "malformed proof"
+ val (proof_top, (contra_ls, contra_proof)) =
+ first_pass (proof, ([concl_l], []))
+ val backpatch_label = the_default ([], []) oo AList.lookup (op =) o fst
+ fun backpatch_labels patches ls =
+ fold merge_fact_sets (map (backpatch_label patches) ls) ([], [])
+ fun second_pass end_qs ([], assums, patches) =
+ ([Have (end_qs, no_label, concl_t,
+ ByMetis (backpatch_labels patches (map snd assums)))], patches)
+ | second_pass end_qs (Assume (l, t) :: proof, assums, patches) =
+ second_pass end_qs (proof, (t, l) :: assums, patches)
+ | second_pass end_qs (Have (qs, l, t, ByMetis (ls, ss)) :: proof, assums,
+ patches) =
+ (if member (op =) (snd (snd patches)) l andalso
+ not (member (op =) (fst (snd patches)) l) andalso
+ not (AList.defined (op =) (fst patches) l) then
+ second_pass end_qs (proof, assums, patches ||> apsnd (append ls))
+ else case List.partition (member (op =) contra_ls) ls of
+ ([contra_l], co_ls) =>
+ if member (op =) qs Show then
+ second_pass end_qs (proof, assums,
+ patches |>> cons (contra_l, (co_ls, ss)))
+ else
+ second_pass end_qs
+ (proof, assums,
+ patches |>> cons (contra_l, (l :: co_ls, ss)))
+ |>> cons (if member (op =) (fst (snd patches)) l then
+ Assume (l, negate_term t)
+ else
+ Have (qs, l, negate_term t,
+ ByMetis (backpatch_label patches l)))
+ | (contra_ls as _ :: _, co_ls) =>
+ let
+ val proofs =
+ map_filter
+ (fn l =>
+ if l = concl_l then
+ NONE
+ else
+ let
+ val drop_ls = filter (curry (op <>) l) contra_ls
+ in
+ second_pass []
+ (proof, assums,
+ patches ||> apfst (insert (op =) l)
+ ||> apsnd (union (op =) drop_ls))
+ |> fst |> SOME
+ end) contra_ls
+ val (assumes, facts) =
+ if member (op =) (fst (snd patches)) l then
+ ([Assume (l, negate_term t)], (l :: co_ls, ss))
+ else
+ ([], (co_ls, ss))
+ in
+ (case join_proofs proofs of
+ SOME (l, t, proofs, proof_tail) =>
+ Have (case_split_qualifiers proofs @
+ (if null proof_tail then end_qs else []), l, t,
+ smart_case_split proofs facts) :: proof_tail
+ | NONE =>
+ [Have (case_split_qualifiers proofs @ end_qs, no_label,
+ concl_t, smart_case_split proofs facts)],
+ patches)
+ |>> append assumes
+ end
+ | _ => raise Fail "malformed proof")
+ | second_pass _ _ = raise Fail "malformed proof"
+ val proof_bottom =
+ second_pass [Show] (contra_proof, [], ([], ([], []))) |> fst
+ in proof_top @ proof_bottom end
+
+(* FIXME: Still needed? Probably not. *)
+val kill_duplicate_assumptions_in_proof =
+ let
+ fun relabel_facts subst =
+ apfst (map (fn l => AList.lookup (op =) subst l |> the_default l))
+ fun do_step (step as Assume (l, t)) (proof, subst, assums) =
+ (case AList.lookup (op aconv) assums t of
+ SOME l' => (proof, (l, l') :: subst, assums)
+ | NONE => (step :: proof, subst, (t, l) :: assums))
+ | do_step (Have (qs, l, t, by)) (proof, subst, assums) =
+ (Have (qs, l, t,
+ case by of
+ ByMetis facts => ByMetis (relabel_facts subst facts)
+ | CaseSplit (proofs, facts) =>
+ CaseSplit (map do_proof proofs, relabel_facts subst facts)) ::
+ proof, subst, assums)
+ | do_step step (proof, subst, assums) = (step :: proof, subst, assums)
+ and do_proof proof = fold do_step proof ([], [], []) |> #1 |> rev
+ in do_proof end
+
+val then_chain_proof =
+ let
+ fun aux _ [] = []
+ | aux _ ((step as Assume (l, _)) :: proof) = step :: aux l proof
+ | aux l' (Have (qs, l, t, by) :: proof) =
+ (case by of
+ ByMetis (ls, ss) =>
+ Have (if member (op =) ls l' then
+ (Then :: qs, l, t,
+ ByMetis (filter_out (curry (op =) l') ls, ss))
+ else
+ (qs, l, t, ByMetis (ls, ss)))
+ | CaseSplit (proofs, facts) =>
+ Have (qs, l, t, CaseSplit (map (aux no_label) proofs, facts))) ::
+ aux l proof
+ | aux _ (step :: proof) = step :: aux no_label proof
+ in aux no_label end
+
+fun kill_useless_labels_in_proof proof =
+ let
+ val used_ls = used_labels_of proof
+ fun do_label l = if member (op =) used_ls l then l else no_label
+ fun do_step (Assume (l, t)) = Assume (do_label l, t)
+ | do_step (Have (qs, l, t, by)) =
+ Have (qs, do_label l, t,
+ case by of
+ CaseSplit (proofs, facts) =>
+ CaseSplit (map (map do_step) proofs, facts)
+ | _ => by)
+ | do_step step = step
+ in map do_step proof end
+
+fun prefix_for_depth n = replicate_string (n + 1)
+
+val relabel_proof =
+ let
+ fun aux _ _ _ [] = []
+ | aux subst depth (next_assum, next_fact) (Assume (l, t) :: proof) =
+ if l = no_label then
+ Assume (l, t) :: aux subst depth (next_assum, next_fact) proof
+ else
+ let val l' = (prefix_for_depth depth assum_prefix, next_assum) in
+ Assume (l', t) ::
+ aux ((l, l') :: subst) depth (next_assum + 1, next_fact) proof
+ end
+ | aux subst depth (next_assum, next_fact) (Have (qs, l, t, by) :: proof) =
+ let
+ val (l', subst, next_fact) =
+ if l = no_label then
+ (l, subst, next_fact)
+ else
+ let
+ val l' = (prefix_for_depth depth fact_prefix, next_fact)
+ in (l', (l, l') :: subst, next_fact + 1) end
+ val relabel_facts =
+ apfst (maps (the_list o AList.lookup (op =) subst))
+ val by =
+ case by of
+ ByMetis facts => ByMetis (relabel_facts facts)
+ | CaseSplit (proofs, facts) =>
+ CaseSplit (map (aux subst (depth + 1) (1, 1)) proofs,
+ relabel_facts facts)
+ in
+ Have (qs, l', t, by) ::
+ aux subst depth (next_assum, next_fact) proof
+ end
+ | aux subst depth nextp (step :: proof) =
+ step :: aux subst depth nextp proof
+ in aux [] 0 (1, 1) end
+
+fun string_for_proof ctxt0 full_types i n =
+ let
+ val ctxt = ctxt0
+ |> Config.put show_free_types false
+ |> Config.put show_types true
+ fun fix_print_mode f x =
+ Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN)
+ (print_mode_value ())) f x
+ fun do_indent ind = replicate_string (ind * indent_size) " "
+ fun do_free (s, T) =
+ maybe_quote s ^ " :: " ^
+ maybe_quote (fix_print_mode (Syntax.string_of_typ ctxt) T)
+ fun do_label l = if l = no_label then "" else string_for_label l ^ ": "
+ fun do_have qs =
+ (if member (op =) qs Moreover then "moreover " else "") ^
+ (if member (op =) qs Ultimately then "ultimately " else "") ^
+ (if member (op =) qs Then then
+ if member (op =) qs Show then "thus" else "hence"
+ else
+ if member (op =) qs Show then "show" else "have")
+ val do_term = maybe_quote o fix_print_mode (Syntax.string_of_term ctxt)
+ fun do_facts (ls, ss) =
+ metis_command full_types 1 1
+ (ls |> sort_distinct (prod_ord string_ord int_ord),
+ ss |> sort_distinct string_ord)
+ and do_step ind (Fix xs) =
+ do_indent ind ^ "fix " ^ space_implode " and " (map do_free xs) ^ "\n"
+ | do_step ind (Let (t1, t2)) =
+ do_indent ind ^ "let " ^ do_term t1 ^ " = " ^ do_term t2 ^ "\n"
+ | do_step ind (Assume (l, t)) =
+ do_indent ind ^ "assume " ^ do_label l ^ do_term t ^ "\n"
+ | do_step ind (Have (qs, l, t, ByMetis facts)) =
+ do_indent ind ^ do_have qs ^ " " ^
+ do_label l ^ do_term t ^ " " ^ do_facts facts ^ "\n"
+ | do_step ind (Have (qs, l, t, CaseSplit (proofs, facts))) =
+ space_implode (do_indent ind ^ "moreover\n")
+ (map (do_block ind) proofs) ^
+ do_indent ind ^ do_have qs ^ " " ^ do_label l ^ do_term t ^ " " ^
+ do_facts facts ^ "\n"
+ and do_steps prefix suffix ind steps =
+ let val s = implode (map (do_step ind) steps) in
+ replicate_string (ind * indent_size - size prefix) " " ^ prefix ^
+ String.extract (s, ind * indent_size,
+ SOME (size s - ind * indent_size - 1)) ^
+ suffix ^ "\n"
+ end
+ and do_block ind proof = do_steps "{ " " }" (ind + 1) proof
+ (* One-step proofs are pointless; better use the Metis one-liner
+ directly. *)
+ and do_proof [Have (_, _, _, ByMetis _)] = ""
+ | do_proof proof =
+ (if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^
+ do_indent 0 ^ "proof -\n" ^ do_steps "" "" 1 proof ^ do_indent 0 ^
+ (if n <> 1 then "next" else "qed")
+ in do_proof end
+
+fun isar_proof_text (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
+ (other_params as (_, full_types, _, tstplike_proof,
+ axiom_names, goal, i)) =
+ let
+ val (params, hyp_ts, concl_t) = strip_subgoal goal i
+ val frees = fold Term.add_frees (concl_t :: hyp_ts) []
+ val tfrees = fold Term.add_tfrees (concl_t :: hyp_ts) []
+ val n = Logic.count_prems (prop_of goal)
+ val (one_line_proof, lemma_names) = metis_proof_text other_params
+ fun isar_proof_for () =
+ case isar_proof_from_tstplike_proof pool ctxt full_types tfrees
+ isar_shrink_factor tstplike_proof conjecture_shape axiom_names
+ params frees
+ |> redirect_proof hyp_ts concl_t
+ |> kill_duplicate_assumptions_in_proof
+ |> then_chain_proof
+ |> kill_useless_labels_in_proof
+ |> relabel_proof
+ |> string_for_proof ctxt full_types i n of
+ "" => "\nNo structured proof available."
+ | proof => "\n\nStructured proof:\n" ^ Markup.markup Markup.sendback proof
+ val isar_proof =
+ if debug then
+ isar_proof_for ()
+ else
+ try isar_proof_for ()
+ |> the_default "\nWarning: The Isar proof construction failed."
+ in (one_line_proof ^ isar_proof, lemma_names) end
+
+fun proof_text isar_proof isar_params other_params =
+ (if isar_proof then isar_proof_text isar_params else metis_proof_text)
+ other_params
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_atp_translate.ML Sun Oct 24 15:11:24 2010 -0700
@@ -0,0 +1,534 @@
+(* Title: HOL/Tools/Sledgehammer/sledgehammer_translate.ML
+ Author: Fabian Immler, TU Muenchen
+ Author: Makarius
+ Author: Jasmin Blanchette, TU Muenchen
+
+Translation of HOL to FOL for Sledgehammer.
+*)
+
+signature SLEDGEHAMMER_ATP_TRANSLATE =
+sig
+ type 'a problem = 'a ATP_Problem.problem
+ type prepared_formula
+
+ val axiom_prefix : string
+ val conjecture_prefix : string
+ val prepare_axiom :
+ Proof.context -> (string * 'a) * thm
+ -> term * ((string * 'a) * prepared_formula) option
+ val prepare_atp_problem :
+ Proof.context -> bool -> bool -> bool -> bool -> term list -> term
+ -> (term * ((string * 'a) * prepared_formula) option) list
+ -> string problem * string Symtab.table * int * (string * 'a) list vector
+end;
+
+structure Sledgehammer_ATP_Translate : SLEDGEHAMMER_ATP_TRANSLATE =
+struct
+
+open ATP_Problem
+open Metis_Translate
+open Sledgehammer_Util
+
+val axiom_prefix = "ax_"
+val conjecture_prefix = "conj_"
+val helper_prefix = "help_"
+val class_rel_clause_prefix = "clrel_";
+val arity_clause_prefix = "arity_"
+val tfree_prefix = "tfree_"
+
+(* Freshness almost guaranteed! *)
+val sledgehammer_weak_prefix = "Sledgehammer:"
+
+type prepared_formula =
+ {name: string,
+ kind: kind,
+ combformula: (name, combterm) formula,
+ ctypes_sorts: typ list}
+
+fun mk_anot phi = AConn (ANot, [phi])
+fun mk_aconn c phi1 phi2 = AConn (c, [phi1, phi2])
+fun mk_ahorn [] phi = phi
+ | mk_ahorn (phi :: phis) psi =
+ AConn (AImplies, [fold (mk_aconn AAnd) phis phi, psi])
+
+fun combformula_for_prop thy =
+ let
+ val do_term = combterm_from_term thy ~1
+ fun do_quant bs q s T t' =
+ let val s = Name.variant (map fst bs) s in
+ do_formula ((s, T) :: bs) t'
+ #>> (fn phi => AQuant (q, [`make_bound_var s], phi))
+ end
+ and do_conn bs c t1 t2 =
+ do_formula bs t1 ##>> do_formula bs t2
+ #>> (fn (phi1, phi2) => AConn (c, [phi1, phi2]))
+ and do_formula bs t =
+ case t of
+ @{const Not} $ t1 =>
+ do_formula bs t1 #>> (fn phi => AConn (ANot, [phi]))
+ | Const (@{const_name All}, _) $ Abs (s, T, t') =>
+ do_quant bs AForall s T t'
+ | Const (@{const_name Ex}, _) $ Abs (s, T, t') =>
+ do_quant bs AExists s T t'
+ | @{const HOL.conj} $ t1 $ t2 => do_conn bs AAnd t1 t2
+ | @{const HOL.disj} $ t1 $ t2 => do_conn bs AOr t1 t2
+ | @{const HOL.implies} $ t1 $ t2 => do_conn bs AImplies t1 t2
+ | Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
+ do_conn bs AIff t1 t2
+ | _ => (fn ts => do_term bs (Envir.eta_contract t)
+ |>> AAtom ||> union (op =) ts)
+ in do_formula [] end
+
+val presimplify_term = prop_of o Meson.presimplify oo Skip_Proof.make_thm
+
+fun concealed_bound_name j = sledgehammer_weak_prefix ^ Int.toString j
+fun conceal_bounds Ts t =
+ subst_bounds (map (Free o apfst concealed_bound_name)
+ (0 upto length Ts - 1 ~~ Ts), t)
+fun reveal_bounds Ts =
+ subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j))
+ (0 upto length Ts - 1 ~~ Ts))
+
+(* Removes the lambdas from an equation of the form "t = (%x. u)".
+ (Cf. "extensionalize_theorem" in "Meson_Clausify".) *)
+fun extensionalize_term t =
+ let
+ fun aux j (@{const Trueprop} $ t') = @{const Trueprop} $ aux j t'
+ | aux j (t as Const (s, Type (_, [Type (_, [_, T']),
+ Type (_, [_, res_T])]))
+ $ t2 $ Abs (var_s, var_T, t')) =
+ if s = @{const_name HOL.eq} orelse s = @{const_name "=="} then
+ let val var_t = Var ((var_s, j), var_T) in
+ Const (s, T' --> T' --> res_T)
+ $ betapply (t2, var_t) $ subst_bound (var_t, t')
+ |> aux (j + 1)
+ end
+ else
+ t
+ | aux _ t = t
+ in aux (maxidx_of_term t + 1) t end
+
+fun introduce_combinators_in_term ctxt kind t =
+ let val thy = ProofContext.theory_of ctxt in
+ if Meson.is_fol_term thy t then
+ t
+ else
+ let
+ fun aux Ts t =
+ case t of
+ @{const Not} $ t1 => @{const Not} $ aux Ts t1
+ | (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') =>
+ t0 $ Abs (s, T, aux (T :: Ts) t')
+ | (t0 as Const (@{const_name All}, _)) $ t1 =>
+ aux Ts (t0 $ eta_expand Ts t1 1)
+ | (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') =>
+ t0 $ Abs (s, T, aux (T :: Ts) t')
+ | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
+ aux Ts (t0 $ eta_expand Ts t1 1)
+ | (t0 as @{const HOL.conj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
+ | (t0 as @{const HOL.disj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
+ | (t0 as @{const HOL.implies}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
+ | (t0 as Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])))
+ $ t1 $ t2 =>
+ t0 $ aux Ts t1 $ aux Ts t2
+ | _ => if not (exists_subterm (fn Abs _ => true | _ => false) t) then
+ t
+ else
+ t |> conceal_bounds Ts
+ |> Envir.eta_contract
+ |> cterm_of thy
+ |> Meson_Clausify.introduce_combinators_in_cterm
+ |> prop_of |> Logic.dest_equals |> snd
+ |> reveal_bounds Ts
+ val (t, ctxt') = Variable.import_terms true [t] ctxt |>> the_single
+ in t |> aux [] |> singleton (Variable.export_terms ctxt' ctxt) end
+ handle THM _ =>
+ (* A type variable of sort "{}" will make abstraction fail. *)
+ if kind = Conjecture then HOLogic.false_const
+ else HOLogic.true_const
+ end
+
+(* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the
+ same in Sledgehammer to prevent the discovery of unreplable proofs. *)
+fun freeze_term t =
+ let
+ fun aux (t $ u) = aux t $ aux u
+ | aux (Abs (s, T, t)) = Abs (s, T, aux t)
+ | aux (Var ((s, i), T)) =
+ Free (sledgehammer_weak_prefix ^ s ^ "_" ^ string_of_int i, T)
+ | aux t = t
+ in t |> exists_subterm is_Var t ? aux end
+
+(* "Object_Logic.atomize_term" isn't as powerful as it could be; for example,
+ it leaves metaequalities over "prop"s alone. *)
+val atomize_term =
+ let
+ fun aux (@{const Trueprop} $ t1) = t1
+ | aux (Const (@{const_name all}, _) $ Abs (s, T, t')) =
+ HOLogic.all_const T $ Abs (s, T, aux t')
+ | aux (@{const "==>"} $ t1 $ t2) = HOLogic.mk_imp (pairself aux (t1, t2))
+ | aux (Const (@{const_name "=="}, Type (_, [@{typ prop}, _])) $ t1 $ t2) =
+ HOLogic.eq_const HOLogic.boolT $ aux t1 $ aux t2
+ | aux (Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2) =
+ HOLogic.eq_const T $ t1 $ t2
+ | aux _ = raise Fail "aux"
+ in perhaps (try aux) end
+
+(* making axiom and conjecture formulas *)
+fun make_formula ctxt presimp name kind t =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val t = t |> Envir.beta_eta_contract
+ |> transform_elim_term
+ |> atomize_term
+ val need_trueprop = (fastype_of t = HOLogic.boolT)
+ val t = t |> need_trueprop ? HOLogic.mk_Trueprop
+ |> extensionalize_term
+ |> presimp ? presimplify_term thy
+ |> perhaps (try (HOLogic.dest_Trueprop))
+ |> introduce_combinators_in_term ctxt kind
+ |> kind <> Axiom ? freeze_term
+ val (combformula, ctypes_sorts) = combformula_for_prop thy t []
+ in
+ {name = name, combformula = combformula, kind = kind,
+ ctypes_sorts = ctypes_sorts}
+ end
+
+fun make_axiom ctxt presimp ((name, loc), th) =
+ case make_formula ctxt presimp name Axiom (prop_of th) of
+ {combformula = AAtom (CombConst (("c_True", _), _, _)), ...} => NONE
+ | formula => SOME ((name, loc), formula)
+fun make_conjecture ctxt ts =
+ let val last = length ts - 1 in
+ map2 (fn j => make_formula ctxt true (Int.toString j)
+ (if j = last then Conjecture else Hypothesis))
+ (0 upto last) ts
+ end
+
+(** Helper facts **)
+
+fun count_combterm (CombConst ((s, _), _, _)) =
+ Symtab.map_entry s (Integer.add 1)
+ | count_combterm (CombVar _) = I
+ | count_combterm (CombApp (t1, t2)) = fold count_combterm [t1, t2]
+fun count_combformula (AQuant (_, _, phi)) = count_combformula phi
+ | count_combformula (AConn (_, phis)) = fold count_combformula phis
+ | count_combformula (AAtom tm) = count_combterm tm
+fun count_prepared_formula ({combformula, ...} : prepared_formula) =
+ count_combformula combformula
+
+val optional_helpers =
+ [(["c_COMBI"], @{thms Meson.COMBI_def}),
+ (["c_COMBK"], @{thms Meson.COMBK_def}),
+ (["c_COMBB"], @{thms Meson.COMBB_def}),
+ (["c_COMBC"], @{thms Meson.COMBC_def}),
+ (["c_COMBS"], @{thms Meson.COMBS_def})]
+val optional_typed_helpers =
+ [(["c_True", "c_False", "c_If"], @{thms True_or_False}),
+ (["c_If"], @{thms if_True if_False})]
+val mandatory_helpers = @{thms Metis.fequal_def}
+
+val init_counters =
+ [optional_helpers, optional_typed_helpers] |> maps (maps fst)
+ |> sort_distinct string_ord |> map (rpair 0) |> Symtab.make
+
+fun get_helper_facts ctxt is_FO full_types conjectures axioms =
+ let
+ val ct =
+ fold (fold count_prepared_formula) [conjectures, axioms] init_counters
+ fun is_needed c = the (Symtab.lookup ct c) > 0
+ fun baptize th = ((Thm.get_name_hint th, false), th)
+ in
+ (optional_helpers
+ |> full_types ? append optional_typed_helpers
+ |> maps (fn (ss, ths) =>
+ if exists is_needed ss then map baptize ths else [])) @
+ (if is_FO then [] else map baptize mandatory_helpers)
+ |> map_filter (Option.map snd o make_axiom ctxt false)
+ end
+
+fun prepare_axiom ctxt (ax as (_, th)) = (prop_of th, make_axiom ctxt true ax)
+
+fun prepare_formulas ctxt full_types hyp_ts concl_t axioms =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val (axiom_ts, prepared_axioms) = ListPair.unzip axioms
+ (* Remove existing axioms from the conjecture, as this can dramatically
+ boost an ATP's performance (for some reason). *)
+ val hyp_ts = hyp_ts |> filter_out (member (op aconv) axiom_ts)
+ val goal_t = Logic.list_implies (hyp_ts, concl_t)
+ val is_FO = Meson.is_fol_term thy goal_t
+ val subs = tfree_classes_of_terms [goal_t]
+ val supers = tvar_classes_of_terms axiom_ts
+ val tycons = type_consts_of_terms thy (goal_t :: axiom_ts)
+ (* TFrees in the conjecture; TVars in the axioms *)
+ val conjectures = make_conjecture ctxt (hyp_ts @ [concl_t])
+ val (axiom_names, axioms) = ListPair.unzip (map_filter I prepared_axioms)
+ val helper_facts = get_helper_facts ctxt is_FO full_types conjectures axioms
+ val (supers', arity_clauses) = make_arity_clauses thy tycons supers
+ val class_rel_clauses = make_class_rel_clauses thy subs supers'
+ in
+ (axiom_names |> map single |> Vector.fromList,
+ (conjectures, axioms, helper_facts, class_rel_clauses, arity_clauses))
+ end
+
+fun wrap_type ty t = ATerm ((type_wrapper_name, type_wrapper_name), [ty, t])
+
+fun fo_term_for_combtyp (CombTVar name) = ATerm (name, [])
+ | fo_term_for_combtyp (CombTFree name) = ATerm (name, [])
+ | fo_term_for_combtyp (CombType (name, tys)) =
+ ATerm (name, map fo_term_for_combtyp tys)
+
+fun fo_literal_for_type_literal (TyLitVar (class, name)) =
+ (true, ATerm (class, [ATerm (name, [])]))
+ | fo_literal_for_type_literal (TyLitFree (class, name)) =
+ (true, ATerm (class, [ATerm (name, [])]))
+
+fun formula_for_fo_literal (pos, t) = AAtom t |> not pos ? mk_anot
+
+fun fo_term_for_combterm full_types =
+ let
+ fun aux top_level u =
+ let
+ val (head, args) = strip_combterm_comb u
+ val (x, ty_args) =
+ case head of
+ CombConst (name as (s, s'), _, ty_args) =>
+ let val ty_args = if full_types then [] else ty_args in
+ if s = "equal" then
+ if top_level andalso length args = 2 then (name, [])
+ else (("c_fequal", @{const_name Metis.fequal}), ty_args)
+ else if top_level then
+ case s of
+ "c_False" => (("$false", s'), [])
+ | "c_True" => (("$true", s'), [])
+ | _ => (name, ty_args)
+ else
+ (name, ty_args)
+ end
+ | CombVar (name, _) => (name, [])
+ | CombApp _ => raise Fail "impossible \"CombApp\""
+ val t = ATerm (x, map fo_term_for_combtyp ty_args @
+ map (aux false) args)
+ in
+ if full_types then wrap_type (fo_term_for_combtyp (combtyp_of u)) t else t
+ end
+ in aux true end
+
+fun formula_for_combformula full_types =
+ let
+ fun aux (AQuant (q, xs, phi)) = AQuant (q, xs, aux phi)
+ | aux (AConn (c, phis)) = AConn (c, map aux phis)
+ | aux (AAtom tm) = AAtom (fo_term_for_combterm full_types tm)
+ in aux end
+
+fun formula_for_axiom full_types
+ ({combformula, ctypes_sorts, ...} : prepared_formula) =
+ mk_ahorn (map (formula_for_fo_literal o fo_literal_for_type_literal)
+ (type_literals_for_types ctypes_sorts))
+ (formula_for_combformula full_types combformula)
+
+fun problem_line_for_fact prefix full_types (formula as {name, kind, ...}) =
+ Fof (prefix ^ ascii_of name, kind, formula_for_axiom full_types formula)
+
+fun problem_line_for_class_rel_clause (ClassRelClause {name, subclass,
+ superclass, ...}) =
+ let val ty_arg = ATerm (("T", "T"), []) in
+ Fof (class_rel_clause_prefix ^ ascii_of name, Axiom,
+ AConn (AImplies, [AAtom (ATerm (subclass, [ty_arg])),
+ AAtom (ATerm (superclass, [ty_arg]))]))
+ end
+
+fun fo_literal_for_arity_literal (TConsLit (c, t, args)) =
+ (true, ATerm (c, [ATerm (t, map (fn arg => ATerm (arg, [])) args)]))
+ | fo_literal_for_arity_literal (TVarLit (c, sort)) =
+ (false, ATerm (c, [ATerm (sort, [])]))
+
+fun problem_line_for_arity_clause (ArityClause {name, conclLit, premLits,
+ ...}) =
+ Fof (arity_clause_prefix ^ ascii_of name, Axiom,
+ mk_ahorn (map (formula_for_fo_literal o apfst not
+ o fo_literal_for_arity_literal) premLits)
+ (formula_for_fo_literal
+ (fo_literal_for_arity_literal conclLit)))
+
+fun problem_line_for_conjecture full_types
+ ({name, kind, combformula, ...} : prepared_formula) =
+ Fof (conjecture_prefix ^ name, kind,
+ formula_for_combformula full_types combformula)
+
+fun free_type_literals_for_conjecture ({ctypes_sorts, ...} : prepared_formula) =
+ map fo_literal_for_type_literal (type_literals_for_types ctypes_sorts)
+
+fun problem_line_for_free_type j lit =
+ Fof (tfree_prefix ^ string_of_int j, Hypothesis, formula_for_fo_literal lit)
+fun problem_lines_for_free_types conjectures =
+ let
+ val litss = map free_type_literals_for_conjecture conjectures
+ val lits = fold (union (op =)) litss []
+ in map2 problem_line_for_free_type (0 upto length lits - 1) lits end
+
+(** "hBOOL" and "hAPP" **)
+
+type const_info = {min_arity: int, max_arity: int, sub_level: bool}
+
+fun consider_term top_level (ATerm ((s, _), ts)) =
+ (if is_atp_variable s then
+ I
+ else
+ let val n = length ts in
+ Symtab.map_default
+ (s, {min_arity = n, max_arity = 0, sub_level = false})
+ (fn {min_arity, max_arity, sub_level} =>
+ {min_arity = Int.min (n, min_arity),
+ max_arity = Int.max (n, max_arity),
+ sub_level = sub_level orelse not top_level})
+ end)
+ #> fold (consider_term (top_level andalso s = type_wrapper_name)) ts
+fun consider_formula (AQuant (_, _, phi)) = consider_formula phi
+ | consider_formula (AConn (_, phis)) = fold consider_formula phis
+ | consider_formula (AAtom tm) = consider_term true tm
+
+fun consider_problem_line (Fof (_, _, phi)) = consider_formula phi
+fun consider_problem problem = fold (fold consider_problem_line o snd) problem
+
+fun const_table_for_problem explicit_apply problem =
+ if explicit_apply then NONE
+ else SOME (Symtab.empty |> consider_problem problem)
+
+fun min_arity_of thy full_types NONE s =
+ (if s = "equal" orelse s = type_wrapper_name orelse
+ String.isPrefix type_const_prefix s orelse
+ String.isPrefix class_prefix s then
+ 16383 (* large number *)
+ else if full_types then
+ 0
+ else case strip_prefix_and_unascii const_prefix s of
+ SOME s' => num_type_args thy (invert_const s')
+ | NONE => 0)
+ | min_arity_of _ _ (SOME the_const_tab) s =
+ case Symtab.lookup the_const_tab s of
+ SOME ({min_arity, ...} : const_info) => min_arity
+ | NONE => 0
+
+fun full_type_of (ATerm ((s, _), [ty, _])) =
+ if s = type_wrapper_name then ty else raise Fail "expected type wrapper"
+ | full_type_of _ = raise Fail "expected type wrapper"
+
+fun list_hAPP_rev _ t1 [] = t1
+ | list_hAPP_rev NONE t1 (t2 :: ts2) =
+ ATerm (`I "hAPP", [list_hAPP_rev NONE t1 ts2, t2])
+ | list_hAPP_rev (SOME ty) t1 (t2 :: ts2) =
+ let val ty' = ATerm (`make_fixed_type_const @{type_name fun},
+ [full_type_of t2, ty]) in
+ ATerm (`I "hAPP", [wrap_type ty' (list_hAPP_rev (SOME ty') t1 ts2), t2])
+ end
+
+fun repair_applications_in_term thy full_types const_tab =
+ let
+ fun aux opt_ty (ATerm (name as (s, _), ts)) =
+ if s = type_wrapper_name then
+ case ts of
+ [t1, t2] => ATerm (name, [aux NONE t1, aux (SOME t1) t2])
+ | _ => raise Fail "malformed type wrapper"
+ else
+ let
+ val ts = map (aux NONE) ts
+ val (ts1, ts2) = chop (min_arity_of thy full_types const_tab s) ts
+ in list_hAPP_rev opt_ty (ATerm (name, ts1)) (rev ts2) end
+ in aux NONE end
+
+fun boolify t = ATerm (`I "hBOOL", [t])
+
+(* True if the constant ever appears outside of the top-level position in
+ literals, or if it appears with different arities (e.g., because of different
+ type instantiations). If false, the constant always receives all of its
+ arguments and is used as a predicate. *)
+fun is_predicate NONE s =
+ s = "equal" orelse s = "$false" orelse s = "$true" orelse
+ String.isPrefix type_const_prefix s orelse String.isPrefix class_prefix s
+ | is_predicate (SOME the_const_tab) s =
+ case Symtab.lookup the_const_tab s of
+ SOME {min_arity, max_arity, sub_level} =>
+ not sub_level andalso min_arity = max_arity
+ | NONE => false
+
+fun repair_predicates_in_term const_tab (t as ATerm ((s, _), ts)) =
+ if s = type_wrapper_name then
+ case ts of
+ [_, t' as ATerm ((s', _), _)] =>
+ if is_predicate const_tab s' then t' else boolify t
+ | _ => raise Fail "malformed type wrapper"
+ else
+ t |> not (is_predicate const_tab s) ? boolify
+
+fun close_universally phi =
+ let
+ fun term_vars bounds (ATerm (name as (s, _), tms)) =
+ (is_atp_variable s andalso not (member (op =) bounds name))
+ ? insert (op =) name
+ #> fold (term_vars bounds) tms
+ fun formula_vars bounds (AQuant (_, xs, phi)) =
+ formula_vars (xs @ bounds) phi
+ | formula_vars bounds (AConn (_, phis)) = fold (formula_vars bounds) phis
+ | formula_vars bounds (AAtom tm) = term_vars bounds tm
+ in
+ case formula_vars [] phi [] of [] => phi | xs => AQuant (AForall, xs, phi)
+ end
+
+fun repair_formula thy explicit_forall full_types const_tab =
+ let
+ fun aux (AQuant (q, xs, phi)) = AQuant (q, xs, aux phi)
+ | aux (AConn (c, phis)) = AConn (c, map aux phis)
+ | aux (AAtom tm) =
+ AAtom (tm |> repair_applications_in_term thy full_types const_tab
+ |> repair_predicates_in_term const_tab)
+ in aux #> explicit_forall ? close_universally end
+
+fun repair_problem_line thy explicit_forall full_types const_tab
+ (Fof (ident, kind, phi)) =
+ Fof (ident, kind, repair_formula thy explicit_forall full_types const_tab phi)
+fun repair_problem_with_const_table thy =
+ map o apsnd o map ooo repair_problem_line thy
+
+fun repair_problem thy explicit_forall full_types explicit_apply problem =
+ repair_problem_with_const_table thy explicit_forall full_types
+ (const_table_for_problem explicit_apply problem) problem
+
+fun prepare_atp_problem ctxt readable_names explicit_forall full_types
+ explicit_apply hyp_ts concl_t axioms =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val (axiom_names, (conjectures, axioms, helper_facts, class_rel_clauses,
+ arity_clauses)) =
+ prepare_formulas ctxt full_types hyp_ts concl_t axioms
+ val axiom_lines = map (problem_line_for_fact axiom_prefix full_types) axioms
+ val helper_lines =
+ map (problem_line_for_fact helper_prefix full_types) helper_facts
+ val conjecture_lines =
+ map (problem_line_for_conjecture full_types) conjectures
+ val tfree_lines = problem_lines_for_free_types conjectures
+ val class_rel_lines =
+ map problem_line_for_class_rel_clause class_rel_clauses
+ val arity_lines = map problem_line_for_arity_clause arity_clauses
+ (* Reordering these might or might not confuse the proof reconstruction
+ code or the SPASS Flotter hack. *)
+ val problem =
+ [("Relevant facts", axiom_lines),
+ ("Class relationships", class_rel_lines),
+ ("Arity declarations", arity_lines),
+ ("Helper facts", helper_lines),
+ ("Conjectures", conjecture_lines),
+ ("Type variables", tfree_lines)]
+ |> repair_problem thy explicit_forall full_types explicit_apply
+ val (problem, pool) = nice_atp_problem readable_names problem
+ val conjecture_offset =
+ length axiom_lines + length class_rel_lines + length arity_lines
+ + length helper_lines
+ in
+ (problem,
+ case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty,
+ conjecture_offset, axiom_names)
+ end
+
+end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_filter.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_filter.ML Sun Oct 24 15:11:24 2010 -0700
@@ -9,35 +9,38 @@
sig
datatype locality = General | Intro | Elim | Simp | Local | Assum | Chained
+ type relevance_fudge =
+ {worse_irrel_freq : real,
+ higher_order_irrel_weight : real,
+ abs_rel_weight : real,
+ abs_irrel_weight : real,
+ skolem_irrel_weight : real,
+ theory_const_rel_weight : real,
+ theory_const_irrel_weight : real,
+ intro_bonus : real,
+ elim_bonus : real,
+ simp_bonus : real,
+ local_bonus : real,
+ assum_bonus : real,
+ chained_bonus : real,
+ max_imperfect : real,
+ max_imperfect_exp : real,
+ threshold_divisor : real,
+ ridiculous_threshold : real}
+
type relevance_override =
- {add: (Facts.ref * Attrib.src list) list,
- del: (Facts.ref * Attrib.src list) list,
- only: bool}
+ {add : (Facts.ref * Attrib.src list) list,
+ del : (Facts.ref * Attrib.src list) list,
+ only : bool}
val trace : bool Unsynchronized.ref
- val worse_irrel_freq : real Unsynchronized.ref
- val higher_order_irrel_weight : real Unsynchronized.ref
- val abs_rel_weight : real Unsynchronized.ref
- val abs_irrel_weight : real Unsynchronized.ref
- val skolem_irrel_weight : real Unsynchronized.ref
- val theory_const_rel_weight : real Unsynchronized.ref
- val theory_const_irrel_weight : real Unsynchronized.ref
- val intro_bonus : real Unsynchronized.ref
- val elim_bonus : real Unsynchronized.ref
- val simp_bonus : real Unsynchronized.ref
- val local_bonus : real Unsynchronized.ref
- val assum_bonus : real Unsynchronized.ref
- val chained_bonus : real Unsynchronized.ref
- val max_imperfect : real Unsynchronized.ref
- val max_imperfect_exp : real Unsynchronized.ref
- val threshold_divisor : real Unsynchronized.ref
- val ridiculous_threshold : real Unsynchronized.ref
val name_thm_pairs_from_ref :
Proof.context -> unit Symtab.table -> thm list
-> Facts.ref * Attrib.src list -> ((string * locality) * thm) list
val relevant_facts :
- Proof.context -> bool -> real * real -> int -> relevance_override
- -> thm list -> term list -> term -> ((string * locality) * thm) list
+ Proof.context -> bool -> real * real -> int -> string list
+ -> relevance_fudge -> relevance_override -> thm list -> term list -> term
+ -> ((string * locality) * thm) list
end;
structure Sledgehammer_Filter : SLEDGEHAMMER_FILTER =
@@ -52,31 +55,31 @@
val term_patterns = false
val respect_no_atp = true
-(* FUDGE *)
-val worse_irrel_freq = Unsynchronized.ref 100.0
-val higher_order_irrel_weight = Unsynchronized.ref 1.05
-val abs_rel_weight = Unsynchronized.ref 0.5
-val abs_irrel_weight = Unsynchronized.ref 2.0
-val skolem_irrel_weight = Unsynchronized.ref 0.75
-val theory_const_rel_weight = Unsynchronized.ref 0.5
-val theory_const_irrel_weight = Unsynchronized.ref 0.25
-val intro_bonus = Unsynchronized.ref 0.15
-val elim_bonus = Unsynchronized.ref 0.15
-val simp_bonus = Unsynchronized.ref 0.15
-val local_bonus = Unsynchronized.ref 0.55
-val assum_bonus = Unsynchronized.ref 1.05
-val chained_bonus = Unsynchronized.ref 1.5
-val max_imperfect = Unsynchronized.ref 11.5
-val max_imperfect_exp = Unsynchronized.ref 1.0
-val threshold_divisor = Unsynchronized.ref 2.0
-val ridiculous_threshold = Unsynchronized.ref 0.1
-
datatype locality = General | Intro | Elim | Simp | Local | Assum | Chained
+type relevance_fudge =
+ {worse_irrel_freq : real,
+ higher_order_irrel_weight : real,
+ abs_rel_weight : real,
+ abs_irrel_weight : real,
+ skolem_irrel_weight : real,
+ theory_const_rel_weight : real,
+ theory_const_irrel_weight : real,
+ intro_bonus : real,
+ elim_bonus : real,
+ simp_bonus : real,
+ local_bonus : real,
+ assum_bonus : real,
+ chained_bonus : real,
+ max_imperfect : real,
+ max_imperfect_exp : real,
+ threshold_divisor : real,
+ ridiculous_threshold : real}
+
type relevance_override =
- {add: (Facts.ref * Attrib.src list) list,
- del: (Facts.ref * Attrib.src list) list,
- only: bool}
+ {add : (Facts.ref * Attrib.src list) list,
+ del : (Facts.ref * Attrib.src list) list,
+ only : bool}
val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator
val abs_name = sledgehammer_prefix ^ "abs"
@@ -172,16 +175,10 @@
fun string_for_hyper_pconst (s, ps) =
s ^ "{" ^ commas (map string_for_ptype ps) ^ "}"
-(* These are typically simplified away by "Meson.presimplify". Equality is
- handled specially via "fequal". *)
-val boring_consts =
- [@{const_name False}, @{const_name True}, @{const_name If}, @{const_name Let},
- @{const_name HOL.eq}]
-
(* Add a pconstant to the table, but a [] entry means a standard
connective, which we ignore.*)
-fun add_pconst_to_table also_skolem (c, p) =
- if member (op =) boring_consts c orelse
+fun add_pconst_to_table irrelevant_consts also_skolem (c, p) =
+ if member (op =) irrelevant_consts c orelse
(not also_skolem andalso String.isPrefix skolem_prefix c) then
I
else
@@ -189,14 +186,15 @@
fun is_formula_type T = (T = HOLogic.boolT orelse T = propT)
-fun pconsts_in_terms thy also_skolems pos ts =
+fun pconsts_in_terms thy irrelevant_consts also_skolems pos ts =
let
val flip = Option.map not
(* We include free variables, as well as constants, to handle locales. For
each quantifiers that must necessarily be skolemized by the ATP, we
introduce a fresh constant to simulate the effect of Skolemization. *)
fun do_const const (s, T) ts =
- add_pconst_to_table also_skolems (rich_pconst thy const (s, T) ts)
+ add_pconst_to_table irrelevant_consts also_skolems
+ (rich_pconst thy const (s, T) ts)
#> fold do_term ts
and do_term t =
case strip_comb t of
@@ -204,13 +202,14 @@
| (Free x, ts) => do_const false x ts
| (Abs (_, T, t'), ts) =>
(null ts
- ? add_pconst_to_table true (abs_name, PType (order_of_type T + 1, [])))
+ ? add_pconst_to_table irrelevant_consts true
+ (abs_name, PType (order_of_type T + 1, [])))
#> fold do_term (t' :: ts)
| (_, ts) => fold do_term ts
fun do_quantifier will_surely_be_skolemized abs_T body_t =
do_formula pos body_t
#> (if also_skolems andalso will_surely_be_skolemized then
- add_pconst_to_table true
+ add_pconst_to_table irrelevant_consts true
(gensym skolem_prefix, PType (order_of_type abs_T, []))
else
I)
@@ -254,9 +253,11 @@
(*Inserts a dummy "constant" referring to the theory name, so that relevance
takes the given theory into account.*)
-fun theory_const_prop_of th =
- if exists (curry (op <) 0.0) [!theory_const_rel_weight,
- !theory_const_irrel_weight] then
+fun theory_const_prop_of ({theory_const_rel_weight,
+ theory_const_irrel_weight, ...} : relevance_fudge)
+ th =
+ if exists (curry (op <) 0.0) [theory_const_rel_weight,
+ theory_const_irrel_weight] then
let
val name = Context.theory_name (theory_of_thm th)
val t = Const (name ^ theory_const_suffix, @{typ bool})
@@ -282,7 +283,7 @@
structure PType_Tab = Table(type key = ptype val ord = ptype_ord)
-fun count_axiom_consts thy =
+fun count_axiom_consts thy fudge =
let
fun do_const const (s, T) ts =
(* Two-dimensional table update. Constant maps to types maps to count. *)
@@ -295,7 +296,7 @@
| (Free x, ts) => do_const false x ts
| (Abs (_, _, t'), ts) => fold do_term (t' :: ts)
| (_, ts) => fold do_term ts
- in do_term o theory_const_prop_of o snd end
+ in do_term o theory_const_prop_of fudge o snd end
(**** Actual Filtering Code ****)
@@ -322,11 +323,12 @@
very rare constants and very common ones -- the former because they can't
lead to the inclusion of too many new facts, and the latter because they are
so common as to be of little interest. *)
-fun irrel_weight_for order freq =
- let val (k, x) = !worse_irrel_freq |> `Real.ceil in
+fun irrel_weight_for ({worse_irrel_freq, higher_order_irrel_weight, ...}
+ : relevance_fudge) order freq =
+ let val (k, x) = worse_irrel_freq |> `Real.ceil in
(if freq < k then Math.ln (Real.fromInt (freq + 1)) / Math.ln x
else rel_weight_for order freq / rel_weight_for order k)
- * pow_int (!higher_order_irrel_weight) (order - 1)
+ * pow_int higher_order_irrel_weight (order - 1)
end
(* Computes a constant's weight, as determined by its frequency. *)
@@ -337,22 +339,25 @@
else if String.isSuffix theory_const_suffix s then theory_const_weight
else weight_for m (pconst_freq (match_ptype o f) const_tab c)
-fun rel_pconst_weight const_tab =
- generic_pconst_weight (!abs_rel_weight) 0.0 (!theory_const_rel_weight)
+fun rel_pconst_weight ({abs_rel_weight, theory_const_rel_weight, ...}
+ : relevance_fudge) const_tab =
+ generic_pconst_weight abs_rel_weight 0.0 theory_const_rel_weight
rel_weight_for I const_tab
-fun irrel_pconst_weight const_tab =
- generic_pconst_weight (!abs_irrel_weight) (!skolem_irrel_weight)
- (!theory_const_irrel_weight) irrel_weight_for swap const_tab
+fun irrel_pconst_weight (fudge as {abs_irrel_weight, skolem_irrel_weight,
+ theory_const_irrel_weight, ...}) const_tab =
+ generic_pconst_weight abs_irrel_weight skolem_irrel_weight
+ theory_const_irrel_weight (irrel_weight_for fudge) swap
+ const_tab
-fun locality_bonus General = 0.0
- | locality_bonus Intro = !intro_bonus
- | locality_bonus Elim = !elim_bonus
- | locality_bonus Simp = !simp_bonus
- | locality_bonus Local = !local_bonus
- | locality_bonus Assum = !assum_bonus
- | locality_bonus Chained = !chained_bonus
+fun locality_bonus (_ : relevance_fudge) General = 0.0
+ | locality_bonus {intro_bonus, ...} Intro = intro_bonus
+ | locality_bonus {elim_bonus, ...} Elim = elim_bonus
+ | locality_bonus {simp_bonus, ...} Simp = simp_bonus
+ | locality_bonus {local_bonus, ...} Local = local_bonus
+ | locality_bonus {assum_bonus, ...} Assum = assum_bonus
+ | locality_bonus {chained_bonus, ...} Chained = chained_bonus
-fun axiom_weight loc const_tab relevant_consts axiom_consts =
+fun axiom_weight fudge loc const_tab relevant_consts axiom_consts =
case axiom_consts |> List.partition (pconst_hyper_mem I relevant_consts)
||> filter_out (pconst_hyper_mem swap relevant_consts) of
([], _) => 0.0
@@ -360,15 +365,15 @@
let
val irrel = irrel |> filter_out (pconst_mem swap rel)
val rel_weight =
- 0.0 |> fold (curry (op +) o rel_pconst_weight const_tab) rel
+ 0.0 |> fold (curry (op +) o rel_pconst_weight fudge const_tab) rel
val irrel_weight =
- ~ (locality_bonus loc)
- |> fold (curry (op +) o irrel_pconst_weight const_tab) irrel
+ ~ (locality_bonus fudge loc)
+ |> fold (curry (op +) o irrel_pconst_weight fudge const_tab) irrel
val res = rel_weight / (rel_weight + irrel_weight)
in if Real.isFinite res then res else 0.0 end
(* FIXME: experiment
-fun debug_axiom_weight loc const_tab relevant_consts axiom_consts =
+fun debug_axiom_weight fudge loc const_tab relevant_consts axiom_consts =
case axiom_consts |> List.partition (pconst_hyper_mem I relevant_consts)
||> filter_out (pconst_hyper_mem swap relevant_consts) of
([], _) => 0.0
@@ -378,19 +383,20 @@
val rels_weight =
0.0 |> fold (curry (op +) o rel_pconst_weight const_tab) rel
val irrels_weight =
- ~ (locality_bonus loc)
- |> fold (curry (op +) o irrel_pconst_weight const_tab) irrel
+ ~ (locality_bonus fudge loc)
+ |> fold (curry (op +) o irrel_pconst_weight fudge const_tab) irrel
val _ = tracing (PolyML.makestring ("REL: ", map (`(rel_pconst_weight const_tab)) rel))
-val _ = tracing (PolyML.makestring ("IRREL: ", map (`(irrel_pconst_weight const_tab)) irrel))
+val _ = tracing (PolyML.makestring ("IRREL: ", map (`(irrel_pconst_weight fudge const_tab)) irrel))
val res = rels_weight / (rels_weight + irrels_weight)
in if Real.isFinite res then res else 0.0 end
*)
-fun pconsts_in_axiom thy t =
+fun pconsts_in_axiom thy irrelevant_consts t =
Symtab.fold (fn (s, pss) => fold (cons o pair s) pss)
- (pconsts_in_terms thy true (SOME true) [t]) []
-fun pair_consts_axiom thy axiom =
- case axiom |> snd |> theory_const_prop_of |> pconsts_in_axiom thy of
+ (pconsts_in_terms thy irrelevant_consts true (SOME true) [t]) []
+fun pair_consts_axiom thy irrelevant_consts fudge axiom =
+ case axiom |> snd |> theory_const_prop_of fudge
+ |> pconsts_in_axiom thy irrelevant_consts of
[] => NONE
| consts => SOME ((axiom, consts), NONE)
@@ -398,12 +404,13 @@
(((unit -> string) * locality) * thm) * (string * ptype) list
fun take_most_relevant max_relevant remaining_max
- (candidates : (annotated_thm * real) list) =
+ ({max_imperfect, max_imperfect_exp, ...} : relevance_fudge)
+ (candidates : (annotated_thm * real) list) =
let
val max_imperfect =
- Real.ceil (Math.pow (!max_imperfect,
+ Real.ceil (Math.pow (max_imperfect,
Math.pow (Real.fromInt remaining_max
- / Real.fromInt max_relevant, !max_imperfect_exp)))
+ / Real.fromInt max_relevant, max_imperfect_exp)))
val (perfect, imperfect) =
candidates |> sort (Real.compare o swap o pairself snd)
|> take_prefix (fn (_, w) => w > 0.99999)
@@ -419,22 +426,23 @@
(accepts, more_rejects @ rejects)
end
-fun if_empty_replace_with_locality thy axioms loc tab =
+fun if_empty_replace_with_locality thy irrelevant_consts axioms loc tab =
if Symtab.is_empty tab then
- pconsts_in_terms thy false (SOME false)
+ pconsts_in_terms thy irrelevant_consts false (SOME false)
(map_filter (fn ((_, loc'), th) =>
if loc' = loc then SOME (prop_of th) else NONE) axioms)
else
tab
-fun relevance_filter ctxt threshold0 decay max_relevant
- ({add, del, ...} : relevance_override) axioms goal_ts =
+fun relevance_filter ctxt threshold0 decay max_relevant irrelevant_consts
+ (fudge as {threshold_divisor, ridiculous_threshold, ...})
+ ({add, del, ...} : relevance_override) axioms goal_ts =
let
val thy = ProofContext.theory_of ctxt
- val const_tab = fold (count_axiom_consts thy) axioms Symtab.empty
+ val const_tab = fold (count_axiom_consts thy fudge) axioms Symtab.empty
val goal_const_tab =
- pconsts_in_terms thy false (SOME false) goal_ts
- |> fold (if_empty_replace_with_locality thy axioms)
+ pconsts_in_terms thy irrelevant_consts false (SOME false) goal_ts
+ |> fold (if_empty_replace_with_locality thy irrelevant_consts axioms)
[Chained, Assum, Local]
val add_ths = Attrib.eval_thms ctxt add
val del_ths = Attrib.eval_thms ctxt del
@@ -450,19 +458,20 @@
else NONE) rejects
fun relevant [] rejects [] =
(* Nothing has been added this iteration. *)
- if j = 0 andalso threshold >= !ridiculous_threshold then
+ if j = 0 andalso threshold >= ridiculous_threshold then
(* First iteration? Try again. *)
- iter 0 max_relevant (threshold / !threshold_divisor) rel_const_tab
+ iter 0 max_relevant (threshold / threshold_divisor) rel_const_tab
hopeless hopeful
else
game_over (rejects @ hopeless)
| relevant candidates rejects [] =
let
val (accepts, more_rejects) =
- take_most_relevant max_relevant remaining_max candidates
+ take_most_relevant max_relevant remaining_max fudge candidates
val rel_const_tab' =
rel_const_tab
- |> fold (add_pconst_to_table false) (maps (snd o fst) accepts)
+ |> fold (add_pconst_to_table irrelevant_consts false)
+ (maps (snd o fst) accepts)
fun is_dirty (c, _) =
Symtab.lookup rel_const_tab' c <> Symtab.lookup rel_const_tab c
val (hopeful_rejects, hopeless_rejects) =
@@ -499,11 +508,12 @@
val weight =
case cached_weight of
SOME w => w
- | NONE => axiom_weight loc const_tab rel_const_tab axiom_consts
+ | NONE => axiom_weight fudge loc const_tab rel_const_tab
+ axiom_consts
(* FIXME: experiment
val name = fst (fst (fst ax)) ()
val _ = if String.isSubstring "positive_minus" name orelse String.isSubstring "not_exp_le_zero" name then
-tracing ("*** " ^ name ^ PolyML.makestring (debug_axiom_weight loc const_tab rel_const_tab axiom_consts))
+tracing ("*** " ^ name ^ PolyML.makestring (debug_axiom_weight fudge loc const_tab rel_const_tab axiom_consts))
else
()
*)
@@ -524,7 +534,7 @@
end
in
axioms |> filter_out (member Thm.eq_thm del_ths o snd)
- |> map_filter (pair_consts_axiom thy)
+ |> map_filter (pair_consts_axiom thy irrelevant_consts fudge)
|> iter 0 max_relevant threshold0 goal_const_tab []
|> tap (fn res => trace_msg (fn () =>
"Total relevant: " ^ Int.toString (length res)))
@@ -708,7 +718,9 @@
(Term.add_vars t [] |> sort_wrt (fst o fst))
|> fst
-fun all_name_thms_pairs ctxt reserved full_types add_ths chained_ths =
+fun all_name_thms_pairs ctxt reserved full_types
+ ({intro_bonus, elim_bonus, simp_bonus, ...} : relevance_fudge) add_ths
+ chained_ths =
let
val thy = ProofContext.theory_of ctxt
val global_facts = Global_Theory.facts_of thy
@@ -718,7 +730,7 @@
fun is_assum th = exists (fn ct => prop_of th aconv term_of ct) assms
val is_chained = member Thm.eq_thm chained_ths
val (intros, elims, simps) =
- if exists (curry (op <) 0.0) [!intro_bonus, !elim_bonus, !simp_bonus] then
+ if exists (curry (op <) 0.0) [intro_bonus, elim_bonus, simp_bonus] then
clasimpset_rules_of ctxt
else
(Termtab.empty, Termtab.empty, Termtab.empty)
@@ -773,7 +785,8 @@
| NONE => ""
end),
let val t = prop_of th in
- if is_chained th then Chained
+ if is_chained th then
+ Chained
else if global then
if Termtab.defined intros t then Intro
else if Termtab.defined elims t then Elim
@@ -801,8 +814,8 @@
(***************************************************************)
fun relevant_facts ctxt full_types (threshold0, threshold1) max_relevant
- (relevance_override as {add, only, ...}) chained_ths hyp_ts
- concl_t =
+ irrelevant_consts fudge (override as {add, only, ...})
+ chained_ths hyp_ts concl_t =
let
val decay = Math.pow ((1.0 - threshold1) / (1.0 - threshold0),
1.0 / Real.fromInt (max_relevant + 1))
@@ -813,7 +826,7 @@
maps (map (fn ((name, loc), th) => ((K name, loc), (true, th)))
o name_thm_pairs_from_ref ctxt reserved chained_ths) add
else
- all_name_thms_pairs ctxt reserved full_types add_ths chained_ths)
+ all_name_thms_pairs ctxt reserved full_types fudge add_ths chained_ths)
|> name_thm_pairs ctxt (respect_no_atp andalso not only)
|> uniquify
in
@@ -825,8 +838,8 @@
max_relevant = 0 then
[]
else
- relevance_filter ctxt threshold0 decay max_relevant relevance_override
- axioms (concl_t :: hyp_ts))
+ relevance_filter ctxt threshold0 decay max_relevant irrelevant_consts
+ fudge override axioms (concl_t :: hyp_ts))
|> map (apfst (apfst (fn f => f ())))
end
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML Sun Oct 24 15:11:24 2010 -0700
@@ -9,10 +9,10 @@
type params = Sledgehammer.params
val auto : bool Unsynchronized.ref
- val atps : string Unsynchronized.ref
+ val provers : string Unsynchronized.ref
val timeout : int Unsynchronized.ref
val full_types : bool Unsynchronized.ref
- val default_params : theory -> (string * string) list -> params
+ val default_params : Proof.context -> (string * string) list -> params
val setup : theory -> theory
end;
@@ -49,14 +49,14 @@
(*** parameters ***)
-val atps = Unsynchronized.ref ""
+val provers = Unsynchronized.ref ""
val timeout = Unsynchronized.ref 30
val full_types = Unsynchronized.ref false
val _ =
ProofGeneralPgip.add_preference Preferences.category_proof
- (Preferences.string_pref atps
- "Sledgehammer: ATPs"
+ (Preferences.string_pref provers
+ "Sledgehammer: Provers"
"Default automatic provers (separated by whitespace)")
val _ =
@@ -84,7 +84,9 @@
("isar_shrink_factor", "1")]
val alias_params =
- [("atp", "atps")]
+ [("prover", "provers"),
+ ("atps", "provers"), (* FIXME: legacy *)
+ ("atp", "provers")] (* FIXME: legacy *)
val negated_alias_params =
[("non_blocking", "blocking"),
("no_debug", "debug"),
@@ -98,7 +100,7 @@
["debug", "verbose", "overlord", "full_types", "isar_proof",
"isar_shrink_factor", "timeout"]
-val property_dependent_params = ["atps", "full_types", "timeout"]
+val property_dependent_params = ["provers", "full_types", "timeout"]
fun is_known_raw_param s =
AList.defined (op =) default_default_params s orelse
@@ -128,16 +130,57 @@
val extend = I
fun merge p : T = AList.merge (op =) (K true) p)
+(* FIXME: dummy *)
+fun is_smt_solver_installed ctxt = true
+
+fun remotify_prover_if_available_and_not_already_remote thy name =
+ if String.isPrefix remote_prefix name then
+ SOME name
+ else
+ let val remote_name = remote_prefix ^ name in
+ if is_prover_available thy remote_name then SOME remote_name else NONE
+ end
+
+fun remotify_prover_if_not_installed ctxt name =
+ let val thy = ProofContext.theory_of ctxt in
+ if is_prover_available thy name andalso is_prover_installed ctxt name then
+ SOME name
+ else
+ remotify_prover_if_available_and_not_already_remote thy name
+ end
+
+fun avoid_too_many_local_threads _ _ [] = []
+ | avoid_too_many_local_threads thy 0 provers =
+ map_filter (remotify_prover_if_available_and_not_already_remote thy) provers
+ | avoid_too_many_local_threads thy n (prover :: provers) =
+ let val n = if String.isPrefix remote_prefix prover then n else n - 1 in
+ prover :: avoid_too_many_local_threads thy n provers
+ end
+
+(* The first ATP of the list is used by Auto Sledgehammer. Because of the low
+ timeout, it makes sense to put SPASS first. *)
+fun default_provers_param_value ctxt =
+ let val thy = ProofContext.theory_of ctxt in
+ [spassN, eN, vampireN, sine_eN (* FIXME: , smtN *)]
+ |> map_filter (remotify_prover_if_not_installed ctxt)
+ |> avoid_too_many_local_threads thy (Thread.numProcessors ())
+ |> space_implode " "
+ end
+
val set_default_raw_param = Data.map o AList.update (op =) o unalias_raw_param
-fun default_raw_params thy =
- Data.get thy
- |> fold (AList.default (op =))
- [("atps", [case !atps of "" => default_atps_param_value () | s => s]),
- ("full_types", [if !full_types then "true" else "false"]),
- ("timeout", let val timeout = !timeout in
- [if timeout <= 0 then "none"
- else string_of_int timeout ^ " s"]
- end)]
+fun default_raw_params ctxt =
+ let val thy = ProofContext.theory_of ctxt in
+ Data.get thy
+ |> fold (AList.default (op =))
+ [("provers", [case !provers of
+ "" => default_provers_param_value ctxt
+ | s => s]),
+ ("full_types", [if !full_types then "true" else "false"]),
+ ("timeout", let val timeout = !timeout in
+ [if timeout <= 0 then "none"
+ else string_of_int timeout ^ " s"]
+ end)]
+ end
val infinity_time_in_secs = 60 * 60 * 24 * 365
val the_timeout = the_default (Time.fromSeconds infinity_time_in_secs)
@@ -180,7 +223,8 @@
val debug = not auto andalso lookup_bool "debug"
val verbose = debug orelse (not auto andalso lookup_bool "verbose")
val overlord = lookup_bool "overlord"
- val atps = lookup_string "atps" |> space_explode " " |> auto ? single o hd
+ val provers = lookup_string "provers" |> space_explode " "
+ |> auto ? single o hd
val full_types = lookup_bool "full_types"
val explicit_apply = lookup_bool "explicit_apply"
val relevance_thresholds =
@@ -193,13 +237,14 @@
val expect = lookup_string "expect"
in
{blocking = blocking, debug = debug, verbose = verbose, overlord = overlord,
- atps = atps, full_types = full_types, explicit_apply = explicit_apply,
+ provers = provers, full_types = full_types,
+ explicit_apply = explicit_apply,
relevance_thresholds = relevance_thresholds, max_relevant = max_relevant,
isar_proof = isar_proof, isar_shrink_factor = isar_shrink_factor,
timeout = timeout, expect = expect}
end
-fun get_params auto thy = extract_params auto (default_raw_params thy)
+fun get_params auto ctxt = extract_params auto (default_raw_params ctxt)
fun default_params thy = get_params false thy o map (apsnd single)
(* Sledgehammer the given subgoal *)
@@ -210,9 +255,9 @@
val runN = "run"
val minimizeN = "minimize"
val messagesN = "messages"
-val available_atpsN = "available_atps"
-val running_atpsN = "running_atps"
-val kill_atpsN = "kill_atps"
+val available_proversN = "available_provers"
+val running_proversN = "running_provers"
+val kill_proversN = "kill_provers"
val refresh_tptpN = "refresh_tptp"
val is_raw_param_relevant_for_minimize =
@@ -220,8 +265,8 @@
fun string_for_raw_param (key, values) =
key ^ (case space_implode " " values of "" => "" | value => " = " ^ value)
-fun minimize_command override_params i atp_name fact_names =
- sledgehammerN ^ " " ^ minimizeN ^ " [atp = " ^ atp_name ^
+fun minimize_command override_params i prover_name fact_names =
+ sledgehammerN ^ " " ^ minimizeN ^ " [prover = " ^ prover_name ^
(override_params |> filter is_raw_param_relevant_for_minimize
|> implode o map (prefix ", " o string_for_raw_param)) ^
"] (" ^ space_implode " " fact_names ^ ")" ^
@@ -229,27 +274,28 @@
fun hammer_away override_params subcommand opt_i relevance_override state =
let
+ val ctxt = Proof.context_of state
val thy = Proof.theory_of state
val _ = app check_raw_param override_params
in
if subcommand = runN then
let val i = the_default 1 opt_i in
- run_sledgehammer (get_params false thy override_params) false i
+ run_sledgehammer (get_params false ctxt override_params) false i
relevance_override (minimize_command override_params i)
state
|> K ()
end
else if subcommand = minimizeN then
- run_minimize (get_params false thy override_params) (the_default 1 opt_i)
+ run_minimize (get_params false ctxt override_params) (the_default 1 opt_i)
(#add relevance_override) state
else if subcommand = messagesN then
messages opt_i
- else if subcommand = available_atpsN then
- available_atps thy
- else if subcommand = running_atpsN then
- running_atps ()
- else if subcommand = kill_atpsN then
- kill_atps ()
+ else if subcommand = available_proversN then
+ available_provers thy
+ else if subcommand = running_proversN then
+ running_provers ()
+ else if subcommand = kill_proversN then
+ kill_provers ()
else if subcommand = refresh_tptpN then
refresh_systems_on_tptp ()
else
@@ -266,13 +312,15 @@
Toplevel.theory
(fold set_default_raw_param params
#> tap (fn thy =>
- writeln ("Default parameters for Sledgehammer:\n" ^
- (case rev (default_raw_params thy) of
- [] => "none"
- | params =>
- (map check_raw_param params;
- params |> map string_for_raw_param
- |> sort_strings |> cat_lines)))))
+ let val ctxt = ProofContext.init_global thy in
+ writeln ("Default parameters for Sledgehammer:\n" ^
+ (case default_raw_params ctxt |> rev of
+ [] => "none"
+ | params =>
+ (map check_raw_param params;
+ params |> map string_for_raw_param
+ |> sort_strings |> cat_lines)))
+ end))
val parse_key = Scan.repeat1 Parse.typ_group >> space_implode " "
val parse_value = Scan.repeat1 Parse.xname
@@ -307,8 +355,8 @@
if not (!auto) then
(false, state)
else
- let val thy = Proof.theory_of state in
- run_sledgehammer (get_params true thy []) true 1 no_relevance_override
+ let val ctxt = Proof.context_of state in
+ run_sledgehammer (get_params true ctxt []) true 1 no_relevance_override
(minimize_command [] 1) state
end
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML Sun Oct 24 15:11:24 2010 -0700
@@ -2,7 +2,7 @@
Author: Philipp Meyer, TU Muenchen
Author: Jasmin Blanchette, TU Muenchen
-Minimization of theorem list for Metis using automatic theorem provers.
+Minimization of fact list for Metis using ATPs.
*)
signature SLEDGEHAMMER_MINIMIZE =
@@ -10,7 +10,7 @@
type locality = Sledgehammer_Filter.locality
type params = Sledgehammer.params
- val minimize_theorems :
+ val minimize_facts :
params -> int -> int -> Proof.state -> ((string * locality) * thm list) list
-> ((string * locality) * thm list) list option * string
val run_minimize :
@@ -23,8 +23,6 @@
open ATP_Proof
open Sledgehammer_Util
open Sledgehammer_Filter
-open Sledgehammer_Translate
-open Sledgehammer_Reconstruct
open Sledgehammer
(* wrapper for calling external prover *)
@@ -34,9 +32,9 @@
| string_for_failure Interrupted = "Interrupted."
| string_for_failure _ = "Unknown error."
-fun n_theorems names =
+fun n_facts names =
let val n = length names in
- string_of_int n ^ " theorem" ^ plural_s n ^
+ string_of_int n ^ " fact" ^ plural_s n ^
(if n > 0 then
": " ^ (names |> map fst
|> sort_distinct string_ord |> space_implode " ")
@@ -44,46 +42,44 @@
"")
end
-fun test_theorems ({debug, verbose, overlord, atps, full_types, isar_proof,
- isar_shrink_factor, ...} : params)
- (prover : prover) explicit_apply timeout subgoal state
- axioms =
+fun test_facts ({debug, verbose, overlord, provers, full_types, isar_proof,
+ isar_shrink_factor, ...} : params) (prover : prover)
+ explicit_apply timeout i n state axioms =
let
val _ =
- priority ("Testing " ^ n_theorems (map fst axioms) ^ "...")
+ priority ("Testing " ^ n_facts (map fst axioms) ^ "...")
val params =
{blocking = true, debug = debug, verbose = verbose, overlord = overlord,
- atps = atps, full_types = full_types, explicit_apply = explicit_apply,
- relevance_thresholds = (1.01, 1.01),
- max_relevant = SOME 65536 (* a large number *), isar_proof = isar_proof,
+ provers = provers, full_types = full_types,
+ explicit_apply = explicit_apply, relevance_thresholds = (1.01, 1.01),
+ max_relevant = NONE, isar_proof = isar_proof,
isar_shrink_factor = isar_shrink_factor, timeout = timeout, expect = ""}
- val axioms = maps (fn (n, ths) => map (pair n) ths) axioms
- val {context = ctxt, goal, ...} = Proof.goal state
+ val axioms =
+ axioms |> maps (fn (n, ths) => ths |> map (Unprepared o pair n))
+ val {goal, ...} = Proof.goal state
val problem =
- {state = state, goal = goal, subgoal = subgoal,
- axioms = map (prepare_axiom ctxt) axioms, only = true}
- val result as {outcome, used_thm_names, ...} = prover params (K "") problem
+ {state = state, goal = goal, subgoal = i, subgoal_count = n,
+ axioms = axioms, only = true}
+ val result as {outcome, used_axioms, ...} = prover params (K "") problem
in
priority (case outcome of
- NONE =>
- if length used_thm_names = length axioms then
- "Found proof."
- else
- "Found proof with " ^ n_theorems used_thm_names ^ "."
- | SOME failure => string_for_failure failure);
+ SOME failure => string_for_failure failure
+ | NONE =>
+ if length used_axioms = length axioms then "Found proof."
+ else "Found proof with " ^ n_facts used_axioms ^ ".");
result
end
(* minimalization of thms *)
-fun filter_used_facts used = filter (member (op =) used o fst)
+fun filter_used_axioms used = filter (member (op =) used o fst)
fun sublinear_minimize _ [] p = p
| sublinear_minimize test (x :: xs) (seen, result) =
case test (xs @ seen) of
- result as {outcome = NONE, used_thm_names, ...} : prover_result =>
- sublinear_minimize test (filter_used_facts used_thm_names xs)
- (filter_used_facts used_thm_names seen, result)
+ result as {outcome = NONE, used_axioms, ...} : prover_result =>
+ sublinear_minimize test (filter_used_axioms used_axioms xs)
+ (filter_used_axioms used_axioms seen, result)
| _ => sublinear_minimize test xs (x :: seen, result)
(* Give the ATP some slack. The ATP gets further slack because the Sledgehammer
@@ -91,48 +87,40 @@
timeout. *)
val fudge_msecs = 1000
-fun minimize_theorems {atps = [], ...} _ _ _ _ = error "No ATP is set."
- | minimize_theorems (params as {debug, atps = atp :: _, full_types,
- isar_proof, isar_shrink_factor, timeout, ...})
- i (_ : int) state axioms =
+fun minimize_facts {provers = [], ...} _ _ _ _ = error "No prover is set."
+ | minimize_facts (params as {provers = prover_name :: _, timeout, ...}) i n
+ state axioms =
let
val thy = Proof.theory_of state
- val prover = get_prover_fun thy atp
+ val prover = get_prover thy false prover_name
val msecs = Time.toMilliseconds timeout
- val _ = priority ("Sledgehammer minimize: ATP " ^ quote atp ^ ".")
- val {context = ctxt, goal, ...} = Proof.goal state
+ val _ = priority ("Sledgehammer minimize: " ^ quote prover_name ^ ".")
+ val {goal, ...} = Proof.goal state
val (_, hyp_ts, concl_t) = strip_subgoal goal i
val explicit_apply =
not (forall (Meson.is_fol_term thy)
(concl_t :: hyp_ts @ maps (map prop_of o snd) axioms))
fun do_test timeout =
- test_theorems params prover explicit_apply timeout i state
+ test_facts params prover explicit_apply timeout i n state
val timer = Timer.startRealTimer ()
in
(case do_test timeout axioms of
- result as {outcome = NONE, pool, used_thm_names,
- conjecture_shape, ...} =>
+ result as {outcome = NONE, used_axioms, ...} =>
let
val time = Timer.checkRealTimer timer
val new_timeout =
Int.min (msecs, Time.toMilliseconds time + fudge_msecs)
|> Time.fromMilliseconds
- val (min_thms, {tstplike_proof, axiom_names, ...}) =
+ val (min_thms, {message, ...}) =
sublinear_minimize (do_test new_timeout)
- (filter_used_facts used_thm_names axioms) ([], result)
+ (filter_used_axioms used_axioms axioms) ([], result)
val n = length min_thms
val _ = priority (cat_lines
- ["Minimized: " ^ string_of_int n ^ " theorem" ^ plural_s n] ^
+ ["Minimized: " ^ string_of_int n ^ " fact" ^ plural_s n] ^
(case length (filter (curry (op =) Chained o snd o fst) min_thms) of
0 => ""
| n => " (including " ^ Int.toString n ^ " chained)") ^ ".")
- in
- (SOME min_thms,
- proof_text isar_proof
- (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
- ("Minimized command", full_types, K "", tstplike_proof,
- axiom_names, goal, i) |> fst)
- end
+ in (SOME min_thms, message) end
| {outcome = SOME TimedOut, ...} =>
(NONE, "Timeout: You can increase the time limit using the \"timeout\" \
\option (e.g., \"timeout = " ^
@@ -159,7 +147,8 @@
case subgoal_count state of
0 => priority "No subgoal!"
| n =>
- (kill_atps (); priority (#2 (minimize_theorems params i n state axioms)))
+ (kill_provers ();
+ priority (#2 (minimize_facts params i n state axioms)))
end
end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML Sun Oct 24 03:43:12 2010 -0700
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,941 +0,0 @@
-(* Title: HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML
- Author: Lawrence C. Paulson, Cambridge University Computer Laboratory
- Author: Claire Quigley, Cambridge University Computer Laboratory
- Author: Jasmin Blanchette, TU Muenchen
-
-Proof reconstruction for Sledgehammer.
-*)
-
-signature SLEDGEHAMMER_RECONSTRUCT =
-sig
- type locality = Sledgehammer_Filter.locality
- type minimize_command = string list -> string
- type metis_params =
- string * bool * minimize_command * string * (string * locality) list vector
- * thm * int
- type isar_params =
- string Symtab.table * bool * int * Proof.context * int list list
- type text_result = string * (string * locality) list
-
- val repair_conjecture_shape_and_axiom_names :
- string -> int list list -> (string * locality) list vector
- -> int list list * (string * locality) list vector
- val metis_proof_text : metis_params -> text_result
- val isar_proof_text : isar_params -> metis_params -> text_result
- val proof_text : bool -> isar_params -> metis_params -> text_result
-end;
-
-structure Sledgehammer_Reconstruct : SLEDGEHAMMER_RECONSTRUCT =
-struct
-
-open ATP_Problem
-open ATP_Proof
-open Metis_Translate
-open Sledgehammer_Util
-open Sledgehammer_Filter
-open Sledgehammer_Translate
-
-type minimize_command = string list -> string
-type metis_params =
- string * bool * minimize_command * string * (string * locality) list vector
- * thm * int
-type isar_params =
- string Symtab.table * bool * int * Proof.context * int list list
-type text_result = string * (string * locality) list
-
-fun is_head_digit s = Char.isDigit (String.sub (s, 0))
-val scan_integer = Scan.many1 is_head_digit >> (the o Int.fromString o implode)
-
-fun find_first_in_list_vector vec key =
- Vector.foldl (fn (ps, NONE) => AList.lookup (op =) ps key
- | (_, value) => value) NONE vec
-
-
-(** SPASS's Flotter hack **)
-
-(* This is a hack required for keeping track of axioms after they have been
- clausified by SPASS's Flotter tool. The "ATP/scripts/spass" script is also
- part of this hack. *)
-
-val set_ClauseFormulaRelationN = "set_ClauseFormulaRelation"
-
-fun extract_clause_sequence output =
- let
- val tokens_of = String.tokens (not o Char.isAlphaNum)
- fun extract_num ("clause" :: (ss as _ :: _)) =
- Int.fromString (List.last ss)
- | extract_num _ = NONE
- in output |> split_lines |> map_filter (extract_num o tokens_of) end
-
-val parse_clause_formula_pair =
- $$ "(" |-- scan_integer --| $$ ","
- -- (Symbol.scan_id ::: Scan.repeat ($$ "," |-- Symbol.scan_id)) --| $$ ")"
- --| Scan.option ($$ ",")
-val parse_clause_formula_relation =
- Scan.this_string set_ClauseFormulaRelationN |-- $$ "("
- |-- Scan.repeat parse_clause_formula_pair
-val extract_clause_formula_relation =
- Substring.full #> Substring.position set_ClauseFormulaRelationN
- #> snd #> Substring.position "." #> fst #> Substring.string
- #> explode #> filter_out Symbol.is_blank #> parse_clause_formula_relation
- #> fst
-
-fun repair_conjecture_shape_and_axiom_names output conjecture_shape
- axiom_names =
- if String.isSubstring set_ClauseFormulaRelationN output then
- let
- val j0 = hd (hd conjecture_shape)
- val seq = extract_clause_sequence output
- val name_map = extract_clause_formula_relation output
- fun renumber_conjecture j =
- conjecture_prefix ^ string_of_int (j - j0)
- |> AList.find (fn (s, ss) => member (op =) ss s) name_map
- |> map (fn s => find_index (curry (op =) s) seq + 1)
- fun names_for_number j =
- j |> AList.lookup (op =) name_map |> these
- |> map_filter (try (unprefix axiom_prefix)) |> map unascii_of
- |> map (fn name =>
- (name, name |> find_first_in_list_vector axiom_names
- |> the)
- handle Option.Option =>
- error ("No such fact: " ^ quote name ^ "."))
- in
- (conjecture_shape |> map (maps renumber_conjecture),
- seq |> map names_for_number |> Vector.fromList)
- end
- else
- (conjecture_shape, axiom_names)
-
-
-(** Soft-core proof reconstruction: Metis one-liner **)
-
-fun string_for_label (s, num) = s ^ string_of_int num
-
-fun metis_using [] = ""
- | metis_using ls =
- "using " ^ space_implode " " (map string_for_label ls) ^ " "
-fun metis_apply _ 1 = "by "
- | metis_apply 1 _ = "apply "
- | metis_apply i _ = "prefer " ^ string_of_int i ^ " apply "
-fun metis_name full_types = if full_types then "metisFT" else "metis"
-fun metis_call full_types [] = metis_name full_types
- | metis_call full_types ss =
- "(" ^ metis_name full_types ^ " " ^ space_implode " " ss ^ ")"
-fun metis_command full_types i n (ls, ss) =
- metis_using ls ^ metis_apply i n ^ metis_call full_types ss
-fun metis_line banner full_types i n ss =
- banner ^ ": " ^
- Markup.markup Markup.sendback (metis_command full_types i n ([], ss)) ^ "."
-fun minimize_line _ [] = ""
- | minimize_line minimize_command ss =
- case minimize_command ss of
- "" => ""
- | command =>
- "\nTo minimize the number of lemmas, try this: " ^
- Markup.markup Markup.sendback command ^ "."
-
-fun resolve_axiom axiom_names ((_, SOME s)) =
- (case strip_prefix_and_unascii axiom_prefix s of
- SOME s' => (case find_first_in_list_vector axiom_names s' of
- SOME x => [(s', x)]
- | NONE => [])
- | NONE => [])
- | resolve_axiom axiom_names (num, NONE) =
- case Int.fromString num of
- SOME j =>
- if j > 0 andalso j <= Vector.length axiom_names then
- Vector.sub (axiom_names, j - 1)
- else
- []
- | NONE => []
-
-fun add_fact axiom_names (Inference (name, _, [])) =
- append (resolve_axiom axiom_names name)
- | add_fact _ _ = I
-
-fun used_facts_in_tstplike_proof axiom_names =
- atp_proof_from_tstplike_string #> rpair [] #-> fold (add_fact axiom_names)
-
-fun used_facts axiom_names =
- used_facts_in_tstplike_proof axiom_names
- #> List.partition (curry (op =) Chained o snd)
- #> pairself (sort_distinct (string_ord o pairself fst))
-
-fun metis_proof_text (banner, full_types, minimize_command,
- tstplike_proof, axiom_names, goal, i) =
- let
- val (chained_lemmas, other_lemmas) =
- used_facts axiom_names tstplike_proof
- val n = Logic.count_prems (prop_of goal)
- in
- (metis_line banner full_types i n (map fst other_lemmas) ^
- minimize_line minimize_command (map fst (other_lemmas @ chained_lemmas)),
- other_lemmas @ chained_lemmas)
- end
-
-
-(** Hard-core proof reconstruction: structured Isar proofs **)
-
-(* Simple simplifications to ensure that sort annotations don't leave a trail of
- spurious "True"s. *)
-fun s_not @{const False} = @{const True}
- | s_not @{const True} = @{const False}
- | s_not (@{const Not} $ t) = t
- | s_not t = @{const Not} $ t
-fun s_conj (@{const True}, t2) = t2
- | s_conj (t1, @{const True}) = t1
- | s_conj p = HOLogic.mk_conj p
-fun s_disj (@{const False}, t2) = t2
- | s_disj (t1, @{const False}) = t1
- | s_disj p = HOLogic.mk_disj p
-fun s_imp (@{const True}, t2) = t2
- | s_imp (t1, @{const False}) = s_not t1
- | s_imp p = HOLogic.mk_imp p
-fun s_iff (@{const True}, t2) = t2
- | s_iff (t1, @{const True}) = t1
- | s_iff (t1, t2) = HOLogic.eq_const HOLogic.boolT $ t1 $ t2
-
-fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda v t
-fun exists_of v t = HOLogic.exists_const (fastype_of v) $ lambda v t
-
-fun negate_term (Const (@{const_name All}, T) $ Abs (s, T', t')) =
- Const (@{const_name Ex}, T) $ Abs (s, T', negate_term t')
- | negate_term (Const (@{const_name Ex}, T) $ Abs (s, T', t')) =
- Const (@{const_name All}, T) $ Abs (s, T', negate_term t')
- | negate_term (@{const HOL.implies} $ t1 $ t2) =
- @{const HOL.conj} $ t1 $ negate_term t2
- | negate_term (@{const HOL.conj} $ t1 $ t2) =
- @{const HOL.disj} $ negate_term t1 $ negate_term t2
- | negate_term (@{const HOL.disj} $ t1 $ t2) =
- @{const HOL.conj} $ negate_term t1 $ negate_term t2
- | negate_term (@{const Not} $ t) = t
- | negate_term t = @{const Not} $ t
-
-val indent_size = 2
-val no_label = ("", ~1)
-
-val raw_prefix = "X"
-val assum_prefix = "A"
-val fact_prefix = "F"
-
-fun resolve_conjecture conjecture_shape (num, s_opt) =
- let
- val k = case try (unprefix conjecture_prefix) (the_default "" s_opt) of
- SOME s => Int.fromString s |> the_default ~1
- | NONE => case Int.fromString num of
- SOME j => find_index (exists (curry (op =) j))
- conjecture_shape
- | NONE => ~1
- in if k >= 0 then [k] else [] end
-
-fun is_axiom conjecture_shape = not o null o resolve_axiom conjecture_shape
-fun is_conjecture conjecture_shape = not o null o resolve_conjecture conjecture_shape
-
-fun raw_label_for_name conjecture_shape name =
- case resolve_conjecture conjecture_shape name of
- [j] => (conjecture_prefix, j)
- | _ => case Int.fromString (fst name) of
- SOME j => (raw_prefix, j)
- | NONE => (raw_prefix ^ fst name, 0)
-
-(**** INTERPRETATION OF TSTP SYNTAX TREES ****)
-
-exception FO_TERM of string fo_term list
-exception FORMULA of (string, string fo_term) formula list
-exception SAME of unit
-
-(* Type variables are given the basic sort "HOL.type". Some will later be
- constrained by information from type literals, or by type inference. *)
-fun type_from_fo_term tfrees (u as ATerm (a, us)) =
- let val Ts = map (type_from_fo_term tfrees) us in
- case strip_prefix_and_unascii type_const_prefix a of
- SOME b => Type (invert_const b, Ts)
- | NONE =>
- if not (null us) then
- raise FO_TERM [u] (* only "tconst"s have type arguments *)
- else case strip_prefix_and_unascii tfree_prefix a of
- SOME b =>
- let val s = "'" ^ b in
- TFree (s, AList.lookup (op =) tfrees s |> the_default HOLogic.typeS)
- end
- | NONE =>
- case strip_prefix_and_unascii tvar_prefix a of
- SOME b => TVar (("'" ^ b, 0), HOLogic.typeS)
- | NONE =>
- (* Variable from the ATP, say "X1" *)
- Type_Infer.param 0 (a, HOLogic.typeS)
- end
-
-(* Type class literal applied to a type. Returns triple of polarity, class,
- type. *)
-fun type_constraint_from_term pos tfrees (u as ATerm (a, us)) =
- case (strip_prefix_and_unascii class_prefix a,
- map (type_from_fo_term tfrees) us) of
- (SOME b, [T]) => (pos, b, T)
- | _ => raise FO_TERM [u]
-
-(** Accumulate type constraints in a formula: negative type literals **)
-fun add_var (key, z) = Vartab.map_default (key, []) (cons z)
-fun add_type_constraint (false, cl, TFree (a ,_)) = add_var ((a, ~1), cl)
- | add_type_constraint (false, cl, TVar (ix, _)) = add_var (ix, cl)
- | add_type_constraint _ = I
-
-fun repair_atp_variable_name f s =
- let
- fun subscript_name s n = s ^ nat_subscript n
- val s = String.map f s
- in
- case space_explode "_" s of
- [_] => (case take_suffix Char.isDigit (String.explode s) of
- (cs1 as _ :: _, cs2 as _ :: _) =>
- subscript_name (String.implode cs1)
- (the (Int.fromString (String.implode cs2)))
- | (_, _) => s)
- | [s1, s2] => (case Int.fromString s2 of
- SOME n => subscript_name s1 n
- | NONE => s)
- | _ => s
- end
-
-(* First-order translation. No types are known for variables. "HOLogic.typeT"
- should allow them to be inferred. *)
-fun raw_term_from_pred thy full_types tfrees =
- let
- fun aux opt_T extra_us u =
- case u of
- ATerm ("hBOOL", [u1]) => aux (SOME @{typ bool}) [] u1
- | ATerm ("hAPP", [u1, u2]) => aux opt_T (u2 :: extra_us) u1
- | ATerm (a, us) =>
- if a = type_wrapper_name then
- case us of
- [typ_u, term_u] =>
- aux (SOME (type_from_fo_term tfrees typ_u)) extra_us term_u
- | _ => raise FO_TERM us
- else case strip_prefix_and_unascii const_prefix a of
- SOME "equal" =>
- let val ts = map (aux NONE []) us in
- if length ts = 2 andalso hd ts aconv List.last ts then
- (* Vampire is keen on producing these. *)
- @{const True}
- else
- list_comb (Const (@{const_name HOL.eq}, HOLogic.typeT), ts)
- end
- | SOME b =>
- let
- val c = invert_const b
- val num_type_args = num_type_args thy c
- val (type_us, term_us) =
- chop (if full_types then 0 else num_type_args) us
- (* Extra args from "hAPP" come after any arguments given directly to
- the constant. *)
- val term_ts = map (aux NONE []) term_us
- val extra_ts = map (aux NONE []) extra_us
- val t =
- Const (c, if full_types then
- case opt_T of
- SOME T => map fastype_of term_ts ---> T
- | NONE =>
- if num_type_args = 0 then
- Sign.const_instance thy (c, [])
- else
- raise Fail ("no type information for " ^ quote c)
- else
- Sign.const_instance thy (c,
- map (type_from_fo_term tfrees) type_us))
- in list_comb (t, term_ts @ extra_ts) end
- | NONE => (* a free or schematic variable *)
- let
- val ts = map (aux NONE []) (us @ extra_us)
- val T = map fastype_of ts ---> HOLogic.typeT
- val t =
- case strip_prefix_and_unascii fixed_var_prefix a of
- SOME b => Free (b, T)
- | NONE =>
- case strip_prefix_and_unascii schematic_var_prefix a of
- SOME b => Var ((b, 0), T)
- | NONE =>
- if is_atp_variable a then
- Var ((repair_atp_variable_name Char.toLower a, 0), T)
- else
- (* Skolem constants? *)
- Var ((repair_atp_variable_name Char.toUpper a, 0), T)
- in list_comb (t, ts) end
- in aux (SOME HOLogic.boolT) [] end
-
-fun term_from_pred thy full_types tfrees pos (u as ATerm (s, _)) =
- if String.isPrefix class_prefix s then
- add_type_constraint (type_constraint_from_term pos tfrees u)
- #> pair @{const True}
- else
- pair (raw_term_from_pred thy full_types tfrees u)
-
-val combinator_table =
- [(@{const_name Meson.COMBI}, @{thm Meson.COMBI_def_raw}),
- (@{const_name Meson.COMBK}, @{thm Meson.COMBK_def_raw}),
- (@{const_name Meson.COMBB}, @{thm Meson.COMBB_def_raw}),
- (@{const_name Meson.COMBC}, @{thm Meson.COMBC_def_raw}),
- (@{const_name Meson.COMBS}, @{thm Meson.COMBS_def_raw})]
-
-fun uncombine_term (t1 $ t2) = betapply (pairself uncombine_term (t1, t2))
- | uncombine_term (Abs (s, T, t')) = Abs (s, T, uncombine_term t')
- | uncombine_term (t as Const (x as (s, _))) =
- (case AList.lookup (op =) combinator_table s of
- SOME thm => thm |> prop_of |> specialize_type @{theory} x |> Logic.dest_equals |> snd
- | NONE => t)
- | uncombine_term t = t
-
-(* Update schematic type variables with detected sort constraints. It's not
- totally clear when this code is necessary. *)
-fun repair_tvar_sorts (t, tvar_tab) =
- let
- fun do_type (Type (a, Ts)) = Type (a, map do_type Ts)
- | do_type (TVar (xi, s)) =
- TVar (xi, the_default s (Vartab.lookup tvar_tab xi))
- | do_type (TFree z) = TFree z
- fun do_term (Const (a, T)) = Const (a, do_type T)
- | do_term (Free (a, T)) = Free (a, do_type T)
- | do_term (Var (xi, T)) = Var (xi, do_type T)
- | do_term (t as Bound _) = t
- | do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t)
- | do_term (t1 $ t2) = do_term t1 $ do_term t2
- in t |> not (Vartab.is_empty tvar_tab) ? do_term end
-
-fun quantify_over_var quant_of var_s t =
- let
- val vars = [] |> Term.add_vars t |> filter (fn ((s, _), _) => s = var_s)
- |> map Var
- in fold_rev quant_of vars t end
-
-(* Interpret an ATP formula as a HOL term, extracting sort constraints as they
- appear in the formula. *)
-fun prop_from_formula thy full_types tfrees phi =
- let
- fun do_formula pos phi =
- case phi of
- AQuant (_, [], phi) => do_formula pos phi
- | AQuant (q, x :: xs, phi') =>
- do_formula pos (AQuant (q, xs, phi'))
- #>> quantify_over_var (case q of
- AForall => forall_of
- | AExists => exists_of)
- (repair_atp_variable_name Char.toLower x)
- | AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not
- | AConn (c, [phi1, phi2]) =>
- do_formula (pos |> c = AImplies ? not) phi1
- ##>> do_formula pos phi2
- #>> (case c of
- AAnd => s_conj
- | AOr => s_disj
- | AImplies => s_imp
- | AIf => s_imp o swap
- | AIff => s_iff
- | ANotIff => s_not o s_iff)
- | AAtom tm => term_from_pred thy full_types tfrees pos tm
- | _ => raise FORMULA [phi]
- in repair_tvar_sorts (do_formula true phi Vartab.empty) end
-
-fun check_formula ctxt =
- Type.constraint HOLogic.boolT
- #> Syntax.check_term (ProofContext.set_mode ProofContext.mode_schematic ctxt)
-
-
-(**** Translation of TSTP files to Isar Proofs ****)
-
-fun unvarify_term (Var ((s, 0), T)) = Free (s, T)
- | unvarify_term t = raise TERM ("unvarify_term: non-Var", [t])
-
-fun decode_line full_types tfrees (Definition (name, phi1, phi2)) ctxt =
- let
- val thy = ProofContext.theory_of ctxt
- val t1 = prop_from_formula thy full_types tfrees phi1
- val vars = snd (strip_comb t1)
- val frees = map unvarify_term vars
- val unvarify_args = subst_atomic (vars ~~ frees)
- val t2 = prop_from_formula thy full_types tfrees phi2
- val (t1, t2) =
- HOLogic.eq_const HOLogic.typeT $ t1 $ t2
- |> unvarify_args |> uncombine_term |> check_formula ctxt
- |> HOLogic.dest_eq
- in
- (Definition (name, t1, t2),
- fold Variable.declare_term (maps OldTerm.term_frees [t1, t2]) ctxt)
- end
- | decode_line full_types tfrees (Inference (name, u, deps)) ctxt =
- let
- val thy = ProofContext.theory_of ctxt
- val t = u |> prop_from_formula thy full_types tfrees
- |> uncombine_term |> check_formula ctxt
- in
- (Inference (name, t, deps),
- fold Variable.declare_term (OldTerm.term_frees t) ctxt)
- end
-fun decode_lines ctxt full_types tfrees lines =
- fst (fold_map (decode_line full_types tfrees) lines ctxt)
-
-fun is_same_inference _ (Definition _) = false
- | is_same_inference t (Inference (_, t', _)) = t aconv t'
-
-(* No "real" literals means only type information (tfree_tcs, clsrel, or
- clsarity). *)
-val is_only_type_information = curry (op aconv) HOLogic.true_const
-
-fun replace_one_dependency (old, new) dep =
- if is_same_step (dep, old) then new else [dep]
-fun replace_dependencies_in_line _ (line as Definition _) = line
- | replace_dependencies_in_line p (Inference (name, t, deps)) =
- Inference (name, t, fold (union (op =) o replace_one_dependency p) deps [])
-
-(* Discard axioms; consolidate adjacent lines that prove the same formula, since
- they differ only in type information.*)
-fun add_line _ _ (line as Definition _) lines = line :: lines
- | add_line conjecture_shape axiom_names (Inference (name, t, [])) lines =
- (* No dependencies: axiom, conjecture, or (for Vampire) internal axioms or
- definitions. *)
- if is_axiom axiom_names name then
- (* Axioms are not proof lines. *)
- if is_only_type_information t then
- map (replace_dependencies_in_line (name, [])) lines
- (* Is there a repetition? If so, replace later line by earlier one. *)
- else case take_prefix (not o is_same_inference t) lines of
- (_, []) => lines (* no repetition of proof line *)
- | (pre, Inference (name', _, _) :: post) =>
- pre @ map (replace_dependencies_in_line (name', [name])) post
- else if is_conjecture conjecture_shape name then
- Inference (name, negate_term t, []) :: lines
- else
- map (replace_dependencies_in_line (name, [])) lines
- | add_line _ _ (Inference (name, t, deps)) lines =
- (* Type information will be deleted later; skip repetition test. *)
- if is_only_type_information t then
- Inference (name, t, deps) :: lines
- (* Is there a repetition? If so, replace later line by earlier one. *)
- else case take_prefix (not o is_same_inference t) lines of
- (* FIXME: Doesn't this code risk conflating proofs involving different
- types? *)
- (_, []) => Inference (name, t, deps) :: lines
- | (pre, Inference (name', t', _) :: post) =>
- Inference (name, t', deps) ::
- pre @ map (replace_dependencies_in_line (name', [name])) post
-
-(* Recursively delete empty lines (type information) from the proof. *)
-fun add_nontrivial_line (Inference (name, t, [])) lines =
- if is_only_type_information t then delete_dependency name lines
- else Inference (name, t, []) :: lines
- | add_nontrivial_line line lines = line :: lines
-and delete_dependency name lines =
- fold_rev add_nontrivial_line
- (map (replace_dependencies_in_line (name, [])) lines) []
-
-(* ATPs sometimes reuse free variable names in the strangest ways. Removing
- offending lines often does the trick. *)
-fun is_bad_free frees (Free x) = not (member (op =) frees x)
- | is_bad_free _ _ = false
-
-fun add_desired_line _ _ _ _ (line as Definition (name, _, _)) (j, lines) =
- (j, line :: map (replace_dependencies_in_line (name, [])) lines)
- | add_desired_line isar_shrink_factor conjecture_shape axiom_names frees
- (Inference (name, t, deps)) (j, lines) =
- (j + 1,
- if is_axiom axiom_names name orelse
- is_conjecture conjecture_shape name orelse
- (* the last line must be kept *)
- j = 0 orelse
- (not (is_only_type_information t) andalso
- null (Term.add_tvars t []) andalso
- not (exists_subterm (is_bad_free frees) t) andalso
- length deps >= 2 andalso j mod isar_shrink_factor = 0 andalso
- (* kill next to last line, which usually results in a trivial step *)
- j <> 1) then
- Inference (name, t, deps) :: lines (* keep line *)
- else
- map (replace_dependencies_in_line (name, deps)) lines) (* drop line *)
-
-(** Isar proof construction and manipulation **)
-
-fun merge_fact_sets (ls1, ss1) (ls2, ss2) =
- (union (op =) ls1 ls2, union (op =) ss1 ss2)
-
-type label = string * int
-type facts = label list * string list
-
-datatype isar_qualifier = Show | Then | Moreover | Ultimately
-
-datatype isar_step =
- Fix of (string * typ) list |
- Let of term * term |
- Assume of label * term |
- Have of isar_qualifier list * label * term * byline
-and byline =
- ByMetis of facts |
- CaseSplit of isar_step list list * facts
-
-fun smart_case_split [] facts = ByMetis facts
- | smart_case_split proofs facts = CaseSplit (proofs, facts)
-
-fun add_fact_from_dependency conjecture_shape axiom_names name =
- if is_axiom axiom_names name then
- apsnd (union (op =) (map fst (resolve_axiom axiom_names name)))
- else
- apfst (insert (op =) (raw_label_for_name conjecture_shape name))
-
-fun step_for_line _ _ _ (Definition (_, t1, t2)) = Let (t1, t2)
- | step_for_line conjecture_shape _ _ (Inference (name, t, [])) =
- Assume (raw_label_for_name conjecture_shape name, t)
- | step_for_line conjecture_shape axiom_names j (Inference (name, t, deps)) =
- Have (if j = 1 then [Show] else [],
- raw_label_for_name conjecture_shape name,
- fold_rev forall_of (map Var (Term.add_vars t [])) t,
- ByMetis (fold (add_fact_from_dependency conjecture_shape axiom_names)
- deps ([], [])))
-
-fun repair_name "$true" = "c_True"
- | repair_name "$false" = "c_False"
- | repair_name "$$e" = "c_equal" (* seen in Vampire proofs *)
- | repair_name "equal" = "c_equal" (* needed by SPASS? *)
- | repair_name s =
- if String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s then
- "c_equal" (* seen in Vampire proofs *)
- else
- s
-
-fun isar_proof_from_tstplike_proof pool ctxt full_types tfrees isar_shrink_factor
- tstplike_proof conjecture_shape axiom_names params frees =
- let
- val lines =
- tstplike_proof
- |> atp_proof_from_tstplike_string
- |> nasty_atp_proof pool
- |> map_term_names_in_atp_proof repair_name
- |> decode_lines ctxt full_types tfrees
- |> rpair [] |-> fold_rev (add_line conjecture_shape axiom_names)
- |> rpair [] |-> fold_rev add_nontrivial_line
- |> rpair (0, []) |-> fold_rev (add_desired_line isar_shrink_factor
- conjecture_shape axiom_names frees)
- |> snd
- in
- (if null params then [] else [Fix params]) @
- map2 (step_for_line conjecture_shape axiom_names) (length lines downto 1)
- lines
- end
-
-(* When redirecting proofs, we keep information about the labels seen so far in
- the "backpatches" data structure. The first component indicates which facts
- should be associated with forthcoming proof steps. The second component is a
- pair ("assum_ls", "drop_ls"), where "assum_ls" are the labels that should
- become assumptions and "drop_ls" are the labels that should be dropped in a
- case split. *)
-type backpatches = (label * facts) list * (label list * label list)
-
-fun used_labels_of_step (Have (_, _, _, by)) =
- (case by of
- ByMetis (ls, _) => ls
- | CaseSplit (proofs, (ls, _)) =>
- fold (union (op =) o used_labels_of) proofs ls)
- | used_labels_of_step _ = []
-and used_labels_of proof = fold (union (op =) o used_labels_of_step) proof []
-
-fun new_labels_of_step (Fix _) = []
- | new_labels_of_step (Let _) = []
- | new_labels_of_step (Assume (l, _)) = [l]
- | new_labels_of_step (Have (_, l, _, _)) = [l]
-val new_labels_of = maps new_labels_of_step
-
-val join_proofs =
- let
- fun aux _ [] = NONE
- | aux proof_tail (proofs as (proof1 :: _)) =
- if exists null proofs then
- NONE
- else if forall (curry (op =) (hd proof1) o hd) (tl proofs) then
- aux (hd proof1 :: proof_tail) (map tl proofs)
- else case hd proof1 of
- Have ([], l, t, _) => (* FIXME: should we really ignore the "by"? *)
- if forall (fn Have ([], l', t', _) :: _ => (l, t) = (l', t')
- | _ => false) (tl proofs) andalso
- not (exists (member (op =) (maps new_labels_of proofs))
- (used_labels_of proof_tail)) then
- SOME (l, t, map rev proofs, proof_tail)
- else
- NONE
- | _ => NONE
- in aux [] o map rev end
-
-fun case_split_qualifiers proofs =
- case length proofs of
- 0 => []
- | 1 => [Then]
- | _ => [Ultimately]
-
-fun redirect_proof hyp_ts concl_t proof =
- let
- (* The first pass outputs those steps that are independent of the negated
- conjecture. The second pass flips the proof by contradiction to obtain a
- direct proof, introducing case splits when an inference depends on
- several facts that depend on the negated conjecture. *)
- val concl_l = (conjecture_prefix, length hyp_ts)
- fun first_pass ([], contra) = ([], contra)
- | first_pass ((step as Fix _) :: proof, contra) =
- first_pass (proof, contra) |>> cons step
- | first_pass ((step as Let _) :: proof, contra) =
- first_pass (proof, contra) |>> cons step
- | first_pass ((step as Assume (l as (_, j), _)) :: proof, contra) =
- if l = concl_l then first_pass (proof, contra ||> cons step)
- else first_pass (proof, contra) |>> cons (Assume (l, nth hyp_ts j))
- | first_pass (Have (qs, l, t, ByMetis (ls, ss)) :: proof, contra) =
- let val step = Have (qs, l, t, ByMetis (ls, ss)) in
- if exists (member (op =) (fst contra)) ls then
- first_pass (proof, contra |>> cons l ||> cons step)
- else
- first_pass (proof, contra) |>> cons step
- end
- | first_pass _ = raise Fail "malformed proof"
- val (proof_top, (contra_ls, contra_proof)) =
- first_pass (proof, ([concl_l], []))
- val backpatch_label = the_default ([], []) oo AList.lookup (op =) o fst
- fun backpatch_labels patches ls =
- fold merge_fact_sets (map (backpatch_label patches) ls) ([], [])
- fun second_pass end_qs ([], assums, patches) =
- ([Have (end_qs, no_label, concl_t,
- ByMetis (backpatch_labels patches (map snd assums)))], patches)
- | second_pass end_qs (Assume (l, t) :: proof, assums, patches) =
- second_pass end_qs (proof, (t, l) :: assums, patches)
- | second_pass end_qs (Have (qs, l, t, ByMetis (ls, ss)) :: proof, assums,
- patches) =
- (if member (op =) (snd (snd patches)) l andalso
- not (member (op =) (fst (snd patches)) l) andalso
- not (AList.defined (op =) (fst patches) l) then
- second_pass end_qs (proof, assums, patches ||> apsnd (append ls))
- else case List.partition (member (op =) contra_ls) ls of
- ([contra_l], co_ls) =>
- if member (op =) qs Show then
- second_pass end_qs (proof, assums,
- patches |>> cons (contra_l, (co_ls, ss)))
- else
- second_pass end_qs
- (proof, assums,
- patches |>> cons (contra_l, (l :: co_ls, ss)))
- |>> cons (if member (op =) (fst (snd patches)) l then
- Assume (l, negate_term t)
- else
- Have (qs, l, negate_term t,
- ByMetis (backpatch_label patches l)))
- | (contra_ls as _ :: _, co_ls) =>
- let
- val proofs =
- map_filter
- (fn l =>
- if l = concl_l then
- NONE
- else
- let
- val drop_ls = filter (curry (op <>) l) contra_ls
- in
- second_pass []
- (proof, assums,
- patches ||> apfst (insert (op =) l)
- ||> apsnd (union (op =) drop_ls))
- |> fst |> SOME
- end) contra_ls
- val (assumes, facts) =
- if member (op =) (fst (snd patches)) l then
- ([Assume (l, negate_term t)], (l :: co_ls, ss))
- else
- ([], (co_ls, ss))
- in
- (case join_proofs proofs of
- SOME (l, t, proofs, proof_tail) =>
- Have (case_split_qualifiers proofs @
- (if null proof_tail then end_qs else []), l, t,
- smart_case_split proofs facts) :: proof_tail
- | NONE =>
- [Have (case_split_qualifiers proofs @ end_qs, no_label,
- concl_t, smart_case_split proofs facts)],
- patches)
- |>> append assumes
- end
- | _ => raise Fail "malformed proof")
- | second_pass _ _ = raise Fail "malformed proof"
- val proof_bottom =
- second_pass [Show] (contra_proof, [], ([], ([], []))) |> fst
- in proof_top @ proof_bottom end
-
-(* FIXME: Still needed? Probably not. *)
-val kill_duplicate_assumptions_in_proof =
- let
- fun relabel_facts subst =
- apfst (map (fn l => AList.lookup (op =) subst l |> the_default l))
- fun do_step (step as Assume (l, t)) (proof, subst, assums) =
- (case AList.lookup (op aconv) assums t of
- SOME l' => (proof, (l, l') :: subst, assums)
- | NONE => (step :: proof, subst, (t, l) :: assums))
- | do_step (Have (qs, l, t, by)) (proof, subst, assums) =
- (Have (qs, l, t,
- case by of
- ByMetis facts => ByMetis (relabel_facts subst facts)
- | CaseSplit (proofs, facts) =>
- CaseSplit (map do_proof proofs, relabel_facts subst facts)) ::
- proof, subst, assums)
- | do_step step (proof, subst, assums) = (step :: proof, subst, assums)
- and do_proof proof = fold do_step proof ([], [], []) |> #1 |> rev
- in do_proof end
-
-val then_chain_proof =
- let
- fun aux _ [] = []
- | aux _ ((step as Assume (l, _)) :: proof) = step :: aux l proof
- | aux l' (Have (qs, l, t, by) :: proof) =
- (case by of
- ByMetis (ls, ss) =>
- Have (if member (op =) ls l' then
- (Then :: qs, l, t,
- ByMetis (filter_out (curry (op =) l') ls, ss))
- else
- (qs, l, t, ByMetis (ls, ss)))
- | CaseSplit (proofs, facts) =>
- Have (qs, l, t, CaseSplit (map (aux no_label) proofs, facts))) ::
- aux l proof
- | aux _ (step :: proof) = step :: aux no_label proof
- in aux no_label end
-
-fun kill_useless_labels_in_proof proof =
- let
- val used_ls = used_labels_of proof
- fun do_label l = if member (op =) used_ls l then l else no_label
- fun do_step (Assume (l, t)) = Assume (do_label l, t)
- | do_step (Have (qs, l, t, by)) =
- Have (qs, do_label l, t,
- case by of
- CaseSplit (proofs, facts) =>
- CaseSplit (map (map do_step) proofs, facts)
- | _ => by)
- | do_step step = step
- in map do_step proof end
-
-fun prefix_for_depth n = replicate_string (n + 1)
-
-val relabel_proof =
- let
- fun aux _ _ _ [] = []
- | aux subst depth (next_assum, next_fact) (Assume (l, t) :: proof) =
- if l = no_label then
- Assume (l, t) :: aux subst depth (next_assum, next_fact) proof
- else
- let val l' = (prefix_for_depth depth assum_prefix, next_assum) in
- Assume (l', t) ::
- aux ((l, l') :: subst) depth (next_assum + 1, next_fact) proof
- end
- | aux subst depth (next_assum, next_fact) (Have (qs, l, t, by) :: proof) =
- let
- val (l', subst, next_fact) =
- if l = no_label then
- (l, subst, next_fact)
- else
- let
- val l' = (prefix_for_depth depth fact_prefix, next_fact)
- in (l', (l, l') :: subst, next_fact + 1) end
- val relabel_facts =
- apfst (maps (the_list o AList.lookup (op =) subst))
- val by =
- case by of
- ByMetis facts => ByMetis (relabel_facts facts)
- | CaseSplit (proofs, facts) =>
- CaseSplit (map (aux subst (depth + 1) (1, 1)) proofs,
- relabel_facts facts)
- in
- Have (qs, l', t, by) ::
- aux subst depth (next_assum, next_fact) proof
- end
- | aux subst depth nextp (step :: proof) =
- step :: aux subst depth nextp proof
- in aux [] 0 (1, 1) end
-
-fun string_for_proof ctxt0 full_types i n =
- let
- val ctxt = ctxt0
- |> Config.put show_free_types false
- |> Config.put show_types true
- fun fix_print_mode f x =
- Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN)
- (print_mode_value ())) f x
- fun do_indent ind = replicate_string (ind * indent_size) " "
- fun do_free (s, T) =
- maybe_quote s ^ " :: " ^
- maybe_quote (fix_print_mode (Syntax.string_of_typ ctxt) T)
- fun do_label l = if l = no_label then "" else string_for_label l ^ ": "
- fun do_have qs =
- (if member (op =) qs Moreover then "moreover " else "") ^
- (if member (op =) qs Ultimately then "ultimately " else "") ^
- (if member (op =) qs Then then
- if member (op =) qs Show then "thus" else "hence"
- else
- if member (op =) qs Show then "show" else "have")
- val do_term = maybe_quote o fix_print_mode (Syntax.string_of_term ctxt)
- fun do_facts (ls, ss) =
- metis_command full_types 1 1
- (ls |> sort_distinct (prod_ord string_ord int_ord),
- ss |> sort_distinct string_ord)
- and do_step ind (Fix xs) =
- do_indent ind ^ "fix " ^ space_implode " and " (map do_free xs) ^ "\n"
- | do_step ind (Let (t1, t2)) =
- do_indent ind ^ "let " ^ do_term t1 ^ " = " ^ do_term t2 ^ "\n"
- | do_step ind (Assume (l, t)) =
- do_indent ind ^ "assume " ^ do_label l ^ do_term t ^ "\n"
- | do_step ind (Have (qs, l, t, ByMetis facts)) =
- do_indent ind ^ do_have qs ^ " " ^
- do_label l ^ do_term t ^ " " ^ do_facts facts ^ "\n"
- | do_step ind (Have (qs, l, t, CaseSplit (proofs, facts))) =
- space_implode (do_indent ind ^ "moreover\n")
- (map (do_block ind) proofs) ^
- do_indent ind ^ do_have qs ^ " " ^ do_label l ^ do_term t ^ " " ^
- do_facts facts ^ "\n"
- and do_steps prefix suffix ind steps =
- let val s = implode (map (do_step ind) steps) in
- replicate_string (ind * indent_size - size prefix) " " ^ prefix ^
- String.extract (s, ind * indent_size,
- SOME (size s - ind * indent_size - 1)) ^
- suffix ^ "\n"
- end
- and do_block ind proof = do_steps "{ " " }" (ind + 1) proof
- (* One-step proofs are pointless; better use the Metis one-liner
- directly. *)
- and do_proof [Have (_, _, _, ByMetis _)] = ""
- | do_proof proof =
- (if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^
- do_indent 0 ^ "proof -\n" ^ do_steps "" "" 1 proof ^ do_indent 0 ^
- (if n <> 1 then "next" else "qed")
- in do_proof end
-
-fun isar_proof_text (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
- (other_params as (_, full_types, _, tstplike_proof,
- axiom_names, goal, i)) =
- let
- val (params, hyp_ts, concl_t) = strip_subgoal goal i
- val frees = fold Term.add_frees (concl_t :: hyp_ts) []
- val tfrees = fold Term.add_tfrees (concl_t :: hyp_ts) []
- val n = Logic.count_prems (prop_of goal)
- val (one_line_proof, lemma_names) = metis_proof_text other_params
- fun isar_proof_for () =
- case isar_proof_from_tstplike_proof pool ctxt full_types tfrees
- isar_shrink_factor tstplike_proof conjecture_shape axiom_names
- params frees
- |> redirect_proof hyp_ts concl_t
- |> kill_duplicate_assumptions_in_proof
- |> then_chain_proof
- |> kill_useless_labels_in_proof
- |> relabel_proof
- |> string_for_proof ctxt full_types i n of
- "" => "\nNo structured proof available."
- | proof => "\n\nStructured proof:\n" ^ Markup.markup Markup.sendback proof
- val isar_proof =
- if debug then
- isar_proof_for ()
- else
- try isar_proof_for ()
- |> the_default "\nWarning: The Isar proof construction failed."
- in (one_line_proof ^ isar_proof, lemma_names) end
-
-fun proof_text isar_proof isar_params other_params =
- (if isar_proof then isar_proof_text isar_params else metis_proof_text)
- other_params
-
-end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_translate.ML Sun Oct 24 03:43:12 2010 -0700
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,533 +0,0 @@
-(* Title: HOL/Tools/Sledgehammer/sledgehammer_translate.ML
- Author: Fabian Immler, TU Muenchen
- Author: Makarius
- Author: Jasmin Blanchette, TU Muenchen
-
-Translation of HOL to FOL for Sledgehammer.
-*)
-
-signature SLEDGEHAMMER_TRANSLATE =
-sig
- type 'a problem = 'a ATP_Problem.problem
- type fol_formula
-
- val axiom_prefix : string
- val conjecture_prefix : string
- val prepare_axiom :
- Proof.context -> (string * 'a) * thm
- -> term * ((string * 'a) * fol_formula) option
- val prepare_problem :
- Proof.context -> bool -> bool -> bool -> bool -> term list -> term
- -> (term * ((string * 'a) * fol_formula) option) list
- -> string problem * string Symtab.table * int * (string * 'a) list vector
-end;
-
-structure Sledgehammer_Translate : SLEDGEHAMMER_TRANSLATE =
-struct
-
-open ATP_Problem
-open Metis_Translate
-open Sledgehammer_Util
-
-val axiom_prefix = "ax_"
-val conjecture_prefix = "conj_"
-val helper_prefix = "help_"
-val class_rel_clause_prefix = "clrel_";
-val arity_clause_prefix = "arity_"
-val tfree_prefix = "tfree_"
-
-(* Freshness almost guaranteed! *)
-val sledgehammer_weak_prefix = "Sledgehammer:"
-
-type fol_formula =
- {name: string,
- kind: kind,
- combformula: (name, combterm) formula,
- ctypes_sorts: typ list}
-
-fun mk_anot phi = AConn (ANot, [phi])
-fun mk_aconn c phi1 phi2 = AConn (c, [phi1, phi2])
-fun mk_ahorn [] phi = phi
- | mk_ahorn (phi :: phis) psi =
- AConn (AImplies, [fold (mk_aconn AAnd) phis phi, psi])
-
-fun combformula_for_prop thy =
- let
- val do_term = combterm_from_term thy ~1
- fun do_quant bs q s T t' =
- let val s = Name.variant (map fst bs) s in
- do_formula ((s, T) :: bs) t'
- #>> (fn phi => AQuant (q, [`make_bound_var s], phi))
- end
- and do_conn bs c t1 t2 =
- do_formula bs t1 ##>> do_formula bs t2
- #>> (fn (phi1, phi2) => AConn (c, [phi1, phi2]))
- and do_formula bs t =
- case t of
- @{const Not} $ t1 =>
- do_formula bs t1 #>> (fn phi => AConn (ANot, [phi]))
- | Const (@{const_name All}, _) $ Abs (s, T, t') =>
- do_quant bs AForall s T t'
- | Const (@{const_name Ex}, _) $ Abs (s, T, t') =>
- do_quant bs AExists s T t'
- | @{const HOL.conj} $ t1 $ t2 => do_conn bs AAnd t1 t2
- | @{const HOL.disj} $ t1 $ t2 => do_conn bs AOr t1 t2
- | @{const HOL.implies} $ t1 $ t2 => do_conn bs AImplies t1 t2
- | Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
- do_conn bs AIff t1 t2
- | _ => (fn ts => do_term bs (Envir.eta_contract t)
- |>> AAtom ||> union (op =) ts)
- in do_formula [] end
-
-val presimplify_term = prop_of o Meson.presimplify oo Skip_Proof.make_thm
-
-fun concealed_bound_name j = sledgehammer_weak_prefix ^ Int.toString j
-fun conceal_bounds Ts t =
- subst_bounds (map (Free o apfst concealed_bound_name)
- (0 upto length Ts - 1 ~~ Ts), t)
-fun reveal_bounds Ts =
- subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j))
- (0 upto length Ts - 1 ~~ Ts))
-
-(* Removes the lambdas from an equation of the form "t = (%x. u)".
- (Cf. "extensionalize_theorem" in "Meson_Clausify".) *)
-fun extensionalize_term t =
- let
- fun aux j (@{const Trueprop} $ t') = @{const Trueprop} $ aux j t'
- | aux j (t as Const (s, Type (_, [Type (_, [_, T']),
- Type (_, [_, res_T])]))
- $ t2 $ Abs (var_s, var_T, t')) =
- if s = @{const_name HOL.eq} orelse s = @{const_name "=="} then
- let val var_t = Var ((var_s, j), var_T) in
- Const (s, T' --> T' --> res_T)
- $ betapply (t2, var_t) $ subst_bound (var_t, t')
- |> aux (j + 1)
- end
- else
- t
- | aux _ t = t
- in aux (maxidx_of_term t + 1) t end
-
-fun introduce_combinators_in_term ctxt kind t =
- let val thy = ProofContext.theory_of ctxt in
- if Meson.is_fol_term thy t then
- t
- else
- let
- fun aux Ts t =
- case t of
- @{const Not} $ t1 => @{const Not} $ aux Ts t1
- | (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') =>
- t0 $ Abs (s, T, aux (T :: Ts) t')
- | (t0 as Const (@{const_name All}, _)) $ t1 =>
- aux Ts (t0 $ eta_expand Ts t1 1)
- | (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') =>
- t0 $ Abs (s, T, aux (T :: Ts) t')
- | (t0 as Const (@{const_name Ex}, _)) $ t1 =>
- aux Ts (t0 $ eta_expand Ts t1 1)
- | (t0 as @{const HOL.conj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
- | (t0 as @{const HOL.disj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
- | (t0 as @{const HOL.implies}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
- | (t0 as Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])))
- $ t1 $ t2 =>
- t0 $ aux Ts t1 $ aux Ts t2
- | _ => if not (exists_subterm (fn Abs _ => true | _ => false) t) then
- t
- else
- t |> conceal_bounds Ts
- |> Envir.eta_contract
- |> cterm_of thy
- |> Meson_Clausify.introduce_combinators_in_cterm
- |> prop_of |> Logic.dest_equals |> snd
- |> reveal_bounds Ts
- val (t, ctxt') = Variable.import_terms true [t] ctxt |>> the_single
- in t |> aux [] |> singleton (Variable.export_terms ctxt' ctxt) end
- handle THM _ =>
- (* A type variable of sort "{}" will make abstraction fail. *)
- if kind = Conjecture then HOLogic.false_const
- else HOLogic.true_const
- end
-
-(* Metis's use of "resolve_tac" freezes the schematic variables. We simulate the
- same in Sledgehammer to prevent the discovery of unreplable proofs. *)
-fun freeze_term t =
- let
- fun aux (t $ u) = aux t $ aux u
- | aux (Abs (s, T, t)) = Abs (s, T, aux t)
- | aux (Var ((s, i), T)) =
- Free (sledgehammer_weak_prefix ^ s ^ "_" ^ string_of_int i, T)
- | aux t = t
- in t |> exists_subterm is_Var t ? aux end
-
-(* "Object_Logic.atomize_term" isn't as powerful as it could be; for example,
- it leaves metaequalities over "prop"s alone. *)
-val atomize_term =
- let
- fun aux (@{const Trueprop} $ t1) = t1
- | aux (Const (@{const_name all}, _) $ Abs (s, T, t')) =
- HOLogic.all_const T $ Abs (s, T, aux t')
- | aux (@{const "==>"} $ t1 $ t2) = HOLogic.mk_imp (pairself aux (t1, t2))
- | aux (Const (@{const_name "=="}, Type (_, [@{typ prop}, _])) $ t1 $ t2) =
- HOLogic.eq_const HOLogic.boolT $ aux t1 $ aux t2
- | aux (Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2) =
- HOLogic.eq_const T $ t1 $ t2
- | aux _ = raise Fail "aux"
- in perhaps (try aux) end
-
-(* making axiom and conjecture formulas *)
-fun make_formula ctxt presimp name kind t =
- let
- val thy = ProofContext.theory_of ctxt
- val t = t |> Envir.beta_eta_contract
- |> transform_elim_term
- |> atomize_term
- val need_trueprop = (fastype_of t = HOLogic.boolT)
- val t = t |> need_trueprop ? HOLogic.mk_Trueprop
- |> extensionalize_term
- |> presimp ? presimplify_term thy
- |> perhaps (try (HOLogic.dest_Trueprop))
- |> introduce_combinators_in_term ctxt kind
- |> kind <> Axiom ? freeze_term
- val (combformula, ctypes_sorts) = combformula_for_prop thy t []
- in
- {name = name, combformula = combformula, kind = kind,
- ctypes_sorts = ctypes_sorts}
- end
-
-fun make_axiom ctxt presimp ((name, loc), th) =
- case make_formula ctxt presimp name Axiom (prop_of th) of
- {combformula = AAtom (CombConst (("c_True", _), _, _)), ...} => NONE
- | formula => SOME ((name, loc), formula)
-fun make_conjecture ctxt ts =
- let val last = length ts - 1 in
- map2 (fn j => make_formula ctxt true (Int.toString j)
- (if j = last then Conjecture else Hypothesis))
- (0 upto last) ts
- end
-
-(** Helper facts **)
-
-fun count_combterm (CombConst ((s, _), _, _)) =
- Symtab.map_entry s (Integer.add 1)
- | count_combterm (CombVar _) = I
- | count_combterm (CombApp (t1, t2)) = fold count_combterm [t1, t2]
-fun count_combformula (AQuant (_, _, phi)) = count_combformula phi
- | count_combformula (AConn (_, phis)) = fold count_combformula phis
- | count_combformula (AAtom tm) = count_combterm tm
-fun count_fol_formula ({combformula, ...} : fol_formula) =
- count_combformula combformula
-
-val optional_helpers =
- [(["c_COMBI"], @{thms Meson.COMBI_def}),
- (["c_COMBK"], @{thms Meson.COMBK_def}),
- (["c_COMBB"], @{thms Meson.COMBB_def}),
- (["c_COMBC"], @{thms Meson.COMBC_def}),
- (["c_COMBS"], @{thms Meson.COMBS_def})]
-val optional_typed_helpers =
- [(["c_True", "c_False", "c_If"], @{thms True_or_False}),
- (["c_If"], @{thms if_True if_False})]
-val mandatory_helpers = @{thms Metis.fequal_def}
-
-val init_counters =
- [optional_helpers, optional_typed_helpers] |> maps (maps fst)
- |> sort_distinct string_ord |> map (rpair 0) |> Symtab.make
-
-fun get_helper_facts ctxt is_FO full_types conjectures axioms =
- let
- val ct = fold (fold count_fol_formula) [conjectures, axioms] init_counters
- fun is_needed c = the (Symtab.lookup ct c) > 0
- fun baptize th = ((Thm.get_name_hint th, false), th)
- in
- (optional_helpers
- |> full_types ? append optional_typed_helpers
- |> maps (fn (ss, ths) =>
- if exists is_needed ss then map baptize ths else [])) @
- (if is_FO then [] else map baptize mandatory_helpers)
- |> map_filter (Option.map snd o make_axiom ctxt false)
- end
-
-fun prepare_axiom ctxt (ax as (_, th)) = (prop_of th, make_axiom ctxt true ax)
-
-fun prepare_formulas ctxt full_types hyp_ts concl_t axioms =
- let
- val thy = ProofContext.theory_of ctxt
- val (axiom_ts, prepared_axioms) = ListPair.unzip axioms
- (* Remove existing axioms from the conjecture, as this can dramatically
- boost an ATP's performance (for some reason). *)
- val hyp_ts = hyp_ts |> filter_out (member (op aconv) axiom_ts)
- val goal_t = Logic.list_implies (hyp_ts, concl_t)
- val is_FO = Meson.is_fol_term thy goal_t
- val subs = tfree_classes_of_terms [goal_t]
- val supers = tvar_classes_of_terms axiom_ts
- val tycons = type_consts_of_terms thy (goal_t :: axiom_ts)
- (* TFrees in the conjecture; TVars in the axioms *)
- val conjectures = make_conjecture ctxt (hyp_ts @ [concl_t])
- val (axiom_names, axioms) = ListPair.unzip (map_filter I prepared_axioms)
- val helper_facts = get_helper_facts ctxt is_FO full_types conjectures axioms
- val (supers', arity_clauses) = make_arity_clauses thy tycons supers
- val class_rel_clauses = make_class_rel_clauses thy subs supers'
- in
- (axiom_names |> map single |> Vector.fromList,
- (conjectures, axioms, helper_facts, class_rel_clauses, arity_clauses))
- end
-
-fun wrap_type ty t = ATerm ((type_wrapper_name, type_wrapper_name), [ty, t])
-
-fun fo_term_for_combtyp (CombTVar name) = ATerm (name, [])
- | fo_term_for_combtyp (CombTFree name) = ATerm (name, [])
- | fo_term_for_combtyp (CombType (name, tys)) =
- ATerm (name, map fo_term_for_combtyp tys)
-
-fun fo_literal_for_type_literal (TyLitVar (class, name)) =
- (true, ATerm (class, [ATerm (name, [])]))
- | fo_literal_for_type_literal (TyLitFree (class, name)) =
- (true, ATerm (class, [ATerm (name, [])]))
-
-fun formula_for_fo_literal (pos, t) = AAtom t |> not pos ? mk_anot
-
-fun fo_term_for_combterm full_types =
- let
- fun aux top_level u =
- let
- val (head, args) = strip_combterm_comb u
- val (x, ty_args) =
- case head of
- CombConst (name as (s, s'), _, ty_args) =>
- let val ty_args = if full_types then [] else ty_args in
- if s = "equal" then
- if top_level andalso length args = 2 then (name, [])
- else (("c_fequal", @{const_name Metis.fequal}), ty_args)
- else if top_level then
- case s of
- "c_False" => (("$false", s'), [])
- | "c_True" => (("$true", s'), [])
- | _ => (name, ty_args)
- else
- (name, ty_args)
- end
- | CombVar (name, _) => (name, [])
- | CombApp _ => raise Fail "impossible \"CombApp\""
- val t = ATerm (x, map fo_term_for_combtyp ty_args @
- map (aux false) args)
- in
- if full_types then wrap_type (fo_term_for_combtyp (combtyp_of u)) t else t
- end
- in aux true end
-
-fun formula_for_combformula full_types =
- let
- fun aux (AQuant (q, xs, phi)) = AQuant (q, xs, aux phi)
- | aux (AConn (c, phis)) = AConn (c, map aux phis)
- | aux (AAtom tm) = AAtom (fo_term_for_combterm full_types tm)
- in aux end
-
-fun formula_for_axiom full_types
- ({combformula, ctypes_sorts, ...} : fol_formula) =
- mk_ahorn (map (formula_for_fo_literal o fo_literal_for_type_literal)
- (type_literals_for_types ctypes_sorts))
- (formula_for_combformula full_types combformula)
-
-fun problem_line_for_fact prefix full_types (formula as {name, kind, ...}) =
- Fof (prefix ^ ascii_of name, kind, formula_for_axiom full_types formula)
-
-fun problem_line_for_class_rel_clause (ClassRelClause {name, subclass,
- superclass, ...}) =
- let val ty_arg = ATerm (("T", "T"), []) in
- Fof (class_rel_clause_prefix ^ ascii_of name, Axiom,
- AConn (AImplies, [AAtom (ATerm (subclass, [ty_arg])),
- AAtom (ATerm (superclass, [ty_arg]))]))
- end
-
-fun fo_literal_for_arity_literal (TConsLit (c, t, args)) =
- (true, ATerm (c, [ATerm (t, map (fn arg => ATerm (arg, [])) args)]))
- | fo_literal_for_arity_literal (TVarLit (c, sort)) =
- (false, ATerm (c, [ATerm (sort, [])]))
-
-fun problem_line_for_arity_clause (ArityClause {name, conclLit, premLits,
- ...}) =
- Fof (arity_clause_prefix ^ ascii_of name, Axiom,
- mk_ahorn (map (formula_for_fo_literal o apfst not
- o fo_literal_for_arity_literal) premLits)
- (formula_for_fo_literal
- (fo_literal_for_arity_literal conclLit)))
-
-fun problem_line_for_conjecture full_types
- ({name, kind, combformula, ...} : fol_formula) =
- Fof (conjecture_prefix ^ name, kind,
- formula_for_combformula full_types combformula)
-
-fun free_type_literals_for_conjecture ({ctypes_sorts, ...} : fol_formula) =
- map fo_literal_for_type_literal (type_literals_for_types ctypes_sorts)
-
-fun problem_line_for_free_type j lit =
- Fof (tfree_prefix ^ string_of_int j, Hypothesis, formula_for_fo_literal lit)
-fun problem_lines_for_free_types conjectures =
- let
- val litss = map free_type_literals_for_conjecture conjectures
- val lits = fold (union (op =)) litss []
- in map2 problem_line_for_free_type (0 upto length lits - 1) lits end
-
-(** "hBOOL" and "hAPP" **)
-
-type const_info = {min_arity: int, max_arity: int, sub_level: bool}
-
-fun consider_term top_level (ATerm ((s, _), ts)) =
- (if is_atp_variable s then
- I
- else
- let val n = length ts in
- Symtab.map_default
- (s, {min_arity = n, max_arity = 0, sub_level = false})
- (fn {min_arity, max_arity, sub_level} =>
- {min_arity = Int.min (n, min_arity),
- max_arity = Int.max (n, max_arity),
- sub_level = sub_level orelse not top_level})
- end)
- #> fold (consider_term (top_level andalso s = type_wrapper_name)) ts
-fun consider_formula (AQuant (_, _, phi)) = consider_formula phi
- | consider_formula (AConn (_, phis)) = fold consider_formula phis
- | consider_formula (AAtom tm) = consider_term true tm
-
-fun consider_problem_line (Fof (_, _, phi)) = consider_formula phi
-fun consider_problem problem = fold (fold consider_problem_line o snd) problem
-
-fun const_table_for_problem explicit_apply problem =
- if explicit_apply then NONE
- else SOME (Symtab.empty |> consider_problem problem)
-
-fun min_arity_of thy full_types NONE s =
- (if s = "equal" orelse s = type_wrapper_name orelse
- String.isPrefix type_const_prefix s orelse
- String.isPrefix class_prefix s then
- 16383 (* large number *)
- else if full_types then
- 0
- else case strip_prefix_and_unascii const_prefix s of
- SOME s' => num_type_args thy (invert_const s')
- | NONE => 0)
- | min_arity_of _ _ (SOME the_const_tab) s =
- case Symtab.lookup the_const_tab s of
- SOME ({min_arity, ...} : const_info) => min_arity
- | NONE => 0
-
-fun full_type_of (ATerm ((s, _), [ty, _])) =
- if s = type_wrapper_name then ty else raise Fail "expected type wrapper"
- | full_type_of _ = raise Fail "expected type wrapper"
-
-fun list_hAPP_rev _ t1 [] = t1
- | list_hAPP_rev NONE t1 (t2 :: ts2) =
- ATerm (`I "hAPP", [list_hAPP_rev NONE t1 ts2, t2])
- | list_hAPP_rev (SOME ty) t1 (t2 :: ts2) =
- let val ty' = ATerm (`make_fixed_type_const @{type_name fun},
- [full_type_of t2, ty]) in
- ATerm (`I "hAPP", [wrap_type ty' (list_hAPP_rev (SOME ty') t1 ts2), t2])
- end
-
-fun repair_applications_in_term thy full_types const_tab =
- let
- fun aux opt_ty (ATerm (name as (s, _), ts)) =
- if s = type_wrapper_name then
- case ts of
- [t1, t2] => ATerm (name, [aux NONE t1, aux (SOME t1) t2])
- | _ => raise Fail "malformed type wrapper"
- else
- let
- val ts = map (aux NONE) ts
- val (ts1, ts2) = chop (min_arity_of thy full_types const_tab s) ts
- in list_hAPP_rev opt_ty (ATerm (name, ts1)) (rev ts2) end
- in aux NONE end
-
-fun boolify t = ATerm (`I "hBOOL", [t])
-
-(* True if the constant ever appears outside of the top-level position in
- literals, or if it appears with different arities (e.g., because of different
- type instantiations). If false, the constant always receives all of its
- arguments and is used as a predicate. *)
-fun is_predicate NONE s =
- s = "equal" orelse s = "$false" orelse s = "$true" orelse
- String.isPrefix type_const_prefix s orelse String.isPrefix class_prefix s
- | is_predicate (SOME the_const_tab) s =
- case Symtab.lookup the_const_tab s of
- SOME {min_arity, max_arity, sub_level} =>
- not sub_level andalso min_arity = max_arity
- | NONE => false
-
-fun repair_predicates_in_term const_tab (t as ATerm ((s, _), ts)) =
- if s = type_wrapper_name then
- case ts of
- [_, t' as ATerm ((s', _), _)] =>
- if is_predicate const_tab s' then t' else boolify t
- | _ => raise Fail "malformed type wrapper"
- else
- t |> not (is_predicate const_tab s) ? boolify
-
-fun close_universally phi =
- let
- fun term_vars bounds (ATerm (name as (s, _), tms)) =
- (is_atp_variable s andalso not (member (op =) bounds name))
- ? insert (op =) name
- #> fold (term_vars bounds) tms
- fun formula_vars bounds (AQuant (_, xs, phi)) =
- formula_vars (xs @ bounds) phi
- | formula_vars bounds (AConn (_, phis)) = fold (formula_vars bounds) phis
- | formula_vars bounds (AAtom tm) = term_vars bounds tm
- in
- case formula_vars [] phi [] of [] => phi | xs => AQuant (AForall, xs, phi)
- end
-
-fun repair_formula thy explicit_forall full_types const_tab =
- let
- fun aux (AQuant (q, xs, phi)) = AQuant (q, xs, aux phi)
- | aux (AConn (c, phis)) = AConn (c, map aux phis)
- | aux (AAtom tm) =
- AAtom (tm |> repair_applications_in_term thy full_types const_tab
- |> repair_predicates_in_term const_tab)
- in aux #> explicit_forall ? close_universally end
-
-fun repair_problem_line thy explicit_forall full_types const_tab
- (Fof (ident, kind, phi)) =
- Fof (ident, kind, repair_formula thy explicit_forall full_types const_tab phi)
-fun repair_problem_with_const_table thy =
- map o apsnd o map ooo repair_problem_line thy
-
-fun repair_problem thy explicit_forall full_types explicit_apply problem =
- repair_problem_with_const_table thy explicit_forall full_types
- (const_table_for_problem explicit_apply problem) problem
-
-fun prepare_problem ctxt readable_names explicit_forall full_types
- explicit_apply hyp_ts concl_t axioms =
- let
- val thy = ProofContext.theory_of ctxt
- val (axiom_names, (conjectures, axioms, helper_facts, class_rel_clauses,
- arity_clauses)) =
- prepare_formulas ctxt full_types hyp_ts concl_t axioms
- val axiom_lines = map (problem_line_for_fact axiom_prefix full_types) axioms
- val helper_lines =
- map (problem_line_for_fact helper_prefix full_types) helper_facts
- val conjecture_lines =
- map (problem_line_for_conjecture full_types) conjectures
- val tfree_lines = problem_lines_for_free_types conjectures
- val class_rel_lines =
- map problem_line_for_class_rel_clause class_rel_clauses
- val arity_lines = map problem_line_for_arity_clause arity_clauses
- (* Reordering these might or might not confuse the proof reconstruction
- code or the SPASS Flotter hack. *)
- val problem =
- [("Relevant facts", axiom_lines),
- ("Class relationships", class_rel_lines),
- ("Arity declarations", arity_lines),
- ("Helper facts", helper_lines),
- ("Conjectures", conjecture_lines),
- ("Type variables", tfree_lines)]
- |> repair_problem thy explicit_forall full_types explicit_apply
- val (problem, pool) = nice_atp_problem readable_names problem
- val conjecture_offset =
- length axiom_lines + length class_rel_lines + length arity_lines
- + length helper_lines
- in
- (problem,
- case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty,
- conjecture_offset, axiom_names)
- end
-
-end;
--- a/src/HOL/Tools/semiring_normalizer.ML Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/Tools/semiring_normalizer.ML Sun Oct 24 15:11:24 2010 -0700
@@ -826,7 +826,7 @@
end;
val nat_exp_ss =
- HOL_basic_ss addsimps (@{thms nat_number} @ @{thms nat_arith} @ @{thms arith_simps} @ @{thms rel_simps})
+ HOL_basic_ss addsimps (@{thms eval_nat_numeral} @ @{thms nat_arith} @ @{thms arith_simps} @ @{thms rel_simps})
addsimps [@{thm Let_def}, @{thm if_False}, @{thm if_True}, @{thm Nat.add_0}, @{thm add_Suc}];
fun simple_cterm_ord t u = Term_Ord.term_ord (term_of t, term_of u) = LESS;
--- a/src/HOL/ex/HarmonicSeries.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/ex/HarmonicSeries.thy Sun Oct 24 15:11:24 2010 -0700
@@ -153,7 +153,7 @@
(auto simp: atLeastSucAtMost_greaterThanAtMost)
also have
"\<dots> = ((\<Sum>n\<in>{2..2::nat}. 1/real n) + 1/(real (1::nat)))"
- by (simp add: nat_number)
+ by (simp add: eval_nat_numeral)
also have
"\<dots> = 1/(real (2::nat)) + 1/(real (1::nat))" by simp
finally have
--- a/src/HOL/ex/NatSum.thy Sun Oct 24 03:43:12 2010 -0700
+++ b/src/HOL/ex/NatSum.thy Sun Oct 24 15:11:24 2010 -0700
@@ -70,9 +70,9 @@
have "(\<Sum>i = 0..Suc n. i)^2 =
(\<Sum>i = 0..n. i^3) + (2*(\<Sum>i = 0..n. i)*(n+1) + (n+1)^2)"
(is "_ = ?A + ?B")
- using Suc by(simp add:nat_number)
+ using Suc by(simp add:eval_nat_numeral)
also have "?B = (n+1)^3"
- using sum_of_naturals by(simp add:nat_number)
+ using sum_of_naturals by(simp add:eval_nat_numeral)
also have "?A + (n+1)^3 = (\<Sum>i=0..Suc n. i^3)" by simp
finally show ?case .
qed