--- a/Admin/Benchmarks/HOL-datatype/ROOT.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/Admin/Benchmarks/HOL-datatype/ROOT.ML Wed Sep 01 17:19:47 2010 +0200
@@ -5,11 +5,11 @@
val tests = ["Brackin", "Instructions", "SML", "Verilog"];
-Unsynchronized.set timing;
+timing := true;
-warning "\nset quick_and_dirty\n"; Unsynchronized.set quick_and_dirty;
+warning "\nset quick_and_dirty\n"; quick_and_dirty := true;
use_thys tests;
-warning "\nreset quick_and_dirty\n"; Unsynchronized.reset quick_and_dirty;
+warning "\nreset quick_and_dirty\n"; quick_and_dirty := false;
List.app Thy_Info.remove_thy tests;
use_thys tests;
--- a/Admin/Benchmarks/HOL-record/ROOT.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/Admin/Benchmarks/HOL-record/ROOT.ML Wed Sep 01 17:19:47 2010 +0200
@@ -5,10 +5,10 @@
val tests = ["RecordBenchmark"];
-Unsynchronized.set timing;
+timing := true;
-warning "\nset quick_and_dirty\n"; Unsynchronized.set quick_and_dirty;
+warning "\nset quick_and_dirty\n"; quick_and_dirty := true;
use_thys tests;
-warning "\nreset quick_and_dirty\n"; Unsynchronized.reset quick_and_dirty;
+warning "\nreset quick_and_dirty\n"; quick_and_dirty := false;
use_thys tests;
--- a/Admin/Benchmarks/HOL-record/RecordBenchmark.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/Admin/Benchmarks/HOL-record/RecordBenchmark.thy Wed Sep 01 17:19:47 2010 +0200
@@ -8,7 +8,7 @@
imports Main
begin
-ML {* Unsynchronized.set Record.timing *}
+ML {* Record.timing := true *}
record many_A =
A000::nat
--- a/Admin/update-keywords Thu Aug 26 13:15:37 2010 +0200
+++ b/Admin/update-keywords Wed Sep 01 17:19:47 2010 +0200
@@ -11,9 +11,9 @@
cd "$ISABELLE_HOME/etc"
isabelle keywords \
- "$LOG/Pure.gz" "$LOG/Pure-ProofGeneral.gz" "$LOG/HOL.gz" "$LOG/HOLCF.gz" \
- "$LOG/HOL-Boogie.gz" "$LOG/HOL-Nominal.gz" "$LOG/HOL-Statespace.gz"
+ "$LOG/Pure.gz" "$LOG/HOL.gz" "$LOG/HOLCF.gz" "$LOG/HOL-Boogie.gz" \
+ "$LOG/HOL-Nominal.gz" "$LOG/HOL-Statespace.gz"
isabelle keywords -k ZF \
- "$LOG/Pure.gz" "$LOG/Pure-ProofGeneral.gz" "$LOG/FOL.gz" "$LOG/ZF.gz"
+ "$LOG/Pure.gz" "$LOG/FOL.gz" "$LOG/ZF.gz"
--- a/NEWS Thu Aug 26 13:15:37 2010 +0200
+++ b/NEWS Wed Sep 01 17:19:47 2010 +0200
@@ -23,6 +23,28 @@
at the cost of clarity of file dependencies. Recall that Isabelle/ML
files exclusively use the .ML extension. Minor INCOMPATIBILTY.
+* Various options that affect document antiquotations are now properly
+handled within the context via configuration options, instead of
+unsynchronized references. There are both ML Config.T entities and
+Isar declaration attributes to access these.
+
+ ML: Isar:
+
+ Thy_Output.display thy_output_display
+ Thy_Output.quotes thy_output_quotes
+ Thy_Output.indent thy_output_indent
+ Thy_Output.source thy_output_source
+ Thy_Output.break thy_output_break
+
+Note that the corresponding "..._default" references may be only
+changed globally at the ROOT session setup, but *not* within a theory.
+
+* ML structure Unsynchronized never opened, not even in Isar
+interaction mode as before. Old Unsynchronized.set etc. have been
+discontinued -- use plain := instead. This should be *rare* anyway,
+since modern tools always work via official context data, notably
+configuration options.
+
*** Pure ***
@@ -40,6 +62,13 @@
*** HOL ***
+* Renamed class eq and constant eq (for code generation) to class equal
+and constant equal, plus renaming of related facts and various tuning.
+INCOMPATIBILITY.
+
+* Scala (2.8 or higher) has been added to the target languages of
+the code generator.
+
* Dropped type classes mult_mono and mult_mono1. INCOMPATIBILITY.
* Theory SetsAndFunctions has been split into Function_Algebras and Set_Algebras;
@@ -104,6 +133,10 @@
Trueprop ~> HOL.Trueprop
True ~> HOL.True
False ~> HOL.False
+ op & ~> HOL.conj
+ op | ~> HOL.disj
+ op --> ~> HOL.implies
+ op = ~> HOL.eq
Not ~> HOL.Not
The ~> HOL.The
All ~> HOL.All
--- a/doc-src/Classes/Thy/Classes.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Classes/Thy/Classes.thy Wed Sep 01 17:19:47 2010 +0200
@@ -8,14 +8,14 @@
Type classes were introduced by Wadler and Blott \cite{wadler89how}
into the Haskell language to allow for a reasonable implementation
of overloading\footnote{throughout this tutorial, we are referring
- to classical Haskell 1.0 type classes, not considering
- later additions in expressiveness}.
- As a canonical example, a polymorphic equality function
- @{text "eq \<Colon> \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"} which is overloaded on different
- types for @{text "\<alpha>"}, which is achieved by splitting introduction
- of the @{text eq} function from its overloaded definitions by means
- of @{text class} and @{text instance} declarations:
- \footnote{syntax here is a kind of isabellized Haskell}
+ to classical Haskell 1.0 type classes, not considering later
+ additions in expressiveness}. As a canonical example, a polymorphic
+ equality function @{text "eq \<Colon> \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool"} which is overloaded on
+ different types for @{text "\<alpha>"}, which is achieved by splitting
+ introduction of the @{text eq} function from its overloaded
+ definitions by means of @{text class} and @{text instance}
+ declarations: \footnote{syntax here is a kind of isabellized
+ Haskell}
\begin{quote}
@@ -41,14 +41,14 @@
these annotations are assertions that a particular polymorphic type
provides definitions for overloaded functions.
- Indeed, type classes not only allow for simple overloading
- but form a generic calculus, an instance of order-sorted
- algebra \cite{nipkow-sorts93,Nipkow-Prehofer:1993,Wenzel:1997:TPHOL}.
+ Indeed, type classes not only allow for simple overloading but form
+ a generic calculus, an instance of order-sorted algebra
+ \cite{nipkow-sorts93,Nipkow-Prehofer:1993,Wenzel:1997:TPHOL}.
- From a software engineering point of view, type classes
- roughly correspond to interfaces in object-oriented languages like Java;
- so, it is naturally desirable that type classes do not only
- provide functions (class parameters) but also state specifications
+ From a software engineering point of view, type classes roughly
+ correspond to interfaces in object-oriented languages like Java; so,
+ it is naturally desirable that type classes do not only provide
+ functions (class parameters) but also state specifications
implementations must obey. For example, the @{text "class eq"}
above could be given the following specification, demanding that
@{text "class eq"} is an equivalence relation obeying reflexivity,
@@ -65,11 +65,10 @@
\end{quote}
- \noindent From a theoretical point of view, type classes are lightweight
- modules; Haskell type classes may be emulated by
- SML functors \cite{classes_modules}.
- Isabelle/Isar offers a discipline of type classes which brings
- all those aspects together:
+ \noindent From a theoretical point of view, type classes are
+ lightweight modules; Haskell type classes may be emulated by SML
+ functors \cite{classes_modules}. Isabelle/Isar offers a discipline
+ of type classes which brings all those aspects together:
\begin{enumerate}
\item specifying abstract parameters together with
@@ -81,15 +80,15 @@
locales \cite{kammueller-locales}.
\end{enumerate}
- \noindent Isar type classes also directly support code generation
- in a Haskell like fashion. Internally, they are mapped to more primitive
- Isabelle concepts \cite{Haftmann-Wenzel:2006:classes}.
+ \noindent Isar type classes also directly support code generation in
+ a Haskell like fashion. Internally, they are mapped to more
+ primitive Isabelle concepts \cite{Haftmann-Wenzel:2006:classes}.
- This tutorial demonstrates common elements of structured specifications
- and abstract reasoning with type classes by the algebraic hierarchy of
- semigroups, monoids and groups. Our background theory is that of
- Isabelle/HOL \cite{isa-tutorial}, for which some
- familiarity is assumed.
+ This tutorial demonstrates common elements of structured
+ specifications and abstract reasoning with type classes by the
+ algebraic hierarchy of semigroups, monoids and groups. Our
+ background theory is that of Isabelle/HOL \cite{isa-tutorial}, for
+ which some familiarity is assumed.
*}
section {* A simple algebra example \label{sec:example} *}
@@ -107,25 +106,24 @@
assumes assoc: "(x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"
text {*
- \noindent This @{command class} specification consists of two
- parts: the \qn{operational} part names the class parameter
- (@{element "fixes"}), the \qn{logical} part specifies properties on them
- (@{element "assumes"}). The local @{element "fixes"} and
- @{element "assumes"} are lifted to the theory toplevel,
- yielding the global
+ \noindent This @{command class} specification consists of two parts:
+ the \qn{operational} part names the class parameter (@{element
+ "fixes"}), the \qn{logical} part specifies properties on them
+ (@{element "assumes"}). The local @{element "fixes"} and @{element
+ "assumes"} are lifted to the theory toplevel, yielding the global
parameter @{term [source] "mult \<Colon> \<alpha>\<Colon>semigroup \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>"} and the
- global theorem @{fact "semigroup.assoc:"}~@{prop [source] "\<And>x y
- z \<Colon> \<alpha>\<Colon>semigroup. (x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"}.
+ global theorem @{fact "semigroup.assoc:"}~@{prop [source] "\<And>x y z \<Colon>
+ \<alpha>\<Colon>semigroup. (x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"}.
*}
subsection {* Class instantiation \label{sec:class_inst} *}
text {*
- The concrete type @{typ int} is made a @{class semigroup}
- instance by providing a suitable definition for the class parameter
- @{text "(\<otimes>)"} and a proof for the specification of @{fact assoc}.
- This is accomplished by the @{command instantiation} target:
+ The concrete type @{typ int} is made a @{class semigroup} instance
+ by providing a suitable definition for the class parameter @{text
+ "(\<otimes>)"} and a proof for the specification of @{fact assoc}. This is
+ accomplished by the @{command instantiation} target:
*}
instantiation %quote int :: semigroup
@@ -143,22 +141,22 @@
end %quote
text {*
- \noindent @{command instantiation} defines class parameters
- at a particular instance using common specification tools (here,
- @{command definition}). The concluding @{command instance}
- opens a proof that the given parameters actually conform
- to the class specification. Note that the first proof step
- is the @{method default} method,
- which for such instance proofs maps to the @{method intro_classes} method.
- This reduces an instance judgement to the relevant primitive
- proof goals; typically it is the first method applied
- in an instantiation proof.
+ \noindent @{command instantiation} defines class parameters at a
+ particular instance using common specification tools (here,
+ @{command definition}). The concluding @{command instance} opens a
+ proof that the given parameters actually conform to the class
+ specification. Note that the first proof step is the @{method
+ default} method, which for such instance proofs maps to the @{method
+ intro_classes} method. This reduces an instance judgement to the
+ relevant primitive proof goals; typically it is the first method
+ applied in an instantiation proof.
- From now on, the type-checker will consider @{typ int}
- as a @{class semigroup} automatically, i.e.\ any general results
- are immediately available on concrete instances.
+ From now on, the type-checker will consider @{typ int} as a @{class
+ semigroup} automatically, i.e.\ any general results are immediately
+ available on concrete instances.
- \medskip Another instance of @{class semigroup} yields the natural numbers:
+ \medskip Another instance of @{class semigroup} yields the natural
+ numbers:
*}
instantiation %quote nat :: semigroup
@@ -177,21 +175,20 @@
end %quote
text {*
- \noindent Note the occurence of the name @{text mult_nat}
- in the primrec declaration; by default, the local name of
- a class operation @{text f} to be instantiated on type constructor
- @{text \<kappa>} is mangled as @{text f_\<kappa>}. In case of uncertainty,
- these names may be inspected using the @{command "print_context"} command
- or the corresponding ProofGeneral button.
+ \noindent Note the occurence of the name @{text mult_nat} in the
+ primrec declaration; by default, the local name of a class operation
+ @{text f} to be instantiated on type constructor @{text \<kappa>} is
+ mangled as @{text f_\<kappa>}. In case of uncertainty, these names may be
+ inspected using the @{command "print_context"} command or the
+ corresponding ProofGeneral button.
*}
subsection {* Lifting and parametric types *}
text {*
- Overloaded definitions given at a class instantiation
- may include recursion over the syntactic structure of types.
- As a canonical example, we model product semigroups
- using our simple algebra:
+ Overloaded definitions given at a class instantiation may include
+ recursion over the syntactic structure of types. As a canonical
+ example, we model product semigroups using our simple algebra:
*}
instantiation %quote prod :: (semigroup, semigroup) semigroup
@@ -211,21 +208,19 @@
text {*
\noindent Associativity of product semigroups is established using
the definition of @{text "(\<otimes>)"} on products and the hypothetical
- associativity of the type components; these hypotheses
- are legitimate due to the @{class semigroup} constraints imposed
- on the type components by the @{command instance} proposition.
- Indeed, this pattern often occurs with parametric types
- and type classes.
+ associativity of the type components; these hypotheses are
+ legitimate due to the @{class semigroup} constraints imposed on the
+ type components by the @{command instance} proposition. Indeed,
+ this pattern often occurs with parametric types and type classes.
*}
subsection {* Subclassing *}
text {*
- We define a subclass @{text monoidl} (a semigroup with a left-hand neutral)
- by extending @{class semigroup}
- with one additional parameter @{text neutral} together
- with its characteristic property:
+ We define a subclass @{text monoidl} (a semigroup with a left-hand
+ neutral) by extending @{class semigroup} with one additional
+ parameter @{text neutral} together with its characteristic property:
*}
class %quote monoidl = semigroup +
@@ -233,10 +228,10 @@
assumes neutl: "\<one> \<otimes> x = x"
text {*
- \noindent Again, we prove some instances, by
- providing suitable parameter definitions and proofs for the
- additional specifications. Observe that instantiations
- for types with the same arity may be simultaneous:
+ \noindent Again, we prove some instances, by providing suitable
+ parameter definitions and proofs for the additional specifications.
+ Observe that instantiations for types with the same arity may be
+ simultaneous:
*}
instantiation %quote nat and int :: monoidl
@@ -309,8 +304,8 @@
end %quote
text {*
- \noindent To finish our small algebra example, we add a @{text group} class
- with a corresponding instance:
+ \noindent To finish our small algebra example, we add a @{text
+ group} class with a corresponding instance:
*}
class %quote group = monoidl +
@@ -338,9 +333,9 @@
subsection {* A look behind the scenes *}
text {*
- The example above gives an impression how Isar type classes work
- in practice. As stated in the introduction, classes also provide
- a link to Isar's locale system. Indeed, the logical core of a class
+ The example above gives an impression how Isar type classes work in
+ practice. As stated in the introduction, classes also provide a
+ link to Isar's locale system. Indeed, the logical core of a class
is nothing other than a locale:
*}
@@ -402,13 +397,14 @@
qed
text {*
- \noindent Here the \qt{@{keyword "in"} @{class group}} target specification
- indicates that the result is recorded within that context for later
- use. This local theorem is also lifted to the global one @{fact
- "group.left_cancel:"} @{prop [source] "\<And>x y z \<Colon> \<alpha>\<Colon>group. x \<otimes> y = x \<otimes>
- z \<longleftrightarrow> y = z"}. Since type @{text "int"} has been made an instance of
- @{text "group"} before, we may refer to that fact as well: @{prop
- [source] "\<And>x y z \<Colon> int. x \<otimes> y = x \<otimes> z \<longleftrightarrow> y = z"}.
+ \noindent Here the \qt{@{keyword "in"} @{class group}} target
+ specification indicates that the result is recorded within that
+ context for later use. This local theorem is also lifted to the
+ global one @{fact "group.left_cancel:"} @{prop [source] "\<And>x y z \<Colon>
+ \<alpha>\<Colon>group. x \<otimes> y = x \<otimes> z \<longleftrightarrow> y = z"}. Since type @{text "int"} has been
+ made an instance of @{text "group"} before, we may refer to that
+ fact as well: @{prop [source] "\<And>x y z \<Colon> int. x \<otimes> y = x \<otimes> z \<longleftrightarrow> y =
+ z"}.
*}
@@ -424,15 +420,14 @@
text {*
\noindent If the locale @{text group} is also a class, this local
- definition is propagated onto a global definition of
- @{term [source] "pow_nat \<Colon> nat \<Rightarrow> \<alpha>\<Colon>monoid \<Rightarrow> \<alpha>\<Colon>monoid"}
- with corresponding theorems
+ definition is propagated onto a global definition of @{term [source]
+ "pow_nat \<Colon> nat \<Rightarrow> \<alpha>\<Colon>monoid \<Rightarrow> \<alpha>\<Colon>monoid"} with corresponding theorems
@{thm pow_nat.simps [no_vars]}.
- \noindent As you can see from this example, for local
- definitions you may use any specification tool
- which works together with locales, such as Krauss's recursive function package
+ \noindent As you can see from this example, for local definitions
+ you may use any specification tool which works together with
+ locales, such as Krauss's recursive function package
\cite{krauss2006}.
*}
@@ -440,19 +435,17 @@
subsection {* A functor analogy *}
text {*
- We introduced Isar classes by analogy to type classes in
- functional programming; if we reconsider this in the
- context of what has been said about type classes and locales,
- we can drive this analogy further by stating that type
- classes essentially correspond to functors that have
- a canonical interpretation as type classes.
- There is also the possibility of other interpretations.
- For example, @{text list}s also form a monoid with
- @{text append} and @{term "[]"} as operations, but it
- seems inappropriate to apply to lists
- the same operations as for genuinely algebraic types.
- In such a case, we can simply make a particular interpretation
- of monoids for lists:
+ We introduced Isar classes by analogy to type classes in functional
+ programming; if we reconsider this in the context of what has been
+ said about type classes and locales, we can drive this analogy
+ further by stating that type classes essentially correspond to
+ functors that have a canonical interpretation as type classes.
+ There is also the possibility of other interpretations. For
+ example, @{text list}s also form a monoid with @{text append} and
+ @{term "[]"} as operations, but it seems inappropriate to apply to
+ lists the same operations as for genuinely algebraic types. In such
+ a case, we can simply make a particular interpretation of monoids
+ for lists:
*}
interpretation %quote list_monoid: monoid append "[]"
@@ -510,12 +503,10 @@
qed
text {*
- The logical proof is carried out on the locale level.
- Afterwards it is propagated
- to the type system, making @{text group} an instance of
- @{text monoid} by adding an additional edge
- to the graph of subclass relations
- (\figref{fig:subclass}).
+ The logical proof is carried out on the locale level. Afterwards it
+ is propagated to the type system, making @{text group} an instance
+ of @{text monoid} by adding an additional edge to the graph of
+ subclass relations (\figref{fig:subclass}).
\begin{figure}[htbp]
\begin{center}
@@ -547,8 +538,8 @@
\end{center}
\end{figure}
- For illustration, a derived definition
- in @{text group} using @{text pow_nat}
+ For illustration, a derived definition in @{text group} using @{text
+ pow_nat}
*}
definition %quote (in group) pow_int :: "int \<Rightarrow> \<alpha> \<Rightarrow> \<alpha>" where
@@ -557,17 +548,17 @@
else (pow_nat (nat (- k)) x)\<div>)"
text {*
- \noindent yields the global definition of
- @{term [source] "pow_int \<Colon> int \<Rightarrow> \<alpha>\<Colon>group \<Rightarrow> \<alpha>\<Colon>group"}
- with the corresponding theorem @{thm pow_int_def [no_vars]}.
+ \noindent yields the global definition of @{term [source] "pow_int \<Colon>
+ int \<Rightarrow> \<alpha>\<Colon>group \<Rightarrow> \<alpha>\<Colon>group"} with the corresponding theorem @{thm
+ pow_int_def [no_vars]}.
*}
subsection {* A note on syntax *}
text {*
- As a convenience, class context syntax allows references
- to local class operations and their global counterparts
- uniformly; type inference resolves ambiguities. For example:
+ As a convenience, class context syntax allows references to local
+ class operations and their global counterparts uniformly; type
+ inference resolves ambiguities. For example:
*}
context %quote semigroup
@@ -581,11 +572,11 @@
term %quote "x \<otimes> y" -- {* example 3 *}
text {*
- \noindent Here in example 1, the term refers to the local class operation
- @{text "mult [\<alpha>]"}, whereas in example 2 the type constraint
- enforces the global class operation @{text "mult [nat]"}.
- In the global context in example 3, the reference is
- to the polymorphic global class operation @{text "mult [?\<alpha> \<Colon> semigroup]"}.
+ \noindent Here in example 1, the term refers to the local class
+ operation @{text "mult [\<alpha>]"}, whereas in example 2 the type
+ constraint enforces the global class operation @{text "mult [nat]"}.
+ In the global context in example 3, the reference is to the
+ polymorphic global class operation @{text "mult [?\<alpha> \<Colon> semigroup]"}.
*}
section {* Further issues *}
@@ -593,16 +584,14 @@
subsection {* Type classes and code generation *}
text {*
- Turning back to the first motivation for type classes,
- namely overloading, it is obvious that overloading
- stemming from @{command class} statements and
- @{command instantiation}
- targets naturally maps to Haskell type classes.
- The code generator framework \cite{isabelle-codegen}
- takes this into account. If the target language (e.g.~SML)
- lacks type classes, then they
- are implemented by an explicit dictionary construction.
- As example, let's go back to the power function:
+ Turning back to the first motivation for type classes, namely
+ overloading, it is obvious that overloading stemming from @{command
+ class} statements and @{command instantiation} targets naturally
+ maps to Haskell type classes. The code generator framework
+ \cite{isabelle-codegen} takes this into account. If the target
+ language (e.g.~SML) lacks type classes, then they are implemented by
+ an explicit dictionary construction. As example, let's go back to
+ the power function:
*}
definition %quote example :: int where
@@ -619,11 +608,18 @@
*}
text %quote {*@{code_stmts example (SML)}*}
+text {*
+ \noindent In Scala, implicts are used as dictionaries:
+*}
+(*<*)code_include %invisible Scala "Natural" -(*>*)
+text %quote {*@{code_stmts example (Scala)}*}
+
+
subsection {* Inspecting the type class universe *}
text {*
- To facilitate orientation in complex subclass structures,
- two diagnostics commands are provided:
+ To facilitate orientation in complex subclass structures, two
+ diagnostics commands are provided:
\begin{description}
--- a/doc-src/Classes/Thy/document/Classes.tex Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Classes/Thy/document/Classes.tex Wed Sep 01 17:19:47 2010 +0200
@@ -26,14 +26,14 @@
Type classes were introduced by Wadler and Blott \cite{wadler89how}
into the Haskell language to allow for a reasonable implementation
of overloading\footnote{throughout this tutorial, we are referring
- to classical Haskell 1.0 type classes, not considering
- later additions in expressiveness}.
- As a canonical example, a polymorphic equality function
- \isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} which is overloaded on different
- types for \isa{{\isasymalpha}}, which is achieved by splitting introduction
- of the \isa{eq} function from its overloaded definitions by means
- of \isa{class} and \isa{instance} declarations:
- \footnote{syntax here is a kind of isabellized Haskell}
+ to classical Haskell 1.0 type classes, not considering later
+ additions in expressiveness}. As a canonical example, a polymorphic
+ equality function \isa{eq\ {\isasymColon}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ bool} which is overloaded on
+ different types for \isa{{\isasymalpha}}, which is achieved by splitting
+ introduction of the \isa{eq} function from its overloaded
+ definitions by means of \isa{class} and \isa{instance}
+ declarations: \footnote{syntax here is a kind of isabellized
+ Haskell}
\begin{quote}
@@ -59,14 +59,14 @@
these annotations are assertions that a particular polymorphic type
provides definitions for overloaded functions.
- Indeed, type classes not only allow for simple overloading
- but form a generic calculus, an instance of order-sorted
- algebra \cite{nipkow-sorts93,Nipkow-Prehofer:1993,Wenzel:1997:TPHOL}.
+ Indeed, type classes not only allow for simple overloading but form
+ a generic calculus, an instance of order-sorted algebra
+ \cite{nipkow-sorts93,Nipkow-Prehofer:1993,Wenzel:1997:TPHOL}.
- From a software engineering point of view, type classes
- roughly correspond to interfaces in object-oriented languages like Java;
- so, it is naturally desirable that type classes do not only
- provide functions (class parameters) but also state specifications
+ From a software engineering point of view, type classes roughly
+ correspond to interfaces in object-oriented languages like Java; so,
+ it is naturally desirable that type classes do not only provide
+ functions (class parameters) but also state specifications
implementations must obey. For example, the \isa{class\ eq}
above could be given the following specification, demanding that
\isa{class\ eq} is an equivalence relation obeying reflexivity,
@@ -83,11 +83,10 @@
\end{quote}
- \noindent From a theoretical point of view, type classes are lightweight
- modules; Haskell type classes may be emulated by
- SML functors \cite{classes_modules}.
- Isabelle/Isar offers a discipline of type classes which brings
- all those aspects together:
+ \noindent From a theoretical point of view, type classes are
+ lightweight modules; Haskell type classes may be emulated by SML
+ functors \cite{classes_modules}. Isabelle/Isar offers a discipline
+ of type classes which brings all those aspects together:
\begin{enumerate}
\item specifying abstract parameters together with
@@ -99,15 +98,15 @@
locales \cite{kammueller-locales}.
\end{enumerate}
- \noindent Isar type classes also directly support code generation
- in a Haskell like fashion. Internally, they are mapped to more primitive
- Isabelle concepts \cite{Haftmann-Wenzel:2006:classes}.
+ \noindent Isar type classes also directly support code generation in
+ a Haskell like fashion. Internally, they are mapped to more
+ primitive Isabelle concepts \cite{Haftmann-Wenzel:2006:classes}.
- This tutorial demonstrates common elements of structured specifications
- and abstract reasoning with type classes by the algebraic hierarchy of
- semigroups, monoids and groups. Our background theory is that of
- Isabelle/HOL \cite{isa-tutorial}, for which some
- familiarity is assumed.%
+ This tutorial demonstrates common elements of structured
+ specifications and abstract reasoning with type classes by the
+ algebraic hierarchy of semigroups, monoids and groups. Our
+ background theory is that of Isabelle/HOL \cite{isa-tutorial}, for
+ which some familiarity is assumed.%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -142,12 +141,9 @@
\endisadelimquote
%
\begin{isamarkuptext}%
-\noindent This \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} specification consists of two
- parts: the \qn{operational} part names the class parameter
- (\hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}}), the \qn{logical} part specifies properties on them
- (\hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}}). The local \hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}} and
- \hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}} are lifted to the theory toplevel,
- yielding the global
+\noindent This \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} specification consists of two parts:
+ the \qn{operational} part names the class parameter (\hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}}), the \qn{logical} part specifies properties on them
+ (\hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}}). The local \hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}} and \hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}} are lifted to the theory toplevel, yielding the global
parameter \isa{{\isachardoublequote}mult\ {\isasymColon}\ {\isasymalpha}{\isasymColon}semigroup\ {\isasymRightarrow}\ {\isasymalpha}\ {\isasymRightarrow}\ {\isasymalpha}{\isachardoublequote}} and the
global theorem \hyperlink{fact.semigroup.assoc:}{\mbox{\isa{semigroup{\isachardot}assoc{\isacharcolon}}}}~\isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ {\isasymalpha}{\isasymColon}semigroup{\isachardot}\ {\isacharparenleft}x\ {\isasymotimes}\ y{\isacharparenright}\ {\isasymotimes}\ z\ {\isacharequal}\ x\ {\isasymotimes}\ {\isacharparenleft}y\ {\isasymotimes}\ z{\isacharparenright}{\isachardoublequote}}.%
\end{isamarkuptext}%
@@ -158,10 +154,9 @@
\isamarkuptrue%
%
\begin{isamarkuptext}%
-The concrete type \isa{int} is made a \isa{semigroup}
- instance by providing a suitable definition for the class parameter
- \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} and a proof for the specification of \hyperlink{fact.assoc}{\mbox{\isa{assoc}}}.
- This is accomplished by the \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} target:%
+The concrete type \isa{int} is made a \isa{semigroup} instance
+ by providing a suitable definition for the class parameter \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} and a proof for the specification of \hyperlink{fact.assoc}{\mbox{\isa{assoc}}}. This is
+ accomplished by the \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} target:%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -204,22 +199,19 @@
\endisadelimquote
%
\begin{isamarkuptext}%
-\noindent \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} defines class parameters
- at a particular instance using common specification tools (here,
- \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}). The concluding \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}}
- opens a proof that the given parameters actually conform
- to the class specification. Note that the first proof step
- is the \hyperlink{method.default}{\mbox{\isa{default}}} method,
- which for such instance proofs maps to the \hyperlink{method.intro-classes}{\mbox{\isa{intro{\isacharunderscore}classes}}} method.
- This reduces an instance judgement to the relevant primitive
- proof goals; typically it is the first method applied
- in an instantiation proof.
+\noindent \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} defines class parameters at a
+ particular instance using common specification tools (here,
+ \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}). The concluding \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} opens a
+ proof that the given parameters actually conform to the class
+ specification. Note that the first proof step is the \hyperlink{method.default}{\mbox{\isa{default}}} method, which for such instance proofs maps to the \hyperlink{method.intro-classes}{\mbox{\isa{intro{\isacharunderscore}classes}}} method. This reduces an instance judgement to the
+ relevant primitive proof goals; typically it is the first method
+ applied in an instantiation proof.
- From now on, the type-checker will consider \isa{int}
- as a \isa{semigroup} automatically, i.e.\ any general results
- are immediately available on concrete instances.
+ From now on, the type-checker will consider \isa{int} as a \isa{semigroup} automatically, i.e.\ any general results are immediately
+ available on concrete instances.
- \medskip Another instance of \isa{semigroup} yields the natural numbers:%
+ \medskip Another instance of \isa{semigroup} yields the natural
+ numbers:%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -259,12 +251,12 @@
\endisadelimquote
%
\begin{isamarkuptext}%
-\noindent Note the occurence of the name \isa{mult{\isacharunderscore}nat}
- in the primrec declaration; by default, the local name of
- a class operation \isa{f} to be instantiated on type constructor
- \isa{{\isasymkappa}} is mangled as \isa{f{\isacharunderscore}{\isasymkappa}}. In case of uncertainty,
- these names may be inspected using the \hyperlink{command.print-context}{\mbox{\isa{\isacommand{print{\isacharunderscore}context}}}} command
- or the corresponding ProofGeneral button.%
+\noindent Note the occurence of the name \isa{mult{\isacharunderscore}nat} in the
+ primrec declaration; by default, the local name of a class operation
+ \isa{f} to be instantiated on type constructor \isa{{\isasymkappa}} is
+ mangled as \isa{f{\isacharunderscore}{\isasymkappa}}. In case of uncertainty, these names may be
+ inspected using the \hyperlink{command.print-context}{\mbox{\isa{\isacommand{print{\isacharunderscore}context}}}} command or the
+ corresponding ProofGeneral button.%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -273,10 +265,9 @@
\isamarkuptrue%
%
\begin{isamarkuptext}%
-Overloaded definitions given at a class instantiation
- may include recursion over the syntactic structure of types.
- As a canonical example, we model product semigroups
- using our simple algebra:%
+Overloaded definitions given at a class instantiation may include
+ recursion over the syntactic structure of types. As a canonical
+ example, we model product semigroups using our simple algebra:%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -318,11 +309,10 @@
\begin{isamarkuptext}%
\noindent Associativity of product semigroups is established using
the definition of \isa{{\isacharparenleft}{\isasymotimes}{\isacharparenright}} on products and the hypothetical
- associativity of the type components; these hypotheses
- are legitimate due to the \isa{semigroup} constraints imposed
- on the type components by the \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} proposition.
- Indeed, this pattern often occurs with parametric types
- and type classes.%
+ associativity of the type components; these hypotheses are
+ legitimate due to the \isa{semigroup} constraints imposed on the
+ type components by the \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} proposition. Indeed,
+ this pattern often occurs with parametric types and type classes.%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -331,10 +321,9 @@
\isamarkuptrue%
%
\begin{isamarkuptext}%
-We define a subclass \isa{monoidl} (a semigroup with a left-hand neutral)
- by extending \isa{semigroup}
- with one additional parameter \isa{neutral} together
- with its characteristic property:%
+We define a subclass \isa{monoidl} (a semigroup with a left-hand
+ neutral) by extending \isa{semigroup} with one additional
+ parameter \isa{neutral} together with its characteristic property:%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -355,10 +344,10 @@
\endisadelimquote
%
\begin{isamarkuptext}%
-\noindent Again, we prove some instances, by
- providing suitable parameter definitions and proofs for the
- additional specifications. Observe that instantiations
- for types with the same arity may be simultaneous:%
+\noindent Again, we prove some instances, by providing suitable
+ parameter definitions and proofs for the additional specifications.
+ Observe that instantiations for types with the same arity may be
+ simultaneous:%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -505,8 +494,7 @@
\endisadelimquote
%
\begin{isamarkuptext}%
-\noindent To finish our small algebra example, we add a \isa{group} class
- with a corresponding instance:%
+\noindent To finish our small algebra example, we add a \isa{group} class with a corresponding instance:%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -563,9 +551,9 @@
\isamarkuptrue%
%
\begin{isamarkuptext}%
-The example above gives an impression how Isar type classes work
- in practice. As stated in the introduction, classes also provide
- a link to Isar's locale system. Indeed, the logical core of a class
+The example above gives an impression how Isar type classes work in
+ practice. As stated in the introduction, classes also provide a
+ link to Isar's locale system. Indeed, the logical core of a class
is nothing other than a locale:%
\end{isamarkuptext}%
\isamarkuptrue%
@@ -780,10 +768,12 @@
\endisadelimquote
%
\begin{isamarkuptext}%
-\noindent Here the \qt{\hyperlink{keyword.in}{\mbox{\isa{\isakeyword{in}}}} \isa{group}} target specification
- indicates that the result is recorded within that context for later
- use. This local theorem is also lifted to the global one \hyperlink{fact.group.left-cancel:}{\mbox{\isa{group{\isachardot}left{\isacharunderscore}cancel{\isacharcolon}}}} \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ {\isasymalpha}{\isasymColon}group{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}. Since type \isa{int} has been made an instance of
- \isa{group} before, we may refer to that fact as well: \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ int{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}.%
+\noindent Here the \qt{\hyperlink{keyword.in}{\mbox{\isa{\isakeyword{in}}}} \isa{group}} target
+ specification indicates that the result is recorded within that
+ context for later use. This local theorem is also lifted to the
+ global one \hyperlink{fact.group.left-cancel:}{\mbox{\isa{group{\isachardot}left{\isacharunderscore}cancel{\isacharcolon}}}} \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ {\isasymalpha}{\isasymColon}group{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}. Since type \isa{int} has been
+ made an instance of \isa{group} before, we may refer to that
+ fact as well: \isa{{\isachardoublequote}{\isasymAnd}x\ y\ z\ {\isasymColon}\ int{\isachardot}\ x\ {\isasymotimes}\ y\ {\isacharequal}\ x\ {\isasymotimes}\ z\ {\isasymlongleftrightarrow}\ y\ {\isacharequal}\ z{\isachardoublequote}}.%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -814,16 +804,14 @@
%
\begin{isamarkuptext}%
\noindent If the locale \isa{group} is also a class, this local
- definition is propagated onto a global definition of
- \isa{{\isachardoublequote}pow{\isacharunderscore}nat\ {\isasymColon}\ nat\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoid\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoid{\isachardoublequote}}
- with corresponding theorems
+ definition is propagated onto a global definition of \isa{{\isachardoublequote}pow{\isacharunderscore}nat\ {\isasymColon}\ nat\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoid\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}monoid{\isachardoublequote}} with corresponding theorems
\isa{pow{\isacharunderscore}nat\ {\isadigit{0}}\ x\ {\isacharequal}\ {\isasymone}\isasep\isanewline%
pow{\isacharunderscore}nat\ {\isacharparenleft}Suc\ n{\isacharparenright}\ x\ {\isacharequal}\ x\ {\isasymotimes}\ pow{\isacharunderscore}nat\ n\ x}.
- \noindent As you can see from this example, for local
- definitions you may use any specification tool
- which works together with locales, such as Krauss's recursive function package
+ \noindent As you can see from this example, for local definitions
+ you may use any specification tool which works together with
+ locales, such as Krauss's recursive function package
\cite{krauss2006}.%
\end{isamarkuptext}%
\isamarkuptrue%
@@ -833,19 +821,17 @@
\isamarkuptrue%
%
\begin{isamarkuptext}%
-We introduced Isar classes by analogy to type classes in
- functional programming; if we reconsider this in the
- context of what has been said about type classes and locales,
- we can drive this analogy further by stating that type
- classes essentially correspond to functors that have
- a canonical interpretation as type classes.
- There is also the possibility of other interpretations.
- For example, \isa{list}s also form a monoid with
- \isa{append} and \isa{{\isacharbrackleft}{\isacharbrackright}} as operations, but it
- seems inappropriate to apply to lists
- the same operations as for genuinely algebraic types.
- In such a case, we can simply make a particular interpretation
- of monoids for lists:%
+We introduced Isar classes by analogy to type classes in functional
+ programming; if we reconsider this in the context of what has been
+ said about type classes and locales, we can drive this analogy
+ further by stating that type classes essentially correspond to
+ functors that have a canonical interpretation as type classes.
+ There is also the possibility of other interpretations. For
+ example, \isa{list}s also form a monoid with \isa{append} and
+ \isa{{\isacharbrackleft}{\isacharbrackright}} as operations, but it seems inappropriate to apply to
+ lists the same operations as for genuinely algebraic types. In such
+ a case, we can simply make a particular interpretation of monoids
+ for lists:%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -969,12 +955,10 @@
\endisadelimquote
%
\begin{isamarkuptext}%
-The logical proof is carried out on the locale level.
- Afterwards it is propagated
- to the type system, making \isa{group} an instance of
- \isa{monoid} by adding an additional edge
- to the graph of subclass relations
- (\figref{fig:subclass}).
+The logical proof is carried out on the locale level. Afterwards it
+ is propagated to the type system, making \isa{group} an instance
+ of \isa{monoid} by adding an additional edge to the graph of
+ subclass relations (\figref{fig:subclass}).
\begin{figure}[htbp]
\begin{center}
@@ -1006,8 +990,7 @@
\end{center}
\end{figure}
- For illustration, a derived definition
- in \isa{group} using \isa{pow{\isacharunderscore}nat}%
+ For illustration, a derived definition in \isa{group} using \isa{pow{\isacharunderscore}nat}%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -1029,9 +1012,7 @@
\endisadelimquote
%
\begin{isamarkuptext}%
-\noindent yields the global definition of
- \isa{{\isachardoublequote}pow{\isacharunderscore}int\ {\isasymColon}\ int\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group{\isachardoublequote}}
- with the corresponding theorem \isa{pow{\isacharunderscore}int\ k\ x\ {\isacharequal}\ {\isacharparenleft}if\ {\isadigit{0}}\ {\isasymle}\ k\ then\ pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ k{\isacharparenright}\ x\ else\ {\isacharparenleft}pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ {\isacharparenleft}{\isacharminus}\ k{\isacharparenright}{\isacharparenright}\ x{\isacharparenright}{\isasymdiv}{\isacharparenright}}.%
+\noindent yields the global definition of \isa{{\isachardoublequote}pow{\isacharunderscore}int\ {\isasymColon}\ int\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group\ {\isasymRightarrow}\ {\isasymalpha}{\isasymColon}group{\isachardoublequote}} with the corresponding theorem \isa{pow{\isacharunderscore}int\ k\ x\ {\isacharequal}\ {\isacharparenleft}if\ {\isadigit{0}}\ {\isasymle}\ k\ then\ pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ k{\isacharparenright}\ x\ else\ {\isacharparenleft}pow{\isacharunderscore}nat\ {\isacharparenleft}nat\ {\isacharparenleft}{\isacharminus}\ k{\isacharparenright}{\isacharparenright}\ x{\isacharparenright}{\isasymdiv}{\isacharparenright}}.%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -1040,9 +1021,9 @@
\isamarkuptrue%
%
\begin{isamarkuptext}%
-As a convenience, class context syntax allows references
- to local class operations and their global counterparts
- uniformly; type inference resolves ambiguities. For example:%
+As a convenience, class context syntax allows references to local
+ class operations and their global counterparts uniformly; type
+ inference resolves ambiguities. For example:%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -1082,11 +1063,11 @@
\endisadelimquote
%
\begin{isamarkuptext}%
-\noindent Here in example 1, the term refers to the local class operation
- \isa{mult\ {\isacharbrackleft}{\isasymalpha}{\isacharbrackright}}, whereas in example 2 the type constraint
- enforces the global class operation \isa{mult\ {\isacharbrackleft}nat{\isacharbrackright}}.
- In the global context in example 3, the reference is
- to the polymorphic global class operation \isa{mult\ {\isacharbrackleft}{\isacharquery}{\isasymalpha}\ {\isasymColon}\ semigroup{\isacharbrackright}}.%
+\noindent Here in example 1, the term refers to the local class
+ operation \isa{mult\ {\isacharbrackleft}{\isasymalpha}{\isacharbrackright}}, whereas in example 2 the type
+ constraint enforces the global class operation \isa{mult\ {\isacharbrackleft}nat{\isacharbrackright}}.
+ In the global context in example 3, the reference is to the
+ polymorphic global class operation \isa{mult\ {\isacharbrackleft}{\isacharquery}{\isasymalpha}\ {\isasymColon}\ semigroup{\isacharbrackright}}.%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -1099,16 +1080,13 @@
\isamarkuptrue%
%
\begin{isamarkuptext}%
-Turning back to the first motivation for type classes,
- namely overloading, it is obvious that overloading
- stemming from \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} statements and
- \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}}
- targets naturally maps to Haskell type classes.
- The code generator framework \cite{isabelle-codegen}
- takes this into account. If the target language (e.g.~SML)
- lacks type classes, then they
- are implemented by an explicit dictionary construction.
- As example, let's go back to the power function:%
+Turning back to the first motivation for type classes, namely
+ overloading, it is obvious that overloading stemming from \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} statements and \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} targets naturally
+ maps to Haskell type classes. The code generator framework
+ \cite{isabelle-codegen} takes this into account. If the target
+ language (e.g.~SML) lacks type classes, then they are implemented by
+ an explicit dictionary construction. As example, let's go back to
+ the power function:%
\end{isamarkuptext}%
\isamarkuptrue%
%
@@ -1328,13 +1306,121 @@
%
\endisadelimquote
%
+\begin{isamarkuptext}%
+\noindent In Scala, implicts are used as dictionaries:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadeliminvisible
+%
+\endisadeliminvisible
+%
+\isataginvisible
+%
+\endisataginvisible
+{\isafoldinvisible}%
+%
+\isadeliminvisible
+%
+\endisadeliminvisible
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\isatagquote
+%
+\begin{isamarkuptext}%
+\isatypewriter%
+\noindent%
+\hspace*{0pt}object Example {\char123}\\
+\hspace*{0pt}\\
+\hspace*{0pt}import /*implicits*/ Example.semigroup{\char95}int,~Example.monoidl{\char95}int,\\
+\hspace*{0pt} ~Example.monoid{\char95}int,~Example.group{\char95}int\\
+\hspace*{0pt}\\
+\hspace*{0pt}abstract sealed class nat\\
+\hspace*{0pt}final case object Zero{\char95}nat extends nat\\
+\hspace*{0pt}final case class Suc(a:~nat) extends nat\\
+\hspace*{0pt}\\
+\hspace*{0pt}def nat{\char95}aux(i:~BigInt,~n:~nat):~nat =\\
+\hspace*{0pt} ~(if (i <= BigInt(0)) n else nat{\char95}aux(i - BigInt(1),~Suc(n)))\\
+\hspace*{0pt}\\
+\hspace*{0pt}def nat(i:~BigInt):~nat = nat{\char95}aux(i,~Zero{\char95}nat)\\
+\hspace*{0pt}\\
+\hspace*{0pt}trait semigroup[A] {\char123}\\
+\hspace*{0pt} ~val `Example+mult`:~(A,~A) => A\\
+\hspace*{0pt}{\char125}\\
+\hspace*{0pt}def mult[A](a:~A,~b:~A)(implicit A:~semigroup[A]):~A =\\
+\hspace*{0pt} ~A.`Example+mult`(a,~b)\\
+\hspace*{0pt}\\
+\hspace*{0pt}trait monoidl[A] extends semigroup[A] {\char123}\\
+\hspace*{0pt} ~val `Example+neutral`:~A\\
+\hspace*{0pt}{\char125}\\
+\hspace*{0pt}def neutral[A](implicit A:~monoidl[A]):~A = A.`Example+neutral`\\
+\hspace*{0pt}\\
+\hspace*{0pt}trait monoid[A] extends monoidl[A] {\char123}\\
+\hspace*{0pt}{\char125}\\
+\hspace*{0pt}\\
+\hspace*{0pt}trait group[A] extends monoid[A] {\char123}\\
+\hspace*{0pt} ~val `Example+inverse`:~A => A\\
+\hspace*{0pt}{\char125}\\
+\hspace*{0pt}def inverse[A](a:~A)(implicit A:~group[A]):~A = A.`Example+inverse`(a)\\
+\hspace*{0pt}\\
+\hspace*{0pt}def pow{\char95}nat[A:~monoid](xa0:~nat,~x:~A):~A = (xa0,~x) match {\char123}\\
+\hspace*{0pt} ~case (Zero{\char95}nat,~x) => neutral[A]\\
+\hspace*{0pt} ~case (Suc(n),~x) => mult[A](x,~pow{\char95}nat[A](n,~x))\\
+\hspace*{0pt}{\char125}\\
+\hspace*{0pt}\\
+\hspace*{0pt}def pow{\char95}int[A:~group](k:~BigInt,~x:~A):~A =\\
+\hspace*{0pt} ~(if (BigInt(0) <= k) pow{\char95}nat[A](nat(k),~x)\\
+\hspace*{0pt} ~~~else inverse[A](pow{\char95}nat[A](nat((- k)),~x)))\\
+\hspace*{0pt}\\
+\hspace*{0pt}def mult{\char95}int(i:~BigInt,~j:~BigInt):~BigInt = i + j\\
+\hspace*{0pt}\\
+\hspace*{0pt}implicit def semigroup{\char95}int:~semigroup[BigInt] = new semigroup[BigInt] {\char123}\\
+\hspace*{0pt} ~val `Example+mult` = (a:~BigInt,~b:~BigInt) => mult{\char95}int(a,~b)\\
+\hspace*{0pt}{\char125}\\
+\hspace*{0pt}\\
+\hspace*{0pt}def neutral{\char95}int:~BigInt = BigInt(0)\\
+\hspace*{0pt}\\
+\hspace*{0pt}implicit def monoidl{\char95}int:~monoidl[BigInt] = new monoidl[BigInt] {\char123}\\
+\hspace*{0pt} ~val `Example+neutral` = neutral{\char95}int\\
+\hspace*{0pt} ~val `Example+mult` = (a:~BigInt,~b:~BigInt) => mult{\char95}int(a,~b)\\
+\hspace*{0pt}{\char125}\\
+\hspace*{0pt}\\
+\hspace*{0pt}implicit def monoid{\char95}int:~monoid[BigInt] = new monoid[BigInt] {\char123}\\
+\hspace*{0pt} ~val `Example+neutral` = neutral{\char95}int\\
+\hspace*{0pt} ~val `Example+mult` = (a:~BigInt,~b:~BigInt) => mult{\char95}int(a,~b)\\
+\hspace*{0pt}{\char125}\\
+\hspace*{0pt}\\
+\hspace*{0pt}def inverse{\char95}int(i:~BigInt):~BigInt = (- i)\\
+\hspace*{0pt}\\
+\hspace*{0pt}implicit def group{\char95}int:~group[BigInt] = new group[BigInt] {\char123}\\
+\hspace*{0pt} ~val `Example+inverse` = (a:~BigInt) => inverse{\char95}int(a)\\
+\hspace*{0pt} ~val `Example+neutral` = neutral{\char95}int\\
+\hspace*{0pt} ~val `Example+mult` = (a:~BigInt,~b:~BigInt) => mult{\char95}int(a,~b)\\
+\hspace*{0pt}{\char125}\\
+\hspace*{0pt}\\
+\hspace*{0pt}def example:~BigInt = pow{\char95}int[BigInt](BigInt(10),~BigInt(- 2))\\
+\hspace*{0pt}\\
+\hspace*{0pt}{\char125}~/* object Example */%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagquote
+{\isafoldquote}%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
\isamarkupsubsection{Inspecting the type class universe%
}
\isamarkuptrue%
%
\begin{isamarkuptext}%
-To facilitate orientation in complex subclass structures,
- two diagnostics commands are provided:
+To facilitate orientation in complex subclass structures, two
+ diagnostics commands are provided:
\begin{description}
--- a/doc-src/Codegen/Thy/Foundations.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Codegen/Thy/Foundations.thy Wed Sep 01 17:19:47 2010 +0200
@@ -220,12 +220,12 @@
text {*
\noindent Obviously, polymorphic equality is implemented the Haskell
way using a type class. How is this achieved? HOL introduces an
- explicit class @{class eq} with a corresponding operation @{const
- eq_class.eq} such that @{thm eq [no_vars]}. The preprocessing
- framework does the rest by propagating the @{class eq} constraints
+ explicit class @{class equal} with a corresponding operation @{const
+ HOL.equal} such that @{thm equal [no_vars]}. The preprocessing
+ framework does the rest by propagating the @{class equal} constraints
through all dependent code equations. For datatypes, instances of
- @{class eq} are implicitly derived when possible. For other types,
- you may instantiate @{text eq} manually like any other type class.
+ @{class equal} are implicitly derived when possible. For other types,
+ you may instantiate @{text equal} manually like any other type class.
*}
--- a/doc-src/Codegen/Thy/Inductive_Predicate.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Codegen/Thy/Inductive_Predicate.thy Wed Sep 01 17:19:47 2010 +0200
@@ -7,7 +7,7 @@
inductive %invisible append where
"append [] ys ys"
-| "append xs ys zs ==> append (x # xs) ys (x # zs)"
+| "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)"
lemma %invisible append: "append xs ys zs = (xs @ ys = zs)"
by (induct xs arbitrary: ys zs) (auto elim: append.cases intro: append.intros)
@@ -95,9 +95,9 @@
"append_i_i_o"}). You can specify your own names as follows:
*}
-code_pred %quote (modes: i => i => o => bool as concat,
- o => o => i => bool as split,
- i => o => i => bool as suffix) append .
+code_pred %quote (modes: i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool as concat,
+ o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool as split,
+ i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool as suffix) append .
subsection {* Alternative introduction rules *}
@@ -220,8 +220,8 @@
"values"} and the number of elements.
*}
-values %quote [mode: i => o => bool] 20 "{n. tranclp succ 10 n}"
-values %quote [mode: o => i => bool] 10 "{n. tranclp succ n 10}"
+values %quote [mode: i \<Rightarrow> o \<Rightarrow> bool] 20 "{n. tranclp succ 10 n}"
+values %quote [mode: o \<Rightarrow> i \<Rightarrow> bool] 10 "{n. tranclp succ n 10}"
subsection {* Embedding into functional code within Isabelle/HOL *}
--- a/doc-src/Codegen/Thy/Introduction.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Codegen/Thy/Introduction.thy Wed Sep 01 17:19:47 2010 +0200
@@ -8,8 +8,9 @@
This tutorial introduces the code generator facilities of @{text
"Isabelle/HOL"}. It allows to turn (a certain class of) HOL
specifications into corresponding executable code in the programming
- languages @{text SML} \cite{SML}, @{text OCaml} \cite{OCaml} and
- @{text Haskell} \cite{haskell-revised-report}.
+ languages @{text SML} \cite{SML}, @{text OCaml} \cite{OCaml},
+ @{text Haskell} \cite{haskell-revised-report} and @{text Scala}
+ \cite{scala-overview-tech-report}.
To profit from this tutorial, some familiarity and experience with
@{theory HOL} \cite{isa-tutorial} and its basic theories is assumed.
@@ -78,8 +79,8 @@
target language identifier and a freely chosen module name. A file
name denotes the destination to store the generated code. Note that
the semantics of the destination depends on the target language: for
- @{text SML} and @{text OCaml} it denotes a \emph{file}, for @{text
- Haskell} it denotes a \emph{directory} where a file named as the
+ @{text SML}, @{text OCaml} and @{text Scala} it denotes a \emph{file},
+ for @{text Haskell} it denotes a \emph{directory} where a file named as the
module name (with extension @{text ".hs"}) is written:
*}
--- a/doc-src/Codegen/Thy/Setup.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Codegen/Thy/Setup.thy Wed Sep 01 17:19:47 2010 +0200
@@ -27,6 +27,6 @@
setup {* Code_Target.set_default_code_width 74 *}
-ML_command {* Unsynchronized.reset unique_names *}
+ML_command {* unique_names := false *}
end
--- a/doc-src/Codegen/Thy/document/Adaptation.tex Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Codegen/Thy/document/Adaptation.tex Wed Sep 01 17:19:47 2010 +0200
@@ -240,7 +240,7 @@
\hspace*{0pt}structure Example :~sig\\
\hspace*{0pt} ~datatype nat = Zero{\char95}nat | Suc of nat\\
\hspace*{0pt} ~datatype boola = True | False\\
-\hspace*{0pt} ~val anda :~boola -> boola -> boola\\
+\hspace*{0pt} ~val conj :~boola -> boola -> boola\\
\hspace*{0pt} ~val less{\char95}nat :~nat -> nat -> boola\\
\hspace*{0pt} ~val less{\char95}eq{\char95}nat :~nat -> nat -> boola\\
\hspace*{0pt} ~val in{\char95}interval :~nat * nat -> nat -> boola\\
@@ -250,17 +250,17 @@
\hspace*{0pt}\\
\hspace*{0pt}datatype boola = True | False;\\
\hspace*{0pt}\\
-\hspace*{0pt}fun anda p True = p\\
-\hspace*{0pt} ~| anda p False = False\\
-\hspace*{0pt} ~| anda True p = p\\
-\hspace*{0pt} ~| anda False p = False;\\
+\hspace*{0pt}fun conj p True = p\\
+\hspace*{0pt} ~| conj p False = False\\
+\hspace*{0pt} ~| conj True p = p\\
+\hspace*{0pt} ~| conj False p = False;\\
\hspace*{0pt}\\
\hspace*{0pt}fun less{\char95}nat m (Suc n) = less{\char95}eq{\char95}nat m n\\
\hspace*{0pt} ~| less{\char95}nat n Zero{\char95}nat = False\\
\hspace*{0pt}and less{\char95}eq{\char95}nat (Suc m) n = less{\char95}nat m n\\
\hspace*{0pt} ~| less{\char95}eq{\char95}nat Zero{\char95}nat n = True;\\
\hspace*{0pt}\\
-\hspace*{0pt}fun in{\char95}interval (k,~l) n = anda (less{\char95}eq{\char95}nat k n) (less{\char95}eq{\char95}nat n l);\\
+\hspace*{0pt}fun in{\char95}interval (k,~l) n = conj (less{\char95}eq{\char95}nat k n) (less{\char95}eq{\char95}nat n l);\\
\hspace*{0pt}\\
\hspace*{0pt}end;~(*struct Example*)%
\end{isamarkuptext}%
--- a/doc-src/Codegen/Thy/document/Inductive_Predicate.tex Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Codegen/Thy/document/Inductive_Predicate.tex Wed Sep 01 17:19:47 2010 +0200
@@ -212,9 +212,9 @@
%
\isatagquote
\isacommand{code{\isacharunderscore}pred}\isamarkupfalse%
-\ {\isacharparenleft}modes{\isacharcolon}\ i\ {\isacharequal}{\isachargreater}\ i\ {\isacharequal}{\isachargreater}\ o\ {\isacharequal}{\isachargreater}\ bool\ as\ concat{\isacharcomma}\isanewline
-\ \ o\ {\isacharequal}{\isachargreater}\ o\ {\isacharequal}{\isachargreater}\ i\ {\isacharequal}{\isachargreater}\ bool\ as\ split{\isacharcomma}\isanewline
-\ \ i\ {\isacharequal}{\isachargreater}\ o\ {\isacharequal}{\isachargreater}\ i\ {\isacharequal}{\isachargreater}\ bool\ as\ suffix{\isacharparenright}\ append\ \isacommand{{\isachardot}}\isamarkupfalse%
+\ {\isacharparenleft}modes{\isacharcolon}\ i\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ bool\ as\ concat{\isacharcomma}\isanewline
+\ \ o\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ bool\ as\ split{\isacharcomma}\isanewline
+\ \ i\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ bool\ as\ suffix{\isacharparenright}\ append\ \isacommand{{\isachardot}}\isamarkupfalse%
%
\endisatagquote
{\isafoldquote}%
@@ -422,9 +422,9 @@
%
\isatagquote
\isacommand{values}\isamarkupfalse%
-\ {\isacharbrackleft}mode{\isacharcolon}\ i\ {\isacharequal}{\isachargreater}\ o\ {\isacharequal}{\isachargreater}\ bool{\isacharbrackright}\ {\isadigit{2}}{\isadigit{0}}\ {\isachardoublequoteopen}{\isacharbraceleft}n{\isachardot}\ tranclp\ succ\ {\isadigit{1}}{\isadigit{0}}\ n{\isacharbraceright}{\isachardoublequoteclose}\isanewline
+\ {\isacharbrackleft}mode{\isacharcolon}\ i\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ bool{\isacharbrackright}\ {\isadigit{2}}{\isadigit{0}}\ {\isachardoublequoteopen}{\isacharbraceleft}n{\isachardot}\ tranclp\ succ\ {\isadigit{1}}{\isadigit{0}}\ n{\isacharbraceright}{\isachardoublequoteclose}\isanewline
\isacommand{values}\isamarkupfalse%
-\ {\isacharbrackleft}mode{\isacharcolon}\ o\ {\isacharequal}{\isachargreater}\ i\ {\isacharequal}{\isachargreater}\ bool{\isacharbrackright}\ {\isadigit{1}}{\isadigit{0}}\ {\isachardoublequoteopen}{\isacharbraceleft}n{\isachardot}\ tranclp\ succ\ n\ {\isadigit{1}}{\isadigit{0}}{\isacharbraceright}{\isachardoublequoteclose}%
+\ {\isacharbrackleft}mode{\isacharcolon}\ o\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ bool{\isacharbrackright}\ {\isadigit{1}}{\isadigit{0}}\ {\isachardoublequoteopen}{\isacharbraceleft}n{\isachardot}\ tranclp\ succ\ n\ {\isadigit{1}}{\isadigit{0}}{\isacharbraceright}{\isachardoublequoteclose}%
\endisatagquote
{\isafoldquote}%
%
--- a/doc-src/Codegen/Thy/document/Introduction.tex Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Codegen/Thy/document/Introduction.tex Wed Sep 01 17:19:47 2010 +0200
@@ -25,8 +25,9 @@
\begin{isamarkuptext}%
This tutorial introduces the code generator facilities of \isa{Isabelle{\isacharslash}HOL}. It allows to turn (a certain class of) HOL
specifications into corresponding executable code in the programming
- languages \isa{SML} \cite{SML}, \isa{OCaml} \cite{OCaml} and
- \isa{Haskell} \cite{haskell-revised-report}.
+ languages \isa{SML} \cite{SML}, \isa{OCaml} \cite{OCaml},
+ \isa{Haskell} \cite{haskell-revised-report} and \isa{Scala}
+ \cite{scala-overview-tech-report}.
To profit from this tutorial, some familiarity and experience with
\hyperlink{theory.HOL}{\mbox{\isa{HOL}}} \cite{isa-tutorial} and its basic theories is assumed.%
@@ -191,7 +192,8 @@
target language identifier and a freely chosen module name. A file
name denotes the destination to store the generated code. Note that
the semantics of the destination depends on the target language: for
- \isa{SML} and \isa{OCaml} it denotes a \emph{file}, for \isa{Haskell} it denotes a \emph{directory} where a file named as the
+ \isa{SML}, \isa{OCaml} and \isa{Scala} it denotes a \emph{file},
+ for \isa{Haskell} it denotes a \emph{directory} where a file named as the
module name (with extension \isa{{\isachardot}hs}) is written:%
\end{isamarkuptext}%
\isamarkuptrue%
--- a/doc-src/Codegen/Thy/pictures/architecture.tex Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Codegen/Thy/pictures/architecture.tex Wed Sep 01 17:19:47 2010 +0200
@@ -30,8 +30,8 @@
\node (seri) at (1.5, 0) [process, positive] {serialisation};
\node (SML) at (2.5, 3) [entity, positive] {\sys{SML}};
\node (OCaml) at (2.5, 2) [entity, positive] {\sys{OCaml}};
- \node (further) at (2.5, 1) [entity, positive] {(\ldots)};
- \node (Haskell) at (2.5, 0) [entity, positive] {\sys{Haskell}};
+ \node (Haskell) at (2.5, 1) [entity, positive] {\sys{Haskell}};
+ \node (Scala) at (2.5, 0) [entity, positive] {\sys{Scala}};
\draw [semithick] (spec) -- (spec_user_join);
\draw [semithick] (user) -- (spec_user_join);
\draw [-diamond, semithick] (spec_user_join) -- (raw);
@@ -41,8 +41,8 @@
\draw [arrow] (iml) -- (seri);
\draw [arrow] (seri) -- (SML);
\draw [arrow] (seri) -- (OCaml);
- \draw [arrow, dashed] (seri) -- (further);
\draw [arrow] (seri) -- (Haskell);
+ \draw [arrow] (seri) -- (Scala);
\end{tikzpicture}
}
--- a/doc-src/Codegen/codegen.tex Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Codegen/codegen.tex Wed Sep 01 17:19:47 2010 +0200
@@ -22,7 +22,7 @@
\begin{abstract}
\noindent This tutorial introduces the code generator facilities of Isabelle/HOL.
They empower the user to turn HOL specifications into corresponding executable
- programs in the languages SML, OCaml and Haskell.
+ programs in the languages SML, OCaml, Haskell and Scala.
\end{abstract}
\thispagestyle{empty}\clearpage
--- a/doc-src/IsarOverview/Isar/ROOT.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/IsarOverview/Isar/ROOT.ML Wed Sep 01 17:19:47 2010 +0200
@@ -1,3 +1,3 @@
-Unsynchronized.set quick_and_dirty;
+quick_and_dirty := true;
use_thys ["Logic", "Induction"];
--- a/doc-src/IsarRef/Thy/HOL_Specific.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/IsarRef/Thy/HOL_Specific.thy Wed Sep 01 17:19:47 2010 +0200
@@ -968,7 +968,8 @@
\medskip One framework generates code from functional programs
(including overloading using type classes) to SML \cite{SML}, OCaml
- \cite{OCaml} and Haskell \cite{haskell-revised-report}.
+ \cite{OCaml}, Haskell \cite{haskell-revised-report} and Scala
+ \cite{scala-overview-tech-report}.
Conceptually, code generation is split up in three steps:
\emph{selection} of code theorems, \emph{translation} into an
abstract executable view and \emph{serialization} to a specific
@@ -1015,7 +1016,7 @@
class: nameref
;
- target: 'OCaml' | 'SML' | 'Haskell'
+ target: 'SML' | 'OCaml' | 'Haskell' | 'Scala'
;
'code' ( 'del' | 'abstype' | 'abstract' ) ?
@@ -1088,7 +1089,7 @@
after the @{keyword "module_name"} keyword; then \emph{all} code is
placed in this module.
- For \emph{SML} and \emph{OCaml}, the file specification refers to a
+ For \emph{SML}, \emph{OCaml} and \emph{Scala} the file specification refers to a
single file; for \emph{Haskell}, it refers to a whole directory,
where code is generated in multiple files reflecting the module
hierarchy. Omitting the file specification denotes standard
--- a/doc-src/IsarRef/Thy/ROOT-HOLCF.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/IsarRef/Thy/ROOT-HOLCF.ML Wed Sep 01 17:19:47 2010 +0200
@@ -1,4 +1,4 @@
-Unsynchronized.set Thy_Output.source;
+Thy_Output.source_default := true;
use "../../antiquote_setup.ML";
use_thy "HOLCF_Specific";
--- a/doc-src/IsarRef/Thy/ROOT-ZF.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/IsarRef/Thy/ROOT-ZF.ML Wed Sep 01 17:19:47 2010 +0200
@@ -1,4 +1,4 @@
-Unsynchronized.set Thy_Output.source;
+Thy_Output.source_default := true;
use "../../antiquote_setup.ML";
use_thy "ZF_Specific";
--- a/doc-src/IsarRef/Thy/ROOT.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/IsarRef/Thy/ROOT.ML Wed Sep 01 17:19:47 2010 +0200
@@ -1,5 +1,5 @@
-Unsynchronized.set quick_and_dirty;
-Unsynchronized.set Thy_Output.source;
+quick_and_dirty := true;
+Thy_Output.source_default := true;
use "../../antiquote_setup.ML";
use_thys [
--- a/doc-src/IsarRef/Thy/document/HOL_Specific.tex Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/IsarRef/Thy/document/HOL_Specific.tex Wed Sep 01 17:19:47 2010 +0200
@@ -984,7 +984,8 @@
\medskip One framework generates code from functional programs
(including overloading using type classes) to SML \cite{SML}, OCaml
- \cite{OCaml} and Haskell \cite{haskell-revised-report}.
+ \cite{OCaml}, Haskell \cite{haskell-revised-report} and Scala
+ \cite{scala-overview-tech-report}.
Conceptually, code generation is split up in three steps:
\emph{selection} of code theorems, \emph{translation} into an
abstract executable view and \emph{serialization} to a specific
@@ -1031,7 +1032,7 @@
class: nameref
;
- target: 'OCaml' | 'SML' | 'Haskell'
+ target: 'SML' | 'OCaml' | 'Haskell' | 'Scala'
;
'code' ( 'del' | 'abstype' | 'abstract' ) ?
@@ -1103,7 +1104,7 @@
after the \hyperlink{keyword.module-name}{\mbox{\isa{\isakeyword{module{\isacharunderscore}name}}}} keyword; then \emph{all} code is
placed in this module.
- For \emph{SML} and \emph{OCaml}, the file specification refers to a
+ For \emph{SML}, \emph{OCaml} and \emph{Scala} the file specification refers to a
single file; for \emph{Haskell}, it refers to a whole directory,
where code is generated in multiple files reflecting the module
hierarchy. Omitting the file specification denotes standard
--- a/doc-src/LaTeXsugar/Sugar/Sugar.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/LaTeXsugar/Sugar/Sugar.thy Wed Sep 01 17:19:47 2010 +0200
@@ -132,7 +132,7 @@
This \verb!no_vars! business can become a bit tedious.
If you would rather never see question marks, simply put
\begin{quote}
-@{ML "Unsynchronized.reset show_question_marks"}\verb!;!
+@{ML "show_question_marks := false"}\verb!;!
\end{quote}
at the beginning of your file \texttt{ROOT.ML}.
The rest of this document is produced with this flag set to \texttt{false}.
@@ -144,7 +144,7 @@
turning the last digit into a subscript: write \verb!x\<^isub>1! and
obtain the much nicer @{text"x\<^isub>1"}. *}
-(*<*)ML "Unsynchronized.reset show_question_marks"(*>*)
+(*<*)ML "show_question_marks := false"(*>*)
subsection {*Qualified names*}
@@ -155,7 +155,7 @@
short names (no qualifiers) by setting \verb!short_names!, typically
in \texttt{ROOT.ML}:
\begin{quote}
-@{ML "Unsynchronized.set short_names"}\verb!;!
+@{ML "short_names := true"}\verb!;!
\end{quote}
*}
--- a/doc-src/LaTeXsugar/Sugar/document/Sugar.tex Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/LaTeXsugar/Sugar/document/Sugar.tex Wed Sep 01 17:19:47 2010 +0200
@@ -173,7 +173,7 @@
This \verb!no_vars! business can become a bit tedious.
If you would rather never see question marks, simply put
\begin{quote}
-\verb|Unsynchronized.reset show_question_marks|\verb!;!
+\verb|show_question_marks := false|\verb!;!
\end{quote}
at the beginning of your file \texttt{ROOT.ML}.
The rest of this document is produced with this flag set to \texttt{false}.
@@ -211,7 +211,7 @@
short names (no qualifiers) by setting \verb!short_names!, typically
in \texttt{ROOT.ML}:
\begin{quote}
-\verb|Unsynchronized.set short_names|\verb!;!
+\verb|short_names := true|\verb!;!
\end{quote}%
\end{isamarkuptext}%
\isamarkuptrue%
--- a/doc-src/Main/Docs/Main_Doc.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Main/Docs/Main_Doc.thy Wed Sep 01 17:19:47 2010 +0200
@@ -10,9 +10,9 @@
Syntax.pretty_typ ctxt T)
val _ = Thy_Output.antiquotation "term_type_only" (Args.term -- Args.typ_abbrev)
- (fn {source, context, ...} => fn arg =>
- Thy_Output.output
- (Thy_Output.maybe_pretty_source (pretty_term_type_only context) source [arg]));
+ (fn {source, context = ctxt, ...} => fn arg =>
+ Thy_Output.output ctxt
+ (Thy_Output.maybe_pretty_source pretty_term_type_only ctxt source [arg]));
*}
(*>*)
text{*
--- a/doc-src/Sledgehammer/sledgehammer.tex Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/Sledgehammer/sledgehammer.tex Wed Sep 01 17:19:47 2010 +0200
@@ -447,33 +447,29 @@
\label{relevance-filter}
\begin{enum}
-\opdefault{relevance\_threshold}{int}{40}
-Specifies the threshold above which facts are considered relevant by the
-relevance filter. The option ranges from 0 to 100, where 0 means that all
-theorems are relevant.
+\opdefault{relevance\_thresholds}{int\_pair}{45~85}
+Specifies the thresholds above which facts are considered relevant by the
+relevance filter. The first threshold is used for the first iteration of the
+relevance filter and the second threshold is used for the last iteration (if it
+is reached). The effective threshold is quadratically interpolated for the other
+iterations. Each threshold ranges from 0 to 100, where 0 means that all theorems
+are relevant and 100 only theorems that refer to previously seen constants.
-\opdefault{relevance\_convergence}{int}{31}
-Specifies the convergence factor, expressed as a percentage, used by the
-relevance filter. This factor is used by the relevance filter to scale down the
-relevance of facts at each iteration of the filter.
-
-\opdefault{max\_relevant\_per\_iter}{int\_or\_smart}{\textit{smart}}
-Specifies the maximum number of facts that may be added during one iteration of
-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 50.
+\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 E and Vampire,
-because empirical results suggest that these are the best settings.
+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.
-
+%\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.
\end{enum}
\subsection{Output Format}
--- a/doc-src/System/Thy/ROOT.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/System/Thy/ROOT.ML Wed Sep 01 17:19:47 2010 +0200
@@ -1,4 +1,4 @@
-Unsynchronized.set Thy_Output.source;
+Thy_Output.source_default := true;
use "../../antiquote_setup.ML";
use_thys ["Basics", "Interfaces", "Presentation", "Misc"];
--- a/doc-src/TutorialI/Documents/Documents.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/TutorialI/Documents/Documents.thy Wed Sep 01 17:19:47 2010 +0200
@@ -144,7 +144,7 @@
definition xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "\<oplus>" 60)
where "A \<oplus> B \<equiv> (A \<and> \<not> B) \<or> (\<not> A \<and> B)"
(*<*)
-local
+setup {* Sign.local_path *}
(*>*)
text {*
@@ -169,7 +169,7 @@
notation (xsymbols) xor (infixl "\<oplus>\<ignore>" 60)
(*<*)
-local
+setup {* Sign.local_path *}
(*>*)
text {*\noindent
--- a/doc-src/TutorialI/Documents/document/Documents.tex Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/TutorialI/Documents/document/Documents.tex Wed Sep 01 17:19:47 2010 +0200
@@ -168,6 +168,19 @@
\isacommand{definition}\isamarkupfalse%
\ xor\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}bool\ {\isasymRightarrow}\ bool\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \ \ \ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymoplus}{\isachardoublequoteclose}\ {\isadigit{6}}{\isadigit{0}}{\isacharparenright}\isanewline
\isakeyword{where}\ {\isachardoublequoteopen}A\ {\isasymoplus}\ B\ {\isasymequiv}\ {\isacharparenleft}A\ {\isasymand}\ {\isasymnot}\ B{\isacharparenright}\ {\isasymor}\ {\isacharparenleft}{\isasymnot}\ A\ {\isasymand}\ B{\isacharparenright}{\isachardoublequoteclose}%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
\begin{isamarkuptext}%
\noindent The X-Symbol package within Proof~General provides several
input methods to enter \isa{{\isasymoplus}} in the text. If all fails one may
@@ -200,6 +213,19 @@
\isanewline
\isacommand{notation}\isamarkupfalse%
\ {\isacharparenleft}xsymbols{\isacharparenright}\ xor\ {\isacharparenleft}\isakeyword{infixl}\ {\isachardoublequoteopen}{\isasymoplus}{\isasymignore}{\isachardoublequoteclose}\ {\isadigit{6}}{\isadigit{0}}{\isacharparenright}%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+%
\begin{isamarkuptext}%
\noindent
The \commdx{notation} command associates a mixfix
--- a/doc-src/TutorialI/Misc/Itrev.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/TutorialI/Misc/Itrev.thy Wed Sep 01 17:19:47 2010 +0200
@@ -2,7 +2,7 @@
theory Itrev
imports Main
begin
-ML"Unsynchronized.reset unique_names"
+ML"unique_names := false"
(*>*)
section{*Induction Heuristics*}
@@ -141,6 +141,6 @@
\index{induction heuristics|)}
*}
(*<*)
-ML"Unsynchronized.set unique_names"
+ML"unique_names := true"
end
(*>*)
--- a/doc-src/TutorialI/Rules/Basic.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/TutorialI/Rules/Basic.thy Wed Sep 01 17:19:47 2010 +0200
@@ -187,7 +187,7 @@
text{*unification failure trace *}
-ML "Unsynchronized.set trace_unify_fail"
+ML "trace_unify_fail := true"
lemma "P(a, f(b, g(e,a), b), a) \<Longrightarrow> P(a, f(b, g(c,a), b), a)"
txt{*
@@ -212,7 +212,7 @@
*}
oops
-ML "Unsynchronized.reset trace_unify_fail"
+ML "trace_unify_fail := false"
text{*Quantifiers*}
--- a/doc-src/TutorialI/Rules/Primes.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/TutorialI/Rules/Primes.thy Wed Sep 01 17:19:47 2010 +0200
@@ -1,4 +1,3 @@
-(* ID: $Id$ *)
(* EXTRACT from HOL/ex/Primes.thy*)
(*Euclid's algorithm
@@ -10,7 +9,7 @@
ML "Pretty.margin_default := 64"
-ML "Thy_Output.indent := 5" (*that is, Doc/TutorialI/settings.ML*)
+declare [[thy_output_indent = 5]] (*that is, Doc/TutorialI/settings.ML*)
text {*Now in Basic.thy!
--- a/doc-src/TutorialI/Sets/Examples.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/TutorialI/Sets/Examples.thy Wed Sep 01 17:19:47 2010 +0200
@@ -1,7 +1,6 @@
-(* ID: $Id$ *)
theory Examples imports Main Binomial begin
-ML "Unsynchronized.reset eta_contract"
+ML "eta_contract := false"
ML "Pretty.margin_default := 64"
text{*membership, intersection *}
--- a/doc-src/TutorialI/Types/Numbers.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/TutorialI/Types/Numbers.thy Wed Sep 01 17:19:47 2010 +0200
@@ -3,7 +3,7 @@
begin
ML "Pretty.margin_default := 64"
-ML "Thy_Output.indent := 0" (*we don't want 5 for listing theorems*)
+declare [[thy_output_indent = 0]] (*we don't want 5 for listing theorems*)
text{*
--- a/doc-src/TutorialI/Types/document/Numbers.tex Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/TutorialI/Types/document/Numbers.tex Wed Sep 01 17:19:47 2010 +0200
@@ -26,16 +26,16 @@
%
\isatagML
\isacommand{ML}\isamarkupfalse%
-\ {\isachardoublequoteopen}Pretty{\isachardot}margin{\isacharunderscore}default\ {\isacharcolon}{\isacharequal}\ {\isadigit{6}}{\isadigit{4}}{\isachardoublequoteclose}\isanewline
-\isacommand{ML}\isamarkupfalse%
-\ {\isachardoublequoteopen}Thy{\isacharunderscore}Output{\isachardot}indent\ {\isacharcolon}{\isacharequal}\ {\isadigit{0}}{\isachardoublequoteclose}%
+\ {\isachardoublequoteopen}Pretty{\isachardot}margin{\isacharunderscore}default\ {\isacharcolon}{\isacharequal}\ {\isadigit{6}}{\isadigit{4}}{\isachardoublequoteclose}%
\endisatagML
{\isafoldML}%
%
\isadelimML
+\isanewline
%
\endisadelimML
-%
+\isacommand{declare}\isamarkupfalse%
+\ {\isacharbrackleft}{\isacharbrackleft}thy{\isacharunderscore}output{\isacharunderscore}indent\ {\isacharequal}\ {\isadigit{0}}{\isacharbrackright}{\isacharbrackright}%
\begin{isamarkuptext}%
numeric literals; default simprules; can re-orient%
\end{isamarkuptext}%
--- a/doc-src/TutorialI/settings.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/TutorialI/settings.ML Wed Sep 01 17:19:47 2010 +0200
@@ -1,3 +1,1 @@
-(* $Id$ *)
-
-Thy_Output.indent := 5;
+Thy_Output.indent_default := 5;
--- a/doc-src/antiquote_setup.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/antiquote_setup.ML Wed Sep 01 17:19:47 2010 +0200
@@ -71,8 +71,8 @@
in
"\\indexdef{}{" ^ kind' ^ "}{" ^ clean_string txt1 ^ "}" ^
(txt'
- |> (if ! Thy_Output.quotes then quote else I)
- |> (if ! Thy_Output.display then enclose "\\begin{verbatim}\n" "\n\\end{verbatim}"
+ |> (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}"
else split_lines #> map verbatim #> space_implode "\\isasep\\isanewline%\n"))
end);
@@ -93,18 +93,18 @@
(Scan.repeat (Attrib.thm -- Scan.lift (Args.parens Args.name)))
(fn {context = ctxt, ...} =>
map (apfst (Thy_Output.pretty_thm ctxt))
- #> (if ! Thy_Output.quotes then map (apfst Pretty.quote) else I)
- #> (if ! Thy_Output.display
+ #> (if Config.get ctxt Thy_Output.quotes then map (apfst Pretty.quote) else I)
+ #> (if Config.get ctxt Thy_Output.display
then
map (fn (p, name) =>
- Output.output (Pretty.string_of (Pretty.indent (! Thy_Output.indent) p)) ^
- "\\rulename{" ^ Output.output (Pretty.str_of (Thy_Output.pretty_text name)) ^ "}")
+ Output.output (Pretty.string_of (Pretty.indent (Config.get ctxt Thy_Output.indent) p)) ^
+ "\\rulename{" ^ Output.output (Pretty.str_of (Thy_Output.pretty_text ctxt name)) ^ "}")
#> space_implode "\\par\\smallskip%\n"
#> enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
else
map (fn (p, name) =>
Output.output (Pretty.str_of p) ^
- "\\rulename{" ^ Output.output (Pretty.str_of (Thy_Output.pretty_text name)) ^ "}")
+ "\\rulename{" ^ Output.output (Pretty.str_of (Thy_Output.pretty_text ctxt name)) ^ "}")
#> space_implode "\\par\\smallskip%\n"
#> enclose "\\isa{" "}"));
@@ -112,7 +112,8 @@
(* theory file *)
val _ = Thy_Output.antiquotation "thy_file" (Scan.lift Args.name)
- (fn _ => fn name => (Thy_Load.check_thy Path.current name; Thy_Output.output [Pretty.str name]));
+ (fn {context = ctxt, ...} =>
+ fn name => (Thy_Load.check_thy Path.current name; Thy_Output.output ctxt [Pretty.str name]));
(* Isabelle/Isar entities (with index) *)
@@ -152,8 +153,9 @@
idx ^
(Output.output name
|> (if markup = "" then I else enclose ("\\" ^ markup ^ "{") "}")
- |> (if ! Thy_Output.quotes then quote else I)
- |> (if ! Thy_Output.display then enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
+ |> (if Config.get ctxt Thy_Output.quotes then quote else I)
+ |> (if Config.get ctxt Thy_Output.display
+ then enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
else hyper o enclose "\\mbox{\\isa{" "}}"))
else error ("Bad " ^ kind ^ " " ^ quote name)
end);
--- a/doc-src/manual.bib Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/manual.bib Wed Sep 01 17:19:47 2010 +0200
@@ -984,6 +984,14 @@
%O
+@TechReport{scala-overview-tech-report,
+ author = {Martin Odersky and al.},
+ title = {An Overview of the Scala Programming Language},
+ institution = {EPFL Lausanne, Switzerland},
+ year = 2004,
+ number = {IC/2004/64}
+}
+
@Manual{pvs-language,
title = {The {PVS} specification language},
author = {S. Owre and N. Shankar and J. M. Rushby},
--- a/doc-src/more_antiquote.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/more_antiquote.ML Wed Sep 01 17:19:47 2010 +0200
@@ -95,7 +95,7 @@
|> snd
|> map_filter (fn (_, (some_thm, proper)) => if proper then some_thm else NONE)
|> map (holize o no_vars ctxt o AxClass.overload thy);
- in Thy_Output.output (Thy_Output.maybe_pretty_source (pretty_thm ctxt) src thms) end;
+ in Thy_Output.output ctxt (Thy_Output.maybe_pretty_source pretty_thm ctxt src thms) end;
in
@@ -124,12 +124,13 @@
in
val _ = Thy_Output.antiquotation "code_stmts"
- (parse_const_terms -- Scan.repeat parse_names -- Scan.lift (Args.parens Args.name))
- (fn {context = ctxt, ...} => fn ((mk_cs, mk_stmtss), target) =>
+ (parse_const_terms -- Scan.repeat parse_names
+ -- Scan.lift (Args.parens (Args.name -- Scan.option Parse.int)))
+ (fn {context = ctxt, ...} => fn ((mk_cs, mk_stmtss), (target, some_width)) =>
let val thy = ProofContext.theory_of ctxt in
- Code_Target.code_of thy
- target NONE "Example" (mk_cs thy)
+ Code_Target.present_code thy (mk_cs thy)
(fn naming => maps (fn f => f thy naming) mk_stmtss)
+ target some_width "Example" []
|> typewriter
end);
--- a/doc-src/rail.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/doc-src/rail.ML Wed Sep 01 17:19:47 2010 +0200
@@ -97,8 +97,9 @@
(idx ^
(Output.output name
|> (if markup = "" then I else enclose ("\\" ^ markup ^ "{") "}")
- |> (if ! Thy_Output.quotes then quote else I)
- |> (if ! Thy_Output.display then enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
+ |> (if Config.get ctxt Thy_Output.quotes then quote else I)
+ |> (if Config.get ctxt Thy_Output.display
+ then enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
else hyper o enclose "\\mbox{\\isa{" "}}")), style)
else ("Bad " ^ kind ^ " " ^ name, false)
end;
--- a/etc/isar-keywords-ZF.el Thu Aug 26 13:15:37 2010 +0200
+++ b/etc/isar-keywords-ZF.el Wed Sep 01 17:19:47 2010 +0200
@@ -1,6 +1,6 @@
;;
;; Keyword classification tables for Isabelle/Isar.
-;; Generated from Pure + Pure-ProofGeneral + FOL + ZF.
+;; Generated from Pure + FOL + ZF.
;; *** DO NOT EDIT *** DO NOT EDIT *** DO NOT EDIT ***
;;
--- a/etc/isar-keywords.el Thu Aug 26 13:15:37 2010 +0200
+++ b/etc/isar-keywords.el Wed Sep 01 17:19:47 2010 +0200
@@ -1,6 +1,6 @@
;;
;; Keyword classification tables for Isabelle/Isar.
-;; Generated from Pure + Pure-ProofGeneral + HOL + HOLCF + HOL-Boogie + HOL-Nominal + HOL-Statespace.
+;; Generated from Pure + HOL + HOLCF + HOL-Boogie + HOL-Nominal + HOL-Statespace.
;; *** DO NOT EDIT *** DO NOT EDIT *** DO NOT EDIT ***
;;
@@ -245,6 +245,7 @@
"thus"
"thy_deps"
"translations"
+ "try"
"txt"
"txt_raw"
"typ"
@@ -398,6 +399,7 @@
"thm"
"thm_deps"
"thy_deps"
+ "try"
"typ"
"unused_thms"
"value"
--- a/src/CCL/ROOT.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/CCL/ROOT.ML Wed Sep 01 17:19:47 2010 +0200
@@ -8,6 +8,6 @@
evaluation to weak head-normal form.
*)
-Unsynchronized.set eta_contract;
+eta_contract := true;
use_thys ["Wfd", "Fix"];
--- a/src/FOLP/IFOLP.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/FOLP/IFOLP.thy Wed Sep 01 17:19:47 2010 +0200
@@ -63,20 +63,22 @@
syntax "_Proof" :: "[p,o]=>prop" ("(_ /: _)" [51, 10] 5)
-ML {*
-
-(*show_proofs:=true displays the proof terms -- they are ENORMOUS*)
-val show_proofs = Unsynchronized.ref false;
-
-fun proof_tr [p,P] = Const (@{const_name Proof}, dummyT) $ P $ p;
-
-fun proof_tr' [P,p] =
- if ! show_proofs then Const (@{syntax_const "_Proof"}, dummyT) $ p $ P
- else P (*this case discards the proof term*);
+parse_translation {*
+ let fun proof_tr [p, P] = Const (@{const_syntax Proof}, dummyT) $ P $ p
+ in [(@{syntax_const "_Proof"}, proof_tr)] end
*}
-parse_translation {* [(@{syntax_const "_Proof"}, proof_tr)] *}
-print_translation {* [(@{const_syntax Proof}, proof_tr')] *}
+(*show_proofs = true displays the proof terms -- they are ENORMOUS*)
+ML {* val (show_proofs, setup_show_proofs) = Attrib.config_bool "show_proofs" (K false) *}
+setup setup_show_proofs
+
+print_translation (advanced) {*
+ let
+ fun proof_tr' ctxt [P, p] =
+ if Config.get ctxt show_proofs then Const (@{syntax_const "_Proof"}, dummyT) $ p $ P
+ else P
+ in [(@{const_syntax Proof}, proof_tr')] end
+*}
axioms
--- a/src/HOL/Auth/Event.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Auth/Event.thy Wed Sep 01 17:19:47 2010 +0200
@@ -22,14 +22,6 @@
consts
bad :: "agent set" -- {* compromised agents *}
- knows :: "agent => event list => msg set"
-
-
-text{*The constant "spies" is retained for compatibility's sake*}
-
-abbreviation (input)
- spies :: "event list => msg set" where
- "spies == knows Spy"
text{*Spy has access to his own key for spoof messages, but Server is secure*}
specification (bad)
@@ -37,9 +29,10 @@
Server_not_bad [iff]: "Server \<notin> bad"
by (rule exI [of _ "{Spy}"], simp)
-primrec
+primrec knows :: "agent => event list => msg set"
+where
knows_Nil: "knows A [] = initState A"
- knows_Cons:
+| knows_Cons:
"knows A (ev # evs) =
(if A = Spy then
(case ev of
@@ -62,14 +55,20 @@
therefore the oops case must use Notes
*)
-consts
- (*Set of items that might be visible to somebody:
+text{*The constant "spies" is retained for compatibility's sake*}
+
+abbreviation (input)
+ spies :: "event list => msg set" where
+ "spies == knows Spy"
+
+
+(*Set of items that might be visible to somebody:
complement of the set of fresh items*)
- used :: "event list => msg set"
-primrec
+primrec used :: "event list => msg set"
+where
used_Nil: "used [] = (UN B. parts (initState B))"
- used_Cons: "used (ev # evs) =
+| used_Cons: "used (ev # evs) =
(case ev of
Says A B X => parts {X} \<union> used evs
| Gets A X => used evs
--- a/src/HOL/Auth/NS_Public_Bad.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Auth/NS_Public_Bad.thy Wed Sep 01 17:19:47 2010 +0200
@@ -203,7 +203,7 @@
apply clarify
apply (frule_tac A' = A in
Says_imp_knows_Spy [THEN parts.Inj, THEN unique_NB], auto)
-apply (rename_tac C B' evs3)
+apply (rename_tac evs3 B' C)
txt{*This is the attack!
@{subgoals[display,indent=0,margin=65]}
*}
--- a/src/HOL/Auth/Yahalom.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Auth/Yahalom.thy Wed Sep 01 17:19:47 2010 +0200
@@ -197,6 +197,7 @@
apply (erule yahalom.induct,
drule_tac [7] YM4_analz_knows_Spy, analz_freshK, spy_analz, blast)
apply (simp only: Says_Server_not_range analz_image_freshK_simps)
+apply safe
done
lemma analz_insert_freshK:
--- a/src/HOL/Boogie/Tools/boogie_commands.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Boogie/Tools/boogie_commands.ML Wed Sep 01 17:19:47 2010 +0200
@@ -91,7 +91,7 @@
| _ => (pair ts, K I))
val discharge = fold (Boogie_VCs.discharge o pair vc_name)
- fun after_qed [thms] = ProofContext.theory (discharge (vcs ~~ thms))
+ fun after_qed [thms] = ProofContext.background_theory (discharge (vcs ~~ thms))
| after_qed _ = I
in
ProofContext.init_global thy
--- a/src/HOL/Boogie/Tools/boogie_loader.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Boogie/Tools/boogie_loader.ML Wed Sep 01 17:19:47 2010 +0200
@@ -504,7 +504,7 @@
in
Const (@{const_name If}, [@{typ bool}, T, T] ---> T) $ c $ t1 $ t2
end) ||
- binexp "implies" (binop @{term "op -->"}) ||
+ binexp "implies" (binop @{term HOL.implies}) ||
scan_line "distinct" num :|-- scan_count exp >>
(fn [] => @{term True}
| ts as (t :: _) => mk_distinct (Term.fastype_of t) ts) ||
--- a/src/HOL/Boogie/Tools/boogie_tactics.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Boogie/Tools/boogie_tactics.ML Wed Sep 01 17:19:47 2010 +0200
@@ -50,11 +50,11 @@
local
- fun explode_conj (@{term "op &"} $ t $ u) = explode_conj t @ explode_conj u
+ fun explode_conj (@{term HOL.conj} $ t $ u) = explode_conj t @ explode_conj u
| explode_conj t = [t]
- fun splt (ts, @{term "op -->"} $ t $ u) = splt (ts @ explode_conj t, u)
- | splt (ts, @{term "op &"} $ t $ u) = splt (ts, t) @ splt (ts, u)
+ fun splt (ts, @{term HOL.implies} $ t $ u) = splt (ts @ explode_conj t, u)
+ | splt (ts, @{term HOL.conj} $ t $ u) = splt (ts, t) @ splt (ts, u)
| splt (ts, @{term assert_at} $ _ $ t) = [(ts, t)]
| splt (_, @{term True}) = []
| splt tp = [tp]
--- a/src/HOL/Boogie/Tools/boogie_vcs.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Boogie/Tools/boogie_vcs.ML Wed Sep 01 17:19:47 2010 +0200
@@ -59,12 +59,12 @@
fun vc_of @{term True} = NONE
| vc_of (@{term assert_at} $ Free (n, _) $ t) =
SOME (Assert ((n, t), True))
- | vc_of (@{term "op -->"} $ @{term True} $ u) = vc_of u
- | vc_of (@{term "op -->"} $ t $ u) =
+ | vc_of (@{term HOL.implies} $ @{term True} $ u) = vc_of u
+ | vc_of (@{term HOL.implies} $ t $ u) =
vc_of u |> Option.map (assume t)
- | vc_of (@{term "op &"} $ (@{term assert_at} $ Free (n, _) $ t) $ u) =
+ | vc_of (@{term HOL.conj} $ (@{term assert_at} $ Free (n, _) $ t) $ u) =
SOME (vc_of u |> the_default True |> assert (n, t))
- | vc_of (@{term "op &"} $ t $ u) =
+ | vc_of (@{term HOL.conj} $ t $ u) =
(case (vc_of t, vc_of u) of
(NONE, r) => r
| (l, NONE) => l
@@ -74,9 +74,9 @@
val prop_of_vc =
let
- fun mk_conj t u = @{term "op &"} $ t $ u
+ fun mk_conj t u = @{term HOL.conj} $ t $ u
- fun term_of (Assume (t, v)) = @{term "op -->"} $ t $ term_of v
+ fun term_of (Assume (t, v)) = @{term HOL.implies} $ t $ term_of v
| term_of (Assert ((n, t), v)) =
mk_conj (@{term assert_at} $ Free (n, @{typ bool}) $ t) (term_of v)
| term_of (Ignore v) = term_of v
--- a/src/HOL/Code_Evaluation.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Code_Evaluation.thy Wed Sep 01 17:19:47 2010 +0200
@@ -162,7 +162,7 @@
subsubsection {* Code generator setup *}
lemmas [code del] = term.recs term.cases term.size
-lemma [code, code del]: "eq_class.eq (t1\<Colon>term) t2 \<longleftrightarrow> eq_class.eq t1 t2" ..
+lemma [code, code del]: "HOL.equal (t1\<Colon>term) t2 \<longleftrightarrow> HOL.equal t1 t2" ..
lemma [code, code del]: "(term_of \<Colon> typerep \<Rightarrow> term) = term_of" ..
lemma [code, code del]: "(term_of \<Colon> term \<Rightarrow> term) = term_of" ..
--- a/src/HOL/Code_Numeral.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Code_Numeral.thy Wed Sep 01 17:19:47 2010 +0200
@@ -115,12 +115,12 @@
lemmas [code del] = code_numeral.recs code_numeral.cases
lemma [code]:
- "eq_class.eq k l \<longleftrightarrow> eq_class.eq (nat_of k) (nat_of l)"
- by (cases k, cases l) (simp add: eq)
+ "HOL.equal k l \<longleftrightarrow> HOL.equal (nat_of k) (nat_of l)"
+ by (cases k, cases l) (simp add: equal)
lemma [code nbe]:
- "eq_class.eq (k::code_numeral) k \<longleftrightarrow> True"
- by (rule HOL.eq_refl)
+ "HOL.equal (k::code_numeral) k \<longleftrightarrow> True"
+ by (rule equal_refl)
subsection {* Code numerals as datatype of ints *}
@@ -301,7 +301,7 @@
(Haskell "Integer")
(Scala "BigInt")
-code_instance code_numeral :: eq
+code_instance code_numeral :: equal
(Haskell -)
setup {*
@@ -342,7 +342,7 @@
(Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))")
(Eval "!(fn n => fn m =>/ if m = 0/ then (0, n) else/ (n div m, n mod m))")
-code_const "eq_class.eq \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
(SML "!((_ : Int.int) = _)")
(OCaml "Big'_int.eq'_big'_int")
(Haskell infixl 4 "==")
--- a/src/HOL/Decision_Procs/Approximation.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Decision_Procs/Approximation.thy Wed Sep 01 17:19:47 2010 +0200
@@ -3305,7 +3305,7 @@
(Const (@{const_name Set.member}, _) $
Free (name, _) $ _)) = name
| variable_of_bound (Const (@{const_name Trueprop}, _) $
- (Const (@{const_name "op ="}, _) $
+ (Const (@{const_name HOL.eq}, _) $
Free (name, _) $ _)) = name
| variable_of_bound t = raise TERM ("variable_of_bound", [t])
@@ -3422,7 +3422,7 @@
ML {*
fun calculated_subterms (@{const Trueprop} $ t) = calculated_subterms t
- | calculated_subterms (@{const "op -->"} $ _ $ t) = calculated_subterms t
+ | calculated_subterms (@{const HOL.implies} $ _ $ t) = calculated_subterms t
| calculated_subterms (@{term "op <= :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) = [t1, t2]
| calculated_subterms (@{term "op < :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) = [t1, t2]
| calculated_subterms (@{term "op : :: real \<Rightarrow> real set \<Rightarrow> bool"} $ t1 $
--- a/src/HOL/Decision_Procs/Cooper.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Decision_Procs/Cooper.thy Wed Sep 01 17:19:47 2010 +0200
@@ -1952,11 +1952,11 @@
| NONE => error "num_of_term: unsupported dvd")
| fm_of_term ps vs (@{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ t1 $ t2) =
@{code Iff} (fm_of_term ps vs t1, fm_of_term ps vs t2)
- | fm_of_term ps vs (@{term "op &"} $ t1 $ t2) =
+ | fm_of_term ps vs (@{term HOL.conj} $ t1 $ t2) =
@{code And} (fm_of_term ps vs t1, fm_of_term ps vs t2)
- | fm_of_term ps vs (@{term "op |"} $ t1 $ t2) =
+ | fm_of_term ps vs (@{term HOL.disj} $ t1 $ t2) =
@{code Or} (fm_of_term ps vs t1, fm_of_term ps vs t2)
- | fm_of_term ps vs (@{term "op -->"} $ t1 $ t2) =
+ | fm_of_term ps vs (@{term HOL.implies} $ t1 $ t2) =
@{code Imp} (fm_of_term ps vs t1, fm_of_term ps vs t2)
| fm_of_term ps vs (@{term "Not"} $ t') =
@{code NOT} (fm_of_term ps vs t')
@@ -2016,7 +2016,7 @@
fun term_bools acc t =
let
- val is_op = member (op =) [@{term "op &"}, @{term "op |"}, @{term "op -->"}, @{term "op = :: bool => _"},
+ val is_op = member (op =) [@{term HOL.conj}, @{term HOL.disj}, @{term HOL.implies}, @{term "op = :: bool => _"},
@{term "op = :: int => _"}, @{term "op < :: int => _"},
@{term "op <= :: int => _"}, @{term "Not"}, @{term "All :: (int => _) => _"},
@{term "Ex :: (int => _) => _"}, @{term "True"}, @{term "False"}]
--- a/src/HOL/Decision_Procs/Dense_Linear_Order.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Decision_Procs/Dense_Linear_Order.thy Wed Sep 01 17:19:47 2010 +0200
@@ -519,9 +519,9 @@
val [lt, le] = map (Morphism.term phi) [@{term "op \<sqsubset>"}, @{term "op \<sqsubseteq>"}]
fun h x t =
case term_of t of
- Const(@{const_name "op ="}, _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
+ Const(@{const_name HOL.eq}, _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
else Ferrante_Rackoff_Data.Nox
- | @{term "Not"}$(Const(@{const_name "op ="}, _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
+ | @{term "Not"}$(Const(@{const_name HOL.eq}, _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
else Ferrante_Rackoff_Data.Nox
| b$y$z => if Term.could_unify (b, lt) then
if term_of x aconv y then Ferrante_Rackoff_Data.Lt
@@ -771,7 +771,7 @@
in rth end
| _ => Thm.reflexive ct)
-| Const(@{const_name "op ="},_)$_$Const(@{const_name Groups.zero},_) =>
+| Const(@{const_name HOL.eq},_)$_$Const(@{const_name Groups.zero},_) =>
(case whatis x (Thm.dest_arg1 ct) of
("c*x+t",[c,t]) =>
let
@@ -835,7 +835,7 @@
val rth = Thm.transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
in rth end
-| Const(@{const_name "op ="},_)$a$b =>
+| Const(@{const_name HOL.eq},_)$a$b =>
let val (ca,cb) = Thm.dest_binop ct
val T = ctyp_of_term ca
val th = instantiate' [SOME T] [SOME ca, SOME cb] eq_iff_diff_eq_0
@@ -844,7 +844,7 @@
(Semiring_Normalizer.semiring_normalize_ord_conv @{context} (earlier vs)))) th
val rth = Thm.transitive nth (xnormalize_conv ctxt vs (Thm.rhs_of nth))
in rth end
-| @{term "Not"} $(Const(@{const_name "op ="},_)$a$b) => Conv.arg_conv (field_isolate_conv phi ctxt vs) ct
+| @{term "Not"} $(Const(@{const_name HOL.eq},_)$a$b) => Conv.arg_conv (field_isolate_conv phi ctxt vs) ct
| _ => Thm.reflexive ct
end;
@@ -852,9 +852,9 @@
let
fun h x t =
case term_of t of
- Const(@{const_name "op ="}, _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
+ Const(@{const_name HOL.eq}, _)$y$z => if term_of x aconv y then Ferrante_Rackoff_Data.Eq
else Ferrante_Rackoff_Data.Nox
- | @{term "Not"}$(Const(@{const_name "op ="}, _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
+ | @{term "Not"}$(Const(@{const_name HOL.eq}, _)$y$z) => if term_of x aconv y then Ferrante_Rackoff_Data.NEq
else Ferrante_Rackoff_Data.Nox
| Const(@{const_name Orderings.less},_)$y$z =>
if term_of x aconv y then Ferrante_Rackoff_Data.Lt
--- a/src/HOL/Decision_Procs/Ferrack.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Decision_Procs/Ferrack.thy Wed Sep 01 17:19:47 2010 +0200
@@ -1996,9 +1996,9 @@
@{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
| fm_of_term vs (@{term "op \<longleftrightarrow> :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ t1 $ t2) =
@{code Iff} (fm_of_term vs t1, fm_of_term vs t2)
- | fm_of_term vs (@{term "op &"} $ t1 $ t2) = @{code And} (fm_of_term vs t1, fm_of_term vs t2)
- | fm_of_term vs (@{term "op |"} $ t1 $ t2) = @{code Or} (fm_of_term vs t1, fm_of_term vs t2)
- | fm_of_term vs (@{term "op -->"} $ t1 $ t2) = @{code Imp} (fm_of_term vs t1, fm_of_term vs t2)
+ | fm_of_term vs (@{term HOL.conj} $ t1 $ t2) = @{code And} (fm_of_term vs t1, fm_of_term vs t2)
+ | fm_of_term vs (@{term HOL.disj} $ t1 $ t2) = @{code Or} (fm_of_term vs t1, fm_of_term vs t2)
+ | fm_of_term vs (@{term HOL.implies} $ t1 $ t2) = @{code Imp} (fm_of_term vs t1, fm_of_term vs t2)
| fm_of_term vs (@{term "Not"} $ t') = @{code NOT} (fm_of_term vs t')
| fm_of_term vs (Const (@{const_name Ex}, _) $ Abs (xn, xT, p)) =
@{code E} (fm_of_term (("", dummyT) :: vs) p)
--- a/src/HOL/Decision_Procs/MIR.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Decision_Procs/MIR.thy Wed Sep 01 17:19:47 2010 +0200
@@ -5837,11 +5837,11 @@
@{code Dvd} (HOLogic.dest_numeral t1, num_of_term vs t2)
| fm_of_term vs (@{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ t1 $ t2) =
@{code Iff} (fm_of_term vs t1, fm_of_term vs t2)
- | fm_of_term vs (@{term "op &"} $ t1 $ t2) =
+ | fm_of_term vs (@{term HOL.conj} $ t1 $ t2) =
@{code And} (fm_of_term vs t1, fm_of_term vs t2)
- | fm_of_term vs (@{term "op |"} $ t1 $ t2) =
+ | fm_of_term vs (@{term HOL.disj} $ t1 $ t2) =
@{code Or} (fm_of_term vs t1, fm_of_term vs t2)
- | fm_of_term vs (@{term "op -->"} $ t1 $ t2) =
+ | fm_of_term vs (@{term HOL.implies} $ t1 $ t2) =
@{code Imp} (fm_of_term vs t1, fm_of_term vs t2)
| fm_of_term vs (@{term "Not"} $ t') =
@{code NOT} (fm_of_term vs t')
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Wed Sep 01 17:19:47 2010 +0200
@@ -912,7 +912,7 @@
definition "lt p = (case p of CP (C c) \<Rightarrow> if 0>\<^sub>N c then T else F| _ \<Rightarrow> Lt p)"
definition "le p = (case p of CP (C c) \<Rightarrow> if 0\<ge>\<^sub>N c then T else F | _ \<Rightarrow> Le p)"
-definition "eq p = (case p of CP (C c) \<Rightarrow> if c = 0\<^sub>N then T else F | _ \<Rightarrow> Eq p)"
+definition eq where "eq p = (case p of CP (C c) \<Rightarrow> if c = 0\<^sub>N then T else F | _ \<Rightarrow> Eq p)"
definition "neq p = not (eq p)"
lemma lt: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (lt p) = Ifm vs bs (Lt p)"
@@ -2954,13 +2954,13 @@
fun powt rT = Const(@{const_name "power"}, [rT,@{typ "nat"}] ---> rT);
val brT = [bT, bT] ---> bT;
val nott = @{term "Not"};
-val conjt = @{term "op &"};
-val disjt = @{term "op |"};
-val impt = @{term "op -->"};
+val conjt = @{term HOL.conj};
+val disjt = @{term HOL.disj};
+val impt = @{term HOL.implies};
val ifft = @{term "op = :: bool => _"}
fun llt rT = Const(@{const_name Orderings.less},rrT rT);
fun lle rT = Const(@{const_name Orderings.less},rrT rT);
-fun eqt rT = Const(@{const_name "op ="},rrT rT);
+fun eqt rT = Const(@{const_name HOL.eq},rrT rT);
fun rz rT = Const(@{const_name Groups.zero},rT);
fun dest_nat t = case t of
@@ -3018,10 +3018,10 @@
Const(@{const_name True},_) => @{code T}
| Const(@{const_name False},_) => @{code F}
| Const(@{const_name Not},_)$p => @{code NOT} (fm_of_term m m' p)
- | Const(@{const_name "op &"},_)$p$q => @{code And} (fm_of_term m m' p, fm_of_term m m' q)
- | Const(@{const_name "op |"},_)$p$q => @{code Or} (fm_of_term m m' p, fm_of_term m m' q)
- | Const(@{const_name "op -->"},_)$p$q => @{code Imp} (fm_of_term m m' p, fm_of_term m m' q)
- | Const(@{const_name "op ="},ty)$p$q =>
+ | Const(@{const_name HOL.conj},_)$p$q => @{code And} (fm_of_term m m' p, fm_of_term m m' q)
+ | Const(@{const_name HOL.disj},_)$p$q => @{code Or} (fm_of_term m m' p, fm_of_term m m' q)
+ | Const(@{const_name HOL.implies},_)$p$q => @{code Imp} (fm_of_term m m' p, fm_of_term m m' q)
+ | Const(@{const_name HOL.eq},ty)$p$q =>
if domain_type ty = bT then @{code Iff} (fm_of_term m m' p, fm_of_term m m' q)
else @{code Eq} (@{code Sub} (tm_of_term m m' p, tm_of_term m m' q))
| Const(@{const_name Orderings.less},_)$p$q =>
--- a/src/HOL/Decision_Procs/commutative_ring_tac.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Decision_Procs/commutative_ring_tac.ML Wed Sep 01 17:19:47 2010 +0200
@@ -65,7 +65,7 @@
(* reification of the equation *)
val cr_sort = @{sort "comm_ring_1"};
-fun reif_eq thy (eq as Const(@{const_name "op ="}, Type("fun", [T, _])) $ lhs $ rhs) =
+fun reif_eq thy (eq as Const(@{const_name HOL.eq}, Type("fun", [T, _])) $ lhs $ rhs) =
if Sign.of_sort thy (T, cr_sort) then
let
val fs = OldTerm.term_frees eq;
--- a/src/HOL/Decision_Procs/ferrante_rackoff.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Decision_Procs/ferrante_rackoff.ML Wed Sep 01 17:19:47 2010 +0200
@@ -33,12 +33,12 @@
{isolate_conv = icv, whatis = wi, simpset = simpset}):entry) =
let
fun uset (vars as (x::vs)) p = case term_of p of
- Const(@{const_name "op &"}, _)$ _ $ _ =>
+ Const(@{const_name HOL.conj}, _)$ _ $ _ =>
let
val ((b,l),r) = Thm.dest_comb p |>> Thm.dest_comb
val (lS,lth) = uset vars l val (rS, rth) = uset vars r
in (lS@rS, Drule.binop_cong_rule b lth rth) end
- | Const(@{const_name "op |"}, _)$ _ $ _ =>
+ | Const(@{const_name HOL.disj}, _)$ _ $ _ =>
let
val ((b,l),r) = Thm.dest_comb p |>> Thm.dest_comb
val (lS,lth) = uset vars l val (rS, rth) = uset vars r
@@ -122,12 +122,12 @@
fun decomp_mpinf fm =
case term_of fm of
- Const(@{const_name "op &"},_)$_$_ =>
+ Const(@{const_name HOL.conj},_)$_$_ =>
let val (p,q) = Thm.dest_binop fm
in ([p,q], myfwd (minf_conj,pinf_conj, nmi_conj, npi_conj,ld_conj)
(Thm.cabs x p) (Thm.cabs x q))
end
- | Const(@{const_name "op |"},_)$_$_ =>
+ | Const(@{const_name HOL.disj},_)$_$_ =>
let val (p,q) = Thm.dest_binop fm
in ([p,q],myfwd (minf_disj, pinf_disj, nmi_disj, npi_disj,ld_disj)
(Thm.cabs x p) (Thm.cabs x q))
@@ -175,15 +175,15 @@
let
fun h bounds tm =
(case term_of tm of
- Const (@{const_name "op ="}, T) $ _ $ _ =>
+ Const (@{const_name HOL.eq}, T) $ _ $ _ =>
if domain_type T = HOLogic.boolT then find_args bounds tm
else Thm.dest_fun2 tm
| Const (@{const_name Not}, _) $ _ => h bounds (Thm.dest_arg tm)
| Const (@{const_name All}, _) $ _ => find_body bounds (Thm.dest_arg tm)
| Const (@{const_name Ex}, _) $ _ => find_body bounds (Thm.dest_arg tm)
- | Const (@{const_name "op &"}, _) $ _ $ _ => find_args bounds tm
- | Const (@{const_name "op |"}, _) $ _ $ _ => find_args bounds tm
- | Const (@{const_name "op -->"}, _) $ _ $ _ => find_args bounds tm
+ | Const (@{const_name HOL.conj}, _) $ _ $ _ => find_args bounds tm
+ | Const (@{const_name HOL.disj}, _) $ _ $ _ => find_args bounds tm
+ | Const (@{const_name HOL.implies}, _) $ _ $ _ => find_args bounds tm
| Const ("==>", _) $ _ $ _ => find_args bounds tm
| Const ("==", _) $ _ $ _ => find_args bounds tm
| Const ("all", _) $ _ => find_body bounds (Thm.dest_arg tm)
--- a/src/HOL/Decision_Procs/langford.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Decision_Procs/langford.ML Wed Sep 01 17:19:47 2010 +0200
@@ -69,28 +69,28 @@
val all_conjuncts =
let fun h acc ct =
case term_of ct of
- @{term "op &"}$_$_ => h (h acc (Thm.dest_arg ct)) (Thm.dest_arg1 ct)
+ @{term HOL.conj}$_$_ => h (h acc (Thm.dest_arg ct)) (Thm.dest_arg1 ct)
| _ => ct::acc
in h [] end;
fun conjuncts ct =
case term_of ct of
- @{term "op &"}$_$_ => (Thm.dest_arg1 ct)::(conjuncts (Thm.dest_arg ct))
+ @{term HOL.conj}$_$_ => (Thm.dest_arg1 ct)::(conjuncts (Thm.dest_arg ct))
| _ => [ct];
fun fold1 f = foldr1 (uncurry f);
-val list_conj = fold1 (fn c => fn c' => Thm.capply (Thm.capply @{cterm "op &"} c) c') ;
+val list_conj = fold1 (fn c => fn c' => Thm.capply (Thm.capply @{cterm HOL.conj} c) c') ;
fun mk_conj_tab th =
let fun h acc th =
case prop_of th of
- @{term "Trueprop"}$(@{term "op &"}$p$q) =>
+ @{term "Trueprop"}$(@{term HOL.conj}$p$q) =>
h (h acc (th RS conjunct2)) (th RS conjunct1)
| @{term "Trueprop"}$p => (p,th)::acc
in fold (Termtab.insert Thm.eq_thm) (h [] th) Termtab.empty end;
-fun is_conj (@{term "op &"}$_$_) = true
+fun is_conj (@{term HOL.conj}$_$_) = true
| is_conj _ = false;
fun prove_conj tab cjs =
@@ -116,7 +116,7 @@
fun contains x ct = member (op aconv) (OldTerm.term_frees (term_of ct)) (term_of x);
fun is_eqx x eq = case term_of eq of
- Const(@{const_name "op ="},_)$l$r => l aconv term_of x orelse r aconv term_of x
+ Const(@{const_name HOL.eq},_)$l$r => l aconv term_of x orelse r aconv term_of x
| _ => false ;
local
@@ -176,16 +176,16 @@
let
fun h bounds tm =
(case term_of tm of
- Const (@{const_name "op ="}, T) $ _ $ _ =>
+ Const (@{const_name HOL.eq}, T) $ _ $ _ =>
if domain_type T = HOLogic.boolT then find_args bounds tm
else Thm.dest_fun2 tm
| Const (@{const_name Not}, _) $ _ => h bounds (Thm.dest_arg tm)
| Const (@{const_name All}, _) $ _ => find_body bounds (Thm.dest_arg tm)
| Const ("all", _) $ _ => find_body bounds (Thm.dest_arg tm)
| Const (@{const_name Ex}, _) $ _ => find_body bounds (Thm.dest_arg tm)
- | Const (@{const_name "op &"}, _) $ _ $ _ => find_args bounds tm
- | Const (@{const_name "op |"}, _) $ _ $ _ => find_args bounds tm
- | Const (@{const_name "op -->"}, _) $ _ $ _ => find_args bounds tm
+ | Const (@{const_name HOL.conj}, _) $ _ $ _ => find_args bounds tm
+ | Const (@{const_name HOL.disj}, _) $ _ $ _ => find_args bounds tm
+ | Const (@{const_name HOL.implies}, _) $ _ $ _ => find_args bounds tm
| Const ("==>", _) $ _ $ _ => find_args bounds tm
| Const ("==", _) $ _ $ _ => find_args bounds tm
| Const (@{const_name Trueprop}, _) $ _ => h bounds (Thm.dest_arg tm)
--- a/src/HOL/HOL.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/HOL.thy Wed Sep 01 17:19:47 2010 +0200
@@ -30,6 +30,7 @@
"~~/src/Tools/induct.ML"
("~~/src/Tools/induct_tacs.ML")
("Tools/recfun_codegen.ML")
+ "Tools/try.ML"
begin
setup {* Intuitionistic.method_setup @{binding iprover} *}
@@ -57,18 +58,12 @@
False :: bool
Not :: "bool => bool" ("~ _" [40] 40)
-setup Sign.root_path
+ conj :: "[bool, bool] => bool" (infixr "&" 35)
+ disj :: "[bool, bool] => bool" (infixr "|" 30)
+ implies :: "[bool, bool] => bool" (infixr "-->" 25)
-consts
- "op &" :: "[bool, bool] => bool" (infixr "&" 35)
- "op |" :: "[bool, bool] => bool" (infixr "|" 30)
- "op -->" :: "[bool, bool] => bool" (infixr "-->" 25)
+ eq :: "['a, 'a] => bool" (infixl "=" 50)
- "op =" :: "['a, 'a] => bool" (infixl "=" 50)
-
-setup Sign.local_path
-
-consts
The :: "('a => bool) => 'a"
All :: "('a => bool) => bool" (binder "ALL " 10)
Ex :: "('a => bool) => bool" (binder "EX " 10)
@@ -78,7 +73,7 @@
subsubsection {* Additional concrete syntax *}
notation (output)
- "op =" (infix "=" 50)
+ eq (infix "=" 50)
abbreviation
not_equal :: "['a, 'a] => bool" (infixl "~=" 50) where
@@ -89,15 +84,15 @@
notation (xsymbols)
Not ("\<not> _" [40] 40) and
- "op &" (infixr "\<and>" 35) and
- "op |" (infixr "\<or>" 30) and
- "op -->" (infixr "\<longrightarrow>" 25) and
+ conj (infixr "\<and>" 35) and
+ disj (infixr "\<or>" 30) and
+ implies (infixr "\<longrightarrow>" 25) and
not_equal (infix "\<noteq>" 50)
notation (HTML output)
Not ("\<not> _" [40] 40) and
- "op &" (infixr "\<and>" 35) and
- "op |" (infixr "\<or>" 30) and
+ conj (infixr "\<and>" 35) and
+ disj (infixr "\<or>" 30) and
not_equal (infix "\<noteq>" 50)
abbreviation (iff)
@@ -183,8 +178,8 @@
True_or_False: "(P=True) | (P=False)"
finalconsts
- "op ="
- "op -->"
+ eq
+ implies
The
definition If :: "bool \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("(if (_)/ then (_)/ else (_))" [0, 0, 10] 10) where
@@ -864,7 +859,7 @@
setup {*
let
- fun non_bool_eq (@{const_name "op ="}, Type (_, [T, _])) = T <> @{typ bool}
+ fun non_bool_eq (@{const_name HOL.eq}, Type (_, [T, _])) = T <> @{typ bool}
| non_bool_eq _ = false;
val hyp_subst_tac' =
SUBGOAL (fn (goal, i) =>
@@ -930,7 +925,7 @@
(
val thy = @{theory}
type claset = Classical.claset
- val equality_name = @{const_name "op ="}
+ val equality_name = @{const_name HOL.eq}
val not_name = @{const_name Not}
val notE = @{thm notE}
val ccontr = @{thm ccontr}
@@ -1578,7 +1573,7 @@
val atomize_conjL = @{thm atomize_conjL}
val atomize_disjL = @{thm atomize_disjL}
val operator_names =
- [@{const_name "op |"}, @{const_name "op &"}, @{const_name Ex}]
+ [@{const_name HOL.disj}, @{const_name HOL.conj}, @{const_name Ex}]
);
*}
@@ -1737,8 +1732,8 @@
"True" ("true")
"False" ("false")
"Not" ("Bool.not")
- "op |" ("(_ orelse/ _)")
- "op &" ("(_ andalso/ _)")
+ HOL.disj ("(_ orelse/ _)")
+ HOL.conj ("(_ andalso/ _)")
"If" ("(if _/ then _/ else _)")
setup {*
@@ -1746,8 +1741,8 @@
fun eq_codegen thy defs dep thyname b t gr =
(case strip_comb t of
- (Const (@{const_name "op ="}, Type (_, [Type ("fun", _), _])), _) => NONE
- | (Const (@{const_name "op ="}, _), [t, u]) =>
+ (Const (@{const_name HOL.eq}, Type (_, [Type ("fun", _), _])), _) => NONE
+ | (Const (@{const_name HOL.eq}, _), [t, u]) =>
let
val (pt, gr') = Codegen.invoke_codegen thy defs dep thyname false t gr;
val (pu, gr'') = Codegen.invoke_codegen thy defs dep thyname false u gr';
@@ -1756,7 +1751,7 @@
SOME (Codegen.parens
(Pretty.block [pt, Codegen.str " =", Pretty.brk 1, pu]), gr''')
end
- | (t as Const (@{const_name "op ="}, _), ts) => SOME (Codegen.invoke_codegen
+ | (t as Const (@{const_name HOL.eq}, _), ts) => SOME (Codegen.invoke_codegen
thy defs dep thyname b (Codegen.eta_expand t ts 2) gr)
| _ => NONE);
@@ -1775,31 +1770,30 @@
subsubsection {* Equality *}
-class eq =
- fixes eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
- assumes eq_equals: "eq x y \<longleftrightarrow> x = y"
+class equal =
+ fixes equal :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+ assumes equal_eq: "equal x y \<longleftrightarrow> x = y"
begin
-lemma eq [code_unfold, code_inline del]: "eq = (op =)"
- by (rule ext eq_equals)+
+lemma equal [code_unfold, code_inline del]: "equal = (op =)"
+ by (rule ext equal_eq)+
-lemma eq_refl: "eq x x \<longleftrightarrow> True"
- unfolding eq by rule+
+lemma equal_refl: "equal x x \<longleftrightarrow> True"
+ unfolding equal by rule+
-lemma equals_eq: "(op =) \<equiv> eq"
- by (rule eq_reflection) (rule ext, rule ext, rule sym, rule eq_equals)
-
-declare equals_eq [symmetric, code_post]
+lemma eq_equal: "(op =) \<equiv> equal"
+ by (rule eq_reflection) (rule ext, rule ext, rule sym, rule equal_eq)
end
-declare equals_eq [code]
+declare eq_equal [symmetric, code_post]
+declare eq_equal [code]
setup {*
Code_Preproc.map_pre (fn simpset =>
- simpset addsimprocs [Simplifier.simproc_global_i @{theory} "eq" [@{term "op ="}]
+ simpset addsimprocs [Simplifier.simproc_global_i @{theory} "equal" [@{term HOL.eq}]
(fn thy => fn _ => fn Const (_, T) => case strip_type T
- of (Type _ :: _, _) => SOME @{thm equals_eq}
+ of (Type _ :: _, _) => SOME @{thm eq_equal}
| _ => NONE)])
*}
@@ -1839,51 +1833,49 @@
and "(P \<longrightarrow> False) \<longleftrightarrow> \<not> P"
and "(P \<longrightarrow> True) \<longleftrightarrow> True" by simp_all
-instantiation itself :: (type) eq
+instantiation itself :: (type) equal
begin
-definition eq_itself :: "'a itself \<Rightarrow> 'a itself \<Rightarrow> bool" where
- "eq_itself x y \<longleftrightarrow> x = y"
+definition equal_itself :: "'a itself \<Rightarrow> 'a itself \<Rightarrow> bool" where
+ "equal_itself x y \<longleftrightarrow> x = y"
instance proof
-qed (fact eq_itself_def)
+qed (fact equal_itself_def)
end
-lemma eq_itself_code [code]:
- "eq_class.eq TYPE('a) TYPE('a) \<longleftrightarrow> True"
- by (simp add: eq)
+lemma equal_itself_code [code]:
+ "equal TYPE('a) TYPE('a) \<longleftrightarrow> True"
+ by (simp add: equal)
text {* Equality *}
declare simp_thms(6) [code nbe]
setup {*
- Sign.add_const_constraint (@{const_name eq}, SOME @{typ "'a\<Colon>type \<Rightarrow> 'a \<Rightarrow> bool"})
+ Sign.add_const_constraint (@{const_name equal}, SOME @{typ "'a\<Colon>type \<Rightarrow> 'a \<Rightarrow> bool"})
*}
-lemma equals_alias_cert: "OFCLASS('a, eq_class) \<equiv> ((op = :: 'a \<Rightarrow> 'a \<Rightarrow> bool) \<equiv> eq)" (is "?ofclass \<equiv> ?eq")
+lemma equal_alias_cert: "OFCLASS('a, equal_class) \<equiv> ((op = :: 'a \<Rightarrow> 'a \<Rightarrow> bool) \<equiv> equal)" (is "?ofclass \<equiv> ?equal")
proof
assume "PROP ?ofclass"
- show "PROP ?eq"
- by (tactic {* ALLGOALS (rtac (Thm.unconstrainT @{thm equals_eq})) *})
+ show "PROP ?equal"
+ by (tactic {* ALLGOALS (rtac (Thm.unconstrainT @{thm eq_equal})) *})
(fact `PROP ?ofclass`)
next
- assume "PROP ?eq"
+ assume "PROP ?equal"
show "PROP ?ofclass" proof
- qed (simp add: `PROP ?eq`)
+ qed (simp add: `PROP ?equal`)
qed
setup {*
- Sign.add_const_constraint (@{const_name eq}, SOME @{typ "'a\<Colon>eq \<Rightarrow> 'a \<Rightarrow> bool"})
+ Sign.add_const_constraint (@{const_name equal}, SOME @{typ "'a\<Colon>equal \<Rightarrow> 'a \<Rightarrow> bool"})
*}
setup {*
- Nbe.add_const_alias @{thm equals_alias_cert}
+ Nbe.add_const_alias @{thm equal_alias_cert}
*}
-hide_const (open) eq
-
text {* Cases *}
lemma Let_case_cert:
@@ -1904,9 +1896,10 @@
code_abort undefined
+
subsubsection {* Generic code generator target languages *}
-text {* type bool *}
+text {* type @{typ bool} *}
code_type bool
(SML "bool")
@@ -1914,7 +1907,7 @@
(Haskell "Bool")
(Scala "Boolean")
-code_const True and False and Not and "op &" and "op |" and If
+code_const True and False and Not and HOL.conj and HOL.disj and If
(SML "true" and "false" and "not"
and infixl 1 "andalso" and infixl 0 "orelse"
and "!(if (_)/ then (_)/ else (_))")
@@ -1924,7 +1917,7 @@
(Haskell "True" and "False" and "not"
and infixl 3 "&&" and infixl 2 "||"
and "!(if (_)/ then (_)/ else (_))")
- (Scala "true" and "false" and "'!/ _"
+ (Scala "true" and "false" and "'! _"
and infixl 3 "&&" and infixl 1 "||"
and "!(if ((_))/ (_)/ else (_))")
@@ -1939,13 +1932,13 @@
text {* using built-in Haskell equality *}
-code_class eq
+code_class equal
(Haskell "Eq")
-code_const "eq_class.eq"
+code_const "HOL.equal"
(Haskell infixl 4 "==")
-code_const "op ="
+code_const HOL.eq
(Haskell infixl 4 "==")
text {* undefined *}
@@ -2134,4 +2127,6 @@
*}
+hide_const (open) eq equal
+
end
--- a/src/HOL/Imperative_HOL/Heap_Monad.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Imperative_HOL/Heap_Monad.thy Wed Sep 01 17:19:47 2010 +0200
@@ -477,34 +477,44 @@
subsubsection {* Scala *}
code_include Scala "Heap"
-{*import collection.mutable.ArraySeq
-import Natural._
-
-def bind[A, B](f: Unit => A, g: A => Unit => B): Unit => B = (_: Unit) => g (f ()) ()
+{*object Heap {
+ def bind[A, B](f: Unit => A, g: A => Unit => B): Unit => B = (_: Unit) => g (f ()) ()
+}
class Ref[A](x: A) {
var value = x
}
object Ref {
- def apply[A](x: A): Ref[A] = new Ref[A](x)
- def lookup[A](r: Ref[A]): A = r.value
- def update[A](r: Ref[A], x: A): Unit = { r.value = x }
+ def apply[A](x: A): Ref[A] =
+ new Ref[A](x)
+ def lookup[A](r: Ref[A]): A =
+ r.value
+ def update[A](r: Ref[A], x: A): Unit =
+ { r.value = x }
}
object Array {
- def alloc[A](n: Natural)(x: A): ArraySeq[A] = ArraySeq.fill(n.as_Int)(x)
- def make[A](n: Natural)(f: Natural => A): ArraySeq[A] = ArraySeq.tabulate(n.as_Int)((k: Int) => f(Natural(k)))
- def len[A](a: ArraySeq[A]): Natural = Natural(a.length)
- def nth[A](a: ArraySeq[A], n: Natural): A = a(n.as_Int)
- def upd[A](a: ArraySeq[A], n: Natural, x: A): Unit = a.update(n.as_Int, x)
- def freeze[A](a: ArraySeq[A]): List[A] = a.toList
-}*}
+ import collection.mutable.ArraySeq
+ def alloc[A](n: Natural)(x: A): ArraySeq[A] =
+ ArraySeq.fill(n.as_Int)(x)
+ def make[A](n: Natural)(f: Natural => A): ArraySeq[A] =
+ ArraySeq.tabulate(n.as_Int)((k: Int) => f(Natural(k)))
+ def len[A](a: ArraySeq[A]): Natural =
+ Natural(a.length)
+ def nth[A](a: ArraySeq[A], n: Natural): A =
+ a(n.as_Int)
+ def upd[A](a: ArraySeq[A], n: Natural, x: A): Unit =
+ a.update(n.as_Int, x)
+ def freeze[A](a: ArraySeq[A]): List[A] =
+ a.toList
+}
+*}
-code_reserved Scala bind Ref Array
+code_reserved Scala Heap Ref Array
code_type Heap (Scala "Unit/ =>/ _")
-code_const bind (Scala "bind")
+code_const bind (Scala "Heap.bind")
code_const return (Scala "('_: Unit)/ =>/ _")
code_const Heap_Monad.raise' (Scala "!error((_))")
--- a/src/HOL/Import/Generate-HOL/GenHOL4Base.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Import/Generate-HOL/GenHOL4Base.thy Wed Sep 01 17:19:47 2010 +0200
@@ -17,8 +17,8 @@
T > True
F > False
"!" > All
- "/\\" > "op &"
- "\\/" > "op |"
+ "/\\" > HOL.conj
+ "\\/" > HOL.disj
"?" > Ex
"?!" > Ex1
"~" > Not
--- a/src/HOL/Import/Generate-HOLLight/GenHOLLight.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Import/Generate-HOLLight/GenHOLLight.thy Wed Sep 01 17:19:47 2010 +0200
@@ -53,10 +53,10 @@
F > False
ONE_ONE > HOL4Setup.ONE_ONE
ONTO > Fun.surj
- "=" > "op ="
- "==>" > "op -->"
- "/\\" > "op &"
- "\\/" > "op |"
+ "=" > HOL.eq
+ "==>" > HOL.implies
+ "/\\" > HOL.conj
+ "\\/" > HOL.disj
"!" > All
"?" > Ex
"?!" > Ex1
--- a/src/HOL/Import/HOL/bool.imp Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Import/HOL/bool.imp Wed Sep 01 17:19:47 2010 +0200
@@ -14,7 +14,7 @@
const_maps
"~" > "HOL.Not"
"bool_case" > "Datatype.bool.bool_case"
- "\\/" > "op |"
+ "\\/" > HOL.disj
"TYPE_DEFINITION" > "HOL4Setup.TYPE_DEFINITION"
"T" > "HOL.True"
"RES_SELECT" > "HOL4Base.bool.RES_SELECT"
@@ -30,7 +30,7 @@
"ARB" > "HOL4Base.bool.ARB"
"?!" > "HOL.Ex1"
"?" > "HOL.Ex"
- "/\\" > "op &"
+ "/\\" > HOL.conj
"!" > "HOL.All"
thm_maps
--- a/src/HOL/Import/HOLLight/hollight.imp Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Import/HOLLight/hollight.imp Wed Sep 01 17:19:47 2010 +0200
@@ -115,7 +115,7 @@
"_10303" > "HOLLight.hollight._10303"
"_10302" > "HOLLight.hollight._10302"
"_0" > "0" :: "nat"
- "\\/" > "op |"
+ "\\/" > HOL.disj
"ZRECSPACE" > "HOLLight.hollight.ZRECSPACE"
"ZIP" > "HOLLight.hollight.ZIP"
"ZCONSTR" > "HOLLight.hollight.ZCONSTR"
@@ -233,11 +233,11 @@
"?" > "HOL.Ex"
">=" > "HOLLight.hollight.>="
">" > "HOLLight.hollight.>"
- "==>" > "op -->"
- "=" > "op ="
+ "==>" > HOL.implies
+ "=" > HOL.eq
"<=" > "HOLLight.hollight.<="
"<" > "HOLLight.hollight.<"
- "/\\" > "op &"
+ "/\\" > HOL.conj
"-" > "Groups.minus" :: "nat => nat => nat"
"," > "Product_Type.Pair"
"+" > "Groups.plus" :: "nat => nat => nat"
--- a/src/HOL/Import/hol4rews.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Import/hol4rews.ML Wed Sep 01 17:19:47 2010 +0200
@@ -627,8 +627,8 @@
thy |> add_hol4_type_mapping "min" "bool" false @{type_name bool}
|> add_hol4_type_mapping "min" "fun" false "fun"
|> add_hol4_type_mapping "min" "ind" false @{type_name ind}
- |> add_hol4_const_mapping "min" "=" false @{const_name "op ="}
- |> add_hol4_const_mapping "min" "==>" false @{const_name "op -->"}
+ |> add_hol4_const_mapping "min" "=" false @{const_name HOL.eq}
+ |> add_hol4_const_mapping "min" "==>" false @{const_name HOL.implies}
|> add_hol4_const_mapping "min" "@" false @{const_name "Eps"}
in
val hol4_setup =
--- a/src/HOL/Import/import_syntax.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Import/import_syntax.ML Wed Sep 01 17:19:47 2010 +0200
@@ -148,11 +148,11 @@
val _ = TextIO.closeIn is
val orig_source = Source.of_string inp
val symb_source = Symbol.source {do_recover = false} orig_source
- val lexes = Unsynchronized.ref
- (Scan.make_lexicon (map Symbol.explode ["import_segment","ignore_thms","import","end",">","::","const_maps","const_moves","thm_maps","const_renames","type_maps","def_maps"]),
+ val lexes =
+ (Scan.make_lexicon
+ (map Symbol.explode ["import_segment","ignore_thms","import","end",">","::","const_maps","const_moves","thm_maps","const_renames","type_maps","def_maps"]),
Scan.empty_lexicon)
- val get_lexes = fn () => !lexes
- val token_source = Token.source {do_recover = NONE} get_lexes Position.start symb_source
+ val token_source = Token.source {do_recover = NONE} (K lexes) Position.start symb_source
val token_list = filter_out (Token.is_kind Token.Space) (Source.exhaust token_source)
val import_segmentP = Parse.$$$ "import_segment" |-- import_segment
val type_mapsP = Parse.$$$ "type_maps" |-- type_maps
--- a/src/HOL/Import/proof_kernel.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Import/proof_kernel.ML Wed Sep 01 17:19:47 2010 +0200
@@ -1205,7 +1205,7 @@
fun non_trivial_term_consts t = fold_aterms
(fn Const (c, _) =>
if c = @{const_name Trueprop} orelse c = @{const_name All}
- orelse c = @{const_name "op -->"} orelse c = @{const_name "op &"} orelse c = @{const_name "op ="}
+ orelse c = @{const_name HOL.implies} orelse c = @{const_name HOL.conj} orelse c = @{const_name HOL.eq}
then I else insert (op =) c
| _ => I) t [];
@@ -1213,11 +1213,11 @@
let
fun add_consts (Const (c, _), cs) =
(case c of
- @{const_name "op ="} => insert (op =) "==" cs
- | @{const_name "op -->"} => insert (op =) "==>" cs
+ @{const_name HOL.eq} => insert (op =) "==" cs
+ | @{const_name HOL.implies} => insert (op =) "==>" cs
| @{const_name All} => cs
| "all" => cs
- | @{const_name "op &"} => cs
+ | @{const_name HOL.conj} => cs
| @{const_name Trueprop} => cs
| _ => insert (op =) c cs)
| add_consts (t $ u, cs) = add_consts (t, add_consts (u, cs))
@@ -1476,10 +1476,10 @@
fun mk_COMB th1 th2 thy =
let
val (f,g) = case concl_of th1 of
- _ $ (Const(@{const_name "op ="},_) $ f $ g) => (f,g)
+ _ $ (Const(@{const_name HOL.eq},_) $ f $ g) => (f,g)
| _ => raise ERR "mk_COMB" "First theorem not an equality"
val (x,y) = case concl_of th2 of
- _ $ (Const(@{const_name "op ="},_) $ x $ y) => (x,y)
+ _ $ (Const(@{const_name HOL.eq},_) $ x $ y) => (x,y)
| _ => raise ERR "mk_COMB" "Second theorem not an equality"
val fty = type_of f
val (fd,fr) = dom_rng fty
@@ -1521,7 +1521,7 @@
val th1 = norm_hyps th1
val th2 = norm_hyps th2
val (l,r) = case concl_of th of
- _ $ (Const(@{const_name "op |"},_) $ l $ r) => (l,r)
+ _ $ (Const(@{const_name HOL.disj},_) $ l $ r) => (l,r)
| _ => raise ERR "DISJ_CASES" "Conclusion not a disjunction"
val th1' = rearrange thy (HOLogic.mk_Trueprop l) th1
val th2' = rearrange thy (HOLogic.mk_Trueprop r) th2
@@ -1788,7 +1788,7 @@
val cv = cterm_of thy v
val th1 = implies_elim_all (beta_eta_thm th)
val (f,g) = case concl_of th1 of
- _ $ (Const(@{const_name "op ="},_) $ f $ g) => (Term.lambda v f,Term.lambda v g)
+ _ $ (Const(@{const_name HOL.eq},_) $ f $ g) => (Term.lambda v f,Term.lambda v g)
| _ => raise ERR "mk_ABS" "Bad conclusion"
val (fd,fr) = dom_rng (type_of f)
val abs_thm' = Drule.instantiate' [SOME (ctyp_of thy fd), SOME (ctyp_of thy fr)] [SOME (cterm_of thy f), SOME (cterm_of thy g)] abs_thm
@@ -1860,7 +1860,7 @@
val _ = if_debug pth hth
val th1 = implies_elim_all (beta_eta_thm th)
val a = case concl_of th1 of
- _ $ (Const(@{const_name "op -->"},_) $ a $ Const(@{const_name False},_)) => a
+ _ $ (Const(@{const_name HOL.implies},_) $ a $ Const(@{const_name False},_)) => a
| _ => raise ERR "NOT_INTRO" "Conclusion of bad form"
val ca = cterm_of thy a
val th2 = Thm.equal_elim (Drule.instantiate' [] [SOME ca] not_intro_thm) th1
--- a/src/HOL/Import/shuffler.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Import/shuffler.ML Wed Sep 01 17:19:47 2010 +0200
@@ -60,14 +60,14 @@
fun mk_meta_eq th =
(case concl_of th of
- Const(@{const_name Trueprop},_) $ (Const(@{const_name "op ="},_) $ _ $ _) => th RS eq_reflection
+ Const(@{const_name Trueprop},_) $ (Const(@{const_name HOL.eq},_) $ _ $ _) => th RS eq_reflection
| Const("==",_) $ _ $ _ => th
| _ => raise THM("Not an equality",0,[th]))
handle _ => raise THM("Couldn't make meta equality",0,[th]) (* FIXME avoid handle _ *)
fun mk_obj_eq th =
(case concl_of th of
- Const(@{const_name Trueprop},_) $ (Const(@{const_name "op ="},_) $ _ $ _) => th
+ Const(@{const_name Trueprop},_) $ (Const(@{const_name HOL.eq},_) $ _ $ _) => th
| Const("==",_) $ _ $ _ => th RS meta_eq_to_obj_eq
| _ => raise THM("Not an equality",0,[th]))
handle _ => raise THM("Couldn't make object equality",0,[th]) (* FIXME avoid handle _ *)
--- a/src/HOL/Int.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Int.thy Wed Sep 01 17:19:47 2010 +0200
@@ -2222,42 +2222,42 @@
mult_bin_simps
arith_extra_simps(4) [where 'a = int]
-instantiation int :: eq
+instantiation int :: equal
begin
definition
- "eq_class.eq k l \<longleftrightarrow> k - l = (0\<Colon>int)"
-
-instance by default (simp add: eq_int_def)
+ "HOL.equal k l \<longleftrightarrow> k - l = (0\<Colon>int)"
+
+instance by default (simp add: equal_int_def)
end
lemma eq_number_of_int_code [code]:
- "eq_class.eq (number_of k \<Colon> int) (number_of l) \<longleftrightarrow> eq_class.eq k l"
- unfolding eq_int_def number_of_is_id ..
+ "HOL.equal (number_of k \<Colon> int) (number_of l) \<longleftrightarrow> HOL.equal k l"
+ unfolding equal_int_def number_of_is_id ..
lemma eq_int_code [code]:
- "eq_class.eq Int.Pls Int.Pls \<longleftrightarrow> True"
- "eq_class.eq Int.Pls Int.Min \<longleftrightarrow> False"
- "eq_class.eq Int.Pls (Int.Bit0 k2) \<longleftrightarrow> eq_class.eq Int.Pls k2"
- "eq_class.eq Int.Pls (Int.Bit1 k2) \<longleftrightarrow> False"
- "eq_class.eq Int.Min Int.Pls \<longleftrightarrow> False"
- "eq_class.eq Int.Min Int.Min \<longleftrightarrow> True"
- "eq_class.eq Int.Min (Int.Bit0 k2) \<longleftrightarrow> False"
- "eq_class.eq Int.Min (Int.Bit1 k2) \<longleftrightarrow> eq_class.eq Int.Min k2"
- "eq_class.eq (Int.Bit0 k1) Int.Pls \<longleftrightarrow> eq_class.eq k1 Int.Pls"
- "eq_class.eq (Int.Bit1 k1) Int.Pls \<longleftrightarrow> False"
- "eq_class.eq (Int.Bit0 k1) Int.Min \<longleftrightarrow> False"
- "eq_class.eq (Int.Bit1 k1) Int.Min \<longleftrightarrow> eq_class.eq k1 Int.Min"
- "eq_class.eq (Int.Bit0 k1) (Int.Bit0 k2) \<longleftrightarrow> eq_class.eq k1 k2"
- "eq_class.eq (Int.Bit0 k1) (Int.Bit1 k2) \<longleftrightarrow> False"
- "eq_class.eq (Int.Bit1 k1) (Int.Bit0 k2) \<longleftrightarrow> False"
- "eq_class.eq (Int.Bit1 k1) (Int.Bit1 k2) \<longleftrightarrow> eq_class.eq k1 k2"
- unfolding eq_equals by simp_all
+ "HOL.equal Int.Pls Int.Pls \<longleftrightarrow> True"
+ "HOL.equal Int.Pls Int.Min \<longleftrightarrow> False"
+ "HOL.equal Int.Pls (Int.Bit0 k2) \<longleftrightarrow> HOL.equal Int.Pls k2"
+ "HOL.equal Int.Pls (Int.Bit1 k2) \<longleftrightarrow> False"
+ "HOL.equal Int.Min Int.Pls \<longleftrightarrow> False"
+ "HOL.equal Int.Min Int.Min \<longleftrightarrow> True"
+ "HOL.equal Int.Min (Int.Bit0 k2) \<longleftrightarrow> False"
+ "HOL.equal Int.Min (Int.Bit1 k2) \<longleftrightarrow> HOL.equal Int.Min k2"
+ "HOL.equal (Int.Bit0 k1) Int.Pls \<longleftrightarrow> HOL.equal k1 Int.Pls"
+ "HOL.equal (Int.Bit1 k1) Int.Pls \<longleftrightarrow> False"
+ "HOL.equal (Int.Bit0 k1) Int.Min \<longleftrightarrow> False"
+ "HOL.equal (Int.Bit1 k1) Int.Min \<longleftrightarrow> HOL.equal k1 Int.Min"
+ "HOL.equal (Int.Bit0 k1) (Int.Bit0 k2) \<longleftrightarrow> HOL.equal k1 k2"
+ "HOL.equal (Int.Bit0 k1) (Int.Bit1 k2) \<longleftrightarrow> False"
+ "HOL.equal (Int.Bit1 k1) (Int.Bit0 k2) \<longleftrightarrow> False"
+ "HOL.equal (Int.Bit1 k1) (Int.Bit1 k2) \<longleftrightarrow> HOL.equal k1 k2"
+ unfolding equal_eq by simp_all
lemma eq_int_refl [code nbe]:
- "eq_class.eq (k::int) k \<longleftrightarrow> True"
- by (rule HOL.eq_refl)
+ "HOL.equal (k::int) k \<longleftrightarrow> True"
+ by (rule equal_refl)
lemma less_eq_number_of_int_code [code]:
"(number_of k \<Colon> int) \<le> number_of l \<longleftrightarrow> k \<le> l"
--- a/src/HOL/IsaMakefile Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/IsaMakefile Wed Sep 01 17:19:47 2010 +0200
@@ -110,6 +110,7 @@
$(SRC)/Tools/Code/code_eval.ML \
$(SRC)/Tools/Code/code_haskell.ML \
$(SRC)/Tools/Code/code_ml.ML \
+ $(SRC)/Tools/Code/code_namespace.ML \
$(SRC)/Tools/Code/code_preproc.ML \
$(SRC)/Tools/Code/code_printer.ML \
$(SRC)/Tools/Code/code_scala.ML \
@@ -213,6 +214,7 @@
Tools/sat_funcs.ML \
Tools/sat_solver.ML \
Tools/split_rule.ML \
+ Tools/try.ML \
Tools/typedef.ML \
Transitive_Closure.thy \
Typedef.thy \
@@ -1321,7 +1323,9 @@
Predicate_Compile_Examples/ROOT.ML \
Predicate_Compile_Examples/Predicate_Compile_Examples.thy \
Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy \
- Predicate_Compile_Examples/Code_Prolog_Examples.thy
+ Predicate_Compile_Examples/Code_Prolog_Examples.thy \
+ Predicate_Compile_Examples/Hotel_Example.thy \
+ Predicate_Compile_Examples/Lambda_Example.thy
@$(ISABELLE_TOOL) usedir $(OUT)/HOL Predicate_Compile_Examples
--- a/src/HOL/Lazy_Sequence.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Lazy_Sequence.thy Wed Sep 01 17:19:47 2010 +0200
@@ -39,10 +39,14 @@
"size xq = (case yield xq of None => 0 | Some (x, xq') => size xq' + 1)"
by (cases xq) auto
-lemma [code]: "eq_class.eq xq yq = (case (yield xq, yield yq) of
- (None, None) => True | (Some (x, xq'), Some (y, yq')) => (HOL.eq x y) \<and> (eq_class.eq xq yq) | _ => False)"
-apply (cases xq) apply (cases yq) apply (auto simp add: eq_class.eq_equals)
-apply (cases yq) apply (auto simp add: eq_class.eq_equals) done
+lemma [code]: "HOL.equal xq yq = (case (yield xq, yield yq) of
+ (None, None) => True | (Some (x, xq'), Some (y, yq')) => (HOL.equal x y) \<and> (HOL.equal xq yq) | _ => False)"
+apply (cases xq) apply (cases yq) apply (auto simp add: equal_eq)
+apply (cases yq) apply (auto simp add: equal_eq) done
+
+lemma [code nbe]:
+ "HOL.equal (x :: 'a lazy_sequence) x \<longleftrightarrow> True"
+ by (fact equal_refl)
lemma seq_case [code]:
"lazy_sequence_case f g xq = (case (yield xq) of None => f | Some (x, xq') => g x xq')"
--- a/src/HOL/Library/AssocList.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/AssocList.thy Wed Sep 01 17:19:47 2010 +0200
@@ -701,7 +701,44 @@
"Mapping.bulkload vs = Mapping (map (\<lambda>n. (n, vs ! n)) [0..<length vs])"
by (rule mapping_eqI) (simp add: map_of_map_restrict expand_fun_eq)
-lemma [code, code del]:
- "HOL.eq (x :: (_, _) mapping) y \<longleftrightarrow> x = y" by (fact eq_equals) (*FIXME*)
+lemma map_of_eqI: (*FIXME move to Map.thy*)
+ assumes set_eq: "set (map fst xs) = set (map fst ys)"
+ assumes map_eq: "\<forall>k\<in>set (map fst xs). map_of xs k = map_of ys k"
+ shows "map_of xs = map_of ys"
+proof (rule ext)
+ fix k show "map_of xs k = map_of ys k"
+ proof (cases "map_of xs k")
+ case None then have "k \<notin> set (map fst xs)" by (simp add: map_of_eq_None_iff)
+ with set_eq have "k \<notin> set (map fst ys)" by simp
+ then have "map_of ys k = None" by (simp add: map_of_eq_None_iff)
+ with None show ?thesis by simp
+ next
+ case (Some v) then have "k \<in> set (map fst xs)" by (auto simp add: dom_map_of_conv_image_fst [symmetric])
+ with map_eq show ?thesis by auto
+ qed
+qed
+
+lemma map_of_eq_dom: (*FIXME move to Map.thy*)
+ assumes "map_of xs = map_of ys"
+ shows "fst ` set xs = fst ` set ys"
+proof -
+ from assms have "dom (map_of xs) = dom (map_of ys)" by simp
+ then show ?thesis by (simp add: dom_map_of_conv_image_fst)
+qed
+
+lemma equal_Mapping [code]:
+ "HOL.equal (Mapping xs) (Mapping ys) \<longleftrightarrow>
+ (let ks = map fst xs; ls = map fst ys
+ in (\<forall>l\<in>set ls. l \<in> set ks) \<and> (\<forall>k\<in>set ks. k \<in> set ls \<and> map_of xs k = map_of ys k))"
+proof -
+ have aux: "\<And>a b xs. (a, b) \<in> set xs \<Longrightarrow> a \<in> fst ` set xs" by (auto simp add: image_def intro!: bexI)
+ show ?thesis
+ by (auto intro!: map_of_eqI simp add: Let_def equal Mapping_def)
+ (auto dest!: map_of_eq_dom intro: aux)
+qed
+
+lemma [code nbe]:
+ "HOL.equal (x :: ('a, 'b) mapping) x \<longleftrightarrow> True"
+ by (fact equal_refl)
end
--- a/src/HOL/Library/Code_Char.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Code_Char.thy Wed Sep 01 17:19:47 2010 +0200
@@ -19,7 +19,7 @@
#> String_Code.add_literal_list_string "Haskell"
*}
-code_instance char :: eq
+code_instance char :: equal
(Haskell -)
code_reserved SML
@@ -31,7 +31,7 @@
code_reserved Scala
char
-code_const "eq_class.eq \<Colon> char \<Rightarrow> char \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> char \<Rightarrow> char \<Rightarrow> bool"
(SML "!((_ : char) = _)")
(OCaml "!((_ : char) = _)")
(Haskell infixl 4 "==")
--- a/src/HOL/Library/Code_Integer.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Code_Integer.thy Wed Sep 01 17:19:47 2010 +0200
@@ -21,7 +21,7 @@
(Scala "BigInt")
(Eval "int")
-code_instance int :: eq
+code_instance int :: equal
(Haskell -)
setup {*
@@ -51,14 +51,14 @@
(SML "IntInf.- ((_), 1)")
(OCaml "Big'_int.pred'_big'_int")
(Haskell "!(_/ -/ 1)")
- (Scala "!(_/ -/ 1)")
+ (Scala "!(_ -/ 1)")
(Eval "!(_/ -/ 1)")
code_const Int.succ
(SML "IntInf.+ ((_), 1)")
(OCaml "Big'_int.succ'_big'_int")
(Haskell "!(_/ +/ 1)")
- (Scala "!(_/ +/ 1)")
+ (Scala "!(_ +/ 1)")
(Eval "!(_/ +/ 1)")
code_const "op + \<Colon> int \<Rightarrow> int \<Rightarrow> int"
@@ -96,7 +96,7 @@
(Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))")
(Eval "Integer.div'_mod/ (abs _)/ (abs _)")
-code_const "eq_class.eq \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> int \<Rightarrow> int \<Rightarrow> bool"
(SML "!((_ : IntInf.int) = _)")
(OCaml "Big'_int.eq'_big'_int")
(Haskell infixl 4 "==")
--- a/src/HOL/Library/Code_Natural.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Code_Natural.thy Wed Sep 01 17:19:47 2010 +0200
@@ -52,12 +52,11 @@
| otherwise = (Natural k, Natural l) where (k, l) = quotRem n m;
};*}
+
code_reserved Haskell Natural
-code_include Scala "Natural" {*
-import scala.Math
-
-object Natural {
+code_include Scala "Natural"
+{*object Natural {
def apply(numeral: BigInt): Natural = new Natural(numeral max 0)
def apply(numeral: Int): Natural = Natural(BigInt(numeral))
@@ -111,7 +110,7 @@
false Code_Printer.literal_alternative_numeral) ["Haskell", "Scala"]
*}
-code_instance code_numeral :: eq
+code_instance code_numeral :: equal
(Haskell -)
code_const "op + \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
@@ -130,7 +129,7 @@
(Haskell "divMod")
(Scala infixl 8 "/%")
-code_const "eq_class.eq \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
(Haskell infixl 4 "==")
(Scala infixl 5 "==")
--- a/src/HOL/Library/Code_Prolog.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Code_Prolog.thy Wed Sep 01 17:19:47 2010 +0200
@@ -9,5 +9,10 @@
uses "~~/src/HOL/Tools/Predicate_Compile/code_prolog.ML"
begin
+section {* Setup for Numerals *}
+
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name number_of}] *}
+setup {* Predicate_Compile_Data.keep_functions [@{const_name number_of}] *}
+
end
--- a/src/HOL/Library/Dlist.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Dlist.thy Wed Sep 01 17:19:47 2010 +0200
@@ -109,16 +109,20 @@
text {* Equality *}
-instantiation dlist :: (eq) eq
+instantiation dlist :: (equal) equal
begin
-definition "HOL.eq dxs dys \<longleftrightarrow> HOL.eq (list_of_dlist dxs) (list_of_dlist dys)"
+definition "HOL.equal dxs dys \<longleftrightarrow> HOL.equal (list_of_dlist dxs) (list_of_dlist dys)"
instance proof
-qed (simp add: eq_dlist_def eq list_of_dlist_inject)
+qed (simp add: equal_dlist_def equal list_of_dlist_inject)
end
+lemma [code nbe]:
+ "HOL.equal (dxs :: 'a::equal dlist) dxs \<longleftrightarrow> True"
+ by (fact equal_refl)
+
section {* Induction principle and case distinction *}
--- a/src/HOL/Library/Efficient_Nat.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Efficient_Nat.thy Wed Sep 01 17:19:47 2010 +0200
@@ -55,12 +55,12 @@
by (simp add: prod_fun_def split_def pdivmod_def nat_div_distrib nat_mod_distrib divmod_nat_div_mod)
lemma eq_nat_code [code]:
- "eq_class.eq n m \<longleftrightarrow> eq_class.eq (of_nat n \<Colon> int) (of_nat m)"
- by (simp add: eq)
+ "HOL.equal n m \<longleftrightarrow> HOL.equal (of_nat n \<Colon> int) (of_nat m)"
+ by (simp add: equal)
lemma eq_nat_refl [code nbe]:
- "eq_class.eq (n::nat) n \<longleftrightarrow> True"
- by (rule HOL.eq_refl)
+ "HOL.equal (n::nat) n \<longleftrightarrow> True"
+ by (rule equal_refl)
lemma less_eq_nat_code [code]:
"n \<le> m \<longleftrightarrow> (of_nat n \<Colon> int) \<le> of_nat m"
@@ -242,8 +242,8 @@
and @{typ int}.
*}
-code_include Haskell "Nat" {*
-newtype Nat = Nat Integer deriving (Eq, Show, Read);
+code_include Haskell "Nat"
+{*newtype Nat = Nat Integer deriving (Eq, Show, Read);
instance Num Nat where {
fromInteger k = Nat (if k >= 0 then k else 0);
@@ -280,10 +280,8 @@
code_reserved Haskell Nat
-code_include Scala "Nat" {*
-import scala.Math
-
-object Nat {
+code_include Scala "Nat"
+{*object Nat {
def apply(numeral: BigInt): Nat = new Nat(numeral max 0)
def apply(numeral: Int): Nat = Nat(BigInt(numeral))
@@ -332,7 +330,7 @@
(Haskell "Nat.Nat")
(Scala "Nat")
-code_instance nat :: eq
+code_instance nat :: equal
(Haskell -)
text {*
@@ -442,7 +440,7 @@
(Scala infixl 8 "/%")
(Eval "Integer.div'_mod")
-code_const "eq_class.eq \<Colon> nat \<Rightarrow> nat \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> nat \<Rightarrow> nat \<Rightarrow> bool"
(SML "!((_ : IntInf.int) = _)")
(OCaml "Big'_int.eq'_big'_int")
(Haskell infixl 4 "==")
--- a/src/HOL/Library/Enum.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Enum.thy Wed Sep 01 17:19:47 2010 +0200
@@ -35,17 +35,21 @@
subsection {* Equality and order on functions *}
-instantiation "fun" :: (enum, eq) eq
+instantiation "fun" :: (enum, equal) equal
begin
definition
- "eq_class.eq f g \<longleftrightarrow> (\<forall>x \<in> set enum. f x = g x)"
+ "HOL.equal f g \<longleftrightarrow> (\<forall>x \<in> set enum. f x = g x)"
instance proof
-qed (simp_all add: eq_fun_def enum_all expand_fun_eq)
+qed (simp_all add: equal_fun_def enum_all expand_fun_eq)
end
+lemma [code nbe]:
+ "HOL.equal (f :: _ \<Rightarrow> _) f \<longleftrightarrow> True"
+ by (fact equal_refl)
+
lemma order_fun [code]:
fixes f g :: "'a\<Colon>enum \<Rightarrow> 'b\<Colon>order"
shows "f \<le> g \<longleftrightarrow> list_all (\<lambda>x. f x \<le> g x) enum"
@@ -169,7 +173,7 @@
end
-lemma enum_fun_code [code]: "enum = (let enum_a = (enum \<Colon> 'a\<Colon>{enum, eq} list)
+lemma enum_fun_code [code]: "enum = (let enum_a = (enum \<Colon> 'a\<Colon>{enum, equal} list)
in map (\<lambda>ys. the o map_of (zip enum_a ys)) (n_lists (length enum_a) enum))"
by (simp add: enum_fun_def Let_def)
--- a/src/HOL/Library/Fset.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Fset.thy Wed Sep 01 17:19:47 2010 +0200
@@ -227,17 +227,21 @@
"A < B \<longleftrightarrow> A \<le> B \<and> \<not> B \<le> (A :: 'a fset)"
by (fact less_le_not_le)
-instantiation fset :: (type) eq
+instantiation fset :: (type) equal
begin
definition
- "eq_fset A B \<longleftrightarrow> A \<le> B \<and> B \<le> (A :: 'a fset)"
+ "HOL.equal A B \<longleftrightarrow> A \<le> B \<and> B \<le> (A :: 'a fset)"
instance proof
-qed (simp add: eq_fset_def set_eq [symmetric])
+qed (simp add: equal_fset_def set_eq [symmetric])
end
+lemma [code nbe]:
+ "HOL.equal (A :: 'a fset) A \<longleftrightarrow> True"
+ by (fact equal_refl)
+
subsection {* Functorial operations *}
--- a/src/HOL/Library/List_Prefix.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/List_Prefix.thy Wed Sep 01 17:19:47 2010 +0200
@@ -81,9 +81,9 @@
by (auto simp add: prefix_def)
lemma less_eq_list_code [code]:
- "([]\<Colon>'a\<Colon>{eq, ord} list) \<le> xs \<longleftrightarrow> True"
- "(x\<Colon>'a\<Colon>{eq, ord}) # xs \<le> [] \<longleftrightarrow> False"
- "(x\<Colon>'a\<Colon>{eq, ord}) # xs \<le> y # ys \<longleftrightarrow> x = y \<and> xs \<le> ys"
+ "([]\<Colon>'a\<Colon>{equal, ord} list) \<le> xs \<longleftrightarrow> True"
+ "(x\<Colon>'a\<Colon>{equal, ord}) # xs \<le> [] \<longleftrightarrow> False"
+ "(x\<Colon>'a\<Colon>{equal, ord}) # xs \<le> y # ys \<longleftrightarrow> x = y \<and> xs \<le> ys"
by simp_all
lemma same_prefix_prefix [simp]: "(xs @ ys \<le> xs @ zs) = (ys \<le> zs)"
--- a/src/HOL/Library/List_lexord.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/List_lexord.thy Wed Sep 01 17:19:47 2010 +0200
@@ -103,15 +103,15 @@
end
lemma less_list_code [code]:
- "xs < ([]\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> False"
- "[] < (x\<Colon>'a\<Colon>{eq, order}) # xs \<longleftrightarrow> True"
- "(x\<Colon>'a\<Colon>{eq, order}) # xs < y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs < ys"
+ "xs < ([]\<Colon>'a\<Colon>{equal, order} list) \<longleftrightarrow> False"
+ "[] < (x\<Colon>'a\<Colon>{equal, order}) # xs \<longleftrightarrow> True"
+ "(x\<Colon>'a\<Colon>{equal, order}) # xs < y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs < ys"
by simp_all
lemma less_eq_list_code [code]:
- "x # xs \<le> ([]\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> False"
- "[] \<le> (xs\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> True"
- "(x\<Colon>'a\<Colon>{eq, order}) # xs \<le> y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs \<le> ys"
+ "x # xs \<le> ([]\<Colon>'a\<Colon>{equal, order} list) \<longleftrightarrow> False"
+ "[] \<le> (xs\<Colon>'a\<Colon>{equal, order} list) \<longleftrightarrow> True"
+ "(x\<Colon>'a\<Colon>{equal, order}) # xs \<le> y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs \<le> ys"
by simp_all
end
--- a/src/HOL/Library/Mapping.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Mapping.thy Wed Sep 01 17:19:47 2010 +0200
@@ -280,14 +280,14 @@
code_datatype empty update
-instantiation mapping :: (type, type) eq
+instantiation mapping :: (type, type) equal
begin
definition [code del]:
- "HOL.eq m n \<longleftrightarrow> lookup m = lookup n"
+ "HOL.equal m n \<longleftrightarrow> lookup m = lookup n"
instance proof
-qed (simp add: eq_mapping_def)
+qed (simp add: equal_mapping_def)
end
--- a/src/HOL/Library/Multiset.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Multiset.thy Wed Sep 01 17:19:47 2010 +0200
@@ -938,17 +938,21 @@
qed
qed
-instantiation multiset :: (eq) eq
+instantiation multiset :: (equal) equal
begin
definition
- "HOL.eq A B \<longleftrightarrow> (A::'a multiset) \<le> B \<and> B \<le> A"
+ "HOL.equal A B \<longleftrightarrow> (A::'a multiset) \<le> B \<and> B \<le> A"
instance proof
-qed (simp add: eq_multiset_def eq_iff)
+qed (simp add: equal_multiset_def eq_iff)
end
+lemma [code nbe]:
+ "HOL.equal (A :: 'a::equal multiset) A \<longleftrightarrow> True"
+ by (fact equal_refl)
+
definition (in term_syntax)
bagify :: "('a\<Colon>typerep \<times> nat) list \<times> (unit \<Rightarrow> Code_Evaluation.term)
\<Rightarrow> 'a multiset \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
--- a/src/HOL/Library/Nested_Environment.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Nested_Environment.thy Wed Sep 01 17:19:47 2010 +0200
@@ -521,22 +521,21 @@
text {* Environments and code generation *}
lemma [code, code del]:
- fixes e1 e2 :: "('b\<Colon>eq, 'a\<Colon>eq, 'c\<Colon>eq) env"
- shows "eq_class.eq e1 e2 \<longleftrightarrow> eq_class.eq e1 e2" ..
+ "(HOL.equal :: (_, _, _) env \<Rightarrow> _) = HOL.equal" ..
-lemma eq_env_code [code]:
- fixes x y :: "'a\<Colon>eq"
- and f g :: "'c\<Colon>{eq, finite} \<Rightarrow> ('b\<Colon>eq, 'a, 'c) env option"
- shows "eq_class.eq (Env x f) (Env y g) \<longleftrightarrow>
- eq_class.eq x y \<and> (\<forall>z\<in>UNIV. case f z
+lemma equal_env_code [code]:
+ fixes x y :: "'a\<Colon>equal"
+ and f g :: "'c\<Colon>{equal, finite} \<Rightarrow> ('b\<Colon>equal, 'a, 'c) env option"
+ shows "HOL.equal (Env x f) (Env y g) \<longleftrightarrow>
+ HOL.equal x y \<and> (\<forall>z\<in>UNIV. case f z
of None \<Rightarrow> (case g z
of None \<Rightarrow> True | Some _ \<Rightarrow> False)
| Some a \<Rightarrow> (case g z
- of None \<Rightarrow> False | Some b \<Rightarrow> eq_class.eq a b))" (is ?env)
- and "eq_class.eq (Val a) (Val b) \<longleftrightarrow> eq_class.eq a b"
- and "eq_class.eq (Val a) (Env y g) \<longleftrightarrow> False"
- and "eq_class.eq (Env x f) (Val b) \<longleftrightarrow> False"
-proof (unfold eq)
+ of None \<Rightarrow> False | Some b \<Rightarrow> HOL.equal a b))" (is ?env)
+ and "HOL.equal (Val a) (Val b) \<longleftrightarrow> HOL.equal a b"
+ and "HOL.equal (Val a) (Env y g) \<longleftrightarrow> False"
+ and "HOL.equal (Env x f) (Val b) \<longleftrightarrow> False"
+proof (unfold equal)
have "f = g \<longleftrightarrow> (\<forall>z. case f z
of None \<Rightarrow> (case g z
of None \<Rightarrow> True | Some _ \<Rightarrow> False)
@@ -562,6 +561,10 @@
of None \<Rightarrow> False | Some b \<Rightarrow> a = b))" by simp
qed simp_all
+lemma [code nbe]:
+ "HOL.equal (x :: (_, _, _) env) x \<longleftrightarrow> True"
+ by (fact equal_refl)
+
lemma [code, code del]:
"(Code_Evaluation.term_of :: ('a::{term_of, type}, 'b::{term_of, type}, 'c::{term_of, type}) env \<Rightarrow> term) = Code_Evaluation.term_of" ..
--- a/src/HOL/Library/OptionalSugar.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/OptionalSugar.thy Wed Sep 01 17:19:47 2010 +0200
@@ -32,7 +32,7 @@
(* deprecated, use thm with style instead, will be removed *)
(* aligning equations *)
notation (tab output)
- "op =" ("(_) \<^raw:}\putisatab\isa{\ >=\<^raw:}\putisatab\isa{> (_)" [50,49] 50) and
+ "HOL.eq" ("(_) \<^raw:}\putisatab\isa{\ >=\<^raw:}\putisatab\isa{> (_)" [50,49] 50) and
"==" ("(_) \<^raw:}\putisatab\isa{\ >\<equiv>\<^raw:}\putisatab\isa{> (_)")
(* Let *)
--- a/src/HOL/Library/Polynomial.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Polynomial.thy Wed Sep 01 17:19:47 2010 +0200
@@ -1505,23 +1505,27 @@
declare pCons_0_0 [code_post]
-instantiation poly :: ("{zero,eq}") eq
+instantiation poly :: ("{zero, equal}") equal
begin
definition
- "eq_class.eq (p::'a poly) q \<longleftrightarrow> p = q"
+ "HOL.equal (p::'a poly) q \<longleftrightarrow> p = q"
-instance
- by default (rule eq_poly_def)
+instance proof
+qed (rule equal_poly_def)
end
lemma eq_poly_code [code]:
- "eq_class.eq (0::_ poly) (0::_ poly) \<longleftrightarrow> True"
- "eq_class.eq (0::_ poly) (pCons b q) \<longleftrightarrow> eq_class.eq 0 b \<and> eq_class.eq 0 q"
- "eq_class.eq (pCons a p) (0::_ poly) \<longleftrightarrow> eq_class.eq a 0 \<and> eq_class.eq p 0"
- "eq_class.eq (pCons a p) (pCons b q) \<longleftrightarrow> eq_class.eq a b \<and> eq_class.eq p q"
-unfolding eq by simp_all
+ "HOL.equal (0::_ poly) (0::_ poly) \<longleftrightarrow> True"
+ "HOL.equal (0::_ poly) (pCons b q) \<longleftrightarrow> HOL.equal 0 b \<and> HOL.equal 0 q"
+ "HOL.equal (pCons a p) (0::_ poly) \<longleftrightarrow> HOL.equal a 0 \<and> HOL.equal p 0"
+ "HOL.equal (pCons a p) (pCons b q) \<longleftrightarrow> HOL.equal a b \<and> HOL.equal p q"
+ by (simp_all add: equal)
+
+lemma [code nbe]:
+ "HOL.equal (p :: _ poly) p \<longleftrightarrow> True"
+ by (fact equal_refl)
lemmas coeff_code [code] =
coeff_0 coeff_pCons_0 coeff_pCons_Suc
--- a/src/HOL/Library/Product_ord.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Product_ord.thy Wed Sep 01 17:19:47 2010 +0200
@@ -22,8 +22,8 @@
end
lemma [code]:
- "(x1\<Colon>'a\<Colon>{ord, eq}, y1) \<le> (x2, y2) \<longleftrightarrow> x1 < x2 \<or> x1 \<le> x2 \<and> y1 \<le> y2"
- "(x1\<Colon>'a\<Colon>{ord, eq}, y1) < (x2, y2) \<longleftrightarrow> x1 < x2 \<or> x1 \<le> x2 \<and> y1 < y2"
+ "(x1\<Colon>'a\<Colon>{ord, equal}, y1) \<le> (x2, y2) \<longleftrightarrow> x1 < x2 \<or> x1 \<le> x2 \<and> y1 \<le> y2"
+ "(x1\<Colon>'a\<Colon>{ord, equal}, y1) < (x2, y2) \<longleftrightarrow> x1 < x2 \<or> x1 \<le> x2 \<and> y1 < y2"
unfolding prod_le_def prod_less_def by simp_all
instance prod :: (preorder, preorder) preorder proof
--- a/src/HOL/Library/RBT.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/RBT.thy Wed Sep 01 17:19:47 2010 +0200
@@ -222,12 +222,14 @@
"Mapping.bulkload vs = Mapping (bulkload (List.map (\<lambda>n. (n, vs ! n)) [0..<length vs]))"
by (rule mapping_eqI) (simp add: map_of_map_restrict expand_fun_eq)
-lemma [code, code del]:
- "HOL.eq (x :: (_, _) mapping) y \<longleftrightarrow> x = y" by (fact eq_equals) (*FIXME*)
+lemma equal_Mapping [code]:
+ "HOL.equal (Mapping t1) (Mapping t2) \<longleftrightarrow> entries t1 = entries t2"
+ by (simp add: equal Mapping_def entries_lookup)
-lemma eq_Mapping [code]:
- "HOL.eq (Mapping t1) (Mapping t2) \<longleftrightarrow> entries t1 = entries t2"
- by (simp add: eq Mapping_def entries_lookup)
+lemma [code nbe]:
+ "HOL.equal (x :: (_, _) mapping) x \<longleftrightarrow> True"
+ by (fact equal_refl)
+
hide_const (open) impl_of lookup empty insert delete
entries keys bulkload map_entry map fold
--- a/src/HOL/Library/Sum_Of_Squares.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Sum_Of_Squares.thy Wed Sep 01 17:19:47 2010 +0200
@@ -28,6 +28,7 @@
without calling an external prover.
*}
+setup Sum_Of_Squares.setup
setup SOS_Wrapper.setup
text {* Tests *}
--- a/src/HOL/Library/Sum_Of_Squares/sos_wrapper.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Sum_Of_Squares/sos_wrapper.ML Wed Sep 01 17:19:47 2010 +0200
@@ -8,8 +8,8 @@
sig
datatype prover_result = Success | Failure | Error
val setup: theory -> theory
- val destdir: string Unsynchronized.ref
- val prover_name: string Unsynchronized.ref
+ val dest_dir: string Config.T
+ val prover_name: string Config.T
end
structure SOS_Wrapper: SOS_WRAPPER =
@@ -22,14 +22,9 @@
| str_of_result Error = "Error"
-(*** output control ***)
-
-fun debug s = if ! Sum_Of_Squares.debugging then writeln s else ()
-
-
(*** calling provers ***)
-val destdir = Unsynchronized.ref ""
+val (dest_dir, setup_dest_dir) = Attrib.config_string "sos_dest_dir" (K "")
fun filename dir name =
let
@@ -43,12 +38,12 @@
else error ("No such directory: " ^ dir)
end
-fun run_solver name cmd find_failure input =
+fun run_solver ctxt name cmd find_failure input =
let
val _ = warning ("Calling solver: " ^ name)
(* create input file *)
- val dir = ! destdir
+ val dir = Config.get ctxt dest_dir
val input_file = filename dir "sos_in"
val _ = File.write input_file input
@@ -71,7 +66,10 @@
(File.rm input_file; if File.exists output_file then File.rm output_file else ())
else ()
- val _ = debug ("Solver output:\n" ^ output)
+ val _ =
+ if Config.get ctxt Sum_Of_Squares.trace
+ then writeln ("Solver output:\n" ^ output)
+ else ()
val _ = warning (str_of_result res ^ ": " ^ res_msg)
in
@@ -120,13 +118,13 @@
| prover "csdp" = csdp
| prover name = error ("Unknown prover: " ^ name)
-val prover_name = Unsynchronized.ref "remote_csdp"
+val (prover_name, setup_prover_name) = Attrib.config_string "sos_prover_name" (K "remote_csdp")
-fun call_solver opt_name =
+fun call_solver ctxt opt_name =
let
- val name = the_default (! prover_name) opt_name
+ val name = the_default (Config.get ctxt prover_name) opt_name
val (cmd, find_failure) = prover name
- in run_solver name (Path.explode cmd) find_failure end
+ in run_solver ctxt name (Path.explode cmd) find_failure end
(* certificate output *)
@@ -143,9 +141,13 @@
fun sos_solver print method = SIMPLE_METHOD' o Sum_Of_Squares.sos_tac print method
val setup =
+ setup_dest_dir #>
+ setup_prover_name #>
Method.setup @{binding sos}
(Scan.lift (Scan.option Parse.xname)
- >> (sos_solver print_cert o Sum_Of_Squares.Prover o call_solver))
+ >> (fn opt_name => fn ctxt =>
+ sos_solver print_cert
+ (Sum_Of_Squares.Prover (call_solver ctxt opt_name)) ctxt))
"prove universal problems over the reals using sums of squares" #>
Method.setup @{binding sos_cert}
(Scan.lift Parse.string
--- a/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML Wed Sep 01 17:19:47 2010 +0200
@@ -9,7 +9,8 @@
sig
datatype proof_method = Certificate of RealArith.pss_tree | Prover of string -> string
val sos_tac: (RealArith.pss_tree -> unit) -> proof_method -> Proof.context -> int -> tactic
- val debugging: bool Unsynchronized.ref
+ val trace: bool Config.T
+ val setup: theory -> theory
exception Failure of string;
end
@@ -49,7 +50,8 @@
val pow2 = rat_pow rat_2;
val pow10 = rat_pow rat_10;
-val debugging = Unsynchronized.ref false;
+val (trace, setup_trace) = Attrib.config_bool "sos_trace" (K false);
+val setup = setup_trace;
exception Sanity;
@@ -1059,7 +1061,7 @@
(* Positiv- and Nullstellensatz. Flag "linf" forces a linear representation. *)
-fun real_positivnullstellensatz_general prover linf d eqs leqs pol =
+fun real_positivnullstellensatz_general ctxt prover linf d eqs leqs pol =
let
val vars = fold_rev (union (op aconvc) o poly_variables)
(pol :: eqs @ map fst leqs) []
@@ -1129,7 +1131,7 @@
fun find_rounding d =
let
val _ =
- if !debugging
+ if Config.get ctxt trace
then writeln ("Trying rounding with limit "^Rat.string_of_rat d ^ "\n")
else ()
val vec = nice_vector d raw_vec
@@ -1245,7 +1247,7 @@
let val e = multidegree pol
val k = if e = 0 then 0 else d div e
val eq' = map fst eq
- in tryfind (fn i => (d,i,real_positivnullstellensatz_general prover false d eq' leq
+ in tryfind (fn i => (d,i,real_positivnullstellensatz_general ctxt prover false d eq' leq
(poly_neg(poly_pow pol i))))
(0 upto k)
end
@@ -1356,7 +1358,7 @@
val known_sos_constants =
[@{term "op ==>"}, @{term "Trueprop"},
- @{term "op -->"}, @{term "op &"}, @{term "op |"},
+ @{term HOL.implies}, @{term HOL.conj}, @{term HOL.disj},
@{term "Not"}, @{term "op = :: bool => _"},
@{term "All :: (real => _) => _"}, @{term "Ex :: (real => _) => _"},
@{term "op = :: real => _"}, @{term "op < :: real => _"},
--- a/src/HOL/Library/positivstellensatz.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/positivstellensatz.ML Wed Sep 01 17:19:47 2010 +0200
@@ -164,21 +164,6 @@
thm list * thm list * thm list -> thm * pss_tree
type cert_conv = cterm -> thm * pss_tree
-val my_eqs = Unsynchronized.ref ([] : thm list);
-val my_les = Unsynchronized.ref ([] : thm list);
-val my_lts = Unsynchronized.ref ([] : thm list);
-val my_proof = Unsynchronized.ref (Axiom_eq 0);
-val my_context = Unsynchronized.ref @{context};
-
-val my_mk_numeric = Unsynchronized.ref ((K @{cterm True}) :Rat.rat -> cterm);
-val my_numeric_eq_conv = Unsynchronized.ref no_conv;
-val my_numeric_ge_conv = Unsynchronized.ref no_conv;
-val my_numeric_gt_conv = Unsynchronized.ref no_conv;
-val my_poly_conv = Unsynchronized.ref no_conv;
-val my_poly_neg_conv = Unsynchronized.ref no_conv;
-val my_poly_add_conv = Unsynchronized.ref no_conv;
-val my_poly_mul_conv = Unsynchronized.ref no_conv;
-
(* Some useful derived rules *)
fun deduct_antisym_rule tha thb =
@@ -362,11 +347,6 @@
poly_conv,poly_neg_conv,poly_add_conv,poly_mul_conv,
absconv1,absconv2,prover) =
let
- val _ = my_context := ctxt
- val _ = (my_mk_numeric := mk_numeric ; my_numeric_eq_conv := numeric_eq_conv ;
- my_numeric_ge_conv := numeric_ge_conv; my_numeric_gt_conv := numeric_gt_conv ;
- my_poly_conv := poly_conv; my_poly_neg_conv := poly_neg_conv;
- my_poly_add_conv := poly_add_conv; my_poly_mul_conv := poly_mul_conv)
val pre_ss = HOL_basic_ss addsimps simp_thms@ ex_simps@ all_simps@[@{thm not_all},@{thm not_ex},ex_disj_distrib, all_conj_distrib, @{thm if_bool_eq_disj}]
val prenex_ss = HOL_basic_ss addsimps prenex_simps
val skolemize_ss = HOL_basic_ss addsimps [choice_iff]
@@ -408,7 +388,6 @@
fun hol_of_positivstellensatz(eqs,les,lts) proof =
let
- val _ = (my_eqs := eqs ; my_les := les ; my_lts := lts ; my_proof := proof)
fun translate prf = case prf of
Axiom_eq n => nth eqs n
| Axiom_le n => nth les n
@@ -439,8 +418,8 @@
val is_req = is_binop @{cterm "op =:: real => _"}
val is_ge = is_binop @{cterm "op <=:: real => _"}
val is_gt = is_binop @{cterm "op <:: real => _"}
- val is_conj = is_binop @{cterm "op &"}
- val is_disj = is_binop @{cterm "op |"}
+ val is_conj = is_binop @{cterm HOL.conj}
+ val is_disj = is_binop @{cterm HOL.disj}
fun conj_pair th = (th RS @{thm conjunct1}, th RS @{thm conjunct2})
fun disj_cases th th1 th2 =
let val (p,q) = Thm.dest_binop (concl th)
@@ -484,7 +463,7 @@
val th_p = poly_conv(Thm.dest_arg(Thm.dest_arg1(Thm.rhs_of th)))
val th_x = Drule.arg_cong_rule @{cterm "uminus :: real => _"} th_p
val th_n = fconv_rule (arg_conv poly_neg_conv) th_x
- val th' = Drule.binop_cong_rule @{cterm "op |"}
+ val th' = Drule.binop_cong_rule @{cterm HOL.disj}
(Drule.arg_cong_rule (Thm.capply @{cterm "op <::real=>_"} @{cterm "0::real"}) th_p)
(Drule.arg_cong_rule (Thm.capply @{cterm "op <::real=>_"} @{cterm "0::real"}) th_n)
in Thm.transitive th th'
@@ -542,12 +521,12 @@
let
val nnf_norm_conv' =
nnf_conv then_conv
- literals_conv [@{term "op &"}, @{term "op |"}] []
+ literals_conv [@{term HOL.conj}, @{term HOL.disj}] []
(Conv.cache_conv
(first_conv [real_lt_conv, real_le_conv,
real_eq_conv, real_not_lt_conv,
real_not_le_conv, real_not_eq_conv, all_conv]))
- fun absremover ct = (literals_conv [@{term "op &"}, @{term "op |"}] []
+ fun absremover ct = (literals_conv [@{term HOL.conj}, @{term HOL.disj}] []
(try_conv (absconv1 then_conv binop_conv (arg_conv poly_conv))) then_conv
try_conv (absconv2 then_conv nnf_norm_conv' then_conv binop_conv absremover)) ct
val nct = Thm.capply @{cterm Trueprop} (Thm.capply @{cterm "Not"} ct)
--- a/src/HOL/Library/reflection.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Library/reflection.ML Wed Sep 01 17:19:47 2010 +0200
@@ -82,7 +82,7 @@
fun rearrange congs =
let
fun P (_, th) =
- let val @{term "Trueprop"}$(Const (@{const_name "op ="},_) $l$_) = concl_of th
+ let val @{term "Trueprop"}$(Const (@{const_name HOL.eq},_) $l$_) = concl_of th
in can dest_Var l end
val (yes,no) = List.partition P congs
in no @ yes end
--- a/src/HOL/List.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/List.thy Wed Sep 01 17:19:47 2010 +0200
@@ -4721,8 +4721,8 @@
by (simp add: null_def)
lemma equal_Nil_null [code_unfold]:
- "eq_class.eq xs [] \<longleftrightarrow> null xs"
- by (simp add: eq eq_Nil_null)
+ "HOL.equal xs [] \<longleftrightarrow> null xs"
+ by (simp add: equal eq_Nil_null)
definition maps :: "('a \<Rightarrow> 'b list) \<Rightarrow> 'a list \<Rightarrow> 'b list" where
[code_post]: "maps f xs = concat (map f xs)"
@@ -4821,10 +4821,10 @@
(Haskell "[]")
(Scala "!Nil")
-code_instance list :: eq
+code_instance list :: equal
(Haskell -)
-code_const "eq_class.eq \<Colon> 'a\<Colon>eq list \<Rightarrow> 'a list \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
(Haskell infixl 4 "==")
code_reserved SML
--- a/src/HOL/Matrix/Compute_Oracle/compute.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Matrix/Compute_Oracle/compute.ML Wed Sep 01 17:19:47 2010 +0200
@@ -371,7 +371,7 @@
fun merge_shyps shyps1 shyps2 = Sorttab.keys (add_shyps shyps2 (add_shyps shyps1 Sorttab.empty))
val (_, export_oracle) = Context.>>> (Context.map_theory_result
- (Thm.add_oracle (Binding.name "compute", fn (thy, hyps, shyps, prop) =>
+ (Thm.add_oracle (@{binding compute}, fn (thy, hyps, shyps, prop) =>
let
val shyptab = add_shyps shyps Sorttab.empty
fun delete s shyptab = Sorttab.delete s shyptab handle Sorttab.UNDEF _ => shyptab
--- a/src/HOL/Mirabelle/Mirabelle_Test.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Mirabelle/Mirabelle_Test.thy Wed Sep 01 17:19:47 2010 +0200
@@ -12,6 +12,7 @@
"Tools/mirabelle_quickcheck.ML"
"Tools/mirabelle_refute.ML"
"Tools/mirabelle_sledgehammer.ML"
+ "Tools/mirabelle_sledgehammer_filter.ML"
begin
text {*
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Wed Sep 01 17:19:47 2010 +0200
@@ -290,10 +290,12 @@
| NONE => get_prover (default_atp_name ()))
end
+type locality = Sledgehammer_Fact_Filter.locality
+
local
datatype sh_result =
- SH_OK of int * int * (string * bool) list |
+ SH_OK of int * int * (string * locality) list |
SH_FAIL of int * int |
SH_ERROR
@@ -355,15 +357,16 @@
case result of
SH_OK (time_isa, time_atp, names) =>
let
- fun get_thms (name, chained) =
- ((name, chained), thms_of_name (Proof.context_of st) name)
+ fun get_thms (name, Sledgehammer_Fact_Filter.Chained) = NONE
+ | get_thms (name, loc) =
+ SOME ((name, loc), thms_of_name (Proof.context_of st) name)
in
change_data id inc_sh_success;
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);
- named_thms := SOME (map get_thms names);
+ named_thms := SOME (map_filter get_thms names);
log (sh_tag id ^ "succeeded (" ^ string_of_int time_isa ^ "+" ^
string_of_int time_atp ^ ") [" ^ prover_name ^ "]:\n" ^ msg)
end
@@ -445,7 +448,7 @@
then () else
let
val named_thms =
- Unsynchronized.ref (NONE : ((string * bool) * thm list) list option)
+ Unsynchronized.ref (NONE : ((string * locality) * thm list) list option)
val minimize = AList.defined (op =) args minimizeK
val metis_ft = AList.defined (op =) args metis_ftK
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML Wed Sep 01 17:19:47 2010 +0200
@@ -0,0 +1,169 @@
+(* Title: HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML
+ Author: Jasmin Blanchette, TU Munich
+*)
+
+structure Mirabelle_Sledgehammer_Filter : MIRABELLE_ACTION =
+struct
+
+val relevance_filter_args =
+ [("worse_irrel_freq", Sledgehammer_Fact_Filter.worse_irrel_freq),
+ ("higher_order_irrel_weight",
+ Sledgehammer_Fact_Filter.higher_order_irrel_weight),
+ ("abs_rel_weight", Sledgehammer_Fact_Filter.abs_rel_weight),
+ ("abs_irrel_weight", Sledgehammer_Fact_Filter.abs_irrel_weight),
+ ("skolem_irrel_weight", Sledgehammer_Fact_Filter.skolem_irrel_weight),
+ ("intro_bonus", Sledgehammer_Fact_Filter.intro_bonus),
+ ("elim_bonus", Sledgehammer_Fact_Filter.elim_bonus),
+ ("simp_bonus", Sledgehammer_Fact_Filter.simp_bonus),
+ ("local_bonus", Sledgehammer_Fact_Filter.local_bonus),
+ ("chained_bonus", Sledgehammer_Fact_Filter.chained_bonus),
+ ("max_imperfect", Sledgehammer_Fact_Filter.max_imperfect),
+ ("max_imperfect_exp", Sledgehammer_Fact_Filter.max_imperfect_exp),
+ ("threshold_divisor", Sledgehammer_Fact_Filter.threshold_divisor),
+ ("ridiculous_threshold", Sledgehammer_Fact_Filter.ridiculous_threshold)]
+
+structure Prooftab =
+ Table(type key = int * int val ord = prod_ord int_ord int_ord);
+
+val proof_table = Unsynchronized.ref Prooftab.empty
+
+val num_successes = Unsynchronized.ref ([] : (int * int) list)
+val num_failures = Unsynchronized.ref ([] : (int * int) list)
+val num_found_proofs = Unsynchronized.ref ([] : (int * int) list)
+val num_lost_proofs = Unsynchronized.ref ([] : (int * int) list)
+val num_found_facts = Unsynchronized.ref ([] : (int * int) list)
+val num_lost_facts = Unsynchronized.ref ([] : (int * int) list)
+
+fun get id c = the_default 0 (AList.lookup (op =) (!c) id)
+fun add id c n =
+ c := (case AList.lookup (op =) (!c) id of
+ SOME m => AList.update (op =) (id, m + n) (!c)
+ | NONE => (id, n) :: !c)
+
+fun init proof_file _ thy =
+ let
+ fun do_line line =
+ case line |> space_explode ":" of
+ [line_num, col_num, proof] =>
+ SOME (pairself (the o Int.fromString) (line_num, col_num),
+ proof |> space_explode " " |> filter_out (curry (op =) ""))
+ | _ => NONE
+ val proofs = File.read (Path.explode proof_file)
+ val proof_tab =
+ proofs |> space_explode "\n"
+ |> map_filter do_line
+ |> AList.coalesce (op =)
+ |> Prooftab.make
+ in proof_table := proof_tab; thy end
+
+fun percentage a b = if b = 0 then "N/A" else string_of_int (a * 100 div b)
+fun percentage_alt a b = percentage a (a + b)
+
+fun done id ({log, ...} : Mirabelle.done_args) =
+ if get id num_successes + get id num_failures > 0 then
+ (log "";
+ log ("Number of overall successes: " ^
+ string_of_int (get id num_successes));
+ log ("Number of overall failures: " ^ string_of_int (get id num_failures));
+ log ("Overall success rate: " ^
+ percentage_alt (get id num_successes) (get id num_failures) ^ "%");
+ log ("Number of found proofs: " ^ string_of_int (get id num_found_proofs));
+ log ("Number of lost proofs: " ^ string_of_int (get id num_lost_proofs));
+ log ("Proof found rate: " ^
+ percentage_alt (get id num_found_proofs) (get id num_lost_proofs) ^
+ "%");
+ log ("Number of found facts: " ^ string_of_int (get id num_found_facts));
+ log ("Number of lost facts: " ^ string_of_int (get id num_lost_facts));
+ log ("Fact found rate: " ^
+ percentage_alt (get id num_found_facts) (get id num_lost_facts) ^
+ "%"))
+ else
+ ()
+
+val default_max_relevant = 300
+
+fun with_index (i, s) = s ^ "@" ^ string_of_int i
+
+fun action args id ({pre, pos, log, ...} : Mirabelle.run_args) =
+ case (Position.line_of pos, Position.column_of pos) of
+ (SOME line_num, SOME col_num) =>
+ (case Prooftab.lookup (!proof_table) (line_num, col_num) of
+ 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 {relevance_thresholds, full_types, max_relevant, theory_relevant,
+ ...} = Sledgehammer_Isar.default_params thy args
+ val subgoal = 1
+ val (_, hyp_ts, concl_t) = Sledgehammer_Util.strip_subgoal goal subgoal
+ val facts =
+ Sledgehammer_Fact_Filter.relevant_facts ctxt full_types
+ relevance_thresholds
+ (the_default default_max_relevant max_relevant)
+ (the_default false theory_relevant)
+ {add = [], del = [], only = false} facts hyp_ts concl_t
+ |> map (fst o fst)
+ val (found_facts, lost_facts) =
+ List.concat proofs |> sort_distinct string_ord
+ |> map (fn fact => (find_index (curry (op =) fact) facts, fact))
+ |> List.partition (curry (op <=) 0 o fst)
+ |>> sort (prod_ord int_ord string_ord) ||> map snd
+ val found_proofs = filter (forall (member (op =) facts)) proofs
+ val n = length found_proofs
+ val _ =
+ if n = 0 then
+ (add id num_failures 1; log "Failure")
+ else
+ (add id num_successes 1;
+ add id num_found_proofs n;
+ log ("Success (" ^ string_of_int n ^ " of " ^
+ string_of_int (length proofs) ^ " proofs)"))
+ val _ = add id num_lost_proofs (length proofs - n)
+ val _ = add id num_found_facts (length found_facts)
+ val _ = add id num_lost_facts (length lost_facts)
+ val _ =
+ if null found_facts then
+ ()
+ else
+ let
+ val found_weight =
+ Real.fromInt (fold (fn (n, _) =>
+ Integer.add (n * n)) found_facts 0)
+ / Real.fromInt (length found_facts)
+ |> Math.sqrt |> Real.ceil
+ in
+ log ("Found facts (among " ^ string_of_int (length facts) ^
+ ", weight " ^ string_of_int found_weight ^ "): " ^
+ commas (map with_index found_facts))
+ end
+ val _ = if null lost_facts then
+ ()
+ else
+ log ("Lost facts (among " ^ string_of_int (length facts) ^
+ "): " ^ commas lost_facts)
+ in () end
+ | NONE => log "No known proof")
+ | _ => ()
+
+val proof_fileK = "proof_file"
+
+fun invoke args =
+ let
+ val (pf_args, other_args) =
+ args |> List.partition (curry (op =) proof_fileK o fst)
+ val proof_file = case pf_args of
+ [] => error "No \"proof_file\" specified"
+ | (_, s) :: _ => s
+ in Mirabelle.register (init proof_file, action other_args, done) end
+
+end;
+
+(* Workaround to keep the "mirabelle.pl" script happy *)
+structure Mirabelle_Sledgehammer_filter = Mirabelle_Sledgehammer_Filter;
--- a/src/HOL/Mirabelle/lib/scripts/mirabelle.pl Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Mirabelle/lib/scripts/mirabelle.pl Wed Sep 01 17:19:47 2010 +0200
@@ -51,7 +51,11 @@
}
my $tools = "";
if ($#action_files >= 0) {
- $tools = "uses " . join(" ", @action_files);
+ # uniquify
+ my $s = join ("\n", @action_files);
+ my @action_files = split(/\n/, $s . "\n" . $s);
+ %action_files = sort(@action_files);
+ $tools = "uses " . join(" ", sort(keys(%action_files)));
}
open(SETUP_FILE, ">$setup_file") || die "Could not create file '$setup_file'";
@@ -71,7 +75,7 @@
END
-foreach (split(/:/, $actions)) {
+foreach (reverse(split(/:/, $actions))) {
if (m/([^[]*)(?:\[(.*)\])?/) {
my ($name, $settings_str) = ($1, $2 || "");
$name =~ s/^([a-z])/\U$1/;
--- a/src/HOL/Multivariate_Analysis/Gauge_Measure.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Gauge_Measure.thy Wed Sep 01 17:19:47 2010 +0200
@@ -311,7 +311,7 @@
shows "(\<Union> f) has_gmeasure (setsum m f)" using assms
proof induct case (insert x s)
have *:"(x \<inter> \<Union>s) = \<Union>{x \<inter> y| y. y\<in>s}"by auto
- show ?case unfolding Union_insert ring_class.setsum.insert[OF insert(1-2)]
+ show ?case unfolding Union_insert setsum.insert [OF insert(1-2)]
apply(rule has_gmeasure_negligible_union) unfolding *
apply(rule insert) defer apply(rule insert) apply(rule insert) defer
apply(rule insert) prefer 4 apply(rule negligible_unions)
--- a/src/HOL/Mutabelle/Mutabelle.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Mutabelle/Mutabelle.thy Wed Sep 01 17:19:47 2010 +0200
@@ -4,7 +4,7 @@
begin
ML {*
-val comms = [@{const_name "op ="}, @{const_name "op |"}, @{const_name "op &"}];
+val comms = [@{const_name HOL.eq}, @{const_name HOL.disj}, @{const_name HOL.conj}];
val forbidden =
[(@{const_name Power.power}, "'a"),
--- a/src/HOL/Mutabelle/mutabelle_extra.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Mutabelle/mutabelle_extra.ML Wed Sep 01 17:19:47 2010 +0200
@@ -187,7 +187,7 @@
val nitpick_mtd = ("nitpick", invoke_nitpick)
*)
-val comms = [@{const_name "op ="}, @{const_name "op |"}, @{const_name "op &"}]
+val comms = [@{const_name HOL.eq}, @{const_name HOL.disj}, @{const_name HOL.conj}]
val forbidden =
[(* (@{const_name "power"}, "'a"), *)
@@ -202,7 +202,7 @@
(@{const_name "top_fun_inst.top_fun"}, "'a"),
(@{const_name "Pure.term"}, "'a"),
(@{const_name "top_class.top"}, "'a"),
- (@{const_name "eq_class.eq"}, "'a"),
+ (@{const_name "HOL.equal"}, "'a"),
(@{const_name "Quotient.Quot_True"}, "'a")(*,
(@{const_name "uminus"}, "'a"),
(@{const_name "Nat.size"}, "'a"),
@@ -237,7 +237,7 @@
@{const_name "top_fun_inst.top_fun"},
@{const_name "Pure.term"},
@{const_name "top_class.top"},
- @{const_name "eq_class.eq"},
+ @{const_name "HOL.equal"},
@{const_name "Quotient.Quot_True"}]
fun is_forbidden_mutant t =
--- a/src/HOL/Nominal/nominal_datatype.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Nominal/nominal_datatype.ML Wed Sep 01 17:19:47 2010 +0200
@@ -183,7 +183,7 @@
end;
fun mk_not_sym ths = maps (fn th => case prop_of th of
- _ $ (Const (@{const_name Not}, _) $ (Const (@{const_name "op ="}, _) $ _ $ _)) => [th, th RS not_sym]
+ _ $ (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ _)) => [th, th RS not_sym]
| _ => [th]) ths;
fun fresh_const T U = Const ("Nominal.fresh", T --> U --> HOLogic.boolT);
@@ -1200,7 +1200,7 @@
(constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs);
val tnames = Datatype_Prop.make_tnames recTs;
val zs = Name.variant_list tnames (replicate (length descr'') "z");
- val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
+ val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
(map (fn ((((i, _), T), tname), z) =>
make_pred fsT i T $ Free (z, fsT) $ Free (tname, T))
(descr'' ~~ recTs ~~ tnames ~~ zs)));
@@ -1215,7 +1215,7 @@
map (make_ind_prem fsT' (fn T => fn t => fn u => HOLogic.Not $
HOLogic.mk_mem (t, the (AList.lookup op = fresh_fs T) $ u)) i T)
(constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs);
- val ind_concl' = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
+ val ind_concl' = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
(map (fn ((((i, _), T), tname), z) =>
make_pred fsT' i T $ Free (z, fsT') $ Free (tname, T))
(descr'' ~~ recTs ~~ tnames ~~ zs)));
@@ -1225,7 +1225,7 @@
(Datatype_Prop.indexify_names (replicate (length dt_atomTs) "pi") ~~
map mk_permT dt_atomTs) @ [("z", fsT')];
val aux_ind_Ts = rev (map snd aux_ind_vars);
- val aux_ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
+ val aux_ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
(map (fn (((i, _), T), tname) =>
HOLogic.list_all (aux_ind_vars, make_pred fsT' i T $ Bound 0 $
fold_rev (mk_perm aux_ind_Ts) (map Bound (length dt_atomTs downto 1))
--- a/src/HOL/Nominal/nominal_inductive.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Nominal/nominal_inductive.ML Wed Sep 01 17:19:47 2010 +0200
@@ -71,7 +71,7 @@
| add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs
| add_binders thy i _ bs = bs;
-fun split_conj f names (Const (@{const_name "op &"}, _) $ p $ q) _ = (case head_of p of
+fun split_conj f names (Const (@{const_name HOL.conj}, _) $ p $ q) _ = (case head_of p of
Const (name, _) =>
if member (op =) names name then SOME (f p q) else NONE
| _ => NONE)
@@ -89,7 +89,7 @@
(* where "id" protects the subformula from simplification *)
(*********************************************************************)
-fun inst_conj_all names ps pis (Const (@{const_name "op &"}, _) $ p $ q) _ =
+fun inst_conj_all names ps pis (Const (@{const_name HOL.conj}, _) $ p $ q) _ =
(case head_of p of
Const (name, _) =>
if member (op =) names name then SOME (HOLogic.mk_conj (p,
--- a/src/HOL/Nominal/nominal_inductive2.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Nominal/nominal_inductive2.ML Wed Sep 01 17:19:47 2010 +0200
@@ -76,7 +76,7 @@
| add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs
| add_binders thy i _ bs = bs;
-fun split_conj f names (Const (@{const_name "op &"}, _) $ p $ q) _ = (case head_of p of
+fun split_conj f names (Const (@{const_name HOL.conj}, _) $ p $ q) _ = (case head_of p of
Const (name, _) =>
if member (op =) names name then SOME (f p q) else NONE
| _ => NONE)
@@ -94,7 +94,7 @@
(* where "id" protects the subformula from simplification *)
(*********************************************************************)
-fun inst_conj_all names ps pis (Const (@{const_name "op &"}, _) $ p $ q) _ =
+fun inst_conj_all names ps pis (Const (@{const_name HOL.conj}, _) $ p $ q) _ =
(case head_of p of
Const (name, _) =>
if member (op =) names name then SOME (HOLogic.mk_conj (p,
--- a/src/HOL/Nominal/nominal_thmdecls.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Nominal/nominal_thmdecls.ML Wed Sep 01 17:19:47 2010 +0200
@@ -18,8 +18,6 @@
val eqvt_force_del: attribute
val setup: theory -> theory
val get_eqvt_thms: Proof.context -> thm list
-
- val NOMINAL_EQVT_DEBUG : bool Unsynchronized.ref
end;
structure NominalThmDecls: NOMINAL_THMDECLS =
@@ -44,29 +42,31 @@
(* equality-lemma can be derived. *)
exception EQVT_FORM of string
-val NOMINAL_EQVT_DEBUG = Unsynchronized.ref false
+val (nominal_eqvt_debug, setup_nominal_eqvt_debug) =
+ Attrib.config_bool "nominal_eqvt_debug" (K false);
-fun tactic (msg, tac) =
- if !NOMINAL_EQVT_DEBUG
+fun tactic ctxt (msg, tac) =
+ if Config.get ctxt nominal_eqvt_debug
then tac THEN' (K (print_tac ("after " ^ msg)))
else tac
fun prove_eqvt_tac ctxt orig_thm pi pi' =
-let
- val mypi = Thm.cterm_of ctxt pi
- val T = fastype_of pi'
- val mypifree = Thm.cterm_of ctxt (Const (@{const_name "rev"}, T --> T) $ pi')
- val perm_pi_simp = PureThy.get_thms ctxt "perm_pi_simp"
-in
- EVERY1 [tactic ("iffI applied", rtac @{thm iffI}),
- tactic ("remove pi with perm_boolE", dtac @{thm perm_boolE}),
- tactic ("solve with orig_thm", etac orig_thm),
- tactic ("applies orig_thm instantiated with rev pi",
- dtac (Drule.cterm_instantiate [(mypi,mypifree)] orig_thm)),
- tactic ("getting rid of the pi on the right", rtac @{thm perm_boolI}),
- tactic ("getting rid of all remaining perms",
- full_simp_tac (HOL_basic_ss addsimps perm_pi_simp))]
-end;
+ let
+ val thy = ProofContext.theory_of ctxt
+ val mypi = Thm.cterm_of thy pi
+ val T = fastype_of pi'
+ val mypifree = Thm.cterm_of thy (Const (@{const_name "rev"}, T --> T) $ pi')
+ val perm_pi_simp = PureThy.get_thms thy "perm_pi_simp"
+ in
+ EVERY1 [tactic ctxt ("iffI applied", rtac @{thm iffI}),
+ tactic ctxt ("remove pi with perm_boolE", dtac @{thm perm_boolE}),
+ tactic ctxt ("solve with orig_thm", etac orig_thm),
+ tactic ctxt ("applies orig_thm instantiated with rev pi",
+ dtac (Drule.cterm_instantiate [(mypi,mypifree)] orig_thm)),
+ tactic ctxt ("getting rid of the pi on the right", rtac @{thm perm_boolI}),
+ tactic ctxt ("getting rid of all remaining perms",
+ full_simp_tac (HOL_basic_ss addsimps perm_pi_simp))]
+ end;
fun get_derived_thm ctxt hyp concl orig_thm pi typi =
let
@@ -78,7 +78,7 @@
val _ = writeln (Syntax.string_of_term ctxt' goal_term);
in
Goal.prove ctxt' [] [] goal_term
- (fn _ => prove_eqvt_tac thy orig_thm pi' pi'') |>
+ (fn _ => prove_eqvt_tac ctxt' orig_thm pi' pi'') |>
singleton (ProofContext.export ctxt' ctxt)
end
@@ -147,7 +147,7 @@
else raise EQVT_FORM "Type Implication"
end
(* case: eqvt-lemma is of the equational form *)
- | (Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op ="}, _) $
+ | (Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $
(Const (@{const_name "perm"},typrm) $ Var (pi,typi) $ lhs) $ rhs)) =>
(if (apply_pi lhs (pi,typi)) = rhs
then [orig_thm]
@@ -170,11 +170,12 @@
val get_eqvt_thms = Context.Proof #> Data.get;
val setup =
- Attrib.setup @{binding eqvt} (Attrib.add_del eqvt_add eqvt_del)
- "equivariance theorem declaration"
- #> Attrib.setup @{binding eqvt_force} (Attrib.add_del eqvt_force_add eqvt_force_del)
- "equivariance theorem declaration (without checking the form of the lemma)"
- #> PureThy.add_thms_dynamic (Binding.name "eqvts", Data.get)
+ setup_nominal_eqvt_debug #>
+ Attrib.setup @{binding eqvt} (Attrib.add_del eqvt_add eqvt_del)
+ "equivariance theorem declaration" #>
+ Attrib.setup @{binding eqvt_force} (Attrib.add_del eqvt_force_add eqvt_force_del)
+ "equivariance theorem declaration (without checking the form of the lemma)" #>
+ PureThy.add_thms_dynamic (@{binding eqvts}, Data.get);
end;
--- a/src/HOL/Option.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Option.thy Wed Sep 01 17:19:47 2010 +0200
@@ -99,8 +99,8 @@
by (simp add: is_none_def)
lemma [code_unfold]:
- "eq_class.eq x None \<longleftrightarrow> is_none x"
- by (simp add: eq is_none_none)
+ "HOL.equal x None \<longleftrightarrow> is_none x"
+ by (simp add: equal is_none_none)
hide_const (open) is_none
@@ -116,10 +116,10 @@
(Haskell "Nothing" and "Just")
(Scala "!None" and "Some")
-code_instance option :: eq
+code_instance option :: equal
(Haskell -)
-code_const "eq_class.eq \<Colon> 'a\<Colon>eq option \<Rightarrow> 'a option \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> 'a option \<Rightarrow> 'a option \<Rightarrow> bool"
(Haskell infixl 4 "==")
code_reserved SML
--- a/src/HOL/Orderings.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Orderings.thy Wed Sep 01 17:19:47 2010 +0200
@@ -640,8 +640,8 @@
let
val All_binder = Syntax.binder_name @{const_syntax All};
val Ex_binder = Syntax.binder_name @{const_syntax Ex};
- val impl = @{const_syntax "op -->"};
- val conj = @{const_syntax "op &"};
+ val impl = @{const_syntax HOL.implies};
+ val conj = @{const_syntax HOL.conj};
val less = @{const_syntax less};
val less_eq = @{const_syntax less_eq};
--- a/src/HOL/Predicate.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Predicate.thy Wed Sep 01 17:19:47 2010 +0200
@@ -808,8 +808,12 @@
lemma eq_pred_code [code]:
fixes P Q :: "'a pred"
- shows "eq_class.eq P Q \<longleftrightarrow> P \<le> Q \<and> Q \<le> P"
- unfolding eq by auto
+ shows "HOL.equal P Q \<longleftrightarrow> P \<le> Q \<and> Q \<le> P"
+ by (auto simp add: equal)
+
+lemma [code nbe]:
+ "HOL.equal (x :: 'a pred) x \<longleftrightarrow> True"
+ by (fact equal_refl)
lemma [code]:
"pred_case f P = f (eval P)"
--- a/src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy Wed Sep 01 17:19:47 2010 +0200
@@ -4,15 +4,42 @@
section {* Example append *}
+
inductive append
where
"append [] ys ys"
| "append xs ys zs ==> append (x # xs) ys (x # zs)"
-code_pred append .
+setup {* Code_Prolog.map_code_options (K
+ {ensure_groundness = false,
+ limited_types = [],
+ limited_predicates = [],
+ replacing = [],
+ manual_reorder = [],
+ prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+values "{(x, y, z). append x y z}"
values 3 "{(x, y, z). append x y z}"
+setup {* Code_Prolog.map_code_options (K
+ {ensure_groundness = false,
+ limited_types = [],
+ limited_predicates = [],
+ replacing = [],
+ manual_reorder = [],
+ prolog_system = Code_Prolog.YAP}) *}
+
+values "{(x, y, z). append x y z}"
+
+setup {* Code_Prolog.map_code_options (K
+ {ensure_groundness = false,
+ limited_types = [],
+ limited_predicates = [],
+ replacing = [],
+ manual_reorder = [],
+ prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
section {* Example queens *}
@@ -71,9 +98,7 @@
where
"queen [1,2,3,4,5,6,7,8,9] ys ==> queen_9 ys"
-code_pred queen_9 .
-
-values 150 "{ys. queen_9 ys}"
+values 10 "{ys. queen_9 ys}"
section {* Example symbolic derivation *}
@@ -163,14 +188,41 @@
where
"d (Mult (Mult (Mult (Mult (Mult (Mult (Mult (Mult (Mult x x) x) x) x) x) x) x) x) x) x e ==> times10 e"
-code_pred ops8 .
-code_pred divide10 .
-code_pred log10 .
-code_pred times10 .
-
values "{e. ops8 e}"
values "{e. divide10 e}"
values "{e. log10 e}"
values "{e. times10 e}"
+section {* Example negation *}
+
+datatype abc = A | B | C
+
+inductive notB :: "abc => bool"
+where
+ "y \<noteq> B \<Longrightarrow> notB y"
+
+setup {* Code_Prolog.map_code_options (K
+ {ensure_groundness = true,
+ limited_types = [],
+ limited_predicates = [],
+ replacing = [],
+ manual_reorder = [],
+ prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+values 2 "{y. notB y}"
+
+inductive notAB :: "abc * abc \<Rightarrow> bool"
+where
+ "y \<noteq> A \<Longrightarrow> z \<noteq> B \<Longrightarrow> notAB (y, z)"
+
+values 5 "{y. notAB y}"
+
+section {* Example prolog conform variable names *}
+
+inductive equals :: "abc => abc => bool"
+where
+ "equals y' y'"
+
+values 1 "{(y, z). equals y z}"
+
end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Predicate_Compile_Examples/Context_Free_Grammar_Example.thy Wed Sep 01 17:19:47 2010 +0200
@@ -0,0 +1,169 @@
+theory Context_Free_Grammar_Example
+imports Code_Prolog
+begin
+
+declare mem_def[code_pred_inline]
+
+
+subsection {* Alternative rules for length *}
+
+definition size_list :: "'a list => nat"
+where "size_list = size"
+
+lemma size_list_simps:
+ "size_list [] = 0"
+ "size_list (x # xs) = Suc (size_list xs)"
+by (auto simp add: size_list_def)
+
+declare size_list_simps[code_pred_def]
+declare size_list_def[symmetric, code_pred_inline]
+
+
+setup {* Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck) *}
+
+datatype alphabet = a | b
+
+inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
+ "[] \<in> S\<^isub>1"
+| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
+| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
+| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
+| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
+| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
+
+lemma
+ "w \<in> S\<^isub>1 \<Longrightarrow> w = []"
+quickcheck[generator = prolog, iterations=1, expect = counterexample]
+oops
+
+definition "filter_a = filter (\<lambda>x. x = a)"
+
+lemma [code_pred_def]: "filter_a [] = []"
+unfolding filter_a_def by simp
+
+lemma [code_pred_def]: "filter_a (x#xs) = (if x = a then x # filter_a xs else filter_a xs)"
+unfolding filter_a_def by simp
+
+declare filter_a_def[symmetric, code_pred_inline]
+
+definition "filter_b = filter (\<lambda>x. x = b)"
+
+lemma [code_pred_def]: "filter_b [] = []"
+unfolding filter_b_def by simp
+
+lemma [code_pred_def]: "filter_b (x#xs) = (if x = b then x # filter_b xs else filter_b xs)"
+unfolding filter_b_def by simp
+
+declare filter_b_def[symmetric, code_pred_inline]
+
+setup {* Code_Prolog.map_code_options (K
+ {ensure_groundness = true,
+ limited_types = [],
+ limited_predicates = [(["s1", "a1", "b1"], 2)],
+ replacing = [(("s1", "limited_s1"), "quickcheck")],
+ manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])],
+ prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+
+theorem S\<^isub>1_sound:
+"w \<in> S\<^isub>1 \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+quickcheck[generator = prolog, iterations=1, expect = counterexample]
+oops
+
+
+inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
+ "[] \<in> S\<^isub>2"
+| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
+| "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2"
+| "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
+| "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
+| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
+
+
+setup {* Code_Prolog.map_code_options (K
+ {ensure_groundness = true,
+ limited_types = [],
+ limited_predicates = [(["s2", "a2", "b2"], 3)],
+ replacing = [(("s2", "limited_s2"), "quickcheck")],
+ manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])],
+ prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+
+theorem S\<^isub>2_sound:
+"w \<in> S\<^isub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+quickcheck[generator=prolog, iterations=1, expect = counterexample]
+oops
+
+inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
+ "[] \<in> S\<^isub>3"
+| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3"
+| "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3"
+| "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3"
+| "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3"
+| "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3"
+
+
+setup {* Code_Prolog.map_code_options (K
+ {ensure_groundness = true,
+ limited_types = [],
+ limited_predicates = [(["s3", "a3", "b3"], 6)],
+ replacing = [(("s3", "limited_s3"), "quickcheck")],
+ manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])],
+ prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+lemma S\<^isub>3_sound:
+"w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+quickcheck[generator=prolog, iterations=1, size=1, expect = no_counterexample]
+oops
+
+
+(*
+setup {* Code_Prolog.map_code_options (K
+ {ensure_groundness = true,
+ limited_types = [],
+ limited_predicates = [],
+ replacing = [],
+ manual_reorder = [],
+ prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+
+theorem S\<^isub>3_complete:
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>3"
+quickcheck[generator = prolog, size=1, iterations=1]
+oops
+*)
+
+inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
+ "[] \<in> S\<^isub>4"
+| "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4"
+| "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4"
+| "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4"
+| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
+| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
+| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
+
+
+setup {* Code_Prolog.map_code_options (K
+ {ensure_groundness = true,
+ limited_types = [],
+ limited_predicates = [(["s4", "a4", "b4"], 6)],
+ replacing = [(("s4", "limited_s4"), "quickcheck")],
+ manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])],
+ prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+
+theorem S\<^isub>4_sound:
+"w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+quickcheck[generator = prolog, size=1, iterations=1, expect = no_counterexample]
+oops
+
+(*
+theorem S\<^isub>4_complete:
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4"
+oops
+*)
+
+hide_const a b
+
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Predicate_Compile_Examples/Hotel_Example.thy Wed Sep 01 17:19:47 2010 +0200
@@ -0,0 +1,129 @@
+theory Hotel_Example
+imports Predicate_Compile_Alternative_Defs Code_Prolog
+begin
+
+datatype guest = Guest0 | Guest1
+datatype key = Key0 | Key1 | Key2 | Key3
+datatype room = Room0
+
+types card = "key * key"
+
+datatype event =
+ Check_in guest room card | Enter guest room card | Exit guest room
+
+definition initk :: "room \<Rightarrow> key"
+ where "initk = (%r. Key0)"
+
+declare initk_def[code_pred_def, code]
+
+primrec owns :: "event list \<Rightarrow> room \<Rightarrow> guest option"
+where
+ "owns [] r = None"
+| "owns (e#s) r = (case e of
+ Check_in g r' c \<Rightarrow> if r' = r then Some g else owns s r |
+ Enter g r' c \<Rightarrow> owns s r |
+ Exit g r' \<Rightarrow> owns s r)"
+
+primrec currk :: "event list \<Rightarrow> room \<Rightarrow> key"
+where
+ "currk [] r = initk r"
+| "currk (e#s) r = (let k = currk s r in
+ case e of Check_in g r' (k1, k2) \<Rightarrow> if r' = r then k2 else k
+ | Enter g r' c \<Rightarrow> k
+ | Exit g r \<Rightarrow> k)"
+
+primrec issued :: "event list \<Rightarrow> key set"
+where
+ "issued [] = range initk"
+| "issued (e#s) = issued s \<union>
+ (case e of Check_in g r (k1, k2) \<Rightarrow> {k2} | Enter g r c \<Rightarrow> {} | Exit g r \<Rightarrow> {})"
+
+primrec cards :: "event list \<Rightarrow> guest \<Rightarrow> card set"
+where
+ "cards [] g = {}"
+| "cards (e#s) g = (let C = cards s g in
+ case e of Check_in g' r c \<Rightarrow> if g' = g then insert c C
+ else C
+ | Enter g r c \<Rightarrow> C
+ | Exit g r \<Rightarrow> C)"
+
+primrec roomk :: "event list \<Rightarrow> room \<Rightarrow> key"
+where
+ "roomk [] r = initk r"
+| "roomk (e#s) r = (let k = roomk s r in
+ case e of Check_in g r' c \<Rightarrow> k
+ | Enter g r' (x,y) \<Rightarrow> if r' = r (*& x = k*) then y else k
+ | Exit g r \<Rightarrow> k)"
+
+primrec isin :: "event list \<Rightarrow> room \<Rightarrow> guest set"
+where
+ "isin [] r = {}"
+| "isin (e#s) r = (let G = isin s r in
+ case e of Check_in g r c \<Rightarrow> G
+ | Enter g r' c \<Rightarrow> if r' = r then {g} \<union> G else G
+ | Exit g r' \<Rightarrow> if r'=r then G - {g} else G)"
+
+primrec hotel :: "event list \<Rightarrow> bool"
+where
+ "hotel [] = True"
+| "hotel (e # s) = (hotel s & (case e of
+ Check_in g r (k,k') \<Rightarrow> k = currk s r \<and> k' \<notin> issued s |
+ Enter g r (k,k') \<Rightarrow> (k,k') : cards s g & (roomk s r : {k, k'}) |
+ Exit g r \<Rightarrow> g : isin s r))"
+
+lemma issued_nil: "issued [] = {Key0}"
+by (auto simp add: initk_def)
+
+lemmas issued_simps[code, code_pred_def] = issued_nil issued.simps(2)
+
+declare Let_def[code_pred_inline]
+
+lemma [code_pred_inline]: "insert == (%y A x. y = x | A x)"
+by (auto simp add: insert_iff[unfolded mem_def] expand_fun_eq intro!: eq_reflection)
+
+lemma [code_pred_inline]: "(op -) == (%A B x. A x \<and> \<not> B x)"
+by (auto simp add: Diff_iff[unfolded mem_def] expand_fun_eq intro!: eq_reflection)
+
+setup {* Code_Prolog.map_code_options (K
+ {ensure_groundness = true,
+ limited_types = [],
+ limited_predicates = [],
+ replacing = [],
+ manual_reorder = [],
+ prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+values 40 "{s. hotel s}"
+
+
+setup {* Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck) *}
+
+lemma "\<lbrakk> hotel s; g \<in> isin s r \<rbrakk> \<Longrightarrow> owns s r = Some g"
+quickcheck[generator = code, iterations = 100000, report]
+quickcheck[generator = prolog, iterations = 1, expect = counterexample]
+oops
+
+
+definition no_Check_in :: "event list \<Rightarrow> room \<Rightarrow> bool" where(*>*)
+[code del]: "no_Check_in s r \<equiv> \<not>(\<exists>g c. Check_in g r c \<in> set s)"
+
+
+definition feels_safe :: "event list \<Rightarrow> room \<Rightarrow> bool"
+where
+ "feels_safe s r = (\<exists>s\<^isub>1 s\<^isub>2 s\<^isub>3 g c c'.
+ s = s\<^isub>3 @ [Enter g r c] @ s\<^isub>2 @ [Check_in g r c'] @ s\<^isub>1 \<and>
+ no_Check_in (s\<^isub>3 @ s\<^isub>2) r \<and> isin (s\<^isub>2 @ [Check_in g r c] @ s\<^isub>1) r = {})"
+
+setup {* Code_Prolog.map_code_options (K
+ {ensure_groundness = true,
+ limited_types = [],
+ limited_predicates = [(["hotel"], 5)],
+ replacing = [(("hotel", "limited_hotel"), "quickcheck")],
+ manual_reorder = [],
+ prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+lemma
+ "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
+quickcheck[generator = prolog, iterations = 1, expect = counterexample]
+oops
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Predicate_Compile_Examples/Lambda_Example.thy Wed Sep 01 17:19:47 2010 +0200
@@ -0,0 +1,126 @@
+theory Lambda_Example
+imports Code_Prolog
+begin
+
+subsection {* Lambda *}
+
+datatype type =
+ Atom nat
+ | Fun type type (infixr "\<Rightarrow>" 200)
+
+datatype dB =
+ Var nat
+ | App dB dB (infixl "\<degree>" 200)
+ | Abs type dB
+
+primrec
+ nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
+where
+ "[]\<langle>i\<rangle> = None"
+| "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
+
+inductive nth_el1 :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool"
+where
+ "nth_el1 (x # xs) 0 x"
+| "nth_el1 xs i y \<Longrightarrow> nth_el1 (x # xs) (Suc i) y"
+
+inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" ("_ \<turnstile> _ : _" [50, 50, 50] 50)
+ where
+ Var [intro!]: "nth_el1 env x T \<Longrightarrow> env \<turnstile> Var x : T"
+ | Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)"
+ | App [intro!]: "env \<turnstile> s : U \<Rightarrow> T \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U"
+
+primrec
+ lift :: "[dB, nat] => dB"
+where
+ "lift (Var i) k = (if i < k then Var i else Var (i + 1))"
+ | "lift (s \<degree> t) k = lift s k \<degree> lift t k"
+ | "lift (Abs T s) k = Abs T (lift s (k + 1))"
+
+primrec pred :: "nat => nat"
+where
+ "pred (Suc i) = i"
+
+primrec
+ subst :: "[dB, dB, nat] => dB" ("_[_'/_]" [300, 0, 0] 300)
+where
+ subst_Var: "(Var i)[s/k] =
+ (if k < i then Var (pred i) else if i = k then s else Var i)"
+ | subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
+ | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"
+
+inductive beta :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<beta>" 50)
+ where
+ beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
+ | appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
+ | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
+ | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"
+
+subsection {* Inductive definitions for ordering on naturals *}
+
+inductive less_nat
+where
+ "less_nat 0 (Suc y)"
+| "less_nat x y ==> less_nat (Suc x) (Suc y)"
+
+lemma less_nat[code_pred_inline]:
+ "x < y = less_nat x y"
+apply (rule iffI)
+apply (induct x arbitrary: y)
+apply (case_tac y) apply (auto intro: less_nat.intros)
+apply (case_tac y)
+apply (auto intro: less_nat.intros)
+apply (induct rule: less_nat.induct)
+apply auto
+done
+
+lemma [code_pred_inline]: "(x::nat) + 1 = Suc x"
+by simp
+
+section {* Manual setup to find counterexample *}
+
+setup {* Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck) *}
+
+setup {* Code_Prolog.map_code_options (K
+ { ensure_groundness = true,
+ limited_types = [(@{typ nat}, 1), (@{typ "type"}, 1), (@{typ dB}, 1), (@{typ "type list"}, 1)],
+ limited_predicates = [(["typing"], 2), (["nthel1"], 2)],
+ replacing = [(("typing", "limited_typing"), "quickcheck"),
+ (("nthel1", "limited_nthel1"), "lim_typing")],
+ manual_reorder = [],
+ prolog_system = Code_Prolog.SWI_PROLOG}) *}
+
+lemma
+ "\<Gamma> \<turnstile> t : U \<Longrightarrow> t \<rightarrow>\<^sub>\<beta> t' \<Longrightarrow> \<Gamma> \<turnstile> t' : U"
+quickcheck[generator = prolog, iterations = 1, expect = counterexample]
+oops
+
+text {* Verifying that the found counterexample really is one by means of a proof *}
+
+lemma
+assumes
+ "t' = Var 0"
+ "U = Atom 0"
+ "\<Gamma> = [Atom 1]"
+ "t = App (Abs (Atom 0) (Var 1)) (Var 0)"
+shows
+ "\<Gamma> \<turnstile> t : U"
+ "t \<rightarrow>\<^sub>\<beta> t'"
+ "\<not> \<Gamma> \<turnstile> t' : U"
+proof -
+ from assms show "\<Gamma> \<turnstile> t : U"
+ by (auto intro!: typing.intros nth_el1.intros)
+next
+ from assms have "t \<rightarrow>\<^sub>\<beta> (Var 1)[Var 0/0]"
+ by (auto simp only: intro: beta.intros)
+ moreover
+ from assms have "(Var 1)[Var 0/0] = t'" by simp
+ ultimately show "t \<rightarrow>\<^sub>\<beta> t'" by simp
+next
+ from assms show "\<not> \<Gamma> \<turnstile> t' : U"
+ by (auto elim: typing.cases nth_el1.cases)
+qed
+
+
+end
+
--- a/src/HOL/Predicate_Compile_Examples/ROOT.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Predicate_Compile_Examples/ROOT.ML Wed Sep 01 17:19:47 2010 +0200
@@ -1,2 +1,2 @@
use_thys ["Predicate_Compile_Examples", "Predicate_Compile_Quickcheck_Examples", "Specialisation_Examples"];
-if getenv "PROLOG_HOME" = "" then () else use_thys ["Code_Prolog_Examples"];
+if getenv "PROLOG_HOME" = "" then () else use_thys ["Code_Prolog_Examples", "Hotel_Example", "Lambda_Example"];
--- a/src/HOL/Product_Type.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Product_Type.thy Wed Sep 01 17:19:47 2010 +0200
@@ -21,17 +21,17 @@
-- "prefer plain propositional version"
lemma
- shows [code]: "eq_class.eq False P \<longleftrightarrow> \<not> P"
- and [code]: "eq_class.eq True P \<longleftrightarrow> P"
- and [code]: "eq_class.eq P False \<longleftrightarrow> \<not> P"
- and [code]: "eq_class.eq P True \<longleftrightarrow> P"
- and [code nbe]: "eq_class.eq P P \<longleftrightarrow> True"
- by (simp_all add: eq)
+ shows [code]: "HOL.equal False P \<longleftrightarrow> \<not> P"
+ and [code]: "HOL.equal True P \<longleftrightarrow> P"
+ and [code]: "HOL.equal P False \<longleftrightarrow> \<not> P"
+ and [code]: "HOL.equal P True \<longleftrightarrow> P"
+ and [code nbe]: "HOL.equal P P \<longleftrightarrow> True"
+ by (simp_all add: equal)
-code_const "eq_class.eq \<Colon> bool \<Rightarrow> bool \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> bool \<Rightarrow> bool \<Rightarrow> bool"
(Haskell infixl 4 "==")
-code_instance bool :: eq
+code_instance bool :: equal
(Haskell -)
@@ -92,7 +92,7 @@
end
lemma [code]:
- "eq_class.eq (u\<Colon>unit) v \<longleftrightarrow> True" unfolding eq unit_eq [of u] unit_eq [of v] by rule+
+ "HOL.equal (u\<Colon>unit) v \<longleftrightarrow> True" unfolding equal unit_eq [of u] unit_eq [of v] by rule+
code_type unit
(SML "unit")
@@ -106,10 +106,10 @@
(Haskell "()")
(Scala "()")
-code_instance unit :: eq
+code_instance unit :: equal
(Haskell -)
-code_const "eq_class.eq \<Colon> unit \<Rightarrow> unit \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> unit \<Rightarrow> unit \<Rightarrow> bool"
(Haskell infixl 4 "==")
code_reserved SML
@@ -277,10 +277,10 @@
(Haskell "!((_),/ (_))")
(Scala "!((_),/ (_))")
-code_instance prod :: eq
+code_instance prod :: equal
(Haskell -)
-code_const "eq_class.eq \<Colon> 'a\<Colon>eq \<times> 'b\<Colon>eq \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> 'a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool"
(Haskell infixl 4 "==")
types_code
--- a/src/HOL/Prolog/prolog.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Prolog/prolog.ML Wed Sep 01 17:19:47 2010 +0200
@@ -2,7 +2,7 @@
Author: David von Oheimb (based on a lecture on Lambda Prolog by Nadathur)
*)
-Unsynchronized.set Proof.show_main_goal;
+Proof.show_main_goal := true;
structure Prolog =
struct
@@ -11,12 +11,12 @@
fun isD t = case t of
Const(@{const_name Trueprop},_)$t => isD t
- | Const(@{const_name "op &"} ,_)$l$r => isD l andalso isD r
- | Const(@{const_name "op -->"},_)$l$r => isG l andalso isD r
+ | Const(@{const_name HOL.conj} ,_)$l$r => isD l andalso isD r
+ | Const(@{const_name HOL.implies},_)$l$r => isG l andalso isD r
| Const( "==>",_)$l$r => isG l andalso isD r
| Const(@{const_name All},_)$Abs(s,_,t) => isD t
| Const("all",_)$Abs(s,_,t) => isD t
- | Const(@{const_name "op |"},_)$_$_ => false
+ | Const(@{const_name HOL.disj},_)$_$_ => false
| Const(@{const_name Ex} ,_)$_ => false
| Const(@{const_name Not},_)$_ => false
| Const(@{const_name True},_) => false
@@ -30,9 +30,9 @@
and
isG t = case t of
Const(@{const_name Trueprop},_)$t => isG t
- | Const(@{const_name "op &"} ,_)$l$r => isG l andalso isG r
- | Const(@{const_name "op |"} ,_)$l$r => isG l andalso isG r
- | Const(@{const_name "op -->"},_)$l$r => isD l andalso isG r
+ | Const(@{const_name HOL.conj} ,_)$l$r => isG l andalso isG r
+ | Const(@{const_name HOL.disj} ,_)$l$r => isG l andalso isG r
+ | Const(@{const_name HOL.implies},_)$l$r => isD l andalso isG r
| Const( "==>",_)$l$r => isD l andalso isG r
| Const(@{const_name All},_)$Abs(_,_,t) => isG t
| Const("all",_)$Abs(_,_,t) => isG t
@@ -53,8 +53,8 @@
fun at thm = case concl_of thm of
_$(Const(@{const_name All} ,_)$Abs(s,_,_))=> at(thm RS
(read_instantiate ctxt [(("x", 0), "?" ^ (if s="P" then "PP" else s))] spec))
- | _$(Const(@{const_name "op &"},_)$_$_) => at(thm RS conjunct1)@at(thm RS conjunct2)
- | _$(Const(@{const_name "op -->"},_)$_$_) => at(thm RS mp)
+ | _$(Const(@{const_name HOL.conj},_)$_$_) => at(thm RS conjunct1)@at(thm RS conjunct2)
+ | _$(Const(@{const_name HOL.implies},_)$_$_) => at(thm RS mp)
| _ => [thm]
in map zero_var_indexes (at thm) end;
@@ -72,7 +72,7 @@
-- is nice, but cannot instantiate unknowns in the assumptions *)
fun hyp_resolve_tac i st = let
fun ap (Const(@{const_name All},_)$Abs(_,_,t))=(case ap t of (k,a,t) => (k+1,a ,t))
- | ap (Const(@{const_name "op -->"},_)$_$t) =(case ap t of (k,_,t) => (k,true ,t))
+ | ap (Const(@{const_name HOL.implies},_)$_$t) =(case ap t of (k,_,t) => (k,true ,t))
| ap t = (0,false,t);
(*
fun rep_goal (Const ("all",_)$Abs (_,_,t)) = rep_goal t
--- a/src/HOL/Quickcheck.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Quickcheck.thy Wed Sep 01 17:19:47 2010 +0200
@@ -97,7 +97,7 @@
\<Rightarrow> Random.seed \<Rightarrow> (('a\<Colon>term_of \<Rightarrow> 'b\<Colon>typerep) \<times> (unit \<Rightarrow> term)) \<times> Random.seed" where
"random_fun_lift f = random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Evaluation.term_of f Random.split_seed"
-instantiation "fun" :: ("{eq, term_of}", random) random
+instantiation "fun" :: ("{equal, term_of}", random) random
begin
definition random_fun :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> Random.seed" where
--- a/src/HOL/Quotient.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Quotient.thy Wed Sep 01 17:19:47 2010 +0200
@@ -651,6 +651,16 @@
shows "(R1 ===> (R1 ===> R2) ===> R2) Let Let"
by auto
+lemma mem_rsp:
+ shows "(R1 ===> (R1 ===> R2) ===> R2) op \<in> op \<in>"
+ by (simp add: mem_def)
+
+lemma mem_prs:
+ assumes a1: "Quotient R1 Abs1 Rep1"
+ and a2: "Quotient R2 Abs2 Rep2"
+ shows "(Rep1 ---> (Abs1 ---> Rep2) ---> Abs2) op \<in> = op \<in>"
+ by (simp add: expand_fun_eq mem_def Quotient_abs_rep[OF a1] Quotient_abs_rep[OF a2])
+
locale quot_type =
fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
and Abs :: "('a \<Rightarrow> bool) \<Rightarrow> 'b"
@@ -721,8 +731,8 @@
declare [[map "fun" = (fun_map, fun_rel)]]
lemmas [quot_thm] = fun_quotient
-lemmas [quot_respect] = quot_rel_rsp if_rsp o_rsp let_rsp
-lemmas [quot_preserve] = if_prs o_prs let_prs
+lemmas [quot_respect] = quot_rel_rsp if_rsp o_rsp let_rsp mem_rsp
+lemmas [quot_preserve] = if_prs o_prs let_prs mem_prs
lemmas [quot_equiv] = identity_equivp
@@ -773,20 +783,20 @@
method_setup lifting =
{* Attrib.thms >> (fn thms => fn ctxt =>
- SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.lift_tac ctxt thms))) *}
+ SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.lift_tac ctxt [] thms))) *}
{* lifts theorems to quotient types *}
method_setup lifting_setup =
{* Attrib.thm >> (fn thm => fn ctxt =>
- SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.lift_procedure_tac ctxt thm))) *}
+ SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.lift_procedure_tac ctxt [] thm))) *}
{* sets up the three goals for the quotient lifting procedure *}
method_setup descending =
- {* Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.descend_tac ctxt))) *}
+ {* Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.descend_tac ctxt []))) *}
{* decends theorems to the raw level *}
method_setup descending_setup =
- {* Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.descend_procedure_tac ctxt))) *}
+ {* Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.descend_procedure_tac ctxt []))) *}
{* sets up the three goals for the decending theorems *}
method_setup regularize =
--- a/src/HOL/Rat.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Rat.thy Wed Sep 01 17:19:47 2010 +0200
@@ -1083,18 +1083,18 @@
by (cases "0::int" a rule: linorder_cases) (simp_all add: quotient_of_Fract)
qed
-instantiation rat :: eq
+instantiation rat :: equal
begin
definition [code]:
- "eq_class.eq a b \<longleftrightarrow> quotient_of a = quotient_of b"
+ "HOL.equal a b \<longleftrightarrow> quotient_of a = quotient_of b"
instance proof
-qed (simp add: eq_rat_def quotient_of_inject_eq)
+qed (simp add: equal_rat_def quotient_of_inject_eq)
lemma rat_eq_refl [code nbe]:
- "eq_class.eq (r::rat) r \<longleftrightarrow> True"
- by (rule HOL.eq_refl)
+ "HOL.equal (r::rat) r \<longleftrightarrow> True"
+ by (rule equal_refl)
end
--- a/src/HOL/RealDef.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/RealDef.thy Wed Sep 01 17:19:47 2010 +0200
@@ -1697,19 +1697,21 @@
"Ratreal (number_of r / number_of s) = number_of r / number_of s"
unfolding Ratreal_number_of_quotient [symmetric] Ratreal_def of_rat_divide ..
-instantiation real :: eq
+instantiation real :: equal
begin
-definition "eq_class.eq (x\<Colon>real) y \<longleftrightarrow> x - y = 0"
+definition "HOL.equal (x\<Colon>real) y \<longleftrightarrow> x - y = 0"
-instance by default (simp add: eq_real_def)
+instance proof
+qed (simp add: equal_real_def)
-lemma real_eq_code [code]: "eq_class.eq (Ratreal x) (Ratreal y) \<longleftrightarrow> eq_class.eq x y"
- by (simp add: eq_real_def eq)
+lemma real_equal_code [code]:
+ "HOL.equal (Ratreal x) (Ratreal y) \<longleftrightarrow> HOL.equal x y"
+ by (simp add: equal_real_def equal)
-lemma real_eq_refl [code nbe]:
- "eq_class.eq (x::real) x \<longleftrightarrow> True"
- by (rule HOL.eq_refl)
+lemma [code nbe]:
+ "HOL.equal (x::real) x \<longleftrightarrow> True"
+ by (rule equal_refl)
end
--- a/src/HOL/Set.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Set.thy Wed Sep 01 17:19:47 2010 +0200
@@ -219,8 +219,8 @@
val Type (set_type, _) = @{typ "'a set"}; (* FIXME 'a => bool (!?!) *)
val All_binder = Syntax.binder_name @{const_syntax All};
val Ex_binder = Syntax.binder_name @{const_syntax Ex};
- val impl = @{const_syntax "op -->"};
- val conj = @{const_syntax "op &"};
+ val impl = @{const_syntax HOL.implies};
+ val conj = @{const_syntax HOL.conj};
val sbset = @{const_syntax subset};
val sbset_eq = @{const_syntax subset_eq};
@@ -268,8 +268,8 @@
fun setcompr_tr [e, idts, b] =
let
- val eq = Syntax.const @{const_syntax "op ="} $ Bound (nvars idts) $ e;
- val P = Syntax.const @{const_syntax "op &"} $ eq $ b;
+ val eq = Syntax.const @{const_syntax HOL.eq} $ Bound (nvars idts) $ e;
+ val P = Syntax.const @{const_syntax HOL.conj} $ eq $ b;
val exP = ex_tr [idts, P];
in Syntax.const @{const_syntax Collect} $ Term.absdummy (dummyT, exP) end;
@@ -288,8 +288,8 @@
fun setcompr_tr' [Abs (abs as (_, _, P))] =
let
fun check (Const (@{const_syntax Ex}, _) $ Abs (_, _, P), n) = check (P, n + 1)
- | check (Const (@{const_syntax "op &"}, _) $
- (Const (@{const_syntax "op ="}, _) $ Bound m $ e) $ P, n) =
+ | check (Const (@{const_syntax HOL.conj}, _) $
+ (Const (@{const_syntax HOL.eq}, _) $ Bound m $ e) $ P, n) =
n > 0 andalso m = n andalso not (loose_bvar1 (P, n)) andalso
subset (op =) (0 upto (n - 1), add_loose_bnos (e, 0, []))
| check _ = false;
@@ -305,7 +305,7 @@
val M = Syntax.const @{syntax_const "_Coll"} $ x $ t;
in
case t of
- Const (@{const_syntax "op &"}, _) $
+ Const (@{const_syntax HOL.conj}, _) $
(Const (@{const_syntax Set.member}, _) $
(Const (@{syntax_const "_bound"}, _) $ Free (yN, _)) $ A) $ P =>
if xN = yN then Syntax.const @{syntax_const "_Collect"} $ x $ A $ P else M
--- a/src/HOL/Statespace/DistinctTreeProver.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Statespace/DistinctTreeProver.thy Wed Sep 01 17:19:47 2010 +0200
@@ -32,17 +32,18 @@
subsection {* Distinctness of Nodes *}
-consts set_of:: "'a tree \<Rightarrow> 'a set"
-primrec
-"set_of Tip = {}"
-"set_of (Node l x d r) = (if d then {} else {x}) \<union> set_of l \<union> set_of r"
+primrec set_of :: "'a tree \<Rightarrow> 'a set"
+where
+ "set_of Tip = {}"
+| "set_of (Node l x d r) = (if d then {} else {x}) \<union> set_of l \<union> set_of r"
-consts all_distinct:: "'a tree \<Rightarrow> bool"
-primrec
-"all_distinct Tip = True"
-"all_distinct (Node l x d r) = ((d \<or> (x \<notin> set_of l \<and> x \<notin> set_of r)) \<and>
- set_of l \<inter> set_of r = {} \<and>
- all_distinct l \<and> all_distinct r)"
+primrec all_distinct :: "'a tree \<Rightarrow> bool"
+where
+ "all_distinct Tip = True"
+| "all_distinct (Node l x d r) =
+ ((d \<or> (x \<notin> set_of l \<and> x \<notin> set_of r)) \<and>
+ set_of l \<inter> set_of r = {} \<and>
+ all_distinct l \<and> all_distinct r)"
text {* Given a binary tree @{term "t"} for which
@{const all_distinct} holds, given two different nodes contained in the tree,
@@ -99,19 +100,19 @@
*}
-consts "delete" :: "'a \<Rightarrow> 'a tree \<Rightarrow> 'a tree option"
-primrec
-"delete x Tip = None"
-"delete x (Node l y d r) = (case delete x l of
- Some l' \<Rightarrow>
- (case delete x r of
- Some r' \<Rightarrow> Some (Node l' y (d \<or> (x=y)) r')
- | None \<Rightarrow> Some (Node l' y (d \<or> (x=y)) r))
- | None \<Rightarrow>
- (case (delete x r) of
- Some r' \<Rightarrow> Some (Node l y (d \<or> (x=y)) r')
- | None \<Rightarrow> if x=y \<and> \<not>d then Some (Node l y True r)
- else None))"
+primrec delete :: "'a \<Rightarrow> 'a tree \<Rightarrow> 'a tree option"
+where
+ "delete x Tip = None"
+| "delete x (Node l y d r) = (case delete x l of
+ Some l' \<Rightarrow>
+ (case delete x r of
+ Some r' \<Rightarrow> Some (Node l' y (d \<or> (x=y)) r')
+ | None \<Rightarrow> Some (Node l' y (d \<or> (x=y)) r))
+ | None \<Rightarrow>
+ (case (delete x r) of
+ Some r' \<Rightarrow> Some (Node l y (d \<or> (x=y)) r')
+ | None \<Rightarrow> if x=y \<and> \<not>d then Some (Node l y True r)
+ else None))"
lemma delete_Some_set_of: "\<And>t'. delete x t = Some t' \<Longrightarrow> set_of t' \<subseteq> set_of t"
@@ -293,15 +294,15 @@
qed
-consts "subtract" :: "'a tree \<Rightarrow> 'a tree \<Rightarrow> 'a tree option"
-primrec
-"subtract Tip t = Some t"
-"subtract (Node l x b r) t =
- (case delete x t of
- Some t' \<Rightarrow> (case subtract l t' of
- Some t'' \<Rightarrow> subtract r t''
- | None \<Rightarrow> None)
- | None \<Rightarrow> None)"
+primrec subtract :: "'a tree \<Rightarrow> 'a tree \<Rightarrow> 'a tree option"
+where
+ "subtract Tip t = Some t"
+| "subtract (Node l x b r) t =
+ (case delete x t of
+ Some t' \<Rightarrow> (case subtract l t' of
+ Some t'' \<Rightarrow> subtract r t''
+ | None \<Rightarrow> None)
+ | None \<Rightarrow> None)"
lemma subtract_Some_set_of_res:
"\<And>t\<^isub>2 t. subtract t\<^isub>1 t\<^isub>2 = Some t \<Longrightarrow> set_of t \<subseteq> set_of t\<^isub>2"
--- a/src/HOL/Statespace/StateFun.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Statespace/StateFun.thy Wed Sep 01 17:19:47 2010 +0200
@@ -33,10 +33,10 @@
lemma K_statefun_cong [cong]: "K_statefun c x = K_statefun c x"
by (rule refl)
-definition lookup:: "('v \<Rightarrow> 'a) \<Rightarrow> 'n \<Rightarrow> ('n \<Rightarrow> 'v) \<Rightarrow> 'a"
+definition lookup :: "('v \<Rightarrow> 'a) \<Rightarrow> 'n \<Rightarrow> ('n \<Rightarrow> 'v) \<Rightarrow> 'a"
where "lookup destr n s = destr (s n)"
-definition update::
+definition update ::
"('v \<Rightarrow> 'a1) \<Rightarrow> ('a2 \<Rightarrow> 'v) \<Rightarrow> 'n \<Rightarrow> ('a1 \<Rightarrow> 'a2) \<Rightarrow> ('n \<Rightarrow> 'v) \<Rightarrow> ('n \<Rightarrow> 'v)"
where "update destr constr n f s = s(n := constr (f (destr (s n))))"
--- a/src/HOL/Statespace/StateSpaceEx.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Statespace/StateSpaceEx.thy Wed Sep 01 17:19:47 2010 +0200
@@ -1,14 +1,12 @@
(* Title: StateSpaceEx.thy
- ID: $Id$
Author: Norbert Schirmer, TU Muenchen
*)
header {* Examples \label{sec:Examples} *}
theory StateSpaceEx
imports StateSpaceLocale StateSpaceSyntax
+begin
-begin
-(* FIXME: Use proper keywords file *)
(*<*)
syntax
"_statespace_updates" :: "('a \<Rightarrow> 'b) \<Rightarrow> updbinds \<Rightarrow> ('a \<Rightarrow> 'b)" ("_\<langle>_\<rangle>" [900,0] 900)
--- a/src/HOL/Statespace/StateSpaceLocale.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Statespace/StateSpaceLocale.thy Wed Sep 01 17:19:47 2010 +0200
@@ -1,5 +1,4 @@
(* Title: StateSpaceLocale.thy
- ID: $Id$
Author: Norbert Schirmer, TU Muenchen
*)
@@ -18,7 +17,7 @@
locale project_inject =
fixes project :: "'value \<Rightarrow> 'a"
- and "inject":: "'a \<Rightarrow> 'value"
+ and inject :: "'a \<Rightarrow> 'value"
assumes project_inject_cancel [statefun_simp]: "project (inject x) = x"
lemma (in project_inject)
--- a/src/HOL/Statespace/StateSpaceSyntax.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Statespace/StateSpaceSyntax.thy Wed Sep 01 17:19:47 2010 +0200
@@ -5,7 +5,6 @@
header {* Syntax for State Space Lookup and Update \label{sec:StateSpaceSyntax}*}
theory StateSpaceSyntax
imports StateSpaceLocale
-
begin
text {* The state space syntax is kept in an extra theory so that you
--- a/src/HOL/Statespace/distinct_tree_prover.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Statespace/distinct_tree_prover.ML Wed Sep 01 17:19:47 2010 +0200
@@ -343,7 +343,7 @@
end handle Option => NONE)
fun distinctTree_tac names ctxt
- (Const (@{const_name Trueprop},_) $ (Const (@{const_name Not}, _) $ (Const (@{const_name "op ="}, _) $ x $ y)), i) =
+ (Const (@{const_name Trueprop},_) $ (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.eq}, _) $ x $ y)), i) =
(case get_fst_success (neq_x_y ctxt x y) names of
SOME neq => rtac neq i
| NONE => no_tac)
@@ -356,7 +356,7 @@
fun distinct_simproc names =
Simplifier.simproc_global @{theory HOL} "DistinctTreeProver.distinct_simproc" ["x = y"]
- (fn thy => fn ss => fn (Const (@{const_name "op ="},_)$x$y) =>
+ (fn thy => fn ss => fn (Const (@{const_name HOL.eq},_)$x$y) =>
case try Simplifier.the_context ss of
SOME ctxt => Option.map (fn neq => neq_to_eq_False OF [neq])
(get_fst_success (neq_x_y ctxt x y) names)
--- a/src/HOL/Statespace/state_fun.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Statespace/state_fun.ML Wed Sep 01 17:19:47 2010 +0200
@@ -53,7 +53,7 @@
val lazy_conj_simproc =
Simplifier.simproc_global @{theory HOL} "lazy_conj_simp" ["P & Q"]
(fn thy => fn ss => fn t =>
- (case t of (Const (@{const_name "op &"},_)$P$Q) =>
+ (case t of (Const (@{const_name HOL.conj},_)$P$Q) =>
let
val P_P' = Simplifier.rewrite ss (cterm_of thy P);
val P' = P_P' |> prop_of |> Logic.dest_equals |> #2
@@ -285,7 +285,7 @@
then Bound 2
else raise TERM ("",[n]);
val sel' = lo $ d $ n' $ s;
- in (Const (@{const_name "op ="},Teq)$sel'$x',hd (binder_types Teq),nT,swap) end;
+ in (Const (@{const_name HOL.eq},Teq)$sel'$x',hd (binder_types Teq),nT,swap) end;
fun dest_state (s as Bound 0) = s
| dest_state (s as (Const (sel,sT)$Bound 0)) =
@@ -295,10 +295,10 @@
| dest_state s =
raise TERM ("StateFun.ex_lookup_eq_simproc: not a record slector",[s]);
- fun dest_sel_eq (Const (@{const_name "op ="},Teq)$
+ fun dest_sel_eq (Const (@{const_name HOL.eq},Teq)$
((lo as (Const ("StateFun.lookup",lT)))$d$n$s)$X) =
(false,Teq,lT,lo,d,n,X,dest_state s)
- | dest_sel_eq (Const (@{const_name "op ="},Teq)$X$
+ | dest_sel_eq (Const (@{const_name HOL.eq},Teq)$X$
((lo as (Const ("StateFun.lookup",lT)))$d$n$s)) =
(true,Teq,lT,lo,d,n,X,dest_state s)
| dest_sel_eq _ = raise TERM ("",[]);
--- a/src/HOL/Statespace/state_space.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Statespace/state_space.ML Wed Sep 01 17:19:47 2010 +0200
@@ -223,7 +223,7 @@
fun distinctTree_tac ctxt
(Const (@{const_name Trueprop},_) $
- (Const (@{const_name Not}, _) $ (Const (@{const_name "op ="}, _) $ (x as Free _)$ (y as Free _))), i) =
+ (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.eq}, _) $ (x as Free _)$ (y as Free _))), i) =
(case (neq_x_y ctxt x y) of
SOME neq => rtac neq i
| NONE => no_tac)
@@ -236,7 +236,7 @@
val distinct_simproc =
Simplifier.simproc_global @{theory HOL} "StateSpace.distinct_simproc" ["x = y"]
- (fn thy => fn ss => (fn (Const (@{const_name "op ="},_)$(x as Free _)$(y as Free _)) =>
+ (fn thy => fn ss => (fn (Const (@{const_name HOL.eq},_)$(x as Free _)$(y as Free _)) =>
(case try Simplifier.the_context ss of
SOME ctxt => Option.map (fn neq => DistinctTreeProver.neq_to_eq_False OF [neq])
(neq_x_y ctxt x y)
@@ -277,28 +277,29 @@
fun comps_of_thm thm = prop_of thm
|> (fn (_$(_$t)) => DistinctTreeProver.dest_tree t) |> map (fst o dest_Free);
- fun type_attr phi (ctxt,thm) =
- (case ctxt of Context.Theory _ => (ctxt,thm)
- | _ =>
+ fun type_attr phi = Thm.declaration_attribute (fn thm => fn context =>
+ (case context of
+ Context.Theory _ => context
+ | Context.Proof ctxt =>
let
- val {declinfo,distinctthm=tt,silent} = (NameSpaceData.get ctxt);
+ val {declinfo,distinctthm=tt,silent} = NameSpaceData.get context;
val all_names = comps_of_thm thm;
fun upd name tt =
- (case (Symtab.lookup tt name) of
+ (case Symtab.lookup tt name of
SOME dthm => if sorted_subset (op =) (comps_of_thm dthm) all_names
then Symtab.update (name,thm) tt else tt
- | NONE => Symtab.update (name,thm) tt)
+ | NONE => Symtab.update (name,thm) tt)
val tt' = tt |> fold upd all_names;
val activate_simproc =
- Output.no_warnings_CRITICAL (* FIXME !?! *)
- (Simplifier.map_ss (fn ss => ss addsimprocs [distinct_simproc]));
- val ctxt' =
- ctxt
+ Simplifier.map_ss
+ (Simplifier.with_context (Context_Position.set_visible false ctxt)
+ (fn ss => ss addsimprocs [distinct_simproc]));
+ val context' =
+ context
|> NameSpaceData.put {declinfo=declinfo,distinctthm=tt',silent=silent}
- |> activate_simproc
- in (ctxt',thm)
- end)
+ |> activate_simproc;
+ in context' end));
val attr = Attrib.internal type_attr;
--- a/src/HOL/String.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/String.thy Wed Sep 01 17:19:47 2010 +0200
@@ -53,7 +53,7 @@
(fn n => fn m => Drule.instantiate' [] [SOME n, SOME m] @{thm nibble_pair_of_char.simps})
nibbles nibbles;
in
- PureThy.note_thmss Thm.definitionK [((Binding.name "nibble_pair_of_char_simps", []), [(thms, [])])]
+ PureThy.note_thmss Thm.definitionK [((@{binding nibble_pair_of_char_simps}, []), [(thms, [])])]
#-> (fn [(_, thms)] => fold_rev Code.add_eqn thms)
end
*}
@@ -183,10 +183,10 @@
fold String_Code.add_literal_string ["SML", "OCaml", "Haskell", "Scala"]
*}
-code_instance literal :: eq
+code_instance literal :: equal
(Haskell -)
-code_const "eq_class.eq \<Colon> literal \<Rightarrow> literal \<Rightarrow> bool"
+code_const "HOL.equal \<Colon> literal \<Rightarrow> literal \<Rightarrow> bool"
(SML "!((_ : string) = _)")
(OCaml "!((_ : string) = _)")
(Haskell infixl 4 "==")
--- a/src/HOL/TLA/Intensional.thy Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/TLA/Intensional.thy Wed Sep 01 17:19:47 2010 +0200
@@ -279,7 +279,7 @@
fun hflatten t =
case (concl_of t) of
- Const _ $ (Const (@{const_name "op -->"}, _) $ _ $ _) => hflatten (t RS mp)
+ Const _ $ (Const (@{const_name HOL.implies}, _) $ _ $ _) => hflatten (t RS mp)
| _ => (hflatten (matchsome conjI t)) handle THM _ => zero_var_indexes t
in
hflatten t
--- a/src/HOL/Tools/ATP/atp_systems.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/ATP/atp_systems.ML Wed Sep 01 17:19:47 2010 +0200
@@ -19,7 +19,7 @@
has_incomplete_mode: bool,
proof_delims: (string * string) list,
known_failures: (failure * string) list,
- default_max_relevant_per_iter: int,
+ default_max_relevant: int,
default_theory_relevant: bool,
explicit_forall: bool,
use_conjecture_for_hypotheses: bool}
@@ -52,7 +52,7 @@
has_incomplete_mode: bool,
proof_delims: (string * string) list,
known_failures: (failure * string) list,
- default_max_relevant_per_iter: int,
+ default_max_relevant: int,
default_theory_relevant: bool,
explicit_forall: bool,
use_conjecture_for_hypotheses: bool}
@@ -125,10 +125,20 @@
priority ("Available ATPs: " ^
commas (sort_strings (Symtab.keys (Data.get thy))) ^ ".")
-fun to_generous_secs time = (Time.toMilliseconds time + 999) div 1000
+fun to_secs bonus time = (Time.toMilliseconds time + bonus + 999) div 1000
(* E prover *)
+(* 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
+ itself, which is in fact much lower. *)
+fun e_bonus () =
+ case getenv "E_VERSION" of
+ "" => 1000
+ | version =>
+ if exists (fn s => String.isPrefix s version) ["0.9", "1.0"] then 1000
+ else 0
+
val tstp_proof_delims =
("# SZS output start CNFRefutation.", "# SZS output end CNFRefutation")
@@ -137,8 +147,7 @@
required_execs = [],
arguments = fn _ => fn timeout =>
"--tstp-in --tstp-out -l5 -xAutoDev -tAutoDev --silent \
- \--soft-cpu-limit=" ^
- string_of_int (to_generous_secs timeout),
+ \--cpu-limit=" ^ string_of_int (to_secs (e_bonus ()) timeout),
has_incomplete_mode = false,
proof_delims = [tstp_proof_delims],
known_failures =
@@ -150,7 +159,7 @@
"# Cannot determine problem status within resource limit"),
(OutOfResources, "SZS status: ResourceOut"),
(OutOfResources, "SZS status ResourceOut")],
- default_max_relevant_per_iter = 80 (* FUDGE *),
+ default_max_relevant = 500 (* FUDGE *),
default_theory_relevant = false,
explicit_forall = false,
use_conjecture_for_hypotheses = true}
@@ -165,7 +174,7 @@
required_execs = [("SPASS_HOME", "SPASS")],
arguments = fn complete => fn timeout =>
("-Auto -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof \
- \-VarWeight=3 -TimeLimit=" ^ string_of_int (to_generous_secs timeout))
+ \-VarWeight=3 -TimeLimit=" ^ string_of_int (to_secs 0 timeout))
|> not complete ? prefix "-SOS=1 ",
has_incomplete_mode = true,
proof_delims = [("Here is a proof", "Formulae used in the proof")],
@@ -177,7 +186,7 @@
(MalformedInput, "Undefined symbol"),
(MalformedInput, "Free Variable"),
(SpassTooOld, "tptp2dfg")],
- default_max_relevant_per_iter = 40 (* FUDGE *),
+ default_max_relevant = 350 (* FUDGE *),
default_theory_relevant = true,
explicit_forall = true,
use_conjecture_for_hypotheses = true}
@@ -190,10 +199,11 @@
val vampire_config : prover_config =
{exec = ("VAMPIRE_HOME", "vampire"),
required_execs = [],
- arguments = fn _ => fn timeout =>
- "--mode casc -t " ^ string_of_int (to_generous_secs timeout) ^
- " --thanks Andrei --input_file",
- has_incomplete_mode = false,
+ arguments = fn complete => fn timeout =>
+ ("--mode casc -t " ^ string_of_int (to_secs 0 timeout) ^
+ " --thanks Andrei --input_file")
+ |> not complete ? prefix "--sos on ",
+ has_incomplete_mode = true,
proof_delims =
[("=========== Refutation ==========",
"======= End of refutation ======="),
@@ -206,7 +216,7 @@
(Unprovable, "Satisfiability detected"),
(Unprovable, "Termination reason: Satisfiable"),
(VampireTooOld, "not a valid option")],
- default_max_relevant_per_iter = 45 (* FUDGE *),
+ default_max_relevant = 400 (* FUDGE *),
default_theory_relevant = false,
explicit_forall = false,
use_conjecture_for_hypotheses = true}
@@ -246,50 +256,50 @@
| SOME sys => sys
fun remote_config system_name system_versions proof_delims known_failures
- default_max_relevant_per_iter default_theory_relevant
- use_conjecture_for_hypotheses =
+ default_max_relevant default_theory_relevant
+ use_conjecture_for_hypotheses : prover_config =
{exec = ("ISABELLE_ATP", "scripts/remote_atp"),
required_execs = [],
arguments = fn _ => fn timeout =>
- " -t " ^ string_of_int (to_generous_secs timeout) ^ " -s " ^
+ " -t " ^ string_of_int (to_secs 0 timeout) ^ " -s " ^
the_system system_name system_versions,
has_incomplete_mode = false,
proof_delims = insert (op =) tstp_proof_delims proof_delims,
known_failures =
known_failures @ known_perl_failures @
[(TimedOut, "says Timeout")],
- default_max_relevant_per_iter = default_max_relevant_per_iter,
+ default_max_relevant = default_max_relevant,
default_theory_relevant = default_theory_relevant,
explicit_forall = true,
use_conjecture_for_hypotheses = use_conjecture_for_hypotheses}
fun remotify_config system_name system_versions
- ({proof_delims, known_failures, default_max_relevant_per_iter,
+ ({proof_delims, known_failures, default_max_relevant,
default_theory_relevant, use_conjecture_for_hypotheses, ...}
: prover_config) : prover_config =
remote_config system_name system_versions proof_delims known_failures
- default_max_relevant_per_iter default_theory_relevant
+ default_max_relevant default_theory_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_per_iter default_theory_relevant
+ default_max_relevant default_theory_relevant
use_conjecture_for_hypotheses =
(remotify_name name,
remote_config system_name system_versions proof_delims known_failures
- default_max_relevant_per_iter default_theory_relevant
+ default_max_relevant default_theory_relevant
use_conjecture_for_hypotheses)
fun remotify_prover (name, config) system_name system_versions =
(remotify_name 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" ["9.9", "0.6", "1.0"]
+val remote_vampire = remotify_prover vampire "Vampire" ["9.0", "1.0", "0.6"]
val remote_sine_e =
- remote_prover "sine_e" "SInE" [] []
- [(Unprovable, "says Unknown")] 150 (* FUDGE *) false true
+ remote_prover "sine_e" "SInE" [] [] [(Unprovable, "says Unknown")]
+ 1000 (* FUDGE *) false true
val remote_snark =
remote_prover "snark" "SNARK---" [] [("refutation.", "end_refutation.")] []
- 50 (* FUDGE *) false true
+ 350 (* FUDGE *) false true
(* Setup *)
--- a/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Wed Sep 01 17:19:47 2010 +0200
@@ -416,7 +416,7 @@
fun prove_case_cong ((t, nchotomy), case_rewrites) =
let
val (Const ("==>", _) $ tm $ _) = t;
- val (Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op ="}, _) $ _ $ Ma)) = tm;
+ val (Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ Ma)) = tm;
val cert = cterm_of thy;
val nchotomy' = nchotomy RS spec;
val [v] = Term.add_vars (concl_of nchotomy') [];
--- a/src/HOL/Tools/Datatype/datatype_aux.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_aux.ML Wed Sep 01 17:19:47 2010 +0200
@@ -120,8 +120,8 @@
fun split_conj_thm th =
((th RS conjunct1)::(split_conj_thm (th RS conjunct2))) handle THM _ => [th];
-val mk_conj = foldr1 (HOLogic.mk_binop @{const_name "op &"});
-val mk_disj = foldr1 (HOLogic.mk_binop @{const_name "op |"});
+val mk_conj = foldr1 (HOLogic.mk_binop @{const_name HOL.conj});
+val mk_disj = foldr1 (HOLogic.mk_binop @{const_name HOL.disj});
fun app_bnds t i = list_comb (t, map Bound (i - 1 downto 0));
--- a/src/HOL/Tools/Datatype/datatype_codegen.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_codegen.ML Wed Sep 01 17:19:47 2010 +0200
@@ -68,7 +68,7 @@
val { descr, index, inject = inject_thms, distinct = distinct_thms, ... } =
Datatype_Data.the_info thy tyco;
val ty = Type (tyco, map TFree vs);
- fun mk_eq (t1, t2) = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
+ fun mk_eq (t1, t2) = Const (@{const_name HOL.equal}, ty --> ty --> HOLogic.boolT)
$ t1 $ t2;
fun true_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.true_const);
fun false_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.false_const);
@@ -83,7 +83,7 @@
val distincts = maps prep_distinct (snd (nth (Datatype_Prop.make_distincts [descr] vs) index));
val refl = HOLogic.mk_Trueprop (true_eq (Free ("x", ty), Free ("x", ty)));
val simpset = Simplifier.global_context thy (HOL_basic_ss addsimps
- (map Simpdata.mk_eq (@{thm eq} :: @{thm eq_True} :: inject_thms @ distinct_thms)));
+ (map Simpdata.mk_eq (@{thm equal} :: @{thm eq_True} :: inject_thms @ distinct_thms)));
fun prove prop = Skip_Proof.prove_global thy [] [] prop (K (ALLGOALS (simp_tac simpset)))
|> Simpdata.mk_eq;
in (map prove (triv_injects @ injects @ distincts), prove refl) end;
@@ -96,7 +96,7 @@
fun mk_side const_name = Const (const_name, ty --> ty --> HOLogic.boolT)
$ Free ("x", ty) $ Free ("y", ty);
val def = HOLogic.mk_Trueprop (HOLogic.mk_eq
- (mk_side @{const_name eq_class.eq}, mk_side @{const_name "op ="}));
+ (mk_side @{const_name HOL.equal}, mk_side @{const_name HOL.eq}));
val def' = Syntax.check_term lthy def;
val ((_, (_, thm)), lthy') = Specification.definition
(NONE, (Attrib.empty_binding, def')) lthy;
@@ -115,7 +115,7 @@
#> snd
in
thy
- |> Class.instantiation (tycos, vs, [HOLogic.class_eq])
+ |> Class.instantiation (tycos, vs, [HOLogic.class_equal])
|> fold_map add_def tycos
|-> (fn def_thms => Class.prove_instantiation_exit_result (map o Morphism.thm)
(fn _ => fn def_thms => tac def_thms) def_thms)
@@ -135,7 +135,7 @@
val case_rewrites = maps (#case_rewrites o Datatype_Data.the_info thy) tycos;
val certs = map (mk_case_cert thy) tycos;
val tycos_eq = filter_out
- (fn tyco => can (Sorts.mg_domain (Sign.classes_of thy) tyco) [HOLogic.class_eq]) tycos;
+ (fn tyco => can (Sorts.mg_domain (Sign.classes_of thy) tyco) [HOLogic.class_equal]) tycos;
in
if null css then thy
else thy
--- a/src/HOL/Tools/Datatype/datatype_data.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_data.ML Wed Sep 01 17:19:47 2010 +0200
@@ -257,7 +257,9 @@
Pretty.str " =" :: Pretty.brk 1 ::
flat (separate [Pretty.brk 1, Pretty.str "| "]
(map (single o pretty_constr) cos)));
- in Thy_Output.output (Thy_Output.maybe_pretty_source (K pretty_datatype) src [()]) end);
+ in
+ Thy_Output.output ctxt (Thy_Output.maybe_pretty_source (K (K pretty_datatype)) ctxt src [()])
+ end);
@@ -428,7 +430,7 @@
unflat rules (map Drule.zero_var_indexes_list raw_thms);
(*FIXME somehow dubious*)
in
- ProofContext.theory_result
+ ProofContext.background_theory_result
(prove_rep_datatype config dt_names alt_names descr vs
raw_inject half_distinct raw_induct)
#-> after_qed
--- a/src/HOL/Tools/Datatype/datatype_prop.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_prop.ML Wed Sep 01 17:19:47 2010 +0200
@@ -70,7 +70,7 @@
val frees' = map Free ((map ((op ^) o (rpair "'")) tnames) ~~ Ts);
in cons (HOLogic.mk_Trueprop (HOLogic.mk_eq
(HOLogic.mk_eq (list_comb (constr_t, frees), list_comb (constr_t, frees')),
- foldr1 (HOLogic.mk_binop @{const_name "op &"})
+ foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
(map HOLogic.mk_eq (frees ~~ frees')))))
end;
in
@@ -149,7 +149,7 @@
val prems = maps (fn ((i, (_, _, constrs)), T) =>
map (make_ind_prem i T) constrs) (descr' ~~ recTs);
val tnames = make_tnames recTs;
- val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name "op &"})
+ val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
(map (fn (((i, _), T), tname) => make_pred i T $ Free (tname, T))
(descr' ~~ recTs ~~ tnames)))
--- a/src/HOL/Tools/Datatype/datatype_realizer.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_realizer.ML Wed Sep 01 17:19:47 2010 +0200
@@ -99,7 +99,7 @@
if member (op =) is i then SOME
(list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Free (tname, T))
else NONE) (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names ~~ tnames));
- val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name "op &"})
+ val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop @{const_name HOL.conj})
(map (fn ((((i, _), T), U), tname) =>
make_pred i U T (mk_proj i is r) (Free (tname, T)))
(descr ~~ recTs ~~ rec_result_Ts ~~ tnames)));
--- a/src/HOL/Tools/Function/function_common.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Function/function_common.ML Wed Sep 01 17:19:47 2010 +0200
@@ -164,7 +164,7 @@
structure Termination_Simps = Named_Thms
(
val name = "termination_simp"
- val description = "Simplification rule for termination proofs"
+ val description = "simplification rules for termination proofs"
)
@@ -175,7 +175,7 @@
type T = Proof.context -> tactic
val empty = (fn _ => error "Termination prover not configured")
val extend = I
- fun merge (a, b) = b (* FIXME ? *)
+ fun merge (a, _) = a
)
val set_termination_prover = TerminationProver.put
--- a/src/HOL/Tools/Function/function_core.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Function/function_core.ML Wed Sep 01 17:19:47 2010 +0200
@@ -860,9 +860,9 @@
rtac (instantiate' [] [SOME (cterm_of thy lhs_acc)] case_split) 1
THEN (rtac (Thm.forall_elim_vars 0 psimp) THEN_ALL_NEW assume_tac) 1
THEN (simp_default_tac (simpset_of ctxt) 1)
- THEN (etac not_acc_down 1)
- THEN ((etac R_cases)
- THEN_ALL_NEW (simp_default_tac (simpset_of ctxt))) 1)
+ THEN TRY ((etac not_acc_down 1)
+ THEN ((etac R_cases)
+ THEN_ALL_NEW (simp_default_tac (simpset_of ctxt))) 1))
|> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
end
in
--- a/src/HOL/Tools/Function/measure_functions.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Function/measure_functions.ML Wed Sep 01 17:19:47 2010 +0200
@@ -20,7 +20,7 @@
(
val name = "measure_function"
val description =
- "Rules that guide the heuristic generation of measure functions"
+ "rules that guide the heuristic generation of measure functions"
);
fun mk_is_measure t =
--- a/src/HOL/Tools/Function/scnp_reconstruct.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Function/scnp_reconstruct.ML Wed Sep 01 17:19:47 2010 +0200
@@ -68,7 +68,7 @@
type T = multiset_setup option
val empty = NONE
val extend = I;
- fun merge (v1, v2) = if is_some v2 then v2 else v1 (* FIXME prefer v1 !?! *)
+ fun merge (v1, v2) = if is_some v1 then v1 else v2
)
val multiset_setup = Multiset_Setup.put o SOME
--- a/src/HOL/Tools/Function/termination.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Function/termination.ML Wed Sep 01 17:19:47 2010 +0200
@@ -148,7 +148,7 @@
val cs = Function_Lib.dest_binop_list @{const_name Lattices.sup} rel
fun collect_pats (Const (@{const_name Collect}, _) $ Abs (_, _, c)) =
let
- val (Const (@{const_name "op &"}, _) $ (Const (@{const_name "op ="}, _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ _)
+ val (Const (@{const_name HOL.conj}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ _)
= Term.strip_qnt_body @{const_name Ex} c
in cons r o cons l end
in
@@ -185,7 +185,7 @@
val vs = Term.strip_qnt_vars @{const_name Ex} c
(* FIXME: throw error "dest_call" for malformed terms *)
- val (Const (@{const_name "op &"}, _) $ (Const (@{const_name "op ="}, _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ Gam)
+ val (Const (@{const_name HOL.conj}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ Gam)
= Term.strip_qnt_body @{const_name Ex} c
val (p, l') = dest_inj sk l
val (q, r') = dest_inj sk r
--- a/src/HOL/Tools/Nitpick/minipick.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Nitpick/minipick.ML Wed Sep 01 17:19:47 2010 +0200
@@ -123,16 +123,16 @@
Exist (decls_for SRep card Ts T, to_F (T :: Ts) t')
| (t0 as Const (@{const_name Ex}, _)) $ t1 =>
to_F Ts (t0 $ eta_expand Ts t1 1)
- | Const (@{const_name "op ="}, _) $ t1 $ t2 =>
+ | Const (@{const_name HOL.eq}, _) $ t1 $ t2 =>
RelEq (to_R_rep Ts t1, to_R_rep Ts t2)
| Const (@{const_name ord_class.less_eq},
Type (@{type_name fun},
[Type (@{type_name fun}, [_, @{typ bool}]), _]))
$ t1 $ t2 =>
Subset (to_R_rep Ts t1, to_R_rep Ts t2)
- | @{const "op &"} $ t1 $ t2 => And (to_F Ts t1, to_F Ts t2)
- | @{const "op |"} $ t1 $ t2 => Or (to_F Ts t1, to_F Ts t2)
- | @{const "op -->"} $ t1 $ t2 => Implies (to_F Ts t1, to_F Ts t2)
+ | @{const HOL.conj} $ t1 $ t2 => And (to_F Ts t1, to_F Ts t2)
+ | @{const HOL.disj} $ t1 $ t2 => Or (to_F Ts t1, to_F Ts t2)
+ | @{const HOL.implies} $ t1 $ t2 => Implies (to_F Ts t1, to_F Ts t2)
| t1 $ t2 => Subset (to_S_rep Ts t2, to_R_rep Ts t1)
| Free _ => raise SAME ()
| Term.Var _ => raise SAME ()
@@ -165,20 +165,20 @@
@{const Not} => to_R_rep Ts (eta_expand Ts t 1)
| Const (@{const_name All}, _) => to_R_rep Ts (eta_expand Ts t 1)
| Const (@{const_name Ex}, _) => to_R_rep Ts (eta_expand Ts t 1)
- | Const (@{const_name "op ="}, _) $ _ => to_R_rep Ts (eta_expand Ts t 1)
- | Const (@{const_name "op ="}, _) => to_R_rep Ts (eta_expand Ts t 2)
+ | Const (@{const_name HOL.eq}, _) $ _ => to_R_rep Ts (eta_expand Ts t 1)
+ | Const (@{const_name HOL.eq}, _) => to_R_rep Ts (eta_expand Ts t 2)
| Const (@{const_name ord_class.less_eq},
Type (@{type_name fun},
[Type (@{type_name fun}, [_, @{typ bool}]), _])) $ _ =>
to_R_rep Ts (eta_expand Ts t 1)
| Const (@{const_name ord_class.less_eq}, _) =>
to_R_rep Ts (eta_expand Ts t 2)
- | @{const "op &"} $ _ => to_R_rep Ts (eta_expand Ts t 1)
- | @{const "op &"} => to_R_rep Ts (eta_expand Ts t 2)
- | @{const "op |"} $ _ => to_R_rep Ts (eta_expand Ts t 1)
- | @{const "op |"} => to_R_rep Ts (eta_expand Ts t 2)
- | @{const "op -->"} $ _ => to_R_rep Ts (eta_expand Ts t 1)
- | @{const "op -->"} => to_R_rep Ts (eta_expand Ts t 2)
+ | @{const HOL.conj} $ _ => to_R_rep Ts (eta_expand Ts t 1)
+ | @{const HOL.conj} => to_R_rep Ts (eta_expand Ts t 2)
+ | @{const HOL.disj} $ _ => to_R_rep Ts (eta_expand Ts t 1)
+ | @{const HOL.disj} => to_R_rep Ts (eta_expand Ts t 2)
+ | @{const HOL.implies} $ _ => to_R_rep Ts (eta_expand Ts t 1)
+ | @{const HOL.implies} => to_R_rep Ts (eta_expand Ts t 2)
| Const (@{const_name bot_class.bot},
T as Type (@{type_name fun}, [_, @{typ bool}])) =>
empty_n_ary_rel (arity_of RRep card T)
--- a/src/HOL/Tools/Nitpick/nitpick.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick.ML Wed Sep 01 17:19:47 2010 +0200
@@ -182,7 +182,7 @@
fun none_true assigns = forall (not_equal (SOME true) o snd) assigns
val syntactic_sorts =
- @{sort "{default,zero,one,plus,minus,uminus,times,inverse,abs,sgn,ord,eq}"} @
+ @{sort "{default,zero,one,plus,minus,uminus,times,inverse,abs,sgn,ord,equal}"} @
@{sort number}
fun has_tfree_syntactic_sort (TFree (_, S as _ :: _)) =
subset (op =) (S, syntactic_sorts)
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Wed Sep 01 17:19:47 2010 +0200
@@ -386,13 +386,13 @@
if t0 = conn_t then strip_connective t0 t2 @ strip_connective t0 t1 else [t]
| strip_connective _ t = [t]
fun strip_any_connective (t as (t0 $ _ $ _)) =
- if t0 = @{const "op &"} orelse t0 = @{const "op |"} then
+ if t0 = @{const HOL.conj} orelse t0 = @{const HOL.disj} then
(strip_connective t0 t, t0)
else
([t], @{const Not})
| strip_any_connective t = ([t], @{const Not})
-val conjuncts_of = strip_connective @{const "op &"}
-val disjuncts_of = strip_connective @{const "op |"}
+val conjuncts_of = strip_connective @{const HOL.conj}
+val disjuncts_of = strip_connective @{const HOL.disj}
(* When you add constants to these lists, make sure to handle them in
"Nitpick_Nut.nut_from_term", and perhaps in "Nitpick_Mono.consider_term" as
@@ -408,10 +408,10 @@
(@{const_name True}, 0),
(@{const_name All}, 1),
(@{const_name Ex}, 1),
- (@{const_name "op ="}, 1),
- (@{const_name "op &"}, 2),
- (@{const_name "op |"}, 2),
- (@{const_name "op -->"}, 2),
+ (@{const_name HOL.eq}, 1),
+ (@{const_name HOL.conj}, 2),
+ (@{const_name HOL.disj}, 2),
+ (@{const_name HOL.implies}, 2),
(@{const_name If}, 3),
(@{const_name Let}, 2),
(@{const_name Pair}, 2),
@@ -1275,7 +1275,7 @@
forall is_Var args andalso not (has_duplicates (op =) args)
| _ => false
fun do_eq (Const (@{const_name "=="}, _) $ t1 $ _) = do_lhs t1
- | do_eq (@{const Trueprop} $ (Const (@{const_name "op ="}, _) $ t1 $ _)) =
+ | do_eq (@{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ t1 $ _)) =
do_lhs t1
| do_eq _ = false
in do_eq end
@@ -1347,7 +1347,7 @@
@{const "==>"} $ _ $ t2 => term_under_def t2
| Const (@{const_name "=="}, _) $ t1 $ _ => term_under_def t1
| @{const Trueprop} $ t1 => term_under_def t1
- | Const (@{const_name "op ="}, _) $ t1 $ _ => term_under_def t1
+ | Const (@{const_name HOL.eq}, _) $ t1 $ _ => term_under_def t1
| Abs (_, _, t') => term_under_def t'
| t1 $ _ => term_under_def t1
| _ => t
@@ -1371,7 +1371,7 @@
val (lhs, rhs) =
case t of
Const (@{const_name "=="}, _) $ t1 $ t2 => (t1, t2)
- | @{const Trueprop} $ (Const (@{const_name "op ="}, _) $ t1 $ t2) =>
+ | @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ t1 $ t2) =>
(t1, t2)
| _ => raise TERM ("Nitpick_HOL.normalized_rhs_of", [t])
val args = strip_comb lhs |> snd
@@ -1453,8 +1453,8 @@
| @{const "==>"} $ _ $ t2 => lhs_of_equation t2
| @{const Trueprop} $ t1 => lhs_of_equation t1
| Const (@{const_name All}, _) $ Abs (_, _, t1) => lhs_of_equation t1
- | Const (@{const_name "op ="}, _) $ t1 $ _ => SOME t1
- | @{const "op -->"} $ _ $ t2 => lhs_of_equation t2
+ | Const (@{const_name HOL.eq}, _) $ t1 $ _ => SOME t1
+ | @{const HOL.implies} $ _ $ t2 => lhs_of_equation t2
| _ => NONE
fun is_constr_pattern _ (Bound _) = true
| is_constr_pattern _ (Var _) = true
@@ -1807,7 +1807,7 @@
(betapply (t2, var_t))
end
| extensional_equal _ T t1 t2 =
- Const (@{const_name "op ="}, T --> T --> bool_T) $ t1 $ t2
+ Const (@{const_name HOL.eq}, T --> T --> bool_T) $ t1 $ t2
fun equationalize_term ctxt tag t =
let
@@ -1816,7 +1816,7 @@
in
Logic.list_implies (prems,
case concl of
- @{const Trueprop} $ (Const (@{const_name "op ="}, Type (_, [T, _]))
+ @{const Trueprop} $ (Const (@{const_name HOL.eq}, Type (_, [T, _]))
$ t1 $ t2) =>
@{const Trueprop} $ extensional_equal j T t1 t2
| @{const Trueprop} $ t' =>
@@ -2148,8 +2148,8 @@
fun repair_rec j (Const (@{const_name Ex}, T1) $ Abs (s2, T2, t2')) =
Const (@{const_name Ex}, T1)
$ Abs (s2, T2, repair_rec (j + 1) t2')
- | repair_rec j (@{const "op &"} $ t1 $ t2) =
- @{const "op &"} $ repair_rec j t1 $ repair_rec j t2
+ | repair_rec j (@{const HOL.conj} $ t1 $ t2) =
+ @{const HOL.conj} $ repair_rec j t1 $ repair_rec j t2
| repair_rec j t =
let val (head, args) = strip_comb t in
if head = Bound j then
@@ -2290,7 +2290,7 @@
| simps => simps
fun is_equational_fun_surely_complete hol_ctxt x =
case equational_fun_axioms hol_ctxt x of
- [@{const Trueprop} $ (Const (@{const_name "op ="}, _) $ t1 $ _)] =>
+ [@{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ t1 $ _)] =>
strip_comb t1 |> snd |> forall is_Var
| _ => false
--- a/src/HOL/Tools/Nitpick/nitpick_model.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_model.ML Wed Sep 01 17:19:47 2010 +0200
@@ -590,7 +590,7 @@
if co then
Const (@{const_name The}, (T --> bool_T) --> T)
$ Abs (cyclic_co_val_name (), T,
- Const (@{const_name "op ="}, T --> T --> bool_T)
+ Const (@{const_name HOL.eq}, T --> T --> bool_T)
$ Bound 0 $ abstract_over (var, t))
else
cyclic_atom ()
@@ -849,7 +849,7 @@
(** Model reconstruction **)
fun unfold_outer_the_binders (t as Const (@{const_name The}, _)
- $ Abs (s, T, Const (@{const_name "op ="}, _)
+ $ Abs (s, T, Const (@{const_name HOL.eq}, _)
$ Bound 0 $ t')) =
betapply (Abs (s, T, t'), t) |> unfold_outer_the_binders
| unfold_outer_the_binders t = t
--- a/src/HOL/Tools/Nitpick/nitpick_mono.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_mono.ML Wed Sep 01 17:19:47 2010 +0200
@@ -222,7 +222,7 @@
| fin_fun_body _ _ (t as Const (@{const_name None}, _)) = SOME t
| fin_fun_body dom_T ran_T
((t0 as Const (@{const_name If}, _))
- $ (t1 as Const (@{const_name "op ="}, _) $ Bound 0 $ t1')
+ $ (t1 as Const (@{const_name HOL.eq}, _) $ Bound 0 $ t1')
$ t2 $ t3) =
(if loose_bvar1 (t1', 0) then
NONE
@@ -650,7 +650,7 @@
Bound 0)))) accum
|>> mtype_of_mterm
end
- | @{const_name "op ="} => do_equals T accum
+ | @{const_name HOL.eq} => do_equals T accum
| @{const_name The} =>
(trace_msg (K "*** The"); raise UNSOLVABLE ())
| @{const_name Eps} =>
@@ -760,7 +760,7 @@
do_term (incr_boundvars ~1 t1') accum
else
raise SAME ()
- | (t11 as Const (@{const_name "op ="}, _)) $ Bound 0 $ t13 =>
+ | (t11 as Const (@{const_name HOL.eq}, _)) $ Bound 0 $ t13 =>
if not (loose_bvar1 (t13, 0)) then
do_term (incr_boundvars ~1 (t11 $ t13)) accum
else
@@ -774,10 +774,10 @@
(MAbs (s, T, M, S Minus, m'), accum |>> pop_bound)
end))
| (t0 as Const (@{const_name All}, _))
- $ Abs (s', T', (t10 as @{const "op -->"}) $ (t11 $ Bound 0) $ t12) =>
+ $ Abs (s', T', (t10 as @{const HOL.implies}) $ (t11 $ Bound 0) $ t12) =>
do_bounded_quantifier t0 s' T' t10 t11 t12 accum
| (t0 as Const (@{const_name Ex}, _))
- $ Abs (s', T', (t10 as @{const "op &"}) $ (t11 $ Bound 0) $ t12) =>
+ $ Abs (s', T', (t10 as @{const HOL.conj}) $ (t11 $ Bound 0) $ t12) =>
do_bounded_quantifier t0 s' T' t10 t11 t12 accum
| Const (@{const_name Let}, _) $ t1 $ t2 =>
do_term (betapply (t2, t1)) accum
@@ -876,19 +876,19 @@
do_term (@{const Not}
$ (HOLogic.eq_const (domain_type T0) $ t1
$ Abs (Name.uu, T1, @{const False}))) accum)
- | Const (x as (@{const_name "op ="}, _)) $ t1 $ t2 =>
+ | Const (x as (@{const_name HOL.eq}, _)) $ t1 $ t2 =>
do_equals x t1 t2
| Const (@{const_name Let}, _) $ t1 $ t2 =>
do_formula sn (betapply (t2, t1)) accum
| (t0 as Const (s0, _)) $ t1 $ t2 =>
if s0 = @{const_name "==>"} orelse
s0 = @{const_name Pure.conjunction} orelse
- s0 = @{const_name "op &"} orelse
- s0 = @{const_name "op |"} orelse
- s0 = @{const_name "op -->"} then
+ s0 = @{const_name HOL.conj} orelse
+ s0 = @{const_name HOL.disj} orelse
+ s0 = @{const_name HOL.implies} then
let
val impl = (s0 = @{const_name "==>"} orelse
- s0 = @{const_name "op -->"})
+ s0 = @{const_name HOL.implies})
val (m1, accum) = do_formula (sn |> impl ? negate) t1 accum
val (m2, accum) = do_formula sn t2 accum
in
@@ -973,10 +973,10 @@
do_conjunction t0 t1 t2 accum
| (t0 as Const (@{const_name All}, _)) $ Abs (s0, T1, t1) =>
do_all t0 s0 T1 t1 accum
- | Const (x as (@{const_name "op ="}, _)) $ t1 $ t2 =>
+ | Const (x as (@{const_name HOL.eq}, _)) $ t1 $ t2 =>
consider_general_equals mdata true x t1 t2 accum
- | (t0 as @{const "op &"}) $ t1 $ t2 => do_conjunction t0 t1 t2 accum
- | (t0 as @{const "op -->"}) $ t1 $ t2 => do_implies t0 t1 t2 accum
+ | (t0 as @{const HOL.conj}) $ t1 $ t2 => do_conjunction t0 t1 t2 accum
+ | (t0 as @{const HOL.implies}) $ t1 $ t2 => do_implies t0 t1 t2 accum
| _ => raise TERM ("Nitpick_Mono.consider_definitional_axiom.\
\do_formula", [t])
in do_formula t end
@@ -1069,7 +1069,7 @@
Abs (Name.uu, set_T', @{const True})
| _ => Const (s, T')
else if s = @{const_name "=="} orelse
- s = @{const_name "op ="} then
+ s = @{const_name HOL.eq} then
let
val T =
case T' of
--- a/src/HOL/Tools/Nitpick/nitpick_nut.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_nut.ML Wed Sep 01 17:19:47 2010 +0200
@@ -447,7 +447,7 @@
val t1 = incr_boundvars n t1
val t2 = incr_boundvars n t2
val xs = map Bound (n - 1 downto 0)
- val equation = Const (@{const_name "op ="},
+ val equation = Const (@{const_name HOL.eq},
body_T --> body_T --> bool_T)
$ betapplys (t1, xs) $ betapplys (t2, xs)
val t =
@@ -515,14 +515,14 @@
do_description_operator The @{const_name undefined_fast_The} x t1
| (Const (x as (@{const_name Eps}, _)), [t1]) =>
do_description_operator Eps @{const_name undefined_fast_Eps} x t1
- | (Const (@{const_name "op ="}, T), [t1]) =>
+ | (Const (@{const_name HOL.eq}, T), [t1]) =>
Op1 (SingletonSet, range_type T, Any, sub t1)
- | (Const (@{const_name "op ="}, T), [t1, t2]) => sub_equals T t1 t2
- | (Const (@{const_name "op &"}, _), [t1, t2]) =>
+ | (Const (@{const_name HOL.eq}, T), [t1, t2]) => sub_equals T t1 t2
+ | (Const (@{const_name HOL.conj}, _), [t1, t2]) =>
Op2 (And, bool_T, Any, sub' t1, sub' t2)
- | (Const (@{const_name "op |"}, _), [t1, t2]) =>
+ | (Const (@{const_name HOL.disj}, _), [t1, t2]) =>
Op2 (Or, bool_T, Any, sub t1, sub t2)
- | (Const (@{const_name "op -->"}, _), [t1, t2]) =>
+ | (Const (@{const_name HOL.implies}, _), [t1, t2]) =>
Op2 (Or, bool_T, Any, Op1 (Not, bool_T, Any, sub t1), sub' t2)
| (Const (@{const_name If}, T), [t1, t2, t3]) =>
Op3 (If, nth_range_type 3 T, Any, sub t1, sub t2, sub t3)
--- a/src/HOL/Tools/Nitpick/nitpick_preproc.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_preproc.ML Wed Sep 01 17:19:47 2010 +0200
@@ -41,9 +41,9 @@
fun aux def (Const (@{const_name "=="}, _) $ t1 $ t2) =
aux def t1 andalso aux false t2
| aux def (@{const "==>"} $ t1 $ t2) = aux false t1 andalso aux def t2
- | aux def (Const (@{const_name "op ="}, _) $ t1 $ t2) =
+ | aux def (Const (@{const_name HOL.eq}, _) $ t1 $ t2) =
aux def t1 andalso aux false t2
- | aux def (@{const "op -->"} $ t1 $ t2) = aux false t1 andalso aux def t2
+ | aux def (@{const HOL.implies} $ t1 $ t2) = aux false t1 andalso aux def t2
| aux def (t1 $ t2) = aux def t1 andalso aux def t2
| aux def (t as Const (s, _)) =
(not def orelse t <> @{const Suc}) andalso
@@ -149,7 +149,7 @@
case t of
@{const Trueprop} $ t1 => box_var_in_def new_Ts old_Ts t1 z
| Const (s0, _) $ t1 $ _ =>
- if s0 = @{const_name "=="} orelse s0 = @{const_name "op ="} then
+ if s0 = @{const_name "=="} orelse s0 = @{const_name HOL.eq} then
let
val (t', args) = strip_comb t1
val T' = fastype_of1 (new_Ts, do_term new_Ts old_Ts Neut t')
@@ -209,16 +209,16 @@
do_quantifier new_Ts old_Ts polar s0 T0 s1 T1 t1
| Const (s0 as @{const_name Ex}, T0) $ Abs (s1, T1, t1) =>
do_quantifier new_Ts old_Ts polar s0 T0 s1 T1 t1
- | Const (s0 as @{const_name "op ="}, T0) $ t1 $ t2 =>
+ | Const (s0 as @{const_name HOL.eq}, T0) $ t1 $ t2 =>
do_equals new_Ts old_Ts s0 T0 t1 t2
- | @{const "op &"} $ t1 $ t2 =>
- @{const "op &"} $ do_term new_Ts old_Ts polar t1
+ | @{const HOL.conj} $ t1 $ t2 =>
+ @{const HOL.conj} $ do_term new_Ts old_Ts polar t1
$ do_term new_Ts old_Ts polar t2
- | @{const "op |"} $ t1 $ t2 =>
- @{const "op |"} $ do_term new_Ts old_Ts polar t1
+ | @{const HOL.disj} $ t1 $ t2 =>
+ @{const HOL.disj} $ do_term new_Ts old_Ts polar t1
$ do_term new_Ts old_Ts polar t2
- | @{const "op -->"} $ t1 $ t2 =>
- @{const "op -->"} $ do_term new_Ts old_Ts (flip_polarity polar) t1
+ | @{const HOL.implies} $ t1 $ t2 =>
+ @{const HOL.implies} $ do_term new_Ts old_Ts (flip_polarity polar) t1
$ do_term new_Ts old_Ts polar t2
| Const (x as (s, T)) =>
if is_descr s then
@@ -332,9 +332,9 @@
do_eq_or_imp Ts true def t0 t1 t2 seen
| (t0 as @{const "==>"}) $ t1 $ t2 =>
if def then (t, []) else do_eq_or_imp Ts false def t0 t1 t2 seen
- | (t0 as Const (@{const_name "op ="}, _)) $ t1 $ t2 =>
+ | (t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2 =>
do_eq_or_imp Ts true def t0 t1 t2 seen
- | (t0 as @{const "op -->"}) $ t1 $ t2 =>
+ | (t0 as @{const HOL.implies}) $ t1 $ t2 =>
do_eq_or_imp Ts false def t0 t1 t2 seen
| Abs (s, T, t') =>
let val (t', seen) = do_term (T :: Ts) def t' [] seen in
@@ -399,9 +399,9 @@
aux_eq careful true t0 t1 t2
| aux careful ((t0 as @{const "==>"}) $ t1 $ t2) =
t0 $ aux false t1 $ aux careful t2
- | aux careful ((t0 as Const (@{const_name "op ="}, _)) $ t1 $ t2) =
+ | aux careful ((t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2) =
aux_eq careful true t0 t1 t2
- | aux careful ((t0 as @{const "op -->"}) $ t1 $ t2) =
+ | aux careful ((t0 as @{const HOL.implies}) $ t1 $ t2) =
t0 $ aux false t1 $ aux careful t2
| aux careful (Abs (s, T, t')) = Abs (s, T, aux careful t')
| aux careful (t1 $ t2) = aux careful t1 $ aux careful t2
@@ -449,7 +449,7 @@
(** Destruction of universal and existential equalities **)
fun curry_assms (@{const "==>"} $ (@{const Trueprop}
- $ (@{const "op &"} $ t1 $ t2)) $ t3) =
+ $ (@{const HOL.conj} $ t1 $ t2)) $ t3) =
curry_assms (Logic.list_implies ([t1, t2] |> map HOLogic.mk_Trueprop, t3))
| curry_assms (@{const "==>"} $ t1 $ t2) =
@{const "==>"} $ curry_assms t1 $ curry_assms t2
@@ -464,9 +464,9 @@
and aux_implies prems zs t1 t2 =
case t1 of
Const (@{const_name "=="}, _) $ Var z $ t' => aux_eq prems zs z t' t1 t2
- | @{const Trueprop} $ (Const (@{const_name "op ="}, _) $ Var z $ t') =>
+ | @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ Var z $ t') =>
aux_eq prems zs z t' t1 t2
- | @{const Trueprop} $ (Const (@{const_name "op ="}, _) $ t' $ Var z) =>
+ | @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ t' $ Var z) =>
aux_eq prems zs z t' t1 t2
| _ => aux (t1 :: prems) (Term.add_vars t1 zs) t2
and aux_eq prems zs z t' t1 t2 =
@@ -499,7 +499,7 @@
handle SAME () => do_term (t :: seen) ts
in
case t of
- Const (@{const_name "op ="}, _) $ t1 $ t2 => do_eq true t1 t2
+ Const (@{const_name HOL.eq}, _) $ t1 $ t2 => do_eq true t1 t2
| _ => do_term (t :: seen) ts
end
in do_term end
@@ -604,12 +604,12 @@
do_quantifier s0 T0 s1 T1 t1
| Const (s0 as @{const_name Ex}, T0) $ Abs (s1, T1, t1) =>
do_quantifier s0 T0 s1 T1 t1
- | @{const "op &"} $ t1 $ t2 =>
+ | @{const HOL.conj} $ t1 $ t2 =>
s_conj (pairself (aux ss Ts js skolemizable polar) (t1, t2))
- | @{const "op |"} $ t1 $ t2 =>
+ | @{const HOL.disj} $ t1 $ t2 =>
s_disj (pairself (aux ss Ts js skolemizable polar) (t1, t2))
- | @{const "op -->"} $ t1 $ t2 =>
- @{const "op -->"} $ aux ss Ts js skolemizable (flip_polarity polar) t1
+ | @{const HOL.implies} $ t1 $ t2 =>
+ @{const HOL.implies} $ aux ss Ts js skolemizable (flip_polarity polar) t1
$ aux ss Ts js skolemizable polar t2
| (t0 as Const (@{const_name Let}, _)) $ t1 $ t2 =>
t0 $ t1 $ aux ss Ts js skolemizable polar t2
@@ -620,8 +620,8 @@
let
val gfp = (fixpoint_kind_of_const thy def_table x = Gfp)
val (pref, connective) =
- if gfp then (lbfp_prefix, @{const "op |"})
- else (ubfp_prefix, @{const "op &"})
+ if gfp then (lbfp_prefix, @{const HOL.disj})
+ else (ubfp_prefix, @{const HOL.conj})
fun pos () = unrolled_inductive_pred_const hol_ctxt gfp x
|> aux ss Ts js skolemizable polar
fun neg () = Const (pref ^ s, T)
@@ -653,7 +653,7 @@
fun params_in_equation (@{const "==>"} $ _ $ t2) = params_in_equation t2
| params_in_equation (@{const Trueprop} $ t1) = params_in_equation t1
- | params_in_equation (Const (@{const_name "op ="}, _) $ t1 $ _) =
+ | params_in_equation (Const (@{const_name HOL.eq}, _) $ t1 $ _) =
snd (strip_comb t1)
| params_in_equation _ = []
@@ -1105,7 +1105,7 @@
case t of
(t0 as Const (@{const_name All}, T0)) $ Abs (s, T1, t1) =>
(case t1 of
- (t10 as @{const "op &"}) $ t11 $ t12 =>
+ (t10 as @{const HOL.conj}) $ t11 $ t12 =>
t10 $ distribute_quantifiers (t0 $ Abs (s, T1, t11))
$ distribute_quantifiers (t0 $ Abs (s, T1, t12))
| (t10 as @{const Not}) $ t11 =>
@@ -1118,10 +1118,10 @@
t0 $ Abs (s, T1, distribute_quantifiers t1))
| (t0 as Const (@{const_name Ex}, T0)) $ Abs (s, T1, t1) =>
(case distribute_quantifiers t1 of
- (t10 as @{const "op |"}) $ t11 $ t12 =>
+ (t10 as @{const HOL.disj}) $ t11 $ t12 =>
t10 $ distribute_quantifiers (t0 $ Abs (s, T1, t11))
$ distribute_quantifiers (t0 $ Abs (s, T1, t12))
- | (t10 as @{const "op -->"}) $ t11 $ t12 =>
+ | (t10 as @{const HOL.implies}) $ t11 $ t12 =>
t10 $ distribute_quantifiers (Const (@{const_name All}, T0)
$ Abs (s, T1, t11))
$ distribute_quantifiers (t0 $ Abs (s, T1, t12))
--- a/src/HOL/Tools/Predicate_Compile/code_prolog.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/code_prolog.ML Wed Sep 01 17:19:47 2010 +0200
@@ -6,23 +6,39 @@
signature CODE_PROLOG =
sig
+ datatype prolog_system = SWI_PROLOG | YAP
+ type code_options =
+ {ensure_groundness : bool,
+ limited_types : (typ * int) list,
+ limited_predicates : (string list * int) list,
+ replacing : ((string * string) * string) list,
+ manual_reorder : ((string * int) * int list) list,
+ prolog_system : prolog_system}
+ val code_options_of : theory -> code_options
+ val map_code_options : (code_options -> code_options) -> theory -> theory
+
datatype arith_op = Plus | Minus
datatype prol_term = Var of string | Cons of string | AppF of string * prol_term list
| Number of int | ArithOp of arith_op * prol_term list;
datatype prem = Conj of prem list
| Rel of string * prol_term list | NotRel of string * prol_term list
| Eq of prol_term * prol_term | NotEq of prol_term * prol_term
- | ArithEq of prol_term * prol_term | NotArithEq of prol_term * prol_term;
-
+ | ArithEq of prol_term * prol_term | NotArithEq of prol_term * prol_term
+ | Ground of string * typ;
+
type clause = ((string * prol_term list) * prem);
type logic_program = clause list;
type constant_table = (string * string) list
-
- val generate : Proof.context -> string list -> (logic_program * constant_table)
+
+ val generate : bool -> Proof.context -> string -> (logic_program * constant_table)
val write_program : logic_program -> string
- val run : logic_program -> string -> string list -> int option -> prol_term list list
+ val run : prolog_system -> logic_program -> string -> string list -> int option -> prol_term list list
+
+ val quickcheck : Proof.context -> bool -> term -> int -> term list option * (bool list * bool)
val trace : bool Unsynchronized.ref
+
+ val replace : ((string * string) * string) -> logic_program -> logic_program
end;
structure Code_Prolog : CODE_PROLOG =
@@ -33,6 +49,45 @@
val trace = Unsynchronized.ref false
fun tracing s = if !trace then Output.tracing s else ()
+
+(* code generation options *)
+
+datatype prolog_system = SWI_PROLOG | YAP
+
+type code_options =
+ {ensure_groundness : bool,
+ limited_types : (typ * int) list,
+ limited_predicates : (string list * int) list,
+ replacing : ((string * string) * string) list,
+ manual_reorder : ((string * int) * int list) list,
+ prolog_system : prolog_system}
+
+structure Options = Theory_Data
+(
+ type T = code_options
+ val empty = {ensure_groundness = false,
+ limited_types = [], limited_predicates = [], replacing = [], manual_reorder = [],
+ prolog_system = SWI_PROLOG}
+ val extend = I;
+ fun merge
+ ({ensure_groundness = ensure_groundness1, limited_types = limited_types1,
+ limited_predicates = limited_predicates1, replacing = replacing1,
+ manual_reorder = manual_reorder1, prolog_system = prolog_system1},
+ {ensure_groundness = ensure_groundness2, limited_types = limited_types2,
+ limited_predicates = limited_predicates2, replacing = replacing2,
+ manual_reorder = manual_reorder2, prolog_system = prolog_system2}) =
+ {ensure_groundness = ensure_groundness1 orelse ensure_groundness2,
+ limited_types = AList.merge (op =) (K true) (limited_types1, limited_types2),
+ limited_predicates = AList.merge (op =) (K true) (limited_predicates1, limited_predicates2),
+ manual_reorder = AList.merge (op =) (K true) (manual_reorder1, manual_reorder2),
+ replacing = Library.merge (op =) (replacing1, replacing2),
+ prolog_system = prolog_system1};
+);
+
+val code_options_of = Options.get
+
+val map_code_options = Options.map
+
(* general string functions *)
val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o explode;
@@ -45,6 +100,21 @@
datatype prol_term = Var of string | Cons of string | AppF of string * prol_term list
| Number of int | ArithOp of arith_op * prol_term list;
+fun dest_Var (Var v) = v
+
+fun add_vars (Var v) = insert (op =) v
+ | add_vars (ArithOp (_, ts)) = fold add_vars ts
+ | add_vars (AppF (_, ts)) = fold add_vars ts
+ | add_vars _ = I
+
+fun map_vars f (Var v) = Var (f v)
+ | map_vars f (ArithOp (opr, ts)) = ArithOp (opr, map (map_vars f) ts)
+ | map_vars f (AppF (fs, ts)) = AppF (fs, map (map_vars f) ts)
+ | map_vars f t = t
+
+fun maybe_AppF (c, []) = Cons c
+ | maybe_AppF (c, xs) = AppF (c, xs)
+
fun is_Var (Var _) = true
| is_Var _ = false
@@ -61,26 +131,61 @@
datatype prem = Conj of prem list
| Rel of string * prol_term list | NotRel of string * prol_term list
| Eq of prol_term * prol_term | NotEq of prol_term * prol_term
- | ArithEq of prol_term * prol_term | NotArithEq of prol_term * prol_term;
-
+ | ArithEq of prol_term * prol_term | NotArithEq of prol_term * prol_term
+ | Ground of string * typ;
+
fun dest_Rel (Rel (c, ts)) = (c, ts)
-
+
+fun map_term_prem f (Conj prems) = Conj (map (map_term_prem f) prems)
+ | map_term_prem f (Rel (r, ts)) = Rel (r, map f ts)
+ | map_term_prem f (NotRel (r, ts)) = NotRel (r, map f ts)
+ | map_term_prem f (Eq (l, r)) = Eq (f l, f r)
+ | map_term_prem f (NotEq (l, r)) = NotEq (f l, f r)
+ | map_term_prem f (ArithEq (l, r)) = ArithEq (f l, f r)
+ | map_term_prem f (NotArithEq (l, r)) = NotArithEq (f l, f r)
+ | map_term_prem f (Ground (v, T)) = Ground (dest_Var (f (Var v)), T)
+
+fun fold_prem_terms f (Conj prems) = fold (fold_prem_terms f) prems
+ | fold_prem_terms f (Rel (_, ts)) = fold f ts
+ | fold_prem_terms f (NotRel (_, ts)) = fold f ts
+ | fold_prem_terms f (Eq (l, r)) = f l #> f r
+ | fold_prem_terms f (NotEq (l, r)) = f l #> f r
+ | fold_prem_terms f (ArithEq (l, r)) = f l #> f r
+ | fold_prem_terms f (NotArithEq (l, r)) = f l #> f r
+ | fold_prem_terms f (Ground (v, T)) = f (Var v)
+
type clause = ((string * prol_term list) * prem);
type logic_program = clause list;
(* translation from introduction rules to internal representation *)
+fun mk_conform f empty avoid name =
+ let
+ fun dest_Char (Symbol.Char c) = c
+ val name' = space_implode "" (map (dest_Char o Symbol.decode)
+ (filter (fn s => Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s)
+ (Symbol.explode name)))
+ val name'' = f (if name' = "" then empty else name')
+ in (if member (op =) avoid name'' then Name.variant avoid name'' else name'') end
+
(** constant table **)
type constant_table = (string * string) list
(* assuming no clashing *)
-fun mk_constant_table consts =
- AList.make (first_lower o Long_Name.base_name) consts
-
fun declare_consts consts constant_table =
- fold (fn c => AList.update (op =) (c, first_lower (Long_Name.base_name c))) consts constant_table
+ let
+ fun update' c table =
+ if AList.defined (op =) table c then table else
+ let
+ val c' = mk_conform first_lower "pred" (map snd table) (Long_Name.base_name c)
+ in
+ AList.update (op =) (c, c') table
+ end
+ in
+ fold update' consts constant_table
+ end
fun translate_const constant_table c =
case AList.lookup (op =) constant_table c of
@@ -96,16 +201,23 @@
case inv_lookup (op =) constant_table c of
SOME c' => c'
| NONE => error ("No constant corresponding to " ^ c)
-
+
(** translation of terms, literals, premises, and clauses **)
fun translate_arith_const @{const_name "Groups.plus_class.plus"} = SOME Plus
| translate_arith_const @{const_name "Groups.minus_class.minus"} = SOME Minus
| translate_arith_const _ = NONE
+fun mk_nat_term constant_table n =
+ let
+ val zero = translate_const constant_table @{const_name "Groups.zero_class.zero"}
+ val Suc = translate_const constant_table @{const_name "Suc"}
+ in funpow n (fn t => AppF (Suc, [t])) (Cons zero) end
+
fun translate_term ctxt constant_table t =
case try HOLogic.dest_number t of
SOME (@{typ "int"}, n) => Number n
+ | SOME (@{typ "nat"}, n) => mk_nat_term constant_table n
| NONE =>
(case strip_comb t of
(Free (v, T), []) => Var v
@@ -119,12 +231,12 @@
fun translate_literal ctxt constant_table t =
case strip_comb t of
- (Const (@{const_name "op ="}, _), [l, r]) =>
+ (Const (@{const_name HOL.eq}, _), [l, r]) =>
let
val l' = translate_term ctxt constant_table l
val r' = translate_term ctxt constant_table r
in
- (if is_Var l' andalso is_arith_term r' then ArithEq else Eq) (l', r')
+ (if is_Var l' andalso is_arith_term r' andalso not (is_Var r') then ArithEq else Eq) (l', r')
end
| (Const (c, _), args) =>
Rel (translate_const constant_table c, map (translate_term ctxt constant_table) args)
@@ -134,11 +246,16 @@
| NegRel_of (Eq eq) = NotEq eq
| NegRel_of (ArithEq eq) = NotArithEq eq
-fun translate_prem ctxt constant_table t =
+fun mk_groundness_prems t = map Ground (Term.add_frees t [])
+
+fun translate_prem ensure_groundness ctxt constant_table t =
case try HOLogic.dest_not t of
- SOME t => NegRel_of (translate_literal ctxt constant_table t)
+ SOME t =>
+ if ensure_groundness then
+ Conj (mk_groundness_prems t @ [NegRel_of (translate_literal ctxt constant_table t)])
+ else
+ NegRel_of (translate_literal ctxt constant_table t)
| NONE => translate_literal ctxt constant_table t
-
fun imp_prems_conv cv ct =
case Thm.term_of ct of
@@ -156,64 +273,230 @@
(Trueprop_conv (Conv.try_conv (Conv.rewr_conv @{thm Predicate.eq_is_eq}))))
(Thm.transfer thy rule)
-fun translate_intros ctxt gr const constant_table =
+fun translate_intros ensure_groundness ctxt gr const constant_table =
let
val intros = map (preprocess_intro (ProofContext.theory_of ctxt)) (Graph.get_node gr const)
val (intros', ctxt') = Variable.import_terms true (map prop_of intros) ctxt
val constant_table' = declare_consts (fold Term.add_const_names intros' []) constant_table
+ |> declare_consts [@{const_name "Groups.zero_class.zero"}, @{const_name "Suc"}]
fun translate_intro intro =
let
val head = HOLogic.dest_Trueprop (Logic.strip_imp_concl intro)
- val prems = map HOLogic.dest_Trueprop (Logic.strip_imp_prems intro)
- val prems' = Conj (map (translate_prem ctxt' constant_table') prems)
+ val prems = map HOLogic.dest_Trueprop (Logic.strip_imp_prems intro)
+ val prems' = Conj (map (translate_prem ensure_groundness ctxt' constant_table') prems)
val clause = (dest_Rel (translate_literal ctxt' constant_table' head), prems')
in clause end
- in (map translate_intro intros', constant_table') end
+ val res = (map translate_intro intros', constant_table')
+ in res end
+
+fun depending_preds_of (key, intros) =
+ fold Term.add_const_names (map Thm.prop_of intros) []
-fun generate ctxt const =
- 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 scc = strong_conn_of gr const
- val constant_table = mk_constant_table (flat scc)
+fun add_edges edges_of key G =
+ let
+ fun extend' key (G, visited) =
+ case try (Graph.get_node G) key of
+ SOME v =>
+ let
+ val new_edges = filter (fn k => is_some (try (Graph.get_node G) k)) (edges_of (key, v))
+ val (G', visited') = fold extend'
+ (subtract (op =) (key :: visited) new_edges) (G, key :: visited)
+ in
+ (fold (Graph.add_edge o (pair key)) new_edges G', visited')
+ end
+ | NONE => (G, visited)
in
- apfst flat (fold_map (translate_intros ctxt gr) (flat scc) constant_table)
+ fst (extend' key (G, []))
end
-(* transform logic program *)
-
-(** ensure groundness of terms before negation **)
+fun generate ensure_groundness ctxt const =
+ 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' = add_edges depending_preds_of const gr
+ val scc = strong_conn_of gr' [const]
+ val constant_table = declare_consts (flat scc) []
+ in
+ apfst flat (fold_map (translate_intros ensure_groundness ctxt gr) (flat scc) constant_table)
+ end
+
+(* implementation for fully enumerating predicates and
+ for size-limited predicates for enumerating the values of a datatype upto a specific size *)
-fun add_vars (Var x) vs = insert (op =) x vs
- | add_vars (Cons c) vs = vs
- | add_vars (AppF (f, args)) vs = fold add_vars args vs
+fun add_ground_typ (Conj prems) = fold add_ground_typ prems
+ | add_ground_typ (Ground (_, T)) = insert (op =) T
+ | add_ground_typ _ = I
+
+fun mk_relname (Type (Tcon, Targs)) =
+ first_lower (Long_Name.base_name Tcon) ^ space_implode "_" (map mk_relname Targs)
+ | mk_relname _ = raise Fail "unexpected type"
+
+fun mk_lim_relname T = "lim_" ^ mk_relname T
+
+(* This is copied from "pat_completeness.ML" *)
+fun inst_constrs_of thy (T as Type (name, _)) =
+ map (fn (Cn,CT) =>
+ Envir.subst_term_types (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT)))
+ (the (Datatype.get_constrs thy name))
+ | inst_constrs_of thy T = raise TYPE ("inst_constrs_of", [T], [])
-fun string_of_typ (Type (s, Ts)) = Long_Name.base_name s
+fun is_recursive_constr T (Const (constr_name, T')) = member (op =) (binder_types T') T
+
+fun mk_ground_impl ctxt limited_types (T as Type (Tcon, Targs)) (seen, constant_table) =
+ if member (op =) seen T then ([], (seen, constant_table))
+ else
+ let
+ val (limited, size) = case AList.lookup (op =) limited_types T of
+ SOME s => (true, s)
+ | NONE => (false, 0)
+ val rel_name = (if limited then mk_lim_relname else mk_relname) T
+ fun mk_impl (Const (constr_name, cT), recursive) (seen, constant_table) =
+ let
+ val constant_table' = declare_consts [constr_name] constant_table
+ val Ts = binder_types cT
+ val (rec_clauses, (seen', constant_table'')) =
+ fold_map (mk_ground_impl ctxt limited_types) Ts (seen, constant_table')
+ val vars = map (fn i => Var ("x" ^ string_of_int i)) (1 upto (length Ts))
+ val lim_var =
+ if limited then
+ if recursive then [AppF ("suc", [Var "Lim"])]
+ else [Var "Lim"]
+ else []
+ fun mk_prem v T' =
+ if limited andalso T' = T then Rel (mk_lim_relname T', [Var "Lim", v])
+ else Rel (mk_relname T', [v])
+ val clause =
+ ((rel_name, lim_var @ [maybe_AppF (translate_const constant_table'' constr_name, vars)]),
+ Conj (map2 mk_prem vars Ts))
+ in
+ (clause :: flat rec_clauses, (seen', constant_table''))
+ end
+ val constrs = inst_constrs_of (ProofContext.theory_of ctxt) T
+ val constrs' = (constrs ~~ map (is_recursive_constr T) constrs)
+ |> (fn cs => filter_out snd cs @ filter snd cs)
+ val (clauses, constant_table') =
+ apfst flat (fold_map mk_impl constrs' (T :: seen, constant_table))
+ val size_term = funpow size (fn t => AppF ("suc", [t])) (Cons "zero")
+ in
+ ((if limited then
+ cons ((mk_relname T, [Var "x"]), Rel (mk_lim_relname T, [size_term, Var "x"]))
+ else I) clauses, constant_table')
+ end
+ | mk_ground_impl ctxt _ T (seen, constant_table) =
+ raise Fail ("unexpected type :" ^ Syntax.string_of_typ ctxt T)
-fun mk_groundness_prems ts =
+fun replace_ground (Conj prems) = Conj (map replace_ground prems)
+ | replace_ground (Ground (x, T)) =
+ Rel (mk_relname T, [Var x])
+ | replace_ground p = p
+
+fun add_ground_predicates ctxt limited_types (p, constant_table) =
let
- val vars = fold add_vars ts []
- fun mk_ground v =
- Rel ("ground", [Var v])
+ val ground_typs = fold (add_ground_typ o snd) p []
+ val (grs, (_, constant_table')) = fold_map (mk_ground_impl ctxt limited_types) ground_typs ([], constant_table)
+ val p' = map (apsnd replace_ground) p
in
- map mk_ground vars
+ ((flat grs) @ p', constant_table')
end
-fun ensure_groundness_prem (NotRel (c, ts)) = Conj (mk_groundness_prems ts @ [NotRel (c, ts)])
- | ensure_groundness_prem (NotEq (l, r)) = Conj (mk_groundness_prems [l, r] @ [NotEq (l, r)])
- | ensure_groundness_prem (Conj ps) = Conj (map ensure_groundness_prem ps)
- | ensure_groundness_prem p = p
+(* make depth-limited version of predicate *)
+
+fun mk_lim_rel_name rel_name = "lim_" ^ rel_name
+
+fun mk_depth_limited rel_names ((rel_name, ts), prem) =
+ let
+ fun has_positive_recursive_prems (Conj prems) = exists has_positive_recursive_prems prems
+ | has_positive_recursive_prems (Rel (rel, ts)) = member (op =) rel_names rel
+ | has_positive_recursive_prems _ = false
+ fun mk_lim_prem (Conj prems) = Conj (map mk_lim_prem prems)
+ | mk_lim_prem (p as Rel (rel, ts)) =
+ if member (op =) rel_names rel then Rel (mk_lim_rel_name rel, Var "Lim" :: ts) else p
+ | mk_lim_prem p = p
+ in
+ if has_positive_recursive_prems prem then
+ ((mk_lim_rel_name rel_name, (AppF ("suc", [Var "Lim"])) :: ts), mk_lim_prem prem)
+ else
+ ((mk_lim_rel_name rel_name, (Var "Lim") :: ts), prem)
+ end
+
+fun add_limited_predicates limited_predicates =
+ let
+ fun add (rel_names, limit) (p, constant_table) =
+ let
+ val clauses = filter (fn ((rel, _), _) => member (op =) rel_names rel) p
+ val clauses' = map (mk_depth_limited rel_names) clauses
+ fun nat_term_of n = funpow n (fn t => AppF ("suc", [t])) (Cons "zero")
+ fun mk_entry_clause rel_name =
+ let
+ val nargs = length (snd (fst
+ (the (find_first (fn ((rel, _), _) => rel = rel_name) clauses))))
+ val vars = map (fn i => Var ("x" ^ string_of_int i)) (1 upto nargs)
+ in
+ (("limited_" ^ rel_name, vars), Rel ("lim_" ^ rel_name, nat_term_of limit :: vars))
+ end
+ in (p @ (map mk_entry_clause rel_names) @ clauses', constant_table) end
+ in
+ fold add limited_predicates
+ end
+
+
+(* replace predicates in clauses *)
-fun ensure_groundness_before_negation p =
- map (apsnd ensure_groundness_prem) p
+(* replace (A, B, C) p = replace A by B in clauses of C *)
+fun replace ((from, to), location) p =
+ let
+ fun replace_prem (Conj prems) = Conj (map replace_prem prems)
+ | replace_prem (r as Rel (rel, ts)) =
+ if rel = from then Rel (to, ts) else r
+ | replace_prem r = r
+ in
+ map (fn ((rel, args), prem) => ((rel, args), (if rel = location then replace_prem else I) prem)) p
+ end
+
+
+(* reorder manually : reorder premises of ith clause of predicate p by a permutation perm *)
+
+fun reorder_manually reorder p =
+ let
+ fun reorder' (clause as ((rel, args), prem)) seen =
+ let
+ val seen' = AList.map_default (op =) (rel, 0) (fn x => x + 1) seen
+ val i = the (AList.lookup (op =) seen' rel)
+ val perm = AList.lookup (op =) reorder (rel, i)
+ val prem' = (case perm of
+ SOME p => (case prem of Conj prems => Conj (map (nth prems) p) | _ => prem)
+ | NONE => prem)
+ in (((rel, args), prem'), seen') end
+ in
+ fst (fold_map reorder' p [])
+ end
+(* rename variables to prolog-friendly names *)
+
+fun rename_vars_term renaming = map_vars (fn v => the (AList.lookup (op =) renaming v))
+
+fun rename_vars_prem renaming = map_term_prem (rename_vars_term renaming)
+
+fun is_prolog_conform v =
+ forall (fn s => Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s) (Symbol.explode v)
+
+fun mk_renaming v renaming =
+ (v, mk_conform first_upper "Var" (map snd renaming) v) :: renaming
+
+fun rename_vars_clause ((rel, args), prem) =
+ let
+ val vars = fold_prem_terms add_vars prem (fold add_vars args [])
+ val renaming = fold mk_renaming vars []
+ in ((rel, map (rename_vars_term renaming) args), rename_vars_prem renaming prem) end
+
+val rename_vars_program = map rename_vars_clause
(* code printer *)
fun write_arith_op Plus = "+"
| write_arith_op Minus = "-"
-fun write_term (Var v) = first_upper v
+fun write_term (Var v) = v
| write_term (Cons c) = c
| write_term (AppF (f, args)) = f ^ "(" ^ space_implode ", " (map write_term args) ^ ")"
| write_term (ArithOp (oper, [a1, a2])) = write_term a1 ^ " " ^ write_arith_op oper ^ " " ^ write_term a2
@@ -236,14 +519,16 @@
fun write_program p =
cat_lines (map write_clause p)
-(** query templates **)
+(* query templates *)
-fun query_first rel vnames =
+(** query and prelude for swi-prolog **)
+
+fun swi_prolog_query_first rel vnames =
"eval :- once(" ^ rel ^ "(" ^ space_implode ", " vnames ^ ")),\n" ^
"writef('" ^ space_implode ";" (map (fn v => v ^ " = %w") vnames) ^
"\\n', [" ^ space_implode ", " vnames ^ "]).\n"
-fun query_firstn n rel vnames =
+fun swi_prolog_query_firstn n rel vnames =
"eval :- findnsols(" ^ string_of_int n ^ ", (" ^ space_implode ", " vnames ^ "), " ^
rel ^ "(" ^ space_implode ", " vnames ^ "), Sols), writelist(Sols).\n" ^
"writelist([]).\n" ^
@@ -251,19 +536,52 @@
"writef('" ^ space_implode ";" (map (fn v => v ^ " = %w") vnames) ^
"\\n', [" ^ space_implode ", " vnames ^ "]), writelist(T).\n"
-val prelude =
+val swi_prolog_prelude =
"#!/usr/bin/swipl -q -t main -f\n\n" ^
":- use_module(library('dialect/ciao/aggregates')).\n" ^
- ":- style_check(-singleton).\n\n" ^
+ ":- style_check(-singleton).\n" ^
+ ":- style_check(-discontiguous).\n" ^
+ ":- style_check(-atom).\n\n" ^
"main :- catch(eval, E, (print_message(error, E), fail)), halt.\n" ^
"main :- halt(1).\n"
+(** query and prelude for yap **)
+
+fun yap_query_first rel vnames =
+ "eval :- once(" ^ rel ^ "(" ^ space_implode ", " vnames ^ ")),\n" ^
+ "format('" ^ space_implode ";" (map (fn v => v ^ " = ~w") vnames) ^
+ "\\n', [" ^ space_implode ", " vnames ^ "]).\n"
+
+val yap_prelude =
+ "#!/usr/bin/yap -L\n\n" ^
+ ":- initialization(eval).\n"
+
+(* system-dependent query, prelude and invocation *)
+
+fun query system nsols =
+ case system of
+ SWI_PROLOG =>
+ (case nsols of NONE => swi_prolog_query_first | SOME n => swi_prolog_query_firstn n)
+ | YAP =>
+ case nsols of NONE => yap_query_first | SOME n =>
+ error "No support for querying multiple solutions in the prolog system yap"
+
+fun prelude system =
+ case system of SWI_PROLOG => swi_prolog_prelude | YAP => yap_prelude
+
+fun invoke system file_name =
+ let
+ val cmd =
+ case system of SWI_PROLOG => "/usr/local/bin/swipl -f " | YAP => "/usr/local/bin/yap -L "
+ in fst (bash_output (cmd ^ file_name)) end
+
(* parsing prolog solution *)
+
val scan_number =
Scan.many1 Symbol.is_ascii_digit
val scan_atom =
- Scan.many1 (fn s => Symbol.is_ascii_lower s orelse Symbol.is_ascii_quasi s)
+ Scan.many1 (fn s => Symbol.is_ascii_lower s orelse Symbol.is_ascii_digit s orelse Symbol.is_ascii_quasi s)
val scan_var =
Scan.many1
@@ -303,28 +621,30 @@
(l :: r :: []) => parse_term (unprefix " " r)
| _ => raise Fail "unexpected equation in prolog output"
fun parse_solution s = map dest_eq (space_explode ";" s)
+ val sols = case space_explode "\n" sol of [] => [] | s => fst (split_last s)
in
- map parse_solution (fst (split_last (space_explode "\n" sol)))
+ map parse_solution sols
end
(* calling external interpreter and getting results *)
-fun run p query_rel vnames nsols =
+fun run system p query_rel vnames nsols =
let
- val cmd = Path.named_root
- val query = case nsols of NONE => query_first | SOME n => query_firstn n
- val prog = prelude ^ query query_rel vnames ^ write_program p
+ val p' = rename_vars_program p
+ val _ = tracing "Renaming variable names..."
+ val renaming = fold mk_renaming vnames []
+ val vnames' = map (fn v => the (AList.lookup (op =) renaming v)) vnames
+ val prog = prelude system ^ query system nsols query_rel vnames' ^ write_program p'
val _ = tracing ("Generated prolog program:\n" ^ prog)
- val prolog_file = File.tmp_path (Path.basic "prolog_file")
- val _ = File.write prolog_file prog
- val (solution, _) = bash_output ("/usr/local/bin/swipl -f " ^ File.shell_path prolog_file)
+ val solution = Cache_IO.with_tmp_file "prolog_file" (fn prolog_file =>
+ (File.write prolog_file prog; invoke system (Path.implode prolog_file)))
val _ = tracing ("Prolog returned solution(s):\n" ^ solution)
val tss = parse_solutions solution
in
tss
end
-(* values command *)
+(* restoring types in terms *)
fun restore_term ctxt constant_table (Var s, T) = Free (s, T)
| restore_term ctxt constant_table (Number n, @{typ "int"}) = HOLogic.mk_number @{typ "int"} n
@@ -343,8 +663,33 @@
map (restore_term ctxt constant_table) (args ~~ argsT'))
end
+(* values command *)
+
+val preprocess_options = Predicate_Compile_Aux.Options {
+ expected_modes = NONE,
+ proposed_modes = NONE,
+ proposed_names = [],
+ show_steps = false,
+ show_intermediate_results = false,
+ show_proof_trace = false,
+ show_modes = false,
+ show_mode_inference = false,
+ show_compilation = false,
+ show_caught_failures = false,
+ skip_proof = true,
+ no_topmost_reordering = false,
+ function_flattening = true,
+ specialise = false,
+ fail_safe_function_flattening = false,
+ no_higher_order_predicate = [],
+ inductify = false,
+ detect_switches = true,
+ compilation = Predicate_Compile_Aux.Pred
+}
+
fun values ctxt soln t_compr =
let
+ val options = code_options_of (ProofContext.theory_of ctxt)
val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t
| _ => error ("Not a set comprehension: " ^ Syntax.string_of_term ctxt t_compr);
val (body, Ts, fp) = HOLogic.strip_psplits split;
@@ -360,10 +705,24 @@
case try (map (fst o dest_Free)) all_args of
SOME vs => vs
| NONE => error ("Not only free variables in " ^ commas (map (Syntax.string_of_term ctxt) all_args))
+ val _ = tracing "Preprocessing specification..."
+ val T = Sign.the_const_type (ProofContext.theory_of ctxt) name
+ val t = Const (name, T)
+ val thy' =
+ Theory.copy (ProofContext.theory_of ctxt)
+ |> Predicate_Compile.preprocess preprocess_options t
+ val ctxt' = ProofContext.init_global thy'
val _ = tracing "Generating prolog program..."
- val (p, constant_table) = generate ctxt [name]
+ val (p, constant_table) = generate (#ensure_groundness options) ctxt' name
+ |> (if #ensure_groundness options then
+ add_ground_predicates ctxt' (#limited_types options)
+ else I)
+ |> add_limited_predicates (#limited_predicates options)
+ |> apfst (fold replace (#replacing options))
+ |> apfst (reorder_manually (#manual_reorder options))
val _ = tracing "Running prolog program..."
- val tss = run p (translate_const constant_table name) (map first_upper vnames) soln
+ val tss = run (#prolog_system options)
+ p (translate_const constant_table name) (map first_upper vnames) soln
val _ = tracing "Restoring terms..."
val empty = Const("Orderings.bot_class.bot", fastype_of t_compr)
fun mk_insert x S =
@@ -379,17 +738,18 @@
mk_set_compr (t :: in_insert) ts xs
else
let
- val uu as (uuN, uuT) = singleton (Variable.variant_frees ctxt [t]) ("uu", fastype_of t)
+ val uu as (uuN, uuT) = singleton (Variable.variant_frees ctxt' [t]) ("uu", fastype_of t)
val set_compr =
HOLogic.mk_Collect (uuN, uuT, fold (fn (s, T) => fn t => HOLogic.mk_exists (s, T, t))
frees (HOLogic.mk_conj (HOLogic.mk_eq (Free uu, t), @{term "True"})))
in
- set_compr :: (if null in_insert then xs else (fold mk_insert in_insert empty) :: xs)
+ mk_set_compr [] ts
+ (set_compr :: (if null in_insert then xs else (fold mk_insert in_insert empty) :: xs))
end
end
in
foldl1 (HOLogic.mk_binop @{const_name sup}) (mk_set_compr []
- (map (fn ts => HOLogic.mk_tuple (map (restore_term ctxt constant_table) (ts ~~ Ts))) tss) [])
+ (map (fn ts => HOLogic.mk_tuple (map (restore_term ctxt' constant_table) (ts ~~ Ts))) tss) [])
end
fun values_cmd print_modes soln raw_t state =
@@ -416,4 +776,57 @@
>> (fn ((print_modes, soln), t) => Toplevel.keep
(values_cmd print_modes soln t))); (*FIXME does not preserve the previous functionality*)
+(* quickcheck generator *)
+
+(* FIXME: large copy of Predicate_Compile_Quickcheck - refactor out commons *)
+
+fun strip_imp_prems (Const(@{const_name HOL.implies}, _) $ A $ B) = A :: strip_imp_prems B
+ | strip_imp_prems _ = [];
+
+fun strip_imp_concl (Const(@{const_name HOL.implies}, _) $ A $ B) = strip_imp_concl B
+ | strip_imp_concl A = A : term;
+
+fun strip_horn A = (strip_imp_prems A, strip_imp_concl A);
+
+fun quickcheck ctxt report t size =
+ let
+ val options = code_options_of (ProofContext.theory_of ctxt)
+ val thy = Theory.copy (ProofContext.theory_of ctxt)
+ val (vs, t') = strip_abs t
+ val vs' = Variable.variant_frees ctxt [] vs
+ val Ts = map snd vs'
+ val t'' = subst_bounds (map Free (rev vs'), t')
+ val (prems, concl) = strip_horn t''
+ val constname = "quickcheck"
+ val full_constname = Sign.full_bname thy constname
+ val constT = Ts ---> @{typ bool}
+ val thy1 = Sign.add_consts_i [(Binding.name constname, constT, NoSyn)] thy
+ val const = Const (full_constname, constT)
+ val t = Logic.list_implies
+ (map HOLogic.mk_Trueprop (prems @ [HOLogic.mk_not concl]),
+ HOLogic.mk_Trueprop (list_comb (Const (full_constname, constT), map Free vs')))
+ val tac = fn _ => Skip_Proof.cheat_tac thy1
+ val intro = Goal.prove (ProofContext.init_global thy1) (map fst vs') [] t tac
+ val thy2 = Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1
+ val thy3 = Predicate_Compile.preprocess preprocess_options const thy2
+ val ctxt' = ProofContext.init_global thy3
+ val _ = tracing "Generating prolog program..."
+ val (p, constant_table) = generate true ctxt' full_constname
+ |> add_ground_predicates ctxt' (#limited_types options)
+ |> add_limited_predicates (#limited_predicates options)
+ |> apfst (fold replace (#replacing options))
+ |> apfst (reorder_manually (#manual_reorder options))
+ val _ = tracing "Running prolog program..."
+ val tss = run (#prolog_system options)
+ p (translate_const constant_table full_constname) (map fst vs') (SOME 1)
+ val _ = tracing "Restoring terms..."
+ val res =
+ case tss of
+ [ts] => SOME (map (restore_term ctxt' constant_table) (ts ~~ Ts))
+ | _ => NONE
+ val empty_report = ([], false)
+ in
+ (res, empty_report)
+ end;
+
end;
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile.ML Wed Sep 01 17:19:47 2010 +0200
@@ -176,9 +176,9 @@
val t = Const (const, T)
val options = extract_options (((expected_modes, proposed_modes), raw_options), const)
in
- if (is_inductify options) then
+ if is_inductify options then
let
- val lthy' = Local_Theory.theory (preprocess options t) lthy
+ val lthy' = Local_Theory.background_theory (preprocess options t) lthy
val const =
case Predicate_Compile_Fun.pred_of_function (ProofContext.theory_of lthy') const of
SOME c => c
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Wed Sep 01 17:19:47 2010 +0200
@@ -405,13 +405,13 @@
(* general syntactic functions *)
(*Like dest_conj, but flattens conjunctions however nested*)
-fun conjuncts_aux (Const (@{const_name "op &"}, _) $ t $ t') conjs = conjuncts_aux t (conjuncts_aux t' conjs)
+fun conjuncts_aux (Const (@{const_name HOL.conj}, _) $ t $ t') conjs = conjuncts_aux t (conjuncts_aux t' conjs)
| conjuncts_aux t conjs = t::conjs;
fun conjuncts t = conjuncts_aux t [];
fun is_equationlike_term (Const ("==", _) $ _ $ _) = true
- | is_equationlike_term (Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op ="}, _) $ _ $ _)) = true
+ | is_equationlike_term (Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ _)) = true
| is_equationlike_term _ = false
val is_equationlike = is_equationlike_term o prop_of
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Wed Sep 01 17:19:47 2010 +0200
@@ -63,6 +63,19 @@
val add_random_dseq_equations : options -> string list -> theory -> theory
val add_new_random_dseq_equations : options -> string list -> theory -> theory
val mk_tracing : string -> term -> term
+ val prepare_intrs : options -> compilation -> theory -> string list -> thm list ->
+ ((string * typ) list * string list * string list * (string * mode list) list *
+ (string * (Term.term list * Predicate_Compile_Aux.indprem list) list) list)
+ type mode_analysis_options = {use_random : bool, reorder_premises : bool, infer_pos_and_neg_modes : bool}
+ datatype mode_derivation = Mode_App of mode_derivation * mode_derivation | Context of mode
+ | Mode_Pair of mode_derivation * mode_derivation | Term of mode
+ 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 -> compilation -> (string * typ) list -> (string * mode list) list ->
+ string list -> (string * (Term.term list * Predicate_Compile_Aux.indprem list) list) list ->
+ theory -> ((moded_clause list pred_mode_table * string list) * theory)
end;
structure Predicate_Compile_Core : PREDICATE_COMPILE_CORE =
@@ -524,7 +537,7 @@
fun dest_conjunct_prem th =
case HOLogic.dest_Trueprop (prop_of th) of
- (Const (@{const_name "op &"}, _) $ t $ t') =>
+ (Const (@{const_name HOL.conj}, _) $ t $ t') =>
dest_conjunct_prem (th RS @{thm conjunct1})
@ dest_conjunct_prem (th RS @{thm conjunct2})
| _ => [th]
@@ -587,7 +600,7 @@
fun preprocess_elim ctxt elimrule =
let
- fun replace_eqs (Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op ="}, T) $ lhs $ rhs)) =
+ fun replace_eqs (Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, T) $ lhs $ rhs)) =
HOLogic.mk_Trueprop (Const (@{const_name Predicate.eq}, T) $ lhs $ rhs)
| replace_eqs t = t
val thy = ProofContext.theory_of ctxt
@@ -730,9 +743,7 @@
type T = (mode * (compilation_funs -> typ -> term)) list Symtab.table;
val empty = Symtab.empty;
val extend = I;
- val merge = Symtab.merge ((K true)
- : ((mode * (compilation_funs -> typ -> term)) list *
- (mode * (compilation_funs -> typ -> term)) list -> bool));
+ fun merge data : T = Symtab.merge (K true) data;
);
fun alternative_compilation_of_global thy pred_name mode =
@@ -3033,12 +3044,13 @@
"adding alternative introduction rules for code generation of inductive predicates"
(* TODO: make Theory_Data to Generic_Data & remove duplication of local theory and theory *)
+(* FIXME ... this is important to avoid changing the background theory below *)
fun generic_code_pred prep_const options raw_const lthy =
let
val thy = ProofContext.theory_of lthy
val const = prep_const thy raw_const
val ctxt = ProofContext.init_global thy
- val lthy' = Local_Theory.theory (PredData.map
+ val lthy' = Local_Theory.background_theory (PredData.map
(extend (fetch_pred_data ctxt) (depending_preds_of ctxt) const)) lthy
val thy' = ProofContext.theory_of lthy'
val ctxt' = ProofContext.init_global thy'
@@ -3063,7 +3075,7 @@
val global_thms = ProofContext.export goal_ctxt
(ProofContext.init_global (ProofContext.theory_of goal_ctxt)) (map the_single thms)
in
- goal_ctxt |> Local_Theory.theory (fold set_elim global_thms #>
+ goal_ctxt |> Local_Theory.background_theory (fold set_elim global_thms #>
((case compilation options of
Pred => add_equations
| DSeq => add_dseq_equations
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Wed Sep 01 17:19:47 2010 +0200
@@ -111,7 +111,7 @@
fun mk_meta_equation th =
case prop_of th of
- Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op ="}, _) $ _ $ _) => th RS @{thm eq_reflection}
+ Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.eq}, _) $ _ $ _) => th RS @{thm eq_reflection}
| _ => th
val meta_fun_cong = @{lemma "f == g ==> f x == g x" by simp}
@@ -217,12 +217,12 @@
@{const_name "==>"},
@{const_name Trueprop},
@{const_name Not},
- @{const_name "op ="},
- @{const_name "op -->"},
+ @{const_name HOL.eq},
+ @{const_name HOL.implies},
@{const_name All},
@{const_name Ex},
- @{const_name "op &"},
- @{const_name "op |"}]
+ @{const_name HOL.conj},
+ @{const_name HOL.disj}]
fun special_cases (c, T) = member (op =) [
@{const_name Product_Type.Unity},
@@ -275,7 +275,6 @@
else
let
val specs = get_specification options thy t
- (*|> Predicate_Compile_Set.unfold_set_notation*)
(*val _ = print_specification options thy constname specs*)
val us = defiants_of specs
in
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML Wed Sep 01 17:19:47 2010 +0200
@@ -21,8 +21,7 @@
structure Fun_Pred = Theory_Data
(
type T = (term * term) Item_Net.T;
- val empty = Item_Net.init ((op aconv o pairself fst) : (term * term) * (term * term) -> bool)
- (single o fst);
+ val empty : T = Item_Net.init (op aconv o pairself fst) (single o fst);
val extend = I;
val merge = Item_Net.merge;
)
@@ -352,13 +351,17 @@
|> map (fn (resargs, (names', prems')) =>
let
val prem' = HOLogic.mk_Trueprop (mk_lit (list_comb (P, resargs)))
- in (prem'::prems', names') end)
+ in (prems' @ [prem'], names') end)
end
val intro_ts' = folds_map rewrite prems frees
|> maps (fn (prems', frees') =>
rewrite concl frees'
- |> map (fn (concl'::conclprems, _) =>
- Logic.list_implies ((flat prems') @ conclprems, concl')))
+ |> map (fn (conclprems, _) =>
+ let
+ val (conclprems', concl') = split_last conclprems
+ in
+ Logic.list_implies ((flat prems') @ conclprems', concl')
+ end))
(*val _ = tracing ("Rewritten intro to " ^
commas (map (Syntax.string_of_term_global thy) intro_ts'))*)
in
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Wed Sep 01 17:19:47 2010 +0200
@@ -89,8 +89,8 @@
fun is_compound ((Const (@{const_name Not}, _)) $ t) =
error "is_compound: Negation should not occur; preprocessing is defect"
| is_compound ((Const (@{const_name Ex}, _)) $ _) = true
- | is_compound ((Const (@{const_name "op |"}, _)) $ _ $ _) = true
- | is_compound ((Const (@{const_name "op &"}, _)) $ _ $ _) =
+ | is_compound ((Const (@{const_name HOL.disj}, _)) $ _ $ _) = true
+ | is_compound ((Const (@{const_name HOL.conj}, _)) $ _ $ _) =
error "is_compound: Conjunction should not occur; preprocessing is defect"
| is_compound _ = false
@@ -250,7 +250,7 @@
fun split_conjs thy t =
let
- fun split_conjunctions (Const (@{const_name "op &"}, _) $ t1 $ t2) =
+ fun split_conjunctions (Const (@{const_name HOL.conj}, _) $ t1 $ t2) =
(split_conjunctions t1) @ (split_conjunctions t2)
| split_conjunctions t = [t]
in
@@ -259,7 +259,8 @@
fun rewrite_intros thy =
Simplifier.full_simplify (HOL_basic_ss addsimps [@{thm all_not_ex}])
- #> Simplifier.full_simplify (HOL_basic_ss addsimps @{thms bool_simps} addsimps @{thms nnf_simps})
+ #> Simplifier.full_simplify
+ (HOL_basic_ss addsimps (tl @{thms bool_simps}) addsimps @{thms nnf_simps})
#> map_term thy (maps_premises (split_conjs thy))
fun print_specs options thy msg ths =
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML Wed Sep 01 17:19:47 2010 +0200
@@ -168,10 +168,10 @@
mk_split_lambda' xs t
end;
-fun strip_imp_prems (Const(@{const_name "op -->"}, _) $ A $ B) = A :: strip_imp_prems B
+fun strip_imp_prems (Const(@{const_name HOL.implies}, _) $ A $ B) = A :: strip_imp_prems B
| strip_imp_prems _ = [];
-fun strip_imp_concl (Const(@{const_name "op -->"}, _) $ A $ B) = strip_imp_concl B
+fun strip_imp_concl (Const(@{const_name HOL.implies}, _) $ A $ B) = strip_imp_concl B
| strip_imp_concl A = A : term;
fun strip_horn A = (strip_imp_prems A, strip_imp_concl A);
@@ -185,10 +185,9 @@
fun compile_term compilation options ctxt t =
let
- val ctxt' = ProofContext.theory (Context.copy_thy) ctxt
- val thy = (ProofContext.theory_of ctxt')
+ val thy = Theory.copy (ProofContext.theory_of ctxt)
val (vs, t') = strip_abs t
- val vs' = Variable.variant_frees ctxt' [] vs
+ val vs' = Variable.variant_frees ctxt [] vs
val t'' = subst_bounds (map Free (rev vs'), t')
val (prems, concl) = strip_horn t''
val constname = "pred_compile_quickcheck"
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_specialisation.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_specialisation.ML Wed Sep 01 17:19:47 2010 +0200
@@ -18,8 +18,7 @@
structure Specialisations = Theory_Data
(
type T = (term * term) Item_Net.T;
- val empty = Item_Net.init ((op aconv o pairself fst) : (term * term) * (term * term) -> bool)
- (single o fst);
+ val empty : T = Item_Net.init (op aconv o pairself fst) (single o fst);
val extend = I;
val merge = Item_Net.merge;
)
--- a/src/HOL/Tools/Qelim/cooper.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Qelim/cooper.ML Wed Sep 01 17:19:47 2010 +0200
@@ -28,7 +28,7 @@
@{term "op * :: int => _"}, @{term "op * :: nat => _"},
@{term "op div :: int => _"}, @{term "op div :: nat => _"},
@{term "op mod :: int => _"}, @{term "op mod :: nat => _"},
- @{term "op &"}, @{term "op |"}, @{term "op -->"},
+ @{term HOL.conj}, @{term HOL.disj}, @{term HOL.implies},
@{term "op = :: int => _"}, @{term "op = :: nat => _"}, @{term "op = :: bool => _"},
@{term "op < :: int => _"}, @{term "op < :: nat => _"},
@{term "op <= :: int => _"}, @{term "op <= :: nat => _"},
@@ -120,10 +120,10 @@
fun whatis x ct =
( case (term_of ct) of
- Const(@{const_name "op &"},_)$_$_ => And (Thm.dest_binop ct)
-| Const (@{const_name "op |"},_)$_$_ => Or (Thm.dest_binop ct)
-| Const (@{const_name "op ="},_)$y$_ => if term_of x aconv y then Eq (Thm.dest_arg ct) else Nox
-| Const (@{const_name Not},_) $ (Const (@{const_name "op ="},_)$y$_) =>
+ Const(@{const_name HOL.conj},_)$_$_ => And (Thm.dest_binop ct)
+| Const (@{const_name HOL.disj},_)$_$_ => Or (Thm.dest_binop ct)
+| Const (@{const_name HOL.eq},_)$y$_ => if term_of x aconv y then Eq (Thm.dest_arg ct) else Nox
+| Const (@{const_name Not},_) $ (Const (@{const_name HOL.eq},_)$y$_) =>
if term_of x aconv y then NEq (funpow 2 Thm.dest_arg ct) else Nox
| Const (@{const_name Orderings.less}, _) $ y$ z =>
if term_of x aconv y then Lt (Thm.dest_arg ct)
@@ -274,7 +274,7 @@
| lin vs (Const (@{const_name Not},T)$t) = Const (@{const_name Not},T)$ (lin vs t)
| lin (vs as x::_) (Const(@{const_name Rings.dvd},_)$d$t) =
HOLogic.mk_binrel @{const_name Rings.dvd} (numeral1 abs d, lint vs t)
- | lin (vs as x::_) ((b as Const(@{const_name "op ="},_))$s$t) =
+ | lin (vs as x::_) ((b as Const(@{const_name HOL.eq},_))$s$t) =
(case lint vs (subC$t$s) of
(t as a$(m$c$y)$r) =>
if x <> y then b$zero$t
@@ -345,7 +345,7 @@
case (term_of t) of
Const(s,_)$(Const(@{const_name Groups.times},_)$c$y)$ _ =>
if x aconv y andalso member (op =)
- ["op =", @{const_name Orderings.less}, @{const_name Orderings.less_eq}] s
+ [@{const_name HOL.eq}, @{const_name Orderings.less}, @{const_name Orderings.less_eq}] s
then (ins (dest_number c) acc,dacc) else (acc,dacc)
| Const(s,_)$_$(Const(@{const_name Groups.times},_)$c$y) =>
if x aconv y andalso member (op =)
@@ -353,8 +353,8 @@
then (ins (dest_number c) acc, dacc) else (acc,dacc)
| Const(@{const_name Rings.dvd},_)$_$(Const(@{const_name Groups.plus},_)$(Const(@{const_name Groups.times},_)$c$y)$_) =>
if x aconv y then (acc,ins (dest_number c) dacc) else (acc,dacc)
- | Const(@{const_name "op &"},_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
- | Const(@{const_name "op |"},_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
+ | Const(@{const_name HOL.conj},_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
+ | Const(@{const_name HOL.disj},_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
| Const (@{const_name Not},_)$_ => h (acc,dacc) (Thm.dest_arg t)
| _ => (acc, dacc)
val (cs,ds) = h ([],[]) p
@@ -382,12 +382,12 @@
end
fun unit_conv t =
case (term_of t) of
- Const(@{const_name "op &"},_)$_$_ => Conv.binop_conv unit_conv t
- | Const(@{const_name "op |"},_)$_$_ => Conv.binop_conv unit_conv t
+ Const(@{const_name HOL.conj},_)$_$_ => Conv.binop_conv unit_conv t
+ | Const(@{const_name HOL.disj},_)$_$_ => Conv.binop_conv unit_conv t
| Const (@{const_name Not},_)$_ => Conv.arg_conv unit_conv t
| Const(s,_)$(Const(@{const_name Groups.times},_)$c$y)$ _ =>
if x=y andalso member (op =)
- ["op =", @{const_name Orderings.less}, @{const_name Orderings.less_eq}] s
+ [@{const_name HOL.eq}, @{const_name Orderings.less}, @{const_name Orderings.less_eq}] s
then cv (l div dest_number c) t else Thm.reflexive t
| Const(s,_)$_$(Const(@{const_name Groups.times},_)$c$y) =>
if x=y andalso member (op =)
@@ -569,7 +569,7 @@
fun add_bools t =
let
val ops = [@{term "op = :: int => _"}, @{term "op < :: int => _"}, @{term "op <= :: int => _"},
- @{term "op &"}, @{term "op |"}, @{term "op -->"}, @{term "op = :: bool => _"},
+ @{term HOL.conj}, @{term HOL.disj}, @{term HOL.implies}, @{term "op = :: bool => _"},
@{term "Not"}, @{term "All :: (int => _) => _"},
@{term "Ex :: (int => _) => _"}, @{term "True"}, @{term "False"}];
val is_op = member (op =) ops;
@@ -612,11 +612,11 @@
fun fm_of_term ps vs (Const (@{const_name True}, _)) = Proc.T
| fm_of_term ps vs (Const (@{const_name False}, _)) = Proc.F
- | fm_of_term ps vs (Const (@{const_name "op &"}, _) $ t1 $ t2) =
+ | fm_of_term ps vs (Const (@{const_name HOL.conj}, _) $ t1 $ t2) =
Proc.And (fm_of_term ps vs t1, fm_of_term ps vs t2)
- | fm_of_term ps vs (Const (@{const_name "op |"}, _) $ t1 $ t2) =
+ | fm_of_term ps vs (Const (@{const_name HOL.disj}, _) $ t1 $ t2) =
Proc.Or (fm_of_term ps vs t1, fm_of_term ps vs t2)
- | fm_of_term ps vs (Const (@{const_name "op -->"}, _) $ t1 $ t2) =
+ | fm_of_term ps vs (Const (@{const_name HOL.implies}, _) $ t1 $ t2) =
Proc.Imp (fm_of_term ps vs t1, fm_of_term ps vs t2)
| fm_of_term ps vs (@{term "op = :: bool => _ "} $ t1 $ t2) =
Proc.Iff (fm_of_term ps vs t1, fm_of_term ps vs t2)
@@ -679,15 +679,17 @@
end;
-val (_, oracle) = Context.>>> (Context.map_theory_result (Thm.add_oracle (Binding.name "cooper",
- (fn (ctxt, t) => (Thm.cterm_of (ProofContext.theory_of ctxt) o Logic.mk_equals o pairself HOLogic.mk_Trueprop)
- (t, procedure t)))));
+val (_, oracle) = Context.>>> (Context.map_theory_result
+ (Thm.add_oracle (@{binding cooper},
+ (fn (ctxt, t) =>
+ (Thm.cterm_of (ProofContext.theory_of ctxt) o Logic.mk_equals o pairself HOLogic.mk_Trueprop)
+ (t, procedure t)))));
val comp_ss = HOL_ss addsimps @{thms semiring_norm};
fun strip_objimp ct =
(case Thm.term_of ct of
- Const (@{const_name "op -->"}, _) $ _ $ _ =>
+ Const (@{const_name HOL.implies}, _) $ _ $ _ =>
let val (A, B) = Thm.dest_binop ct
in A :: strip_objimp B end
| _ => [ct]);
@@ -712,7 +714,7 @@
val qs = filter P ps
val q = if P c then c else @{cterm "False"}
val ng = fold_rev (fn (a,v) => fn t => Thm.capply a (Thm.cabs v t)) qvs
- (fold_rev (fn p => fn q => Thm.capply (Thm.capply @{cterm "op -->"} p) q) qs q)
+ (fold_rev (fn p => fn q => Thm.capply (Thm.capply @{cterm HOL.implies} p) q) qs q)
val g = Thm.capply (Thm.capply @{cterm "op ==>"} (Thm.capply @{cterm "Trueprop"} ng)) p'
val ntac = (case qs of [] => q aconvc @{cterm "False"}
| _ => false)
--- a/src/HOL/Tools/Qelim/qelim.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Qelim/qelim.ML Wed Sep 01 17:19:47 2010 +0200
@@ -25,8 +25,8 @@
case (term_of p) of
Const(s,T)$_$_ =>
if domain_type T = HOLogic.boolT
- andalso member (op =) [@{const_name "op &"}, @{const_name "op |"},
- @{const_name "op -->"}, @{const_name "op ="}] s
+ andalso member (op =) [@{const_name HOL.conj}, @{const_name HOL.disj},
+ @{const_name HOL.implies}, @{const_name HOL.eq}] s
then binop_conv (conv env) p
else atcv env p
| Const(@{const_name Not},_)$_ => arg_conv (conv env) p
--- a/src/HOL/Tools/Quotient/quotient_info.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Quotient/quotient_info.ML Wed Sep 01 17:19:47 2010 +0200
@@ -56,10 +56,12 @@
type maps_info = {mapfun: string, relmap: string}
structure MapsData = Theory_Data
- (type T = maps_info Symtab.table
- val empty = Symtab.empty
- val extend = I
- fun merge data = Symtab.merge (K true) data)
+(
+ type T = maps_info Symtab.table
+ val empty = Symtab.empty
+ val extend = I
+ fun merge data = Symtab.merge (K true) data
+)
fun maps_defined thy s =
Symtab.defined (MapsData.get thy) s
@@ -70,7 +72,7 @@
| NONE => raise NotFound
fun maps_update_thy k minfo = MapsData.map (Symtab.update (k, minfo))
-fun maps_update k minfo = ProofContext.theory (maps_update_thy k minfo)
+fun maps_update k minfo = ProofContext.background_theory (maps_update_thy k minfo) (* FIXME *)
fun maps_attribute_aux s minfo = Thm.declaration_attribute
(fn _ => Context.mapping (maps_update_thy s minfo) (maps_update s minfo))
@@ -120,10 +122,12 @@
type quotdata_info = {qtyp: typ, rtyp: typ, equiv_rel: term, equiv_thm: thm}
structure QuotData = Theory_Data
- (type T = quotdata_info Symtab.table
- val empty = Symtab.empty
- val extend = I
- fun merge data = Symtab.merge (K true) data)
+(
+ type T = quotdata_info Symtab.table
+ val empty = Symtab.empty
+ val extend = I
+ fun merge data = Symtab.merge (K true) data
+)
fun transform_quotdata phi {qtyp, rtyp, equiv_rel, equiv_thm} =
{qtyp = Morphism.typ phi qtyp,
@@ -174,10 +178,12 @@
for example given "nat fset" we need to find "'a fset";
but overloaded constants share the same name *)
structure QConstsData = Theory_Data
- (type T = (qconsts_info list) Symtab.table
- val empty = Symtab.empty
- val extend = I
- val merge = Symtab.merge_list qconsts_info_eq)
+(
+ type T = qconsts_info list Symtab.table
+ val empty = Symtab.empty
+ val extend = I
+ val merge = Symtab.merge_list qconsts_info_eq
+)
fun transform_qconsts phi {qconst, rconst, def} =
{qconst = Morphism.term phi qconst,
@@ -229,39 +235,49 @@
(* equivalence relation theorems *)
structure EquivRules = Named_Thms
- (val name = "quot_equiv"
- val description = "Equivalence relation theorems.")
+(
+ val name = "quot_equiv"
+ val description = "equivalence relation theorems"
+)
val equiv_rules_get = EquivRules.get
val equiv_rules_add = EquivRules.add
(* respectfulness theorems *)
structure RspRules = Named_Thms
- (val name = "quot_respect"
- val description = "Respectfulness theorems.")
+(
+ val name = "quot_respect"
+ val description = "respectfulness theorems"
+)
val rsp_rules_get = RspRules.get
val rsp_rules_add = RspRules.add
(* preservation theorems *)
structure PrsRules = Named_Thms
- (val name = "quot_preserve"
- val description = "Preservation theorems.")
+(
+ val name = "quot_preserve"
+ val description = "preservation theorems"
+)
val prs_rules_get = PrsRules.get
val prs_rules_add = PrsRules.add
(* id simplification theorems *)
structure IdSimps = Named_Thms
- (val name = "id_simps"
- val description = "Identity simp rules for maps.")
+(
+ val name = "id_simps"
+ val description = "identity simp rules for maps"
+)
val id_simps_get = IdSimps.get
(* quotient theorems *)
structure QuotientRules = Named_Thms
- (val name = "quot_thm"
- val description = "Quotient theorems.")
+(
+ val name = "quot_thm"
+ val description = "quotient theorems"
+)
val quotient_rules_get = QuotientRules.get
val quotient_rules_add = QuotientRules.add
--- a/src/HOL/Tools/Quotient/quotient_tacs.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Quotient/quotient_tacs.ML Wed Sep 01 17:19:47 2010 +0200
@@ -12,11 +12,11 @@
val all_injection_tac: Proof.context -> int -> tactic
val clean_tac: Proof.context -> int -> tactic
- val descend_procedure_tac: Proof.context -> int -> tactic
- val descend_tac: Proof.context -> int -> tactic
+ val descend_procedure_tac: Proof.context -> thm list -> int -> tactic
+ val descend_tac: Proof.context -> thm list -> int -> tactic
- val lift_procedure_tac: Proof.context -> thm -> int -> tactic
- val lift_tac: Proof.context -> thm list -> int -> tactic
+ val lift_procedure_tac: Proof.context -> thm list -> thm -> int -> tactic
+ val lift_tac: Proof.context -> thm list -> thm list -> int -> tactic
val lifted: Proof.context -> typ list -> thm list -> thm -> thm
val lifted_attrib: attribute
@@ -338,7 +338,7 @@
=> rtac @{thm fun_relI} THEN' quot_true_tac ctxt unlam
(* (op =) (Ball...) (Ball...) ----> (op =) (...) (...) *)
-| (Const (@{const_name "op ="},_) $
+| (Const (@{const_name HOL.eq},_) $
(Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
(Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _))
=> rtac @{thm ball_rsp} THEN' dtac @{thm QT_all}
@@ -350,7 +350,7 @@
=> rtac @{thm fun_relI} THEN' quot_true_tac ctxt unlam
(* (op =) (Bex...) (Bex...) ----> (op =) (...) (...) *)
-| Const (@{const_name "op ="},_) $
+| Const (@{const_name HOL.eq},_) $
(Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
(Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _)
=> rtac @{thm bex_rsp} THEN' dtac @{thm QT_ex}
@@ -370,13 +370,13 @@
(Const(@{const_name Babs},_) $ (Const (@{const_name Respects}, _) $ _) $ _))
=> rtac @{thm babs_rsp} THEN' RANGE [quotient_tac ctxt]
-| Const (@{const_name "op ="},_) $ (R $ _ $ _) $ (_ $ _ $ _) =>
+| Const (@{const_name HOL.eq},_) $ (R $ _ $ _) $ (_ $ _ $ _) =>
(rtac @{thm refl} ORELSE'
(equals_rsp_tac R ctxt THEN' RANGE [
quot_true_tac ctxt (fst o dest_bcomb), quot_true_tac ctxt (snd o dest_bcomb)]))
(* reflexivity of operators arising from Cong_tac *)
-| Const (@{const_name "op ="},_) $ _ $ _ => rtac @{thm refl}
+| Const (@{const_name HOL.eq},_) $ _ $ _ => rtac @{thm refl}
(* respectfulness of constants; in particular of a simple relation *)
| _ $ (Const _) $ (Const _) (* fun_rel, list_rel, etc but not equality *)
@@ -606,9 +606,9 @@
val rtrm' = HOLogic.dest_Trueprop rtrm
val qtrm' = HOLogic.dest_Trueprop qtrm
val reg_goal = regularize_trm_chk ctxt (rtrm', qtrm')
- handle (LIFT_MATCH msg) => lift_match_error ctxt msg rtrm qtrm
+ handle LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm
val inj_goal = inj_repabs_trm_chk ctxt (reg_goal, qtrm')
- handle (LIFT_MATCH msg) => lift_match_error ctxt msg rtrm qtrm
+ handle LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm
in
Drule.instantiate' []
[SOME (cterm_of thy rtrm'),
@@ -618,10 +618,21 @@
end
+(* Since we use Ball and Bex during the lifting and descending,
+ we cannot deal with lemmas containing them, unless we unfold
+ them by default. *)
+
+val default_unfolds = @{thms Ball_def Bex_def}
+
+
(** descending as tactic **)
-fun descend_procedure_tac ctxt =
- Object_Logic.full_atomize_tac
+fun descend_procedure_tac ctxt simps =
+let
+ val ss = (mk_minimal_ss ctxt) addsimps (simps @ default_unfolds)
+in
+ full_simp_tac ss
+ THEN' Object_Logic.full_atomize_tac
THEN' gen_frees_tac ctxt
THEN' SUBGOAL (fn (goal, i) =>
let
@@ -631,11 +642,12 @@
in
rtac rule i
end)
+end
-fun descend_tac ctxt =
+fun descend_tac ctxt simps =
let
val mk_tac_raw =
- descend_procedure_tac ctxt
+ descend_procedure_tac ctxt simps
THEN' RANGE
[Object_Logic.rulify_tac THEN' (K all_tac),
regularize_tac ctxt,
@@ -650,15 +662,20 @@
(* the tactic leaves three subgoals to be proved *)
-fun lift_procedure_tac ctxt rthm =
- Object_Logic.full_atomize_tac
+fun lift_procedure_tac ctxt simps rthm =
+let
+ val ss = (mk_minimal_ss ctxt) addsimps (simps @ default_unfolds)
+in
+ full_simp_tac ss
+ THEN' Object_Logic.full_atomize_tac
THEN' gen_frees_tac ctxt
THEN' SUBGOAL (fn (goal, i) =>
let
(* full_atomize_tac contracts eta redexes,
so we do it also in the original theorem *)
val rthm' =
- rthm |> Drule.eta_contraction_rule
+ rthm |> full_simplify ss
+ |> Drule.eta_contraction_rule
|> Thm.forall_intr_frees
|> atomize_thm
@@ -666,32 +683,29 @@
in
(rtac rule THEN' rtac rthm') i
end)
-
+end
-fun lift_single_tac ctxt rthm =
- lift_procedure_tac ctxt rthm
+fun lift_single_tac ctxt simps rthm =
+ lift_procedure_tac ctxt simps rthm
THEN' RANGE
[ regularize_tac ctxt,
all_injection_tac ctxt,
clean_tac ctxt ]
-fun lift_tac ctxt rthms =
+fun lift_tac ctxt simps rthms =
Goal.conjunction_tac
- THEN' RANGE (map (lift_single_tac ctxt) rthms)
+ THEN' RANGE (map (lift_single_tac ctxt simps) rthms)
(* automated lifting with pre-simplification of the theorems;
for internal usage *)
fun lifted ctxt qtys simps rthm =
let
- val ss = (mk_minimal_ss ctxt) addsimps simps
- val rthm' = asm_full_simplify ss rthm
-
- val ((_, [rthm'']), ctxt') = Variable.import true [rthm'] ctxt
- val goal = derive_qtrm ctxt' qtys (prop_of rthm'')
+ val ((_, [rthm']), ctxt') = Variable.import true [rthm] ctxt
+ val goal = derive_qtrm ctxt' qtys (prop_of rthm')
in
Goal.prove ctxt' [] [] goal
- (K (EVERY1 [asm_full_simp_tac ss, lift_single_tac ctxt' rthm'']))
+ (K (HEADGOAL (lift_single_tac ctxt' simps rthm')))
|> singleton (ProofContext.export ctxt' ctxt)
end
--- a/src/HOL/Tools/Quotient/quotient_term.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Quotient/quotient_term.ML Wed Sep 01 17:19:47 2010 +0200
@@ -267,7 +267,7 @@
map_types (Envir.subst_type ty_inst) trm
end
-fun is_eq (Const (@{const_name "op ="}, _)) = true
+fun is_eq (Const (@{const_name HOL.eq}, _)) = true
| is_eq _ = false
fun mk_rel_compose (trm1, trm2) =
@@ -485,7 +485,7 @@
end
| (Const (@{const_name Ex1}, ty) $ (Abs (_, _,
- (Const (@{const_name "op &"}, _) $ (Const (@{const_name Set.member}, _) $ _ $
+ (Const (@{const_name HOL.conj}, _) $ (Const (@{const_name Set.member}, _) $ _ $
(Const (@{const_name Respects}, _) $ resrel)) $ (t $ _)))),
Const (@{const_name Ex1}, ty') $ t') =>
let
@@ -539,12 +539,12 @@
end
| (* equalities need to be replaced by appropriate equivalence relations *)
- (Const (@{const_name "op ="}, ty), Const (@{const_name "op ="}, ty')) =>
+ (Const (@{const_name HOL.eq}, ty), Const (@{const_name HOL.eq}, ty')) =>
if ty = ty' then rtrm
else equiv_relation ctxt (domain_type ty, domain_type ty')
| (* in this case we just check whether the given equivalence relation is correct *)
- (rel, Const (@{const_name "op ="}, ty')) =>
+ (rel, Const (@{const_name HOL.eq}, ty')) =>
let
val rel_ty = fastype_of rel
val rel' = equiv_relation_chk ctxt (domain_type rel_ty, domain_type ty')
@@ -680,7 +680,7 @@
if T = T' then rtrm
else mk_repabs ctxt (T, T') rtrm
- | (_, Const (@{const_name "op ="}, _)) => rtrm
+ | (_, Const (@{const_name HOL.eq}, _)) => rtrm
| (_, Const (_, T')) =>
let
--- a/src/HOL/Tools/Quotient/quotient_typ.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Quotient/quotient_typ.ML Wed Sep 01 17:19:47 2010 +0200
@@ -77,7 +77,8 @@
Typedef.add_typedef false NONE (qty_name, vs, mx)
(typedef_term rel rty lthy) NONE typedef_tac lthy
*)
- Local_Theory.theory_result
+(* FIXME should really use local typedef here *)
+ Local_Theory.background_theory_result
(Typedef.add_typedef_global false NONE
(qty_name, map (rpair dummyS) vs, mx)
(typedef_term rel rty lthy)
--- a/src/HOL/Tools/SMT/smt_monomorph.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/SMT/smt_monomorph.ML Wed Sep 01 17:19:47 2010 +0200
@@ -16,7 +16,7 @@
val ignored = member (op =) [
@{const_name All}, @{const_name Ex}, @{const_name Let}, @{const_name If},
- @{const_name "op ="}, @{const_name zero_class.zero},
+ @{const_name HOL.eq}, @{const_name zero_class.zero},
@{const_name one_class.one}, @{const_name number_of}]
fun is_const f (n, T) = not (ignored n) andalso f T
--- a/src/HOL/Tools/SMT/smt_normalize.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/SMT/smt_normalize.ML Wed Sep 01 17:19:47 2010 +0200
@@ -252,7 +252,7 @@
fun norm_def ctxt thm =
(case Thm.prop_of thm of
- @{term Trueprop} $ (Const (@{const_name "op ="}, _) $ _ $ Abs _) =>
+ @{term Trueprop} $ (Const (@{const_name HOL.eq}, _) $ _ $ Abs _) =>
norm_def ctxt (thm RS @{thm fun_cong})
| Const (@{const_name "=="}, _) $ _ $ Abs _ =>
norm_def ctxt (thm RS @{thm meta_eq_to_obj_eq})
--- a/src/HOL/Tools/SMT/smt_solver.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/SMT/smt_solver.ML Wed Sep 01 17:19:47 2010 +0200
@@ -310,18 +310,18 @@
(* setup *)
val setup =
- Attrib.setup (Binding.name "smt_solver")
+ Attrib.setup @{binding smt_solver}
(Scan.lift (Parse.$$$ "=" |-- Args.name) >>
(Thm.declaration_attribute o K o select_solver))
"SMT solver configuration" #>
setup_timeout #>
setup_trace #>
setup_fixed_certificates #>
- Attrib.setup (Binding.name "smt_certificates")
+ Attrib.setup @{binding smt_certificates}
(Scan.lift (Parse.$$$ "=" |-- Args.name) >>
(Thm.declaration_attribute o K o select_certificates))
"SMT certificates" #>
- Method.setup (Binding.name "smt") smt_method
+ Method.setup @{binding smt} smt_method
"Applies an SMT solver to the current goal."
--- a/src/HOL/Tools/SMT/smtlib_interface.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/SMT/smtlib_interface.ML Wed Sep 01 17:19:47 2010 +0200
@@ -159,15 +159,15 @@
fun conn @{const_name True} = SOME "true"
| conn @{const_name False} = SOME "false"
| conn @{const_name Not} = SOME "not"
- | conn @{const_name "op &"} = SOME "and"
- | conn @{const_name "op |"} = SOME "or"
- | conn @{const_name "op -->"} = SOME "implies"
- | conn @{const_name "op ="} = SOME "iff"
+ | conn @{const_name HOL.conj} = SOME "and"
+ | conn @{const_name HOL.disj} = SOME "or"
+ | conn @{const_name HOL.implies} = SOME "implies"
+ | conn @{const_name HOL.eq} = SOME "iff"
| conn @{const_name If} = SOME "if_then_else"
| conn _ = NONE
fun pred @{const_name distinct} _ = SOME "distinct"
- | pred @{const_name "op ="} _ = SOME "="
+ | pred @{const_name HOL.eq} _ = SOME "="
| pred @{const_name term_eq} _ = SOME "="
| pred @{const_name less} T = if_int_type T "<"
| pred @{const_name less_eq} T = if_int_type T "<="
--- a/src/HOL/Tools/SMT/z3_interface.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/SMT/z3_interface.ML Wed Sep 01 17:19:47 2010 +0200
@@ -170,13 +170,13 @@
val mk_true = @{cterm "~False"}
val mk_false = @{cterm False}
val mk_not = Thm.capply @{cterm Not}
-val mk_implies = Thm.mk_binop @{cterm "op -->"}
+val mk_implies = Thm.mk_binop @{cterm HOL.implies}
val mk_iff = Thm.mk_binop @{cterm "op = :: bool => _"}
fun mk_nary _ cu [] = cu
| mk_nary ct _ cts = uncurry (fold_rev (Thm.mk_binop ct)) (split_last cts)
-val eq = mk_inst_pair destT1 @{cpat "op ="}
+val eq = mk_inst_pair destT1 @{cpat HOL.eq}
fun mk_eq ct cu = Thm.mk_binop (instT' ct eq) ct cu
val if_term = mk_inst_pair (destT1 o destT2) @{cpat If}
@@ -216,8 +216,8 @@
(Sym ("true", _), []) => SOME mk_true
| (Sym ("false", _), []) => SOME mk_false
| (Sym ("not", _), [ct]) => SOME (mk_not ct)
- | (Sym ("and", _), _) => SOME (mk_nary @{cterm "op &"} mk_true cts)
- | (Sym ("or", _), _) => SOME (mk_nary @{cterm "op |"} mk_false cts)
+ | (Sym ("and", _), _) => SOME (mk_nary @{cterm HOL.conj} mk_true cts)
+ | (Sym ("or", _), _) => SOME (mk_nary @{cterm HOL.disj} mk_false cts)
| (Sym ("implies", _), [ct, cu]) => SOME (mk_implies ct cu)
| (Sym ("iff", _), [ct, cu]) => SOME (mk_iff ct cu)
| (Sym ("~", _), [ct, cu]) => SOME (mk_iff ct cu)
--- a/src/HOL/Tools/SMT/z3_proof_literals.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/SMT/z3_proof_literals.ML Wed Sep 01 17:19:47 2010 +0200
@@ -62,14 +62,14 @@
val is_neg = (fn @{term Not} $ _ => true | _ => false)
fun is_neg' f = (fn @{term Not} $ t => f t | _ => false)
val is_dneg = is_neg' is_neg
-val is_conj = (fn @{term "op &"} $ _ $ _ => true | _ => false)
-val is_disj = (fn @{term "op |"} $ _ $ _ => true | _ => false)
+val is_conj = (fn @{term HOL.conj} $ _ $ _ => true | _ => false)
+val is_disj = (fn @{term HOL.disj} $ _ $ _ => true | _ => false)
fun dest_disj_term' f = (fn
- @{term Not} $ (@{term "op |"} $ t $ u) => SOME (f t, f u)
+ @{term Not} $ (@{term HOL.disj} $ t $ u) => SOME (f t, f u)
| _ => NONE)
-val dest_conj_term = (fn @{term "op &"} $ t $ u => SOME (t, u) | _ => NONE)
+val dest_conj_term = (fn @{term HOL.conj} $ t $ u => SOME (t, u) | _ => NONE)
val dest_disj_term =
dest_disj_term' (fn @{term Not} $ t => t | t => @{term Not} $ t)
@@ -202,11 +202,11 @@
| comp_disj ((true, thm1), (false, thm2)) = comp2 disj3 thm1 thm2
| comp_disj ((true, thm1), (true, thm2)) = comp2 disj4 thm1 thm2
- fun dest_conj (@{term "op &"} $ t $ u) = ((false, t), (false, u))
+ fun dest_conj (@{term HOL.conj} $ t $ u) = ((false, t), (false, u))
| dest_conj t = raise TERM ("dest_conj", [t])
val neg = (fn @{term Not} $ t => (true, t) | t => (false, @{term Not} $ t))
- fun dest_disj (@{term Not} $ (@{term "op |"} $ t $ u)) = (neg t, neg u)
+ fun dest_disj (@{term Not} $ (@{term HOL.disj} $ t $ u)) = (neg t, neg u)
| dest_disj t = raise TERM ("dest_disj", [t])
val dnegE = T.precompose (single o d2 o d1) @{thm notnotD}
@@ -234,7 +234,7 @@
@{term Not} $ (@{term Not} $ t) => (T.compose dnegI, lookup t)
| @{term Not} $ (@{term "op = :: bool => _"} $ t $ (@{term Not} $ u)) =>
(T.compose negIffI, lookup (iff_const $ u $ t))
- | @{term Not} $ ((eq as Const (@{const_name "op ="}, _)) $ t $ u) =>
+ | @{term Not} $ ((eq as Const (@{const_name HOL.eq}, _)) $ t $ u) =>
let fun rewr lit = lit COMP @{thm not_sym}
in (rewr, lookup (@{term Not} $ (eq $ u $ t))) end
| _ =>
--- a/src/HOL/Tools/SMT/z3_proof_reconstruction.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/SMT/z3_proof_reconstruction.ML Wed Sep 01 17:19:47 2010 +0200
@@ -298,13 +298,13 @@
let val (ct1, ct2) = Thm.dest_binop (Thm.dest_arg ct)
in
(case Thm.term_of ct1 of
- @{term Not} $ (@{term "op &"} $ _ $ _) =>
+ @{term Not} $ (@{term HOL.conj} $ _ $ _) =>
prove disjI1 (Thm.dest_arg ct1, true) (ct2, true)
- | @{term "op &"} $ _ $ _ =>
+ | @{term HOL.conj} $ _ $ _ =>
prove disjI3 (Thm.capply @{cterm Not} ct2, false) (ct1, true)
- | @{term Not} $ (@{term "op |"} $ _ $ _) =>
+ | @{term Not} $ (@{term HOL.disj} $ _ $ _) =>
prove disjI3 (Thm.capply @{cterm Not} ct2, false) (ct1, false)
- | @{term "op |"} $ _ $ _ =>
+ | @{term HOL.disj} $ _ $ _ =>
prove disjI2 (Thm.capply @{cterm Not} ct1, false) (ct2, true)
| Const (@{const_name distinct}, _) $ _ =>
let
@@ -477,8 +477,8 @@
fun mono f (cp as (cl, _)) =
(case Term.head_of (Thm.term_of cl) of
- @{term "op &"} => prove_nary L.is_conj (prove_eq_exn f)
- | @{term "op |"} => prove_nary L.is_disj (prove_eq_exn f)
+ @{term HOL.conj} => prove_nary L.is_conj (prove_eq_exn f)
+ | @{term HOL.disj} => prove_nary L.is_disj (prove_eq_exn f)
| Const (@{const_name distinct}, _) => prove_distinct (prove_eq_safe f)
| _ => prove (prove_eq_safe f)) cp
in
--- a/src/HOL/Tools/SMT/z3_proof_tools.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/SMT/z3_proof_tools.ML Wed Sep 01 17:19:47 2010 +0200
@@ -196,10 +196,10 @@
| @{term True} => pair ct
| @{term False} => pair ct
| @{term Not} $ _ => abstr1 cvs ct
- | @{term "op &"} $ _ $ _ => abstr2 cvs ct
- | @{term "op |"} $ _ $ _ => abstr2 cvs ct
- | @{term "op -->"} $ _ $ _ => abstr2 cvs ct
- | Const (@{const_name "op ="}, _) $ _ $ _ => abstr2 cvs ct
+ | @{term HOL.conj} $ _ $ _ => abstr2 cvs ct
+ | @{term HOL.disj} $ _ $ _ => abstr2 cvs ct
+ | @{term HOL.implies} $ _ $ _ => abstr2 cvs ct
+ | Const (@{const_name HOL.eq}, _) $ _ $ _ => abstr2 cvs ct
| Const (@{const_name distinct}, _) $ _ =>
if ext_logic then abs_arg (abs_list abstr fresh_abstraction) cvs ct
else fresh_abstraction cvs ct
--- a/src/HOL/Tools/Sledgehammer/clausifier.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/clausifier.ML Wed Sep 01 17:19:47 2010 +0200
@@ -68,8 +68,8 @@
(*Universal quant: insert a free variable into body and continue*)
let val fname = Name.variant (OldTerm.add_term_names (p,[])) a
in dec_sko (subst_bound (Free(fname,T), p)) rhss end
- | dec_sko (@{const "op &"} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q
- | dec_sko (@{const "op |"} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q
+ | dec_sko (@{const HOL.conj} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q
+ | dec_sko (@{const HOL.disj} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q
| dec_sko (@{const Trueprop} $ p) rhss = dec_sko p rhss
| dec_sko _ rhss = rhss
in dec_sko (prop_of th) [] end;
@@ -83,7 +83,7 @@
(Cf. "extensionalize_term" in "Sledgehammer_Translate".) *)
fun extensionalize_theorem th =
case prop_of th of
- _ $ (Const (@{const_name "op ="}, Type (_, [Type (@{type_name fun}, _), _]))
+ _ $ (Const (@{const_name HOL.eq}, Type (_, [Type (@{type_name fun}, _), _]))
$ _ $ Abs (s, _, _)) => extensionalize_theorem (th RS fun_cong_all)
| _ => th
--- a/src/HOL/Tools/Sledgehammer/metis_clauses.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/metis_clauses.ML Wed Sep 01 17:19:47 2010 +0200
@@ -38,11 +38,10 @@
val const_prefix: string
val type_const_prefix: string
val class_prefix: string
- val union_all: ''a list list -> ''a list
val invert_const: string -> string
val ascii_of: string -> string
- val undo_ascii_of: string -> string
- val strip_prefix_and_undo_ascii: string -> string -> string option
+ val unascii_of: string -> string
+ val strip_prefix_and_unascii: string -> string -> string option
val make_bound_var : string -> string
val make_schematic_var : string * int -> string
val make_fixed_var : string -> string
@@ -95,10 +94,10 @@
val const_trans_table =
Symtab.make [(@{type_name Product_Type.prod}, "prod"),
(@{type_name Sum_Type.sum}, "sum"),
- (@{const_name "op ="}, "equal"),
- (@{const_name "op &"}, "and"),
- (@{const_name "op |"}, "or"),
- (@{const_name "op -->"}, "implies"),
+ (@{const_name HOL.eq}, "equal"),
+ (@{const_name HOL.conj}, "and"),
+ (@{const_name HOL.disj}, "or"),
+ (@{const_name HOL.implies}, "implies"),
(@{const_name Set.member}, "member"),
(@{const_name fequal}, "fequal"),
(@{const_name COMBI}, "COMBI"),
@@ -112,7 +111,7 @@
(* Invert the table of translations between Isabelle and ATPs. *)
val const_trans_table_inv =
- Symtab.update ("fequal", @{const_name "op ="})
+ Symtab.update ("fequal", @{const_name HOL.eq})
(Symtab.make (map swap (Symtab.dest const_trans_table)))
val invert_const = perhaps (Symtab.lookup const_trans_table_inv)
@@ -121,7 +120,7 @@
Alphanumeric characters are left unchanged.
The character _ goes to __
Characters in the range ASCII space to / go to _A to _P, respectively.
- Other printing characters go to _nnn where nnn is the decimal ASCII code.*)
+ Other characters go to _nnn where nnn is the decimal ASCII code.*)
val A_minus_space = Char.ord #"A" - Char.ord #" ";
fun stringN_of_int 0 _ = ""
@@ -132,9 +131,7 @@
else if c = #"_" then "__"
else if #" " <= c andalso c <= #"/"
then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space))
- else if Char.isPrint c
- then ("_" ^ stringN_of_int 3 (Char.ord c)) (*fixed width, in case more digits follow*)
- else ""
+ else ("_" ^ stringN_of_int 3 (Char.ord c)) (*fixed width, in case more digits follow*)
val ascii_of = String.translate ascii_of_c;
@@ -142,29 +139,28 @@
(*We don't raise error exceptions because this code can run inside the watcher.
Also, the errors are "impossible" (hah!)*)
-fun undo_ascii_aux rcs [] = String.implode(rev rcs)
- | undo_ascii_aux rcs [#"_"] = undo_ascii_aux (#"_"::rcs) [] (*ERROR*)
+fun unascii_aux rcs [] = String.implode(rev rcs)
+ | unascii_aux rcs [#"_"] = unascii_aux (#"_"::rcs) [] (*ERROR*)
(*Three types of _ escapes: __, _A to _P, _nnn*)
- | undo_ascii_aux rcs (#"_" :: #"_" :: cs) = undo_ascii_aux (#"_"::rcs) cs
- | undo_ascii_aux rcs (#"_" :: c :: cs) =
+ | unascii_aux rcs (#"_" :: #"_" :: cs) = unascii_aux (#"_"::rcs) cs
+ | unascii_aux rcs (#"_" :: c :: cs) =
if #"A" <= c andalso c<= #"P" (*translation of #" " to #"/"*)
- then undo_ascii_aux (Char.chr(Char.ord c - A_minus_space) :: rcs) cs
+ then unascii_aux (Char.chr(Char.ord c - A_minus_space) :: rcs) cs
else
let val digits = List.take (c::cs, 3) handle Subscript => []
in
case Int.fromString (String.implode digits) of
- NONE => undo_ascii_aux (c:: #"_"::rcs) cs (*ERROR*)
- | SOME n => undo_ascii_aux (Char.chr n :: rcs) (List.drop (cs, 2))
+ NONE => unascii_aux (c:: #"_"::rcs) cs (*ERROR*)
+ | SOME n => unascii_aux (Char.chr n :: rcs) (List.drop (cs, 2))
end
- | undo_ascii_aux rcs (c::cs) = undo_ascii_aux (c::rcs) cs;
-
-val undo_ascii_of = undo_ascii_aux [] o String.explode;
+ | unascii_aux rcs (c::cs) = unascii_aux (c::rcs) cs
+val unascii_of = unascii_aux [] o String.explode
(* If string s has the prefix s1, return the result of deleting it,
un-ASCII'd. *)
-fun strip_prefix_and_undo_ascii s1 s =
+fun strip_prefix_and_unascii s1 s =
if String.isPrefix s1 s then
- SOME (undo_ascii_of (String.extract (s, size s1, NONE)))
+ SOME (unascii_of (String.extract (s, size s1, NONE)))
else
NONE
@@ -189,8 +185,8 @@
SOME c' => c'
| NONE => ascii_of c
-(* "op =" MUST BE "equal" because it's built into ATPs. *)
-fun make_fixed_const @{const_name "op ="} = "equal"
+(* HOL.eq MUST BE "equal" because it's built into ATPs. *)
+fun make_fixed_const @{const_name HOL.eq} = "equal"
| make_fixed_const c = const_prefix ^ lookup_const c
fun make_fixed_type_const c = type_const_prefix ^ lookup_const c
@@ -434,7 +430,7 @@
fun literals_of_term1 args thy (@{const Trueprop} $ P) =
literals_of_term1 args thy P
- | literals_of_term1 args thy (@{const "op |"} $ P $ Q) =
+ | literals_of_term1 args thy (@{const HOL.disj} $ P $ Q) =
literals_of_term1 (literals_of_term1 args thy P) thy Q
| literals_of_term1 (lits, ts) thy P =
let val ((pred, ts'), pol) = predicate_of thy (P, true) in
@@ -514,8 +510,8 @@
(*Type constructors used to instantiate overloaded constants are the only ones needed.*)
fun add_type_consts_in_term thy =
let
- val const_typargs = Sign.const_typargs thy
- fun aux (Const x) = fold (fold_type_consts set_insert) (const_typargs x)
+ fun aux (Const x) =
+ fold (fold_type_consts set_insert) (Sign.const_typargs thy x)
| aux (Abs (_, _, u)) = aux u
| aux (Const (@{const_name skolem_id}, _) $ _) = I
| aux (t $ u) = aux t #> aux u
--- a/src/HOL/Tools/Sledgehammer/metis_tactics.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/metis_tactics.ML Wed Sep 01 17:19:47 2010 +0200
@@ -224,19 +224,19 @@
fun make_tvar s = TVar (("'" ^ s, 0), HOLogic.typeS)
-fun smart_invert_const "fequal" = @{const_name "op ="}
+fun smart_invert_const "fequal" = @{const_name HOL.eq}
| smart_invert_const s = invert_const s
fun hol_type_from_metis_term _ (Metis.Term.Var v) =
- (case strip_prefix_and_undo_ascii tvar_prefix v of
+ (case strip_prefix_and_unascii tvar_prefix v of
SOME w => make_tvar w
| NONE => make_tvar v)
| hol_type_from_metis_term ctxt (Metis.Term.Fn(x, tys)) =
- (case strip_prefix_and_undo_ascii type_const_prefix x of
+ (case strip_prefix_and_unascii type_const_prefix x of
SOME tc => Term.Type (smart_invert_const tc,
map (hol_type_from_metis_term ctxt) tys)
| NONE =>
- case strip_prefix_and_undo_ascii tfree_prefix x of
+ case strip_prefix_and_unascii tfree_prefix x of
SOME tf => mk_tfree ctxt tf
| NONE => raise Fail ("hol_type_from_metis_term: " ^ x));
@@ -246,10 +246,10 @@
val _ = trace_msg (fn () => "hol_term_from_metis_PT: " ^
Metis.Term.toString fol_tm)
fun tm_to_tt (Metis.Term.Var v) =
- (case strip_prefix_and_undo_ascii tvar_prefix v of
+ (case strip_prefix_and_unascii tvar_prefix v of
SOME w => Type (make_tvar w)
| NONE =>
- case strip_prefix_and_undo_ascii schematic_var_prefix v of
+ case strip_prefix_and_unascii schematic_var_prefix v of
SOME w => Term (mk_var (w, HOLogic.typeT))
| NONE => Term (mk_var (v, HOLogic.typeT)) )
(*Var from Metis with a name like _nnn; possibly a type variable*)
@@ -264,9 +264,9 @@
end
| tm_to_tt (Metis.Term.Fn (fname, args)) = applic_to_tt (fname,args)
and applic_to_tt ("=",ts) =
- Term (list_comb(Const (@{const_name "op ="}, HOLogic.typeT), terms_of (map tm_to_tt ts)))
+ Term (list_comb(Const (@{const_name HOL.eq}, HOLogic.typeT), terms_of (map tm_to_tt ts)))
| applic_to_tt (a,ts) =
- case strip_prefix_and_undo_ascii const_prefix a of
+ case strip_prefix_and_unascii const_prefix a of
SOME b =>
let val c = smart_invert_const b
val ntypes = num_type_args thy c
@@ -284,14 +284,14 @@
cat_lines (map (Syntax.string_of_term ctxt) (terms_of tts)))
end
| NONE => (*Not a constant. Is it a type constructor?*)
- case strip_prefix_and_undo_ascii type_const_prefix a of
+ case strip_prefix_and_unascii type_const_prefix a of
SOME b =>
Type (Term.Type (smart_invert_const b, types_of (map tm_to_tt ts)))
| NONE => (*Maybe a TFree. Should then check that ts=[].*)
- case strip_prefix_and_undo_ascii tfree_prefix a of
+ case strip_prefix_and_unascii tfree_prefix a of
SOME b => Type (mk_tfree ctxt b)
| NONE => (*a fixed variable? They are Skolem functions.*)
- case strip_prefix_and_undo_ascii fixed_var_prefix a of
+ case strip_prefix_and_unascii fixed_var_prefix a of
SOME b =>
let val opr = Term.Free(b, HOLogic.typeT)
in apply_list opr (length ts) (map tm_to_tt ts) end
@@ -307,16 +307,16 @@
let val _ = trace_msg (fn () => "hol_term_from_metis_FT: " ^
Metis.Term.toString fol_tm)
fun cvt (Metis.Term.Fn ("ti", [Metis.Term.Var v, _])) =
- (case strip_prefix_and_undo_ascii schematic_var_prefix v of
+ (case strip_prefix_and_unascii schematic_var_prefix v of
SOME w => mk_var(w, dummyT)
| NONE => mk_var(v, dummyT))
| cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn ("=",[]), _])) =
- Const (@{const_name "op ="}, HOLogic.typeT)
+ Const (@{const_name HOL.eq}, HOLogic.typeT)
| cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn (x,[]), ty])) =
- (case strip_prefix_and_undo_ascii const_prefix x of
+ (case strip_prefix_and_unascii const_prefix x of
SOME c => Const (smart_invert_const c, dummyT)
| NONE => (*Not a constant. Is it a fixed variable??*)
- case strip_prefix_and_undo_ascii fixed_var_prefix x of
+ case strip_prefix_and_unascii fixed_var_prefix x of
SOME v => Free (v, hol_type_from_metis_term ctxt ty)
| NONE => raise Fail ("hol_term_from_metis_FT bad constant: " ^ x))
| cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn (".",[tm1,tm2]), _])) =
@@ -325,12 +325,12 @@
cvt tm1 $ cvt tm2
| cvt (Metis.Term.Fn ("{}", [arg])) = cvt arg (*hBOOL*)
| cvt (Metis.Term.Fn ("=", [tm1,tm2])) =
- list_comb(Const (@{const_name "op ="}, HOLogic.typeT), map cvt [tm1,tm2])
+ list_comb(Const (@{const_name HOL.eq}, HOLogic.typeT), map cvt [tm1,tm2])
| cvt (t as Metis.Term.Fn (x, [])) =
- (case strip_prefix_and_undo_ascii const_prefix x of
+ (case strip_prefix_and_unascii const_prefix x of
SOME c => Const (smart_invert_const c, dummyT)
| NONE => (*Not a constant. Is it a fixed variable??*)
- case strip_prefix_and_undo_ascii fixed_var_prefix x of
+ case strip_prefix_and_unascii fixed_var_prefix x of
SOME v => Free (v, dummyT)
| NONE => (trace_msg (fn () => "hol_term_from_metis_FT bad const: " ^ x);
hol_term_from_metis_PT ctxt t))
@@ -410,11 +410,11 @@
" in " ^ Display.string_of_thm ctxt i_th);
NONE)
fun remove_typeinst (a, t) =
- case strip_prefix_and_undo_ascii schematic_var_prefix a of
+ case strip_prefix_and_unascii schematic_var_prefix a of
SOME b => SOME (b, t)
- | NONE => case strip_prefix_and_undo_ascii tvar_prefix a of
+ | NONE => case strip_prefix_and_unascii tvar_prefix a of
SOME _ => NONE (*type instantiations are forbidden!*)
- | NONE => SOME (a,t) (*internal Metis var?*)
+ | NONE => SOME (a,t) (*internal Metis var?*)
val _ = trace_msg (fn () => " isa th: " ^ Display.string_of_thm ctxt i_th)
val substs = map_filter remove_typeinst (Metis.Subst.toList fsubst)
val (vars,rawtms) = ListPair.unzip (map_filter subst_translation substs)
@@ -480,7 +480,7 @@
val c_t = cterm_incr_types thy refl_idx i_t
in cterm_instantiate [(refl_x, c_t)] REFL_THM end;
-fun get_ty_arg_size _ (Const (@{const_name "op ="}, _)) = 0 (*equality has no type arguments*)
+fun get_ty_arg_size _ (Const (@{const_name HOL.eq}, _)) = 0 (*equality has no type arguments*)
| get_ty_arg_size thy (Const (c, _)) = (num_type_args thy c handle TYPE _ => 0)
| get_ty_arg_size _ _ = 0;
@@ -529,13 +529,13 @@
" isa-term: " ^ Syntax.string_of_term ctxt tm ^
" fol-term: " ^ Metis.Term.toString t)
fun path_finder FO tm ps _ = path_finder_FO tm ps
- | path_finder HO (tm as Const(@{const_name "op ="},_) $ _ $ _) (p::ps) _ =
+ | path_finder HO (tm as Const(@{const_name HOL.eq},_) $ _ $ _) (p::ps) _ =
(*equality: not curried, as other predicates are*)
if p=0 then path_finder_HO tm (0::1::ps) (*select first operand*)
else path_finder_HO tm (p::ps) (*1 selects second operand*)
| path_finder HO tm (_ :: ps) (Metis.Term.Fn ("{}", [_])) =
path_finder_HO tm ps (*if not equality, ignore head to skip hBOOL*)
- | path_finder FT (tm as Const(@{const_name "op ="}, _) $ _ $ _) (p::ps)
+ | path_finder FT (tm as Const(@{const_name HOL.eq}, _) $ _ $ _) (p::ps)
(Metis.Term.Fn ("=", [t1,t2])) =
(*equality: not curried, as other predicates are*)
if p=0 then path_finder_FT tm (0::1::ps)
--- a/src/HOL/Tools/Sledgehammer/sledgehammer.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer.ML Wed Sep 01 17:19:47 2010 +0200
@@ -9,6 +9,7 @@
signature SLEDGEHAMMER =
sig
type failure = ATP_Systems.failure
+ type locality = Sledgehammer_Fact_Filter.locality
type relevance_override = Sledgehammer_Fact_Filter.relevance_override
type minimize_command = Sledgehammer_Proof_Reconstruct.minimize_command
type params =
@@ -18,11 +19,9 @@
atps: string list,
full_types: bool,
explicit_apply: bool,
- relevance_threshold: real,
- relevance_convergence: real,
- max_relevant_per_iter: int option,
+ relevance_thresholds: real * real,
+ max_relevant: int option,
theory_relevant: bool option,
- defs_relevant: bool,
isar_proof: bool,
isar_shrink_factor: int,
timeout: Time.time}
@@ -30,16 +29,16 @@
{subgoal: int,
goal: Proof.context * (thm list * thm),
relevance_override: relevance_override,
- axioms: ((string * bool) * thm) list option}
+ axioms: ((string * locality) * thm) list option}
type prover_result =
{outcome: failure option,
message: string,
pool: string Symtab.table,
- used_thm_names: (string * bool) list,
+ used_thm_names: (string * locality) list,
atp_run_time_in_msecs: int,
output: string,
proof: string,
- axiom_names: (string * bool) vector,
+ axiom_names: (string * locality) list vector,
conjecture_shape: int list list}
type prover = params -> minimize_command -> problem -> prover_result
@@ -87,11 +86,9 @@
atps: string list,
full_types: bool,
explicit_apply: bool,
- relevance_threshold: real,
- relevance_convergence: real,
- max_relevant_per_iter: int option,
+ relevance_thresholds: real * real,
+ max_relevant: int option,
theory_relevant: bool option,
- defs_relevant: bool,
isar_proof: bool,
isar_shrink_factor: int,
timeout: Time.time}
@@ -100,17 +97,17 @@
{subgoal: int,
goal: Proof.context * (thm list * thm),
relevance_override: relevance_override,
- axioms: ((string * bool) * thm) list option}
+ axioms: ((string * locality) * thm) list option}
type prover_result =
{outcome: failure option,
message: string,
pool: string Symtab.table,
- used_thm_names: (string * bool) list,
+ used_thm_names: (string * locality) list,
atp_run_time_in_msecs: int,
output: string,
proof: string,
- axiom_names: (string * bool) vector,
+ axiom_names: (string * locality) list vector,
conjecture_shape: int list list}
type prover = params -> minimize_command -> problem -> prover_result
@@ -174,11 +171,11 @@
Scan.this_string set_ClauseFormulaRelationN |-- $$ "("
|-- Scan.repeat parse_clause_formula_pair
val extract_clause_formula_relation =
- Substring.full
- #> Substring.position set_ClauseFormulaRelationN
- #> snd #> Substring.string #> strip_spaces #> explode
- #> parse_clause_formula_relation #> fst
+ Substring.full #> Substring.position set_ClauseFormulaRelationN
+ #> snd #> Substring.string #> strip_spaces_except_between_ident_chars
+ #> explode #> parse_clause_formula_relation #> fst
+(* TODO: move to "Sledgehammer_Proof_Reconstruct.ML" *)
fun repair_conjecture_shape_and_theorem_names output conjecture_shape
axiom_names =
if String.isSubstring set_ClauseFormulaRelationN output then
@@ -190,20 +187,20 @@
val seq = extract_clause_sequence output
val name_map = extract_clause_formula_relation output
fun renumber_conjecture j =
- conjecture_prefix ^ Int.toString (j - j0)
+ 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 name_for_number j =
- let
- val axioms =
- j |> AList.lookup (op =) name_map
- |> these |> map_filter (try (unprefix axiom_prefix))
- |> map undo_ascii_of
- val chained = forall (is_true_for axiom_names) axioms
- in (axioms |> space_implode " ", chained) end
+ 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 name_for_number |> Vector.fromList)
+ seq |> map names_for_number |> Vector.fromList)
end
else
(conjecture_shape, axiom_names)
@@ -213,20 +210,18 @@
fun prover_fun atp_name
{exec, required_execs, arguments, has_incomplete_mode, proof_delims,
- known_failures, default_max_relevant_per_iter, default_theory_relevant,
+ known_failures, default_max_relevant, default_theory_relevant,
explicit_forall, use_conjecture_for_hypotheses}
({debug, verbose, overlord, full_types, explicit_apply,
- relevance_threshold, relevance_convergence,
- max_relevant_per_iter, theory_relevant,
- defs_relevant, isar_proof, isar_shrink_factor, timeout, ...} : params)
+ relevance_thresholds, max_relevant, theory_relevant, isar_proof,
+ isar_shrink_factor, timeout, ...} : params)
minimize_command
- ({subgoal, goal, relevance_override, axioms} : problem) =
+ ({subgoal, goal = (ctxt, (chained_ths, th)), relevance_override,
+ axioms} : problem) =
let
- val (ctxt, (_, th)) = goal;
- val thy = ProofContext.theory_of ctxt
- val (params, hyp_ts, concl_t) = strip_subgoal th subgoal
+ val (_, hyp_ts, concl_t) = strip_subgoal th subgoal
- fun print s = (priority s; if debug then tracing s else ())
+ val print = priority
fun print_v f = () |> verbose ? print o f
fun print_d f = () |> debug ? print o f
@@ -234,15 +229,13 @@
case axioms of
SOME axioms => axioms
| NONE =>
- (relevant_facts full_types relevance_threshold relevance_convergence
- defs_relevant
- (the_default default_max_relevant_per_iter
- max_relevant_per_iter)
+ (relevant_facts ctxt full_types relevance_thresholds
+ (the_default default_max_relevant max_relevant)
(the_default default_theory_relevant theory_relevant)
- relevance_override goal hyp_ts concl_t
+ relevance_override chained_ths hyp_ts concl_t
|> tap ((fn n => print_v (fn () =>
- "Selected " ^ string_of_int n ^ " fact" ^ plural_s n ^
- " for " ^ quote atp_name ^ ".")) o length))
+ "Selected " ^ string_of_int n ^ " fact" ^ plural_s n ^
+ " for " ^ quote atp_name ^ ".")) o length))
(* path to unique problem file *)
val the_dest_dir = if overlord then getenv "ISABELLE_HOME_USER"
@@ -258,9 +251,10 @@
if the_dest_dir = "" then File.tmp_path probfile
else if File.exists (Path.explode the_dest_dir)
then Path.append (Path.explode the_dest_dir) probfile
- else error ("No such directory: " ^ the_dest_dir ^ ".")
+ else error ("No such directory: " ^ quote the_dest_dir ^ ".")
end;
+ val measure_run_time = verbose orelse Config.get ctxt measure_runtime
val command = Path.explode (getenv (fst exec) ^ "/" ^ snd exec)
(* write out problem file and call prover *)
fun command_line complete timeout probfile =
@@ -268,10 +262,8 @@
val core = File.shell_path command ^ " " ^ arguments complete timeout ^
" " ^ File.shell_path probfile
in
- (if Config.get ctxt measure_runtime then
- "TIMEFORMAT='%3U'; { time " ^ core ^ " ; }"
- else
- "exec " ^ core) ^ " 2>&1"
+ (if measure_run_time then "TIMEFORMAT='%3U'; { time " ^ core ^ " ; }"
+ else "exec " ^ core) ^ " 2>&1"
end
fun split_time s =
let
@@ -300,14 +292,11 @@
prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n")
else
I)
- |>> (if Config.get ctxt measure_runtime then split_time
- else rpair 0)
+ |>> (if measure_run_time then split_time else rpair 0)
val (proof, outcome) =
extract_proof_and_outcome complete res_code proof_delims
known_failures output
in (output, msecs, proof, outcome) end
- val _ = print_d (fn () => "Preparing problem for " ^
- quote atp_name ^ "...")
val readable_names = debug andalso overlord
val (problem, pool, conjecture_offset, axiom_names) =
prepare_problem ctxt readable_names explicit_forall full_types
@@ -358,6 +347,11 @@
proof_text isar_proof
(pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
(full_types, minimize_command, proof, axiom_names, th, subgoal)
+ |>> (fn message =>
+ message ^ (if verbose then
+ "\nATP CPU time: " ^ string_of_int msecs ^ " ms."
+ else
+ ""))
| SOME failure => (string_for_failure failure, [])
in
{outcome = outcome, message = message, pool = pool,
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML Thu Aug 26 13:15:37 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML Wed Sep 01 17:19:47 2010 +0200
@@ -5,19 +5,34 @@
signature SLEDGEHAMMER_FACT_FILTER =
sig
+ datatype locality = General | Intro | Elim | Simp | Local | Chained
+
type relevance_override =
{add: Facts.ref list,
del: Facts.ref list,
only: bool}
val trace : bool Unsynchronized.ref
- val name_thms_pair_from_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 intro_bonus : real Unsynchronized.ref
+ val elim_bonus : real Unsynchronized.ref
+ val simp_bonus : real Unsynchronized.ref
+ val local_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
- -> (unit -> string * bool) * thm list
+ -> ((string * locality) * thm) list
val relevant_facts :
- bool -> real -> real -> bool -> int -> bool -> relevance_override
- -> Proof.context * (thm list * 'a) -> term list -> term
- -> ((string * bool) * thm) list
+ Proof.context -> bool -> real * real -> int -> bool -> relevance_override
+ -> thm list -> term list -> term -> ((string * locality) * thm) list
end;
structure Sledgehammer_Fact_Filter : SLEDGEHAMMER_FACT_FILTER =
@@ -28,8 +43,13 @@
val trace = Unsynchronized.ref false
fun trace_msg msg = if !trace then tracing (msg ()) else ()
+(* experimental feature *)
+val term_patterns = false
+
val respect_no_atp = true
+datatype locality = General | Intro | Elim | Simp | Local | Chained
+
type relevance_override =
{add: Facts.ref list,
del: Facts.ref list,
@@ -37,13 +57,22 @@
val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator
-fun name_thms_pair_from_ref ctxt reserved chained_ths xref =
- let val ths = ProofContext.get_fact ctxt xref in
- (fn () => let
- val name = Facts.string_of_ref xref
- val name = name |> Symtab.defined reserved name ? quote
- val chained = forall (member Thm.eq_thm chained_ths) ths
- in (name, chained) end, ths)
+fun repair_name reserved multi j name =
+ (name |> Symtab.defined reserved name ? quote) ^
+ (if multi then "(" ^ Int.toString j ^ ")" else "")
+
+fun name_thm_pairs_from_ref ctxt reserved chained_ths xref =
+ let
+ val ths = ProofContext.get_fact ctxt xref
+ val name = Facts.string_of_ref xref
+ val multi = length ths > 1
+ in
+ (ths, (1, []))
+ |-> fold (fn th => fn (j, rest) =>
+ (j + 1, ((repair_name reserved multi j name,
+ if member Thm.eq_thm chained_ths th then Chained
+ else General), th) :: rest))
+ |> snd
end
(***************************************************************)
@@ -52,104 +81,146 @@
(*** constants with types ***)
-(*An abstraction of Isabelle types*)
-datatype const_typ = CTVar | CType of string * const_typ list
+fun order_of_type (Type (@{type_name fun}, [T1, @{typ bool}])) =
+ order_of_type T1 (* cheat: pretend sets are first-order *)
+ | order_of_type (Type (@{type_name fun}, [T1, T2])) =
+ Int.max (order_of_type T1 + 1, order_of_type T2)
+ | order_of_type (Type (_, Ts)) = fold (Integer.max o order_of_type) Ts 0
+ | order_of_type _ = 0
+
+(* An abstraction of Isabelle types and first-order terms *)
+datatype pattern = PVar | PApp of string * pattern list
+datatype ptype = PType of int * pattern list
+
+fun string_for_pattern PVar = "_"
+ | string_for_pattern (PApp (s, ps)) =
+ if null ps then s else s ^ string_for_patterns ps
+and string_for_patterns ps = "(" ^ commas (map string_for_pattern ps) ^ ")"
+fun string_for_ptype (PType (_, ps)) = string_for_patterns ps
(*Is the second type an instance of the first one?*)
-fun match_type (CType(con1,args1)) (CType(con2,args2)) =
- con1=con2 andalso match_types args1 args2
- | match_type CTVar _ = true
- | match_type _ CTVar = false
-and match_types [] [] = true
- | match_types (a1::as1) (a2::as2) = match_type a1 a2 andalso match_types as1 as2;
+fun match_pattern (PVar, _) = true
+ | match_pattern (PApp _, PVar) = false
+ | match_pattern (PApp (s, ps), PApp (t, qs)) =
+ s = t andalso match_patterns (ps, qs)
+and match_patterns (_, []) = true
+ | match_patterns ([], _) = false
+ | match_patterns (p :: ps, q :: qs) =
+ match_pattern (p, q) andalso match_patterns (ps, qs)
+fun match_ptype (PType (_, ps), PType (_, qs)) = match_patterns (ps, qs)
-(*Is there a unifiable constant?*)
-fun const_mem const_tab (c, c_typ) =
- exists (match_types c_typ) (these (Symtab.lookup const_tab c))
+(* Is there a unifiable constant? *)
+fun pconst_mem f consts (s, ps) =
+ exists (curry (match_ptype o f) ps)
+ (map snd (filter (curry (op =) s o fst) consts))
+fun pconst_hyper_mem f const_tab (s, ps) =
+ exists (curry (match_ptype o f) ps) (these (Symtab.lookup const_tab s))
+
+fun pattern_for_type (Type (s, Ts)) = PApp (s, map pattern_for_type Ts)
+ | pattern_for_type (TFree (s, _)) = PApp (s, [])
+ | pattern_for_type (TVar _) = PVar
-(*Maps a "real" type to a const_typ*)
-fun const_typ_of (Type (c,typs)) = CType (c, map const_typ_of typs)
- | const_typ_of (TFree _) = CTVar
- | const_typ_of (TVar _) = CTVar
+fun pterm thy t =
+ case strip_comb t of
+ (Const x, ts) => PApp (pconst thy true x ts)
+ | (Free x, ts) => PApp (pconst thy false x ts)
+ | (Var x, []) => PVar
+ | _ => PApp ("?", []) (* equivalence class of higher-order constructs *)
+(* Pairs a constant with the list of its type instantiations. *)
+and ptype thy const x ts =
+ (if const then map pattern_for_type (these (try (Sign.const_typargs thy) x))
+ else []) @
+ (if term_patterns then map (pterm thy) ts else [])
+and pconst thy const (s, T) ts = (s, ptype thy const (s, T) ts)
+and rich_ptype thy const (s, T) ts =
+ PType (order_of_type T, ptype thy const (s, T) ts)
+and rich_pconst thy const (s, T) ts = (s, rich_ptype thy const (s, T) ts)
-(*Pairs a constant with the list of its type instantiations (using const_typ)*)
-fun const_with_typ thy (c,typ) =
- let val tvars = Sign.const_typargs thy (c,typ) in
- (c, map const_typ_of tvars) end
- handle TYPE _ => (c, []) (*Variable (locale constant): monomorphic*)
+fun string_for_hyper_pconst (s, ps) =
+ s ^ "{" ^ commas (map string_for_ptype ps) ^ "}"
+
+val abs_name = "Sledgehammer.abs"
+val skolem_prefix = "Sledgehammer.sko"
-(*Add a const/type pair to the table, but a [] entry means a standard connective,
- which we ignore.*)
-fun add_const_to_table (c, ctyps) =
- Symtab.map_default (c, [ctyps])
- (fn [] => [] | ctypss => insert (op =) ctyps ctypss)
+(* 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
+ (not also_skolem andalso String.isPrefix skolem_prefix c) then
+ I
+ else
+ Symtab.map_default (c, [p]) (insert (op =) p)
fun is_formula_type T = (T = HOLogic.boolT orelse T = propT)
-val fresh_prefix = "Sledgehammer.FRESH."
-val flip = Option.map not
-(* These are typically simplified away by "Meson.presimplify". *)
-val boring_consts =
- [@{const_name False}, @{const_name True}, @{const_name If}, @{const_name Let}]
-
-fun get_consts thy pos ts =
+fun pconsts_in_terms thy 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_term t =
- case t of
- Const x => add_const_to_table (const_with_typ thy x)
- | Free (s, _) => add_const_to_table (s, [])
- | t1 $ t2 => fold do_term [t1, t2]
- | Abs (_, _, t') => do_term t'
- | _ => I
- fun do_quantifier will_surely_be_skolemized body_t =
+ fun do_const const (s, T) ts =
+ add_pconst_to_table also_skolems (rich_pconst thy const (s, T) ts)
+ #> fold do_term ts
+ and do_term t =
+ case strip_comb t of
+ (Const x, ts) => do_const true x ts
+ | (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, [])))
+ #> 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 will_surely_be_skolemized then
- add_const_to_table (gensym fresh_prefix, [])
+ #> (if also_skolems andalso will_surely_be_skolemized then
+ add_pconst_to_table true
+ (gensym skolem_prefix, PType (order_of_type abs_T, []))
else
I)
and do_term_or_formula T =
if is_formula_type T then do_formula NONE else do_term
and do_formula pos t =
case t of
- Const (@{const_name all}, _) $ Abs (_, _, body_t) =>
- do_quantifier (pos = SOME false) body_t
+ Const (@{const_name all}, _) $ Abs (_, T, t') =>
+ do_quantifier (pos = SOME false) T t'
| @{const "==>"} $ t1 $ t2 =>
do_formula (flip pos) t1 #> do_formula pos t2
| Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2 =>
fold (do_term_or_formula T) [t1, t2]
| @{const Trueprop} $ t1 => do_formula pos t1
| @{const Not} $ t1 => do_formula (flip pos) t1
- | Const (@{const_name All}, _) $ Abs (_, _, body_t) =>
- do_quantifier (pos = SOME false) body_t
- | Const (@{const_name Ex}, _) $ Abs (_, _, body_t) =>
- do_quantifier (pos = SOME true) body_t
- | @{const "op &"} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
- | @{const "op |"} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
- | @{const "op -->"} $ t1 $ t2 =>
+ | Const (@{const_name All}, _) $ Abs (_, T, t') =>
+ do_quantifier (pos = SOME false) T t'
+ | Const (@{const_name Ex}, _) $ Abs (_, T, t') =>
+ do_quantifier (pos = SOME true) T t'
+ | @{const HOL.conj} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
+ | @{const HOL.disj} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
+ | @{const HOL.implies} $ t1 $ t2 =>
do_formula (flip pos) t1 #> do_formula pos t2
- | Const (@{const_name "op ="}, Type (_, [T, _])) $ t1 $ t2 =>
+ | Const (@{const_name HOL.eq}, Type (_, [T, _])) $ t1 $ t2 =>
fold (do_term_or_formula T) [t1, t2]
| Const (@{const_name If}, Type (_, [_, Type (_, [T, _])]))
$ t1 $ t2 $ t3 =>
do_formula NONE t1 #> fold (do_term_or_formula T) [t2, t3]
- | Const (@{const_name Ex1}, _) $ Abs (_, _, body_t) =>
- do_quantifier (is_some pos) body_t
- | Const (@{const_name Ball}, _) $ t1 $ Abs (_, _, body_t) =>
- do_quantifier (pos = SOME false)
- (HOLogic.mk_imp (incr_boundvars 1 t1 $ Bound 0, body_t))
- | Const (@{const_name Bex}, _) $ t1 $ Abs (_, _, body_t) =>
- do_quantifier (pos = SOME true)
- (HOLogic.mk_conj (incr_boundvars 1 t1 $ Bound 0, body_t))
+ | Const (@{const_name Ex1}, _) $ Abs (_, T, t') =>
+ do_quantifier (is_some pos) T t'
+ | Const (@{const_name Ball}, _) $ t1 $ Abs (_, T, t') =>
+ do_quantifier (pos = SOME false) T
+ (HOLogic.mk_imp (incr_boundvars 1 t1 $ Bound 0, t'))
+ | Const (@{const_name Bex}, _) $ t1 $ Abs (_, T, t') =>
+ do_quantifier (pos = SOME true) T
+ (HOLogic.mk_conj (incr_boundvars 1 t1 $ Bound 0, t'))
| (t0 as Const (_, @{typ bool})) $ t1 =>
do_term t0 #> do_formula pos t1 (* theory constant *)
| _ => do_term t
- in
- Symtab.empty |> fold (Symtab.update o rpair []) boring_consts
- |> fold (do_formula pos) ts
- end
+ in Symtab.empty |> fold (do_formula pos) ts end
(*Inserts a dummy "constant" referring to the theory name, so that relevance
takes the given theory into account.*)
@@ -164,235 +235,292 @@
(**** Constant / Type Frequencies ****)
-(*A two-dimensional symbol table counts frequencies of constants. It's keyed first by
- constant name and second by its list of type instantiations. For the latter, we need
- a linear ordering on type const_typ list.*)
-
-local
+(* A two-dimensional symbol table counts frequencies of constants. It's keyed
+ first by constant name and second by its list of type instantiations. For the
+ latter, we need a linear ordering on "pattern list". *)
-fun cons_nr CTVar = 0
- | cons_nr (CType _) = 1;
-
-in
-
-fun const_typ_ord TU =
- case TU of
- (CType (a, Ts), CType (b, Us)) =>
- (case fast_string_ord(a,b) of EQUAL => dict_ord const_typ_ord (Ts,Us) | ord => ord)
- | (T, U) => int_ord (cons_nr T, cons_nr U);
+fun pattern_ord p =
+ case p of
+ (PVar, PVar) => EQUAL
+ | (PVar, PApp _) => LESS
+ | (PApp _, PVar) => GREATER
+ | (PApp q1, PApp q2) =>
+ prod_ord fast_string_ord (dict_ord pattern_ord) (q1, q2)
+fun ptype_ord (PType p, PType q) =
+ prod_ord (dict_ord pattern_ord) int_ord (swap p, swap q)
-end;
+structure PType_Tab = Table(type key = ptype val ord = ptype_ord)
-structure CTtab = Table(type key = const_typ list val ord = dict_ord const_typ_ord);
-
-fun count_axiom_consts theory_relevant thy (_, th) =
+fun count_axiom_consts theory_relevant thy =
let
- fun do_const (a, T) =
- let val (c, cts) = const_with_typ thy (a, T) in
- (* Two-dimensional table update. Constant maps to types maps to
- count. *)
- CTtab.map_default (cts, 0) (Integer.add 1)
- |> Symtab.map_default (c, CTtab.empty)
- end
- fun do_term (Const x) = do_const x
- | do_term (Free x) = do_const x
- | do_term (t $ u) = do_term t #> do_term u
- | do_term (Abs (_, _, t)) = do_term t
- | do_term _ = I
- in th |> theory_const_prop_of theory_relevant |> do_term end
+ fun do_const const (s, T) ts =
+ (* Two-dimensional table update. Constant maps to types maps to count. *)
+ PType_Tab.map_default (rich_ptype thy const (s, T) ts, 0) (Integer.add 1)
+ |> Symtab.map_default (s, PType_Tab.empty)
+ #> fold do_term ts
+ and do_term t =
+ case strip_comb t of
+ (Const x, ts) => do_const true x ts
+ | (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 theory_relevant o snd end
(**** Actual Filtering Code ****)
+fun pow_int x 0 = 1.0
+ | pow_int x 1 = x
+ | pow_int x n = if n > 0 then x * pow_int x (n - 1) else pow_int x (n + 1) / x
+
(*The frequency of a constant is the sum of those of all instances of its type.*)
-fun const_frequency const_tab (c, cts) =
- CTtab.fold (fn (cts', m) => match_types cts cts' ? Integer.add m)
- (the (Symtab.lookup const_tab c)) 0
- handle Option.Option => 0
+fun pconst_freq match const_tab (c, ps) =
+ PType_Tab.fold (fn (qs, m) => match (ps, qs) ? Integer.add m)
+ (the (Symtab.lookup const_tab c)) 0
(* A surprising number of theorems contain only a few significant constants.
These include all induction rules, and other general theorems. *)
(* "log" seems best in practice. A constant function of one ignores the constant
- frequencies. *)
-fun rel_log (x : real) = 1.0 + 2.0 / Math.ln (x + 1.0)
-fun irrel_log (x : real) = Math.ln (x + 19.0) / 6.4
+ frequencies. Rare constants give more points if they are relevant than less
+ rare ones. *)
+fun rel_weight_for order freq = 1.0 + 2.0 / Math.ln (Real.fromInt freq + 1.0)
+
+(* FUDGE *)
+val worse_irrel_freq = Unsynchronized.ref 100.0
+val higher_order_irrel_weight = Unsynchronized.ref 1.05
+
+(* Irrelevant constants are treated differently. We associate lower penalties to
+ 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
+ (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)
+ end
+
+(* FUDGE *)
+val abs_rel_weight = Unsynchronized.ref 0.5
+val abs_irrel_weight = Unsynchronized.ref 2.0
+val skolem_irrel_weight = Unsynchronized.ref 0.75
(* Computes a constant's weight, as determined by its frequency. *)
-val rel_const_weight = rel_log o real oo const_frequency
-val irrel_const_weight = irrel_log o real oo const_frequency
-(* fun irrel_const_weight _ _ = 1.0 FIXME: OLD CODE *)
+fun generic_pconst_weight abs_weight skolem_weight weight_for f const_tab
+ (c as (s, PType (m, _))) =
+ if s = abs_name then abs_weight
+ else if String.isPrefix skolem_prefix s then skolem_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 rel_weight_for I const_tab
+fun irrel_pconst_weight const_tab =
+ generic_pconst_weight (!abs_irrel_weight) (!skolem_irrel_weight)
+ irrel_weight_for swap const_tab
+
+(* FUDGE *)
+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 chained_bonus = Unsynchronized.ref 1.5
+
+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 Chained = !chained_bonus
-fun axiom_weight const_tab relevant_consts axiom_consts =
- let
- val (rel, irrel) = List.partition (const_mem relevant_consts) axiom_consts
- val rel_weight = fold (curry Real.+ o rel_const_weight const_tab) rel 0.0
- val irrel_weight = fold (curry Real.+ o irrel_const_weight const_tab) irrel 0.0
- val res = rel_weight / (rel_weight + irrel_weight)
- in if Real.isFinite res then res else 0.0 end
+fun axiom_weight 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
+ | (rel, irrel) =>
+ 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
+ val irrel_weight =
+ ~ (locality_bonus loc)
+ |> fold (curry (op +) o irrel_pconst_weight const_tab) irrel
+ val res = rel_weight / (rel_weight + irrel_weight)
+ in if Real.isFinite res then res else 0.0 end
-(* OLD CODE:
-(*Relevant constants are weighted according to frequency,
- but irrelevant constants are simply counted. Otherwise, Skolem functions,
- which are rare, would harm a formula's chances of being picked.*)
-fun axiom_weight const_tab relevant_consts axiom_consts =
- let
- val rel = filter (const_mem relevant_consts) axiom_consts
- val rel_weight = fold (curry Real.+ o rel_const_weight const_tab) rel 0.0
- val res = rel_weight / (rel_weight + real (length axiom_consts - length rel))
- in if Real.isFinite res then res else 0.0 end
+(* FIXME: experiment
+fun debug_axiom_weight 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
+ | (rel, irrel) =>
+ let
+ val irrel = irrel |> filter_out (pconst_mem swap rel)
+ 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
+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 res = rels_weight / (rels_weight + irrels_weight)
+ in if Real.isFinite res then res else 0.0 end
*)
-(*Multiplies out to a list of pairs: 'a * 'b list -> ('a * 'b) list -> ('a * 'b) list*)
-fun add_expand_pairs (x, ys) xys = List.foldl (fn (y,acc) => (x,y)::acc) xys ys
-
-fun consts_of_term thy t =
- Symtab.fold add_expand_pairs (get_consts thy (SOME true) [t]) []
-
+fun pconsts_in_axiom thy t =
+ Symtab.fold (fn (s, pss) => fold (cons o pair s) pss)
+ (pconsts_in_terms thy true (SOME true) [t]) []
fun pair_consts_axiom theory_relevant thy axiom =
- (axiom, axiom |> snd |> theory_const_prop_of theory_relevant
- |> consts_of_term thy)
-
-exception CONST_OR_FREE of unit
-
-fun dest_Const_or_Free (Const x) = x
- | dest_Const_or_Free (Free x) = x
- | dest_Const_or_Free _ = raise CONST_OR_FREE ()
-
-(*Look for definitions of the form f ?x1 ... ?xn = t, but not reversed.*)
-fun defines thy thm gctypes =
- let val tm = prop_of thm
- fun defs lhs rhs =
- let val (rator,args) = strip_comb lhs
- val ct = const_with_typ thy (dest_Const_or_Free rator)
- in
- forall is_Var args andalso const_mem gctypes ct andalso
- subset (op =) (Term.add_vars rhs [], Term.add_vars lhs [])
- end
- handle CONST_OR_FREE () => false
- in
- case tm of
- @{const Trueprop} $ (Const (@{const_name "op ="}, _) $ lhs $ rhs) =>
- defs lhs rhs
- | _ => false
- end;
+ case axiom |> snd |> theory_const_prop_of theory_relevant
+ |> pconsts_in_axiom thy of
+ [] => NONE
+ | consts => SOME ((axiom, consts), NONE)
type annotated_thm =
- ((unit -> string * bool) * thm) * (string * const_typ list) list
+ (((unit -> string) * locality) * thm) * (string * ptype) list
-(*For a reverse sort, putting the largest values first.*)
-fun compare_pairs ((_, w1), (_, w2)) = Real.compare (w2, w1)
+(* FUDGE *)
+val max_imperfect = Unsynchronized.ref 11.5
+val max_imperfect_exp = Unsynchronized.ref 1.0
-(* Limit the number of new facts, to prevent runaway acceptance. *)
-fun take_best max_new (new_pairs : (annotated_thm * real) list) =
- let val nnew = length new_pairs in
- if nnew <= max_new then
- (map #1 new_pairs, [])
- else
- let
- val new_pairs = sort compare_pairs new_pairs
- val accepted = List.take (new_pairs, max_new)
- in
- trace_msg (fn () => ("Number of candidates, " ^ Int.toString nnew ^
- ", exceeds the limit of " ^ Int.toString max_new));
- trace_msg (fn () => ("Effective pass mark: " ^ Real.toString (#2 (List.last accepted))));
- trace_msg (fn () => "Actually passed: " ^
- space_implode ", " (map (fst o (fn f => f ()) o fst o fst o fst) accepted));
- (map #1 accepted, List.drop (new_pairs, max_new))
- end
- end;
+fun take_most_relevant max_relevant remaining_max
+ (candidates : (annotated_thm * real) list) =
+ let
+ val max_imperfect =
+ Real.ceil (Math.pow (!max_imperfect,
+ Math.pow (Real.fromInt remaining_max
+ / 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)
+ val ((accepts, more_rejects), rejects) =
+ chop max_imperfect imperfect |>> append perfect |>> chop remaining_max
+ in
+ trace_msg (fn () =>
+ "Actually passed (" ^ Int.toString (length accepts) ^ " of " ^
+ Int.toString (length candidates) ^ "): " ^
+ (accepts |> map (fn ((((name, _), _), _), weight) =>
+ name () ^ " [" ^ Real.toString weight ^ "]")
+ |> commas));
+ (accepts, more_rejects @ rejects)
+ end
-val threshold_divisor = 2.0
-val ridiculous_threshold = 0.1
+fun if_empty_replace_with_locality thy axioms loc tab =
+ if Symtab.is_empty tab then
+ pconsts_in_terms thy false (SOME false)
+ (map_filter (fn ((_, loc'), th) =>
+ if loc' = loc then SOME (prop_of th) else NONE) axioms)
+ else
+ tab
+
+(* FUDGE *)
+val threshold_divisor = Unsynchronized.ref 2.0
+val ridiculous_threshold = Unsynchronized.ref 0.1
-fun relevance_filter ctxt relevance_threshold relevance_convergence
- defs_relevant max_new theory_relevant
+fun relevance_filter ctxt threshold0 decay max_relevant theory_relevant
({add, del, ...} : relevance_override) axioms goal_ts =
- if relevance_threshold > 1.0 then
- []
- else if relevance_threshold < 0.0 then
- axioms
- else
- let
- val thy = ProofContext.theory_of ctxt
- val const_tab = fold (count_axiom_consts theory_relevant thy) axioms
- Symtab.empty
- val goal_const_tab = get_consts thy (SOME false) goal_ts
- val _ =
- trace_msg (fn () => "Initial constants: " ^
- commas (goal_const_tab
- |> Symtab.dest
- |> filter (curry (op <>) [] o snd)
- |> map fst))
- val add_thms = maps (ProofContext.get_fact ctxt) add
- val del_thms = maps (ProofContext.get_fact ctxt) del
- fun iter j threshold rel_const_tab =
- let
- fun relevant ([], rejects) [] =
- (* Nothing was added this iteration. *)
- if j = 0 andalso threshold >= ridiculous_threshold then
- (* First iteration? Try again. *)
- iter 0 (threshold / threshold_divisor) rel_const_tab
- (map (apsnd SOME) rejects)
+ let
+ val thy = ProofContext.theory_of ctxt
+ val const_tab =
+ fold (count_axiom_consts theory_relevant thy) axioms Symtab.empty
+ val goal_const_tab =
+ pconsts_in_terms thy false (SOME false) goal_ts
+ |> fold (if_empty_replace_with_locality thy axioms) [Chained, Local]
+ val add_thms = maps (ProofContext.get_fact ctxt) add
+ val del_thms = maps (ProofContext.get_fact ctxt) del
+ fun iter j remaining_max threshold rel_const_tab hopeless hopeful =
+ let
+ fun game_over rejects =
+ (* Add "add:" facts. *)
+ if null add_thms then
+ []
+ else
+ map_filter (fn ((p as (_, th), _), _) =>
+ if member Thm.eq_thm add_thms th then SOME p
+ else NONE) rejects
+ fun relevant [] rejects [] =
+ (* Nothing has been added this iteration. *)
+ if j = 0 andalso threshold >= !ridiculous_threshold then
+ (* First iteration? Try again. *)
+ 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
+ val rel_const_tab' =
+ rel_const_tab
+ |> fold (add_pconst_to_table 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) =
+ (rejects @ hopeless, ([], []))
+ |-> fold (fn (ax as (_, consts), old_weight) =>
+ if exists is_dirty consts then
+ apfst (cons (ax, NONE))
+ else
+ apsnd (cons (ax, old_weight)))
+ |>> append (more_rejects
+ |> map (fn (ax as (_, consts), old_weight) =>
+ (ax, if exists is_dirty consts then NONE
+ else SOME old_weight)))
+ val threshold =
+ 1.0 - (1.0 - threshold)
+ * Math.pow (decay, Real.fromInt (length accepts))
+ val remaining_max = remaining_max - length accepts
+ in
+ trace_msg (fn () => "New or updated constants: " ^
+ commas (rel_const_tab' |> Symtab.dest
+ |> subtract (op =) (rel_const_tab |> Symtab.dest)
+ |> map string_for_hyper_pconst));
+ map (fst o fst) accepts @
+ (if remaining_max = 0 then
+ game_over (hopeful_rejects @ map (apsnd SOME) hopeless_rejects)
+ else
+ iter (j + 1) remaining_max threshold rel_const_tab'
+ hopeless_rejects hopeful_rejects)
+ end
+ | relevant candidates rejects
+ (((ax as (((_, loc), th), axiom_consts)), cached_weight)
+ :: hopeful) =
+ let
+ val weight =
+ case cached_weight of
+ SOME w => w
+ | NONE => axiom_weight 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))
+else
+()
+*)
+ in
+ if weight >= threshold then
+ relevant ((ax, weight) :: candidates) rejects hopeful
else
- (* Add "add:" facts. *)
- if null add_thms then
- []
- else
- map_filter (fn ((p as (_, th), _), _) =>
- if member Thm.eq_thm add_thms th then SOME p
- else NONE) rejects
- | relevant (new_pairs, rejects) [] =
- let
- val (new_rels, more_rejects) = take_best max_new new_pairs
- val rel_const_tab' =
- rel_const_tab |> fold add_const_to_table (maps snd new_rels)
- fun is_dirty c =
- const_mem rel_const_tab' c andalso
- not (const_mem rel_const_tab c)
- val rejects =
- more_rejects @ rejects
- |> map (fn (ax as (_, consts), old_weight) =>
- (ax, if exists is_dirty consts then NONE
- else SOME old_weight))
- val threshold =
- threshold + (1.0 - threshold) * relevance_convergence
- in
- trace_msg (fn () => "relevant this iteration: " ^
- Int.toString (length new_rels));
- map #1 new_rels @ iter (j + 1) threshold rel_const_tab' rejects
- end
- | relevant (new_rels, rejects)
- (((ax as ((name, th), axiom_consts)), cached_weight)
- :: rest) =
- let
- val weight =
- case cached_weight of
- SOME w => w
- | NONE => axiom_weight const_tab rel_const_tab axiom_consts
- in
- if weight >= threshold orelse
- (defs_relevant andalso defines thy th rel_const_tab) then
- (trace_msg (fn () =>
- fst (name ()) ^ " passes: " ^ Real.toString weight
- ^ " consts: " ^ commas (map fst axiom_consts));
- relevant ((ax, weight) :: new_rels, rejects) rest)
- else
- relevant (new_rels, (ax, weight) :: rejects) rest
- end
- in
- trace_msg (fn () => "relevant_facts, current threshold: " ^
- Real.toString threshold);
- relevant ([], [])
- end
- in
- axioms |> filter_out (member Thm.eq_thm del_thms o snd)
- |> map (rpair NONE o pair_consts_axiom theory_relevant thy)
- |> iter 0 relevance_threshold goal_const_tab
- |> tap (fn res => trace_msg (fn () =>
+ relevant candidates ((ax, weight) :: rejects) hopeful
+ end
+ in
+ trace_msg (fn () =>
+ "ITERATION " ^ string_of_int j ^ ": current threshold: " ^
+ Real.toString threshold ^ ", constants: " ^
+ commas (rel_const_tab |> Symtab.dest
+ |> filter (curry (op <>) [] o snd)
+ |> map string_for_hyper_pconst));
+ relevant [] [] hopeful
+ end
+ in
+ axioms |> filter_out (member Thm.eq_thm del_thms o snd)
+ |> map_filter (pair_consts_axiom theory_relevant thy)
+ |> iter 0 max_relevant threshold0 goal_const_tab []
+ |> tap (fn res => trace_msg (fn () =>
"Total relevant: " ^ Int.toString (length res)))
- end
+ end
+
(***************************************************************)
(* Retrieving and filtering lemmas *)
@@ -410,9 +538,9 @@
String.isSuffix "_def" a orelse String.isSuffix "_defs" a
end;
-fun make_fact_table xs =
- fold (Termtab.update o `(prop_of o snd)) xs Termtab.empty
-fun make_unique xs = Termtab.fold (cons o snd) (make_fact_table xs) []
+fun mk_fact_table f xs =
+ fold (Termtab.update o `(prop_of o f)) xs Termtab.empty
+fun uniquify xs = Termtab.fold (cons o snd) (mk_fact_table snd xs) []
(* FIXME: put other record thms here, or declare as "no_atp" *)
val multi_base_blacklist =
@@ -453,12 +581,22 @@
val exists_sledgehammer_const =
exists_Const (fn (s, _) => String.isPrefix sledgehammer_prefix s)
-fun is_strange_theorem th =
+(* FIXME: make more reliable *)
+val exists_low_level_class_const =
+ exists_Const (fn (s, _) =>
+ String.isSubstring (Long_Name.separator ^ "class" ^ Long_Name.separator) s)
+
+fun is_metastrange_theorem th =
case head_of (concl_of th) of
Const (a, _) => (a <> @{const_name Trueprop} andalso
a <> @{const_name "=="})
| _ => false
+fun is_that_fact th =
+ String.isSuffix (Long_Name.separator ^ Obtain.thatN) (Thm.get_name_hint th)
+ andalso exists_subterm (fn Free (s, _) => s = Name.skolem Auto_Bind.thesisN
+ | _ => false) (prop_of th)
+
val type_has_top_sort =
exists_subtype (fn TFree (_, []) => true | TVar (_, []) => true | _ => false)
@@ -491,13 +629,13 @@
| (_, @{const "==>"} $ t1 $ t2) =>
do_formula (not pos) t1 andalso
(t2 = @{prop False} orelse do_formula pos t2)
- | (_, @{const "op -->"} $ t1 $ t2) =>
+ | (_, @{const HOL.implies} $ t1 $ t2) =>
do_formula (not pos) t1 andalso
(t2 = @{const False} orelse do_formula pos t2)
| (_, @{const Not} $ t1) => do_formula (not pos) t1
- | (true, @{const "op |"} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
- | (false, @{const "op &"} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
- | (true, Const (@{const_name "op ="}, _) $ t1 $ t2) => do_equals t1 t2
+ | (true, @{const HOL.disj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
+ | (false, @{const HOL.conj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
+ | (true, Const (@{const_name HOL.eq}, _) $ t1 $ t2) => do_equals t1 t2
| (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2
| _ => false
in do_formula true end
@@ -528,27 +666,42 @@
let val t = prop_of thm in
is_formula_too_complex t orelse exists_type type_has_top_sort t orelse
is_dangerous_term full_types t orelse exists_sledgehammer_const t orelse
- is_strange_theorem thm
+ exists_low_level_class_const t orelse is_metastrange_theorem thm orelse
+ is_that_fact thm
end
+fun clasimpset_rules_of ctxt =
+ let
+ val {safeIs, safeEs, hazIs, hazEs, ...} = ctxt |> claset_of |> rep_cs
+ val intros = safeIs @ hazIs
+ val elims = map Classical.classical_rule (safeEs @ hazEs)
+ val simps = ctxt |> simpset_of |> dest_ss |> #simps |> map snd
+ in (mk_fact_table I intros, mk_fact_table I elims, mk_fact_table I simps) end
+
fun all_name_thms_pairs ctxt reserved full_types add_thms chained_ths =
let
- val is_chained = member Thm.eq_thm chained_ths
- val global_facts = PureThy.facts_of (ProofContext.theory_of ctxt)
+ val thy = ProofContext.theory_of ctxt
+ val global_facts = PureThy.facts_of thy
val local_facts = ProofContext.facts_of ctxt
val named_locals = local_facts |> Facts.dest_static []
- (* Unnamed, not chained formulas with schematic variables are omitted,
- because they are rejected by the backticks (`...`) parser for some
- reason. *)
- fun is_bad_unnamed_local th =
- exists (fn (_, ths) => member Thm.eq_thm ths th) named_locals orelse
- (exists_subterm is_Var (prop_of th) andalso not (is_chained th))
+ 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
+ clasimpset_rules_of ctxt
+ else
+ (Termtab.empty, Termtab.empty, Termtab.empty)
+ (* Unnamed nonchained formulas with schematic variables are omitted, because
+ they are rejected by the backticks (`...`) parser for some reason. *)
+ fun is_good_unnamed_local th =
+ not (Thm.has_name_hint th) andalso
+ (not (exists_subterm is_Var (prop_of th)) orelse (is_chained th)) andalso
+ forall (fn (_, ths) => not (member Thm.eq_thm ths th)) named_locals
val unnamed_locals =
- local_facts |> Facts.props |> filter_out is_bad_unnamed_local
- |> map (pair "" o single)
+ union Thm.eq_thm (Facts.props local_facts) chained_ths
+ |> filter is_good_unnamed_local |> map (pair "" o single)
val full_space =
- Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts);
- fun add_valid_facts foldx facts =
+ Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts)
+ fun add_facts global foldx facts =
foldx (fn (name0, ths) =>
if name0 <> "" andalso
forall (not o member Thm.eq_thm add_thms) ths andalso
@@ -563,6 +716,8 @@
fun backquotify th =
"`" ^ Print_Mode.setmp [Print_Mode.input]
(Syntax.string_of_term ctxt) (prop_of th) ^ "`"
+ |> String.translate (fn c => if Char.isPrint c then str c else "")
+ |> simplify_spaces
fun check_thms a =
case try (ProofContext.get_thms ctxt) a of
NONE => false
@@ -575,66 +730,71 @@
not (member Thm.eq_thm add_thms th) then
rest
else
- (fn () =>
- (if name0 = "" then
- th |> backquotify
- else
- let
- val name1 = Facts.extern facts name0
- val name2 = Name_Space.extern full_space name0
- in
- case find_first check_thms [name1, name2, name0] of
- SOME name =>
- let
- val name =
- name |> Symtab.defined reserved name ? quote
- in
- if multi then name ^ "(" ^ Int.toString j ^ ")"
- else name
- end
- | NONE => ""
- end, is_chained th), (multi, th)) :: rest)) ths
+ (((fn () =>
+ if name0 = "" then
+ th |> backquotify
+ else
+ let
+ val name1 = Facts.extern facts name0
+ val name2 = Name_Space.extern full_space name0
+ in
+ case find_first check_thms [name1, name2, name0] of
+ SOME name => repair_name reserved multi j name
+ | NONE => ""
+ end),
+ let val t = prop_of th in
+ if is_chained th then Chained
+ else if not global then Local
+ else if Termtab.defined intros t then Intro
+ else if Termtab.defined elims t then Elim
+ else if Termtab.defined simps t then Simp
+ else General
+ end),
+ (multi, th)) :: rest)) ths
#> snd
end)
in
- [] |> add_valid_facts fold local_facts (unnamed_locals @ named_locals)
- |> add_valid_facts Facts.fold_static global_facts global_facts
+ [] |> add_facts false fold local_facts (unnamed_locals @ named_locals)
<