--- a/NEWS Fri Aug 27 16:04:15 2010 +0200
+++ b/NEWS Fri Aug 27 16:05:46 2010 +0200
@@ -23,6 +23,22 @@
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.
+
*** Pure ***
@@ -40,6 +56,9 @@
*** HOL ***
+* 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 +123,9 @@
Trueprop ~> HOL.Trueprop
True ~> HOL.True
False ~> HOL.False
+ op & ~> HOL.conj
+ op | ~> HOL.disj
+ op --> ~> HOL.implies
Not ~> HOL.Not
The ~> HOL.The
All ~> HOL.All
--- a/doc-src/Classes/Thy/Classes.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/Classes/Thy/Classes.thy Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/Classes/Thy/document/Classes.tex Fri Aug 27 16:05:46 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/Inductive_Predicate.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/Codegen/Thy/Inductive_Predicate.thy Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/Codegen/Thy/Introduction.thy Fri Aug 27 16:05:46 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/document/Adaptation.tex Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/Codegen/Thy/document/Adaptation.tex Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/Codegen/Thy/document/Inductive_Predicate.tex Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/Codegen/Thy/document/Introduction.tex Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/Codegen/Thy/pictures/architecture.tex Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/Codegen/codegen.tex Fri Aug 27 16:05:46 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/IsarRef/Thy/HOL_Specific.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/IsarRef/Thy/HOL_Specific.thy Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/IsarRef/Thy/ROOT-HOLCF.ML Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/IsarRef/Thy/ROOT-ZF.ML Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/IsarRef/Thy/ROOT.ML Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/IsarRef/Thy/document/HOL_Specific.tex Fri Aug 27 16:05:46 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/Main/Docs/Main_Doc.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/Main/Docs/Main_Doc.thy Fri Aug 27 16:05:46 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/System/Thy/ROOT.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/System/Thy/ROOT.ML Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/TutorialI/Documents/Documents.thy Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/TutorialI/Documents/document/Documents.tex Fri Aug 27 16:05:46 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/Rules/Primes.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/TutorialI/Rules/Primes.thy Fri Aug 27 16:05:46 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/Types/Numbers.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/TutorialI/Types/Numbers.thy Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/TutorialI/Types/document/Numbers.tex Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/TutorialI/settings.ML Fri Aug 27 16:05:46 2010 +0200
@@ -1,3 +1,1 @@
-(* $Id$ *)
-
-Thy_Output.indent := 5;
+Thy_Output.indent_default := 5;
--- a/doc-src/antiquote_setup.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/antiquote_setup.ML Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/manual.bib Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/more_antiquote.ML Fri Aug 27 16:05:46 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
--- a/doc-src/rail.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/doc-src/rail.ML Fri Aug 27 16:05:46 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/src/HOL/Auth/Yahalom.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Auth/Yahalom.thy Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Boogie/Tools/boogie_commands.ML Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Boogie/Tools/boogie_loader.ML Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Boogie/Tools/boogie_tactics.ML Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Boogie/Tools/boogie_vcs.ML Fri Aug 27 16:05:46 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/Decision_Procs/Approximation.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Decision_Procs/Approximation.thy Fri Aug 27 16:05:46 2010 +0200
@@ -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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Decision_Procs/Cooper.thy Fri Aug 27 16:05:46 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/Ferrack.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Decision_Procs/Ferrack.thy Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Decision_Procs/MIR.thy Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Fri Aug 27 16:05:46 2010 +0200
@@ -2954,9 +2954,9 @@
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);
@@ -3018,9 +3018,9 @@
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 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 "op ="},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))
--- a/src/HOL/Decision_Procs/ferrante_rackoff.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Decision_Procs/ferrante_rackoff.ML Fri Aug 27 16:05:46 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))
@@ -181,9 +181,9 @@
| 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Decision_Procs/langford.ML Fri Aug 27 16:05:46 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 =
@@ -183,9 +183,9 @@
| 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/HOL.thy Fri Aug 27 16:05:46 2010 +0200
@@ -57,13 +57,13 @@
False :: bool
Not :: "bool => bool" ("~ _" [40] 40)
+ conj :: "[bool, bool] => bool" (infixr "&" 35)
+ disj :: "[bool, bool] => bool" (infixr "|" 30)
+ implies :: "[bool, bool] => bool" (infixr "-->" 25)
+
setup Sign.root_path
consts
- "op &" :: "[bool, bool] => bool" (infixr "&" 35)
- "op |" :: "[bool, bool] => bool" (infixr "|" 30)
- "op -->" :: "[bool, bool] => bool" (infixr "-->" 25)
-
"op =" :: "['a, 'a] => bool" (infixl "=" 50)
setup Sign.local_path
@@ -89,15 +89,15 @@
notation (xsymbols)
Not ("\<not> _" [40] 40) and
- "op &" (infixr "\<and>" 35) and
- "op |" (infixr "\<or>" 30) and
- "op -->" (infixr "\<longrightarrow>" 25) and
+ HOL.conj (infixr "\<and>" 35) and
+ HOL.disj (infixr "\<or>" 30) and
+ HOL.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
+ HOL.conj (infixr "\<and>" 35) and
+ HOL.disj (infixr "\<or>" 30) and
not_equal (infix "\<noteq>" 50)
abbreviation (iff)
@@ -184,7 +184,7 @@
finalconsts
"op ="
- "op -->"
+ HOL.implies
The
definition If :: "bool \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" ("(if (_)/ then (_)/ else (_))" [0, 0, 10] 10) where
@@ -1578,7 +1578,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 +1737,8 @@
"True" ("true")
"False" ("false")
"Not" ("Bool.not")
- "op |" ("(_ orelse/ _)")
- "op &" ("(_ andalso/ _)")
+ HOL.disj ("(_ orelse/ _)")
+ HOL.conj ("(_ andalso/ _)")
"If" ("(if _/ then _/ else _)")
setup {*
@@ -1914,7 +1914,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 +1924,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 (_))")
--- a/src/HOL/Imperative_HOL/Array.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Imperative_HOL/Array.thy Fri Aug 27 16:05:46 2010 +0200
@@ -484,11 +484,11 @@
code_type array (Scala "!collection.mutable.ArraySeq[_]")
code_const Array (Scala "!error(\"bare Array\")")
-code_const Array.new' (Scala "('_: Unit)/ => / Array.alloc((_))((_))")
-code_const Array.make' (Scala "('_: Unit)/ =>/ Array.make((_))((_))")
-code_const Array.len' (Scala "('_: Unit)/ =>/ Array.len((_))")
-code_const Array.nth' (Scala "('_: Unit)/ =>/ Array.nth((_), (_))")
-code_const Array.upd' (Scala "('_: Unit)/ =>/ Array.upd((_), (_), (_))")
-code_const Array.freeze (Scala "('_: Unit)/ =>/ Array.freeze((_))")
+code_const Array.new' (Scala "('_: Unit)/ => / Heap.Array.alloc((_))((_))")
+code_const Array.make' (Scala "('_: Unit)/ =>/ Heap.Array.make((_))((_))")
+code_const Array.len' (Scala "('_: Unit)/ =>/ Heap.Array.len((_))")
+code_const Array.nth' (Scala "('_: Unit)/ =>/ Heap.Array.nth((_), (_))")
+code_const Array.upd' (Scala "('_: Unit)/ =>/ Heap.Array.upd((_), (_), (_))")
+code_const Array.freeze (Scala "('_: Unit)/ =>/ Heap.Array.freeze((_))")
end
--- a/src/HOL/Imperative_HOL/Heap_Monad.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Imperative_HOL/Heap_Monad.thy Fri Aug 27 16:05:46 2010 +0200
@@ -478,7 +478,6 @@
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 ()) ()
@@ -487,24 +486,33 @@
}
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
+ def alloc[A](n: Natural.Nat)(x: A): ArraySeq[A] =
+ ArraySeq.fill(n.as_Int)(x)
+ def make[A](n: Natural.Nat)(f: Natural.Nat => A): ArraySeq[A] =
+ ArraySeq.tabulate(n.as_Int)((k: Int) => f(Natural.Nat(k)))
+ def len[A](a: ArraySeq[A]): Natural.Nat =
+ Natural.Nat(a.length)
+ def nth[A](a: ArraySeq[A], n: Natural.Nat): A =
+ a(n.as_Int)
+ def upd[A](a: ArraySeq[A], n: Natural.Nat, 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_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/Imperative_HOL/Ref.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Imperative_HOL/Ref.thy Fri Aug 27 16:05:46 2010 +0200
@@ -306,10 +306,10 @@
text {* Scala *}
-code_type ref (Scala "!Ref[_]")
+code_type ref (Scala "!Heap.Ref[_]")
code_const Ref (Scala "!error(\"bare Ref\")")
-code_const ref' (Scala "('_: Unit)/ =>/ Ref((_))")
-code_const Ref.lookup (Scala "('_: Unit)/ =>/ Ref.lookup((_))")
-code_const Ref.update (Scala "('_: Unit)/ =>/ Ref.update((_), (_))")
+code_const ref' (Scala "('_: Unit)/ =>/ Heap.Ref((_))")
+code_const Ref.lookup (Scala "('_: Unit)/ =>/ Heap.Ref.lookup((_))")
+code_const Ref.update (Scala "('_: Unit)/ =>/ Heap.Ref.update((_), (_))")
end
--- a/src/HOL/Import/Generate-HOL/GenHOL4Base.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Import/Generate-HOL/GenHOL4Base.thy Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Import/Generate-HOLLight/GenHOLLight.thy Fri Aug 27 16:05:46 2010 +0200
@@ -54,9 +54,9 @@
ONE_ONE > HOL4Setup.ONE_ONE
ONTO > Fun.surj
"=" > "op ="
- "==>" > "op -->"
- "/\\" > "op &"
- "\\/" > "op |"
+ "==>" > HOL.implies
+ "/\\" > HOL.conj
+ "\\/" > HOL.disj
"!" > All
"?" > Ex
"?!" > Ex1
--- a/src/HOL/Import/HOL/bool.imp Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Import/HOL/bool.imp Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Import/HOLLight/hollight.imp Fri Aug 27 16:05:46 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 -->"
+ "==>" > HOL.implies
"=" > "op ="
"<=" > "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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Import/hol4rews.ML Fri Aug 27 16:05:46 2010 +0200
@@ -628,7 +628,7 @@
|> 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.implies}
|> add_hol4_const_mapping "min" "@" false @{const_name "Eps"}
in
val hol4_setup =
--- a/src/HOL/Import/proof_kernel.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Import/proof_kernel.ML Fri Aug 27 16:05:46 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 "op ="}
then I else insert (op =) c
| _ => I) t [];
@@ -1214,10 +1214,10 @@
fun add_consts (Const (c, _), cs) =
(case c of
@{const_name "op ="} => insert (op =) "==" cs
- | @{const_name "op -->"} => 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))
@@ -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
@@ -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/Library/Code_Integer.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Library/Code_Integer.thy Fri Aug 27 16:05:46 2010 +0200
@@ -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"
--- a/src/HOL/Library/Code_Natural.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Library/Code_Natural.thy Fri Aug 27 16:05:46 2010 +0200
@@ -52,50 +52,51 @@
| otherwise = (Natural k, Natural l) where (k, l) = quotRem n m;
};*}
+
code_reserved Haskell Natural
-code_include Scala "Natural" {*
-import scala.Math
+code_include Scala "Natural"
+{*import scala.Math
-object Natural {
+object Nat {
- def apply(numeral: BigInt): Natural = new Natural(numeral max 0)
- def apply(numeral: Int): Natural = Natural(BigInt(numeral))
- def apply(numeral: String): Natural = Natural(BigInt(numeral))
+ def apply(numeral: BigInt): Nat = new Nat(numeral max 0)
+ def apply(numeral: Int): Nat = Nat(BigInt(numeral))
+ def apply(numeral: String): Nat = Nat(BigInt(numeral))
}
-class Natural private(private val value: BigInt) {
+class Nat private(private val value: BigInt) {
override def hashCode(): Int = this.value.hashCode()
override def equals(that: Any): Boolean = that match {
- case that: Natural => this equals that
+ case that: Nat => this equals that
case _ => false
}
override def toString(): String = this.value.toString
- def equals(that: Natural): Boolean = this.value == that.value
+ def equals(that: Nat): Boolean = this.value == that.value
def as_BigInt: BigInt = this.value
def as_Int: Int = if (this.value >= Int.MinValue && this.value <= Int.MaxValue)
this.value.intValue
else error("Int value out of range: " + this.value.toString)
- def +(that: Natural): Natural = new Natural(this.value + that.value)
- def -(that: Natural): Natural = Natural(this.value - that.value)
- def *(that: Natural): Natural = new Natural(this.value * that.value)
+ def +(that: Nat): Nat = new Nat(this.value + that.value)
+ def -(that: Nat): Nat = Nat(this.value - that.value)
+ def *(that: Nat): Nat = new Nat(this.value * that.value)
- def /%(that: Natural): (Natural, Natural) = if (that.value == 0) (new Natural(0), this)
+ def /%(that: Nat): (Nat, Nat) = if (that.value == 0) (new Nat(0), this)
else {
val (k, l) = this.value /% that.value
- (new Natural(k), new Natural(l))
+ (new Nat(k), new Nat(l))
}
- def <=(that: Natural): Boolean = this.value <= that.value
+ def <=(that: Nat): Boolean = this.value <= that.value
- def <(that: Natural): Boolean = this.value < that.value
+ def <(that: Nat): Boolean = this.value < that.value
}
*}
@@ -104,7 +105,7 @@
code_type code_numeral
(Haskell "Natural.Natural")
- (Scala "Natural")
+ (Scala "Natural.Nat")
setup {*
fold (Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral}
--- a/src/HOL/Library/Efficient_Nat.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Library/Efficient_Nat.thy Fri Aug 27 16:05:46 2010 +0200
@@ -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,8 +280,8 @@
code_reserved Haskell Nat
-code_include Scala "Nat" {*
-import scala.Math
+code_include Scala "Nat"
+{*import scala.Math
object Nat {
@@ -330,7 +330,7 @@
code_type nat
(Haskell "Nat.Nat")
- (Scala "Nat")
+ (Scala "Nat.Nat")
code_instance nat :: eq
(Haskell -)
@@ -397,7 +397,7 @@
code_const int and nat
(Haskell "toInteger" and "fromInteger")
- (Scala "!_.as'_BigInt" and "Nat")
+ (Scala "!_.as'_BigInt" and "Nat.Nat")
text {* Conversion from and to code numerals. *}
@@ -405,14 +405,14 @@
(SML "IntInf.toInt")
(OCaml "_")
(Haskell "!(fromInteger/ ./ toInteger)")
- (Scala "!Natural(_.as'_BigInt)")
+ (Scala "!Natural.Nat(_.as'_BigInt)")
(Eval "_")
code_const Code_Numeral.nat_of
(SML "IntInf.fromInt")
(OCaml "_")
(Haskell "!(fromInteger/ ./ toInteger)")
- (Scala "!Nat(_.as'_BigInt)")
+ (Scala "!Nat.Nat(_.as'_BigInt)")
(Eval "_")
text {* Using target language arithmetic operations whenever appropriate *}
--- a/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML Fri Aug 27 16:05:46 2010 +0200
@@ -1356,7 +1356,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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Library/positivstellensatz.ML Fri Aug 27 16:05:46 2010 +0200
@@ -439,8 +439,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 +484,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 +542,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/Multivariate_Analysis/Gauge_Measure.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Gauge_Measure.thy Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Mutabelle/Mutabelle.thy Fri Aug 27 16:05:46 2010 +0200
@@ -4,7 +4,7 @@
begin
ML {*
-val comms = [@{const_name "op ="}, @{const_name "op |"}, @{const_name "op &"}];
+val comms = [@{const_name "op ="}, @{const_name HOL.disj}, @{const_name HOL.conj}];
val forbidden =
[(@{const_name Power.power}, "'a"),
--- a/src/HOL/Mutabelle/mutabelle_extra.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Mutabelle/mutabelle_extra.ML Fri Aug 27 16:05:46 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 "op ="}, @{const_name HOL.disj}, @{const_name HOL.conj}]
val forbidden =
[(* (@{const_name "power"}, "'a"), *)
--- a/src/HOL/Nominal/nominal_datatype.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Nominal/nominal_datatype.ML Fri Aug 27 16:05:46 2010 +0200
@@ -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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Nominal/nominal_inductive.ML Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Nominal/nominal_inductive2.ML Fri Aug 27 16:05:46 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/Orderings.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Orderings.thy Fri Aug 27 16:05:46 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_Compile_Examples/Code_Prolog_Examples.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy Fri Aug 27 16:05:46 2010 +0200
@@ -4,13 +4,24 @@
section {* Example append *}
+
inductive append
where
"append [] ys ys"
| "append xs ys zs ==> append (x # xs) ys (x # zs)"
+ML {* Code_Prolog.options := { ensure_groundness = false, limited_types = [], prolog_system = Code_Prolog.SWI_PROLOG} *}
+
+values "{(x, y, z). append x y z}"
+
values 3 "{(x, y, z). append x y z}"
+ML {* Code_Prolog.options := { ensure_groundness = false, limited_types = [], prolog_system = Code_Prolog.YAP} *}
+
+values "{(x, y, z). append x y z}"
+
+ML {* Code_Prolog.options := { ensure_groundness = false, limited_types = [], prolog_system = Code_Prolog.SWI_PROLOG} *}
+
section {* Example queens *}
@@ -172,7 +183,7 @@
where
"y \<noteq> B \<Longrightarrow> notB y"
-ML {* Code_Prolog.options := {ensure_groundness = true} *}
+ML {* Code_Prolog.options := {ensure_groundness = true, limited_types = [], prolog_system = Code_Prolog.SWI_PROLOG} *}
values 2 "{y. notB y}"
@@ -187,7 +198,7 @@
inductive equals :: "abc => abc => bool"
where
"equals y' y'"
-ML {* set Code_Prolog.trace *}
+
values 1 "{(y, z). equals y z}"
end
--- a/src/HOL/Prolog/prolog.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Prolog/prolog.ML Fri Aug 27 16:05:46 2010 +0200
@@ -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/Set.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Set.thy Fri Aug 27 16:05:46 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};
@@ -269,7 +269,7 @@
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 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,7 +288,7 @@
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 &"}, _) $
+ | check (Const (@{const_syntax HOL.conj}, _) $
(Const (@{const_syntax "op ="}, _) $ 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, []))
@@ -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/state_fun.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Statespace/state_fun.ML Fri Aug 27 16:05:46 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
--- a/src/HOL/TLA/Intensional.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/TLA/Intensional.thy Fri Aug 27 16:05:46 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/Datatype/datatype_aux.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_aux.ML Fri Aug 27 16:05:46 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_data.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_data.ML Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_prop.ML Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_realizer.ML Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Function/function_common.ML Fri Aug 27 16:05:46 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/measure_functions.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Function/measure_functions.ML Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Function/scnp_reconstruct.ML Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Function/termination.ML Fri Aug 27 16:05:46 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 "op ="}, _) $ _ $ (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 "op ="}, _) $ _ $ (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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Nitpick/minipick.ML Fri Aug 27 16:05:46 2010 +0200
@@ -130,9 +130,9 @@
[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 ()
@@ -173,12 +173,12 @@
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_hol.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Fri Aug 27 16:05:46 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
@@ -409,9 +409,9 @@
(@{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.conj}, 2),
+ (@{const_name HOL.disj}, 2),
+ (@{const_name HOL.implies}, 2),
(@{const_name If}, 3),
(@{const_name Let}, 2),
(@{const_name Pair}, 2),
@@ -1454,7 +1454,7 @@
| @{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 HOL.implies} $ _ $ t2 => lhs_of_equation t2
| _ => NONE
fun is_constr_pattern _ (Bound _) = true
| is_constr_pattern _ (Var _) = true
@@ -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
--- a/src/HOL/Tools/Nitpick/nitpick_mono.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_mono.ML Fri Aug 27 16:05:46 2010 +0200
@@ -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
@@ -883,12 +883,12 @@
| (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
@@ -975,8 +975,8 @@
do_all t0 s0 T1 t1 accum
| Const (x as (@{const_name "op ="}, _)) $ 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
--- a/src/HOL/Tools/Nitpick/nitpick_nut.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_nut.ML Fri Aug 27 16:05:46 2010 +0200
@@ -518,11 +518,11 @@
| (Const (@{const_name "op ="}, 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.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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_preproc.ML Fri Aug 27 16:05:46 2010 +0200
@@ -43,7 +43,7 @@
| aux def (@{const "==>"} $ t1 $ t2) = aux false t1 andalso aux def t2
| aux def (Const (@{const_name "op ="}, _) $ 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
@@ -211,14 +211,14 @@
do_quantifier new_Ts old_Ts polar s0 T0 s1 T1 t1
| Const (s0 as @{const_name "op ="}, 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
@@ -334,7 +334,7 @@
if def then (t, []) else do_eq_or_imp Ts false def t0 t1 t2 seen
| (t0 as Const (@{const_name "op ="}, _)) $ 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
@@ -401,7 +401,7 @@
t0 $ aux false t1 $ aux careful t2
| aux careful ((t0 as Const (@{const_name "op ="}, _)) $ 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
@@ -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)
@@ -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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/code_prolog.ML Fri Aug 27 16:05:46 2010 +0200
@@ -6,7 +6,9 @@
signature CODE_PROLOG =
sig
- type code_options = {ensure_groundness : bool}
+ datatype prolog_system = SWI_PROLOG | YAP
+ type code_options =
+ {ensure_groundness : bool, limited_types : (typ * int) list, prolog_system : prolog_system}
val options : code_options ref
datatype arith_op = Plus | Minus
@@ -21,10 +23,10 @@
type clause = ((string * prol_term list) * prem);
type logic_program = clause list;
type constant_table = (string * string) list
-
- val generate : code_options -> Proof.context -> string -> (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)
@@ -42,9 +44,13 @@
(* code generation options *)
-type code_options = {ensure_groundness : bool}
+datatype prolog_system = SWI_PROLOG | YAP
-val options = Unsynchronized.ref {ensure_groundness = false};
+type code_options =
+ {ensure_groundness : bool, limited_types : (typ * int) list, prolog_system : prolog_system}
+
+val options = Unsynchronized.ref {ensure_groundness = false, limited_types = [],
+ prolog_system = SWI_PROLOG};
(* general string functions *)
@@ -190,10 +196,10 @@
fun mk_groundness_prems t = map Ground (Term.add_frees t [])
-fun translate_prem options ctxt constant_table t =
+fun translate_prem ensure_groundness ctxt constant_table t =
case try HOLogic.dest_not t of
SOME t =>
- if #ensure_groundness options then
+ 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)
@@ -215,7 +221,7 @@
(Trueprop_conv (Conv.try_conv (Conv.rewr_conv @{thm Predicate.eq_is_eq}))))
(Thm.transfer thy rule)
-fun translate_intros options 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
@@ -225,33 +231,11 @@
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 options ctxt' constant_table') prems)
+ 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 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 depending_preds_of (key, intros) =
fold Term.add_const_names (map Thm.prop_of intros) []
@@ -272,7 +256,7 @@
fst (extend' key (G, []))
end
-fun generate options ctxt const =
+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)
@@ -281,10 +265,11 @@
val scc = strong_conn_of gr' [const]
val constant_table = mk_constant_table (flat scc)
in
- apfst flat (fold_map (translate_intros options ctxt gr) (flat scc) constant_table)
+ apfst flat (fold_map (translate_intros ensure_groundness ctxt gr) (flat scc) constant_table)
end
-(* add implementation for ground predicates *)
+(* implementation for fully enumerating predicates and
+ for size-limited predicates for enumerating the values of a datatype upto a specific size *)
fun add_ground_typ (Conj prems) = fold add_ground_typ prems
| add_ground_typ (Ground (_, T)) = insert (op =) T
@@ -294,34 +279,58 @@
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 is_recursive_constr T (Const (constr_name, T')) = member (op =) (binder_types T') T
-fun mk_ground_impl ctxt (T as Type (Tcon, Targs)) (seen, constant_table) =
+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 rel_name = mk_relname T
- fun mk_impl (Const (constr_name, T)) (seen, constant_table) =
+ 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) (binder_types T) (seen, constant_table')
- val vars = map (fn i => Var ("x" ^ string_of_int i)) (1 upto (length (binder_types T)))
- fun mk_prem v T = Rel (mk_relname T, [v])
+ 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, [maybe_AppF (translate_const constant_table'' constr_name, vars)]),
- Conj (map2 mk_prem vars (binder_types T)))
+ ((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
- in apfst flat (fold_map mk_impl constrs (T :: seen, constant_table)) end
- | mk_ground_impl ctxt T (seen, constant_table) =
+ 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 replace_ground (Conj prems) = Conj (map replace_ground prems)
@@ -329,15 +338,16 @@
Rel (mk_relname T, [Var x])
| replace_ground p = p
-fun add_ground_predicates ctxt (p, constant_table) =
+fun add_ground_predicates ctxt limited_types (p, constant_table) =
let
val ground_typs = fold (add_ground_typ o snd) p []
- val (grs, (_, constant_table')) = fold_map (mk_ground_impl ctxt) ground_typs ([], constant_table)
+ val (grs, (_, constant_table')) = fold_map (mk_ground_impl ctxt limited_types) ground_typs ([], constant_table)
val p' = map (apsnd replace_ground) p
in
((flat grs) @ p', constant_table')
end
-
+
+
(* rename variables to prolog-friendly names *)
fun rename_vars_term renaming = map_vars (fn v => the (AList.lookup (op =) renaming v))
@@ -396,14 +406,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" ^
@@ -411,7 +423,7 @@
"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" ^
@@ -420,7 +432,38 @@
"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 bash_output (cmd ^ file_name) end
+
(* parsing prolog solution *)
+
val scan_number =
Scan.many1 Symbol.is_ascii_digit
@@ -471,26 +514,25 @@
(* 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 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 ^ query query_rel vnames' ^ write_program p'
+ 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, _) = invoke system (File.shell_path 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
@@ -509,6 +551,30 @@
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 = !options
@@ -530,14 +596,18 @@
val _ = tracing "Preprocessing specification..."
val T = Sign.the_const_type (ProofContext.theory_of ctxt) name
val t = Const (name, T)
- val ctxt' = ProofContext.theory (Context.copy_thy) ctxt
- val thy' = Predicate_Compile.preprocess preprocess_options t (ProofContext.theory_of ctxt')
- val ctxt'' = ProofContext.init_global thy'
+ 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 options ctxt'' name
- |> (if #ensure_groundness options then add_ground_predicates ctxt'' else I)
+ val (p, constant_table) = generate (#ensure_groundness options) ctxt' name
+ |> (if #ensure_groundness options then
+ add_ground_predicates ctxt' (#limited_types options)
+ else I)
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 =
@@ -553,7 +623,7 @@
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"})))
@@ -564,7 +634,7 @@
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 =
@@ -595,20 +665,19 @@
(* FIXME: large copy of Predicate_Compile_Quickcheck - refactor out commons *)
-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);
fun quickcheck ctxt report t size =
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 Ts = map snd vs'
val t'' = subst_bounds (map Free (rev vs'), t')
val (prems, concl) = strip_horn t''
@@ -624,15 +693,15 @@
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 ctxt' = ProofContext.init_global thy3
val _ = tracing "Generating prolog program..."
- val (p, constant_table) = generate {ensure_groundness = true} ctxt'' full_constname
- |> add_ground_predicates ctxt''
+ val (p, constant_table) = generate true ctxt' full_constname
+ |> add_ground_predicates ctxt' (#limited_types (!options))
val _ = tracing "Running prolog program..."
- val [ts] = run p (translate_const constant_table full_constname) (map fst vs')
- (SOME 1)
+ val [ts] = run (#prolog_system (!options))
+ p (translate_const constant_table full_constname) (map fst vs') (SOME 1)
val _ = tracing "Restoring terms..."
- val res = SOME (map (restore_term ctxt'' constant_table) (ts ~~ Ts))
+ val res = SOME (map (restore_term ctxt' constant_table) (ts ~~ Ts))
val empty_report = ([], false)
in
(res, empty_report)
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile.ML Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Fri Aug 27 16:05:46 2010 +0200
@@ -405,7 +405,7 @@
(* 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 [];
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Fri Aug 27 16:05:46 2010 +0200
@@ -524,7 +524,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]
@@ -730,9 +730,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 +3031,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 +3062,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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML Fri Aug 27 16:05:46 2010 +0200
@@ -218,11 +218,11 @@
@{const_name Trueprop},
@{const_name Not},
@{const_name "op ="},
- @{const_name "op -->"},
+ @{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},
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML Fri Aug 27 16:05:46 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;
)
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Fri Aug 27 16:05:46 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
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_specialisation.ML Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Qelim/cooper.ML Fri Aug 27 16:05:46 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,8 +120,8 @@
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 HOL.conj},_)$_$_ => And (Thm.dest_binop ct)
+| Const (@{const_name HOL.disj},_)$_$_ => 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$_) =>
if term_of x aconv y then NEq (funpow 2 Thm.dest_arg ct) else Nox
@@ -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,8 +382,8 @@
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 =)
@@ -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)
@@ -687,7 +687,7 @@
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 +712,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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Qelim/qelim.ML Fri Aug 27 16:05:46 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 "op ="}] 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Quotient/quotient_info.ML Fri Aug 27 16:05:46 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_term.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Quotient/quotient_term.ML Fri Aug 27 16:05:46 2010 +0200
@@ -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
--- a/src/HOL/Tools/Quotient/quotient_typ.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Quotient/quotient_typ.ML Fri Aug 27 16:05:46 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/smtlib_interface.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/SMT/smtlib_interface.ML Fri Aug 27 16:05:46 2010 +0200
@@ -159,9 +159,9 @@
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 HOL.conj} = SOME "and"
+ | conn @{const_name HOL.disj} = SOME "or"
+ | conn @{const_name HOL.implies} = SOME "implies"
| conn @{const_name "op ="} = SOME "iff"
| conn @{const_name If} = SOME "if_then_else"
| conn _ = NONE
--- a/src/HOL/Tools/SMT/z3_interface.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/SMT/z3_interface.ML Fri Aug 27 16:05:46 2010 +0200
@@ -170,7 +170,7 @@
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
@@ -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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/SMT/z3_proof_literals.ML Fri Aug 27 16:05:46 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}
--- a/src/HOL/Tools/SMT/z3_proof_reconstruction.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/SMT/z3_proof_reconstruction.ML Fri Aug 27 16:05:46 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 Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/SMT/z3_proof_tools.ML Fri Aug 27 16:05:46 2010 +0200
@@ -196,9 +196,9 @@
| @{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
+ | @{term HOL.conj} $ _ $ _ => abstr2 cvs ct
+ | @{term HOL.disj} $ _ $ _ => abstr2 cvs ct
+ | @{term HOL.implies} $ _ $ _ => abstr2 cvs ct
| Const (@{const_name "op ="}, _) $ _ $ _ => abstr2 cvs ct
| Const (@{const_name distinct}, _) $ _ =>
if ext_logic then abs_arg (abs_list abstr fresh_abstraction) cvs ct
--- a/src/HOL/Tools/Sledgehammer/clausifier.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/clausifier.ML Fri Aug 27 16:05:46 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;
--- a/src/HOL/Tools/Sledgehammer/metis_clauses.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/metis_clauses.ML Fri Aug 27 16:05:46 2010 +0200
@@ -95,9 +95,9 @@
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.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"),
@@ -430,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
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML Fri Aug 27 16:05:46 2010 +0200
@@ -173,9 +173,9 @@
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 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 =>
fold (do_term_or_formula T) [t1, t2]
@@ -551,12 +551,12 @@
| (_, @{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 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 "op ="}, _) $ t1 $ t2) => do_equals t1 t2
| (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2
| _ => false
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML Fri Aug 27 16:05:46 2010 +0200
@@ -72,12 +72,12 @@
Const (@{const_name Ex}, T) $ Abs (s, T', negate_term t')
| negate_term (Const (@{const_name Ex}, T) $ Abs (s, T', t')) =
Const (@{const_name All}, T) $ Abs (s, T', negate_term t')
- | negate_term (@{const "op -->"} $ t1 $ t2) =
- @{const "op &"} $ t1 $ negate_term t2
- | negate_term (@{const "op &"} $ t1 $ t2) =
- @{const "op |"} $ negate_term t1 $ negate_term t2
- | negate_term (@{const "op |"} $ t1 $ t2) =
- @{const "op &"} $ negate_term t1 $ negate_term t2
+ | negate_term (@{const HOL.implies} $ t1 $ t2) =
+ @{const HOL.conj} $ t1 $ negate_term t2
+ | negate_term (@{const HOL.conj} $ t1 $ t2) =
+ @{const HOL.disj} $ negate_term t1 $ negate_term t2
+ | negate_term (@{const HOL.disj} $ t1 $ t2) =
+ @{const HOL.conj} $ negate_term t1 $ negate_term t2
| negate_term (@{const Not} $ t) = t
| negate_term t = @{const Not} $ t
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_translate.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_translate.ML Fri Aug 27 16:05:46 2010 +0200
@@ -70,9 +70,9 @@
do_quant bs AForall s T t'
| Const (@{const_name Ex}, _) $ Abs (s, T, t') =>
do_quant bs AExists s T t'
- | @{const "op &"} $ t1 $ t2 => do_conn bs AAnd t1 t2
- | @{const "op |"} $ t1 $ t2 => do_conn bs AOr t1 t2
- | @{const "op -->"} $ t1 $ t2 => do_conn bs AImplies t1 t2
+ | @{const HOL.conj} $ t1 $ t2 => do_conn bs AAnd t1 t2
+ | @{const HOL.disj} $ t1 $ t2 => do_conn bs AOr t1 t2
+ | @{const HOL.implies} $ t1 $ t2 => do_conn bs AImplies t1 t2
| Const (@{const_name "op ="}, Type (_, [@{typ bool}, _])) $ t1 $ t2 =>
do_conn bs AIff t1 t2
| _ => (fn ts => do_term bs (Envir.eta_contract t)
@@ -125,9 +125,9 @@
t0 $ Abs (s, T, aux (T :: Ts) t')
| (t0 as Const (@{const_name Ex}, _)) $ t1 =>
aux Ts (t0 $ eta_expand Ts t1 1)
- | (t0 as @{const "op &"}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
- | (t0 as @{const "op |"}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
- | (t0 as @{const "op -->"}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
+ | (t0 as @{const HOL.conj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
+ | (t0 as @{const HOL.disj}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
+ | (t0 as @{const HOL.implies}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2
| (t0 as Const (@{const_name "op ="}, Type (_, [@{typ bool}, _])))
$ t1 $ t2 =>
t0 $ aux Ts t1 $ aux Ts t2
--- a/src/HOL/Tools/TFL/dcterm.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/TFL/dcterm.ML Fri Aug 27 16:05:46 2010 +0200
@@ -79,11 +79,11 @@
in capply c (uncurry cabs r) end;
-local val c = mk_hol_const(@{const_name "op &"}, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
+local val c = mk_hol_const(@{const_name HOL.conj}, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
in fun mk_conj(conj1,conj2) = capply (capply c conj1) conj2
end;
-local val c = mk_hol_const(@{const_name "op |"}, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
+local val c = mk_hol_const(@{const_name HOL.disj}, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
in fun mk_disj(disj1,disj2) = capply (capply c disj1) disj2
end;
@@ -128,9 +128,9 @@
val dest_neg = dest_monop @{const_name Not}
val dest_pair = dest_binop @{const_name Pair}
val dest_eq = dest_binop @{const_name "op ="}
-val dest_imp = dest_binop @{const_name "op -->"}
-val dest_conj = dest_binop @{const_name "op &"}
-val dest_disj = dest_binop @{const_name "op |"}
+val dest_imp = dest_binop @{const_name HOL.implies}
+val dest_conj = dest_binop @{const_name HOL.conj}
+val dest_disj = dest_binop @{const_name HOL.disj}
val dest_select = dest_binder @{const_name Eps}
val dest_exists = dest_binder @{const_name Ex}
val dest_forall = dest_binder @{const_name All}
--- a/src/HOL/Tools/TFL/usyntax.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/TFL/usyntax.ML Fri Aug 27 16:05:46 2010 +0200
@@ -159,7 +159,7 @@
fun mk_imp{ant,conseq} =
- let val c = Const(@{const_name "op -->"},HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
+ let val c = Const(@{const_name HOL.implies},HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
in list_comb(c,[ant,conseq])
end;
@@ -183,12 +183,12 @@
fun mk_conj{conj1,conj2} =
- let val c = Const(@{const_name "op &"},HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
+ let val c = Const(@{const_name HOL.conj},HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
in list_comb(c,[conj1,conj2])
end;
fun mk_disj{disj1,disj2} =
- let val c = Const(@{const_name "op |"},HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
+ let val c = Const(@{const_name HOL.disj},HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT)
in list_comb(c,[disj1,disj2])
end;
@@ -247,7 +247,7 @@
fun dest_eq(Const(@{const_name "op ="},_) $ M $ N) = {lhs=M, rhs=N}
| dest_eq _ = raise USYN_ERR "dest_eq" "not an equality";
-fun dest_imp(Const(@{const_name "op -->"},_) $ M $ N) = {ant=M, conseq=N}
+fun dest_imp(Const(@{const_name HOL.implies},_) $ M $ N) = {ant=M, conseq=N}
| dest_imp _ = raise USYN_ERR "dest_imp" "not an implication";
fun dest_forall(Const(@{const_name All},_) $ (a as Abs _)) = fst (dest_abs [] a)
@@ -259,10 +259,10 @@
fun dest_neg(Const(@{const_name Not},_) $ M) = M
| dest_neg _ = raise USYN_ERR "dest_neg" "not a negation";
-fun dest_conj(Const(@{const_name "op &"},_) $ M $ N) = {conj1=M, conj2=N}
+fun dest_conj(Const(@{const_name HOL.conj},_) $ M $ N) = {conj1=M, conj2=N}
| dest_conj _ = raise USYN_ERR "dest_conj" "not a conjunction";
-fun dest_disj(Const(@{const_name "op |"},_) $ M $ N) = {disj1=M, disj2=N}
+fun dest_disj(Const(@{const_name HOL.disj},_) $ M $ N) = {disj1=M, disj2=N}
| dest_disj _ = raise USYN_ERR "dest_disj" "not a disjunction";
fun mk_pair{fst,snd} =
--- a/src/HOL/Tools/choice_specification.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/choice_specification.ML Fri Aug 27 16:05:46 2010 +0200
@@ -220,8 +220,8 @@
|> process_all (zip3 alt_names rew_imps frees)
end
- fun after_qed [[thm]] = (ProofContext.theory (fn thy =>
- #1 (post_process (add_specification axiomatic (zip3 names cnames overloaded) (thy, thm)))));
+ fun after_qed [[thm]] = ProofContext.background_theory (fn thy =>
+ #1 (post_process (add_specification axiomatic (zip3 names cnames overloaded) (thy, thm))));
in
thy
|> ProofContext.init_global
--- a/src/HOL/Tools/cnf_funcs.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/cnf_funcs.ML Fri Aug 27 16:05:46 2010 +0200
@@ -95,9 +95,9 @@
fun is_atom (Const (@{const_name False}, _)) = false
| is_atom (Const (@{const_name True}, _)) = false
- | is_atom (Const (@{const_name "op &"}, _) $ _ $ _) = false
- | is_atom (Const (@{const_name "op |"}, _) $ _ $ _) = false
- | is_atom (Const (@{const_name "op -->"}, _) $ _ $ _) = false
+ | is_atom (Const (@{const_name HOL.conj}, _) $ _ $ _) = false
+ | is_atom (Const (@{const_name HOL.disj}, _) $ _ $ _) = false
+ | is_atom (Const (@{const_name HOL.implies}, _) $ _ $ _) = false
| is_atom (Const (@{const_name "op ="}, Type ("fun", @{typ bool} :: _)) $ _ $ _) = false
| is_atom (Const (@{const_name Not}, _) $ _) = false
| is_atom _ = true;
@@ -105,7 +105,7 @@
fun is_literal (Const (@{const_name Not}, _) $ x) = is_atom x
| is_literal x = is_atom x;
-fun is_clause (Const (@{const_name "op |"}, _) $ x $ y) = is_clause x andalso is_clause y
+fun is_clause (Const (@{const_name HOL.disj}, _) $ x $ y) = is_clause x andalso is_clause y
| is_clause x = is_literal x;
(* ------------------------------------------------------------------------- *)
@@ -184,21 +184,21 @@
(* Theory.theory -> Term.term -> Thm.thm *)
-fun make_nnf_thm thy (Const (@{const_name "op &"}, _) $ x $ y) =
+fun make_nnf_thm thy (Const (@{const_name HOL.conj}, _) $ x $ y) =
let
val thm1 = make_nnf_thm thy x
val thm2 = make_nnf_thm thy y
in
conj_cong OF [thm1, thm2]
end
- | make_nnf_thm thy (Const (@{const_name "op |"}, _) $ x $ y) =
+ | make_nnf_thm thy (Const (@{const_name HOL.disj}, _) $ x $ y) =
let
val thm1 = make_nnf_thm thy x
val thm2 = make_nnf_thm thy y
in
disj_cong OF [thm1, thm2]
end
- | make_nnf_thm thy (Const (@{const_name "op -->"}, _) $ x $ y) =
+ | make_nnf_thm thy (Const (@{const_name HOL.implies}, _) $ x $ y) =
let
val thm1 = make_nnf_thm thy (HOLogic.Not $ x)
val thm2 = make_nnf_thm thy y
@@ -218,21 +218,21 @@
make_nnf_not_false
| make_nnf_thm thy (Const (@{const_name Not}, _) $ Const (@{const_name True}, _)) =
make_nnf_not_true
- | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name "op &"}, _) $ x $ y)) =
+ | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.conj}, _) $ x $ y)) =
let
val thm1 = make_nnf_thm thy (HOLogic.Not $ x)
val thm2 = make_nnf_thm thy (HOLogic.Not $ y)
in
make_nnf_not_conj OF [thm1, thm2]
end
- | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name "op |"}, _) $ x $ y)) =
+ | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.disj}, _) $ x $ y)) =
let
val thm1 = make_nnf_thm thy (HOLogic.Not $ x)
val thm2 = make_nnf_thm thy (HOLogic.Not $ y)
in
make_nnf_not_disj OF [thm1, thm2]
end
- | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name "op -->"}, _) $ x $ y)) =
+ | make_nnf_thm thy (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.implies}, _) $ x $ y)) =
let
val thm1 = make_nnf_thm thy x
val thm2 = make_nnf_thm thy (HOLogic.Not $ y)
@@ -276,7 +276,7 @@
(* Theory.theory -> Term.term -> Thm.thm *)
-fun simp_True_False_thm thy (Const (@{const_name "op &"}, _) $ x $ y) =
+fun simp_True_False_thm thy (Const (@{const_name HOL.conj}, _) $ x $ y) =
let
val thm1 = simp_True_False_thm thy x
val x' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm1
@@ -298,7 +298,7 @@
conj_cong OF [thm1, thm2] (* (x & y) = (x' & y') *)
end
end
- | simp_True_False_thm thy (Const (@{const_name "op |"}, _) $ x $ y) =
+ | simp_True_False_thm thy (Const (@{const_name HOL.disj}, _) $ x $ y) =
let
val thm1 = simp_True_False_thm thy x
val x' = (snd o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) thm1
@@ -334,24 +334,24 @@
fun make_cnf_thm thy t =
let
(* Term.term -> Thm.thm *)
- fun make_cnf_thm_from_nnf (Const (@{const_name "op &"}, _) $ x $ y) =
+ fun make_cnf_thm_from_nnf (Const (@{const_name HOL.conj}, _) $ x $ y) =
let
val thm1 = make_cnf_thm_from_nnf x
val thm2 = make_cnf_thm_from_nnf y
in
conj_cong OF [thm1, thm2]
end
- | make_cnf_thm_from_nnf (Const (@{const_name "op |"}, _) $ x $ y) =
+ | make_cnf_thm_from_nnf (Const (@{const_name HOL.disj}, _) $ x $ y) =
let
(* produces a theorem "(x' | y') = t'", where x', y', and t' are in CNF *)
- fun make_cnf_disj_thm (Const (@{const_name "op &"}, _) $ x1 $ x2) y' =
+ fun make_cnf_disj_thm (Const (@{const_name HOL.conj}, _) $ x1 $ x2) y' =
let
val thm1 = make_cnf_disj_thm x1 y'
val thm2 = make_cnf_disj_thm x2 y'
in
make_cnf_disj_conj_l OF [thm1, thm2] (* ((x1 & x2) | y') = ((x1 | y')' & (x2 | y')') *)
end
- | make_cnf_disj_thm x' (Const (@{const_name "op &"}, _) $ y1 $ y2) =
+ | make_cnf_disj_thm x' (Const (@{const_name HOL.conj}, _) $ y1 $ y2) =
let
val thm1 = make_cnf_disj_thm x' y1
val thm2 = make_cnf_disj_thm x' y2
@@ -403,27 +403,27 @@
val var_id = Unsynchronized.ref 0 (* properly initialized below *)
fun new_free () =
Free ("cnfx_" ^ string_of_int (Unsynchronized.inc var_id), HOLogic.boolT)
- fun make_cnfx_thm_from_nnf (Const (@{const_name "op &"}, _) $ x $ y) : thm =
+ fun make_cnfx_thm_from_nnf (Const (@{const_name HOL.conj}, _) $ x $ y) : thm =
let
val thm1 = make_cnfx_thm_from_nnf x
val thm2 = make_cnfx_thm_from_nnf y
in
conj_cong OF [thm1, thm2]
end
- | make_cnfx_thm_from_nnf (Const (@{const_name "op |"}, _) $ x $ y) =
+ | make_cnfx_thm_from_nnf (Const (@{const_name HOL.disj}, _) $ x $ y) =
if is_clause x andalso is_clause y then
inst_thm thy [HOLogic.mk_disj (x, y)] iff_refl
else if is_literal y orelse is_literal x then let
(* produces a theorem "(x' | y') = t'", where x', y', and t' are *)
(* almost in CNF, and x' or y' is a literal *)
- fun make_cnfx_disj_thm (Const (@{const_name "op &"}, _) $ x1 $ x2) y' =
+ fun make_cnfx_disj_thm (Const (@{const_name HOL.conj}, _) $ x1 $ x2) y' =
let
val thm1 = make_cnfx_disj_thm x1 y'
val thm2 = make_cnfx_disj_thm x2 y'
in
make_cnf_disj_conj_l OF [thm1, thm2] (* ((x1 & x2) | y') = ((x1 | y')' & (x2 | y')') *)
end
- | make_cnfx_disj_thm x' (Const (@{const_name "op &"}, _) $ y1 $ y2) =
+ | make_cnfx_disj_thm x' (Const (@{const_name HOL.conj}, _) $ y1 $ y2) =
let
val thm1 = make_cnfx_disj_thm x' y1
val thm2 = make_cnfx_disj_thm x' y2
--- a/src/HOL/Tools/groebner.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/groebner.ML Fri Aug 27 16:05:46 2010 +0200
@@ -395,7 +395,7 @@
fun refute_disj rfn tm =
case term_of tm of
- Const(@{const_name "op |"},_)$l$r =>
+ Const(@{const_name HOL.disj},_)$l$r =>
compose_single(refute_disj rfn (dest_arg tm),2,compose_single(refute_disj rfn (dest_arg1 tm),2,disjE))
| _ => rfn tm ;
@@ -454,7 +454,7 @@
val specl = fold_rev (fn x => fn th => instantiate' [] [SOME x] (th RS spec));
val cTrp = @{cterm "Trueprop"};
-val cConj = @{cterm "op &"};
+val cConj = @{cterm HOL.conj};
val (cNot,false_tm) = (@{cterm "Not"}, @{cterm "False"});
val assume_Trueprop = mk_comb cTrp #> assume;
val list_mk_conj = list_mk_binop cConj;
@@ -479,22 +479,22 @@
(* FIXME : copied from cqe.ML -- complex QE*)
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 =
@@ -852,14 +852,14 @@
fun unwind_polys_conv tm =
let
val (vars,bod) = strip_exists tm
- val cjs = striplist (dest_binary @{cterm "op &"}) bod
+ val cjs = striplist (dest_binary @{cterm HOL.conj}) bod
val th1 = (the (get_first (try (isolate_variable vars)) cjs)
handle Option => raise CTERM ("unwind_polys_conv",[tm]))
val eq = Thm.lhs_of th1
- val bod' = list_mk_binop @{cterm "op &"} (eq::(remove op aconvc eq cjs))
+ val bod' = list_mk_binop @{cterm HOL.conj} (eq::(remove op aconvc eq cjs))
val th2 = conj_ac_rule (mk_eq bod bod')
val th3 = transitive th2
- (Drule.binop_cong_rule @{cterm "op &"} th1
+ (Drule.binop_cong_rule @{cterm HOL.conj} th1
(reflexive (Thm.dest_arg (Thm.rhs_of th2))))
val v = Thm.dest_arg1(Thm.dest_arg1(Thm.rhs_of th3))
val vars' = (remove op aconvc v vars) @ [v]
@@ -929,9 +929,9 @@
| Const (@{const_name Not}, _) $ _ => find_term bounds (dest_arg tm)
| Const (@{const_name All}, _) $ _ => find_body bounds (dest_arg tm)
| Const (@{const_name Ex}, _) $ _ => find_body bounds (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
| @{term "op ==>"} $_$_ => find_args bounds tm
| Const("op ==",_)$_$_ => find_args bounds tm
| @{term Trueprop}$_ => find_term bounds (dest_arg tm)
--- a/src/HOL/Tools/hologic.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/hologic.ML Fri Aug 27 16:05:46 2010 +0200
@@ -208,29 +208,29 @@
let val (th1, th2) = conj_elim th
in conj_elims th1 @ conj_elims th2 end handle THM _ => [th];
-val conj = @{term "op &"}
-and disj = @{term "op |"}
-and imp = @{term "op -->"}
-and Not = @{term "Not"};
+val conj = @{term HOL.conj}
+and disj = @{term HOL.disj}
+and imp = @{term implies}
+and Not = @{term Not};
fun mk_conj (t1, t2) = conj $ t1 $ t2
and mk_disj (t1, t2) = disj $ t1 $ t2
and mk_imp (t1, t2) = imp $ t1 $ t2
and mk_not t = Not $ t;
-fun dest_conj (Const ("op &", _) $ t $ t') = t :: dest_conj t'
+fun dest_conj (Const ("HOL.conj", _) $ t $ t') = t :: dest_conj t'
| dest_conj t = [t];
-fun dest_disj (Const ("op |", _) $ t $ t') = t :: dest_disj t'
+fun dest_disj (Const ("HOL.disj", _) $ t $ t') = t :: dest_disj t'
| dest_disj t = [t];
(*Like dest_disj, but flattens disjunctions however nested*)
-fun disjuncts_aux (Const ("op |", _) $ t $ t') disjs = disjuncts_aux t (disjuncts_aux t' disjs)
+fun disjuncts_aux (Const ("HOL.disj", _) $ t $ t') disjs = disjuncts_aux t (disjuncts_aux t' disjs)
| disjuncts_aux t disjs = t::disjs;
fun disjuncts t = disjuncts_aux t [];
-fun dest_imp (Const ("op -->", _) $ A $ B) = (A, B)
+fun dest_imp (Const ("HOL.implies", _) $ A $ B) = (A, B)
| dest_imp t = raise TERM ("dest_imp", [t]);
fun dest_not (Const ("HOL.Not", _) $ t) = t
--- a/src/HOL/Tools/inductive_set.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/inductive_set.ML Fri Aug 27 16:05:46 2010 +0200
@@ -74,9 +74,9 @@
in Drule.instantiate' [] (rev (map (SOME o cterm_of thy o Var) vs))
(p (fold (Logic.all o Var) vs t) f)
end;
- fun mkop @{const_name "op &"} T x =
+ fun mkop @{const_name HOL.conj} T x =
SOME (Const (@{const_name Lattices.inf}, T --> T --> T), x)
- | mkop @{const_name "op |"} T x =
+ | mkop @{const_name HOL.disj} T x =
SOME (Const (@{const_name Lattices.sup}, T --> T --> T), x)
| mkop _ _ _ = NONE;
fun mk_collect p T t =
--- a/src/HOL/Tools/int_arith.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/int_arith.ML Fri Aug 27 16:05:46 2010 +0200
@@ -91,7 +91,7 @@
fun number_of thy T n =
if not (Sign.of_sort thy (T, @{sort number}))
then raise CTERM ("number_of", [])
- else Numeral.mk_cnumber (Thm.ctyp_of thy T) n
+ else Numeral.mk_cnumber (Thm.ctyp_of thy T) n;
val setup =
Lin_Arith.add_inj_thms [@{thm zle_int} RS iffD2, @{thm int_int_eq} RS iffD2]
--- a/src/HOL/Tools/lin_arith.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/lin_arith.ML Fri Aug 27 16:05:46 2010 +0200
@@ -16,8 +16,7 @@
val add_simprocs: simproc list -> Context.generic -> Context.generic
val add_inj_const: string * typ -> Context.generic -> Context.generic
val add_discrete_type: string -> Context.generic -> Context.generic
- val set_number_of: (theory -> typ -> int -> cterm) -> Context.generic ->
- Context.generic
+ val set_number_of: (theory -> typ -> int -> cterm) -> Context.generic -> Context.generic
val setup: Context.generic -> Context.generic
val global_setup: theory -> theory
val split_limit: int Config.T
@@ -46,7 +45,7 @@
val mk_Trueprop = HOLogic.mk_Trueprop;
fun atomize thm = case Thm.prop_of thm of
- Const (@{const_name Trueprop}, _) $ (Const (@{const_name "op &"}, _) $ _ $ _) =>
+ Const (@{const_name Trueprop}, _) $ (Const (@{const_name HOL.conj}, _) $ _ $ _) =>
atomize (thm RS conjunct1) @ atomize (thm RS conjunct2)
| _ => [thm];
@@ -769,29 +768,11 @@
structure Fast_Arith = Fast_Lin_Arith(structure LA_Logic = LA_Logic and LA_Data = LA_Data);
-fun map_inj_thms f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} =
- {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = f inj_thms,
- lessD = lessD, neqE = neqE, simpset = simpset, number_of = number_of};
-
-fun map_lessD f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} =
- {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms,
- lessD = f lessD, neqE = neqE, simpset = simpset, number_of = number_of};
-
-fun map_simpset f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} =
- {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms,
- lessD = lessD, neqE = neqE, simpset = f simpset, number_of = number_of};
-
-fun map_number_of f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} =
- {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms,
- lessD = lessD, neqE = neqE, simpset = simpset, number_of = f number_of};
-
-fun add_inj_thms thms = Fast_Arith.map_data (map_inj_thms (append thms));
-fun add_lessD thm = Fast_Arith.map_data (map_lessD (fn thms => thms @ [thm]));
-fun add_simps thms = Fast_Arith.map_data (map_simpset (fn simpset => simpset addsimps thms));
-fun add_simprocs procs = Fast_Arith.map_data (map_simpset (fn simpset => simpset addsimprocs procs));
-
-fun set_number_of f = Fast_Arith.map_data (map_number_of (K (serial (), f)))
-
+val add_inj_thms = Fast_Arith.add_inj_thms;
+val add_lessD = Fast_Arith.add_lessD;
+val add_simps = Fast_Arith.add_simps;
+val add_simprocs = Fast_Arith.add_simprocs;
+val set_number_of = Fast_Arith.set_number_of;
fun simple_tac ctxt = Fast_Arith.lin_arith_tac ctxt false;
val lin_arith_tac = Fast_Arith.lin_arith_tac;
--- a/src/HOL/Tools/meson.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/meson.ML Fri Aug 27 16:05:46 2010 +0200
@@ -151,8 +151,8 @@
let fun has (Const(a,_)) = false
| has (Const(@{const_name Trueprop},_) $ p) = has p
| has (Const(@{const_name Not},_) $ p) = has p
- | has (Const(@{const_name "op |"},_) $ p $ q) = member (op =) bs @{const_name "op |"} orelse has p orelse has q
- | has (Const(@{const_name "op &"},_) $ p $ q) = member (op =) bs @{const_name "op &"} orelse has p orelse has q
+ | has (Const(@{const_name HOL.disj},_) $ p $ q) = member (op =) bs @{const_name HOL.disj} orelse has p orelse has q
+ | has (Const(@{const_name HOL.conj},_) $ p $ q) = member (op =) bs @{const_name HOL.conj} orelse has p orelse has q
| has (Const(@{const_name All},_) $ Abs(_,_,p)) = member (op =) bs @{const_name All} orelse has p
| has (Const(@{const_name Ex},_) $ Abs(_,_,p)) = member (op =) bs @{const_name Ex} orelse has p
| has _ = false
@@ -162,7 +162,7 @@
(**** Clause handling ****)
fun literals (Const(@{const_name Trueprop},_) $ P) = literals P
- | literals (Const(@{const_name "op |"},_) $ P $ Q) = literals P @ literals Q
+ | literals (Const(@{const_name HOL.disj},_) $ P $ Q) = literals P @ literals Q
| literals (Const(@{const_name Not},_) $ P) = [(false,P)]
| literals P = [(true,P)];
@@ -172,7 +172,7 @@
(*** Tautology Checking ***)
-fun signed_lits_aux (Const (@{const_name "op |"}, _) $ P $ Q) (poslits, neglits) =
+fun signed_lits_aux (Const (@{const_name HOL.disj}, _) $ P $ Q) (poslits, neglits) =
signed_lits_aux Q (signed_lits_aux P (poslits, neglits))
| signed_lits_aux (Const(@{const_name Not},_) $ P) (poslits, neglits) = (poslits, P::neglits)
| signed_lits_aux P (poslits, neglits) = (P::poslits, neglits);
@@ -208,16 +208,16 @@
fun refl_clause_aux 0 th = th
| refl_clause_aux n th =
case HOLogic.dest_Trueprop (concl_of th) of
- (Const (@{const_name "op |"}, _) $ (Const (@{const_name "op |"}, _) $ _ $ _) $ _) =>
+ (Const (@{const_name HOL.disj}, _) $ (Const (@{const_name HOL.disj}, _) $ _ $ _) $ _) =>
refl_clause_aux n (th RS disj_assoc) (*isolate an atom as first disjunct*)
- | (Const (@{const_name "op |"}, _) $ (Const(@{const_name Not},_) $ (Const(@{const_name "op ="},_) $ t $ u)) $ _) =>
+ | (Const (@{const_name HOL.disj}, _) $ (Const(@{const_name Not},_) $ (Const(@{const_name "op ="},_) $ t $ u)) $ _) =>
if eliminable(t,u)
then refl_clause_aux (n-1) (th RS not_refl_disj_D) (*Var inequation: delete*)
else refl_clause_aux (n-1) (th RS disj_comm) (*not between Vars: ignore*)
- | (Const (@{const_name "op |"}, _) $ _ $ _) => refl_clause_aux n (th RS disj_comm)
+ | (Const (@{const_name HOL.disj}, _) $ _ $ _) => refl_clause_aux n (th RS disj_comm)
| _ => (*not a disjunction*) th;
-fun notequal_lits_count (Const (@{const_name "op |"}, _) $ P $ Q) =
+fun notequal_lits_count (Const (@{const_name HOL.disj}, _) $ P $ Q) =
notequal_lits_count P + notequal_lits_count Q
| notequal_lits_count (Const(@{const_name Not},_) $ (Const(@{const_name "op ="},_) $ _ $ _)) = 1
| notequal_lits_count _ = 0;
@@ -268,13 +268,13 @@
(*Estimate the number of clauses in order to detect infeasible theorems*)
fun signed_nclauses b (Const(@{const_name Trueprop},_) $ t) = signed_nclauses b t
| signed_nclauses b (Const(@{const_name Not},_) $ t) = signed_nclauses (not b) t
- | signed_nclauses b (Const(@{const_name "op &"},_) $ t $ u) =
+ | signed_nclauses b (Const(@{const_name HOL.conj},_) $ t $ u) =
if b then sum (signed_nclauses b t) (signed_nclauses b u)
else prod (signed_nclauses b t) (signed_nclauses b u)
- | signed_nclauses b (Const(@{const_name "op |"},_) $ t $ u) =
+ | signed_nclauses b (Const(@{const_name HOL.disj},_) $ t $ u) =
if b then prod (signed_nclauses b t) (signed_nclauses b u)
else sum (signed_nclauses b t) (signed_nclauses b u)
- | signed_nclauses b (Const(@{const_name "op -->"},_) $ t $ u) =
+ | signed_nclauses b (Const(@{const_name HOL.implies},_) $ t $ u) =
if b then prod (signed_nclauses (not b) t) (signed_nclauses b u)
else sum (signed_nclauses (not b) t) (signed_nclauses b u)
| signed_nclauses b (Const(@{const_name "op ="}, Type ("fun", [T, _])) $ t $ u) =
@@ -330,10 +330,10 @@
let val ctxtr = Unsynchronized.ref ctxt (* FIXME ??? *)
fun cnf_aux (th,ths) =
if not (can HOLogic.dest_Trueprop (prop_of th)) then ths (*meta-level: ignore*)
- else if not (has_conns [@{const_name All}, @{const_name Ex}, @{const_name "op &"}] (prop_of th))
+ else if not (has_conns [@{const_name All}, @{const_name Ex}, @{const_name HOL.conj}] (prop_of th))
then nodups ctxt th :: ths (*no work to do, terminate*)
else case head_of (HOLogic.dest_Trueprop (concl_of th)) of
- Const (@{const_name "op &"}, _) => (*conjunction*)
+ Const (@{const_name HOL.conj}, _) => (*conjunction*)
cnf_aux (th RS conjunct1, cnf_aux (th RS conjunct2, ths))
| Const (@{const_name All}, _) => (*universal quantifier*)
let val (th',ctxt') = freeze_spec th (!ctxtr)
@@ -341,7 +341,7 @@
| Const (@{const_name Ex}, _) =>
(*existential quantifier: Insert Skolem functions*)
cnf_aux (apply_skolem_theorem (th, skolem_ths), ths)
- | Const (@{const_name "op |"}, _) =>
+ | Const (@{const_name HOL.disj}, _) =>
(*Disjunction of P, Q: Create new goal of proving ?P | ?Q and solve it using
all combinations of converting P, Q to CNF.*)
let val tac =
@@ -365,7 +365,7 @@
(**** Generation of contrapositives ****)
fun is_left (Const (@{const_name Trueprop}, _) $
- (Const (@{const_name "op |"}, _) $ (Const (@{const_name "op |"}, _) $ _ $ _) $ _)) = true
+ (Const (@{const_name HOL.disj}, _) $ (Const (@{const_name HOL.disj}, _) $ _ $ _) $ _)) = true
| is_left _ = false;
(*Associate disjuctions to right -- make leftmost disjunct a LITERAL*)
@@ -400,8 +400,8 @@
(*Is the string the name of a connective? Really only | and Not can remain,
since this code expects to be called on a clause form.*)
val is_conn = member (op =)
- [@{const_name Trueprop}, @{const_name "op &"}, @{const_name "op |"},
- @{const_name "op -->"}, @{const_name Not},
+ [@{const_name Trueprop}, @{const_name HOL.conj}, @{const_name HOL.disj},
+ @{const_name HOL.implies}, @{const_name Not},
@{const_name All}, @{const_name Ex}, @{const_name Ball}, @{const_name Bex}];
(*True if the term contains a function--not a logical connective--where the type
@@ -429,7 +429,7 @@
fun rigid t = not (is_Var (head_of t));
-fun ok4horn (Const (@{const_name Trueprop},_) $ (Const (@{const_name "op |"}, _) $ t $ _)) = rigid t
+fun ok4horn (Const (@{const_name Trueprop},_) $ (Const (@{const_name HOL.disj}, _) $ t $ _)) = rigid t
| ok4horn (Const (@{const_name Trueprop},_) $ t) = rigid t
| ok4horn _ = false;
--- a/src/HOL/Tools/prop_logic.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/prop_logic.ML Fri Aug 27 16:05:46 2010 +0200
@@ -397,14 +397,14 @@
(False, table)
| aux (Const (@{const_name Not}, _) $ x) table =
apfst Not (aux x table)
- | aux (Const (@{const_name "op |"}, _) $ x $ y) table =
+ | aux (Const (@{const_name HOL.disj}, _) $ x $ y) table =
let
val (fm1, table1) = aux x table
val (fm2, table2) = aux y table1
in
(Or (fm1, fm2), table2)
end
- | aux (Const (@{const_name "op &"}, _) $ x $ y) table =
+ | aux (Const (@{const_name HOL.conj}, _) $ x $ y) table =
let
val (fm1, table1) = aux x table
val (fm2, table2) = aux y table1
--- a/src/HOL/Tools/quickcheck_generators.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/quickcheck_generators.ML Fri Aug 27 16:05:46 2010 +0200
@@ -342,7 +342,7 @@
val bound_max = length Ts - 1;
val bounds = map_index (fn (i, ty) =>
(2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) Ts;
- fun strip_imp (Const(@{const_name "op -->"},_) $ A $ B) = apfst (cons A) (strip_imp B)
+ fun strip_imp (Const(@{const_name HOL.implies},_) $ A $ B) = apfst (cons A) (strip_imp B)
| strip_imp A = ([], A)
val prop' = betapplys (prop, map (fn (i, _, _, _) => Bound i) bounds);
val terms = HOLogic.mk_list @{typ term} (map (fn (_, i, _, _) => Bound i $ @{term "()"}) bounds)
--- a/src/HOL/Tools/record.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/record.ML Fri Aug 27 16:05:46 2010 +0200
@@ -420,7 +420,7 @@
equalities = equalities, extinjects=extinjects, extsplit = extsplit, splits = splits,
extfields = extfields, fieldext = fieldext }: data;
-structure Records_Data = Theory_Data
+structure Data = Theory_Data
(
type T = data;
val empty =
@@ -474,25 +474,22 @@
(* access 'records' *)
-val get_info = Symtab.lookup o #records o Records_Data.get;
+val get_info = Symtab.lookup o #records o Data.get;
fun the_info thy name =
(case get_info thy name of
SOME info => info
| NONE => error ("Unknown record type " ^ quote name));
-fun put_record name info thy =
- let
- val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
- Records_Data.get thy;
- val data = make_data (Symtab.update (name, info) records)
- sel_upd equalities extinjects extsplit splits extfields fieldext;
- in Records_Data.put data thy end;
+fun put_record name info =
+ Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
+ make_data (Symtab.update (name, info) records)
+ sel_upd equalities extinjects extsplit splits extfields fieldext);
(* access 'sel_upd' *)
-val get_sel_upd = #sel_upd o Records_Data.get;
+val get_sel_upd = #sel_upd o Data.get;
val is_selector = Symtab.defined o #selectors o get_sel_upd;
val get_updates = Symtab.lookup o #updates o get_sel_upd;
@@ -517,7 +514,7 @@
val upds = map (suffix updateN) all ~~ all;
val {records, sel_upd = {selectors, updates, simpset, defset, foldcong, unfoldcong},
- equalities, extinjects, extsplit, splits, extfields, fieldext} = Records_Data.get thy;
+ equalities, extinjects, extsplit, splits, extfields, fieldext} = Data.get thy;
val data = make_data records
{selectors = fold Symtab.update_new sels selectors,
updates = fold Symtab.update_new upds updates,
@@ -526,80 +523,61 @@
foldcong = foldcong addcongs folds,
unfoldcong = unfoldcong addcongs unfolds}
equalities extinjects extsplit splits extfields fieldext;
- in Records_Data.put data thy end;
+ in Data.put data thy end;
(* access 'equalities' *)
-fun add_equalities name thm thy =
- let
- val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
- Records_Data.get thy;
- val data = make_data records sel_upd
- (Symtab.update_new (name, thm) equalities) extinjects extsplit splits extfields fieldext;
- in Records_Data.put data thy end;
-
-val get_equalities = Symtab.lookup o #equalities o Records_Data.get;
+fun add_equalities name thm =
+ Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
+ make_data records sel_upd
+ (Symtab.update_new (name, thm) equalities) extinjects extsplit splits extfields fieldext);
+
+val get_equalities = Symtab.lookup o #equalities o Data.get;
(* access 'extinjects' *)
-fun add_extinjects thm thy =
- let
- val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
- Records_Data.get thy;
- val data =
- make_data records sel_upd equalities (insert Thm.eq_thm_prop thm extinjects)
- extsplit splits extfields fieldext;
- in Records_Data.put data thy end;
-
-val get_extinjects = rev o #extinjects o Records_Data.get;
+fun add_extinjects thm =
+ Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
+ make_data records sel_upd equalities (insert Thm.eq_thm_prop thm extinjects)
+ extsplit splits extfields fieldext);
+
+val get_extinjects = rev o #extinjects o Data.get;
(* access 'extsplit' *)
-fun add_extsplit name thm thy =
- let
- val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
- Records_Data.get thy;
- val data =
- make_data records sel_upd equalities extinjects
- (Symtab.update_new (name, thm) extsplit) splits extfields fieldext;
- in Records_Data.put data thy end;
+fun add_extsplit name thm =
+ Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
+ make_data records sel_upd equalities extinjects
+ (Symtab.update_new (name, thm) extsplit) splits extfields fieldext);
(* access 'splits' *)
-fun add_splits name thmP thy =
- let
- val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
- Records_Data.get thy;
- val data =
- make_data records sel_upd equalities extinjects extsplit
- (Symtab.update_new (name, thmP) splits) extfields fieldext;
- in Records_Data.put data thy end;
-
-val get_splits = Symtab.lookup o #splits o Records_Data.get;
+fun add_splits name thmP =
+ Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
+ make_data records sel_upd equalities extinjects extsplit
+ (Symtab.update_new (name, thmP) splits) extfields fieldext);
+
+val get_splits = Symtab.lookup o #splits o Data.get;
(* parent/extension of named record *)
-val get_parent = (Option.join o Option.map #parent) oo (Symtab.lookup o #records o Records_Data.get);
-val get_extension = Option.map #extension oo (Symtab.lookup o #records o Records_Data.get);
+val get_parent = (Option.join o Option.map #parent) oo (Symtab.lookup o #records o Data.get);
+val get_extension = Option.map #extension oo (Symtab.lookup o #records o Data.get);
(* access 'extfields' *)
-fun add_extfields name fields thy =
- let
- val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
- Records_Data.get thy;
- val data =
- make_data records sel_upd equalities extinjects extsplit splits
- (Symtab.update_new (name, fields) extfields) fieldext;
- in Records_Data.put data thy end;
-
-val get_extfields = Symtab.lookup o #extfields o Records_Data.get;
+fun add_extfields name fields =
+ Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
+ make_data records sel_upd equalities extinjects extsplit splits
+ (Symtab.update_new (name, fields) extfields) fieldext);
+
+val get_extfields = Symtab.lookup o #extfields o Data.get;
fun get_extT_fields thy T =
let
@@ -609,7 +587,7 @@
in Long_Name.implode (rev (nm :: rst)) end;
val midx = maxidx_of_typs (moreT :: Ts);
val varifyT = varifyT midx;
- val {records, extfields, ...} = Records_Data.get thy;
+ val {records, extfields, ...} = Data.get thy;
val (fields, (more, _)) = split_last (Symtab.lookup_list extfields name);
val args = map varifyT (snd (#extension (the (Symtab.lookup records recname))));
@@ -628,18 +606,14 @@
(* access 'fieldext' *)
-fun add_fieldext extname_types fields thy =
- let
- val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
- Records_Data.get thy;
- val fieldext' =
- fold (fn field => Symtab.update_new (field, extname_types)) fields fieldext;
- val data =
- make_data records sel_upd equalities extinjects
- extsplit splits extfields fieldext';
- in Records_Data.put data thy end;
-
-val get_fieldext = Symtab.lookup o #fieldext o Records_Data.get;
+fun add_fieldext extname_types fields =
+ Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =>
+ let
+ val fieldext' =
+ fold (fn field => Symtab.update_new (field, extname_types)) fields fieldext;
+ in make_data records sel_upd equalities extinjects extsplit splits extfields fieldext' end);
+
+val get_fieldext = Symtab.lookup o #fieldext o Data.get;
(* parent records *)
@@ -1179,7 +1153,7 @@
((upd as Const (u, Type (_, [_, Type (_, [rT, _])]))) $ k $ r) =>
if is_selector thy s andalso is_some (get_updates thy u) then
let
- val {sel_upd = {updates, ...}, extfields, ...} = Records_Data.get thy;
+ val {sel_upd = {updates, ...}, extfields, ...} = Data.get thy;
fun mk_eq_terms ((upd as Const (u, Type(_, [kT, _]))) $ k $ r) =
(case Symtab.lookup updates u of
@@ -1598,15 +1572,17 @@
[] => (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop
(hd (rev (Logic.strip_assums_hyp (hd (prems_of rule')))))))), true)
| [x] => (head_of x, false));
- val rule'' = cterm_instantiate (map (pairself cert)
- (case rev params of
- [] =>
- (case AList.lookup (op =) (Term.add_frees g []) s of
- NONE => sys_error "try_param_tac: no such variable"
- | SOME T => [(P, if ca then concl else lambda (Free (s, T)) concl), (x, Free (s, T))])
- | (_, T) :: _ =>
- [(P, list_abs (params, if ca then concl else incr_boundvars 1 (Abs (s, T, concl)))),
- (x, list_abs (params, Bound 0))])) rule';
+ val rule'' =
+ cterm_instantiate
+ (map (pairself cert)
+ (case rev params of
+ [] =>
+ (case AList.lookup (op =) (Term.add_frees g []) s of
+ NONE => sys_error "try_param_tac: no such variable"
+ | SOME T => [(P, if ca then concl else lambda (Free (s, T)) concl), (x, Free (s, T))])
+ | (_, T) :: _ =>
+ [(P, list_abs (params, if ca then concl else incr_boundvars 1 (Abs (s, T, concl)))),
+ (x, list_abs (params, Bound 0))])) rule';
in compose_tac (false, rule'', nprems_of rule) i end);
@@ -1635,7 +1611,8 @@
val suff = if i = 0 then ext_typeN else inner_typeN ^ string_of_int i;
val ((_, cons), thy') = thy
|> Iso_Tuple_Support.add_iso_tuple_type
- (Binding.suffix_name suff (Binding.name base_name), alphas_zeta) (fastype_of left, fastype_of right);
+ (Binding.suffix_name suff (Binding.name base_name), alphas_zeta)
+ (fastype_of left, fastype_of right);
in
(cons $ left $ right, (thy', i + 1))
end;
@@ -1778,7 +1755,10 @@
("ext_surjective", surject),
("ext_split", split_meta)]);
- in (((ext_name, ext_type), (ext_tyco, alphas_zeta), extT, induct', inject', surjective', split_meta', ext_def), thm_thy) end;
+ in
+ (((ext_name, ext_type), (ext_tyco, alphas_zeta),
+ extT, induct', inject', surjective', split_meta', ext_def), thm_thy)
+ end;
fun chunks [] [] = []
| chunks [] xs = [xs]
@@ -1795,7 +1775,7 @@
(* mk_recordT *)
-(*builds up the record type from the current extension tpye extT and a list
+(*build up the record type from the current extension tpye extT and a list
of parent extensions, starting with the root of the record hierarchy*)
fun mk_recordT extT =
fold_rev (fn (parent, Ts) => fn T => Type (parent, subst_last T Ts)) extT;
@@ -1834,8 +1814,10 @@
val lhs = HOLogic.mk_random T size;
val tc = HOLogic.mk_return Tm @{typ Random.seed}
(HOLogic.mk_valtermify_app extN params T);
- val rhs = HOLogic.mk_ST
- (map (fn (v, T') => ((HOLogic.mk_random T' size, @{typ Random.seed}), SOME (v, termifyT T'))) params)
+ val rhs =
+ HOLogic.mk_ST
+ (map (fn (v, T') =>
+ ((HOLogic.mk_random T' size, @{typ Random.seed}), SOME (v, termifyT T'))) params)
tc @{typ Random.seed} (SOME Tm, @{typ Random.seed});
val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs));
in
@@ -1860,17 +1842,20 @@
fun add_code ext_tyco vs extT ext simps inject thy =
let
- val eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq)
- (Const (@{const_name eq_class.eq}, extT --> extT --> HOLogic.boolT),
- Const (@{const_name "op ="}, extT --> extT --> HOLogic.boolT));
- fun tac eq_def = Class.intro_classes_tac []
+ val eq =
+ (HOLogic.mk_Trueprop o HOLogic.mk_eq)
+ (Const (@{const_name eq_class.eq}, extT --> extT --> HOLogic.boolT),
+ Const (@{const_name "op ="}, extT --> extT --> HOLogic.boolT));
+ fun tac eq_def =
+ Class.intro_classes_tac []
THEN (Simplifier.rewrite_goals_tac [Simpdata.mk_eq eq_def])
THEN ALLGOALS (rtac @{thm refl});
fun mk_eq thy eq_def = Simplifier.rewrite_rule
[(AxClass.unoverload thy o Thm.symmetric o Simpdata.mk_eq) eq_def] inject;
- fun mk_eq_refl thy = @{thm HOL.eq_refl}
+ fun mk_eq_refl thy =
+ @{thm HOL.eq_refl}
|> Thm.instantiate
- ([pairself (Thm.ctyp_of thy) (TVar (("'a", 0), @{sort eq}), Logic.varifyT_global extT)], [])
+ ([pairself (Thm.ctyp_of thy) (TVar (("'a", 0), @{sort eq}), Logic.varifyT_global extT)], [])
|> AxClass.unoverload thy;
in
thy
@@ -1944,7 +1929,8 @@
val extension_name = full binding;
- val ((ext, (ext_tyco, vs), extT, ext_induct, ext_inject, ext_surjective, ext_split, ext_def), ext_thy) =
+ val ((ext, (ext_tyco, vs),
+ extT, ext_induct, ext_inject, ext_surjective, ext_split, ext_def), ext_thy) =
thy
|> Sign.qualified_path false binding
|> extension_definition extension_name fields alphas_ext zeta moreT more vars;
--- a/src/HOL/Tools/refute.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/refute.ML Fri Aug 27 16:05:46 2010 +0200
@@ -648,9 +648,9 @@
| Const (@{const_name All}, _) => t
| Const (@{const_name Ex}, _) => t
| Const (@{const_name "op ="}, _) => t
- | Const (@{const_name "op &"}, _) => t
- | Const (@{const_name "op |"}, _) => t
- | Const (@{const_name "op -->"}, _) => t
+ | Const (@{const_name HOL.conj}, _) => t
+ | Const (@{const_name HOL.disj}, _) => t
+ | Const (@{const_name HOL.implies}, _) => t
(* sets *)
| Const (@{const_name Collect}, _) => t
| Const (@{const_name Set.member}, _) => t
@@ -824,9 +824,9 @@
| Const (@{const_name All}, T) => collect_type_axioms T axs
| Const (@{const_name Ex}, T) => collect_type_axioms T axs
| Const (@{const_name "op ="}, T) => collect_type_axioms T axs
- | Const (@{const_name "op &"}, _) => axs
- | Const (@{const_name "op |"}, _) => axs
- | Const (@{const_name "op -->"}, _) => axs
+ | Const (@{const_name HOL.conj}, _) => axs
+ | Const (@{const_name HOL.disj}, _) => axs
+ | Const (@{const_name HOL.implies}, _) => axs
(* sets *)
| Const (@{const_name Collect}, T) => collect_type_axioms T axs
| Const (@{const_name Set.member}, T) => collect_type_axioms T axs
@@ -1861,7 +1861,7 @@
SOME (interpret thy model args (eta_expand t 1))
| Const (@{const_name "op ="}, _) =>
SOME (interpret thy model args (eta_expand t 2))
- | Const (@{const_name "op &"}, _) $ t1 $ t2 =>
+ | Const (@{const_name HOL.conj}, _) $ t1 $ t2 =>
(* 3-valued logic *)
let
val (i1, m1, a1) = interpret thy model args t1
@@ -1871,14 +1871,14 @@
in
SOME (Leaf [fmTrue, fmFalse], m2, a2)
end
- | Const (@{const_name "op &"}, _) $ t1 =>
+ | Const (@{const_name HOL.conj}, _) $ t1 =>
SOME (interpret thy model args (eta_expand t 1))
- | Const (@{const_name "op &"}, _) =>
+ | Const (@{const_name HOL.conj}, _) =>
SOME (interpret thy model args (eta_expand t 2))
(* this would make "undef" propagate, even for formulae like *)
(* "False & undef": *)
(* SOME (Node [Node [TT, FF], Node [FF, FF]], model, args) *)
- | Const (@{const_name "op |"}, _) $ t1 $ t2 =>
+ | Const (@{const_name HOL.disj}, _) $ t1 $ t2 =>
(* 3-valued logic *)
let
val (i1, m1, a1) = interpret thy model args t1
@@ -1888,14 +1888,14 @@
in
SOME (Leaf [fmTrue, fmFalse], m2, a2)
end
- | Const (@{const_name "op |"}, _) $ t1 =>
+ | Const (@{const_name HOL.disj}, _) $ t1 =>
SOME (interpret thy model args (eta_expand t 1))
- | Const (@{const_name "op |"}, _) =>
+ | Const (@{const_name HOL.disj}, _) =>
SOME (interpret thy model args (eta_expand t 2))
(* this would make "undef" propagate, even for formulae like *)
(* "True | undef": *)
(* SOME (Node [Node [TT, TT], Node [TT, FF]], model, args) *)
- | Const (@{const_name "op -->"}, _) $ t1 $ t2 => (* similar to "==>" (Pure) *)
+ | Const (@{const_name HOL.implies}, _) $ t1 $ t2 => (* similar to "==>" (Pure) *)
(* 3-valued logic *)
let
val (i1, m1, a1) = interpret thy model args t1
@@ -1905,9 +1905,9 @@
in
SOME (Leaf [fmTrue, fmFalse], m2, a2)
end
- | Const (@{const_name "op -->"}, _) $ t1 =>
+ | Const (@{const_name HOL.implies}, _) $ t1 =>
SOME (interpret thy model args (eta_expand t 1))
- | Const (@{const_name "op -->"}, _) =>
+ | Const (@{const_name HOL.implies}, _) =>
SOME (interpret thy model args (eta_expand t 2))
(* this would make "undef" propagate, even for formulae like *)
(* "False --> undef": *)
--- a/src/HOL/Tools/simpdata.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/simpdata.ML Fri Aug 27 16:05:46 2010 +0200
@@ -12,9 +12,9 @@
(*abstract syntax*)
fun dest_eq ((c as Const(@{const_name "op ="},_)) $ s $ t) = SOME (c, s, t)
| dest_eq _ = NONE;
- fun dest_conj ((c as Const(@{const_name "op &"},_)) $ s $ t) = SOME (c, s, t)
+ fun dest_conj ((c as Const(@{const_name HOL.conj},_)) $ s $ t) = SOME (c, s, t)
| dest_conj _ = NONE;
- fun dest_imp ((c as Const(@{const_name "op -->"},_)) $ s $ t) = SOME (c, s, t)
+ fun dest_imp ((c as Const(@{const_name HOL.implies},_)) $ s $ t) = SOME (c, s, t)
| dest_imp _ = NONE;
val conj = HOLogic.conj
val imp = HOLogic.imp
@@ -159,8 +159,8 @@
(Scan.succeed "Clasimp.clasimpset_of (ML_Context.the_local_context ())");
val mksimps_pairs =
- [(@{const_name "op -->"}, [@{thm mp}]),
- (@{const_name "op &"}, [@{thm conjunct1}, @{thm conjunct2}]),
+ [(@{const_name HOL.implies}, [@{thm mp}]),
+ (@{const_name HOL.conj}, [@{thm conjunct1}, @{thm conjunct2}]),
(@{const_name All}, [@{thm spec}]),
(@{const_name True}, []),
(@{const_name False}, []),
--- a/src/HOL/Tools/typedef.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/Tools/typedef.ML Fri Aug 27 16:05:46 2010 +0200
@@ -186,7 +186,8 @@
||> Thm.term_of;
val ((RepC, AbsC, axiom_name, axiom), axiom_lthy) = set_lthy |>
- Local_Theory.theory_result (primitive_typedef typedef_name newT oldT Rep_name Abs_name A);
+ Local_Theory.background_theory_result
+ (primitive_typedef typedef_name newT oldT Rep_name Abs_name A);
val cert = Thm.cterm_of (ProofContext.theory_of axiom_lthy);
val typedef =
@@ -246,7 +247,7 @@
in
lthy2
|> Local_Theory.declaration true (fn phi => put_info full_tname (transform_info phi info))
- |> Local_Theory.theory (Typedef_Interpretation.data full_tname)
+ |> Local_Theory.background_theory (Typedef_Interpretation.data full_tname)
|> pair (full_tname, info)
end;
--- a/src/HOL/ex/Meson_Test.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/ex/Meson_Test.thy Fri Aug 27 16:05:46 2010 +0200
@@ -10,7 +10,7 @@
below and constants declared in HOL!
*}
-hide_const (open) subset quotient union inter sum
+hide_const (open) implies union inter subset sum quotient
text {*
Test data for the MESON proof procedure
--- a/src/HOL/ex/Numeral.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/ex/Numeral.thy Fri Aug 27 16:05:46 2010 +0200
@@ -97,7 +97,7 @@
structure Dig_Simps = Named_Thms
(
val name = "numeral"
- val description = "Simplification rules for numerals"
+ val description = "simplification rules for numerals"
)
*}
@@ -1033,14 +1033,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"
--- a/src/HOL/ex/SVC_Oracle.thy Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/ex/SVC_Oracle.thy Fri Aug 27 16:05:46 2010 +0200
@@ -88,9 +88,9 @@
else if T = HOLogic.boolT then c $ (fm x) $ (fm y)
else replace (c $ x $ y) (*non-numeric comparison*)
(*abstraction of a formula*)
- and fm ((c as Const(@{const_name "op &"}, _)) $ p $ q) = c $ (fm p) $ (fm q)
- | fm ((c as Const(@{const_name "op |"}, _)) $ p $ q) = c $ (fm p) $ (fm q)
- | fm ((c as Const(@{const_name "op -->"}, _)) $ p $ q) = c $ (fm p) $ (fm q)
+ and fm ((c as Const(@{const_name HOL.conj}, _)) $ p $ q) = c $ (fm p) $ (fm q)
+ | fm ((c as Const(@{const_name HOL.disj}, _)) $ p $ q) = c $ (fm p) $ (fm q)
+ | fm ((c as Const(@{const_name HOL.implies}, _)) $ p $ q) = c $ (fm p) $ (fm q)
| fm ((c as Const(@{const_name Not}, _)) $ p) = c $ (fm p)
| fm ((c as Const(@{const_name True}, _))) = c
| fm ((c as Const(@{const_name False}, _))) = c
--- a/src/HOL/ex/svc_funcs.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOL/ex/svc_funcs.ML Fri Aug 27 16:05:46 2010 +0200
@@ -89,9 +89,9 @@
let fun tag t =
let val (c,ts) = strip_comb t
in case c of
- Const(@{const_name "op &"}, _) => apply c (map tag ts)
- | Const(@{const_name "op |"}, _) => apply c (map tag ts)
- | Const(@{const_name "op -->"}, _) => apply c (map tag ts)
+ Const(@{const_name HOL.conj}, _) => apply c (map tag ts)
+ | Const(@{const_name HOL.disj}, _) => apply c (map tag ts)
+ | Const(@{const_name HOL.implies}, _) => apply c (map tag ts)
| Const(@{const_name Not}, _) => apply c (map tag ts)
| Const(@{const_name True}, _) => (c, false)
| Const(@{const_name False}, _) => (c, false)
@@ -183,11 +183,11 @@
| tm t = Int (lit t)
handle Match => var (t,[])
(*translation of a formula*)
- and fm pos (Const(@{const_name "op &"}, _) $ p $ q) =
+ and fm pos (Const(@{const_name HOL.conj}, _) $ p $ q) =
Buildin("AND", [fm pos p, fm pos q])
- | fm pos (Const(@{const_name "op |"}, _) $ p $ q) =
+ | fm pos (Const(@{const_name HOL.disj}, _) $ p $ q) =
Buildin("OR", [fm pos p, fm pos q])
- | fm pos (Const(@{const_name "op -->"}, _) $ p $ q) =
+ | fm pos (Const(@{const_name HOL.implies}, _) $ p $ q) =
Buildin("=>", [fm (not pos) p, fm pos q])
| fm pos (Const(@{const_name Not}, _) $ p) =
Buildin("NOT", [fm (not pos) p])
--- a/src/HOLCF/Tools/Domain/domain_library.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOLCF/Tools/Domain/domain_library.ML Fri Aug 27 16:05:46 2010 +0200
@@ -121,7 +121,7 @@
val r_inst = read_instantiate ctxt;
fun at thm =
case concl_of thm of
- _$(Const(@{const_name "op &"},_)$_$_) => at(thm RS conjunct1)@at(thm RS conjunct2)
+ _$(Const(@{const_name HOL.conj},_)$_$_) => at(thm RS conjunct1)@at(thm RS conjunct2)
| _$(Const(@{const_name All} ,_)$Abs(s,_,_))=> at(thm RS (r_inst [(("x", 0), "?" ^ s)] spec))
| _ => [thm];
in map zero_var_indexes (at thm) end;
--- a/src/HOLCF/Tools/pcpodef.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/HOLCF/Tools/pcpodef.ML Fri Aug 27 16:05:46 2010 +0200
@@ -326,7 +326,7 @@
val args = map (apsnd (prep_constraint ctxt)) raw_args;
val (goal1, goal2, make_result) =
prepare_cpodef prep_term def name (b, args, mx) set opt_morphs thy;
- fun after_qed [[th1, th2]] = ProofContext.theory (snd o make_result th1 th2)
+ fun after_qed [[th1, th2]] = ProofContext.background_theory (snd o make_result th1 th2)
| after_qed _ = raise Fail "cpodef_proof";
in Proof.theorem NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end;
@@ -337,7 +337,7 @@
val args = map (apsnd (prep_constraint ctxt)) raw_args;
val (goal1, goal2, make_result) =
prepare_pcpodef prep_term def name (b, args, mx) set opt_morphs thy;
- fun after_qed [[th1, th2]] = ProofContext.theory (snd o make_result th1 th2)
+ fun after_qed [[th1, th2]] = ProofContext.background_theory (snd o make_result th1 th2)
| after_qed _ = raise Fail "pcpodef_proof";
in Proof.theorem NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end;
--- a/src/Provers/Arith/fast_lin_arith.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Provers/Arith/fast_lin_arith.ML Fri Aug 27 16:05:46 2010 +0200
@@ -88,13 +88,19 @@
val cut_lin_arith_tac: simpset -> int -> tactic
val lin_arith_tac: Proof.context -> bool -> int -> tactic
val lin_arith_simproc: simpset -> term -> thm option
- val map_data: ({add_mono_thms: thm list, mult_mono_thms: thm list, inj_thms: thm list,
- lessD: thm list, neqE: thm list, simpset: Simplifier.simpset,
- number_of : serial * (theory -> typ -> int -> cterm)}
- -> {add_mono_thms: thm list, mult_mono_thms: thm list, inj_thms: thm list,
- lessD: thm list, neqE: thm list, simpset: Simplifier.simpset,
- number_of : serial * (theory -> typ -> int -> cterm)})
- -> Context.generic -> Context.generic
+ val map_data:
+ ({add_mono_thms: thm list, mult_mono_thms: thm list, inj_thms: thm list,
+ lessD: thm list, neqE: thm list, simpset: Simplifier.simpset,
+ number_of: (theory -> typ -> int -> cterm) option} ->
+ {add_mono_thms: thm list, mult_mono_thms: thm list, inj_thms: thm list,
+ lessD: thm list, neqE: thm list, simpset: Simplifier.simpset,
+ number_of: (theory -> typ -> int -> cterm) option}) ->
+ Context.generic -> Context.generic
+ val add_inj_thms: thm list -> Context.generic -> Context.generic
+ val add_lessD: thm -> Context.generic -> Context.generic
+ val add_simps: thm list -> Context.generic -> Context.generic
+ val add_simprocs: simproc list -> Context.generic -> Context.generic
+ val set_number_of: (theory -> typ -> int -> cterm) -> Context.generic -> Context.generic
val trace: bool Unsynchronized.ref
end;
@@ -105,8 +111,6 @@
(** theory data **)
-fun no_number_of _ _ _ = raise CTERM ("number_of", [])
-
structure Data = Generic_Data
(
type T =
@@ -116,32 +120,57 @@
lessD: thm list,
neqE: thm list,
simpset: Simplifier.simpset,
- number_of : serial * (theory -> typ -> int -> cterm)};
+ number_of: (theory -> typ -> int -> cterm) option};
- val empty = {add_mono_thms = [], mult_mono_thms = [], inj_thms = [],
- lessD = [], neqE = [], simpset = Simplifier.empty_ss,
- number_of = (serial (), no_number_of) } : T;
+ val empty : T =
+ {add_mono_thms = [], mult_mono_thms = [], inj_thms = [],
+ lessD = [], neqE = [], simpset = Simplifier.empty_ss,
+ number_of = NONE};
val extend = I;
fun merge
- ({add_mono_thms= add_mono_thms1, mult_mono_thms= mult_mono_thms1, inj_thms= inj_thms1,
- lessD = lessD1, neqE=neqE1, simpset = simpset1,
- number_of = (number_of1 as (s1, _))},
- {add_mono_thms= add_mono_thms2, mult_mono_thms= mult_mono_thms2, inj_thms= inj_thms2,
- lessD = lessD2, neqE=neqE2, simpset = simpset2,
- number_of = (number_of2 as (s2, _))}) =
+ ({add_mono_thms = add_mono_thms1, mult_mono_thms = mult_mono_thms1, inj_thms = inj_thms1,
+ lessD = lessD1, neqE = neqE1, simpset = simpset1, number_of = number_of1},
+ {add_mono_thms = add_mono_thms2, mult_mono_thms = mult_mono_thms2, inj_thms = inj_thms2,
+ lessD = lessD2, neqE = neqE2, simpset = simpset2, number_of = number_of2}) : T =
{add_mono_thms = Thm.merge_thms (add_mono_thms1, add_mono_thms2),
mult_mono_thms = Thm.merge_thms (mult_mono_thms1, mult_mono_thms2),
inj_thms = Thm.merge_thms (inj_thms1, inj_thms2),
lessD = Thm.merge_thms (lessD1, lessD2),
neqE = Thm.merge_thms (neqE1, neqE2),
simpset = Simplifier.merge_ss (simpset1, simpset2),
- (* FIXME depends on accidental load order !?! *)
- number_of = if s1 > s2 then number_of1 else number_of2};
+ number_of = if is_some number_of1 then number_of1 else number_of2};
);
val map_data = Data.map;
val get_data = Data.get o Context.Proof;
+fun map_inj_thms f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} =
+ {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = f inj_thms,
+ lessD = lessD, neqE = neqE, simpset = simpset, number_of = number_of};
+
+fun map_lessD f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} =
+ {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms,
+ lessD = f lessD, neqE = neqE, simpset = simpset, number_of = number_of};
+
+fun map_simpset f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} =
+ {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms,
+ lessD = lessD, neqE = neqE, simpset = f simpset, number_of = number_of};
+
+fun add_inj_thms thms = map_data (map_inj_thms (append thms));
+fun add_lessD thm = map_data (map_lessD (fn thms => thms @ [thm]));
+fun add_simps thms = map_data (map_simpset (fn simpset => simpset addsimps thms));
+fun add_simprocs procs = map_data (map_simpset (fn simpset => simpset addsimprocs procs));
+
+fun set_number_of f =
+ map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, ...} =>
+ {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms,
+ lessD = lessD, neqE = neqE, simpset = simpset, number_of = SOME f});
+
+fun number_of ctxt =
+ (case Data.get (Context.Proof ctxt) of
+ {number_of = SOME f, ...} => f (ProofContext.theory_of ctxt)
+ | _ => fn _ => fn _ => raise CTERM ("number_of", []));
+
(*** A fast decision procedure ***)
@@ -446,8 +475,8 @@
let
val ctxt = Simplifier.the_context ss;
val thy = ProofContext.theory_of ctxt;
- val {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset,
- number_of = (_, num_of), ...} = get_data ctxt;
+ val {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset, ...} = get_data ctxt;
+ val number_of = number_of ctxt;
val simpset' = Simplifier.inherit_context ss simpset;
fun only_concl f thm =
if Thm.no_prems thm then f (Thm.concl_of thm) else NONE;
@@ -493,7 +522,7 @@
val T = #T (Thm.rep_cterm cv)
in
mth
- |> Thm.instantiate ([], [(cv, num_of thy T n)])
+ |> Thm.instantiate ([], [(cv, number_of T n)])
|> rewrite_concl
|> discharge_prem
handle CTERM _ => mult_by_add n thm
--- a/src/Pure/General/markup.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/General/markup.ML Fri Aug 27 16:05:46 2010 +0200
@@ -99,6 +99,8 @@
val command_spanN: string val command_span: string -> T
val ignored_spanN: string val ignored_span: T
val malformed_spanN: string val malformed_span: T
+ val subgoalsN: string
+ val proof_stateN: string val proof_state: int -> T
val stateN: string val state: T
val subgoalN: string val subgoal: T
val sendbackN: string val sendback: T
@@ -156,16 +158,13 @@
fun markup_int elem prop = (elem, fn i => (elem, [(prop, print_int i)]): T);
-(* name *)
+(* misc properties *)
val nameN = "name";
fun name a = properties [(nameN, a)];
val (bindingN, binding) = markup_string "binding" nameN;
-
-(* kind *)
-
val kindN = "kind";
@@ -305,6 +304,9 @@
(* toplevel *)
+val subgoalsN = "subgoals";
+val (proof_stateN, proof_state) = markup_int "proof_state" subgoalsN;
+
val (stateN, state) = markup_elem "state";
val (subgoalN, subgoal) = markup_elem "subgoal";
val (sendbackN, sendback) = markup_elem "sendback";
--- a/src/Pure/General/markup.scala Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/General/markup.scala Fri Aug 27 16:05:46 2010 +0200
@@ -9,7 +9,7 @@
object Markup
{
- /* integers */
+ /* plain values */
object Int {
def apply(i: scala.Int): String = i.toString
@@ -26,25 +26,33 @@
}
- /* property values */
-
- def get_string(name: String, props: List[(String, String)]): Option[String] =
- props.find(p => p._1 == name).map(_._2)
+ /* named properties */
- def get_long(name: String, props: List[(String, String)]): Option[scala.Long] =
+ class Property(val name: String)
{
- get_string(name, props) match {
- case None => None
- case Some(Long(i)) => Some(i)
- }
+ def apply(value: String): List[(String, String)] = List((name, value))
+ def unapply(props: List[(String, String)]): Option[String] =
+ props.find(_._1 == name).map(_._2)
}
- def get_int(name: String, props: List[(String, String)]): Option[scala.Int] =
+ class Int_Property(name: String)
{
- get_string(name, props) match {
- case None => None
- case Some(Int(i)) => Some(i)
- }
+ def apply(value: scala.Int): List[(String, String)] = List((name, Int(value)))
+ def unapply(props: List[(String, String)]): Option[Int] =
+ props.find(_._1 == name) match {
+ case None => None
+ case Some((_, value)) => Int.unapply(value)
+ }
+ }
+
+ class Long_Property(name: String)
+ {
+ def apply(value: scala.Long): List[(String, String)] = List((name, Long(value)))
+ def unapply(props: List[(String, String)]): Option[Long] =
+ props.find(_._1 == name) match {
+ case None => None
+ case Some((_, value)) => Long.unapply(value)
+ }
}
@@ -53,7 +61,7 @@
val Empty = Markup("", Nil)
- /* name */
+ /* misc properties */
val NAME = "name"
val KIND = "kind"
@@ -86,9 +94,9 @@
/* pretty printing */
- val INDENT = "indent"
+ val Indent = new Int_Property("indent")
val BLOCK = "block"
- val WIDTH = "width"
+ val Width = new Int_Property("width")
val BREAK = "break"
@@ -188,6 +196,9 @@
/* toplevel */
+ val SUBGOALS = "subgoals"
+ val PROOF_STATE = "proof_state"
+
val STATE = "state"
val SUBGOAL = "subgoal"
val SENDBACK = "sendback"
--- a/src/Pure/General/position.scala Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/General/position.scala Fri Aug 27 16:05:46 2010 +0200
@@ -11,22 +11,21 @@
{
type T = List[(String, String)]
- def get_line(pos: T): Option[Int] = Markup.get_int(Markup.LINE, pos)
- def get_column(pos: T): Option[Int] = Markup.get_int(Markup.COLUMN, pos)
- def get_offset(pos: T): Option[Int] = Markup.get_int(Markup.OFFSET, pos)
- def get_end_line(pos: T): Option[Int] = Markup.get_int(Markup.END_LINE, pos)
- def get_end_column(pos: T): Option[Int] = Markup.get_int(Markup.END_COLUMN, pos)
- def get_end_offset(pos: T): Option[Int] = Markup.get_int(Markup.END_OFFSET, pos)
- def get_file(pos: T): Option[String] = Markup.get_string(Markup.FILE, pos)
- def get_id(pos: T): Option[Long] = Markup.get_long(Markup.ID, pos)
+ val Line = new Markup.Int_Property(Markup.LINE)
+ val End_Line = new Markup.Int_Property(Markup.END_LINE)
+ val Offset = new Markup.Int_Property(Markup.OFFSET)
+ val End_Offset = new Markup.Int_Property(Markup.END_OFFSET)
+ val File = new Markup.Property(Markup.FILE)
+ val Id = new Markup.Long_Property(Markup.ID)
- def get_range(pos: T): Option[Text.Range] =
- (get_offset(pos), get_end_offset(pos)) match {
- case (Some(start), Some(stop)) if start <= stop => Some(Text.Range(start, stop))
- case (Some(start), None) => Some(Text.Range(start))
- case (_, _) => None
- }
-
- object Id { def unapply(pos: T): Option[Long] = get_id(pos) }
- object Range { def unapply(pos: T): Option[Text.Range] = get_range(pos) }
+ object Range
+ {
+ def apply(range: Text.Range): T = Offset(range.start) ++ Offset(range.stop)
+ def unapply(pos: T): Option[Text.Range] =
+ (Offset.unapply(pos), End_Offset.unapply(pos)) match {
+ case (Some(start), Some(stop)) if start <= stop => Some(Text.Range(start, stop))
+ case (Some(start), None) => Some(Text.Range(start))
+ case _ => None
+ }
+ }
}
--- a/src/Pure/General/pretty.scala Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/General/pretty.scala Fri Aug 27 16:05:46 2010 +0200
@@ -19,12 +19,11 @@
object Block
{
def apply(i: Int, body: XML.Body): XML.Tree =
- XML.Elem(Markup(Markup.BLOCK, List((Markup.INDENT, Markup.Int(i)))), body)
+ XML.Elem(Markup(Markup.BLOCK, Markup.Indent(i)), body)
def unapply(tree: XML.Tree): Option[(Int, XML.Body)] =
tree match {
- case XML.Elem(
- Markup(Markup.BLOCK, List((Markup.INDENT, Markup.Int(i)))), body) => Some((i, body))
+ case XML.Elem(Markup(Markup.BLOCK, Markup.Indent(i)), body) => Some((i, body))
case _ => None
}
}
@@ -32,12 +31,11 @@
object Break
{
def apply(w: Int): XML.Tree =
- XML.Elem(
- Markup(Markup.BREAK, List((Markup.WIDTH, Markup.Int(w)))), List(XML.Text(Symbol.spaces(w))))
+ XML.Elem(Markup(Markup.BREAK, Markup.Width(w)), List(XML.Text(Symbol.spaces(w))))
def unapply(tree: XML.Tree): Option[Int] =
tree match {
- case XML.Elem(Markup(Markup.BREAK, List((Markup.WIDTH, Markup.Int(w)))), _) => Some(w)
+ case XML.Elem(Markup(Markup.BREAK, Markup.Width(w)), _) => Some(w)
case _ => None
}
}
--- a/src/Pure/Isar/class.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/Isar/class.ML Fri Aug 27 16:05:46 2010 +0200
@@ -368,7 +368,7 @@
fun gen_classrel mk_prop classrel thy =
let
fun after_qed results =
- ProofContext.theory ((fold o fold) AxClass.add_classrel results);
+ ProofContext.background_theory ((fold o fold) AxClass.add_classrel results);
in
thy
|> ProofContext.init_global
@@ -482,7 +482,7 @@
val type_name = sanitize_name o Long_Name.base_name;
-fun define_overloaded (c, U) v (b_def, rhs) = Local_Theory.theory_result
+fun define_overloaded (c, U) v (b_def, rhs) = Local_Theory.background_theory_result
(AxClass.declare_overloaded (c, U) ##>> AxClass.define_overloaded b_def (c, rhs))
##> (map_instantiation o apsnd) (filter_out (fn (_, (v', _)) => v' = v))
##> Local_Theory.target synchronize_inst_syntax;
@@ -572,7 +572,7 @@
val (tycos, vs, sort) = (#arities o the_instantiation) lthy;
val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos;
fun after_qed' results =
- Local_Theory.theory (fold (AxClass.add_arity o Thm.varifyT_global) results)
+ Local_Theory.background_theory (fold (AxClass.add_arity o Thm.varifyT_global) results)
#> after_qed;
in
lthy
@@ -608,7 +608,7 @@
val (tycos, vs, sort) = read_multi_arity thy raw_arities;
val sorts = map snd vs;
val arities = maps (fn tyco => Logic.mk_arities (tyco, sorts, sort)) tycos;
- fun after_qed results = ProofContext.theory
+ fun after_qed results = ProofContext.background_theory
((fold o fold) AxClass.add_arity results);
in
thy
--- a/src/Pure/Isar/class_declaration.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/Isar/class_declaration.ML Fri Aug 27 16:05:46 2010 +0200
@@ -330,7 +330,7 @@
val some_prop = try the_single props;
val some_dep_morph = try the_single (map snd deps);
fun after_qed some_wit =
- ProofContext.theory (Class.register_subclass (sub, sup)
+ ProofContext.background_theory (Class.register_subclass (sub, sup)
some_dep_morph some_wit export)
#> ProofContext.theory_of #> Named_Target.init sub;
in do_proof after_qed some_prop goal_ctxt end;
--- a/src/Pure/Isar/expression.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/Isar/expression.ML Fri Aug 27 16:05:46 2010 +0200
@@ -824,8 +824,9 @@
val goal_ctxt = fold Variable.auto_fixes eqns expr_ctxt;
val export' = Variable.export_morphism goal_ctxt expr_ctxt;
- fun after_qed witss eqns = (ProofContext.theory o Context.theory_map)
- (note_eqns_register deps witss attrss eqns export export');
+ fun after_qed witss eqns =
+ (ProofContext.background_theory o Context.theory_map)
+ (note_eqns_register deps witss attrss eqns export export');
in Element.witness_proof_eqs after_qed propss eqns goal_ctxt end;
@@ -872,7 +873,7 @@
val target = intern thy raw_target;
val target_ctxt = Named_Target.init target thy;
val ((propss, deps, export), goal_ctxt) = prep_expr expression target_ctxt;
- fun after_qed witss = ProofContext.theory
+ fun after_qed witss = ProofContext.background_theory
(fold (fn ((dep, morph), wits) => Locale.add_dependency
target (dep, morph $> Element.satisfy_morphism wits) export) (deps ~~ witss));
in Element.witness_proof after_qed propss goal_ctxt end;
--- a/src/Pure/Isar/generic_target.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/Isar/generic_target.ML Fri Aug 27 16:05:46 2010 +0200
@@ -195,16 +195,16 @@
(Morphism.transform (Local_Theory.global_morphism lthy) decl);
in
lthy
- |> Local_Theory.theory (Context.theory_map global_decl)
+ |> Local_Theory.background_theory (Context.theory_map global_decl)
|> Local_Theory.target (Context.proof_map global_decl)
end;
fun theory_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) lthy =
let
- val (const, lthy2) = lthy |> Local_Theory.theory_result
+ val (const, lthy2) = lthy |> Local_Theory.background_theory_result
(Sign.declare_const ((b, U), mx));
val lhs = list_comb (const, type_params @ term_params);
- val ((_, def), lthy3) = lthy2 |> Local_Theory.theory_result
+ val ((_, def), lthy3) = lthy2 |> Local_Theory.background_theory_result
(Thm.add_def false false (b_def, Logic.mk_equals (lhs, rhs)));
in ((lhs, def), lthy3) end;
@@ -214,12 +214,13 @@
val global_facts' = Attrib.map_facts (Attrib.attribute_i thy) global_facts;
in
lthy
- |> Local_Theory.theory (PureThy.note_thmss kind global_facts' #> snd)
+ |> Local_Theory.background_theory (PureThy.note_thmss kind global_facts' #> snd)
|> Local_Theory.target (ProofContext.note_thmss kind global_facts' #> snd)
end;
-fun theory_abbrev prmode ((b, mx), t) = Local_Theory.theory
- (Sign.add_abbrev (#1 prmode) (b, t) #->
- (fn (lhs, _) => Sign.notation true prmode [(lhs, mx)]));
+fun theory_abbrev prmode ((b, mx), t) =
+ Local_Theory.background_theory
+ (Sign.add_abbrev (#1 prmode) (b, t) #->
+ (fn (lhs, _) => Sign.notation true prmode [(lhs, mx)]));
end;
--- a/src/Pure/Isar/local_theory.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/Isar/local_theory.ML Fri Aug 27 16:05:46 2010 +0200
@@ -21,8 +21,8 @@
val target_of: local_theory -> Proof.context
val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
val raw_theory: (theory -> theory) -> local_theory -> local_theory
- val theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
- val theory: (theory -> theory) -> local_theory -> local_theory
+ val background_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
+ val background_theory: (theory -> theory) -> local_theory -> local_theory
val target_result: (Proof.context -> 'a * Proof.context) -> local_theory -> 'a * local_theory
val target: (Proof.context -> Proof.context) -> local_theory -> local_theory
val map_contexts: (Context.generic -> Context.generic) -> local_theory -> local_theory
@@ -144,7 +144,7 @@
val checkpoint = raw_theory Theory.checkpoint;
-fun theory_result f lthy =
+fun background_theory_result f lthy =
lthy |> raw_theory_result (fn thy =>
thy
|> Sign.map_naming (K (naming_of lthy))
@@ -152,7 +152,7 @@
||> Sign.restore_naming thy
||> Theory.checkpoint);
-fun theory f = #2 o theory_result (f #> pair ());
+fun background_theory f = #2 o background_theory_result (f #> pair ());
fun target_result f lthy =
let
@@ -169,7 +169,7 @@
fun target f = #2 o target_result (f #> pair ());
fun map_contexts f =
- theory (Context.theory_map f) #>
+ background_theory (Context.theory_map f) #>
target (Context.proof_map f) #>
Context.proof_map f;
@@ -261,7 +261,7 @@
(* exit *)
-val exit = ProofContext.theory Theory.checkpoint o operation #exit;
+val exit = ProofContext.background_theory Theory.checkpoint o operation #exit;
val exit_global = ProofContext.theory_of o exit;
fun exit_result f (x, lthy) =
--- a/src/Pure/Isar/locale.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/Isar/locale.ML Fri Aug 27 16:05:46 2010 +0200
@@ -483,7 +483,7 @@
val thy' = (change_locale loc o apsnd) (cons ((name, (morph, export)), serial ())) thy;
val context' = Context.Theory thy';
val (_, regs) = fold_rev (roundup thy' cons export)
- (all_registrations context') (get_idents (context'), []);
+ (registrations_of context' loc) (get_idents (context'), []);
in
thy'
|> fold_rev (fn dep => Context.theory_map (add_registration dep NONE export)) regs
@@ -497,8 +497,8 @@
fun add_thmss loc kind args ctxt =
let
val (Notes args', ctxt') = Element.activate_i (Notes (kind, args)) ctxt;
- val ctxt'' = ctxt' |> ProofContext.theory (
- (change_locale loc o apfst o apsnd) (cons (args', serial ()))
+ val ctxt'' = ctxt' |> ProofContext.background_theory
+ ((change_locale loc o apfst o apsnd) (cons (args', serial ()))
#>
(* Registrations *)
(fn thy => fold_rev (fn (_, morph) =>
@@ -519,7 +519,7 @@
[([Drule.dummy_thm], [])])];
fun add_syntax_declaration loc decl =
- ProofContext.theory ((change_locale loc o apfst o apfst) (cons (decl, serial ())))
+ ProofContext.background_theory ((change_locale loc o apfst o apfst) (cons (decl, serial ())))
#> add_declaration loc decl;
--- a/src/Pure/Isar/named_target.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/Isar/named_target.ML Fri Aug 27 16:05:46 2010 +0200
@@ -118,7 +118,7 @@
(Element.morph_ctxt (Local_Theory.target_morphism lthy)) local_facts;
in
lthy
- |> Local_Theory.theory (PureThy.note_thmss kind global_facts' #> snd)
+ |> Local_Theory.background_theory (PureThy.note_thmss kind global_facts' #> snd)
|> Local_Theory.target (Locale.add_thmss locale kind local_facts')
end
@@ -129,19 +129,21 @@
(* abbrev *)
-fun locale_abbrev ta prmode ((b, mx), t) xs = Local_Theory.theory_result
- (Sign.add_abbrev Print_Mode.internal (b, t)) #->
- (fn (lhs, _) => locale_const_declaration ta prmode
- ((b, mx), Term.list_comb (Logic.unvarify_global lhs, xs)));
+fun locale_abbrev ta prmode ((b, mx), t) xs =
+ Local_Theory.background_theory_result
+ (Sign.add_abbrev Print_Mode.internal (b, t)) #->
+ (fn (lhs, _) => locale_const_declaration ta prmode
+ ((b, mx), Term.list_comb (Logic.unvarify_global lhs, xs)));
fun target_abbrev (ta as Target {target, is_locale, is_class, ...})
prmode (b, mx) (global_rhs, t') xs lthy =
- if is_locale
- then lthy
- |> locale_abbrev ta prmode ((b, if is_class then NoSyn else mx), global_rhs) xs
- |> is_class ? Class.abbrev target prmode ((b, mx), t')
- else lthy
- |> Generic_Target.theory_abbrev prmode ((b, mx), global_rhs);
+ if is_locale then
+ lthy
+ |> locale_abbrev ta prmode ((b, if is_class then NoSyn else mx), global_rhs) xs
+ |> is_class ? Class.abbrev target prmode ((b, mx), t')
+ else
+ lthy
+ |> Generic_Target.theory_abbrev prmode ((b, mx), global_rhs);
(* pretty *)
@@ -200,9 +202,10 @@
fun init target thy = init_target (named_target thy target) thy;
-fun reinit lthy = case peek lthy
- of SOME {target, ...} => init target o Local_Theory.exit_global
- | NONE => error "Not in a named target";
+fun reinit lthy =
+ (case peek lthy of
+ SOME {target, ...} => init target o Local_Theory.exit_global
+ | NONE => error "Not in a named target");
val theory_init = init_target global_target;
--- a/src/Pure/Isar/overloading.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/Isar/overloading.ML Fri Aug 27 16:05:46 2010 +0200
@@ -140,7 +140,7 @@
end
fun define_overloaded (c, U) (v, checked) (b_def, rhs) =
- Local_Theory.theory_result
+ Local_Theory.background_theory_result
(Thm.add_def (not checked) true (b_def, Logic.mk_equals (Const (c, Term.fastype_of rhs), rhs)))
##> map_overloading (filter_out (fn (_, (v', _)) => v' = v))
##> Local_Theory.target synchronize_syntax
--- a/src/Pure/Isar/proof.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/Isar/proof.ML Fri Aug 27 16:05:46 2010 +0200
@@ -42,6 +42,7 @@
val raw_goal: state -> {context: context, facts: thm list, goal: thm}
val goal: state -> {context: context, facts: thm list, goal: thm}
val simple_goal: state -> {context: context, goal: thm}
+ val status_markup: state -> Markup.T
val let_bind: (term list * term) list -> state -> state
val let_bind_cmd: (string list * string) list -> state -> state
val write: Syntax.mode -> (term * mixfix) list -> state -> state
@@ -520,6 +521,11 @@
val (ctxt, (_, goal)) = get_goal (refine_insert facts state);
in {context = ctxt, goal = goal} end;
+fun status_markup state =
+ (case try goal state of
+ SOME {goal, ...} => Markup.proof_state (Thm.nprems_of goal)
+ | NONE => Markup.empty);
+
(*** structured proof commands ***)
--- a/src/Pure/Isar/proof_context.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/Isar/proof_context.ML Fri Aug 27 16:05:46 2010 +0200
@@ -38,8 +38,8 @@
val cases_of: Proof.context -> (string * (Rule_Cases.T * bool)) list
val transfer_syntax: theory -> Proof.context -> Proof.context
val transfer: theory -> Proof.context -> Proof.context
- val theory: (theory -> theory) -> Proof.context -> Proof.context
- val theory_result: (theory -> 'a * theory) -> Proof.context -> 'a * Proof.context
+ val background_theory: (theory -> theory) -> Proof.context -> Proof.context
+ val background_theory_result: (theory -> 'a * theory) -> Proof.context -> 'a * Proof.context
val extern_fact: Proof.context -> string -> xstring
val pretty_term_abbrev: Proof.context -> term -> Pretty.T
val pretty_fact_aux: Proof.context -> bool -> string * thm list -> Pretty.T
@@ -283,9 +283,9 @@
fun transfer thy = Context.raw_transfer thy #> transfer_syntax thy;
-fun theory f ctxt = transfer (f (theory_of ctxt)) ctxt;
+fun background_theory f ctxt = transfer (f (theory_of ctxt)) ctxt;
-fun theory_result f ctxt =
+fun background_theory_result f ctxt =
let val (res, thy') = f (theory_of ctxt)
in (res, ctxt |> transfer thy') end;
--- a/src/Pure/Isar/toplevel.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/Isar/toplevel.ML Fri Aug 27 16:05:46 2010 +0200
@@ -563,13 +563,6 @@
fun status tr m =
setmp_thread_position tr (fn () => Output.status (Markup.markup m "")) ();
-fun async_state (tr as Transition {print, ...}) st =
- if print then
- ignore
- (Future.fork (fn () =>
- setmp_thread_position tr (fn () => Future.status (fn () => print_state false st)) ()))
- else ();
-
fun error_msg tr exn_info =
setmp_thread_position tr (fn () =>
Output.error_msg (ML_Compiler.exn_message (EXCURSION_FAIL exn_info))) ();
@@ -628,6 +621,22 @@
(* managed execution *)
+local
+
+fun async_state (tr as Transition {print, ...}) st =
+ if print then
+ ignore
+ (Future.fork (fn () =>
+ setmp_thread_position tr (fn () => Future.status (fn () => print_state false st)) ()))
+ else ();
+
+fun proof_status tr st =
+ (case try proof_of st of
+ SOME prf => status tr (Proof.status_markup prf)
+ | NONE => ());
+
+in
+
fun run_command thy_name tr st =
(case
(case init_of tr of
@@ -637,13 +646,18 @@
let val int = is_some (init_of tr) in
(case transition int tr st of
SOME (st', NONE) =>
- (status tr Markup.finished; if int then () else async_state tr st'; SOME st')
+ (status tr Markup.finished;
+ proof_status tr st';
+ if int then () else async_state tr st';
+ SOME st')
| SOME (_, SOME (exn as Exn.Interrupt, _)) => reraise exn
| SOME (_, SOME exn_info) => (error_msg tr exn_info; status tr Markup.failed; NONE)
| NONE => (error_msg tr (TERMINATE, at_command tr); status tr Markup.failed; NONE))
end
| Exn.Exn exn => (error_msg tr (exn, at_command tr); status tr Markup.failed; NONE));
+end;
+
(* nested commands *)
--- a/src/Pure/Isar/typedecl.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/Isar/typedecl.ML Fri Aug 27 16:05:46 2010 +0200
@@ -34,7 +34,7 @@
fun basic_decl decl (b, n, mx) lthy =
let val name = Local_Theory.full_name lthy b in
lthy
- |> Local_Theory.theory (decl name)
+ |> Local_Theory.background_theory (decl name)
|> Local_Theory.type_notation true Syntax.mode_default [(Type (name, replicate n dummyT), mx)]
|> Local_Theory.type_alias b name
|> pair name
--- a/src/Pure/ML/ml_compiler_polyml-5.3.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/ML/ml_compiler_polyml-5.3.ML Fri Aug 27 16:05:46 2010 +0200
@@ -44,13 +44,18 @@
fun report_parse_tree depth space =
let
+ val report_text =
+ (case Context.thread_data () of
+ SOME (Context.Proof ctxt) => Context_Position.report_text ctxt
+ | _ => Position.report_text);
+
fun report_decl markup loc decl =
- Position.report_text Markup.ML_ref (position_of loc)
+ report_text Markup.ML_ref (position_of loc)
(Markup.markup (Markup.properties (Position.properties_of (position_of decl)) markup) "");
fun report loc (PolyML.PTtype types) =
PolyML.NameSpace.displayTypeExpression (types, depth, space)
|> pretty_ml |> Pretty.from_ML |> Pretty.string_of
- |> Position.report_text Markup.ML_typing (position_of loc)
+ |> report_text Markup.ML_typing (position_of loc)
| report loc (PolyML.PTdeclaredAt decl) = report_decl Markup.ML_def loc decl
| report loc (PolyML.PTopenedAt decl) = report_decl Markup.ML_open loc decl
| report loc (PolyML.PTstructureAt decl) = report_decl Markup.ML_struct loc decl
--- a/src/Pure/ML/ml_context.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/ML/ml_context.ML Fri Aug 27 16:05:46 2010 +0200
@@ -166,7 +166,9 @@
fun eval verbose pos ants =
let
(*prepare source text*)
- val ((env, body), env_ctxt) = eval_antiquotes (ants, pos) (Context.thread_data ());
+ val ((env, body), env_ctxt) =
+ eval_antiquotes (ants, pos) (Context.thread_data ())
+ ||> Option.map (Context.mapping I (Context_Position.set_visible false));
val _ =
if ! trace then tracing (cat_lines [ML_Lex.flatten env, ML_Lex.flatten body])
else ();
--- a/src/Pure/PIDE/command.scala Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/PIDE/command.scala Fri Aug 27 16:05:46 2010 +0200
@@ -46,11 +46,11 @@
case XML.Elem(Markup(Markup.REPORT, _), msgs) =>
(this /: msgs)((state, msg) =>
msg match {
- case XML.Elem(Markup(name, atts), args)
- if Position.get_id(atts) == Some(command.id) && Position.get_range(atts).isDefined =>
- val range = command.decode(Position.get_range(atts).get)
+ case XML.Elem(Markup(name, atts @ Position.Range(range)), args)
+ if Position.Id.unapply(atts) == Some(command.id) =>
val props = atts.filterNot(p => Markup.POSITION_PROPERTIES(p._1))
- val info = Text.Info[Any](range, XML.Elem(Markup(name, props), args))
+ val info =
+ Text.Info[Any](command.decode(range), XML.Elem(Markup(name, props), args))
state.add_markup(info)
case _ => System.err.println("Ignored report message: " + msg); state
})
--- a/src/Pure/PIDE/markup_tree.scala Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/PIDE/markup_tree.scala Fri Aug 27 16:05:46 2010 +0200
@@ -115,7 +115,10 @@
if (last < stop) parent.restrict(Text.Range(last, stop)) #:: nexts
else nexts
- case Nil => Stream(Text.Info(Text.Range(last, root_range.stop), default))
+ case Nil =>
+ val stop = root_range.stop
+ if (last < stop) Stream(Text.Info(Text.Range(last, stop), default))
+ else Stream.empty
}
}
stream(root_range.start, List((Text.Info(root_range, default), overlapping(root_range))))
--- a/src/Pure/PIDE/text.scala Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/PIDE/text.scala Fri Aug 27 16:05:46 2010 +0200
@@ -33,7 +33,7 @@
def +(i: Offset): Range = map(_ + i)
def -(i: Offset): Range = map(_ - i)
- def is_singleton: Boolean = start == stop
+ def is_singularity: Boolean = start == stop
def contains(i: Offset): Boolean = start == i || start < i && i < stop
def contains(that: Range): Boolean = this.contains(that.start) && that.stop <= this.stop
--- a/src/Pure/Proof/extraction.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/Proof/extraction.ML Fri Aug 27 16:05:46 2010 +0200
@@ -204,7 +204,7 @@
realizers = Symtab.merge_list (eq_set (op =) o pairself #1) (realizers1, realizers2),
defs = Library.merge Thm.eq_thm (defs1, defs2),
expand = Library.merge (op =) (expand1, expand2),
- prep = (case prep1 of NONE => prep2 | _ => prep1)};
+ prep = if is_some prep1 then prep1 else prep2};
);
fun read_condeq thy =
--- a/src/Pure/System/session.scala Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/System/session.scala Fri Aug 27 16:05:46 2010 +0200
@@ -131,15 +131,15 @@
{
raw_protocol.event(result)
- Position.get_id(result.properties) match {
- case Some(state_id) =>
+ result.properties match {
+ case Position.Id(state_id) =>
try {
val (st, state) = global_state.accumulate(state_id, result.message)
global_state = state
indicate_command_change(st.command)
}
catch { case _: Document.State.Fail => bad_result(result) }
- case None =>
+ case _ =>
if (result.is_status) {
result.body match {
case List(Isar_Document.Assign(id, edits)) =>
--- a/src/Pure/Thy/thy_load.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/Thy/thy_load.ML Fri Aug 27 16:05:46 2010 +0200
@@ -195,7 +195,7 @@
val _ = Context.>> Local_Theory.propagate_ml_env;
val provide = provide (src_path, (path, id));
- val _ = Context.>> (Context.mapping provide (Local_Theory.theory provide));
+ val _ = Context.>> (Context.mapping provide (Local_Theory.background_theory provide));
in () end;
fun exec_ml src_path = ML_Context.exec (fn () => use_ml src_path);
--- a/src/Pure/Thy/thy_output.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Pure/Thy/thy_output.ML Fri Aug 27 16:05:46 2010 +0200
@@ -6,31 +6,37 @@
signature THY_OUTPUT =
sig
- val display: bool Unsynchronized.ref
- val quotes: bool Unsynchronized.ref
- val indent: int Unsynchronized.ref
- val source: bool Unsynchronized.ref
- val break: bool Unsynchronized.ref
- val add_commands: (string * (Args.src -> Toplevel.state -> string)) list -> unit
- val add_options: (string * (string -> (unit -> string) -> unit -> string)) list -> unit
+ val display_default: bool Unsynchronized.ref
+ val quotes_default: bool Unsynchronized.ref
+ val indent_default: int Unsynchronized.ref
+ val source_default: bool Unsynchronized.ref
+ val break_default: bool Unsynchronized.ref
+ val display: bool Config.T
+ val quotes: bool Config.T
+ val indent: int Config.T
+ val source: bool Config.T
+ val break: bool Config.T
+ val add_wrapper: ((unit -> string) -> unit -> string) -> Proof.context -> Proof.context
+ val add_option: string -> (string -> Proof.context -> Proof.context) -> unit
val defined_command: string -> bool
val defined_option: string -> bool
val print_antiquotations: unit -> unit
val boolean: string -> bool
val integer: string -> int
val antiquotation: string -> 'a context_parser ->
- ({state: Toplevel.state, source: Args.src, context: Proof.context} -> 'a -> string) -> unit
+ ({source: Args.src, state: Toplevel.state, context: Proof.context} -> 'a -> string) -> unit
datatype markup = Markup | MarkupEnv | Verbatim
val modes: string list Unsynchronized.ref
val eval_antiquote: Scan.lexicon -> Toplevel.state -> Symbol_Pos.text * Position.T -> string
val present_thy: Scan.lexicon -> (string -> string list) -> (markup -> string -> bool) ->
(Toplevel.transition * Toplevel.state) list -> (Token.T, 'a) Source.source -> Buffer.T
- val pretty_text: string -> Pretty.T
+ val pretty_text: Proof.context -> string -> Pretty.T
val pretty_term: Proof.context -> term -> Pretty.T
val pretty_thm: Proof.context -> thm -> Pretty.T
val str_of_source: Args.src -> string
- val maybe_pretty_source: ('a -> Pretty.T) -> Args.src -> 'a list -> Pretty.T list
- val output: Pretty.T list -> string
+ val maybe_pretty_source: (Proof.context -> 'a -> Pretty.T) -> Proof.context ->
+ Args.src -> 'a list -> Pretty.T list
+ val output: Proof.context -> Pretty.T list -> string
end;
structure Thy_Output: THY_OUTPUT =
@@ -38,11 +44,31 @@
(** global options **)
-val display = Unsynchronized.ref false;
-val quotes = Unsynchronized.ref false;
-val indent = Unsynchronized.ref 0;
-val source = Unsynchronized.ref false;
-val break = Unsynchronized.ref false;
+val display_default = Unsynchronized.ref false;
+val quotes_default = Unsynchronized.ref false;
+val indent_default = Unsynchronized.ref 0;
+val source_default = Unsynchronized.ref false;
+val break_default = Unsynchronized.ref false;
+
+val (display, setup_display) = Attrib.config_bool "thy_output_display" (fn _ => ! display_default);
+val (quotes, setup_quotes) = Attrib.config_bool "thy_output_quotes" (fn _ => ! quotes_default);
+val (indent, setup_indent) = Attrib.config_int "thy_output_indent" (fn _ => ! indent_default);
+val (source, setup_source) = Attrib.config_bool "thy_output_source" (fn _ => ! source_default);
+val (break, setup_break) = Attrib.config_bool "thy_output_break" (fn _ => ! break_default);
+
+val _ = Context.>> (Context.map_theory
+ (setup_display #> setup_quotes #> setup_indent #> setup_source #> setup_break));
+
+
+structure Wrappers = Proof_Data
+(
+ type T = ((unit -> string) -> unit -> string) list;
+ fun init _ = [];
+);
+
+fun add_wrapper wrapper = Wrappers.map (cons wrapper);
+
+val wrap = Wrappers.get #> fold (fn wrapper => fn f => wrapper f);
@@ -51,22 +77,23 @@
local
val global_commands =
- Unsynchronized.ref (Symtab.empty: (Args.src -> Toplevel.state -> string) Symtab.table);
+ Unsynchronized.ref
+ (Symtab.empty: (Args.src -> Toplevel.state -> Proof.context -> string) Symtab.table);
val global_options =
- Unsynchronized.ref (Symtab.empty: (string -> (unit -> string) -> unit -> string) Symtab.table);
+ Unsynchronized.ref (Symtab.empty: (string -> Proof.context -> Proof.context) Symtab.table);
-fun add_item kind (name, x) tab =
+fun add_item kind name item tab =
(if not (Symtab.defined tab name) then ()
else warning ("Redefined document antiquotation " ^ kind ^ ": " ^ quote name);
- Symtab.update (name, x) tab);
+ Symtab.update (name, item) tab);
in
-fun add_commands xs =
- CRITICAL (fn () => Unsynchronized.change global_commands (fold (add_item "command") xs));
-fun add_options xs =
- CRITICAL (fn () => Unsynchronized.change global_options (fold (add_item "option") xs));
+fun add_command name cmd =
+ CRITICAL (fn () => Unsynchronized.change global_commands (add_item "command" name cmd));
+fun add_option name opt =
+ CRITICAL (fn () => Unsynchronized.change global_options (add_item "option" name opt));
fun defined_command name = Symtab.defined (! global_commands) name;
fun defined_option name = Symtab.defined (! global_options) name;
@@ -77,18 +104,15 @@
NONE => error ("Unknown document antiquotation: " ^ quote name ^ Position.str_of pos)
| SOME f =>
(Position.report (Markup.doc_antiq name) pos;
- (fn state => f src state handle ERROR msg =>
+ (fn state => fn ctxt => f src state ctxt handle ERROR msg =>
cat_error msg ("The error(s) above occurred in document antiquotation: " ^
quote name ^ Position.str_of pos))))
end;
-fun option (name, s) f () =
+fun option (name, s) ctxt =
(case Symtab.lookup (! global_options) name of
NONE => error ("Unknown document antiquotation option: " ^ quote name)
- | SOME opt => opt s f ());
-
-fun options [] f = f
- | options (opt :: opts) f = option opt (options opts f);
+ | SOME opt => opt s ctxt);
fun print_antiquotations () =
@@ -100,10 +124,11 @@
end;
-fun antiquotation name scan out = add_commands [(name, fn src => fn state =>
- let val (x, ctxt) = Args.context_syntax "document antiquotation"
- scan src (Toplevel.presentation_context_of state)
- in out {source = src, state = state, context = ctxt} x end)];
+fun antiquotation name scan out =
+ add_command name
+ (fn src => fn state => fn ctxt =>
+ let val (x, ctxt') = Args.context_syntax "document antiquotation" scan src ctxt
+ in out {source = src, state = state, context = ctxt'} x end);
@@ -151,10 +176,13 @@
let
fun expand (Antiquote.Text ss) = Symbol_Pos.content ss
| expand (Antiquote.Antiq (ss, (pos, _))) =
- let val (opts, src) = Token.read_antiq lex antiq (ss, pos) in
- options opts (fn () => command src state) (); (*preview errors!*)
- Print_Mode.with_modes (! modes @ Latex.modes)
- (Output.no_warnings_CRITICAL (options opts (fn () => command src state))) ()
+ let
+ val (opts, src) = Token.read_antiq lex antiq (ss, pos);
+ val opts_ctxt = fold option opts (Toplevel.presentation_context_of state);
+ val cmd = wrap opts_ctxt (fn () => command src state opts_ctxt);
+ val _ = cmd (); (*preview errors!*)
+ in
+ Print_Mode.with_modes (! modes @ Latex.modes) (Output.no_warnings_CRITICAL cmd) ()
end
| expand (Antiquote.Open _) = ""
| expand (Antiquote.Close _) = "";
@@ -417,31 +445,31 @@
(* options *)
-val _ = add_options
- [("show_types", setmp_CRITICAL Syntax.show_types o boolean),
- ("show_sorts", setmp_CRITICAL Syntax.show_sorts o boolean),
- ("show_structs", setmp_CRITICAL show_structs o boolean),
- ("show_question_marks", setmp_CRITICAL show_question_marks o boolean),
- ("long_names", setmp_CRITICAL Name_Space.long_names o boolean),
- ("short_names", setmp_CRITICAL Name_Space.short_names o boolean),
- ("unique_names", setmp_CRITICAL Name_Space.unique_names o boolean),
- ("eta_contract", setmp_CRITICAL Syntax.eta_contract o boolean),
- ("display", setmp_CRITICAL display o boolean),
- ("break", setmp_CRITICAL break o boolean),
- ("quotes", setmp_CRITICAL quotes o boolean),
- ("mode", fn s => fn f => Print_Mode.with_modes [s] f),
- ("margin", setmp_CRITICAL Pretty.margin_default o integer),
- ("indent", setmp_CRITICAL indent o integer),
- ("source", setmp_CRITICAL source o boolean),
- ("goals_limit", setmp_CRITICAL goals_limit o integer)];
+val _ = add_option "show_types" (add_wrapper o setmp_CRITICAL Syntax.show_types o boolean);
+val _ = add_option "show_sorts" (add_wrapper o setmp_CRITICAL Syntax.show_sorts o boolean);
+val _ = add_option "show_structs" (add_wrapper o setmp_CRITICAL show_structs o boolean);
+val _ = add_option "show_question_marks" (add_wrapper o setmp_CRITICAL show_question_marks o boolean);
+val _ = add_option "long_names" (add_wrapper o setmp_CRITICAL Name_Space.long_names o boolean);
+val _ = add_option "short_names" (add_wrapper o setmp_CRITICAL Name_Space.short_names o boolean);
+val _ = add_option "unique_names" (add_wrapper o setmp_CRITICAL Name_Space.unique_names o boolean);
+val _ = add_option "eta_contract" (add_wrapper o setmp_CRITICAL Syntax.eta_contract o boolean);
+val _ = add_option "display" (Config.put display o boolean);
+val _ = add_option "break" (Config.put break o boolean);
+val _ = add_option "quotes" (Config.put quotes o boolean);
+val _ = add_option "mode" (add_wrapper o Print_Mode.with_modes o single);
+val _ = add_option "margin" (add_wrapper o setmp_CRITICAL Pretty.margin_default o integer);
+val _ = add_option "indent" (Config.put indent o integer);
+val _ = add_option "source" (Config.put source o boolean);
+val _ = add_option "goals_limit" (add_wrapper o setmp_CRITICAL goals_limit o integer);
(* basic pretty printing *)
-fun tweak_line s =
- if ! display then s else Symbol.strip_blanks s;
+fun tweak_line ctxt s =
+ if Config.get ctxt display then s else Symbol.strip_blanks s;
-val pretty_text = Pretty.chunks o map Pretty.str o map tweak_line o Library.split_lines;
+fun pretty_text ctxt =
+ Pretty.chunks o map Pretty.str o map (tweak_line ctxt) o Library.split_lines;
fun pretty_term ctxt t = Syntax.pretty_term (Variable.auto_fixes t ctxt) t;
@@ -490,19 +518,19 @@
val str_of_source = space_implode " " o map Token.unparse o #2 o #1 o Args.dest_src;
-fun maybe_pretty_source pretty src xs =
- map pretty xs (*always pretty in order to exhibit errors!*)
- |> (if ! source then K [pretty_text (str_of_source src)] else I);
+fun maybe_pretty_source pretty ctxt src xs =
+ map (pretty ctxt) xs (*always pretty in order to exhibit errors!*)
+ |> (if Config.get ctxt source then K [pretty_text ctxt (str_of_source src)] else I);
-fun output prts =
+fun output ctxt prts =
prts
- |> (if ! quotes then map Pretty.quote else I)
- |> (if ! display then
- map (Output.output o Pretty.string_of o Pretty.indent (! indent))
+ |> (if Config.get ctxt quotes then map Pretty.quote else I)
+ |> (if Config.get ctxt display then
+ map (Output.output o Pretty.string_of o Pretty.indent (Config.get ctxt indent))
#> space_implode "\\isasep\\isanewline%\n"
#> enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
else
- map (Output.output o (if ! break then Pretty.string_of else Pretty.str_of))
+ map (Output.output o (if Config.get ctxt break then Pretty.string_of else Pretty.str_of))
#> space_implode "\\isasep\\isanewline%\n"
#> enclose "\\isa{" "}");
@@ -515,11 +543,12 @@
local
fun basic_entities name scan pretty = antiquotation name scan
- (fn {source, context, ...} => output o maybe_pretty_source (pretty context) source);
+ (fn {source, context, ...} => output context o maybe_pretty_source pretty context source);
fun basic_entities_style name scan pretty = antiquotation name scan
(fn {source, context, ...} => fn (style, xs) =>
- output (maybe_pretty_source (fn x => pretty context (style, x)) source xs));
+ output context
+ (maybe_pretty_source (fn ctxt => fn x => pretty ctxt (style, x)) context source xs));
fun basic_entity name scan = basic_entities name (scan >> single);
@@ -533,7 +562,7 @@
val _ = basic_entity "const" (Args.const_proper false) pretty_const;
val _ = basic_entity "abbrev" (Scan.lift Args.name_source) pretty_abbrev;
val _ = basic_entity "typ" Args.typ_abbrev Syntax.pretty_typ;
-val _ = basic_entity "text" (Scan.lift Args.name) (K pretty_text);
+val _ = basic_entity "text" (Scan.lift Args.name) pretty_text;
val _ = basic_entities "prf" Attrib.thms (pretty_prf false);
val _ = basic_entities "full_prf" Attrib.thms (pretty_prf true);
val _ = basic_entity "theory" (Scan.lift Args.name) pretty_theory;
@@ -554,7 +583,7 @@
| _ => error "No proof state");
fun goal_state name main_goal = antiquotation name (Scan.succeed ())
- (fn {state, ...} => fn () => output
+ (fn {state, context, ...} => fn () => output context
[Pretty.chunks (Proof.pretty_goals main_goal (proof_state state))]);
in
@@ -578,7 +607,7 @@
val _ = context
|> Proof.theorem NONE (K I) [[(prop, [])]]
|> Proof.global_terminal_proof methods;
- in output (maybe_pretty_source (pretty_term context) prop_src [prop]) end);
+ in output context (maybe_pretty_source pretty_term context prop_src [prop]) end);
(* ML text *)
@@ -589,8 +618,8 @@
(fn {context, ...} => fn (txt, pos) =>
(ML_Context.eval_in (SOME context) false pos (ml pos txt);
Symbol_Pos.content (Symbol_Pos.explode (txt, pos))
- |> (if ! quotes then quote else I)
- |> (if ! display then enclose "\\begin{verbatim}\n" "\n\\end{verbatim}"
+ |> (if Config.get context quotes then quote else I)
+ |> (if Config.get context display then enclose "\\begin{verbatim}\n" "\n\\end{verbatim}"
else
split_lines
#> map (space_implode "\\verb,|," o map (enclose "\\verb|" "|") o space_explode "|")
--- a/src/Tools/Code/code_haskell.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Tools/Code/code_haskell.ML Fri Aug 27 16:05:46 2010 +0200
@@ -261,9 +261,8 @@
end;
in print_stmt end;
-fun haskell_program_of_program labelled_name module_name module_prefix reserved raw_module_alias program =
+fun haskell_program_of_program labelled_name module_prefix reserved module_alias program =
let
- val module_alias = if is_some module_name then K module_name else raw_module_alias;
val reserved = Name.make_context reserved;
val mk_name_module = mk_name_module reserved module_prefix module_alias program;
fun add_stmt (name, (stmt, deps)) =
@@ -314,15 +313,14 @@
handle Option => error ("Unknown statement name: " ^ labelled_name name);
in (deresolver, hs_program) end;
-fun serialize_haskell module_prefix raw_module_name string_classes labelled_name
- raw_reserved includes raw_module_alias
- syntax_class syntax_tyco syntax_const (code_of_pretty, code_writeln) program stmt_names destination =
+fun serialize_haskell module_prefix module_name string_classes labelled_name
+ raw_reserved includes module_alias
+ syntax_class syntax_tyco syntax_const (code_of_pretty, code_writeln) program
+ (stmt_names, presentation_stmt_names) destination =
let
- val presentation_stmt_names = Code_Target.stmt_names_of_destination destination;
- val module_name = if null presentation_stmt_names then raw_module_name else SOME "Code";
val reserved = fold (insert (op =) o fst) includes raw_reserved;
val (deresolver, hs_program) = haskell_program_of_program labelled_name
- module_name module_prefix reserved raw_module_alias program;
+ module_prefix reserved module_alias program;
val contr_classparam_typs = Code_Thingol.contr_classparam_typs program;
fun deriving_show tyco =
let
@@ -344,11 +342,9 @@
contr_classparam_typs
(if string_classes then deriving_show else K false);
fun print_module name content =
- (name, Pretty.chunks [
+ (name, Pretty.chunks2 [
str ("module " ^ name ^ " where {"),
- str "",
content,
- str "",
str "}"
]);
fun serialize_module1 (module_name', (deps, (stmts, _))) =
--- a/src/Tools/Code/code_ml.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Tools/Code/code_ml.ML Fri Aug 27 16:05:46 2010 +0200
@@ -489,7 +489,7 @@
|> intro_vars ((fold o Code_Thingol.fold_varnames)
(insert (op =)) ts []);
in concat [
- (Pretty.block o Pretty.commas)
+ (Pretty.block o commas)
(map (print_term is_pseudo_fun some_thm vars NOBR) ts),
str "->",
print_term is_pseudo_fun some_thm vars NOBR t
@@ -535,7 +535,7 @@
:: Pretty.brk 1
:: str "match"
:: Pretty.brk 1
- :: (Pretty.block o Pretty.commas) dummy_parms
+ :: (Pretty.block o commas) dummy_parms
:: Pretty.brk 1
:: str "with"
:: Pretty.brk 1
@@ -722,9 +722,8 @@
in
-fun ml_node_of_program labelled_name module_name reserved raw_module_alias program =
+fun ml_node_of_program labelled_name module_name reserved module_alias program =
let
- val module_alias = if is_some module_name then K module_name else raw_module_alias;
val reserved = Name.make_context reserved;
val empty_module = ((reserved, reserved), Graph.empty);
fun map_node [] f = f
@@ -813,7 +812,7 @@
) stmts
#>> (split_list #> apsnd (map_filter I
#> (fn [] => error ("Datatype block without data statement: "
- ^ (commas o map (labelled_name o fst)) stmts)
+ ^ (Library.commas o map (labelled_name o fst)) stmts)
| stmts => ML_Datas stmts)));
fun add_class stmts =
fold_map
@@ -828,7 +827,7 @@
) stmts
#>> (split_list #> apsnd (map_filter I
#> (fn [] => error ("Class block without class statement: "
- ^ (commas o map (labelled_name o fst)) stmts)
+ ^ (Library.commas o map (labelled_name o fst)) stmts)
| [stmt] => ML_Class stmt)));
fun add_stmts ([stmt as (name, Code_Thingol.Fun _)]) =
add_fun stmt
@@ -849,7 +848,7 @@
| add_stmts ((stmts as (_, Code_Thingol.Classinst _)::_)) =
add_funs stmts
| add_stmts stmts = error ("Illegal mutual dependencies: " ^
- (commas o map (labelled_name o fst)) stmts);
+ (Library.commas o map (labelled_name o fst)) stmts);
fun add_stmts' stmts nsp_nodes =
let
val names as (name :: names') = map fst stmts;
@@ -858,9 +857,9 @@
val module_name = (the_single o distinct (op =) o map mk_name_module) module_names
handle Empty =>
error ("Different namespace prefixes for mutual dependencies:\n"
- ^ commas (map labelled_name names)
+ ^ Library.commas (map labelled_name names)
^ "\n"
- ^ commas module_names);
+ ^ Library.commas module_names);
val module_name_path = Long_Name.explode module_name;
fun add_dep name name' =
let
@@ -907,15 +906,14 @@
error ("Unknown statement name: " ^ labelled_name name);
in (deresolver, nodes) end;
-fun serialize_ml target print_module print_stmt raw_module_name with_signatures labelled_name
- reserved includes raw_module_alias _ syntax_tyco syntax_const (code_of_pretty, code_writeln) program stmt_names destination =
+fun serialize_ml target print_module print_stmt module_name with_signatures labelled_name
+ reserved includes module_alias _ syntax_tyco syntax_const (code_of_pretty, code_writeln) program
+ (stmt_names, presentation_stmt_names) =
let
val is_cons = Code_Thingol.is_cons program;
- val presentation_stmt_names = Code_Target.stmt_names_of_destination destination;
val is_presentation = not (null presentation_stmt_names);
- val module_name = if is_presentation then SOME "Code" else raw_module_name;
val (deresolver, nodes) = ml_node_of_program labelled_name module_name
- reserved raw_module_alias program;
+ reserved module_alias program;
val reserved = make_vars reserved;
fun print_node prefix (Dummy _) =
NONE
@@ -939,7 +937,7 @@
in
Code_Target.mk_serialization target
(fn NONE => code_writeln | SOME file => File.write file o code_of_pretty)
- (rpair stmt_names' o code_of_pretty) p destination
+ (rpair stmt_names' o code_of_pretty) p
end;
end; (*local*)
@@ -948,8 +946,8 @@
(** for instrumentalization **)
fun evaluation_code_of thy target struct_name =
- Code_Target.serialize_custom thy (target, fn _ => fn [] =>
- serialize_ml target print_sml_module print_sml_stmt (SOME struct_name) true);
+ Code_Target.serialize_custom thy (target, fn module_name => fn [] =>
+ serialize_ml target print_sml_module print_sml_stmt module_name true) (SOME struct_name);
(** Isar setup **)
--- a/src/Tools/Code/code_printer.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Tools/Code/code_printer.ML Fri Aug 27 16:05:46 2010 +0200
@@ -19,6 +19,7 @@
val concat: Pretty.T list -> Pretty.T
val brackets: Pretty.T list -> Pretty.T
val enclose: string -> string -> Pretty.T list -> Pretty.T
+ val commas: Pretty.T list -> Pretty.T list
val enum: string -> string -> string -> Pretty.T list -> Pretty.T
val enum_default: string -> string -> string -> string -> Pretty.T list -> Pretty.T
val semicolon: Pretty.T list -> Pretty.T
@@ -112,6 +113,7 @@
fun xs @| y = xs @ [y];
val str = Print_Mode.setmp [] Pretty.str;
val concat = Pretty.block o Pretty.breaks;
+val commas = Print_Mode.setmp [] Pretty.commas;
fun enclose l r = Print_Mode.setmp [] (Pretty.enclose l r);
val brackets = enclose "(" ")" o Pretty.breaks;
fun enum sep l r = Print_Mode.setmp [] (Pretty.enum sep l r);
--- a/src/Tools/Code/code_scala.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Tools/Code/code_scala.ML Fri Aug 27 16:05:46 2010 +0200
@@ -25,9 +25,8 @@
(** Scala serializer **)
fun print_scala_stmt labelled_name syntax_tyco syntax_const reserved
- args_num is_singleton_constr deresolve =
+ args_num is_singleton_constr (deresolve, deresolve_full) =
let
- val deresolve_base = Long_Name.base_name o deresolve;
fun lookup_tyvar tyvars = lookup_var tyvars o first_upper;
fun intro_tyvars vs = intro_vars (map (first_upper o fst) vs);
fun print_tyco_expr tyvars fxy (tyco, tys) = applify "[" "]"
@@ -194,29 +193,30 @@
str "match", str "{"], str "}")
(map print_clause eqs)
end;
- val print_method = (str o Library.enclose "`" "+`" o deresolve_base);
+ val print_method = str o Library.enclose "`" "`" o space_implode "+"
+ o Long_Name.explode o deresolve_full;
fun print_stmt (name, Code_Thingol.Fun (_, (((vs, ty), raw_eqs), _))) =
print_def name (vs, ty) (filter (snd o snd) raw_eqs)
| print_stmt (name, Code_Thingol.Datatype (_, (vs, cos))) =
let
val tyvars = intro_tyvars vs reserved;
fun print_co ((co, _), []) =
- concat [str "final", str "case", str "object", (str o deresolve_base) co,
- str "extends", applify "[" "]" I NOBR ((str o deresolve_base) name)
+ concat [str "final", str "case", str "object", (str o deresolve) co,
+ str "extends", applify "[" "]" I NOBR ((str o deresolve) name)
(replicate (length vs) (str "Nothing"))]
| print_co ((co, vs_args), tys) =
concat [applify "(" ")"
(fn (v, arg) => constraint (str v) (print_typ tyvars NOBR arg)) NOBR
(applify "[" "]" (str o lookup_tyvar tyvars) NOBR ((concat o map str)
- ["final", "case", "class", deresolve_base co]) vs_args)
+ ["final", "case", "class", deresolve co]) vs_args)
(Name.names (snd reserved) "a" tys),
str "extends",
applify "[" "]" (str o lookup_tyvar tyvars o fst) NOBR
- ((str o deresolve_base) name) vs
+ ((str o deresolve) name) vs
];
in
Pretty.chunks (applify "[" "]" (str o prefix "+" o lookup_tyvar tyvars o fst)
- NOBR ((concat o map str) ["abstract", "sealed", "class", deresolve_base name]) vs
+ NOBR ((concat o map str) ["abstract", "sealed", "class", deresolve name]) vs
:: map print_co cos)
end
| print_stmt (name, Code_Thingol.Class (_, (v, (super_classes, classparams)))) =
@@ -249,7 +249,7 @@
in
Pretty.chunks (
(Pretty.block_enclose
- (concat ([str "trait", (add_typarg o deresolve_base) name]
+ (concat ([str "trait", (add_typarg o deresolve) name]
@ the_list (print_super_classes super_classes) @ [str "{"]), str "}")
(map print_classparam_val classparams))
:: map print_classparam_def classparams
@@ -281,67 +281,149 @@
end;
in print_stmt end;
-fun scala_program_of_program labelled_name module_name reserved raw_module_alias program =
+local
+
+(* hierarchical module name space *)
+
+datatype node =
+ Dummy
+ | Stmt of Code_Thingol.stmt
+ | Module of (string list * (string * node) Graph.T);
+
+in
+
+fun scala_program_of_program labelled_name reserved module_alias program =
let
- val the_module_name = the_default "Program" module_name;
- val module_alias = K (SOME the_module_name);
- val reserved = Name.make_context reserved;
- fun prepare_stmt (name, stmt) (nsps, stmts) =
+
+ (* building module name hierarchy *)
+ fun alias_fragments name = case module_alias name
+ of SOME name' => Long_Name.explode name'
+ | NONE => map (fn name => fst (yield_singleton Name.variants name reserved))
+ (Long_Name.explode name);
+ val module_names = Graph.fold (insert (op =) o fst o dest_name o fst) program [];
+ val fragments_tab = fold (fn name => Symtab.update
+ (name, alias_fragments name)) module_names Symtab.empty;
+ val dest_name = Code_Printer.dest_name #>> (the o Symtab.lookup fragments_tab);
+
+ (* building empty module hierarchy *)
+ val empty_module = ([], Graph.empty);
+ fun map_module f (Module content) = Module (f content);
+ fun change_module [] = I
+ | change_module (name_fragment :: name_fragments) =
+ apsnd o Graph.map_node name_fragment o apsnd o map_module
+ o change_module name_fragments;
+ fun ensure_module name_fragment (implicits, nodes) =
+ if can (Graph.get_node nodes) name_fragment then (implicits, nodes)
+ else (implicits,
+ nodes |> Graph.new_node (name_fragment, (name_fragment, Module empty_module)));
+ fun allocate_module [] = I
+ | allocate_module (name_fragment :: name_fragments) =
+ ensure_module name_fragment
+ #> (apsnd o Graph.map_node name_fragment o apsnd o map_module o allocate_module) name_fragments;
+ val empty_program = Symtab.fold (fn (_, fragments) => allocate_module fragments)
+ fragments_tab empty_module;
+
+ (* distribute statements over hierarchy *)
+ fun add_stmt name stmt =
let
- val (_, base) = Code_Printer.dest_name name;
- val mk_name_stmt = yield_singleton Name.variants;
- fun add_class ((nsp_class, nsp_object), nsp_common) =
- let
- val (base', nsp_class') = mk_name_stmt base nsp_class
- in (base', ((nsp_class', nsp_object), Name.declare base' nsp_common)) end;
- fun add_object ((nsp_class, nsp_object), nsp_common) =
- let
- val (base', nsp_object') = mk_name_stmt base nsp_object
- in (base', ((nsp_class, nsp_object'), Name.declare base' nsp_common)) end;
- fun add_common upper ((nsp_class, nsp_object), nsp_common) =
+ val (name_fragments, base) = dest_name name;
+ fun is_classinst stmt = case stmt
+ of Code_Thingol.Classinst _ => true
+ | _ => false;
+ val implicit_deps = filter (is_classinst o Graph.get_node program)
+ (Graph.imm_succs program name);
+ in
+ change_module name_fragments (fn (implicits, nodes) =>
+ (union (op =) implicit_deps implicits, Graph.new_node (name, (base, Stmt stmt)) nodes))
+ end;
+ fun add_dependency name name' =
+ let
+ val (name_fragments, base) = dest_name name;
+ val (name_fragments', base') = dest_name name';
+ val (name_fragments_common, (diff, diff')) =
+ chop_prefix (op =) (name_fragments, name_fragments');
+ val dep = if null diff then (name, name') else (hd diff, hd diff')
+ in (change_module name_fragments_common o apsnd) (Graph.add_edge dep) end;
+ val proto_program = empty_program
+ |> Graph.fold (fn (name, (stmt, _)) => add_stmt name stmt) program
+ |> Graph.fold (fn (name, (_, (_, names))) => fold (add_dependency name) names) program;
+
+ (* name declarations *)
+ fun namify_module name_fragment ((nsp_class, nsp_object), nsp_common) =
+ let
+ val declare = Name.declare name_fragment;
+ in (name_fragment, ((declare nsp_class, declare nsp_object), declare nsp_common)) end;
+ fun namify_class base ((nsp_class, nsp_object), nsp_common) =
+ let
+ val (base', nsp_class') = yield_singleton Name.variants base nsp_class
+ in (base', ((nsp_class', nsp_object), Name.declare base' nsp_common)) end;
+ fun namify_object base ((nsp_class, nsp_object), nsp_common) =
+ let
+ val (base', nsp_object') = yield_singleton Name.variants base nsp_object
+ in (base', ((nsp_class, nsp_object'), Name.declare base' nsp_common)) end;
+ fun namify_common upper base ((nsp_class, nsp_object), nsp_common) =
+ let
+ val (base', nsp_common') =
+ yield_singleton Name.variants (if upper then first_upper base else base) nsp_common
+ in
+ (base',
+ ((Name.declare base' nsp_class, Name.declare base' nsp_object), nsp_common'))
+ end;
+ fun namify_stmt (Code_Thingol.Fun _) = namify_object
+ | namify_stmt (Code_Thingol.Datatype _) = namify_class
+ | namify_stmt (Code_Thingol.Datatypecons _) = namify_common true
+ | namify_stmt (Code_Thingol.Class _) = namify_class
+ | namify_stmt (Code_Thingol.Classrel _) = namify_object
+ | namify_stmt (Code_Thingol.Classparam _) = namify_object
+ | namify_stmt (Code_Thingol.Classinst _) = namify_common false;
+ fun make_declarations nsps (implicits, nodes) =
+ let
+ val (module_fragments, stmt_names) = List.partition
+ (fn name_fragment => case Graph.get_node nodes name_fragment
+ of (_, Module _) => true | _ => false) (Graph.keys nodes);
+ fun modify_stmt (Stmt (Code_Thingol.Datatypecons _)) = Dummy
+ | modify_stmt (Stmt (Code_Thingol.Classrel _)) = Dummy
+ | modify_stmt (Stmt (Code_Thingol.Classparam _)) = Dummy
+ | modify_stmt stmt = stmt;
+ fun declare namify modify name (nsps, nodes) =
let
- val (base', nsp_common') =
- mk_name_stmt (if upper then first_upper base else base) nsp_common
- in
- (base',
- ((Name.declare base' nsp_class, Name.declare base' nsp_object), nsp_common'))
- end;
- val add_name = case stmt
- of Code_Thingol.Fun _ => add_object
- | Code_Thingol.Datatype _ => add_class
- | Code_Thingol.Datatypecons _ => add_common true
- | Code_Thingol.Class _ => add_class
- | Code_Thingol.Classrel _ => add_object
- | Code_Thingol.Classparam _ => add_object
- | Code_Thingol.Classinst _ => add_common false;
- fun add_stmt base' = case stmt
- of Code_Thingol.Datatypecons _ => cons (name, (base', NONE))
- | Code_Thingol.Classrel _ => cons (name, (base', NONE))
- | Code_Thingol.Classparam _ => cons (name, (base', NONE))
- | _ => cons (name, (base', SOME stmt));
- in
- nsps
- |> add_name
- |-> (fn base' => rpair (add_stmt base' stmts))
- end;
- val stmts = AList.make (Graph.get_node program) (Graph.strong_conn program |> flat)
- |> filter_out (Code_Thingol.is_case o snd);
- val (_, sca_program) = fold prepare_stmt stmts (((reserved, reserved), reserved), []);
- fun deresolver name = (fst o the o AList.lookup (op =) sca_program) name
- handle Option => error ("Unknown statement name: " ^ labelled_name name);
- in (deresolver, (the_module_name, sca_program)) end;
+ val (base, node) = Graph.get_node nodes name;
+ val (base', nsps') = namify node base nsps;
+ val nodes' = Graph.map_node name (K (base', modify node)) nodes;
+ in (nsps', nodes') end;
+ val (nsps', nodes') = (nsps, nodes)
+ |> fold (declare (K namify_module) I) module_fragments
+ |> fold (declare (namify_stmt o (fn Stmt stmt => stmt)) modify_stmt) stmt_names;
+ val nodes'' = nodes'
+ |> fold (fn name_fragment => (Graph.map_node name_fragment
+ o apsnd o map_module) (make_declarations nsps')) module_fragments;
+ in (implicits, nodes'') end;
+ val (_, sca_program) = make_declarations ((reserved, reserved), reserved) proto_program;
-fun serialize_scala raw_module_name labelled_name
- raw_reserved includes raw_module_alias
+ (* deresolving *)
+ fun deresolver prefix_fragments name =
+ let
+ val (name_fragments, _) = dest_name name;
+ val (_, (_, remainder)) = chop_prefix (op =) (prefix_fragments, name_fragments);
+ val nodes = fold (fn name_fragment => fn nodes => case Graph.get_node nodes name_fragment
+ of (_, Module (_, nodes)) => nodes) name_fragments sca_program;
+ val (base', _) = Graph.get_node nodes name;
+ in Long_Name.implode (remainder @ [base']) end
+ handle Graph.UNDEF _ => error ("Unknown statement name: " ^ labelled_name name);
+
+ in (deresolver, sca_program) end;
+
+fun serialize_scala labelled_name raw_reserved includes module_alias
_ syntax_tyco syntax_const (code_of_pretty, code_writeln)
- program stmt_names destination =
+ program (stmt_names, presentation_stmt_names) destination =
let
- val presentation_stmt_names = Code_Target.stmt_names_of_destination destination;
- val module_name = if null presentation_stmt_names then raw_module_name else SOME "Code";
+
+ (* build program *)
val reserved = fold (insert (op =) o fst) includes raw_reserved;
- val (deresolver, (the_module_name, sca_program)) = scala_program_of_program labelled_name
- module_name reserved raw_module_alias program;
- val reserved = make_vars reserved;
+ val (deresolver, sca_program) = scala_program_of_program labelled_name
+ (Name.make_context reserved) module_alias program;
+
+ (* print statements *)
fun lookup_constr tyco constr = case Graph.get_node program tyco
of Code_Thingol.Datatype (_, (_, constrs)) =>
the (AList.lookup (op = o apsnd fst) constrs constr);
@@ -359,44 +441,52 @@
of Code_Thingol.Datatypecons (_, tyco) => null (lookup_constr tyco c)
| _ => false;
val print_stmt = print_scala_stmt labelled_name syntax_tyco syntax_const
- reserved args_num is_singleton_constr deresolver;
- fun print_module name imports content =
- (name, Pretty.chunks (
- str ("object " ^ name ^ " {")
- :: (if null imports then []
- else str "" :: map (fn name => str ("import " ^ name ^ "._")) imports)
- @ [str "",
- content,
- str "",
- str "}"]
- ));
- fun serialize_module the_module_name sca_program =
+ (make_vars reserved) args_num is_singleton_constr;
+
+ (* print nodes *)
+ fun print_implicit prefix_fragments implicit =
let
- val content = Pretty.chunks2 (map_filter
- (fn (name, (_, SOME stmt)) => SOME (print_stmt (name, stmt))
- | (_, (_, NONE)) => NONE) sca_program);
- in print_module the_module_name (map fst includes) content end;
- fun check_destination destination =
- (File.check destination; destination);
- fun write_module destination (modlname, content) =
- let
- val filename = case modlname
- of "" => Path.explode "Main.scala"
- | _ => (Path.ext "scala" o Path.explode o implode o separate "/"
- o Long_Name.explode) modlname;
- val pathname = Path.append destination filename;
- val _ = File.mkdir_leaf (Path.dir pathname);
- in File.write pathname (code_of_pretty content) end
+ val s = deresolver prefix_fragments implicit;
+ in if length (Long_Name.explode s) = 1 then NONE else SOME (str s) end;
+ fun print_implicits prefix_fragments implicits =
+ case map_filter (print_implicit prefix_fragments) implicits
+ of [] => NONE
+ | ps => (SOME o Pretty.block)
+ (str "import /*implicits*/" :: Pretty.brk 1 :: commas ps);
+ fun print_module prefix_fragments base implicits p = Pretty.chunks2
+ ([str ("object " ^ base ^ " {")] @ the_list (print_implicits prefix_fragments implicits)
+ @ [p, str ("} /* object " ^ base ^ " */")]);
+ fun print_node _ (_, Dummy) = NONE
+ | print_node prefix_fragments (name, Stmt stmt) =
+ if null presentation_stmt_names
+ orelse member (op =) presentation_stmt_names name
+ then SOME (print_stmt (deresolver prefix_fragments, deresolver []) (name, stmt))
+ else NONE
+ | print_node prefix_fragments (name_fragment, Module (implicits, nodes)) =
+ if null presentation_stmt_names
+ then case print_nodes (prefix_fragments @ [name_fragment]) nodes
+ of NONE => NONE
+ | SOME p => SOME (print_module prefix_fragments
+ (Long_Name.base_name name_fragment) implicits p)
+ else print_nodes [] nodes
+ and print_nodes prefix_fragments nodes = let
+ val ps = map_filter (fn name => print_node prefix_fragments (name,
+ snd (Graph.get_node nodes name)))
+ ((rev o flat o Graph.strong_conn) nodes);
+ in if null ps then NONE else SOME (Pretty.chunks2 ps) end;
+
+ (* serialization *)
+ val p_includes = if null presentation_stmt_names
+ then map (fn (base, p) => print_module [] base [] p) includes else [];
+ val p = Pretty.chunks2 (p_includes @ the_list (print_nodes [] sca_program));
in
Code_Target.mk_serialization target
- (fn NONE => K () o map (code_writeln o snd) | SOME file => K () o map
- (write_module (check_destination file)))
- (rpair [] o cat_lines o map (code_of_pretty o snd))
- (map (fn (name, content) => print_module name [] content) includes
- @| serialize_module the_module_name sca_program)
- destination
+ (fn NONE => code_writeln | SOME file => File.write file o code_of_pretty)
+ (rpair [] o code_of_pretty) p destination
end;
+end; (*local*)
+
val literals = let
fun char_scala c = if c = "'" then "\\'"
else if c = "\"" then "\\\""
@@ -412,8 +502,8 @@
literal_char = Library.enclose "'" "'" o char_scala,
literal_string = quote o translate_string char_scala,
literal_numeral = fn k => "BigInt(" ^ numeral_scala k ^ ")",
- literal_positive_numeral = fn k => "Nat(" ^ numeral_scala k ^ ")",
- literal_alternative_numeral = fn k => "Natural(" ^ numeral_scala k ^ ")",
+ literal_positive_numeral = fn k => "Nat.Nat(" ^ numeral_scala k ^ ")",
+ literal_alternative_numeral = fn k => "Natural.Nat(" ^ numeral_scala k ^ ")",
literal_naive_numeral = fn k => "BigInt(" ^ numeral_scala k ^ ")",
literal_list = fn [] => str "Nil" | ps => Pretty.block [str "List", enum "," "(" ")" ps],
infix_cons = (6, "::")
@@ -422,17 +512,17 @@
(** Isar setup **)
-fun isar_serializer module_name =
+fun isar_serializer _ =
Code_Target.parse_args (Scan.succeed ())
- #> (fn () => serialize_scala module_name);
+ #> (fn () => serialize_scala);
val setup =
Code_Target.add_target
(target, { serializer = isar_serializer, literals = literals,
- check = { env_var = "SCALA_HOME", make_destination = I,
+ check = { env_var = "SCALA_HOME", make_destination = fn p => Path.append p (Path.explode "ROOT.scala"),
make_command = fn scala_home => fn p => fn _ =>
"export JAVA_OPTS='-Xms128m -Xmx512m -Xss2m' && "
- ^ Path.implode (Path.append (Path.explode scala_home) (Path.explode "bin/scalac")) ^ " *.scala" } })
+ ^ Path.implode (Path.append (Path.explode scala_home) (Path.explode "bin/scalac")) ^ " " ^ File.shell_path p } })
#> Code_Target.add_syntax_tyco target "fun"
(SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] =>
brackify_infix (1, R) fxy (
--- a/src/Tools/Code/code_simp.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Tools/Code/code_simp.ML Fri Aug 27 16:05:46 2010 +0200
@@ -37,7 +37,8 @@
(* dedicated simpset *)
-structure Simpset = Theory_Data (
+structure Simpset = Theory_Data
+(
type T = simpset;
val empty = empty_ss;
fun extend ss = MetaSimplifier.inherit_context empty_ss ss;
--- a/src/Tools/Code/code_target.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Tools/Code/code_target.ML Fri Aug 27 16:05:46 2010 +0200
@@ -28,7 +28,7 @@
-> 'a -> serialization
val serialize: theory -> string -> int option -> string option -> Token.T list
-> Code_Thingol.naming -> Code_Thingol.program -> string list -> serialization
- val serialize_custom: theory -> string * serializer
+ val serialize_custom: theory -> string * serializer -> string option
-> Code_Thingol.naming -> Code_Thingol.program -> string list -> string * string option list
val the_literals: theory -> string -> literals
val export: serialization -> unit
@@ -111,7 +111,7 @@
-> (string -> Code_Printer.activated_const_syntax option)
-> ((Pretty.T -> string) * (Pretty.T -> unit))
-> Code_Thingol.program
- -> string list (*selected statements*)
+ -> (string list * string list) (*selected statements*)
-> serialization;
datatype description = Fundamental of { serializer: serializer, literals: Code_Printer.literals,
@@ -254,7 +254,7 @@
|>> map_filter I;
fun invoke_serializer thy abortable serializer literals reserved abs_includes
- module_alias class instance tyco const module width args naming program2 names1 =
+ module_alias class instance tyco const module_name width args naming program2 (names1, presentation_names) =
let
val (names_class, class') =
activate_syntax (Code_Thingol.lookup_class naming) class;
@@ -275,14 +275,14 @@
val _ = if null empty_funs then () else error ("No code equations for "
^ commas (map (Sign.extern_const thy) empty_funs));
in
- serializer module args (Code_Thingol.labelled_name thy program2) reserved includes
- (Symtab.lookup module_alias) (Symtab.lookup class')
- (Symtab.lookup tyco') (Symtab.lookup const')
+ serializer module_name args (Code_Thingol.labelled_name thy program2) reserved includes
+ (if is_some module_name then K module_name else Symtab.lookup module_alias)
+ (Symtab.lookup class') (Symtab.lookup tyco') (Symtab.lookup const')
(Code_Printer.string_of_pretty width, Code_Printer.writeln_pretty width)
- program4 names1
+ program4 (names1, presentation_names)
end;
-fun mount_serializer thy alt_serializer target some_width module args naming program names =
+fun mount_serializer thy alt_serializer target some_width raw_module_name args naming program names destination =
let
val ((targets, abortable), default_width) = Targets.get thy;
fun collapse_hierarchy target =
@@ -299,6 +299,9 @@
val (modify, data) = collapse_hierarchy target;
val serializer = the_default (case the_description data
of Fundamental seri => #serializer seri) alt_serializer;
+ val presentation_names = stmt_names_of_destination destination;
+ val module_name = if null presentation_names
+ then raw_module_name else SOME "Code";
val reserved = the_reserved data;
fun select_include names_all (name, (content, cs)) =
if null cs then SOME (name, content)
@@ -308,13 +311,14 @@
then SOME (name, content) else NONE;
fun includes names_all = map_filter (select_include names_all)
((Symtab.dest o the_includes) data);
- val module_alias = the_module_alias data;
+ val module_alias = the_module_alias data
val { class, instance, tyco, const } = the_name_syntax data;
val literals = the_literals thy target;
val width = the_default default_width some_width;
in
invoke_serializer thy abortable serializer literals reserved
- includes module_alias class instance tyco const module width args naming (modify program) names
+ includes module_alias class instance tyco const module_name width args
+ naming (modify program) (names, presentation_names) destination
end;
in
@@ -344,8 +348,8 @@
else Cache_IO.with_tmp_dir "Code_Test" (ext_check env_param)
end;
-fun serialize_custom thy (target_name, seri) naming program names =
- mount_serializer thy (SOME seri) target_name NONE NONE [] naming program names (String [])
+fun serialize_custom thy (target_name, seri) module_name naming program names =
+ mount_serializer thy (SOME seri) target_name NONE module_name [] naming program names (String [])
|> the;
end; (* local *)
--- a/src/Tools/Code/code_thingol.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Tools/Code/code_thingol.ML Fri Aug 27 16:05:46 2010 +0200
@@ -271,9 +271,6 @@
of SOME (tyco, _) => Codegen.thyname_of_type thy tyco
| NONE => Codegen.thyname_of_const thy c);
fun purify_base "==>" = "follows"
- | purify_base "op &" = "and"
- | purify_base "op |" = "or"
- | purify_base "op -->" = "implies"
| purify_base "op =" = "eq"
| purify_base s = Name.desymbolize false s;
fun namify thy get_basename get_thyname name =
--- a/src/Tools/jEdit/src/jedit/isabelle_hyperlinks.scala Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Tools/jEdit/src/jedit/isabelle_hyperlinks.scala Fri Aug 27 16:05:46 2010 +0200
@@ -55,11 +55,11 @@
val Text.Range(begin, end) = snapshot.convert(info_range + command_start)
val line = buffer.getLineOfOffset(begin)
- (Position.get_file(props), Position.get_line(props)) match {
+ (Position.File.unapply(props), Position.Line.unapply(props)) match {
case (Some(ref_file), Some(ref_line)) =>
new External_Hyperlink(begin, end, line, ref_file, ref_line)
case _ =>
- (Position.get_id(props), Position.get_offset(props)) match {
+ (Position.Id.unapply(props), Position.Offset.unapply(props)) match {
case (Some(ref_id), Some(ref_offset)) =>
snapshot.lookup_command(ref_id) match {
case Some(ref_cmd) =>
--- a/src/Tools/quickcheck.ML Fri Aug 27 16:04:15 2010 +0200
+++ b/src/Tools/quickcheck.ML Fri Aug 27 16:05:46 2010 +0200
@@ -78,27 +78,32 @@
datatype test_params = Test_Params of
{ size: int, iterations: int, default_type: typ list, no_assms: bool,
- expect : expectation, report: bool, quiet : bool};
+ expect : expectation, report: bool, quiet : bool};
fun dest_test_params (Test_Params { size, iterations, default_type, no_assms, expect, report, quiet }) =
((size, iterations), ((default_type, no_assms), ((expect, report), quiet)));
+
fun make_test_params ((size, iterations), ((default_type, no_assms), ((expect, report), quiet))) =
Test_Params { size = size, iterations = iterations, default_type = default_type,
- no_assms = no_assms, expect = expect, report = report, quiet = quiet };
+ no_assms = no_assms, expect = expect, report = report, quiet = quiet };
+
fun map_test_params f (Test_Params { size, iterations, default_type, no_assms, expect, report, quiet }) =
make_test_params (f ((size, iterations), ((default_type, no_assms), ((expect, report), quiet))));
-fun merge_test_params (Test_Params { size = size1, iterations = iterations1, default_type = default_type1,
- no_assms = no_assms1, expect = expect1, report = report1, quiet = quiet1 },
+
+fun merge_test_params
+ (Test_Params { size = size1, iterations = iterations1, default_type = default_type1,
+ no_assms = no_assms1, expect = expect1, report = report1, quiet = quiet1 },
Test_Params { size = size2, iterations = iterations2, default_type = default_type2,
- no_assms = no_assms2, expect = expect2, report = report2, quiet = quiet2 }) =
+ no_assms = no_assms2, expect = expect2, report = report2, quiet = quiet2 }) =
make_test_params ((Int.max (size1, size2), Int.max (iterations1, iterations2)),
((merge (op =) (default_type1, default_type2), no_assms1 orelse no_assms2),
((merge_expectation (expect1, expect2), report1 orelse report2), quiet1 orelse quiet2)));
structure Data = Theory_Data
(
- type T = (string * (Proof.context -> bool -> term -> int -> term list option * (bool list * bool))) list
- * test_params;
+ type T =
+ (string * (Proof.context -> bool -> term -> int -> term list option * (bool list * bool))) list
+ * test_params;
val empty = ([], Test_Params
{ size = 10, iterations = 100, default_type = [], no_assms = false,
expect = No_Expectation, report = false, quiet = false});