--- a/CONTRIBUTORS Mon Jun 21 11:37:25 2010 +0200
+++ b/CONTRIBUTORS Mon Jun 21 17:41:57 2010 +0200
@@ -3,6 +3,10 @@
who is listed as an author in one of the source files of this Isabelle
distribution.
+Contributions to this Isabelle version
+--------------------------------------
+
+
Contributions to Isabelle2009-2
--------------------------------------
--- a/NEWS Mon Jun 21 11:37:25 2010 +0200
+++ b/NEWS Mon Jun 21 17:41:57 2010 +0200
@@ -1,6 +1,48 @@
Isabelle NEWS -- history user-relevant changes
==============================================
+New in this Isabelle version
+----------------------------
+
+*** HOL ***
+
+* Some previously unqualified names have been qualified:
+
+ types
+ nat ~> Nat.nat
+ * ~> Product_Type,*
+ + ~> Sum_Type.+
+
+ constants
+ Ball ~> Set.Ball
+ Bex ~> Set.Bex
+ Suc ~> Nat.Suc
+ Pair ~> Product_Type.Pair
+ fst ~> Product_Type.fst
+ snd ~> Product_Type.snd
+ split ~> Product_Type.split
+ curry ~> Product_Type.curry
+
+INCOMPATIBILITY.
+
+* Removed simplifier congruence rule of "prod_case", as has for long
+been the case with "split".
+
+* Datatype package: theorems generated for executable equality
+(class eq) carry proper names and are treated as default code
+equations.
+
+* List.thy: use various operations from the Haskell prelude when
+generating Haskell code.
+
+* code_simp.ML: simplification with rules determined by
+code generator.
+
+* code generator: do not print function definitions for case
+combinators any longer.
+
+
+
New in Isabelle2009-2 (June 2010)
---------------------------------
--- a/doc-src/Codegen/Thy/Further.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/doc-src/Codegen/Thy/Further.thy Mon Jun 21 17:41:57 2010 +0200
@@ -12,6 +12,87 @@
contains exhaustive syntax diagrams.
*}
+subsection {* Locales and interpretation *}
+
+text {*
+ A technical issue comes to surface when generating code from
+ specifications stemming from locale interpretation.
+
+ Let us assume a locale specifying a power operation
+ on arbitrary types:
+*}
+
+locale %quote power =
+ fixes power :: "'a \<Rightarrow> 'b \<Rightarrow> 'b"
+ assumes power_commute: "power x \<circ> power y = power y \<circ> power x"
+begin
+
+text {*
+ \noindent Inside that locale we can lift @{text power} to exponent lists
+ by means of specification relative to that locale:
+*}
+
+primrec %quote powers :: "'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
+ "powers [] = id"
+| "powers (x # xs) = power x \<circ> powers xs"
+
+lemma %quote powers_append:
+ "powers (xs @ ys) = powers xs \<circ> powers ys"
+ by (induct xs) simp_all
+
+lemma %quote powers_power:
+ "powers xs \<circ> power x = power x \<circ> powers xs"
+ by (induct xs)
+ (simp_all del: o_apply id_apply add: o_assoc [symmetric],
+ simp del: o_apply add: o_assoc power_commute)
+
+lemma %quote powers_rev:
+ "powers (rev xs) = powers xs"
+ by (induct xs) (simp_all add: powers_append powers_power)
+
+end %quote
+
+text {*
+ After an interpretation of this locale (say, @{command
+ interpretation} @{text "fun_power:"} @{term [source] "power (\<lambda>n (f ::
+ 'a \<Rightarrow> 'a). f ^^ n)"}), one would expect to have a constant @{text
+ "fun_power.powers :: nat list \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"} for which code
+ can be generated. But this not the case: internally, the term
+ @{text "fun_power.powers"} is an abbreviation for the foundational
+ term @{term [source] "power.powers (\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n)"}
+ (see \cite{isabelle-locale} for the details behind).
+
+ Furtunately, with minor effort the desired behaviour can be achieved.
+ First, a dedicated definition of the constant on which the local @{text "powers"}
+ after interpretation is supposed to be mapped on:
+*}
+
+definition %quote funpows :: "nat list \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where
+ [code del]: "funpows = power.powers (\<lambda>n f. f ^^ n)"
+
+text {*
+ \noindent In general, the pattern is @{text "c = t"} where @{text c} is
+ the name of the future constant and @{text t} the foundational term
+ corresponding to the local constant after interpretation.
+
+ The interpretation itself is enriched with an equation @{text "t = c"}:
+*}
+
+interpretation %quote fun_power: power "\<lambda>n (f :: 'a \<Rightarrow> 'a). f ^^ n" where
+ "power.powers (\<lambda>n f. f ^^ n) = funpows"
+ by unfold_locales
+ (simp_all add: expand_fun_eq funpow_mult mult_commute funpows_def)
+
+text {*
+ \noindent This additional equation is trivially proved by the definition
+ itself.
+
+ After this setup procedure, code generation can continue as usual:
+*}
+
+text %quote {*@{code_stmts funpows (consts) Nat.funpow funpows (Haskell)}*}
+
+
subsection {* Modules *}
text {*
--- a/doc-src/Codegen/Thy/Program.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/doc-src/Codegen/Thy/Program.thy Mon Jun 21 17:41:57 2010 +0200
@@ -134,7 +134,8 @@
text {*
\noindent This is a convenient place to show how explicit dictionary construction
- manifests in generated code (here, the same example in @{text SML}):
+ manifests in generated code (here, the same example in @{text SML})
+ \cite{Haftmann-Nipkow:2010:code}:
*}
text %quote {*@{code_stmts bexp (SML)}*}
--- a/doc-src/Codegen/Thy/document/Further.tex Mon Jun 21 11:37:25 2010 +0200
+++ b/doc-src/Codegen/Thy/document/Further.tex Mon Jun 21 17:41:57 2010 +0200
@@ -33,6 +33,170 @@
\end{isamarkuptext}%
\isamarkuptrue%
%
+\isamarkupsubsection{Locales and interpretation%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+A technical issue comes to surface when generating code from
+ specifications stemming from locale interpretation.
+
+ Let us assume a locale specifying a power operation
+ on arbitrary types:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\isatagquote
+\isacommand{locale}\isamarkupfalse%
+\ power\ {\isacharequal}\isanewline
+\ \ \isakeyword{fixes}\ power\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ {\isasymRightarrow}\ {\isacharprime}b{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{assumes}\ power{\isacharunderscore}commute{\isacharcolon}\ {\isachardoublequoteopen}power\ x\ {\isasymcirc}\ power\ y\ {\isacharequal}\ power\ y\ {\isasymcirc}\ power\ x{\isachardoublequoteclose}\isanewline
+\isakeyword{begin}%
+\endisatagquote
+{\isafoldquote}%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\begin{isamarkuptext}%
+\noindent Inside that locale we can lift \isa{power} to exponent lists
+ by means of specification relative to that locale:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\isatagquote
+\isacommand{primrec}\isamarkupfalse%
+\ powers\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}b\ {\isasymRightarrow}\ {\isacharprime}b{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}powers\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ id{\isachardoublequoteclose}\isanewline
+{\isacharbar}\ {\isachardoublequoteopen}powers\ {\isacharparenleft}x\ {\isacharhash}\ xs{\isacharparenright}\ {\isacharequal}\ power\ x\ {\isasymcirc}\ powers\ xs{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ powers{\isacharunderscore}append{\isacharcolon}\isanewline
+\ \ {\isachardoublequoteopen}powers\ {\isacharparenleft}xs\ {\isacharat}\ ys{\isacharparenright}\ {\isacharequal}\ powers\ xs\ {\isasymcirc}\ powers\ ys{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}induct\ xs{\isacharparenright}\ simp{\isacharunderscore}all\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ powers{\isacharunderscore}power{\isacharcolon}\isanewline
+\ \ {\isachardoublequoteopen}powers\ xs\ {\isasymcirc}\ power\ x\ {\isacharequal}\ power\ x\ {\isasymcirc}\ powers\ xs{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}induct\ xs{\isacharparenright}\isanewline
+\ \ \ \ {\isacharparenleft}simp{\isacharunderscore}all\ del{\isacharcolon}\ o{\isacharunderscore}apply\ id{\isacharunderscore}apply\ add{\isacharcolon}\ o{\isacharunderscore}assoc\ {\isacharbrackleft}symmetric{\isacharbrackright}{\isacharcomma}\isanewline
+\ \ \ \ \ \ simp\ del{\isacharcolon}\ o{\isacharunderscore}apply\ add{\isacharcolon}\ o{\isacharunderscore}assoc\ power{\isacharunderscore}commute{\isacharparenright}\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ powers{\isacharunderscore}rev{\isacharcolon}\isanewline
+\ \ {\isachardoublequoteopen}powers\ {\isacharparenleft}rev\ xs{\isacharparenright}\ {\isacharequal}\ powers\ xs{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}induct\ xs{\isacharparenright}\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ powers{\isacharunderscore}append\ powers{\isacharunderscore}power{\isacharparenright}\isanewline
+\isanewline
+\isacommand{end}\isamarkupfalse%
+%
+\endisatagquote
+{\isafoldquote}%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\begin{isamarkuptext}%
+After an interpretation of this locale (say, \hyperlink{command.interpretation}{\mbox{\isa{\isacommand{interpretation}}}} \isa{fun{\isacharunderscore}power{\isacharcolon}} \isa{{\isachardoublequote}power\ {\isacharparenleft}{\isasymlambda}n\ {\isacharparenleft}f\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}{\isachardot}\ f\ {\isacharcircum}{\isacharcircum}\ n{\isacharparenright}{\isachardoublequote}}), one would expect to have a constant \isa{fun{\isacharunderscore}power{\isachardot}powers\ {\isacharcolon}{\isacharcolon}\ nat\ list\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} for which code
+ can be generated. But this not the case: internally, the term
+ \isa{fun{\isacharunderscore}power{\isachardot}powers} is an abbreviation for the foundational
+ term \isa{{\isachardoublequote}power{\isachardot}powers\ {\isacharparenleft}{\isasymlambda}n\ {\isacharparenleft}f\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}{\isachardot}\ f\ {\isacharcircum}{\isacharcircum}\ n{\isacharparenright}{\isachardoublequote}}
+ (see \cite{isabelle-locale} for the details behind).
+
+ Furtunately, with minor effort the desired behaviour can be achieved.
+ First, a dedicated definition of the constant on which the local \isa{powers}
+ after interpretation is supposed to be mapped on:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\isatagquote
+\isacommand{definition}\isamarkupfalse%
+\ funpows\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequoteopen}nat\ list\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
+\ \ {\isacharbrackleft}code\ del{\isacharbrackright}{\isacharcolon}\ {\isachardoublequoteopen}funpows\ {\isacharequal}\ power{\isachardot}powers\ {\isacharparenleft}{\isasymlambda}n\ f{\isachardot}\ f\ {\isacharcircum}{\isacharcircum}\ n{\isacharparenright}{\isachardoublequoteclose}%
+\endisatagquote
+{\isafoldquote}%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\begin{isamarkuptext}%
+\noindent In general, the pattern is \isa{c\ {\isacharequal}\ t} where \isa{c} is
+ the name of the future constant and \isa{t} the foundational term
+ corresponding to the local constant after interpretation.
+
+ The interpretation itself is enriched with an equation \isa{t\ {\isacharequal}\ c}:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\isatagquote
+\isacommand{interpretation}\isamarkupfalse%
+\ fun{\isacharunderscore}power{\isacharcolon}\ power\ {\isachardoublequoteopen}{\isasymlambda}n\ {\isacharparenleft}f\ {\isacharcolon}{\isacharcolon}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}{\isachardot}\ f\ {\isacharcircum}{\isacharcircum}\ n{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}power{\isachardot}powers\ {\isacharparenleft}{\isasymlambda}n\ f{\isachardot}\ f\ {\isacharcircum}{\isacharcircum}\ n{\isacharparenright}\ {\isacharequal}\ funpows{\isachardoublequoteclose}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ unfold{\isacharunderscore}locales\isanewline
+\ \ \ \ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ expand{\isacharunderscore}fun{\isacharunderscore}eq\ funpow{\isacharunderscore}mult\ mult{\isacharunderscore}commute\ funpows{\isacharunderscore}def{\isacharparenright}%
+\endisatagquote
+{\isafoldquote}%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\begin{isamarkuptext}%
+\noindent This additional equation is trivially proved by the definition
+ itself.
+
+ After this setup procedure, code generation can continue as usual:%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
+\isatagquote
+%
+\begin{isamarkuptext}%
+\isatypewriter%
+\noindent%
+\hspace*{0pt}funpow ::~forall a.~Nat -> (a -> a) -> a -> a;\\
+\hspace*{0pt}funpow Zero{\char95}nat f = id;\\
+\hspace*{0pt}funpow (Suc n) f = f .~funpow n f;\\
+\hspace*{0pt}\\
+\hspace*{0pt}funpows ::~forall a.~[Nat] -> (a -> a) -> a -> a;\\
+\hspace*{0pt}funpows [] = id;\\
+\hspace*{0pt}funpows (x :~xs) = funpow x .~funpows xs;%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\endisatagquote
+{\isafoldquote}%
+%
+\isadelimquote
+%
+\endisadelimquote
+%
\isamarkupsubsection{Modules%
}
\isamarkuptrue%
--- a/doc-src/Codegen/Thy/document/Introduction.tex Mon Jun 21 11:37:25 2010 +0200
+++ b/doc-src/Codegen/Thy/document/Introduction.tex Mon Jun 21 17:41:57 2010 +0200
@@ -234,13 +234,6 @@
\noindent%
\hspace*{0pt}module Example where {\char123}\\
\hspace*{0pt}\\
-\hspace*{0pt}foldla ::~forall a b.~(a -> b -> a) -> a -> [b] -> a;\\
-\hspace*{0pt}foldla f a [] = a;\\
-\hspace*{0pt}foldla f a (x :~xs) = foldla f (f a x) xs;\\
-\hspace*{0pt}\\
-\hspace*{0pt}rev ::~forall a.~[a] -> [a];\\
-\hspace*{0pt}rev xs = foldla ({\char92}~xsa x -> x :~xsa) [] xs;\\
-\hspace*{0pt}\\
\hspace*{0pt}list{\char95}case ::~forall a b.~a -> (b -> [b] -> a) -> [b] -> a;\\
\hspace*{0pt}list{\char95}case f1 f2 (a :~list) = f2 a list;\\
\hspace*{0pt}list{\char95}case f1 f2 [] = f1;\\
--- a/doc-src/Codegen/Thy/document/Program.tex Mon Jun 21 11:37:25 2010 +0200
+++ b/doc-src/Codegen/Thy/document/Program.tex Mon Jun 21 17:41:57 2010 +0200
@@ -323,7 +323,8 @@
%
\begin{isamarkuptext}%
\noindent This is a convenient place to show how explicit dictionary construction
- manifests in generated code (here, the same example in \isa{SML}):%
+ manifests in generated code (here, the same example in \isa{SML})
+ \cite{Haftmann-Nipkow:2010:code}:%
\end{isamarkuptext}%
\isamarkuptrue%
%
--- a/doc-src/Codegen/Thy/examples/Example.hs Mon Jun 21 11:37:25 2010 +0200
+++ b/doc-src/Codegen/Thy/examples/Example.hs Mon Jun 21 17:41:57 2010 +0200
@@ -2,13 +2,6 @@
module Example where {
-foldla :: forall a b. (a -> b -> a) -> a -> [b] -> a;
-foldla f a [] = a;
-foldla f a (x : xs) = foldla f (f a x) xs;
-
-rev :: forall a. [a] -> [a];
-rev xs = foldla (\ xsa x -> x : xsa) [] xs;
-
list_case :: forall a b. a -> (b -> [b] -> a) -> [b] -> a;
list_case f1 f2 (a : list) = f2 a list;
list_case f1 f2 [] = f1;
--- a/doc-src/IsarRef/Thy/HOL_Specific.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/doc-src/IsarRef/Thy/HOL_Specific.thy Mon Jun 21 17:41:57 2010 +0200
@@ -864,9 +864,10 @@
Internally, the evaluation is performed by registered evaluators,
which are invoked sequentially until a result is returned.
Alternatively a specific evaluator can be selected using square
- brackets; available evaluators include @{text nbe} for
- \emph{normalization by evaluation} and \emph{code} for code
- generation in SML.
+ brackets; typical evaluators use the current set of code equations
+ to normalize and include @{text simp} for fully symbolic evaluation
+ using the simplifier, @{text nbe} for \emph{normalization by evaluation}
+ and \emph{code} for code generation in SML.
\item @{command (HOL) "quickcheck"} tests the current goal for
counter examples using a series of arbitrary assignments for its
@@ -965,7 +966,214 @@
instantiates these mechanisms in a way that is amenable to end-user
applications.
- One framework generates code from both functional and relational
+ \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}.
+ 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
+ \emph{target language}. Inductive specifications can be executed
+ using the predicate compiler which operates within HOL.
+ See \cite{isabelle-codegen} for an introduction.
+
+ \begin{matharray}{rcl}
+ @{command_def (HOL) "export_code"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
+ @{attribute_def (HOL) code} & : & @{text attribute} \\
+ @{command_def (HOL) "code_abort"} & : & @{text "theory \<rightarrow> theory"} \\
+ @{command_def (HOL) "code_datatype"} & : & @{text "theory \<rightarrow> theory"} \\
+ @{command_def (HOL) "print_codesetup"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
+ @{attribute_def (HOL) code_inline} & : & @{text attribute} \\
+ @{attribute_def (HOL) code_post} & : & @{text attribute} \\
+ @{command_def (HOL) "print_codeproc"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
+ @{command_def (HOL) "code_thms"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
+ @{command_def (HOL) "code_deps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
+ @{command_def (HOL) "code_const"} & : & @{text "theory \<rightarrow> theory"} \\
+ @{command_def (HOL) "code_type"} & : & @{text "theory \<rightarrow> theory"} \\
+ @{command_def (HOL) "code_class"} & : & @{text "theory \<rightarrow> theory"} \\
+ @{command_def (HOL) "code_instance"} & : & @{text "theory \<rightarrow> theory"} \\
+ @{command_def (HOL) "code_reserved"} & : & @{text "theory \<rightarrow> theory"} \\
+ @{command_def (HOL) "code_monad"} & : & @{text "theory \<rightarrow> theory"} \\
+ @{command_def (HOL) "code_include"} & : & @{text "theory \<rightarrow> theory"} \\
+ @{command_def (HOL) "code_modulename"} & : & @{text "theory \<rightarrow> theory"} \\
+ \end{matharray}
+
+ \begin{rail}
+ 'export\_code' ( constexpr + ) \\
+ ( ( 'in' target ( 'module\_name' string ) ? \\
+ ( 'file' ( string | '-' ) ) ? ( '(' args ')' ) ?) + ) ?
+ ;
+
+ const: term
+ ;
+
+ constexpr: ( const | 'name.*' | '*' )
+ ;
+
+ typeconstructor: nameref
+ ;
+
+ class: nameref
+ ;
+
+ target: 'OCaml' | 'SML' | 'Haskell'
+ ;
+
+ 'code' ( 'del' ) ?
+ ;
+
+ 'code\_abort' ( const + )
+ ;
+
+ 'code\_datatype' ( const + )
+ ;
+
+ 'code_inline' ( 'del' ) ?
+ ;
+
+ 'code_post' ( 'del' ) ?
+ ;
+
+ 'code\_thms' ( constexpr + ) ?
+ ;
+
+ 'code\_deps' ( constexpr + ) ?
+ ;
+
+ 'code\_const' (const + 'and') \\
+ ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
+ ;
+
+ 'code\_type' (typeconstructor + 'and') \\
+ ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
+ ;
+
+ 'code\_class' (class + 'and') \\
+ ( ( '(' target \\ ( string ? + 'and' ) ')' ) + )
+ ;
+
+ 'code\_instance' (( typeconstructor '::' class ) + 'and') \\
+ ( ( '(' target ( '-' ? + 'and' ) ')' ) + )
+ ;
+
+ 'code\_reserved' target ( string + )
+ ;
+
+ 'code\_monad' const const target
+ ;
+
+ 'code\_include' target ( string ( string | '-') )
+ ;
+
+ 'code\_modulename' target ( ( string string ) + )
+ ;
+
+ syntax: string | ( 'infix' | 'infixl' | 'infixr' ) nat string
+ ;
+
+ \end{rail}
+
+ \begin{description}
+
+ \item @{command (HOL) "export_code"} generates code for a given list
+ of constants in the specified target language(s). If no serialization
+ instruction is given, only abstract code is generated internally.
+
+ Constants may be specified by giving them literally, referring to
+ all executable contants within a certain theory by giving @{text
+ "name.*"}, or referring to \emph{all} executable constants currently
+ available by giving @{text "*"}.
+
+ By default, for each involved theory one corresponding name space
+ module is generated. Alternativly, a module name may be specified
+ 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
+ single file; for \emph{Haskell}, it refers to a whole directory,
+ where code is generated in multiple files reflecting the module
+ hierarchy. The file specification ``@{text "-"}'' denotes standard
+ output. For \emph{SML}, omitting the file specification compiles
+ code internally in the context of the current ML session.
+
+ Serializers take an optional list of arguments in parentheses. For
+ \emph{SML} and \emph{OCaml}, ``@{text no_signatures}`` omits
+ explicit module signatures.
+
+ For \emph{Haskell} a module name prefix may be given using the ``@{text
+ "root:"}'' argument; ``@{text string_classes}'' adds a ``@{verbatim
+ "deriving (Read, Show)"}'' clause to each appropriate datatype
+ declaration.
+
+ \item @{attribute (HOL) code} explicitly selects (or with option
+ ``@{text "del"}'' deselects) a code equation for code
+ generation. Usually packages introducing code equations provide
+ a reasonable default setup for selection.
+
+ \item @{command (HOL) "code_abort"} declares constants which are not
+ required to have a definition by means of code equations; if
+ needed these are implemented by program abort instead.
+
+ \item @{command (HOL) "code_datatype"} specifies a constructor set
+ for a logical type.
+
+ \item @{command (HOL) "print_codesetup"} gives an overview on
+ selected code equations and code generator datatypes.
+
+ \item @{attribute (HOL) code_inline} declares (or with
+ option ``@{text "del"}'' removes) inlining theorems which are
+ applied as rewrite rules to any code equation during
+ preprocessing.
+
+ \item @{attribute (HOL) code_post} declares (or with
+ option ``@{text "del"}'' removes) theorems which are
+ applied as rewrite rules to any result of an evaluation.
+
+ \item @{command (HOL) "print_codeproc"} prints the setup
+ of the code generator preprocessor.
+
+ \item @{command (HOL) "code_thms"} prints a list of theorems
+ representing the corresponding program containing all given
+ constants after preprocessing.
+
+ \item @{command (HOL) "code_deps"} visualizes dependencies of
+ theorems representing the corresponding program containing all given
+ constants after preprocessing.
+
+ \item @{command (HOL) "code_const"} associates a list of constants
+ with target-specific serializations; omitting a serialization
+ deletes an existing serialization.
+
+ \item @{command (HOL) "code_type"} associates a list of type
+ constructors with target-specific serializations; omitting a
+ serialization deletes an existing serialization.
+
+ \item @{command (HOL) "code_class"} associates a list of classes
+ with target-specific class names; omitting a serialization deletes
+ an existing serialization. This applies only to \emph{Haskell}.
+
+ \item @{command (HOL) "code_instance"} declares a list of type
+ constructor / class instance relations as ``already present'' for a
+ given target. Omitting a ``@{text "-"}'' deletes an existing
+ ``already present'' declaration. This applies only to
+ \emph{Haskell}.
+
+ \item @{command (HOL) "code_reserved"} declares a list of names as
+ reserved for a given target, preventing it to be shadowed by any
+ generated code.
+
+ \item @{command (HOL) "code_monad"} provides an auxiliary mechanism
+ to generate monadic code for Haskell.
+
+ \item @{command (HOL) "code_include"} adds arbitrary named content
+ (``include'') to generated code. A ``@{text "-"}'' as last argument
+ will remove an already added ``include''.
+
+ \item @{command (HOL) "code_modulename"} declares aliasings from one
+ module name onto another.
+
+ \end{description}
+
+ The other framework generates code from both functional and relational
programs to SML. See \cite{isabelle-HOL} for further information
(this actually covers the new-style theory format as well).
@@ -1011,209 +1219,6 @@
;
\end{rail}
- \medskip The other framework generates code from functional programs
- (including overloading using type classes) to SML \cite{SML}, OCaml
- \cite{OCaml} and Haskell \cite{haskell-revised-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
- \emph{target language}. See \cite{isabelle-codegen} for an
- introduction on how to use it.
-
- \begin{matharray}{rcl}
- @{command_def (HOL) "export_code"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
- @{command_def (HOL) "code_thms"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
- @{command_def (HOL) "code_deps"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
- @{command_def (HOL) "code_datatype"} & : & @{text "theory \<rightarrow> theory"} \\
- @{command_def (HOL) "code_const"} & : & @{text "theory \<rightarrow> theory"} \\
- @{command_def (HOL) "code_type"} & : & @{text "theory \<rightarrow> theory"} \\
- @{command_def (HOL) "code_class"} & : & @{text "theory \<rightarrow> theory"} \\
- @{command_def (HOL) "code_instance"} & : & @{text "theory \<rightarrow> theory"} \\
- @{command_def (HOL) "code_monad"} & : & @{text "theory \<rightarrow> theory"} \\
- @{command_def (HOL) "code_reserved"} & : & @{text "theory \<rightarrow> theory"} \\
- @{command_def (HOL) "code_include"} & : & @{text "theory \<rightarrow> theory"} \\
- @{command_def (HOL) "code_modulename"} & : & @{text "theory \<rightarrow> theory"} \\
- @{command_def (HOL) "code_abort"} & : & @{text "theory \<rightarrow> theory"} \\
- @{command_def (HOL) "print_codesetup"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
- @{command_def (HOL) "print_codeproc"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
- @{attribute_def (HOL) code} & : & @{text attribute} \\
- \end{matharray}
-
- \begin{rail}
- 'export\_code' ( constexpr + ) \\
- ( ( 'in' target ( 'module\_name' string ) ? \\
- ( 'file' ( string | '-' ) ) ? ( '(' args ')' ) ?) + ) ?
- ;
-
- 'code\_thms' ( constexpr + ) ?
- ;
-
- 'code\_deps' ( constexpr + ) ?
- ;
-
- const: term
- ;
-
- constexpr: ( const | 'name.*' | '*' )
- ;
-
- typeconstructor: nameref
- ;
-
- class: nameref
- ;
-
- target: 'OCaml' | 'SML' | 'Haskell'
- ;
-
- 'code\_datatype' ( const + )
- ;
-
- 'code\_const' (const + 'and') \\
- ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
- ;
-
- 'code\_type' (typeconstructor + 'and') \\
- ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
- ;
-
- 'code\_class' (class + 'and') \\
- ( ( '(' target \\ ( string ? + 'and' ) ')' ) + )
- ;
-
- 'code\_instance' (( typeconstructor '::' class ) + 'and') \\
- ( ( '(' target ( '-' ? + 'and' ) ')' ) + )
- ;
-
- 'code\_monad' const const target
- ;
-
- 'code\_reserved' target ( string + )
- ;
-
- 'code\_include' target ( string ( string | '-') )
- ;
-
- 'code\_modulename' target ( ( string string ) + )
- ;
-
- 'code\_abort' ( const + )
- ;
-
- syntax: string | ( 'infix' | 'infixl' | 'infixr' ) nat string
- ;
-
- 'code' ( 'del' ) ?
- ;
-
- 'code_unfold' ( 'del' ) ?
- ;
-
- 'code_post' ( 'del' ) ?
- ;
- \end{rail}
-
- \begin{description}
-
- \item @{command (HOL) "export_code"} is the canonical interface for
- generating and serializing code: for a given list of constants, code
- is generated for the specified target languages. If no serialization
- instruction is given, only abstract code is generated internally.
-
- Constants may be specified by giving them literally, referring to
- all executable contants within a certain theory by giving @{text
- "name.*"}, or referring to \emph{all} executable constants currently
- available by giving @{text "*"}.
-
- By default, for each involved theory one corresponding name space
- module is generated. Alternativly, a module name may be specified
- 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
- single file; for \emph{Haskell}, it refers to a whole directory,
- where code is generated in multiple files reflecting the module
- hierarchy. The file specification ``@{text "-"}'' denotes standard
- output. For \emph{SML}, omitting the file specification compiles
- code internally in the context of the current ML session.
-
- Serializers take an optional list of arguments in parentheses. For
- \emph{SML} and \emph{OCaml}, ``@{text no_signatures}`` omits
- explicit module signatures.
-
- For \emph{Haskell} a module name prefix may be given using the ``@{text
- "root:"}'' argument; ``@{text string_classes}'' adds a ``@{verbatim
- "deriving (Read, Show)"}'' clause to each appropriate datatype
- declaration.
-
- \item @{command (HOL) "code_thms"} prints a list of theorems
- representing the corresponding program containing all given
- constants.
-
- \item @{command (HOL) "code_deps"} visualizes dependencies of
- theorems representing the corresponding program containing all given
- constants.
-
- \item @{command (HOL) "code_datatype"} specifies a constructor set
- for a logical type.
-
- \item @{command (HOL) "code_const"} associates a list of constants
- with target-specific serializations; omitting a serialization
- deletes an existing serialization.
-
- \item @{command (HOL) "code_type"} associates a list of type
- constructors with target-specific serializations; omitting a
- serialization deletes an existing serialization.
-
- \item @{command (HOL) "code_class"} associates a list of classes
- with target-specific class names; omitting a serialization deletes
- an existing serialization. This applies only to \emph{Haskell}.
-
- \item @{command (HOL) "code_instance"} declares a list of type
- constructor / class instance relations as ``already present'' for a
- given target. Omitting a ``@{text "-"}'' deletes an existing
- ``already present'' declaration. This applies only to
- \emph{Haskell}.
-
- \item @{command (HOL) "code_monad"} provides an auxiliary mechanism
- to generate monadic code for Haskell.
-
- \item @{command (HOL) "code_reserved"} declares a list of names as
- reserved for a given target, preventing it to be shadowed by any
- generated code.
-
- \item @{command (HOL) "code_include"} adds arbitrary named content
- (``include'') to generated code. A ``@{text "-"}'' as last argument
- will remove an already added ``include''.
-
- \item @{command (HOL) "code_modulename"} declares aliasings from one
- module name onto another.
-
- \item @{command (HOL) "code_abort"} declares constants which are not
- required to have a definition by means of code equations; if
- needed these are implemented by program abort instead.
-
- \item @{attribute (HOL) code} explicitly selects (or with option
- ``@{text "del"}'' deselects) a code equation for code
- generation. Usually packages introducing code equations provide
- a reasonable default setup for selection.
-
- \item @{attribute (HOL) code_inline} declares (or with
- option ``@{text "del"}'' removes) inlining theorems which are
- applied as rewrite rules to any code equation during
- preprocessing.
-
- \item @{attribute (HOL) code_post} declares (or with
- option ``@{text "del"}'' removes) theorems which are
- applied as rewrite rules to any result of an evaluation.
-
- \item @{command (HOL) "print_codesetup"} gives an overview on
- selected code equations and code generator datatypes.
-
- \item @{command (HOL) "print_codeproc"} prints the setup
- of the code generator preprocessor.
-
- \end{description}
*}
--- a/doc-src/IsarRef/Thy/document/HOL_Specific.tex Mon Jun 21 11:37:25 2010 +0200
+++ b/doc-src/IsarRef/Thy/document/HOL_Specific.tex Mon Jun 21 17:41:57 2010 +0200
@@ -882,9 +882,10 @@
Internally, the evaluation is performed by registered evaluators,
which are invoked sequentially until a result is returned.
Alternatively a specific evaluator can be selected using square
- brackets; available evaluators include \isa{nbe} for
- \emph{normalization by evaluation} and \emph{code} for code
- generation in SML.
+ brackets; typical evaluators use the current set of code equations
+ to normalize and include \isa{simp} for fully symbolic evaluation
+ using the simplifier, \isa{nbe} for \emph{normalization by evaluation}
+ and \emph{code} for code generation in SML.
\item \hyperlink{command.HOL.quickcheck}{\mbox{\isa{\isacommand{quickcheck}}}} tests the current goal for
counter examples using a series of arbitrary assignments for its
@@ -981,7 +982,211 @@
instantiates these mechanisms in a way that is amenable to end-user
applications.
- One framework generates code from both functional and relational
+ \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}.
+ 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
+ \emph{target language}. Inductive specifications can be executed
+ using the predicate compiler which operates within HOL.
+ See \cite{isabelle-codegen} for an introduction.
+
+ \begin{matharray}{rcl}
+ \indexdef{HOL}{command}{export\_code}\hypertarget{command.HOL.export-code}{\hyperlink{command.HOL.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
+ \indexdef{HOL}{attribute}{code}\hypertarget{attribute.HOL.code}{\hyperlink{attribute.HOL.code}{\mbox{\isa{code}}}} & : & \isa{attribute} \\
+ \indexdef{HOL}{command}{code\_abort}\hypertarget{command.HOL.code-abort}{\hyperlink{command.HOL.code-abort}{\mbox{\isa{\isacommand{code{\isacharunderscore}abort}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
+ \indexdef{HOL}{command}{code\_datatype}\hypertarget{command.HOL.code-datatype}{\hyperlink{command.HOL.code-datatype}{\mbox{\isa{\isacommand{code{\isacharunderscore}datatype}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
+ \indexdef{HOL}{command}{print\_codesetup}\hypertarget{command.HOL.print-codesetup}{\hyperlink{command.HOL.print-codesetup}{\mbox{\isa{\isacommand{print{\isacharunderscore}codesetup}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
+ \indexdef{HOL}{attribute}{code\_inline}\hypertarget{attribute.HOL.code-inline}{\hyperlink{attribute.HOL.code-inline}{\mbox{\isa{code{\isacharunderscore}inline}}}} & : & \isa{attribute} \\
+ \indexdef{HOL}{attribute}{code\_post}\hypertarget{attribute.HOL.code-post}{\hyperlink{attribute.HOL.code-post}{\mbox{\isa{code{\isacharunderscore}post}}}} & : & \isa{attribute} \\
+ \indexdef{HOL}{command}{print\_codeproc}\hypertarget{command.HOL.print-codeproc}{\hyperlink{command.HOL.print-codeproc}{\mbox{\isa{\isacommand{print{\isacharunderscore}codeproc}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
+ \indexdef{HOL}{command}{code\_thms}\hypertarget{command.HOL.code-thms}{\hyperlink{command.HOL.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
+ \indexdef{HOL}{command}{code\_deps}\hypertarget{command.HOL.code-deps}{\hyperlink{command.HOL.code-deps}{\mbox{\isa{\isacommand{code{\isacharunderscore}deps}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
+ \indexdef{HOL}{command}{code\_const}\hypertarget{command.HOL.code-const}{\hyperlink{command.HOL.code-const}{\mbox{\isa{\isacommand{code{\isacharunderscore}const}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
+ \indexdef{HOL}{command}{code\_type}\hypertarget{command.HOL.code-type}{\hyperlink{command.HOL.code-type}{\mbox{\isa{\isacommand{code{\isacharunderscore}type}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
+ \indexdef{HOL}{command}{code\_class}\hypertarget{command.HOL.code-class}{\hyperlink{command.HOL.code-class}{\mbox{\isa{\isacommand{code{\isacharunderscore}class}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
+ \indexdef{HOL}{command}{code\_instance}\hypertarget{command.HOL.code-instance}{\hyperlink{command.HOL.code-instance}{\mbox{\isa{\isacommand{code{\isacharunderscore}instance}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
+ \indexdef{HOL}{command}{code\_reserved}\hypertarget{command.HOL.code-reserved}{\hyperlink{command.HOL.code-reserved}{\mbox{\isa{\isacommand{code{\isacharunderscore}reserved}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
+ \indexdef{HOL}{command}{code\_monad}\hypertarget{command.HOL.code-monad}{\hyperlink{command.HOL.code-monad}{\mbox{\isa{\isacommand{code{\isacharunderscore}monad}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
+ \indexdef{HOL}{command}{code\_include}\hypertarget{command.HOL.code-include}{\hyperlink{command.HOL.code-include}{\mbox{\isa{\isacommand{code{\isacharunderscore}include}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
+ \indexdef{HOL}{command}{code\_modulename}\hypertarget{command.HOL.code-modulename}{\hyperlink{command.HOL.code-modulename}{\mbox{\isa{\isacommand{code{\isacharunderscore}modulename}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
+ \end{matharray}
+
+ \begin{rail}
+ 'export\_code' ( constexpr + ) \\
+ ( ( 'in' target ( 'module\_name' string ) ? \\
+ ( 'file' ( string | '-' ) ) ? ( '(' args ')' ) ?) + ) ?
+ ;
+
+ const: term
+ ;
+
+ constexpr: ( const | 'name.*' | '*' )
+ ;
+
+ typeconstructor: nameref
+ ;
+
+ class: nameref
+ ;
+
+ target: 'OCaml' | 'SML' | 'Haskell'
+ ;
+
+ 'code' ( 'del' ) ?
+ ;
+
+ 'code\_abort' ( const + )
+ ;
+
+ 'code\_datatype' ( const + )
+ ;
+
+ 'code_inline' ( 'del' ) ?
+ ;
+
+ 'code_post' ( 'del' ) ?
+ ;
+
+ 'code\_thms' ( constexpr + ) ?
+ ;
+
+ 'code\_deps' ( constexpr + ) ?
+ ;
+
+ 'code\_const' (const + 'and') \\
+ ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
+ ;
+
+ 'code\_type' (typeconstructor + 'and') \\
+ ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
+ ;
+
+ 'code\_class' (class + 'and') \\
+ ( ( '(' target \\ ( string ? + 'and' ) ')' ) + )
+ ;
+
+ 'code\_instance' (( typeconstructor '::' class ) + 'and') \\
+ ( ( '(' target ( '-' ? + 'and' ) ')' ) + )
+ ;
+
+ 'code\_reserved' target ( string + )
+ ;
+
+ 'code\_monad' const const target
+ ;
+
+ 'code\_include' target ( string ( string | '-') )
+ ;
+
+ 'code\_modulename' target ( ( string string ) + )
+ ;
+
+ syntax: string | ( 'infix' | 'infixl' | 'infixr' ) nat string
+ ;
+
+ \end{rail}
+
+ \begin{description}
+
+ \item \hyperlink{command.HOL.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}} generates code for a given list
+ of constants in the specified target language(s). If no serialization
+ instruction is given, only abstract code is generated internally.
+
+ Constants may be specified by giving them literally, referring to
+ all executable contants within a certain theory by giving \isa{{\isachardoublequote}name{\isachardot}{\isacharasterisk}{\isachardoublequote}}, or referring to \emph{all} executable constants currently
+ available by giving \isa{{\isachardoublequote}{\isacharasterisk}{\isachardoublequote}}.
+
+ By default, for each involved theory one corresponding name space
+ module is generated. Alternativly, a module name may be specified
+ 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
+ single file; for \emph{Haskell}, it refers to a whole directory,
+ where code is generated in multiple files reflecting the module
+ hierarchy. The file specification ``\isa{{\isachardoublequote}{\isacharminus}{\isachardoublequote}}'' denotes standard
+ output. For \emph{SML}, omitting the file specification compiles
+ code internally in the context of the current ML session.
+
+ Serializers take an optional list of arguments in parentheses. For
+ \emph{SML} and \emph{OCaml}, ``\isa{no{\isacharunderscore}signatures}`` omits
+ explicit module signatures.
+
+ For \emph{Haskell} a module name prefix may be given using the ``\isa{{\isachardoublequote}root{\isacharcolon}{\isachardoublequote}}'' argument; ``\isa{string{\isacharunderscore}classes}'' adds a ``\verb|deriving (Read, Show)|'' clause to each appropriate datatype
+ declaration.
+
+ \item \hyperlink{attribute.HOL.code}{\mbox{\isa{code}}} explicitly selects (or with option
+ ``\isa{{\isachardoublequote}del{\isachardoublequote}}'' deselects) a code equation for code
+ generation. Usually packages introducing code equations provide
+ a reasonable default setup for selection.
+
+ \item \hyperlink{command.HOL.code-abort}{\mbox{\isa{\isacommand{code{\isacharunderscore}abort}}}} declares constants which are not
+ required to have a definition by means of code equations; if
+ needed these are implemented by program abort instead.
+
+ \item \hyperlink{command.HOL.code-datatype}{\mbox{\isa{\isacommand{code{\isacharunderscore}datatype}}}} specifies a constructor set
+ for a logical type.
+
+ \item \hyperlink{command.HOL.print-codesetup}{\mbox{\isa{\isacommand{print{\isacharunderscore}codesetup}}}} gives an overview on
+ selected code equations and code generator datatypes.
+
+ \item \hyperlink{attribute.HOL.code-inline}{\mbox{\isa{code{\isacharunderscore}inline}}} declares (or with
+ option ``\isa{{\isachardoublequote}del{\isachardoublequote}}'' removes) inlining theorems which are
+ applied as rewrite rules to any code equation during
+ preprocessing.
+
+ \item \hyperlink{attribute.HOL.code-post}{\mbox{\isa{code{\isacharunderscore}post}}} declares (or with
+ option ``\isa{{\isachardoublequote}del{\isachardoublequote}}'' removes) theorems which are
+ applied as rewrite rules to any result of an evaluation.
+
+ \item \hyperlink{command.HOL.print-codeproc}{\mbox{\isa{\isacommand{print{\isacharunderscore}codeproc}}}} prints the setup
+ of the code generator preprocessor.
+
+ \item \hyperlink{command.HOL.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}} prints a list of theorems
+ representing the corresponding program containing all given
+ constants after preprocessing.
+
+ \item \hyperlink{command.HOL.code-deps}{\mbox{\isa{\isacommand{code{\isacharunderscore}deps}}}} visualizes dependencies of
+ theorems representing the corresponding program containing all given
+ constants after preprocessing.
+
+ \item \hyperlink{command.HOL.code-const}{\mbox{\isa{\isacommand{code{\isacharunderscore}const}}}} associates a list of constants
+ with target-specific serializations; omitting a serialization
+ deletes an existing serialization.
+
+ \item \hyperlink{command.HOL.code-type}{\mbox{\isa{\isacommand{code{\isacharunderscore}type}}}} associates a list of type
+ constructors with target-specific serializations; omitting a
+ serialization deletes an existing serialization.
+
+ \item \hyperlink{command.HOL.code-class}{\mbox{\isa{\isacommand{code{\isacharunderscore}class}}}} associates a list of classes
+ with target-specific class names; omitting a serialization deletes
+ an existing serialization. This applies only to \emph{Haskell}.
+
+ \item \hyperlink{command.HOL.code-instance}{\mbox{\isa{\isacommand{code{\isacharunderscore}instance}}}} declares a list of type
+ constructor / class instance relations as ``already present'' for a
+ given target. Omitting a ``\isa{{\isachardoublequote}{\isacharminus}{\isachardoublequote}}'' deletes an existing
+ ``already present'' declaration. This applies only to
+ \emph{Haskell}.
+
+ \item \hyperlink{command.HOL.code-reserved}{\mbox{\isa{\isacommand{code{\isacharunderscore}reserved}}}} declares a list of names as
+ reserved for a given target, preventing it to be shadowed by any
+ generated code.
+
+ \item \hyperlink{command.HOL.code-monad}{\mbox{\isa{\isacommand{code{\isacharunderscore}monad}}}} provides an auxiliary mechanism
+ to generate monadic code for Haskell.
+
+ \item \hyperlink{command.HOL.code-include}{\mbox{\isa{\isacommand{code{\isacharunderscore}include}}}} adds arbitrary named content
+ (``include'') to generated code. A ``\isa{{\isachardoublequote}{\isacharminus}{\isachardoublequote}}'' as last argument
+ will remove an already added ``include''.
+
+ \item \hyperlink{command.HOL.code-modulename}{\mbox{\isa{\isacommand{code{\isacharunderscore}modulename}}}} declares aliasings from one
+ module name onto another.
+
+ \end{description}
+
+ The other framework generates code from both functional and relational
programs to SML. See \cite{isabelle-HOL} for further information
(this actually covers the new-style theory format as well).
@@ -1025,208 +1230,7 @@
'code' (name)?
;
- \end{rail}
-
- \medskip The other framework generates code from functional programs
- (including overloading using type classes) to SML \cite{SML}, OCaml
- \cite{OCaml} and Haskell \cite{haskell-revised-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
- \emph{target language}. See \cite{isabelle-codegen} for an
- introduction on how to use it.
-
- \begin{matharray}{rcl}
- \indexdef{HOL}{command}{export\_code}\hypertarget{command.HOL.export-code}{\hyperlink{command.HOL.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
- \indexdef{HOL}{command}{code\_thms}\hypertarget{command.HOL.code-thms}{\hyperlink{command.HOL.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
- \indexdef{HOL}{command}{code\_deps}\hypertarget{command.HOL.code-deps}{\hyperlink{command.HOL.code-deps}{\mbox{\isa{\isacommand{code{\isacharunderscore}deps}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
- \indexdef{HOL}{command}{code\_datatype}\hypertarget{command.HOL.code-datatype}{\hyperlink{command.HOL.code-datatype}{\mbox{\isa{\isacommand{code{\isacharunderscore}datatype}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
- \indexdef{HOL}{command}{code\_const}\hypertarget{command.HOL.code-const}{\hyperlink{command.HOL.code-const}{\mbox{\isa{\isacommand{code{\isacharunderscore}const}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
- \indexdef{HOL}{command}{code\_type}\hypertarget{command.HOL.code-type}{\hyperlink{command.HOL.code-type}{\mbox{\isa{\isacommand{code{\isacharunderscore}type}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
- \indexdef{HOL}{command}{code\_class}\hypertarget{command.HOL.code-class}{\hyperlink{command.HOL.code-class}{\mbox{\isa{\isacommand{code{\isacharunderscore}class}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
- \indexdef{HOL}{command}{code\_instance}\hypertarget{command.HOL.code-instance}{\hyperlink{command.HOL.code-instance}{\mbox{\isa{\isacommand{code{\isacharunderscore}instance}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
- \indexdef{HOL}{command}{code\_monad}\hypertarget{command.HOL.code-monad}{\hyperlink{command.HOL.code-monad}{\mbox{\isa{\isacommand{code{\isacharunderscore}monad}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
- \indexdef{HOL}{command}{code\_reserved}\hypertarget{command.HOL.code-reserved}{\hyperlink{command.HOL.code-reserved}{\mbox{\isa{\isacommand{code{\isacharunderscore}reserved}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
- \indexdef{HOL}{command}{code\_include}\hypertarget{command.HOL.code-include}{\hyperlink{command.HOL.code-include}{\mbox{\isa{\isacommand{code{\isacharunderscore}include}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
- \indexdef{HOL}{command}{code\_modulename}\hypertarget{command.HOL.code-modulename}{\hyperlink{command.HOL.code-modulename}{\mbox{\isa{\isacommand{code{\isacharunderscore}modulename}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
- \indexdef{HOL}{command}{code\_abort}\hypertarget{command.HOL.code-abort}{\hyperlink{command.HOL.code-abort}{\mbox{\isa{\isacommand{code{\isacharunderscore}abort}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
- \indexdef{HOL}{command}{print\_codesetup}\hypertarget{command.HOL.print-codesetup}{\hyperlink{command.HOL.print-codesetup}{\mbox{\isa{\isacommand{print{\isacharunderscore}codesetup}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
- \indexdef{HOL}{command}{print\_codeproc}\hypertarget{command.HOL.print-codeproc}{\hyperlink{command.HOL.print-codeproc}{\mbox{\isa{\isacommand{print{\isacharunderscore}codeproc}}}}}\isa{{\isachardoublequote}\isactrlsup {\isacharasterisk}{\isachardoublequote}} & : & \isa{{\isachardoublequote}context\ {\isasymrightarrow}{\isachardoublequote}} \\
- \indexdef{HOL}{attribute}{code}\hypertarget{attribute.HOL.code}{\hyperlink{attribute.HOL.code}{\mbox{\isa{code}}}} & : & \isa{attribute} \\
- \end{matharray}
-
- \begin{rail}
- 'export\_code' ( constexpr + ) \\
- ( ( 'in' target ( 'module\_name' string ) ? \\
- ( 'file' ( string | '-' ) ) ? ( '(' args ')' ) ?) + ) ?
- ;
-
- 'code\_thms' ( constexpr + ) ?
- ;
-
- 'code\_deps' ( constexpr + ) ?
- ;
-
- const: term
- ;
-
- constexpr: ( const | 'name.*' | '*' )
- ;
-
- typeconstructor: nameref
- ;
-
- class: nameref
- ;
-
- target: 'OCaml' | 'SML' | 'Haskell'
- ;
-
- 'code\_datatype' ( const + )
- ;
-
- 'code\_const' (const + 'and') \\
- ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
- ;
-
- 'code\_type' (typeconstructor + 'and') \\
- ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
- ;
-
- 'code\_class' (class + 'and') \\
- ( ( '(' target \\ ( string ? + 'and' ) ')' ) + )
- ;
-
- 'code\_instance' (( typeconstructor '::' class ) + 'and') \\
- ( ( '(' target ( '-' ? + 'and' ) ')' ) + )
- ;
-
- 'code\_monad' const const target
- ;
-
- 'code\_reserved' target ( string + )
- ;
-
- 'code\_include' target ( string ( string | '-') )
- ;
-
- 'code\_modulename' target ( ( string string ) + )
- ;
-
- 'code\_abort' ( const + )
- ;
-
- syntax: string | ( 'infix' | 'infixl' | 'infixr' ) nat string
- ;
-
- 'code' ( 'del' ) ?
- ;
-
- 'code_unfold' ( 'del' ) ?
- ;
-
- 'code_post' ( 'del' ) ?
- ;
- \end{rail}
-
- \begin{description}
-
- \item \hyperlink{command.HOL.export-code}{\mbox{\isa{\isacommand{export{\isacharunderscore}code}}}} is the canonical interface for
- generating and serializing code: for a given list of constants, code
- is generated for the specified target languages. If no serialization
- instruction is given, only abstract code is generated internally.
-
- Constants may be specified by giving them literally, referring to
- all executable contants within a certain theory by giving \isa{{\isachardoublequote}name{\isachardot}{\isacharasterisk}{\isachardoublequote}}, or referring to \emph{all} executable constants currently
- available by giving \isa{{\isachardoublequote}{\isacharasterisk}{\isachardoublequote}}.
-
- By default, for each involved theory one corresponding name space
- module is generated. Alternativly, a module name may be specified
- 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
- single file; for \emph{Haskell}, it refers to a whole directory,
- where code is generated in multiple files reflecting the module
- hierarchy. The file specification ``\isa{{\isachardoublequote}{\isacharminus}{\isachardoublequote}}'' denotes standard
- output. For \emph{SML}, omitting the file specification compiles
- code internally in the context of the current ML session.
-
- Serializers take an optional list of arguments in parentheses. For
- \emph{SML} and \emph{OCaml}, ``\isa{no{\isacharunderscore}signatures}`` omits
- explicit module signatures.
-
- For \emph{Haskell} a module name prefix may be given using the ``\isa{{\isachardoublequote}root{\isacharcolon}{\isachardoublequote}}'' argument; ``\isa{string{\isacharunderscore}classes}'' adds a ``\verb|deriving (Read, Show)|'' clause to each appropriate datatype
- declaration.
-
- \item \hyperlink{command.HOL.code-thms}{\mbox{\isa{\isacommand{code{\isacharunderscore}thms}}}} prints a list of theorems
- representing the corresponding program containing all given
- constants.
-
- \item \hyperlink{command.HOL.code-deps}{\mbox{\isa{\isacommand{code{\isacharunderscore}deps}}}} visualizes dependencies of
- theorems representing the corresponding program containing all given
- constants.
-
- \item \hyperlink{command.HOL.code-datatype}{\mbox{\isa{\isacommand{code{\isacharunderscore}datatype}}}} specifies a constructor set
- for a logical type.
-
- \item \hyperlink{command.HOL.code-const}{\mbox{\isa{\isacommand{code{\isacharunderscore}const}}}} associates a list of constants
- with target-specific serializations; omitting a serialization
- deletes an existing serialization.
-
- \item \hyperlink{command.HOL.code-type}{\mbox{\isa{\isacommand{code{\isacharunderscore}type}}}} associates a list of type
- constructors with target-specific serializations; omitting a
- serialization deletes an existing serialization.
-
- \item \hyperlink{command.HOL.code-class}{\mbox{\isa{\isacommand{code{\isacharunderscore}class}}}} associates a list of classes
- with target-specific class names; omitting a serialization deletes
- an existing serialization. This applies only to \emph{Haskell}.
-
- \item \hyperlink{command.HOL.code-instance}{\mbox{\isa{\isacommand{code{\isacharunderscore}instance}}}} declares a list of type
- constructor / class instance relations as ``already present'' for a
- given target. Omitting a ``\isa{{\isachardoublequote}{\isacharminus}{\isachardoublequote}}'' deletes an existing
- ``already present'' declaration. This applies only to
- \emph{Haskell}.
-
- \item \hyperlink{command.HOL.code-monad}{\mbox{\isa{\isacommand{code{\isacharunderscore}monad}}}} provides an auxiliary mechanism
- to generate monadic code for Haskell.
-
- \item \hyperlink{command.HOL.code-reserved}{\mbox{\isa{\isacommand{code{\isacharunderscore}reserved}}}} declares a list of names as
- reserved for a given target, preventing it to be shadowed by any
- generated code.
-
- \item \hyperlink{command.HOL.code-include}{\mbox{\isa{\isacommand{code{\isacharunderscore}include}}}} adds arbitrary named content
- (``include'') to generated code. A ``\isa{{\isachardoublequote}{\isacharminus}{\isachardoublequote}}'' as last argument
- will remove an already added ``include''.
-
- \item \hyperlink{command.HOL.code-modulename}{\mbox{\isa{\isacommand{code{\isacharunderscore}modulename}}}} declares aliasings from one
- module name onto another.
-
- \item \hyperlink{command.HOL.code-abort}{\mbox{\isa{\isacommand{code{\isacharunderscore}abort}}}} declares constants which are not
- required to have a definition by means of code equations; if
- needed these are implemented by program abort instead.
-
- \item \hyperlink{attribute.HOL.code}{\mbox{\isa{code}}} explicitly selects (or with option
- ``\isa{{\isachardoublequote}del{\isachardoublequote}}'' deselects) a code equation for code
- generation. Usually packages introducing code equations provide
- a reasonable default setup for selection.
-
- \item \hyperlink{attribute.HOL.code-inline}{\mbox{\isa{code{\isacharunderscore}inline}}} declares (or with
- option ``\isa{{\isachardoublequote}del{\isachardoublequote}}'' removes) inlining theorems which are
- applied as rewrite rules to any code equation during
- preprocessing.
-
- \item \hyperlink{attribute.HOL.code-post}{\mbox{\isa{code{\isacharunderscore}post}}} declares (or with
- option ``\isa{{\isachardoublequote}del{\isachardoublequote}}'' removes) theorems which are
- applied as rewrite rules to any result of an evaluation.
-
- \item \hyperlink{command.HOL.print-codesetup}{\mbox{\isa{\isacommand{print{\isacharunderscore}codesetup}}}} gives an overview on
- selected code equations and code generator datatypes.
-
- \item \hyperlink{command.HOL.print-codeproc}{\mbox{\isa{\isacommand{print{\isacharunderscore}codeproc}}}} prints the setup
- of the code generator preprocessor.
-
- \end{description}%
+ \end{rail}%
\end{isamarkuptext}%
\isamarkuptrue%
%
--- a/doc-src/Sledgehammer/sledgehammer.tex Mon Jun 21 11:37:25 2010 +0200
+++ b/doc-src/Sledgehammer/sledgehammer.tex Mon Jun 21 17:41:57 2010 +0200
@@ -338,11 +338,13 @@
Weidenbach et al.\ \cite{weidenbach-et-al-2009}. To use SPASS, set the
environment variable \texttt{SPASS\_HOME} to the directory that contains the
\texttt{SPASS} executable, or install the prebuilt SPASS package from Isabelle's
-download page. See \S\ref{installation} for details.
+download page. Sledgehammer requires version 3.5 or above. See
+\S\ref{installation} for details.
-\item[$\bullet$] \textbf{\textit{spass\_tptp}:} Same as the above, except that
-Sledgehammer communicates with SPASS using the TPTP syntax rather than the
-native DFG syntax. This ATP is provided for experimental purposes.
+\item[$\bullet$] \textbf{\textit{spass\_dfg}:} Same as the above, except that
+Sledgehammer communicates with SPASS using the native DFG syntax rather than the
+TPTP syntax. Sledgehammer requires version 3.0 or above. This ATP is provided
+for compatibility reasons.
\item[$\bullet$] \textbf{\textit{vampire}:} Vampire is an ATP developed by
Andrei Voronkov and his colleagues \cite{riazanov-voronkov-2002}. To use
--- a/doc-src/manual.bib Mon Jun 21 11:37:25 2010 +0200
+++ b/doc-src/manual.bib Mon Jun 21 17:41:57 2010 +0200
@@ -130,6 +130,13 @@
@book{Baader-Nipkow,author={Franz Baader and Tobias Nipkow},
title="Term Rewriting and All That",publisher=CUP,year=1998}
+@manual{isabelle-locale,
+ author = {Clemens Ballarin},
+ title = {Tutorial to Locales and Locale Interpretation},
+ institution = TUM,
+ note = {\url{http://isabelle.in.tum.de/doc/locales.pdf}}
+}
+
@InCollection{Barendregt-Geuvers:2001,
author = {H. Barendregt and H. Geuvers},
title = {Proof Assistants using Dependent Type Systems},
@@ -508,15 +515,15 @@
year = {2007}
}
-@TechReport{Haftmann-Nipkow:2007:codegen,
- author = {Florian Haftmann and Tobias Nipkow},
- title = {A Code Generator Framework for {Isabelle/HOL}},
- editor = {Klaus Schneider and Jens Brandt},
- booktitle = {Theorem Proving in Higher Order Logics: Emerging Trends Proceedings},
- month = {08},
- year = {2007},
- institution = {Department of Computer Science, University of Kaiserslautern},
- number = {364/07}
+@inproceedings{Haftmann-Nipkow:2010:code,
+ author = {Florian Haftmann and Tobias Nipkow},
+ title = {Code Generation via Higher-Order Rewrite Systems},
+ booktitle = {Functional and Logic Programming: 10th International Symposium: FLOPS 2010},
+ year = 2010,
+ publisher = {Springer},
+ series = {Lecture Notes in Computer Science},
+ editor = {Matthias Blume and Naoki Kobayashi and Germ{\'a}n Vidal},
+ volume = {6009}
}
@InProceedings{Haftmann-Wenzel:2009,
--- a/src/HOL/Algebra/README.html Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Algebra/README.html Mon Jun 21 17:41:57 2010 +0200
@@ -20,7 +20,7 @@
<H2>GroupTheory, including Sylow's Theorem</H2>
-<P>These proofs are mainly by Florian Kammller. (Later, Larry
+<P>These proofs are mainly by Florian Kammüller. (Later, Larry
Paulson simplified some of the proofs.) These theories were indeed
the original motivation for locales.
--- a/src/HOL/Bali/DeclConcepts.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Bali/DeclConcepts.thy Mon Jun 21 17:41:57 2010 +0200
@@ -244,30 +244,30 @@
lemma mhead_static_simp [simp]: "is_static (mhead m) = is_static m"
by (cases m) (simp add: mhead_def member_is_static_simp)
-constdefs --{* some mnemotic selectors for various pairs *}
-
- "decliface":: "(qtname \<times> ('a::type) decl_scheme) \<Rightarrow> qtname"
- "decliface \<equiv> fst" --{* get the interface component *}
+-- {* some mnemotic selectors for various pairs *}
+
+definition decliface :: "qtname \<times> 'a decl_scheme \<Rightarrow> qtname" where
+ "decliface = fst" --{* get the interface component *}
- "mbr":: "(qtname \<times> memberdecl) \<Rightarrow> memberdecl"
- "mbr \<equiv> snd" --{* get the memberdecl component *}
+definition mbr :: "qtname \<times> memberdecl \<Rightarrow> memberdecl" where
+ "mbr = snd" --{* get the memberdecl component *}
- "mthd":: "('b \<times> 'a) \<Rightarrow> 'a"
- --{* also used for mdecl, mhead *}
- "mthd \<equiv> snd" --{* get the method component *}
+definition mthd :: "'b \<times> 'a \<Rightarrow> 'a" where
+ "mthd = snd" --{* get the method component *}
+ --{* also used for mdecl, mhead *}
- "fld":: "('b \<times> ('a::type) decl_scheme) \<Rightarrow> ('a::type) decl_scheme"
- --{* also used for @{text "((vname \<times> qtname)\<times> field)"} *}
- "fld \<equiv> snd" --{* get the field component *}
+definition fld :: "'b \<times> 'a decl_scheme \<Rightarrow> 'a decl_scheme" where
+ "fld = snd" --{* get the field component *}
+ --{* also used for @{text "((vname \<times> qtname)\<times> field)"} *}
+-- {* some mnemotic selectors for @{text "(vname \<times> qtname)"} *}
-constdefs --{* some mnemotic selectors for @{text "(vname \<times> qtname)"} *}
- fname:: "(vname \<times> 'a) \<Rightarrow> vname" --{* also used for fdecl *}
- "fname \<equiv> fst"
-
- declclassf:: "(vname \<times> qtname) \<Rightarrow> qtname"
- "declclassf \<equiv> snd"
+definition fname:: "vname \<times> 'a \<Rightarrow> vname" where
+ "fname = fst"
+ --{* also used for fdecl *}
+definition declclassf:: "(vname \<times> qtname) \<Rightarrow> qtname" where
+ "declclassf = snd"
lemma decliface_simp[simp]: "decliface (I,m) = I"
@@ -318,12 +318,13 @@
lemma declclassf_simp[simp]:"declclassf (n,c) = c"
by (simp add: declclassf_def)
-constdefs --{* some mnemotic selectors for @{text "(vname \<times> qtname)"} *}
- "fldname" :: "(vname \<times> qtname) \<Rightarrow> vname"
- "fldname \<equiv> fst"
+ --{* some mnemotic selectors for @{text "(vname \<times> qtname)"} *}
- "fldclass" :: "(vname \<times> qtname) \<Rightarrow> qtname"
- "fldclass \<equiv> snd"
+definition fldname :: "vname \<times> qtname \<Rightarrow> vname" where
+ "fldname = fst"
+
+definition fldclass :: "vname \<times> qtname \<Rightarrow> qtname" where
+ "fldclass = snd"
lemma fldname_simp[simp]: "fldname (n,c) = n"
by (simp add: fldname_def)
@@ -337,8 +338,8 @@
text {* Convert a qualified method declaration (qualified with its declaring
class) to a qualified member declaration: @{text methdMembr} *}
-definition methdMembr :: "(qtname \<times> mdecl) \<Rightarrow> (qtname \<times> memberdecl)" where
- "methdMembr m \<equiv> (fst m,mdecl (snd m))"
+definition methdMembr :: "qtname \<times> mdecl \<Rightarrow> qtname \<times> memberdecl" where
+ "methdMembr m = (fst m, mdecl (snd m))"
lemma methdMembr_simp[simp]: "methdMembr (c,m) = (c,mdecl m)"
by (simp add: methdMembr_def)
--- a/src/HOL/Bali/WellType.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Bali/WellType.thy Mon Jun 21 17:41:57 2010 +0200
@@ -103,22 +103,23 @@
"mheads G S (ClassT C) = cmheads G S C"
"mheads G S (ArrayT T) = accObjectmheads G S (ArrayT T)"
-constdefs
+definition
--{* applicable methods, cf. 15.11.2.1 *}
- appl_methds :: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> (emhead \<times> ty list)\<spacespace> set"
- "appl_methds G S rt \<equiv> \<lambda> sig.
+ appl_methds :: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> (emhead \<times> ty list)\<spacespace> set" where
+ "appl_methds G S rt = (\<lambda> sig.
{(mh,pTs') |mh pTs'. mh \<in> mheads G S rt \<lparr>name=name sig,parTs=pTs'\<rparr> \<and>
- G\<turnstile>(parTs sig)[\<preceq>]pTs'}"
+ G\<turnstile>(parTs sig)[\<preceq>]pTs'})"
+definition
--{* more specific methods, cf. 15.11.2.2 *}
- more_spec :: "prog \<Rightarrow> emhead \<times> ty list \<Rightarrow> emhead \<times> ty list \<Rightarrow> bool"
- "more_spec G \<equiv> \<lambda>(mh,pTs). \<lambda>(mh',pTs'). G\<turnstile>pTs[\<preceq>]pTs'"
+ more_spec :: "prog \<Rightarrow> emhead \<times> ty list \<Rightarrow> emhead \<times> ty list \<Rightarrow> bool" where
+ "more_spec G = (\<lambda>(mh,pTs). \<lambda>(mh',pTs'). G\<turnstile>pTs[\<preceq>]pTs')"
(*more_spec G \<equiv>\<lambda>((d,h),pTs). \<lambda>((d',h'),pTs'). G\<turnstile>RefT d\<preceq>RefT d'\<and>G\<turnstile>pTs[\<preceq>]pTs'*)
+definition
--{* maximally specific methods, cf. 15.11.2.2 *}
- max_spec :: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> (emhead \<times> ty list)\<spacespace> set"
-
- "max_spec G S rt sig \<equiv>{m. m \<in>appl_methds G S rt sig \<and>
+ max_spec :: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> (emhead \<times> ty list)\<spacespace> set" where
+ "max_spec G S rt sig = {m. m \<in>appl_methds G S rt sig \<and>
(\<forall>m'\<in>appl_methds G S rt sig. more_spec G m' m \<longrightarrow> m'=m)}"
--- a/src/HOL/Decision_Procs/Approximation.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Decision_Procs/Approximation.thy Mon Jun 21 17:41:57 2010 +0200
@@ -2552,9 +2552,7 @@
where l_eq: "Some (l, u') = approx prec a vs"
and u_eq: "Some (l', u) = approx prec b vs"
and approx_form': "approx_form' prec f (ss ! n) n l u vs ss"
- by (cases "approx prec a vs", simp,
- cases "approx prec b vs", auto) blast
-
+ by (cases "approx prec a vs", simp) (cases "approx prec b vs", auto)
{ assume "xs ! n \<in> { interpret_floatarith a xs .. interpret_floatarith b xs }"
with approx[OF Bound.prems(2) l_eq] and approx[OF Bound.prems(2) u_eq]
have "xs ! n \<in> { real l .. real u}" by auto
@@ -3440,7 +3438,7 @@
fun dest_float (@{const "Float"} $ m $ e) = (snd (HOLogic.dest_number m), snd (HOLogic.dest_number e))
fun dest_ivl (Const (@{const_name "Some"}, _) $
- (Const (@{const_name "Pair"}, _) $ u $ l)) = SOME (dest_float u, dest_float l)
+ (Const (@{const_name Pair}, _) $ u $ l)) = SOME (dest_float u, dest_float l)
| dest_ivl (Const (@{const_name "None"}, _)) = NONE
| dest_ivl t = raise TERM ("dest_result", [t])
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Mon Jun 21 17:41:57 2010 +0200
@@ -2964,7 +2964,7 @@
fun rz rT = Const(@{const_name Groups.zero},rT);
fun dest_nat t = case t of
- Const ("Suc",_)$t' => 1 + dest_nat t'
+ Const (@{const_name Suc}, _) $ t' => 1 + dest_nat t'
| _ => (snd o HOLogic.dest_number) t;
fun num_of_term m t =
--- a/src/HOL/Finite_Set.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Finite_Set.thy Mon Jun 21 17:41:57 2010 +0200
@@ -2180,6 +2180,55 @@
by (auto intro: le_antisym card_inj_on_le)
+subsubsection {* Pigeonhole Principles *}
+
+lemma pigeonhole: "finite(A) \<Longrightarrow> card A > card(f ` A) \<Longrightarrow> ~ inj_on f A "
+by (auto dest: card_image less_irrefl_nat)
+
+lemma pigeonhole_infinite:
+assumes "~ finite A" and "finite(f`A)"
+shows "EX a0:A. ~finite{a:A. f a = f a0}"
+proof -
+ have "finite(f`A) \<Longrightarrow> ~ finite A \<Longrightarrow> EX a0:A. ~finite{a:A. f a = f a0}"
+ proof(induct "f`A" arbitrary: A rule: finite_induct)
+ case empty thus ?case by simp
+ next
+ case (insert b F)
+ show ?case
+ proof cases
+ assume "finite{a:A. f a = b}"
+ hence "~ finite(A - {a:A. f a = b})" using `\<not> finite A` by simp
+ also have "A - {a:A. f a = b} = {a:A. f a \<noteq> b}" by blast
+ finally have "~ finite({a:A. f a \<noteq> b})" .
+ from insert(3)[OF _ this]
+ show ?thesis using insert(2,4) by simp (blast intro: rev_finite_subset)
+ next
+ assume 1: "~finite{a:A. f a = b}"
+ hence "{a \<in> A. f a = b} \<noteq> {}" by force
+ thus ?thesis using 1 by blast
+ qed
+ qed
+ from this[OF assms(2,1)] show ?thesis .
+qed
+
+lemma pigeonhole_infinite_rel:
+assumes "~finite A" and "finite B" and "ALL a:A. EX b:B. R a b"
+shows "EX b:B. ~finite{a:A. R a b}"
+proof -
+ let ?F = "%a. {b:B. R a b}"
+ from finite_Pow_iff[THEN iffD2, OF `finite B`]
+ have "finite(?F ` A)" by(blast intro: rev_finite_subset)
+ from pigeonhole_infinite[where f = ?F, OF assms(1) this]
+ obtain a0 where "a0\<in>A" and 1: "\<not> finite {a\<in>A. ?F a = ?F a0}" ..
+ obtain b0 where "b0 : B" and "R a0 b0" using `a0:A` assms(3) by blast
+ { assume "finite{a:A. R a b0}"
+ then have "finite {a\<in>A. ?F a = ?F a0}"
+ using `b0 : B` `R a0 b0` by(blast intro: rev_finite_subset)
+ }
+ with 1 `b0 : B` show ?thesis by blast
+qed
+
+
subsubsection {* Cardinality of sums *}
lemma card_Plus:
@@ -2210,7 +2259,7 @@
apply (blast elim!: equalityE)
done
-text {* Relates to equivalence classes. Based on a theorem of F. Kammüller. *}
+text {* Relates to equivalence classes. Based on a theorem of F. Kammüller. *}
lemma dvd_partition:
"finite (Union C) ==>
--- a/src/HOL/HOL.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/HOL.thy Mon Jun 21 17:41:57 2010 +0200
@@ -1773,6 +1773,7 @@
setup {*
Code_Preproc.map_pre (K HOL_basic_ss)
#> Code_Preproc.map_post (K HOL_basic_ss)
+ #> Code_Simp.map_ss (K HOL_basic_ss)
*}
subsubsection {* Equality *}
@@ -1800,8 +1801,8 @@
setup {*
Code_Preproc.map_pre (fn simpset =>
simpset addsimprocs [Simplifier.simproc_i @{theory} "eq" [@{term "op ="}]
- (fn thy => fn _ => fn t as Const (_, T) => case strip_type T
- of ((T as Type _) :: _, _) => SOME @{thm equals_eq}
+ (fn thy => fn _ => fn Const (_, T) => case strip_type T
+ of (Type _ :: _, _) => SOME @{thm equals_eq}
| _ => NONE)])
*}
--- a/src/HOL/Hoare/hoare_tac.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Hoare/hoare_tac.ML Mon Jun 21 17:41:57 2010 +0200
@@ -41,9 +41,9 @@
else let val (n, T) = dest_Free x ;
val z = mk_bodyC xs;
val T2 = case z of Free(_, T) => T
- | Const ("Pair", Type ("fun", [_, Type
+ | Const (@{const_name Pair}, Type ("fun", [_, Type
("fun", [_, T])])) $ _ $ _ => T;
- in Const ("Pair", [T, T2] ---> mk_prodT (T, T2)) $ x $ z end;
+ in Const (@{const_name Pair}, [T, T2] ---> mk_prodT (T, T2)) $ x $ z end;
(** maps a subgoal of the form:
VARS x1 ... xn {._.} _ {._.} or to [x1,...,xn]**)
--- a/src/HOL/Imperative_HOL/ex/SatChecker.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Imperative_HOL/ex/SatChecker.thy Mon Jun 21 17:41:57 2010 +0200
@@ -555,7 +555,6 @@
apply auto
apply (drule fold_steps_correct)
apply (simp add: correctArray_def array_ran_def)
-apply (cases "n = 0", auto)
apply (rule implies_empty_inconsistent)
apply (simp add:correctArray_def)
apply (drule bspec)
--- a/src/HOL/Import/Generate-HOL/GenHOL4Base.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Import/Generate-HOL/GenHOL4Base.thy Mon Jun 21 17:41:57 2010 +0200
@@ -123,7 +123,7 @@
import_theory pair;
type_maps
- prod > "*";
+ prod > "Product_Type.*";
const_maps
"," > Pair
--- a/src/HOL/Import/Generate-HOLLight/GenHOLLight.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Import/Generate-HOLLight/GenHOLLight.thy Mon Jun 21 17:41:57 2010 +0200
@@ -38,9 +38,9 @@
bool > bool
fun > fun
N_1 > Product_Type.unit
- prod > "*"
- num > nat
- sum > "+"
+ prod > "Product_Type.*"
+ num > Nat.nat
+ sum > "Sum_Type.+"
(* option > Datatype.option*);
const_renames
--- a/src/HOL/Import/HOL/num.imp Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Import/HOL/num.imp Mon Jun 21 17:41:57 2010 +0200
@@ -3,10 +3,10 @@
import_segment "hol4"
type_maps
- "num" > "nat"
+ "num" > "Nat.nat"
const_maps
- "SUC" > "Suc"
+ "SUC" > "Nat.Suc"
"0" > "0" :: "nat"
thm_maps
--- a/src/HOL/Import/HOL/pair.imp Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Import/HOL/pair.imp Mon Jun 21 17:41:57 2010 +0200
@@ -7,17 +7,17 @@
"LEX" > "LEX_def"
type_maps
- "prod" > "*"
+ "prod" > "Product_Type.*"
const_maps
- "pair_case" > "split"
- "UNCURRY" > "split"
- "SND" > "snd"
+ "pair_case" > "Product_Type.split"
+ "UNCURRY" > "Product_Type.split"
+ "FST" > "Product_Type.fst"
+ "SND" > "Product_Type.snd"
"RPROD" > "HOL4Base.pair.RPROD"
"LEX" > "HOL4Base.pair.LEX"
- "FST" > "fst"
- "CURRY" > "curry"
- "," > "Pair"
+ "CURRY" > "Product_Type.curry"
+ "," > "Product_Type.Pair"
"##" > "prod_fun"
thm_maps
--- a/src/HOL/Import/HOLLight/hollight.imp Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Import/HOLLight/hollight.imp Mon Jun 21 17:41:57 2010 +0200
@@ -6,9 +6,9 @@
"sum" > "+"
"recspace" > "HOLLight.hollight.recspace"
"real" > "HOLLight.hollight.real"
- "prod" > "*"
+ "prod" > "Product_Type.*"
"option" > "HOLLight.hollight.option"
- "num" > "nat"
+ "num" > "Nat.nat"
"nadd" > "HOLLight.hollight.nadd"
"list" > "HOLLight.hollight.list"
"int" > "HOLLight.hollight.int"
@@ -128,10 +128,10 @@
"TL" > "HOLLight.hollight.TL"
"T" > "True"
"SURJ" > "HOLLight.hollight.SURJ"
- "SUC" > "Suc"
+ "SUC" > "Nat.Suc"
"SUBSET" > "HOLLight.hollight.SUBSET"
"SOME" > "HOLLight.hollight.SOME"
- "SND" > "snd"
+ "SND" > "Product_Type.snd"
"SING" > "HOLLight.hollight.SING"
"SETSPEC" > "HOLLight.hollight.SETSPEC"
"REVERSE" > "HOLLight.hollight.REVERSE"
@@ -188,7 +188,7 @@
"GSPEC" > "HOLLight.hollight.GSPEC"
"GEQ" > "HOLLight.hollight.GEQ"
"GABS" > "HOLLight.hollight.GABS"
- "FST" > "fst"
+ "FST" > "Product_Type.fst"
"FNIL" > "HOLLight.hollight.FNIL"
"FINREC" > "HOLLight.hollight.FINREC"
"FINITE" > "HOLLight.hollight.FINITE"
@@ -239,7 +239,7 @@
"<" > "HOLLight.hollight.<"
"/\\" > "op &"
"-" > "Groups.minus" :: "nat => nat => nat"
- "," > "Pair"
+ "," > "Product_Type.Pair"
"+" > "Groups.plus" :: "nat => nat => nat"
"*" > "Groups.times" :: "nat => nat => nat"
"$" > "HOLLight.hollight.$"
--- a/src/HOL/Import/proof_kernel.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Import/proof_kernel.ML Mon Jun 21 17:41:57 2010 +0200
@@ -1946,7 +1946,7 @@
val p3 = string_of_mixfix csyn
val p4 = smart_string_of_cterm crhs
in
- add_dump ("constdefs\n " ^ p1 ^ " :: \"" ^ p2 ^ "\" "^ p3 ^ "\n " ^ p4) thy''
+ add_dump ("definition\n " ^ p1 ^ " :: \"" ^ p2 ^ "\" "^ p3 ^ " where\n " ^ p4) thy''
end
else
add_dump ("consts\n " ^ quotename constname ^ " :: \"" ^
@@ -2139,18 +2139,6 @@
end
handle e => (message "exception in new_type_definition"; print_exn e)
-fun add_dump_constdefs thy defname constname rhs ty =
- let
- val n = quotename constname
- val t = Syntax.string_of_typ_global thy ty
- val syn = string_of_mixfix (mk_syn thy constname)
- (*val eq = smart_string_of_cterm (cterm_of thy (Const(rhs, ty)))*)
- val eq = quote (constname ^ " == "^rhs)
- val d = case defname of NONE => "" | SOME defname => (quotename defname)^" : "
- in
- add_dump ("constdefs\n " ^ n ^ " :: \"" ^ t ^ "\" " ^ syn ^ "\n " ^ d ^ eq) thy
- end
-
fun add_dump_syntax thy name =
let
val n = quotename name
--- a/src/HOL/Inductive.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Inductive.thy Mon Jun 21 17:41:57 2010 +0200
@@ -9,7 +9,6 @@
uses
("Tools/inductive.ML")
"Tools/dseq.ML"
- ("Tools/inductive_codegen.ML")
"Tools/Datatype/datatype_aux.ML"
"Tools/Datatype/datatype_prop.ML"
"Tools/Datatype/datatype_case.ML"
@@ -286,9 +285,6 @@
use "Tools/old_primrec.ML"
use "Tools/primrec.ML"
-use "Tools/inductive_codegen.ML"
-setup InductiveCodegen.setup
-
text{* Lambda-abstractions with pattern matching: *}
syntax
--- a/src/HOL/IsaMakefile Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/IsaMakefile Mon Jun 21 17:41:57 2010 +0200
@@ -112,6 +112,7 @@
$(SRC)/Tools/Code/code_preproc.ML \
$(SRC)/Tools/Code/code_printer.ML \
$(SRC)/Tools/Code/code_scala.ML \
+ $(SRC)/Tools/Code/code_simp.ML \
$(SRC)/Tools/Code/code_target.ML \
$(SRC)/Tools/Code/code_thingol.ML \
$(SRC)/Tools/Code_Generator.thy \
--- a/src/HOL/Library/AssocList.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Library/AssocList.thy Mon Jun 21 17:41:57 2010 +0200
@@ -5,7 +5,7 @@
header {* Map operations implemented on association lists*}
theory AssocList
-imports Main Mapping
+imports Main More_List Mapping
begin
text {*
@@ -79,7 +79,7 @@
by (simp add: update_conv' image_map_upd)
definition updates :: "'key list \<Rightarrow> 'val list \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list" where
- "updates ks vs al = foldl (\<lambda>al (k, v). update k v al) al (zip ks vs)"
+ "updates ks vs = More_List.fold (prod_case update) (zip ks vs)"
lemma updates_simps [simp]:
"updates [] vs ps = ps"
@@ -94,11 +94,10 @@
lemma updates_conv': "map_of (updates ks vs al) = (map_of al)(ks[\<mapsto>]vs)"
proof -
- have "foldl (\<lambda>f (k, v). f(k \<mapsto> v)) (map_of al) (zip ks vs) =
- map_of (foldl (\<lambda>al (k, v). update k v al) al (zip ks vs))"
- by (rule foldl_apply) (auto simp add: expand_fun_eq update_conv')
- then show ?thesis
- by (simp add: updates_def map_upds_fold_map_upd)
+ have "map_of \<circ> More_List.fold (prod_case update) (zip ks vs) =
+ More_List.fold (\<lambda>(k, v) f. f(k \<mapsto> v)) (zip ks vs) \<circ> map_of"
+ by (rule fold_apply) (auto simp add: expand_fun_eq update_conv')
+ then show ?thesis by (auto simp add: updates_def expand_fun_eq map_upds_fold_map_upd foldl_fold split_def)
qed
lemma updates_conv: "map_of (updates ks vs al) k = ((map_of al)(ks[\<mapsto>]vs)) k"
@@ -108,15 +107,14 @@
assumes "distinct (map fst al)"
shows "distinct (map fst (updates ks vs al))"
proof -
- from assms have "distinct (foldl
- (\<lambda>al (k, v). if k \<in> set al then al else al @ [k])
- (map fst al) (zip ks vs))"
- by (rule foldl_invariant) auto
- moreover have "foldl (\<lambda>al (k, v). if k \<in> set al then al else al @ [k])
- (map fst al) (zip ks vs)
- = map fst (foldl (\<lambda>al (k, v). update k v al) al (zip ks vs))"
- by (rule foldl_apply) (simp add: update_keys split_def comp_def)
- ultimately show ?thesis by (simp add: updates_def)
+ have "distinct (More_List.fold
+ (\<lambda>(k, v) al. if k \<in> set al then al else al @ [k])
+ (zip ks vs) (map fst al))"
+ by (rule fold_invariant [of "zip ks vs" "\<lambda>_. True"]) (auto intro: assms)
+ moreover have "map fst \<circ> More_List.fold (prod_case update) (zip ks vs) =
+ More_List.fold (\<lambda>(k, v) al. if k \<in> set al then al else al @ [k]) (zip ks vs) \<circ> map fst"
+ by (rule fold_apply) (simp add: update_keys split_def prod_case_beta comp_def)
+ ultimately show ?thesis by (simp add: updates_def expand_fun_eq)
qed
lemma updates_append1[simp]: "size ks < size vs \<Longrightarrow>
@@ -341,10 +339,10 @@
lemma clearjunk_updates:
"clearjunk (updates ks vs al) = updates ks vs (clearjunk al)"
proof -
- have "foldl (\<lambda>al (k, v). update k v al) (clearjunk al) (zip ks vs) =
- clearjunk (foldl (\<lambda>al (k, v). update k v al) al (zip ks vs))"
- by (rule foldl_apply) (simp add: clearjunk_update expand_fun_eq split_def)
- then show ?thesis by (simp add: updates_def)
+ have "clearjunk \<circ> More_List.fold (prod_case update) (zip ks vs) =
+ More_List.fold (prod_case update) (zip ks vs) \<circ> clearjunk"
+ by (rule fold_apply) (simp add: clearjunk_update prod_case_beta o_def)
+ then show ?thesis by (simp add: updates_def expand_fun_eq)
qed
lemma clearjunk_delete:
@@ -429,8 +427,8 @@
lemma merge_updates:
"merge qs ps = updates (rev (map fst ps)) (rev (map snd ps)) qs"
- by (simp add: merge_def updates_def split_def
- foldr_foldl zip_rev zip_map_fst_snd)
+ by (simp add: merge_def updates_def split_prod_case
+ foldr_fold_rev zip_rev zip_map_fst_snd)
lemma dom_merge: "fst ` set (merge xs ys) = fst ` set xs \<union> fst ` set ys"
by (induct ys arbitrary: xs) (auto simp add: dom_update)
@@ -447,11 +445,11 @@
lemma merge_conv':
"map_of (merge xs ys) = map_of xs ++ map_of ys"
proof -
- have "foldl (\<lambda>m (k, v). m(k \<mapsto> v)) (map_of xs) (rev ys) =
- map_of (foldl (\<lambda>xs (k, v). update k v xs) xs (rev ys)) "
- by (rule foldl_apply) (simp add: expand_fun_eq split_def update_conv')
+ have "map_of \<circ> More_List.fold (prod_case update) (rev ys) =
+ More_List.fold (\<lambda>(k, v) m. m(k \<mapsto> v)) (rev ys) \<circ> map_of"
+ by (rule fold_apply) (simp add: update_conv' prod_case_beta split_def expand_fun_eq)
then show ?thesis
- by (simp add: merge_def map_add_map_of_foldr foldr_foldl split_def)
+ by (simp add: merge_def map_add_map_of_foldr foldr_fold_rev split_prod_case expand_fun_eq)
qed
corollary merge_conv:
--- a/src/HOL/Library/Binomial.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Library/Binomial.thy Mon Jun 21 17:41:57 2010 +0200
@@ -235,7 +235,7 @@
have eq: "insert 0 {1 .. n} = {0..n}" by auto
have th1: "(\<Prod>n\<in>{1\<Colon>nat..n}. a + of_nat n) =
(\<Prod>n\<in>{0\<Colon>nat..n - 1}. a + 1 + of_nat n)"
- apply (rule setprod_reindex_cong[where f = "Suc"])
+ apply (rule setprod_reindex_cong [where f = Suc])
using n0 by (auto simp add: expand_fun_eq field_simps)
have ?thesis apply (simp add: pochhammer_def)
unfolding setprod_insert[OF th0, unfolded eq]
--- a/src/HOL/Library/Code_Char.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Library/Code_Char.thy Mon Jun 21 17:41:57 2010 +0200
@@ -58,12 +58,12 @@
(SML "String.implode")
(OCaml "failwith/ \"implode\"")
(Haskell "_")
- (Scala "List.toString((_))")
+ (Scala "!(\"\" ++/ _)")
code_const explode
(SML "String.explode")
(OCaml "failwith/ \"explode\"")
(Haskell "_")
- (Scala "List.fromString((_))")
+ (Scala "!(_.toList)")
end
--- a/src/HOL/Library/Countable.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Library/Countable.thy Mon Jun 21 17:41:57 2010 +0200
@@ -70,7 +70,7 @@
text {* Sums *}
-instance "+":: (countable, countable) countable
+instance "+" :: (countable, countable) countable
by (rule countable_classI [of "(\<lambda>x. case x of Inl a \<Rightarrow> to_nat (False, to_nat a)
| Inr b \<Rightarrow> to_nat (True, to_nat b))"])
(simp split: sum.split_asm)
--- a/src/HOL/Library/Diagonalize.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Library/Diagonalize.thy Mon Jun 21 17:41:57 2010 +0200
@@ -81,7 +81,7 @@
done
-subsection {* Diagonalization: an injective embedding of two @{typ "nat"}s to one @{typ "nat"} *}
+subsection {* Diagonalization: an injective embedding of two @{typ nat}s to one @{typ nat} *}
definition diagonalize :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
"diagonalize m n = sum (m + n) + m"
--- a/src/HOL/Library/Dlist.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Library/Dlist.thy Mon Jun 21 17:41:57 2010 +0200
@@ -3,7 +3,7 @@
header {* Lists with elements distinct as canonical example for datatype invariants *}
theory Dlist
-imports Main More_List Fset
+imports Main Fset
begin
section {* The type of distinct lists *}
@@ -181,11 +181,11 @@
lemma Set_Dlist [simp]:
"Set (Dlist xs) = Fset (set xs)"
- by (simp add: Set_def Fset.Set_def)
+ by (rule fset_eqI) (simp add: Set_def)
lemma Coset_Dlist [simp]:
"Coset (Dlist xs) = Fset (- set xs)"
- by (simp add: Coset_def Fset.Coset_def)
+ by (rule fset_eqI) (simp add: Coset_def)
lemma member_Set [simp]:
"Fset.member (Set dxs) = List.member (list_of_dlist dxs)"
@@ -197,11 +197,11 @@
lemma Set_dlist_of_list [code]:
"Fset.Set xs = Set (dlist_of_list xs)"
- by simp
+ by (rule fset_eqI) simp
lemma Coset_dlist_of_list [code]:
"Fset.Coset xs = Coset (dlist_of_list xs)"
- by simp
+ by (rule fset_eqI) simp
lemma is_empty_Set [code]:
"Fset.is_empty (Set dxs) \<longleftrightarrow> null dxs"
@@ -233,15 +233,15 @@
lemma compl_code [code]:
"- Set dxs = Coset dxs"
"- Coset dxs = Set dxs"
- by (simp_all add: not_set_compl member_set)
+ by (rule fset_eqI, simp add: member_set not_set_compl)+
lemma map_code [code]:
"Fset.map f (Set dxs) = Set (map f dxs)"
- by (simp add: member_set)
+ by (rule fset_eqI) (simp add: member_set)
lemma filter_code [code]:
"Fset.filter f (Set dxs) = Set (filter f dxs)"
- by (simp add: member_set)
+ by (rule fset_eqI) (simp add: member_set)
lemma forall_Set [code]:
"Fset.forall P (Set xs) \<longleftrightarrow> list_all P (list_of_dlist xs)"
--- a/src/HOL/Library/Efficient_Nat.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Library/Efficient_Nat.thy Mon Jun 21 17:41:57 2010 +0200
@@ -11,7 +11,7 @@
text {*
When generating code for functions on natural numbers, the
canonical representation using @{term "0::nat"} and
- @{term "Suc"} is unsuitable for computations involving large
+ @{term Suc} is unsuitable for computations involving large
numbers. The efficiency of the generated code can be improved
drastically by implementing natural numbers by target-language
integers. To do this, just include this theory.
@@ -362,7 +362,7 @@
Since natural numbers are implemented
using integers in ML, the coercion function @{const "of_nat"} of type
@{typ "nat \<Rightarrow> int"} is simply implemented by the identity function.
- For the @{const "nat"} function for converting an integer to a natural
+ For the @{const nat} function for converting an integer to a natural
number, we give a specific implementation using an ML function that
returns its input value, provided that it is non-negative, and otherwise
returns @{text "0"}.
--- a/src/HOL/Library/Formal_Power_Series.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Library/Formal_Power_Series.thy Mon Jun 21 17:41:57 2010 +0200
@@ -2031,7 +2031,7 @@
done
also have "\<dots> = setsum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (setsum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}"
unfolding fps_deriv_nth
- apply (rule setsum_reindex_cong[where f="Suc"])
+ apply (rule setsum_reindex_cong [where f = Suc])
by (auto simp add: mult_assoc)
finally have th0: "(fps_deriv (a oo b))$n = setsum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (setsum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}" .
--- a/src/HOL/Library/Fset.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Library/Fset.thy Mon Jun 21 17:41:57 2010 +0200
@@ -7,22 +7,30 @@
imports More_Set More_List
begin
-declare mem_def [simp]
-
subsection {* Lifting *}
-datatype 'a fset = Fset "'a set"
+typedef (open) 'a fset = "UNIV :: 'a set set"
+ morphisms member Fset by rule+
-primrec member :: "'a fset \<Rightarrow> 'a set" where
+lemma member_Fset [simp]:
"member (Fset A) = A"
-
-lemma member_inject [simp]:
- "member A = member B \<Longrightarrow> A = B"
- by (cases A, cases B) simp
+ by (rule Fset_inverse) rule
lemma Fset_member [simp]:
"Fset (member A) = A"
- by (cases A) simp
+ by (rule member_inverse)
+
+declare member_inject [simp]
+
+lemma Fset_inject [simp]:
+ "Fset A = Fset B \<longleftrightarrow> A = B"
+ by (simp add: Fset_inject)
+
+lemma fset_eqI:
+ "member A = member B \<Longrightarrow> A = B"
+ by simp
+
+declare mem_def [simp]
definition Set :: "'a list \<Rightarrow> 'a fset" where
"Set xs = Fset (set xs)"
@@ -56,6 +64,27 @@
then show ?thesis by (simp add: image_def)
qed
+definition (in term_syntax)
+ setify :: "'a\<Colon>typerep list \<times> (unit \<Rightarrow> Code_Evaluation.term)
+ \<Rightarrow> 'a fset \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
+ [code_unfold]: "setify xs = Code_Evaluation.valtermify Set {\<cdot>} xs"
+
+notation fcomp (infixl "o>" 60)
+notation scomp (infixl "o\<rightarrow>" 60)
+
+instantiation fset :: (random) random
+begin
+
+definition
+ "Quickcheck.random i = Quickcheck.random i o\<rightarrow> (\<lambda>xs. Pair (setify xs))"
+
+instance ..
+
+end
+
+no_notation fcomp (infixl "o>" 60)
+no_notation scomp (infixl "o\<rightarrow>" 60)
+
subsection {* Lattice instantiation *}
@@ -117,11 +146,11 @@
lemma empty_Set [code]:
"bot = Set []"
- by simp
+ by (simp add: Set_def)
lemma UNIV_Set [code]:
"top = Coset []"
- by simp
+ by (simp add: Coset_def)
definition insert :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
[simp]: "insert x A = Fset (Set.insert x (member A))"
@@ -198,9 +227,16 @@
"A < B \<longleftrightarrow> A \<le> B \<and> \<not> B \<le> (A :: 'a fset)"
by (fact less_le_not_le)
-lemma eq_fset_subfset_eq [code]:
- "eq_class.eq A B \<longleftrightarrow> A \<le> B \<and> B \<le> (A :: 'a fset)"
- by (cases A, cases B) (simp add: eq set_eq)
+instantiation fset :: (type) eq
+begin
+
+definition
+ "eq_fset A B \<longleftrightarrow> A \<le> B \<and> B \<le> (A :: 'a fset)"
+
+instance proof
+qed (simp add: eq_fset_def set_eq [symmetric])
+
+end
subsection {* Functorial operations *}
@@ -276,22 +312,6 @@
end
-subsection {* Misc operations *}
-
-lemma size_fset [code]:
- "fset_size f A = 0"
- "size A = 0"
- by (cases A, simp) (cases A, simp)
-
-lemma fset_case_code [code]:
- "fset_case f A = f (member A)"
- by (cases A) simp
-
-lemma fset_rec_code [code]:
- "fset_rec f A = f (member A)"
- by (cases A) simp
-
-
subsection {* Simplified simprules *}
lemma is_empty_simp [simp]:
@@ -312,7 +332,7 @@
declare mem_def [simp del]
-hide_const (open) is_empty insert remove map filter forall exists card
+hide_const (open) setify is_empty insert remove map filter forall exists card
Inter Union
end
--- a/src/HOL/Library/List_Prefix.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Library/List_Prefix.thy Mon Jun 21 17:41:57 2010 +0200
@@ -10,17 +10,20 @@
subsection {* Prefix order on lists *}
-instantiation list :: (type) order
+instantiation list :: (type) "{order, bot}"
begin
definition
- prefix_def [code del]: "xs \<le> ys = (\<exists>zs. ys = xs @ zs)"
+ prefix_def: "xs \<le> ys \<longleftrightarrow> (\<exists>zs. ys = xs @ zs)"
definition
- strict_prefix_def [code del]: "xs < ys = (xs \<le> ys \<and> xs \<noteq> (ys::'a list))"
+ strict_prefix_def: "xs < ys \<longleftrightarrow> xs \<le> ys \<and> xs \<noteq> (ys::'a list)"
-instance
- by intro_classes (auto simp add: prefix_def strict_prefix_def)
+definition
+ "bot = []"
+
+instance proof
+qed (auto simp add: prefix_def strict_prefix_def bot_list_def)
end
@@ -77,6 +80,12 @@
lemma Cons_prefix_Cons [simp]: "(x # xs \<le> y # ys) = (x = y \<and> xs \<le> ys)"
by (auto simp add: prefix_def)
+lemma less_eq_list_code [code]:
+ "([]\<Colon>'a\<Colon>{eq, ord} list) \<le> xs \<longleftrightarrow> True"
+ "(x\<Colon>'a\<Colon>{eq, ord}) # xs \<le> [] \<longleftrightarrow> False"
+ "(x\<Colon>'a\<Colon>{eq, ord}) # xs \<le> y # ys \<longleftrightarrow> x = y \<and> xs \<le> ys"
+ by simp_all
+
lemma same_prefix_prefix [simp]: "(xs @ ys \<le> xs @ zs) = (ys \<le> zs)"
by (induct xs) simp_all
@@ -125,10 +134,10 @@
lemma prefix_length_less: "xs < ys \<Longrightarrow> length xs < length ys"
by (auto simp: strict_prefix_def prefix_def)
-lemma strict_prefix_simps [simp]:
- "xs < [] = False"
- "[] < (x # xs) = True"
- "(x # xs) < (y # ys) = (x = y \<and> xs < ys)"
+lemma strict_prefix_simps [simp, code]:
+ "xs < [] \<longleftrightarrow> False"
+ "[] < x # xs \<longleftrightarrow> True"
+ "x # xs < y # ys \<longleftrightarrow> x = y \<and> xs < ys"
by (simp_all add: strict_prefix_def cong: conj_cong)
lemma take_strict_prefix: "xs < ys \<Longrightarrow> take n xs < ys"
@@ -301,7 +310,7 @@
by (induct zs) (auto intro!: postfix_appendI postfix_ConsI)
qed
-lemma postfix_to_prefix: "xs >>= ys \<longleftrightarrow> rev ys \<le> rev xs"
+lemma postfix_to_prefix [code]: "xs >>= ys \<longleftrightarrow> rev ys \<le> rev xs"
proof
assume "xs >>= ys"
then obtain zs where "xs = zs @ ys" ..
@@ -370,21 +379,4 @@
qed
qed
-
-subsection {* Executable code *}
-
-lemma less_eq_code [code]:
- "([]\<Colon>'a\<Colon>{eq, ord} list) \<le> xs \<longleftrightarrow> True"
- "(x\<Colon>'a\<Colon>{eq, ord}) # xs \<le> [] \<longleftrightarrow> False"
- "(x\<Colon>'a\<Colon>{eq, ord}) # xs \<le> y # ys \<longleftrightarrow> x = y \<and> xs \<le> ys"
- by simp_all
-
-lemma less_code [code]:
- "xs < ([]\<Colon>'a\<Colon>{eq, ord} list) \<longleftrightarrow> False"
- "[] < (x\<Colon>'a\<Colon>{eq, ord})# xs \<longleftrightarrow> True"
- "(x\<Colon>'a\<Colon>{eq, ord}) # xs < y # ys \<longleftrightarrow> x = y \<and> xs < ys"
- unfolding strict_prefix_def by auto
-
-lemmas [code] = postfix_to_prefix
-
end
--- a/src/HOL/Library/List_lexord.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Library/List_lexord.thy Mon Jun 21 17:41:57 2010 +0200
@@ -12,10 +12,10 @@
begin
definition
- list_less_def [code del]: "(xs::('a::ord) list) < ys \<longleftrightarrow> (xs, ys) \<in> lexord {(u,v). u < v}"
+ list_less_def: "xs < ys \<longleftrightarrow> (xs, ys) \<in> lexord {(u, v). u < v}"
definition
- list_le_def [code del]: "(xs::('a::ord) list) \<le> ys \<longleftrightarrow> (xs < ys \<or> xs = ys)"
+ list_le_def: "(xs :: _ list) \<le> ys \<longleftrightarrow> xs < ys \<or> xs = ys"
instance ..
@@ -91,13 +91,24 @@
lemma Cons_le_Cons [simp]: "a # x \<le> b # y \<longleftrightarrow> a < b \<or> a = b \<and> x \<le> y"
by (unfold list_le_def) auto
-lemma less_code [code]:
+instantiation list :: (order) bot
+begin
+
+definition
+ "bot = []"
+
+instance proof
+qed (simp add: bot_list_def)
+
+end
+
+lemma less_list_code [code]:
"xs < ([]\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> False"
"[] < (x\<Colon>'a\<Colon>{eq, order}) # xs \<longleftrightarrow> True"
"(x\<Colon>'a\<Colon>{eq, order}) # xs < y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs < ys"
by simp_all
-lemma less_eq_code [code]:
+lemma less_eq_list_code [code]:
"x # xs \<le> ([]\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> False"
"[] \<le> (xs\<Colon>'a\<Colon>{eq, order} list) \<longleftrightarrow> True"
"(x\<Colon>'a\<Colon>{eq, order}) # xs \<le> y # ys \<longleftrightarrow> x < y \<or> x = y \<and> xs \<le> ys"
--- a/src/HOL/Library/RBT.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Library/RBT.thy Mon Jun 21 17:41:57 2010 +0200
@@ -143,7 +143,7 @@
by (simp add: map_def lookup_RBT lookup_map lookup_impl_of)
lemma fold_fold:
- "fold f t = (\<lambda>s. foldl (\<lambda>s (k, v). f k v s) s (entries t))"
+ "fold f t = More_List.fold (prod_case f) (entries t)"
by (simp add: fold_def expand_fun_eq RBT_Impl.fold_def entries_impl_of)
lemma is_empty_empty [simp]:
--- a/src/HOL/Library/RBT_Impl.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Library/RBT_Impl.thy Mon Jun 21 17:41:57 2010 +0200
@@ -6,7 +6,7 @@
header {* Implementation of Red-Black Trees *}
theory RBT_Impl
-imports Main
+imports Main More_List
begin
text {*
@@ -1049,7 +1049,7 @@
subsection {* Folding over entries *}
definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c" where
- "fold f t s = foldl (\<lambda>s (k, v). f k v s) s (entries t)"
+ "fold f t = More_List.fold (prod_case f) (entries t)"
lemma fold_simps [simp, code]:
"fold f Empty = id"
@@ -1071,12 +1071,12 @@
proof -
obtain ys where "ys = rev xs" by simp
have "\<And>t. is_rbt t \<Longrightarrow>
- lookup (foldl (\<lambda>t (k, v). insert k v t) t ys) = lookup t ++ map_of (rev ys)"
- by (induct ys) (simp_all add: bulkload_def split_def lookup_insert)
+ lookup (More_List.fold (prod_case insert) ys t) = lookup t ++ map_of (rev ys)"
+ by (induct ys) (simp_all add: bulkload_def lookup_insert prod_case_beta)
from this Empty_is_rbt have
- "lookup (foldl (\<lambda>t (k, v). insert k v t) Empty (rev xs)) = lookup Empty ++ map_of xs"
+ "lookup (More_List.fold (prod_case insert) (rev xs) Empty) = lookup Empty ++ map_of xs"
by (simp add: `ys = rev xs`)
- then show ?thesis by (simp add: bulkload_def foldl_foldr lookup_Empty split_def)
+ then show ?thesis by (simp add: bulkload_def lookup_Empty foldr_fold_rev prod_case_split)
qed
hide_const (open) Empty insert delete entries keys bulkload lookup map_entry map fold union sorted
--- a/src/HOL/Library/SML_Quickcheck.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Library/SML_Quickcheck.thy Mon Jun 21 17:41:57 2010 +0200
@@ -6,7 +6,7 @@
begin
setup {*
- InductiveCodegen.quickcheck_setup #>
+ Inductive_Codegen.quickcheck_setup #>
Quickcheck.add_generator ("SML", Codegen.test_term)
*}
--- a/src/HOL/List.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/List.thy Mon Jun 21 17:41:57 2010 +0200
@@ -5,7 +5,7 @@
header {* The datatype of finite lists *}
theory List
-imports Plain Presburger Sledgehammer Recdef
+imports Plain Quotient Presburger Code_Numeral Sledgehammer Recdef
uses ("Tools/list_code.ML")
begin
@@ -3190,8 +3190,16 @@
lemma set_replicate_conv_if: "set (replicate n x) = (if n = 0 then {} else {x})"
by auto
-lemma in_set_replicateD: "x : set (replicate n y) ==> x = y"
-by (simp add: set_replicate_conv_if split: split_if_asm)
+lemma in_set_replicate[simp]: "(x : set (replicate n y)) = (x = y & n \<noteq> 0)"
+by (simp add: set_replicate_conv_if)
+
+lemma Ball_set_replicate[simp]:
+ "(ALL x : set(replicate n a). P x) = (P a | n=0)"
+by(simp add: set_replicate_conv_if)
+
+lemma Bex_set_replicate[simp]:
+ "(EX x : set(replicate n a). P x) = (P a & n\<noteq>0)"
+by(simp add: set_replicate_conv_if)
lemma replicate_append_same:
"replicate i x @ [x] = x # replicate i x"
@@ -4188,8 +4196,8 @@
primrec -- {*The lexicographic ordering for lists of the specified length*}
lexn :: "('a \<times> 'a) set \<Rightarrow> nat \<Rightarrow> ('a list \<times> 'a list) set" where
- "lexn r 0 = {}"
- | "lexn r (Suc n) = (prod_fun (%(x, xs). x#xs) (%(x, xs). x#xs) ` (r <*lex*> lexn r n)) Int
+ [code del]: "lexn r 0 = {}"
+ | [code del]: "lexn r (Suc n) = (prod_fun (%(x, xs). x#xs) (%(x, xs). x#xs) ` (r <*lex*> lexn r n)) Int
{(xs, ys). length xs = Suc n \<and> length ys = Suc n}"
definition
@@ -4706,7 +4714,11 @@
lemma list_all_length:
"list_all P xs \<longleftrightarrow> (\<forall>n < length xs. P (xs ! n))"
- unfolding list_all_iff by (auto intro: all_nth_imp_all_set)
+unfolding list_all_iff by (auto intro: all_nth_imp_all_set)
+
+lemma list_all_cong[fundef_cong]:
+ "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> list_all f xs = list_all g ys"
+by (simp add: list_all_iff)
lemma list_ex_iff [code_post]:
"list_ex P xs \<longleftrightarrow> (\<exists>x \<in> set xs. P x)"
@@ -4717,7 +4729,11 @@
lemma list_ex_length:
"list_ex P xs \<longleftrightarrow> (\<exists>n < length xs. P (xs ! n))"
- unfolding list_ex_iff set_conv_nth by auto
+unfolding list_ex_iff set_conv_nth by auto
+
+lemma list_ex_cong[fundef_cong]:
+ "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> list_ex f xs = list_ex g ys"
+by (simp add: list_ex_iff)
lemma filtermap_conv:
"filtermap f xs = map (\<lambda>x. the (f x)) (filter (\<lambda>x. f x \<noteq> None) xs)"
@@ -4929,4 +4945,40 @@
"list_ex P [i..j] = (~ all_from_to_int (%x. ~P x) i j)"
by(simp add: all_from_to_int_iff_ball list_ex_iff)
+
+subsubsection {* Use convenient predefined operations *}
+
+code_const "op @"
+ (SML infixr 7 "@")
+ (OCaml infixr 6 "@")
+ (Haskell infixr 5 "++")
+ (Scala infixl 7 "++")
+
+code_const map
+ (Haskell "map")
+
+code_const filter
+ (Haskell "filter")
+
+code_const concat
+ (Haskell "concat")
+
+code_const rev
+ (Haskell "reverse")
+
+code_const zip
+ (Haskell "zip")
+
+code_const takeWhile
+ (Haskell "takeWhile")
+
+code_const dropWhile
+ (Haskell "dropWhile")
+
+code_const hd
+ (Haskell "head")
+
+code_const last
+ (Haskell "last")
+
end
--- a/src/HOL/MicroJava/BV/BVSpec.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/MicroJava/BV/BVSpec.thy Mon Jun 21 17:41:57 2010 +0200
@@ -15,48 +15,53 @@
to type \emph{checking}.
*}
-constdefs
+definition
-- "The program counter will always be inside the method:"
- check_bounded :: "instr list \<Rightarrow> exception_table \<Rightarrow> bool"
- "check_bounded ins et \<equiv>
+ check_bounded :: "instr list \<Rightarrow> exception_table \<Rightarrow> bool" where
+ "check_bounded ins et \<longleftrightarrow>
(\<forall>pc < length ins. \<forall>pc' \<in> set (succs (ins!pc) pc). pc' < length ins) \<and>
(\<forall>e \<in> set et. fst (snd (snd e)) < length ins)"
+definition
-- "The method type only contains declared classes:"
- check_types :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> JVMType.state list \<Rightarrow> bool"
- "check_types G mxs mxr phi \<equiv> set phi \<subseteq> states G mxs mxr"
+ check_types :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> JVMType.state list \<Rightarrow> bool" where
+ "check_types G mxs mxr phi \<longleftrightarrow> set phi \<subseteq> states G mxs mxr"
+definition
-- "An instruction is welltyped if it is applicable and its effect"
-- "is compatible with the type at all successor instructions:"
wt_instr :: "[instr,jvm_prog,ty,method_type,nat,p_count,
- exception_table,p_count] \<Rightarrow> bool"
- "wt_instr i G rT phi mxs max_pc et pc \<equiv>
+ exception_table,p_count] \<Rightarrow> bool" where
+ "wt_instr i G rT phi mxs max_pc et pc \<longleftrightarrow>
app i G mxs rT pc et (phi!pc) \<and>
(\<forall>(pc',s') \<in> set (eff i G pc et (phi!pc)). pc' < max_pc \<and> G \<turnstile> s' <=' phi!pc')"
+definition
-- {* The type at @{text "pc=0"} conforms to the method calling convention: *}
- wt_start :: "[jvm_prog,cname,ty list,nat,method_type] \<Rightarrow> bool"
- "wt_start G C pTs mxl phi ==
+ wt_start :: "[jvm_prog,cname,ty list,nat,method_type] \<Rightarrow> bool" where
+ "wt_start G C pTs mxl phi \<longleftrightarrow>
G \<turnstile> Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err)) <=' phi!0"
+definition
-- "A method is welltyped if the body is not empty, if execution does not"
-- "leave the body, if the method type covers all instructions and mentions"
-- "declared classes only, if the method calling convention is respected, and"
-- "if all instructions are welltyped."
wt_method :: "[jvm_prog,cname,ty list,ty,nat,nat,instr list,
- exception_table,method_type] \<Rightarrow> bool"
- "wt_method G C pTs rT mxs mxl ins et phi \<equiv>
- let max_pc = length ins in
+ exception_table,method_type] \<Rightarrow> bool" where
+ "wt_method G C pTs rT mxs mxl ins et phi \<longleftrightarrow>
+ (let max_pc = length ins in
0 < max_pc \<and>
length phi = length ins \<and>
check_bounded ins et \<and>
check_types G mxs (1+length pTs+mxl) (map OK phi) \<and>
wt_start G C pTs mxl phi \<and>
- (\<forall>pc. pc<max_pc \<longrightarrow> wt_instr (ins!pc) G rT phi mxs max_pc et pc)"
+ (\<forall>pc. pc<max_pc \<longrightarrow> wt_instr (ins!pc) G rT phi mxs max_pc et pc))"
+definition
-- "A program is welltyped if it is wellformed and all methods are welltyped"
- wt_jvm_prog :: "[jvm_prog,prog_type] \<Rightarrow> bool"
- "wt_jvm_prog G phi ==
+ wt_jvm_prog :: "[jvm_prog,prog_type] \<Rightarrow> bool" where
+ "wt_jvm_prog G phi \<longleftrightarrow>
wf_prog (\<lambda>G C (sig,rT,(maxs,maxl,b,et)).
wt_method G C (snd sig) rT maxs maxl b et (phi C sig)) G"
--- a/src/HOL/MicroJava/Comp/CorrCompTp.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/MicroJava/Comp/CorrCompTp.thy Mon Jun 21 17:41:57 2010 +0200
@@ -2372,9 +2372,6 @@
apply (simp add: check_type_simps)
apply (simp only: list_def)
apply (auto simp: err_def)
-apply (subgoal_tac "set (replicate mxl Err) \<subseteq> {Err}")
-apply blast
-apply (rule subset_replicate)
done
--- a/src/HOL/MicroJava/JVM/JVMDefensive.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/MicroJava/JVM/JVMDefensive.thy Mon Jun 21 17:41:57 2010 +0200
@@ -117,10 +117,10 @@
| Normal s' \<Rightarrow> if check G s' then Normal (exec (G, s')) else TypeError"
-constdefs
+definition
exec_all_d :: "jvm_prog \<Rightarrow> jvm_state type_error \<Rightarrow> jvm_state type_error \<Rightarrow> bool"
- ("_ |- _ -jvmd-> _" [61,61,61]60)
- "G |- s -jvmd-> t \<equiv>
+ ("_ |- _ -jvmd-> _" [61,61,61]60) where
+ "G |- s -jvmd-> t \<longleftrightarrow>
(s,t) \<in> ({(s,t). exec_d G s = TypeError \<and> t = TypeError} \<union>
{(s,t). \<exists>t'. exec_d G s = Normal (Some t') \<and> t = Normal t'})\<^sup>*"
--- a/src/HOL/Modelcheck/mucke_oracle.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Modelcheck/mucke_oracle.ML Mon Jun 21 17:41:57 2010 +0200
@@ -832,7 +832,7 @@
(* OUTPUT - relevant *)
(* transforms constructor term to a mucke-suitable output *)
-fun term_OUTPUT sg (Const("Pair",_) $ a $ b) =
+fun term_OUTPUT sg (Const(@{const_name Pair},_) $ a $ b) =
(term_OUTPUT sg a) ^ "_I_" ^ (term_OUTPUT sg b) |
term_OUTPUT sg (a $ b) = (term_OUTPUT sg a) ^ "_I_" ^ (term_OUTPUT sg b) |
term_OUTPUT sg (Const(s,t)) = post_last_dot s |
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Jun 21 17:41:57 2010 +0200
@@ -5266,7 +5266,7 @@
(* class constraint due to continuous_on_inverse *)
shows "compact s \<Longrightarrow> continuous_on s f \<Longrightarrow> (f ` s = t) \<Longrightarrow> inj_on f s
\<Longrightarrow> s homeomorphic t"
- unfolding homeomorphic_def by(metis homeomorphism_compact)
+ unfolding homeomorphic_def by (auto intro: homeomorphism_compact)
text{* Preservation of topological properties. *}
--- a/src/HOL/Nat.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Nat.thy Mon Jun 21 17:41:57 2010 +0200
@@ -39,16 +39,11 @@
Zero_RepI: "Nat Zero_Rep"
| Suc_RepI: "Nat i \<Longrightarrow> Nat (Suc_Rep i)"
-global
-
-typedef (open Nat)
- nat = Nat
+typedef (open Nat) nat = Nat
by (rule exI, unfold mem_def, rule Nat.Zero_RepI)
definition Suc :: "nat => nat" where
- Suc_def: "Suc == (%n. Abs_Nat (Suc_Rep (Rep_Nat n)))"
-
-local
+ "Suc n = Abs_Nat (Suc_Rep (Rep_Nat n))"
instantiation nat :: zero
begin
@@ -1208,9 +1203,9 @@
lemmas [code_unfold] = funpow_code_def [symmetric]
lemma [code]:
+ "funpow (Suc n) f = f o funpow n f"
"funpow 0 f = id"
- "funpow (Suc n) f = f o funpow n f"
- unfolding funpow_code_def by simp_all
+ by (simp_all add: funpow_code_def)
hide_const (open) funpow
@@ -1218,6 +1213,11 @@
"f ^^ (m + n) = f ^^ m \<circ> f ^^ n"
by (induct m) simp_all
+lemma funpow_mult:
+ fixes f :: "'a \<Rightarrow> 'a"
+ shows "(f ^^ m) ^^ n = f ^^ (m * n)"
+ by (induct n) (simp_all add: funpow_add)
+
lemma funpow_swap1:
"f ((f ^^ n) x) = (f ^^ n) (f x)"
proof -
@@ -1457,8 +1457,7 @@
end
lemma mono_iff_le_Suc: "mono f = (\<forall>n. f n \<le> f (Suc n))"
-unfolding mono_def
-by (auto intro:lift_Suc_mono_le[of f])
+ unfolding mono_def by (auto intro: lift_Suc_mono_le [of f])
lemma mono_nat_linear_lb:
"(!!m n::nat. m<n \<Longrightarrow> f m < f n) \<Longrightarrow> f(m)+k \<le> f(m+k)"
--- a/src/HOL/Nitpick.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Nitpick.thy Mon Jun 21 17:41:57 2010 +0200
@@ -214,6 +214,10 @@
definition inverse_frac :: "'a \<Rightarrow> 'a" where
"inverse_frac q \<equiv> frac (denom q) (num q)"
+definition less_frac :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
+[nitpick_simp]:
+"less_frac q r \<longleftrightarrow> num (plus_frac q (uminus_frac r)) < 0"
+
definition less_eq_frac :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
[nitpick_simp]:
"less_eq_frac q r \<longleftrightarrow> num (plus_frac q (uminus_frac r)) \<le> 0"
@@ -245,7 +249,7 @@
wf' wf_wfrec wf_wfrec' wfrec' card' setsum' fold_graph' nat_gcd nat_lcm
int_gcd int_lcm Frac Abs_Frac Rep_Frac zero_frac one_frac num denom
norm_frac frac plus_frac times_frac uminus_frac number_of_frac inverse_frac
- less_eq_frac of_frac
+ less_frac less_eq_frac of_frac
hide_type (open) bisim_iterator fin_fun fun_box pair_box unsigned_bit signed_bit
word
hide_fact (open) If_def Ex1_def split_def rtrancl_def rtranclp_def tranclp_def
@@ -254,6 +258,6 @@
list_size_simp nat_gcd_def nat_lcm_def int_gcd_def int_lcm_def Frac_def
zero_frac_def one_frac_def num_def denom_def norm_frac_def frac_def
plus_frac_def times_frac_def uminus_frac_def number_of_frac_def
- inverse_frac_def less_eq_frac_def of_frac_def
+ inverse_frac_def less_frac_def less_eq_frac_def of_frac_def
end
--- a/src/HOL/Nitpick_Examples/Core_Nits.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Nitpick_Examples/Core_Nits.thy Mon Jun 21 17:41:57 2010 +0200
@@ -1053,50 +1053,50 @@
nitpick_params [full_descrs, max_potential = 1]
-lemma "(THE j. j > Suc 10 \<and> j \<le> 11) \<noteq> 0"
-nitpick [card nat = 4, expect = potential]
-nitpick [card nat = 14, expect = potential] (* unfortunate *)
+lemma "(THE j. j > Suc 2 \<and> j \<le> 3) \<noteq> 0"
+nitpick [card nat = 2, expect = potential]
+nitpick [card nat = 6, expect = potential] (* unfortunate *)
oops
-lemma "(THE j. j > Suc 10 \<and> j \<le> 12) = x \<Longrightarrow> x \<noteq> 0"
-nitpick [card nat = 4, expect = potential]
-nitpick [card nat = 14, expect = none]
+lemma "(THE j. j > Suc 2 \<and> j \<le> 4) = x \<Longrightarrow> x \<noteq> 0"
+nitpick [card nat = 2, expect = potential]
+nitpick [card nat = 6, expect = none]
sorry
-lemma "(THE j. j > Suc 10 \<and> j \<le> 12) = x \<Longrightarrow> x = 12"
-nitpick [card nat = 4, expect = potential]
-nitpick [card nat = 14, expect = none]
+lemma "(THE j. j > Suc 2 \<and> j \<le> 4) = x \<Longrightarrow> x = 4"
+nitpick [card nat = 2, expect = potential]
+nitpick [card nat = 6, expect = none]
sorry
-lemma "(THE j. j > Suc 10 \<and> j \<le> 13) = x \<Longrightarrow> x = 12"
-nitpick [card nat = 14, expect = genuine]
+lemma "(THE j. j > Suc 2 \<and> j \<le> 5) = x \<Longrightarrow> x = 4"
+nitpick [card nat = 6, expect = genuine]
oops
-lemma "(THE j. j > Suc 10 \<and> j \<le> 13) = x \<Longrightarrow> x = 12 \<or> x = 13"
-nitpick [card nat = 14, expect = genuine]
+lemma "(THE j. j > Suc 2 \<and> j \<le> 5) = x \<Longrightarrow> x = 4 \<or> x = 5"
+nitpick [card nat = 6, expect = genuine]
oops
-lemma "(SOME j. j > Suc 10 \<and> j \<le> 11) \<noteq> 0"
-nitpick [card nat = 4, expect = potential]
-nitpick [card nat = 14, expect = genuine]
+lemma "(SOME j. j > Suc 2 \<and> j \<le> 3) \<noteq> 0"
+nitpick [card nat = 2, expect = potential]
+nitpick [card nat = 6, expect = genuine]
oops
-lemma "(SOME j. j > Suc 10 \<and> j \<le> 12) = x \<Longrightarrow> x \<noteq> 0"
-nitpick [card nat = 4, expect = potential]
-nitpick [card nat = 14, expect = none]
+lemma "(SOME j. j > Suc 2 \<and> j \<le> 4) = x \<Longrightarrow> x \<noteq> 0"
+nitpick [card nat = 2, expect = potential]
+nitpick [card nat = 6, expect = none]
oops
-lemma "(SOME j. j > Suc 10 \<and> j \<le> 12) = x \<Longrightarrow> x = 12"
-nitpick [card nat = 4, expect = potential]
-nitpick [card nat = 14, expect = none]
+lemma "(SOME j. j > Suc 2 \<and> j \<le> 4) = x \<Longrightarrow> x = 4"
+nitpick [card nat = 2, expect = potential]
+nitpick [card nat = 6, expect = none]
sorry
-lemma "(SOME j. j > Suc 10 \<and> j \<le> 13) = x \<Longrightarrow> x = 12"
-nitpick [card nat = 14, expect = genuine]
+lemma "(SOME j. j > Suc 2 \<and> j \<le> 5) = x \<Longrightarrow> x = 4"
+nitpick [card nat = 6, expect = genuine]
oops
-lemma "(SOME j. j > Suc 10 \<and> j \<le> 13) = x \<Longrightarrow> x = 12 \<or> x = 13"
-nitpick [card nat = 14, expect = none]
+lemma "(SOME j. j > Suc 2 \<and> j \<le> 5) = x \<Longrightarrow> x = 4 \<or> x = 5"
+nitpick [card nat = 6, expect = none]
sorry
nitpick_params [fast_descrs, max_potential = 0]
--- a/src/HOL/Nitpick_Examples/Integer_Nits.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Nitpick_Examples/Integer_Nits.thy Mon Jun 21 17:41:57 2010 +0200
@@ -187,7 +187,7 @@
lemma "y \<noteq> 0 \<Longrightarrow> x * y div y = (x::int)"
nitpick [unary_ints, expect = none]
-nitpick [binary_ints, expect = none, card = 1\<midarrow>5, bits = 1\<midarrow>5]
+nitpick [binary_ints, card = 1\<midarrow>4, bits = 1\<midarrow>4, expect = none]
sorry
lemma "(x * y < 0) \<longleftrightarrow> (x > 0 \<and> y < 0) \<or> (x < 0 \<and> y > (0::int))"
@@ -213,7 +213,7 @@
"labels (Node x t u) = {x} \<union> labels t \<union> labels u"
lemma "labels (Node x t u) \<noteq> labels (Node y v w)"
-nitpick [expect = genuine]
+nitpick [expect = potential] (* unfortunate *)
nitpick [dont_finitize, expect = potential]
oops
@@ -222,12 +222,12 @@
oops
lemma "card (labels t) > 0"
-nitpick [expect = genuine]
+nitpick [expect = potential] (* unfortunate *)
nitpick [dont_finitize, expect = potential]
oops
lemma "(\<Sum>n \<in> labels t. n + 2) \<ge> 2"
-nitpick [expect = genuine]
+nitpick [expect = potential] (* unfortunate *)
nitpick [dont_finitize, expect = potential]
oops
@@ -237,7 +237,7 @@
sorry
lemma "(\<Sum>i \<in> labels (Node x t u). f i\<Colon>nat) = f x + (\<Sum>i \<in> labels t. f i) + (\<Sum>i \<in> labels u. f i)"
-nitpick [expect = genuine]
+nitpick [expect = none] (* unfortunate *)
nitpick [dont_finitize, expect = none]
oops
--- a/src/HOL/Nitpick_Examples/Manual_Nits.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy Mon Jun 21 17:41:57 2010 +0200
@@ -7,13 +7,17 @@
header {* Examples from the Nitpick Manual *}
+(* The "expect" arguments to Nitpick in this theory and the other example
+ theories are there so that the example can also serve as a regression test
+ suite. *)
+
theory Manual_Nits
imports Main Quotient_Product RealDef
begin
chapter {* 3. First Steps *}
-nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1]
+nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60 s]
subsection {* 3.1. Propositional Logic *}
@@ -379,7 +383,7 @@
theorem S\<^isub>3_sound:
"w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-nitpick [card = 1\<midarrow>6, expect = none]
+nitpick [card = 1\<midarrow>5, expect = none]
sorry
theorem S\<^isub>3_complete:
@@ -398,19 +402,19 @@
theorem S\<^isub>4_sound:
"w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-nitpick [card = 1\<midarrow>6, expect = none]
+nitpick [card = 1\<midarrow>5, expect = none]
sorry
theorem S\<^isub>4_complete:
"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4"
-nitpick [card = 1\<midarrow>6, expect = none]
+nitpick [card = 1\<midarrow>5, expect = none]
sorry
theorem S\<^isub>4_A\<^isub>4_B\<^isub>4_sound_and_complete:
"w \<in> S\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
"w \<in> A\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] + 1"
"w \<in> B\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = b] = length [x \<leftarrow> w. x = a] + 1"
-nitpick [card = 1\<midarrow>6, expect = none]
+nitpick [card = 1\<midarrow>5, expect = none]
sorry
subsection {* 4.2. AA Trees *}
--- a/src/HOL/Nitpick_Examples/Mono_Nits.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Nitpick_Examples/Mono_Nits.thy Mon Jun 21 17:41:57 2010 +0200
@@ -76,8 +76,8 @@
ML {* const @{term "[a::'a set]"} *}
ML {* const @{term "[A \<union> (B::'a set)]"} *}
ML {* const @{term "[A \<union> (B::'a set)] = [C]"} *}
+ML {* const @{term "{(\<lambda>x::'a. x = a)} = C"} *}
-ML {* nonconst @{term "{(\<lambda>x::'a. x = a)} = C"} *}
ML {* nonconst @{term "\<forall>P (a::'a). P a"} *}
ML {* nonconst @{term "\<forall>a::'a. P a"} *}
ML {* nonconst @{term "(\<lambda>a::'a. \<not> A a) = B"} *}
--- a/src/HOL/Nitpick_Examples/Typedef_Nits.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Nitpick_Examples/Typedef_Nits.thy Mon Jun 21 17:41:57 2010 +0200
@@ -95,20 +95,20 @@
nitpick [expect = genuine]
oops
-lemma "Pair a b \<equiv> Abs_Prod (Pair_Rep a b)"
+lemma "Pair a b = Abs_prod (Pair_Rep a b)"
nitpick [card = 1\<midarrow>2, expect = none]
by (rule Pair_def)
-lemma "Pair a b \<equiv> Abs_Prod (Pair_Rep b a)"
+lemma "Pair a b = Abs_prod (Pair_Rep b a)"
nitpick [card = 1\<midarrow>2, expect = none]
nitpick [dont_box, expect = genuine]
oops
-lemma "fst (Abs_Prod (Pair_Rep a b)) = a"
+lemma "fst (Abs_prod (Pair_Rep a b)) = a"
nitpick [card = 2, expect = none]
-by (simp add: Pair_def [THEN symmetric])
+by (simp add: Pair_def [THEN sym])
-lemma "fst (Abs_Prod (Pair_Rep a b)) = b"
+lemma "fst (Abs_prod (Pair_Rep a b)) = b"
nitpick [card = 1\<midarrow>2, expect = none]
nitpick [dont_box, expect = genuine]
oops
@@ -117,19 +117,19 @@
nitpick [expect = none]
apply (rule ccontr)
apply simp
-apply (drule subst [where P = "\<lambda>r. Abs_Prod r = Abs_Prod (Pair_Rep a b)"])
+apply (drule subst [where P = "\<lambda>r. Abs_prod r = Abs_prod (Pair_Rep a b)"])
apply (rule refl)
-by (simp add: Pair_def [THEN symmetric])
+by (simp add: Pair_def [THEN sym])
-lemma "Abs_Prod (Rep_Prod a) = a"
+lemma "Abs_prod (Rep_prod a) = a"
nitpick [card = 2, expect = none]
-by (rule Rep_Prod_inverse)
+by (rule Rep_prod_inverse)
-lemma "Inl \<equiv> \<lambda>a. Abs_Sum (Inl_Rep a)"
+lemma "Inl \<equiv> \<lambda>a. Abs_sum (Inl_Rep a)"
nitpick [card = 1, expect = none]
-by (simp only: Inl_def o_def)
+by (simp add: Inl_def o_def)
-lemma "Inl \<equiv> \<lambda>a. Abs_Sum (Inr_Rep a)"
+lemma "Inl \<equiv> \<lambda>a. Abs_sum (Inr_Rep a)"
nitpick [card = 1, card "'a + 'a" = 2, expect = genuine]
oops
@@ -137,20 +137,20 @@
nitpick [expect = none]
by (rule Inl_Rep_not_Inr_Rep)
-lemma "Abs_Sum (Rep_Sum a) = a"
+lemma "Abs_sum (Rep_sum a) = a"
nitpick [card = 1, expect = none]
nitpick [card = 2, expect = none]
-by (rule Rep_Sum_inverse)
+by (rule Rep_sum_inverse)
lemma "0::nat \<equiv> Abs_Nat Zero_Rep"
nitpick [expect = none]
by (rule Zero_nat_def_raw)
-lemma "Suc \<equiv> \<lambda>n. Abs_Nat (Suc_Rep (Rep_Nat n))"
+lemma "Suc n = Abs_Nat (Suc_Rep (Rep_Nat n))"
nitpick [expect = none]
by (rule Suc_def)
-lemma "Suc \<equiv> \<lambda>n. Abs_Nat (Suc_Rep (Suc_Rep (Rep_Nat n)))"
+lemma "Suc n = Abs_Nat (Suc_Rep (Suc_Rep (Rep_Nat n)))"
nitpick [expect = genuine]
oops
@@ -174,9 +174,9 @@
Xcoord :: int
Ycoord :: int
-lemma "Abs_point_ext_type (Rep_point_ext_type a) = a"
+lemma "make_point_ext_type (dest_point_ext_type a) = a"
nitpick [expect = none]
-by (rule Rep_point_ext_type_inverse)
+by (rule dest_point_ext_type_inverse)
lemma "Fract a b = of_int a / of_int b"
nitpick [card = 1, expect = none]
--- a/src/HOL/Nominal/Examples/Support.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Nominal/Examples/Support.thy Mon Jun 21 17:41:57 2010 +0200
@@ -31,7 +31,7 @@
text {* An atom is either even or odd. *}
lemma even_or_odd:
- fixes n::"nat"
+ fixes n :: nat
shows "\<exists>i. (n = 2*i) \<or> (n=2*i+1)"
by (induct n) (presburger)+
--- a/src/HOL/Nominal/nominal_atoms.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Nominal/nominal_atoms.ML Mon Jun 21 17:41:57 2010 +0200
@@ -543,7 +543,7 @@
|> AxClass.prove_arity ("Nominal.noption",[[cls_name]],[cls_name]) (pt_proof pt_thm_noptn)
|> AxClass.prove_arity ("Option.option",[[cls_name]],[cls_name]) (pt_proof pt_thm_optn)
|> AxClass.prove_arity ("List.list",[[cls_name]],[cls_name]) (pt_proof pt_thm_list)
- |> AxClass.prove_arity ("*",[[cls_name],[cls_name]],[cls_name]) (pt_proof pt_thm_prod)
+ |> AxClass.prove_arity (@{type_name "*"},[[cls_name],[cls_name]],[cls_name]) (pt_proof pt_thm_prod)
|> AxClass.prove_arity ("Nominal.nprod",[[cls_name],[cls_name]],[cls_name])
(pt_proof pt_thm_nprod)
|> AxClass.prove_arity ("Product_Type.unit",[],[cls_name]) (pt_proof pt_thm_unit)
@@ -604,7 +604,7 @@
in
thy
|> AxClass.prove_arity ("Product_Type.unit",[],[cls_name]) (fs_proof fs_thm_unit)
- |> AxClass.prove_arity ("*",[[cls_name],[cls_name]],[cls_name]) (fs_proof fs_thm_prod)
+ |> AxClass.prove_arity (@{type_name "*"},[[cls_name],[cls_name]],[cls_name]) (fs_proof fs_thm_prod)
|> AxClass.prove_arity ("Nominal.nprod",[[cls_name],[cls_name]],[cls_name])
(fs_proof fs_thm_nprod)
|> AxClass.prove_arity ("List.list",[[cls_name]],[cls_name]) (fs_proof fs_thm_list)
@@ -686,7 +686,7 @@
in
thy'
|> AxClass.prove_arity ("Product_Type.unit",[],[cls_name]) (cp_proof cp_thm_unit)
- |> AxClass.prove_arity ("*",[[cls_name],[cls_name]],[cls_name]) (cp_proof cp_thm_prod)
+ |> AxClass.prove_arity (@{type_name "*"}, [[cls_name],[cls_name]],[cls_name]) (cp_proof cp_thm_prod)
|> AxClass.prove_arity ("List.list",[[cls_name]],[cls_name]) (cp_proof cp_thm_list)
|> AxClass.prove_arity ("fun",[[cls_name],[cls_name]],[cls_name]) (cp_proof cp_thm_fun)
|> AxClass.prove_arity ("Option.option",[[cls_name]],[cls_name]) (cp_proof cp_thm_optn)
--- a/src/HOL/Nominal/nominal_datatype.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Nominal/nominal_datatype.ML Mon Jun 21 17:41:57 2010 +0200
@@ -121,7 +121,7 @@
val dj_cp = thm "dj_cp";
-fun dest_permT (Type ("fun", [Type ("List.list", [Type ("*", [T, _])]),
+fun dest_permT (Type ("fun", [Type ("List.list", [Type (@{type_name "*"}, [T, _])]),
Type ("fun", [_, U])])) = (T, U);
fun permTs_of (Const ("Nominal.perm", T) $ t $ u) = fst (dest_permT T) :: permTs_of u
--- a/src/HOL/Nominal/nominal_permeq.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Nominal/nominal_permeq.ML Mon Jun 21 17:41:57 2010 +0200
@@ -103,7 +103,7 @@
case redex of
(* case pi o (f x) == (pi o f) (pi o x) *)
(Const("Nominal.perm",
- Type("fun",[Type("List.list",[Type("*",[Type(n,_),_])]),_])) $ pi $ (f $ x)) =>
+ Type("fun",[Type("List.list",[Type(@{type_name "*"},[Type(n,_),_])]),_])) $ pi $ (f $ x)) =>
(if (applicable_app f) then
let
val name = Long_Name.base_name n
@@ -190,8 +190,8 @@
fun perm_compose_simproc' sg ss redex =
(case redex of
(Const ("Nominal.perm", Type ("fun", [Type ("List.list",
- [Type ("*", [T as Type (tname,_),_])]),_])) $ pi1 $ (Const ("Nominal.perm",
- Type ("fun", [Type ("List.list", [Type ("*", [U as Type (uname,_),_])]),_])) $
+ [Type (@{type_name "*"}, [T as Type (tname,_),_])]),_])) $ pi1 $ (Const ("Nominal.perm",
+ Type ("fun", [Type ("List.list", [Type (@{type_name "*"}, [U as Type (uname,_),_])]),_])) $
pi2 $ t)) =>
let
val tname' = Long_Name.base_name tname
--- a/src/HOL/Nominal/nominal_thmdecls.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Nominal/nominal_thmdecls.ML Mon Jun 21 17:41:57 2010 +0200
@@ -103,7 +103,7 @@
let fun get_pi_aux s =
(case s of
(Const (@{const_name "perm"} ,typrm) $
- (Var (pi,typi as Type(@{type_name "list"}, [Type ("*", [Type (tyatm,[]),_])]))) $
+ (Var (pi,typi as Type(@{type_name "list"}, [Type (@{type_name "*"}, [Type (tyatm,[]),_])]))) $
(Var (n,ty))) =>
let
(* FIXME: this should be an operation the library *)
--- a/src/HOL/Product_Type.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Product_Type.thy Mon Jun 21 17:41:57 2010 +0200
@@ -9,6 +9,7 @@
imports Typedef Inductive Fun
uses
("Tools/split_rule.ML")
+ ("Tools/inductive_codegen.ML")
("Tools/inductive_set.ML")
begin
@@ -128,13 +129,10 @@
definition Pair_Rep :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool" where
"Pair_Rep a b = (\<lambda>x y. x = a \<and> y = b)"
-global
-
-typedef (Prod)
- ('a, 'b) "*" (infixr "*" 20)
- = "{f. \<exists>a b. f = Pair_Rep (a\<Colon>'a) (b\<Colon>'b)}"
+typedef (prod) ('a, 'b) "*" (infixr "*" 20)
+ = "{f. \<exists>a b. f = Pair_Rep (a\<Colon>'a) (b\<Colon>'b)}"
proof
- fix a b show "Pair_Rep a b \<in> ?Prod"
+ fix a b show "Pair_Rep a b \<in> ?prod"
by rule+
qed
@@ -143,35 +141,29 @@
type_notation (HTML output)
"*" ("(_ \<times>/ _)" [21, 20] 20)
-consts
- Pair :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<times> 'b"
-
-local
-
-defs
- Pair_def: "Pair a b == Abs_Prod (Pair_Rep a b)"
+definition Pair :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<times> 'b" where
+ "Pair a b = Abs_prod (Pair_Rep a b)"
rep_datatype (prod) Pair proof -
fix P :: "'a \<times> 'b \<Rightarrow> bool" and p
assume "\<And>a b. P (Pair a b)"
- then show "P p" by (cases p) (auto simp add: Prod_def Pair_def Pair_Rep_def)
+ then show "P p" by (cases p) (auto simp add: prod_def Pair_def Pair_Rep_def)
next
fix a c :: 'a and b d :: 'b
have "Pair_Rep a b = Pair_Rep c d \<longleftrightarrow> a = c \<and> b = d"
by (auto simp add: Pair_Rep_def expand_fun_eq)
- moreover have "Pair_Rep a b \<in> Prod" and "Pair_Rep c d \<in> Prod"
- by (auto simp add: Prod_def)
+ moreover have "Pair_Rep a b \<in> prod" and "Pair_Rep c d \<in> prod"
+ by (auto simp add: prod_def)
ultimately show "Pair a b = Pair c d \<longleftrightarrow> a = c \<and> b = d"
- by (simp add: Pair_def Abs_Prod_inject)
+ by (simp add: Pair_def Abs_prod_inject)
qed
+declare weak_case_cong [cong del]
+
subsubsection {* Tuple syntax *}
-global consts
- split :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c"
-
-local defs
+definition split :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c" where
split_prod_case: "split == prod_case"
text {*
@@ -393,13 +385,11 @@
lemma surj_pair [simp]: "EX x y. p = (x, y)"
by (cases p) simp
-global consts
- fst :: "'a \<times> 'b \<Rightarrow> 'a"
- snd :: "'a \<times> 'b \<Rightarrow> 'b"
+definition fst :: "'a \<times> 'b \<Rightarrow> 'a" where
+ "fst p = (case p of (a, b) \<Rightarrow> a)"
-local defs
- fst_def: "fst p == case p of (a, b) \<Rightarrow> a"
- snd_def: "snd p == case p of (a, b) \<Rightarrow> b"
+definition snd :: "'a \<times> 'b \<Rightarrow> 'b" where
+ "snd p = (case p of (a, b) \<Rightarrow> b)"
lemma fst_conv [simp, code]: "fst (a, b) = a"
unfolding fst_def by simp
@@ -788,11 +778,8 @@
subsubsection {* Derived operations *}
-global consts
- curry :: "('a \<times> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'c"
-
-local defs
- curry_def: "curry == (%c x y. c (Pair x y))"
+definition curry :: "('a \<times> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'c" where
+ "curry = (\<lambda>c x y. c (x, y))"
lemma curry_conv [simp, code]: "curry f a b = f (a, b)"
by (simp add: curry_def)
@@ -1192,6 +1179,9 @@
subsection {* Inductively defined sets *}
+use "Tools/inductive_codegen.ML"
+setup Inductive_Codegen.setup
+
use "Tools/inductive_set.ML"
setup Inductive_Set.setup
--- a/src/HOL/Rat.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Rat.thy Mon Jun 21 17:41:57 2010 +0200
@@ -1182,15 +1182,16 @@
(@{const_name uminus_rat_inst.uminus_rat}, @{const_name Nitpick.uminus_frac}),
(@{const_name number_rat_inst.number_of_rat}, @{const_name Nitpick.number_of_frac}),
(@{const_name inverse_rat_inst.inverse_rat}, @{const_name Nitpick.inverse_frac}),
+ (@{const_name ord_rat_inst.less_rat}, @{const_name Nitpick.less_frac}),
(@{const_name ord_rat_inst.less_eq_rat}, @{const_name Nitpick.less_eq_frac}),
(@{const_name field_char_0_class.of_rat}, @{const_name Nitpick.of_frac}),
(@{const_name field_char_0_class.Rats}, @{const_abbrev UNIV})]
*}
lemmas [nitpick_def] = inverse_rat_inst.inverse_rat
- number_rat_inst.number_of_rat one_rat_inst.one_rat ord_rat_inst.less_eq_rat
- plus_rat_inst.plus_rat times_rat_inst.times_rat uminus_rat_inst.uminus_rat
- zero_rat_inst.zero_rat
+ number_rat_inst.number_of_rat one_rat_inst.one_rat ord_rat_inst.less_rat
+ ord_rat_inst.less_eq_rat plus_rat_inst.plus_rat times_rat_inst.times_rat
+ uminus_rat_inst.uminus_rat zero_rat_inst.zero_rat
subsection{* Float syntax *}
--- a/src/HOL/RealDef.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/RealDef.thy Mon Jun 21 17:41:57 2010 +0200
@@ -1806,12 +1806,13 @@
(@{const_name uminus_real_inst.uminus_real}, @{const_name Nitpick.uminus_frac}),
(@{const_name number_real_inst.number_of_real}, @{const_name Nitpick.number_of_frac}),
(@{const_name inverse_real_inst.inverse_real}, @{const_name Nitpick.inverse_frac}),
+ (@{const_name ord_real_inst.less_real}, @{const_name Nitpick.less_eq_frac}),
(@{const_name ord_real_inst.less_eq_real}, @{const_name Nitpick.less_eq_frac})]
*}
lemmas [nitpick_def] = inverse_real_inst.inverse_real
number_real_inst.number_of_real one_real_inst.one_real
- ord_real_inst.less_eq_real plus_real_inst.plus_real
+ ord_real_inst.less_real ord_real_inst.less_eq_real plus_real_inst.plus_real
times_real_inst.times_real uminus_real_inst.uminus_real
zero_real_inst.zero_real
--- a/src/HOL/Set.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Set.thy Mon Jun 21 17:41:57 2010 +0200
@@ -151,17 +151,11 @@
supset_eq ("op \<supseteq>") and
supset_eq ("(_/ \<supseteq> _)" [50, 51] 50)
-global
-
-consts
- Ball :: "'a set => ('a => bool) => bool" -- "bounded universal quantifiers"
- Bex :: "'a set => ('a => bool) => bool" -- "bounded existential quantifiers"
+definition Ball :: "'a set \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
+ "Ball A P \<longleftrightarrow> (\<forall>x. x \<in> A \<longrightarrow> P x)" -- "bounded universal quantifiers"
-local
-
-defs
- Ball_def: "Ball A P == ALL x. x:A --> P(x)"
- Bex_def: "Bex A P == EX x. x:A & P(x)"
+definition Bex :: "'a set \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
+ "Bex A P \<longleftrightarrow> (\<exists>x. x \<in> A \<and> P x)" -- "bounded existential quantifiers"
syntax
"_Ball" :: "pttrn => 'a set => bool => bool" ("(3ALL _:_./ _)" [0, 0, 10] 10)
--- a/src/HOL/SetInterval.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/SetInterval.thy Mon Jun 21 17:41:57 2010 +0200
@@ -770,7 +770,7 @@
qed
lemma card_less_Suc2: "0 \<notin> M \<Longrightarrow> card {k. Suc k \<in> M \<and> k < i} = card {k \<in> M. k < Suc i}"
-apply (rule card_bij_eq [of "Suc" _ _ "\<lambda>x. x - Suc 0"])
+apply (rule card_bij_eq [of Suc _ _ "\<lambda>x. x - Suc 0"])
apply simp
apply fastsimp
apply auto
--- a/src/HOL/Sledgehammer.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Sledgehammer.thy Mon Jun 21 17:41:57 2010 +0200
@@ -25,6 +25,9 @@
("Tools/Sledgehammer/metis_tactics.ML")
begin
+definition skolem_id :: "'a \<Rightarrow> 'a" where
+[no_atp]: "skolem_id = (\<lambda>x. x)"
+
definition COMBI :: "'a \<Rightarrow> 'a" where
[no_atp]: "COMBI P \<equiv> P"
--- a/src/HOL/Sum_Type.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Sum_Type.thy Mon Jun 21 17:41:57 2010 +0200
@@ -17,21 +17,17 @@
definition Inr_Rep :: "'b \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool \<Rightarrow> bool" where
"Inr_Rep b x y p \<longleftrightarrow> y = b \<and> \<not> p"
-global
-
-typedef (Sum) ('a, 'b) "+" (infixr "+" 10) = "{f. (\<exists>a. f = Inl_Rep (a::'a)) \<or> (\<exists>b. f = Inr_Rep (b::'b))}"
+typedef (sum) ('a, 'b) "+" (infixr "+" 10) = "{f. (\<exists>a. f = Inl_Rep (a::'a)) \<or> (\<exists>b. f = Inr_Rep (b::'b))}"
by auto
-local
-
-lemma Inl_RepI: "Inl_Rep a \<in> Sum"
- by (auto simp add: Sum_def)
+lemma Inl_RepI: "Inl_Rep a \<in> sum"
+ by (auto simp add: sum_def)
-lemma Inr_RepI: "Inr_Rep b \<in> Sum"
- by (auto simp add: Sum_def)
+lemma Inr_RepI: "Inr_Rep b \<in> sum"
+ by (auto simp add: sum_def)
-lemma inj_on_Abs_Sum: "A \<subseteq> Sum \<Longrightarrow> inj_on Abs_Sum A"
- by (rule inj_on_inverseI, rule Abs_Sum_inverse) auto
+lemma inj_on_Abs_sum: "A \<subseteq> sum \<Longrightarrow> inj_on Abs_sum A"
+ by (rule inj_on_inverseI, rule Abs_sum_inverse) auto
lemma Inl_Rep_inject: "inj_on Inl_Rep A"
proof (rule inj_onI)
@@ -49,28 +45,28 @@
by (auto simp add: Inl_Rep_def Inr_Rep_def expand_fun_eq)
definition Inl :: "'a \<Rightarrow> 'a + 'b" where
- "Inl = Abs_Sum \<circ> Inl_Rep"
+ "Inl = Abs_sum \<circ> Inl_Rep"
definition Inr :: "'b \<Rightarrow> 'a + 'b" where
- "Inr = Abs_Sum \<circ> Inr_Rep"
+ "Inr = Abs_sum \<circ> Inr_Rep"
lemma inj_Inl [simp]: "inj_on Inl A"
-by (auto simp add: Inl_def intro!: comp_inj_on Inl_Rep_inject inj_on_Abs_Sum Inl_RepI)
+by (auto simp add: Inl_def intro!: comp_inj_on Inl_Rep_inject inj_on_Abs_sum Inl_RepI)
lemma Inl_inject: "Inl x = Inl y \<Longrightarrow> x = y"
using inj_Inl by (rule injD)
lemma inj_Inr [simp]: "inj_on Inr A"
-by (auto simp add: Inr_def intro!: comp_inj_on Inr_Rep_inject inj_on_Abs_Sum Inr_RepI)
+by (auto simp add: Inr_def intro!: comp_inj_on Inr_Rep_inject inj_on_Abs_sum Inr_RepI)
lemma Inr_inject: "Inr x = Inr y \<Longrightarrow> x = y"
using inj_Inr by (rule injD)
lemma Inl_not_Inr: "Inl a \<noteq> Inr b"
proof -
- from Inl_RepI [of a] Inr_RepI [of b] have "{Inl_Rep a, Inr_Rep b} \<subseteq> Sum" by auto
- with inj_on_Abs_Sum have "inj_on Abs_Sum {Inl_Rep a, Inr_Rep b}" .
- with Inl_Rep_not_Inr_Rep [of a b] inj_on_contraD have "Abs_Sum (Inl_Rep a) \<noteq> Abs_Sum (Inr_Rep b)" by auto
+ from Inl_RepI [of a] Inr_RepI [of b] have "{Inl_Rep a, Inr_Rep b} \<subseteq> sum" by auto
+ with inj_on_Abs_sum have "inj_on Abs_sum {Inl_Rep a, Inr_Rep b}" .
+ with Inl_Rep_not_Inr_Rep [of a b] inj_on_contraD have "Abs_sum (Inl_Rep a) \<noteq> Abs_sum (Inr_Rep b)" by auto
then show ?thesis by (simp add: Inl_def Inr_def)
qed
@@ -81,10 +77,10 @@
assumes "\<And>x::'a. s = Inl x \<Longrightarrow> P"
and "\<And>y::'b. s = Inr y \<Longrightarrow> P"
shows P
-proof (rule Abs_Sum_cases [of s])
+proof (rule Abs_sum_cases [of s])
fix f
- assume "s = Abs_Sum f" and "f \<in> Sum"
- with assms show P by (auto simp add: Sum_def Inl_def Inr_def)
+ assume "s = Abs_sum f" and "f \<in> sum"
+ with assms show P by (auto simp add: sum_def Inl_def Inr_def)
qed
rep_datatype (sum) Inl Inr
--- a/src/HOL/Tools/ATP_Manager/atp_manager.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/ATP_Manager/atp_manager.ML Mon Jun 21 17:41:57 2010 +0200
@@ -34,8 +34,8 @@
axiom_clauses: (thm * (string * int)) list option,
filtered_clauses: (thm * (string * int)) list option}
datatype failure =
- Unprovable | TimedOut | OutOfResources | OldSpass | MalformedInput |
- MalformedOutput | UnknownError
+ Unprovable | IncompleteUnprovable | TimedOut | OutOfResources | OldSpass |
+ MalformedInput | MalformedOutput | UnknownError
type prover_result =
{outcome: failure option,
message: string,
@@ -95,8 +95,8 @@
filtered_clauses: (thm * (string * int)) list option};
datatype failure =
- Unprovable | TimedOut | OutOfResources | OldSpass | MalformedInput |
- MalformedOutput | UnknownError
+ Unprovable | IncompleteUnprovable | TimedOut | OutOfResources | OldSpass |
+ MalformedInput | MalformedOutput | UnknownError
type prover_result =
{outcome: failure option,
@@ -336,8 +336,9 @@
(* start prover thread *)
-fun start_prover_thread (params as {timeout, ...}) birth_time death_time i n
- relevance_override minimize_command proof_state name =
+fun start_prover_thread (params as {full_types, timeout, ...}) birth_time
+ death_time i n relevance_override minimize_command
+ proof_state name =
let
val prover = get_prover (Proof.theory_of proof_state) name
val {context = ctxt, facts, goal} = Proof.goal proof_state;
@@ -353,7 +354,7 @@
filtered_clauses = NONE}
val message =
#message (prover params (minimize_command name) timeout problem)
- handle Sledgehammer_HOL_Clause.TRIVIAL => metis_line i n []
+ handle Sledgehammer_HOL_Clause.TRIVIAL => metis_line full_types i n []
| ERROR message => "Error: " ^ message ^ "\n"
val _ = unregister params message (Thread.self ());
in () end)
--- a/src/HOL/Tools/ATP_Manager/atp_systems.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/ATP_Manager/atp_systems.ML Mon Jun 21 17:41:57 2010 +0200
@@ -87,6 +87,8 @@
| (failure :: _) => ("", SOME failure)
fun string_for_failure Unprovable = "The ATP problem is unprovable."
+ | string_for_failure IncompleteUnprovable =
+ "The ATP cannot prove the problem."
| string_for_failure TimedOut = "Timed out."
| string_for_failure OutOfResources = "The ATP ran out of resources."
| string_for_failure OldSpass =
@@ -220,9 +222,9 @@
case outcome of
NONE =>
proof_text isar_proof
- (pool, debug, full_types, isar_shrink_factor, ctxt,
- conjecture_shape)
- (minimize_command, proof, internal_thm_names, th, subgoal)
+ (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
+ (full_types, minimize_command, proof, internal_thm_names, th,
+ subgoal)
| SOME failure => (string_for_failure failure ^ "\n", [])
in
{outcome = outcome, message = message, pool = pool,
@@ -247,7 +249,7 @@
(relevant_facts full_types respect_no_atp relevance_threshold
relevance_convergence defs_relevant max_axiom_clauses
(the_default prefers_theory_relevant theory_relevant))
- (prepare_clauses false)
+ (prepare_clauses false full_types)
(write_tptp_file (debug andalso overlord)) home_var
executable (arguments timeout) proof_delims known_failures name params
minimize_command
@@ -274,9 +276,9 @@
"======= End of refutation ======="),
("% SZS output start Refutation", "% SZS output end Refutation")],
known_failures =
- [(Unprovable, "Satisfiability detected"),
- (Unprovable, "UNPROVABLE"),
- (Unprovable, "CANNOT PROVE"),
+ [(Unprovable, "UNPROVABLE"),
+ (IncompleteUnprovable, "CANNOT PROVE"),
+ (Unprovable, "Satisfiability detected"),
(OutOfResources, "Refutation not found")],
max_axiom_clauses = 60,
prefers_theory_relevant = false}
@@ -321,7 +323,7 @@
(relevant_facts full_types respect_no_atp relevance_threshold
relevance_convergence defs_relevant max_axiom_clauses
(the_default prefers_theory_relevant theory_relevant))
- (prepare_clauses true) write_dfg_file home_var executable
+ (prepare_clauses true full_types) write_dfg_file home_var executable
(arguments timeout) proof_delims known_failures name params
minimize_command
@@ -329,39 +331,29 @@
(* The "-VarWeight=3" option helps the higher-order problems, probably by
counteracting the presence of "hAPP". *)
-val spass_config : prover_config =
+fun generic_spass_config dfg : prover_config =
{home_var = "SPASS_HOME",
executable = "SPASS",
arguments = fn timeout =>
"-Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof \
- \-VarWeight=3 -TimeLimit=" ^ string_of_int (to_generous_secs timeout),
+ \-VarWeight=3 -TimeLimit=" ^ string_of_int (to_generous_secs timeout)
+ |> not dfg ? prefix "-TPTP ",
proof_delims = [("Here is a proof", "Formulae used in the proof")],
known_failures =
- [(Unprovable, "SPASS beiseite: Completion found"),
+ [(IncompleteUnprovable, "SPASS beiseite: Completion found"),
(TimedOut, "SPASS beiseite: Ran out of time"),
(OutOfResources, "SPASS beiseite: Maximal number of loops exceeded"),
- (MalformedInput, "Undefined symbol")],
+ (MalformedInput, "Undefined symbol"),
+ (MalformedInput, "Free Variable"),
+ (OldSpass, "unrecognized option `-TPTP'"),
+ (OldSpass, "Unrecognized option TPTP")],
max_axiom_clauses = 40,
prefers_theory_relevant = true}
-val spass = dfg_prover "spass" spass_config
-
-(* SPASS 3.7 supports both the DFG and the TPTP syntax, whereas SPASS 3.0
- supports only the DFG syntax. As soon as all Isabelle repository/snapshot
- users have upgraded to 3.7, we can kill "spass" (and all DFG support in
- Sledgehammer) and rename "spass_tptp" "spass". *)
+val spass_dfg_config = generic_spass_config true
+val spass_dfg = dfg_prover "spass_dfg" spass_dfg_config
-val spass_tptp_config =
- {home_var = #home_var spass_config,
- executable = #executable spass_config,
- arguments = prefix "-TPTP " o #arguments spass_config,
- proof_delims = #proof_delims spass_config,
- known_failures =
- #known_failures spass_config @
- [(OldSpass, "unrecognized option `-TPTP'"),
- (OldSpass, "Unrecognized option TPTP")],
- max_axiom_clauses = #max_axiom_clauses spass_config,
- prefers_theory_relevant = #prefers_theory_relevant spass_config}
-val spass_tptp = tptp_prover "spass_tptp" spass_tptp_config
+val spass_config = generic_spass_config false
+val spass = tptp_prover "spass" spass_config
(* remote prover invocation via SystemOnTPTP *)
@@ -423,7 +415,7 @@
remotify (fst vampire)]
val provers =
- [spass, spass_tptp, vampire, e, remote_vampire, remote_spass, remote_e]
+ [spass, spass_dfg, vampire, e, remote_vampire, remote_spass, remote_e]
val prover_setup = fold add_prover provers
val setup =
--- a/src/HOL/Tools/Datatype/datatype_codegen.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_codegen.ML Mon Jun 21 17:41:57 2010 +0200
@@ -105,12 +105,14 @@
in (thm', lthy') end;
fun tac thms = Class.intro_classes_tac []
THEN ALLGOALS (ProofContext.fact_tac thms);
+ fun prefix dtco = Binding.qualify true (Long_Name.base_name dtco) o Binding.qualify true "eq" o Binding.name;
fun add_eq_thms dtco =
Theory.checkpoint
#> `(fn thy => mk_eq_eqns thy dtco)
- #-> (fn (thms, thm) =>
- Code.add_nbe_eqn thm
- #> fold_rev Code.add_eqn thms);
+ #-> (fn (thms, thm) => PureThy.note_thmss Thm.lemmaK
+ [((prefix dtco "refl", [Code.add_nbe_default_eqn_attribute]), [([thm], [])]),
+ ((prefix dtco "simps", [Code.add_default_eqn_attribute]), [(rev thms, [])])])
+ #> snd
in
thy
|> Theory_Target.instantiation (dtcos, vs, [HOLogic.class_eq])
--- a/src/HOL/Tools/Function/function_lib.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Function/function_lib.ML Mon Jun 21 17:41:57 2010 +0200
@@ -24,7 +24,7 @@
case vars of
(Free v) => lambda (Free v) t
| (Var v) => lambda (Var v) t
- | (Const ("Pair", Type ("fun", [Ta, Type ("fun", [Tb, _])]))) $ us $ vs =>
+ | (Const (@{const_name Pair}, Type ("fun", [Ta, Type ("fun", [Tb, _])]))) $ us $ vs =>
(HOLogic.split_const (Ta,Tb, fastype_of t)) $ (tupled_lambda us (tupled_lambda vs t))
| _ => raise Match
--- a/src/HOL/Tools/Function/measure_functions.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Function/measure_functions.ML Mon Jun 21 17:41:57 2010 +0200
@@ -39,17 +39,17 @@
fun constant_0 T = Abs ("x", T, HOLogic.zero)
fun constant_1 T = Abs ("x", T, HOLogic.Suc_zero)
-fun mk_funorder_funs (Type ("+", [fT, sT])) =
+fun mk_funorder_funs (Type (@{type_name "+"}, [fT, sT])) =
map (fn m => SumTree.mk_sumcase fT sT HOLogic.natT m (constant_0 sT)) (mk_funorder_funs fT)
@ map (fn m => SumTree.mk_sumcase fT sT HOLogic.natT (constant_0 fT) m) (mk_funorder_funs sT)
| mk_funorder_funs T = [ constant_1 T ]
-fun mk_ext_base_funs ctxt (Type ("+", [fT, sT])) =
+fun mk_ext_base_funs ctxt (Type (@{type_name "+"}, [fT, sT])) =
map_product (SumTree.mk_sumcase fT sT HOLogic.natT)
(mk_ext_base_funs ctxt fT) (mk_ext_base_funs ctxt sT)
| mk_ext_base_funs ctxt T = find_measures ctxt T
-fun mk_all_measure_funs ctxt (T as Type ("+", _)) =
+fun mk_all_measure_funs ctxt (T as Type (@{type_name "+"}, _)) =
mk_ext_base_funs ctxt T @ mk_funorder_funs T
| mk_all_measure_funs ctxt T = find_measures ctxt T
--- a/src/HOL/Tools/Function/sum_tree.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Function/sum_tree.ML Mon Jun 21 17:41:57 2010 +0200
@@ -17,7 +17,7 @@
{left = (fn f => f o left), right = (fn f => f o right), init = I} len i init
(* Sum types *)
-fun mk_sumT LT RT = Type ("+", [LT, RT])
+fun mk_sumT LT RT = Type (@{type_name "+"}, [LT, RT])
fun mk_sumcase TL TR T l r =
Const (@{const_name sum.sum_case},
(TL --> T) --> (TR --> T) --> mk_sumT TL TR --> T) $ l $ r
@@ -27,18 +27,18 @@
fun mk_inj ST n i =
access_top_down
{ init = (ST, I : term -> term),
- left = (fn (T as Type ("+", [LT, RT]), inj) =>
+ left = (fn (T as Type (@{type_name "+"}, [LT, RT]), inj) =>
(LT, inj o App (Const (@{const_name Inl}, LT --> T)))),
- right =(fn (T as Type ("+", [LT, RT]), inj) =>
+ right =(fn (T as Type (@{type_name "+"}, [LT, RT]), inj) =>
(RT, inj o App (Const (@{const_name Inr}, RT --> T))))} n i
|> snd
fun mk_proj ST n i =
access_top_down
{ init = (ST, I : term -> term),
- left = (fn (T as Type ("+", [LT, RT]), proj) =>
+ left = (fn (T as Type (@{type_name "+"}, [LT, RT]), proj) =>
(LT, App (Const (@{const_name Sum_Type.Projl}, T --> LT)) o proj)),
- right =(fn (T as Type ("+", [LT, RT]), proj) =>
+ right =(fn (T as Type (@{type_name "+"}, [LT, RT]), proj) =>
(RT, App (Const (@{const_name Sum_Type.Projr}, T --> RT)) o proj))} n i
|> snd
--- a/src/HOL/Tools/Function/termination.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Function/termination.ML Mon Jun 21 17:41:57 2010 +0200
@@ -106,7 +106,7 @@
end
(* compute list of types for nodes *)
-fun node_types sk T = dest_tree (fn Type ("+", [LT, RT]) => (LT, RT)) sk T |> map snd
+fun node_types sk T = dest_tree (fn Type (@{type_name "+"}, [LT, RT]) => (LT, RT)) sk T |> map snd
(* find index and raw term *)
fun dest_inj (SLeaf i) trm = (i, trm)
@@ -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 ("op &", _) $ (Const ("op =", _) $ _ $ (Const ("Pair", _) $ r $ l)) $ _)
+ val (Const ("op &", _) $ (Const ("op =", _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ _)
= Term.strip_qnt_body "Ex" c
in cons r o cons l end
in
@@ -185,7 +185,7 @@
val vs = Term.strip_qnt_vars "Ex" c
(* FIXME: throw error "dest_call" for malformed terms *)
- val (Const ("op &", _) $ (Const ("op =", _) $ _ $ (Const ("Pair", _) $ r $ l)) $ Gam)
+ val (Const ("op &", _) $ (Const ("op =", _) $ _ $ (Const (@{const_name Pair}, _) $ r $ l)) $ Gam)
= Term.strip_qnt_body "Ex" c
val (p, l') = dest_inj sk l
val (q, r') = dest_inj sk r
--- a/src/HOL/Tools/Nitpick/minipick.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Nitpick/minipick.ML Mon Jun 21 17:41:57 2010 +0200
@@ -80,7 +80,7 @@
fun kodkod_formula_from_term ctxt card frees =
let
- fun R_rep_from_S_rep (T as Type (@{type_name fun}, [T1, @{typ bool}])) r =
+ fun R_rep_from_S_rep (Type (@{type_name fun}, [T1, @{typ bool}])) r =
let
val jss = atom_schema_of SRep card T1 |> map (rpair 0)
|> all_combinations
@@ -115,7 +115,7 @@
@{const Not} $ t1 => Not (to_F Ts t1)
| @{const False} => False
| @{const True} => True
- | Const (@{const_name All}, _) $ Abs (s, T, t') =>
+ | Const (@{const_name All}, _) $ Abs (_, T, t') =>
All (decls_for SRep card Ts T, to_F (T :: Ts) t')
| (t0 as Const (@{const_name All}, _)) $ t1 =>
to_F Ts (t0 $ eta_expand Ts t1 1)
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Mon Jun 21 17:41:57 2010 +0200
@@ -64,6 +64,10 @@
val iter_var_prefix : string
val strip_first_name_sep : string -> string * string
val original_name : string -> string
+ val abs_var : indexname * typ -> term -> term
+ val s_let : string -> int -> typ -> typ -> (term -> term) -> term -> term
+ val s_betapply : typ list -> term * term -> term
+ val s_betapplys : typ list -> term * term list -> term
val s_conj : term * term -> term
val s_disj : term * term -> term
val strip_any_connective : term -> term list * term
@@ -162,7 +166,6 @@
val is_finite_type : hol_context -> typ -> bool
val is_small_finite_type : hol_context -> typ -> bool
val special_bounds : term list -> (indexname * typ) list
- val abs_var : indexname * typ -> term -> term
val is_funky_typedef : theory -> typ -> bool
val all_axioms_of :
Proof.context -> (term * term) list -> term list * term list * term list
@@ -302,10 +305,55 @@
else
s
-fun s_betapply (Const (@{const_name If}, _) $ @{const True} $ t, _) = t
- | s_betapply (Const (@{const_name If}, _) $ @{const False} $ _, t) = t
- | s_betapply p = betapply p
-val s_betapplys = Library.foldl s_betapply
+fun abs_var ((s, j), T) body = Abs (s, T, abstract_over (Var ((s, j), T), body))
+
+fun let_var s = (nitpick_prefix ^ s, 999)
+val let_inline_threshold = 20
+
+fun s_let s n abs_T body_T f t =
+ if (n - 1) * (size_of_term t - 1) <= let_inline_threshold then
+ f t
+ else
+ let val z = (let_var s, abs_T) in
+ Const (@{const_name Let}, abs_T --> (abs_T --> body_T) --> body_T)
+ $ t $ abs_var z (incr_boundvars 1 (f (Var z)))
+ end
+
+fun loose_bvar1_count (Bound i, k) = if i = k then 1 else 0
+ | loose_bvar1_count (t1 $ t2, k) =
+ loose_bvar1_count (t1, k) + loose_bvar1_count (t2, k)
+ | loose_bvar1_count (Abs (_, _, t), k) = loose_bvar1_count (t, k + 1)
+ | loose_bvar1_count _ = 0
+
+fun s_betapply _ (Const (@{const_name If}, _) $ @{const True} $ t1', _) = t1'
+ | s_betapply _ (Const (@{const_name If}, _) $ @{const False} $ _, t2) = t2
+ | s_betapply Ts (Const (@{const_name Let},
+ Type (_, [bound_T, Type (_, [_, body_T])]))
+ $ t12 $ Abs (s, T, t13'), t2) =
+ let val body_T' = range_type body_T in
+ Const (@{const_name Let}, bound_T --> (bound_T --> body_T') --> body_T')
+ $ t12 $ Abs (s, T, s_betapply (T :: Ts) (t13', incr_boundvars 1 t2))
+ end
+ | s_betapply Ts (t1 as Abs (s1, T1, t1'), t2) =
+ (s_let s1 (loose_bvar1_count (t1', 0)) T1 (fastype_of1 (T1 :: Ts, t1'))
+ (curry betapply t1) t2
+ handle TERM _ => betapply (t1, t2)) (* FIXME: fix all uses *)
+ | s_betapply _ (t1, t2) = t1 $ t2
+fun s_betapplys Ts = Library.foldl (s_betapply Ts)
+
+fun s_beta_norm Ts t =
+ let
+ fun aux _ (Var _) = raise Same.SAME
+ | aux Ts (Abs (s, T, t')) = Abs (s, T, aux (T :: Ts) t')
+ | aux Ts ((t1 as Abs _) $ t2) =
+ Same.commit (aux Ts) (s_betapply Ts (t1, t2))
+ | aux Ts (t1 $ t2) =
+ ((case aux Ts t1 of
+ t1 as Abs _ => Same.commit (aux Ts) (s_betapply Ts (t1, t2))
+ | t1 => t1 $ Same.commit (aux Ts) t2)
+ handle Same.SAME => t1 $ aux Ts t2)
+ | aux _ _ = raise Same.SAME
+ in aux Ts t handle Same.SAME => t end
fun s_conj (t1, @{const True}) = t1
| s_conj (@{const True}, t2) = t2
@@ -344,7 +392,7 @@
(@{const_name True}, 0),
(@{const_name All}, 1),
(@{const_name Ex}, 1),
- (@{const_name "op ="}, 2),
+ (@{const_name "op ="}, 1),
(@{const_name "op &"}, 2),
(@{const_name "op |"}, 2),
(@{const_name "op -->"}, 2),
@@ -355,7 +403,6 @@
(@{const_name fst}, 1),
(@{const_name snd}, 1),
(@{const_name Id}, 0),
- (@{const_name insert}, 2),
(@{const_name converse}, 1),
(@{const_name trancl}, 1),
(@{const_name rel_comp}, 2),
@@ -396,10 +443,7 @@
((@{const_name ord_class.less_eq}, nat_T --> nat_T --> bool_T), 2),
((@{const_name of_nat}, nat_T --> int_T), 0)]
val built_in_set_consts =
- [(@{const_name semilattice_inf_class.inf}, 2),
- (@{const_name semilattice_sup_class.sup}, 2),
- (@{const_name minus_class.minus}, 2),
- (@{const_name ord_class.less_eq}, 2)]
+ [(@{const_name ord_class.less_eq}, 2)]
fun unarize_type @{typ "unsigned_bit word"} = nat_T
| unarize_type @{typ "signed_bit word"} = int_T
@@ -924,8 +968,8 @@
Const x' =>
if x = x' then @{const True}
else if is_constr_like ctxt x' then @{const False}
- else betapply (discr_term_for_constr hol_ctxt x, t)
- | _ => betapply (discr_term_for_constr hol_ctxt x, t)
+ else s_betapply [] (discr_term_for_constr hol_ctxt x, t)
+ | _ => s_betapply [] (discr_term_for_constr hol_ctxt x, t)
fun nth_arg_sel_term_for_constr thy stds (x as (s, T)) n =
let val (arg_Ts, dataT) = strip_type T in
@@ -955,7 +999,8 @@
else if is_constr_like ctxt x' then Const (@{const_name unknown}, res_T)
else raise SAME ()
| _ => raise SAME())
- handle SAME () => betapply (nth_arg_sel_term_for_constr thy stds x n, t)
+ handle SAME () =>
+ s_betapply [] (nth_arg_sel_term_for_constr thy stds x n, t)
end
fun construct_value _ _ x [] = Const x
@@ -1150,8 +1195,6 @@
fun special_bounds ts =
fold Term.add_vars ts [] |> sort (Term_Ord.fast_indexname_ord o pairself fst)
-fun abs_var ((s, j), T) body = Abs (s, T, abstract_over (Var ((s, j), T), body))
-
fun is_funky_typedef_name thy s =
member (op =) [@{type_name unit}, @{type_name "*"}, @{type_name "+"},
@{type_name int}] s orelse
@@ -1309,7 +1352,8 @@
original_name s <> s then
NONE
else
- x |> def_props_for_const thy [(NONE, false)] false table |> List.last
+ x |> def_props_for_const thy [(NONE, false)] false table
+ |> List.last
|> normalized_rhs_of |> Option.map (prefix_abs_vars s)
handle List.Empty => NONE
@@ -1458,25 +1502,44 @@
(** Constant unfolding **)
-fun constr_case_body ctxt stds (j, (x as (_, T))) =
+fun constr_case_body ctxt stds (func_t, (x as (_, T))) =
let val arg_Ts = binder_types T in
- list_comb (Bound j, map2 (select_nth_constr_arg ctxt stds x (Bound 0))
- (index_seq 0 (length arg_Ts)) arg_Ts)
+ s_betapplys [] (func_t, map2 (select_nth_constr_arg ctxt stds x (Bound 0))
+ (index_seq 0 (length arg_Ts)) arg_Ts)
end
-fun add_constr_case (hol_ctxt as {ctxt, stds, ...}) res_T (j, x) res_t =
- Const (@{const_name If}, bool_T --> res_T --> res_T --> res_T)
- $ discriminate_value hol_ctxt x (Bound 0) $ constr_case_body ctxt stds (j, x)
- $ res_t
-fun optimized_case_def (hol_ctxt as {ctxt, stds, ...}) dataT res_T =
+fun add_constr_case res_T (body_t, guard_t) res_t =
+ if res_T = bool_T then
+ s_conj (HOLogic.mk_imp (guard_t, body_t), res_t)
+ else
+ Const (@{const_name If}, bool_T --> res_T --> res_T --> res_T)
+ $ guard_t $ body_t $ res_t
+fun optimized_case_def (hol_ctxt as {ctxt, stds, ...}) dataT res_T func_ts =
let
val xs = datatype_constrs hol_ctxt dataT
- val func_Ts = map ((fn T => binder_types T ---> res_T) o snd) xs
- val (xs', x) = split_last xs
+ val cases =
+ func_ts ~~ xs
+ |> map (fn (func_t, x) =>
+ (constr_case_body ctxt stds (incr_boundvars 1 func_t, x),
+ discriminate_value hol_ctxt x (Bound 0)))
+ |> AList.group (op aconv)
+ |> map (apsnd (List.foldl s_disj @{const False}))
+ |> sort (int_ord o pairself (size_of_term o snd))
+ |> rev
in
- constr_case_body ctxt stds (1, x)
- |> fold_rev (add_constr_case hol_ctxt res_T) (length xs downto 2 ~~ xs')
- |> fold_rev (curry absdummy) (func_Ts @ [dataT])
+ if res_T = bool_T then
+ if forall (member (op =) [@{const False}, @{const True}] o fst) cases then
+ case cases of
+ [(body_t, _)] => body_t
+ | [_, (@{const True}, head_t2)] => head_t2
+ | [_, (@{const False}, head_t2)] => @{const Not} $ head_t2
+ | _ => raise BAD ("Nitpick_HOL.optimized_case_def", "impossible cases")
+ else
+ @{const True} |> fold_rev (add_constr_case res_T) cases
+ else
+ fst (hd cases) |> fold_rev (add_constr_case res_T) (tl cases)
end
+ |> curry absdummy dataT
+
fun optimized_record_get (hol_ctxt as {thy, ctxt, stds, ...}) s rec_T res_T t =
let val constr_x = hd (datatype_constrs hol_ctxt rec_T) in
case no_of_record_field thy s rec_T of
@@ -1504,7 +1567,7 @@
map2 (fn j => fn T =>
let val t = select_nth_constr_arg ctxt stds constr_x rec_t j T in
if j = special_j then
- betapply (fun_t, t)
+ s_betapply [] (fun_t, t)
else if j = n - 1 andalso special_j = ~1 then
optimized_record_update hol_ctxt s
(rec_T |> dest_Type |> snd |> List.last) fun_t t
@@ -1542,12 +1605,13 @@
handle TERM _ => raise SAME ()
else
raise SAME ())
- handle SAME () => betapply (do_term depth Ts t0, do_term depth Ts t1))
+ handle SAME () =>
+ s_betapply [] (do_term depth Ts t0, do_term depth Ts t1))
| Const (@{const_name refl_on}, T) $ Const (@{const_name top}, _) $ t2 =>
do_const depth Ts t (@{const_name refl'}, range_type T) [t2]
| (t0 as Const (@{const_name Sigma}, _)) $ t1 $ (t2 as Abs (_, _, t2')) =>
- betapplys (t0 |> loose_bvar1 (t2', 0) ? do_term depth Ts,
- map (do_term depth Ts) [t1, t2])
+ s_betapplys Ts (t0 |> loose_bvar1 (t2', 0) ? do_term depth Ts,
+ map (do_term depth Ts) [t1, t2])
| Const (x as (@{const_name distinct},
Type (@{type_name fun}, [Type (@{type_name list}, [T']), _])))
$ (t1 as _ $ _) =>
@@ -1560,11 +1624,11 @@
do_term depth Ts t2
else
do_const depth Ts t x [t1, t2, t3]
- | Const x $ t1 $ t2 $ t3 => do_const depth Ts t x [t1, t2, t3]
- | Const x $ t1 $ t2 => do_const depth Ts t x [t1, t2]
- | Const x $ t1 => do_const depth Ts t x [t1]
| Const x => do_const depth Ts t x []
- | t1 $ t2 => betapply (do_term depth Ts t1, do_term depth Ts t2)
+ | t1 $ t2 =>
+ (case strip_comb t of
+ (Const x, ts) => do_const depth Ts t x ts
+ | _ => s_betapply [] (do_term depth Ts t1, do_term depth Ts t2))
| Free _ => t
| Var _ => t
| Bound _ => t
@@ -1585,13 +1649,17 @@
(Const x, ts)
else case AList.lookup (op =) case_names s of
SOME n =>
- let
- val (dataT, res_T) = nth_range_type n T
- |> pairf domain_type range_type
- in
- (optimized_case_def hol_ctxt dataT res_T
- |> do_term (depth + 1) Ts, ts)
- end
+ if length ts < n then
+ (do_term depth Ts (eta_expand Ts t (n - length ts)), [])
+ else
+ let
+ val (dataT, res_T) = nth_range_type n T
+ |> pairf domain_type range_type
+ in
+ (optimized_case_def hol_ctxt dataT res_T
+ (map (do_term depth Ts) (take n ts)),
+ drop n ts)
+ end
| _ =>
if is_constr ctxt stds x then
(Const x, ts)
@@ -1645,11 +1713,14 @@
string_of_int depth ^ ") while expanding " ^
quote s)
else if s = @{const_name wfrec'} then
- (do_term (depth + 1) Ts (betapplys (def, ts)), [])
+ (do_term (depth + 1) Ts (s_betapplys Ts (def, ts)), [])
else
(do_term (depth + 1) Ts def, ts)
| NONE => (Const x, ts)
- in s_betapplys (const, map (do_term depth Ts) ts) |> Envir.beta_norm end
+ in
+ s_betapplys Ts (const, map (do_term depth Ts) ts)
+ |> s_beta_norm Ts
+ end
in do_term 0 [] end
(** Axiom extraction/generation **)
@@ -1796,8 +1867,9 @@
in
[HOLogic.eq_const bool_T $ (bisim_const $ n_var $ x_var $ y_var)
$ (@{term "op |"} $ (HOLogic.eq_const iter_T $ n_var $ zero_const iter_T)
- $ (betapplys (optimized_case_def hol_ctxt T bool_T,
- map case_func xs @ [x_var]))),
+ $ (s_betapply []
+ (optimized_case_def hol_ctxt T bool_T (map case_func xs),
+ x_var))),
HOLogic.eq_const set_T $ (bisim_const $ bisim_max $ x_var)
$ (Const (@{const_name insert}, T --> set_T --> set_T)
$ x_var $ Const (@{const_name bot_class.bot}, set_T))]
@@ -2036,11 +2108,12 @@
val outer_bounds = map Bound (length outer - 1 downto 0)
val cur = Var ((iter_var_prefix, j + 1), iter_T)
val next = suc_const iter_T $ cur
- val rhs = case fp_app of
- Const _ $ t =>
- betapply (t, list_comb (Const x', next :: outer_bounds))
- | _ => raise TERM ("Nitpick_HOL.unrolled_inductive_pred_\
- \const", [fp_app])
+ val rhs =
+ case fp_app of
+ Const _ $ t =>
+ s_betapply [] (t, list_comb (Const x', next :: outer_bounds))
+ | _ => raise TERM ("Nitpick_HOL.unrolled_inductive_pred_const",
+ [fp_app])
val (inner, naked_rhs) = strip_abs rhs
val all = outer @ inner
val bounds = map Bound (length all - 1 downto 0)
@@ -2056,10 +2129,10 @@
val def = the (def_of_const thy def_table x)
val (outer, fp_app) = strip_abs def
val outer_bounds = map Bound (length outer - 1 downto 0)
- val rhs = case fp_app of
- Const _ $ t => betapply (t, list_comb (Const x, outer_bounds))
- | _ => raise TERM ("Nitpick_HOL.raw_inductive_pred_axiom",
- [fp_app])
+ val rhs =
+ case fp_app of
+ Const _ $ t => s_betapply [] (t, list_comb (Const x, outer_bounds))
+ | _ => raise TERM ("Nitpick_HOL.raw_inductive_pred_axiom", [fp_app])
val (inner, naked_rhs) = strip_abs rhs
val all = outer @ inner
val bounds = map Bound (length all - 1 downto 0)
--- a/src/HOL/Tools/Nitpick/nitpick_kodkod.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_kodkod.ML Mon Jun 21 17:41:57 2010 +0200
@@ -866,24 +866,8 @@
| _ => kk_rel_eq r (KK.Atom (j0 + 1))
val formula_from_atom = formula_from_opt_atom Pos
- fun kk_notimplies f1 f2 = kk_and f1 (kk_not f2)
- val kk_or3 =
- double_rel_rel_let kk
- (fn r1 => fn r2 =>
- kk_rel_if (kk_subset true_atom (kk_union r1 r2)) true_atom
- (kk_intersect r1 r2))
- val kk_and3 =
- double_rel_rel_let kk
- (fn r1 => fn r2 =>
- kk_rel_if (kk_subset false_atom (kk_union r1 r2)) false_atom
- (kk_intersect r1 r2))
- fun kk_notimplies3 r1 r2 = kk_and3 r1 (kk_not3 r2)
-
val unpack_formulas =
map (formula_from_atom bool_j0) oo unpack_vect_in_chunks kk 1
- fun kk_vect_set_op connective k r1 r2 =
- fold1 kk_product (map2 (atom_from_formula kk bool_j0 oo connective)
- (unpack_formulas k r1) (unpack_formulas k r2))
fun kk_vect_set_bool_op connective k r1 r2 =
fold1 kk_and (map2 connective (unpack_formulas k r1)
(unpack_formulas k r2))
@@ -1369,14 +1353,6 @@
kk_union (kk_rel_if f1 true_atom KK.None)
(kk_rel_if f2 KK.None false_atom)
end
- | Op2 (Union, _, R, u1, u2) =>
- to_set_op kk_or kk_or3 kk_union kk_union kk_intersect false R u1 u2
- | Op2 (SetDifference, _, R, u1, u2) =>
- to_set_op kk_notimplies kk_notimplies3 kk_difference kk_intersect
- kk_union true R u1 u2
- | Op2 (Intersect, _, R, u1, u2) =>
- to_set_op kk_and kk_and3 kk_intersect kk_intersect kk_union false R
- u1 u2
| Op2 (Composition, _, R, u1, u2) =>
let
val (a_T, b_T) = HOLogic.dest_prodT (domain_type (type_of u1))
@@ -1644,37 +1620,6 @@
(kk_join r2 true_atom)
| _ => raise REP ("Nitpick_Kodkod.to_set_bool_op", [min_R])
end
- and to_set_op connective connective3 set_oper true_set_oper false_set_oper
- neg_second R u1 u2 =
- let
- val min_R = min_rep (rep_of u1) (rep_of u2)
- val r1 = to_rep min_R u1
- val r2 = to_rep min_R u2
- val unopt_R = unopt_rep R
- in
- rel_expr_from_rel_expr kk unopt_R (unopt_rep min_R)
- (case min_R of
- Opt (Vect (k, Atom _)) => kk_vect_set_op connective k r1 r2
- | Vect (k, Atom _) => kk_vect_set_op connective k r1 r2
- | Func (_, Formula Neut) => set_oper r1 r2
- | Func (Unit, _) => connective3 r1 r2
- | Func _ =>
- double_rel_rel_let kk
- (fn r1 => fn r2 =>
- kk_union
- (kk_product
- (true_set_oper (kk_join r1 true_atom)
- (kk_join r2 (atom_for_bool bool_j0
- (not neg_second))))
- true_atom)
- (kk_product
- (false_set_oper (kk_join r1 false_atom)
- (kk_join r2 (atom_for_bool bool_j0
- neg_second)))
- false_atom))
- r1 r2
- | _ => raise REP ("Nitpick_Kodkod.to_set_op", [min_R]))
- end
and to_bit_word_unary_op T R oper =
let
val Ts = strip_type T ||> single |> op @
--- a/src/HOL/Tools/Nitpick/nitpick_mono.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_mono.ML Mon Jun 21 17:41:57 2010 +0200
@@ -54,8 +54,9 @@
exception MTYPE of string * mtyp list * typ list
exception MTERM of string * mterm list
-fun print_g (_ : string) = ()
-(* val print_g = tracing *)
+val debug_mono = false
+
+fun print_g f = () |> debug_mono ? tracing o f
val string_for_var = signed_string_of_int
fun string_for_vars sep [] = "0\<^bsub>" ^ sep ^ "\<^esub>"
@@ -401,10 +402,10 @@
[M1, M2], [])
fun add_mtype_comp cmp M1 M2 ((lits, comps, sexps) : constraint_set) =
- (print_g ("*** Add " ^ string_for_mtype M1 ^ " " ^ string_for_comp_op cmp ^
- " " ^ string_for_mtype M2);
+ (print_g (fn () => "*** Add " ^ string_for_mtype M1 ^ " " ^
+ string_for_comp_op cmp ^ " " ^ string_for_mtype M2);
case do_mtype_comp cmp [] M1 M2 (SOME (lits, comps)) of
- NONE => (print_g "**** Unsolvable"; raise UNSOLVABLE ())
+ NONE => (print_g (K "**** Unsolvable"); raise UNSOLVABLE ())
| SOME (lits, comps) => (lits, comps, sexps))
val add_mtypes_equal = add_mtype_comp Eq
@@ -446,10 +447,11 @@
raise MTYPE ("Nitpick_Mono.do_notin_mtype_fv", [M], [])
fun add_notin_mtype_fv sn M ((lits, comps, sexps) : constraint_set) =
- (print_g ("*** Add " ^ string_for_mtype M ^ " is " ^
- (case sn of Minus => "concrete" | Plus => "complete") ^ ".");
+ (print_g (fn () => "*** Add " ^ string_for_mtype M ^ " is " ^
+ (case sn of Minus => "concrete" | Plus => "complete") ^
+ ".");
case do_notin_mtype_fv sn [] M (SOME (lits, sexps)) of
- NONE => (print_g "**** Unsolvable"; raise UNSOLVABLE ())
+ NONE => (print_g (K "**** Unsolvable"); raise UNSOLVABLE ())
| SOME (lits, sexps) => (lits, comps, sexps))
val add_mtype_is_concrete = add_notin_mtype_fv Minus
@@ -491,15 +493,16 @@
subscript_string_for_vars " \<and> " xs ^ " " ^ string_for_sign_atom a2
fun print_problem lits comps sexps =
- print_g ("*** Problem:\n" ^ cat_lines (map string_for_literal lits @
- map string_for_comp comps @
- map string_for_sign_expr sexps))
+ print_g (fn () => "*** Problem:\n" ^
+ cat_lines (map string_for_literal lits @
+ map string_for_comp comps @
+ map string_for_sign_expr sexps))
fun print_solution lits =
let val (pos, neg) = List.partition (curry (op =) Plus o snd) lits in
- print_g ("*** Solution:\n" ^
- "+: " ^ commas (map (string_for_var o fst) pos) ^ "\n" ^
- "-: " ^ commas (map (string_for_var o fst) neg))
+ print_g (fn () => "*** Solution:\n" ^
+ "+: " ^ commas (map (string_for_var o fst) pos) ^ "\n" ^
+ "-: " ^ commas (map (string_for_var o fst) neg))
end
fun solve max_var (lits, comps, sexps) =
@@ -550,6 +553,12 @@
def_table, ...},
alpha_T, max_fresh, ...}) =
let
+ fun is_enough_eta_expanded t =
+ case strip_comb t of
+ (Const x, ts) =>
+ the_default 0 (arity_of_built_in_const thy stds fast_descrs x)
+ <= length ts
+ | _ => true
val mtype_for = fresh_mtype_for_type mdata false
fun plus_set_mtype_for_dom M =
MFun (M, S (if exists_alpha_sub_mtype M then Plus else Minus), bool_M)
@@ -640,8 +649,10 @@
|>> mtype_of_mterm
end
| @{const_name "op ="} => do_equals T accum
- | @{const_name The} => (print_g "*** The"; raise UNSOLVABLE ())
- | @{const_name Eps} => (print_g "*** Eps"; raise UNSOLVABLE ())
+ | @{const_name The} =>
+ (print_g (K "*** The"); raise UNSOLVABLE ())
+ | @{const_name Eps} =>
+ (print_g (K "*** Eps"); raise UNSOLVABLE ())
| @{const_name If} =>
do_robust_set_operation (range_type T) accum
|>> curry3 MFun bool_M (S Minus)
@@ -650,19 +661,6 @@
| @{const_name snd} => do_nth_pair_sel 1 T accum
| @{const_name Id} =>
(MFun (mtype_for (domain_type T), S Minus, bool_M), accum)
- | @{const_name insert} =>
- let
- val set_T = domain_type (range_type T)
- val M1 = mtype_for (domain_type set_T)
- val M1' = plus_set_mtype_for_dom M1
- val M2 = mtype_for set_T
- val M3 = mtype_for set_T
- in
- (MFun (M1, S Minus, MFun (M2, S Minus, M3)),
- (gamma, cset |> add_mtype_is_concrete M1
- |> add_is_sub_mtype M1' M3
- |> add_is_sub_mtype M2 M3))
- end
| @{const_name converse} =>
let
val x = Unsynchronized.inc max_fresh
@@ -720,25 +718,9 @@
val a_set_M = mtype_for (domain_type T)
val a_M = dest_MFun a_set_M |> #1
in (MFun (a_set_M, S Minus, a_M), accum) end
- else if s = @{const_name minus_class.minus} andalso
- is_set_type (domain_type T) then
- let
- val set_T = domain_type T
- val left_set_M = mtype_for set_T
- val right_set_M = mtype_for set_T
- in
- (MFun (left_set_M, S Minus,
- MFun (right_set_M, S Minus, left_set_M)),
- (gamma, cset |> add_mtype_is_concrete right_set_M
- |> add_is_sub_mtype right_set_M left_set_M))
- end
else if s = @{const_name ord_class.less_eq} andalso
is_set_type (domain_type T) then
do_fragile_set_operation T accum
- else if (s = @{const_name semilattice_inf_class.inf} orelse
- s = @{const_name semilattice_sup_class.sup}) andalso
- is_set_type (domain_type T) then
- do_robust_set_operation T accum
else if is_sel s then
(mtype_for_sel mdata x, accum)
else if is_constr ctxt stds x then
@@ -758,7 +740,7 @@
(M, ({bound_Ts = bound_Ts, bound_Ms = bound_Ms,
frees = (x, M) :: frees, consts = consts}, cset))
end) |>> curry MRaw t
- | Var _ => (print_g "*** Var"; raise UNSOLVABLE ())
+ | Var _ => (print_g (K "*** Var"); raise UNSOLVABLE ())
| Bound j => (MRaw (t, nth bound_Ms j), accum)
| Abs (s, T, t') =>
(case fin_fun_body T (fastype_of1 (T :: bound_Ts, t')) t' of
@@ -771,10 +753,16 @@
| NONE =>
((case t' of
t1' $ Bound 0 =>
- if not (loose_bvar1 (t1', 0)) then
+ if not (loose_bvar1 (t1', 0)) andalso
+ is_enough_eta_expanded t1' then
do_term (incr_boundvars ~1 t1') accum
else
raise SAME ()
+ | (t11 as Const (@{const_name "op ="}, _)) $ Bound 0 $ t13 =>
+ if not (loose_bvar1 (t13, 0)) then
+ do_term (incr_boundvars ~1 (t11 $ t13)) accum
+ else
+ raise SAME ()
| _ => raise SAME ())
handle SAME () =>
let
@@ -803,8 +791,8 @@
val M2 = mtype_of_mterm m2
in (MApp (m1, m2), accum ||> add_is_sub_mtype M2 M11) end
end)
- |> tap (fn (m, _) => print_g (" \<Gamma> \<turnstile> " ^
- string_for_mterm ctxt m))
+ |> tap (fn (m, _) => print_g (fn () => " \<Gamma> \<turnstile> " ^
+ string_for_mterm ctxt m))
in do_term end
fun force_minus_funs 0 _ = I
@@ -902,9 +890,9 @@
| _ => do_term t accum
end
|> tap (fn (m, _) =>
- print_g ("\<Gamma> \<turnstile> " ^
- string_for_mterm ctxt m ^ " : o\<^sup>" ^
- string_for_sign sn))
+ print_g (fn () => "\<Gamma> \<turnstile> " ^
+ string_for_mterm ctxt m ^ " : o\<^sup>" ^
+ string_for_sign sn))
in do_formula end
(* The harmless axiom optimization below is somewhat too aggressive in the face
@@ -987,9 +975,10 @@
Syntax.string_of_term ctxt t ^ " : " ^ string_for_mtype (resolve_mtype lits M)
fun print_mtype_context ctxt lits ({frees, consts, ...} : mtype_context) =
- map (fn (x, M) => string_for_mtype_of_term ctxt lits (Free x) M) frees @
- map (fn (x, M) => string_for_mtype_of_term ctxt lits (Const x) M) consts
- |> cat_lines |> print_g
+ print_g (fn () =>
+ map (fn (x, M) => string_for_mtype_of_term ctxt lits (Free x) M) frees @
+ map (fn (x, M) => string_for_mtype_of_term ctxt lits (Const x) M) consts
+ |> cat_lines)
fun amass f t (ms, accum) =
let val (m, accum) = f t accum in (m :: ms, accum) end
@@ -997,9 +986,9 @@
fun infer which no_harmless (hol_ctxt as {ctxt, ...}) binarize alpha_T
(nondef_ts, def_ts) =
let
- val _ = print_g ("****** " ^ which ^ " analysis: " ^
- string_for_mtype MAlpha ^ " is " ^
- Syntax.string_of_typ ctxt alpha_T)
+ val _ = print_g (fn () => "****** " ^ which ^ " analysis: " ^
+ string_for_mtype MAlpha ^ " is " ^
+ Syntax.string_of_typ ctxt alpha_T)
val mdata as {max_fresh, constr_mcache, ...} =
initial_mdata hol_ctxt binarize no_harmless alpha_T
val accum = (initial_gamma, ([], [], []))
@@ -1064,26 +1053,21 @@
in
case t of
Const (x as (s, _)) =>
- if s = @{const_name insert} then
- case nth_range_type 2 T' of
- set_T' as Type (@{type_name fin_fun}, [elem_T', _]) =>
- Abs (Name.uu, elem_T', Abs (Name.uu, set_T',
- Const (@{const_name If},
- bool_T --> set_T' --> set_T' --> set_T')
- $ (Const (@{const_name is_unknown},
- elem_T' --> bool_T) $ Bound 1)
- $ (Const (@{const_name unknown}, set_T'))
- $ (coerce_term hol_ctxt new_Ts T' T (Const x)
- $ Bound 1 $ Bound 0)))
- | _ => Const (s, T')
- else if s = @{const_name finite} then
+ if s = @{const_name finite} then
case domain_type T' of
set_T' as Type (@{type_name fin_fun}, _) =>
Abs (Name.uu, set_T', @{const True})
| _ => Const (s, T')
else if s = @{const_name "=="} orelse
s = @{const_name "op ="} then
- Const (s, T')
+ let
+ val T =
+ case T' of
+ Type (_, [T1, Type (_, [T2, T3])]) =>
+ T1 --> T2 --> T3
+ | _ => raise TYPE ("Nitpick_Mono.finitize_funs.\
+ \term_from_mterm", [T, T'], [])
+ in coerce_term hol_ctxt new_Ts T' T (Const (s, T)) end
else if is_built_in_const thy stds fast_descrs x then
coerce_term hol_ctxt new_Ts T' T t
else if is_constr ctxt stds x then
--- a/src/HOL/Tools/Nitpick/nitpick_nut.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_nut.ML Mon Jun 21 17:41:57 2010 +0200
@@ -56,9 +56,6 @@
The |
Eps |
Triad |
- Union |
- SetDifference |
- Intersect |
Composition |
Product |
Image |
@@ -176,9 +173,6 @@
The |
Eps |
Triad |
- Union |
- SetDifference |
- Intersect |
Composition |
Product |
Image |
@@ -247,9 +241,6 @@
| string_for_op2 The = "The"
| string_for_op2 Eps = "Eps"
| string_for_op2 Triad = "Triad"
- | string_for_op2 Union = "Union"
- | string_for_op2 SetDifference = "SetDifference"
- | string_for_op2 Intersect = "Intersect"
| string_for_op2 Composition = "Composition"
| string_for_op2 Product = "Product"
| string_for_op2 Image = "Image"
@@ -525,6 +516,8 @@
do_description_operator The @{const_name undefined_fast_The} x t1
| (Const (x as (@{const_name Eps}, _)), [t1]) =>
do_description_operator Eps @{const_name undefined_fast_Eps} x t1
+ | (Const (@{const_name "op ="}, T), [t1]) =>
+ 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]) =>
Op2 (And, bool_T, Any, sub' t1, sub' t2)
@@ -547,13 +540,6 @@
| (Const (@{const_name snd}, T), [t1]) =>
Op1 (Second, range_type T, Any, sub t1)
| (Const (@{const_name Id}, T), []) => Cst (Iden, T, Any)
- | (Const (@{const_name insert}, T), [t1, t2]) =>
- (case t2 of
- Abs (_, _, @{const False}) =>
- Op1 (SingletonSet, nth_range_type 2 T, Any, sub t1)
- | _ =>
- Op2 (Union, nth_range_type 2 T, Any,
- Op1 (SingletonSet, nth_range_type 2 T, Any, sub t1), sub t2))
| (Const (@{const_name converse}, T), [t1]) =>
Op1 (Converse, range_type T, Any, sub t1)
| (Const (@{const_name trancl}, T), [t1]) =>
@@ -585,11 +571,6 @@
| (Const (x as (s as @{const_name plus_class.plus}, T)), []) =>
if is_built_in_const thy stds false x then Cst (Add, T, Any)
else ConstName (s, T, Any)
- | (Const (@{const_name minus_class.minus},
- Type (@{type_name fun},
- [T1 as Type (@{type_name fun}, [_, @{typ bool}]), _])),
- [t1, t2]) =>
- Op2 (SetDifference, T1, Any, sub t1, sub t2)
| (Const (x as (s as @{const_name minus_class.minus}, T)), []) =>
if is_built_in_const thy stds false x then Cst (Subtract, T, Any)
else ConstName (s, T, Any)
@@ -643,16 +624,6 @@
| (Const (@{const_name of_nat},
T as @{typ "unsigned_bit word => signed_bit word"}), []) =>
Cst (NatToInt, T, Any)
- | (Const (@{const_name semilattice_inf_class.inf},
- Type (@{type_name fun},
- [T1 as Type (@{type_name fun}, [_, @{typ bool}]), _])),
- [t1, t2]) =>
- Op2 (Intersect, T1, Any, sub t1, sub t2)
- | (Const (@{const_name semilattice_sup_class.sup},
- Type (@{type_name fun},
- [T1 as Type (@{type_name fun}, [_, @{typ bool}]), _])),
- [t1, t2]) =>
- Op2 (Union, T1, Any, sub t1, sub t2)
| (t0 as Const (x as (s, T)), ts) =>
if is_constr ctxt stds x then
do_construct x ts
@@ -685,15 +656,14 @@
| Op1 (SingletonSet, _, _, _) => true
| Op1 (Converse, _, _, u1) => is_fully_representable_set u1
| Op2 (oper, _, _, u1, u2) =>
- if oper = Union orelse oper = SetDifference orelse oper = Intersect then
- forall is_fully_representable_set [u1, u2]
- else if oper = Apply then
+ if oper = Apply then
case u1 of
ConstName (s, _, _) =>
is_sel_like_and_no_discr s orelse s = @{const_name set}
| _ => false
else
false
+ | Op2 (Lambda, _, _, _, Cst (False, _, _)) => true
| _ => false
fun rep_for_abs_fun scope T =
@@ -904,7 +874,7 @@
end
| Cst (cst, T, _) =>
if cst = Unknown orelse cst = Unrep then
- case (is_boolean_type T, polar) of
+ case (is_boolean_type T, polar |> unsound ? flip_polarity) of
(true, Pos) => Cst (False, T, Formula Pos)
| (true, Neg) => Cst (True, T, Formula Neg)
| _ => Cst (cst, T, best_opt_set_rep_for_type scope T)
--- a/src/HOL/Tools/Nitpick/nitpick_preproc.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_preproc.ML Mon Jun 21 17:41:57 2010 +0200
@@ -91,7 +91,7 @@
fun uncurry_term table t =
let
fun aux (t1 $ t2) args = aux t1 (aux t2 [] :: args)
- | aux (Abs (s, T, t')) args = betapplys (Abs (s, T, aux t' []), args)
+ | aux (Abs (s, T, t')) args = s_betapplys [] (Abs (s, T, aux t' []), args)
| aux (t as Const (s, T)) args =
(case Termtab.lookup table t of
SOME n =>
@@ -111,17 +111,18 @@
val tuple_T = HOLogic.mk_tupleT tuple_arg_Ts
in
if n - j < 2 then
- betapplys (t, args)
+ s_betapplys [] (t, args)
else
- betapplys (Const (uncurry_prefix_for (n - j) j ^ s,
- before_arg_Ts ---> tuple_T --> rest_T),
- before_args @ [mk_flat_tuple tuple_T tuple_args] @
- after_args)
+ s_betapplys []
+ (Const (uncurry_prefix_for (n - j) j ^ s,
+ before_arg_Ts ---> tuple_T --> rest_T),
+ before_args @ [mk_flat_tuple tuple_T tuple_args] @
+ after_args)
end
else
- betapplys (t, args)
- | NONE => betapplys (t, args))
- | aux t args = betapplys (t, args)
+ s_betapplys [] (t, args)
+ | NONE => s_betapplys [] (t, args))
+ | aux t args = s_betapplys [] (t, args)
in aux t [] end
(** Boxing **)
@@ -248,13 +249,13 @@
val T2 = fastype_of1 (new_Ts, t2)
val t2 = coerce_term hol_ctxt new_Ts (hd Ts1) T2 t2
in
- betapply (if s1 = @{type_name fun} then
- t1
- else
- select_nth_constr_arg ctxt stds
- (@{const_name FunBox},
- Type (@{type_name fun}, Ts1) --> T1) t1 0
- (Type (@{type_name fun}, Ts1)), t2)
+ s_betapply new_Ts (if s1 = @{type_name fun} then
+ t1
+ else
+ select_nth_constr_arg ctxt stds
+ (@{const_name FunBox},
+ Type (@{type_name fun}, Ts1) --> T1) t1 0
+ (Type (@{type_name fun}, Ts1)), t2)
end
| t1 $ t2 =>
let
@@ -265,13 +266,13 @@
val T2 = fastype_of1 (new_Ts, t2)
val t2 = coerce_term hol_ctxt new_Ts (hd Ts1) T2 t2
in
- betapply (if s1 = @{type_name fun} then
- t1
- else
- select_nth_constr_arg ctxt stds
- (@{const_name FunBox},
- Type (@{type_name fun}, Ts1) --> T1) t1 0
- (Type (@{type_name fun}, Ts1)), t2)
+ s_betapply new_Ts (if s1 = @{type_name fun} then
+ t1
+ else
+ select_nth_constr_arg ctxt stds
+ (@{const_name FunBox},
+ Type (@{type_name fun}, Ts1) --> T1) t1 0
+ (Type (@{type_name fun}, Ts1)), t2)
end
| Free (s, T) => Free (s, box_type hol_ctxt InExpr T)
| Var (z as (x, T)) =>
@@ -388,18 +389,6 @@
(list_comb (t, args), seen)
in aux [] 0 t [] [] |> fst end
-val let_var_prefix = nitpick_prefix ^ "l"
-val let_inline_threshold = 32
-
-fun hol_let n abs_T body_T f t =
- if n * size_of_term t <= let_inline_threshold then
- f t
- else
- let val z = ((let_var_prefix, 0), abs_T) in
- Const (@{const_name Let}, abs_T --> (abs_T --> body_T) --> body_T)
- $ t $ abs_var z (incr_boundvars 1 (f (Var z)))
- end
-
fun destroy_pulled_out_constrs (hol_ctxt as {ctxt, stds, ...}) axiom t =
let
val num_occs_of_var =
@@ -439,10 +428,10 @@
x = (@{const_name Suc}, nat_T --> nat_T)) andalso
(not careful orelse not (is_Var t1) orelse
String.isPrefix val_var_prefix (fst (fst (dest_Var t1)))) then
- hol_let (n + 1) dataT bool_T
- (fn t1 => discriminate_value hol_ctxt x t1 ::
- map3 (sel_eq x t1) (index_seq 0 n) arg_Ts args
- |> foldr1 s_conj) t1
+ s_let "l" (n + 1) dataT bool_T
+ (fn t1 => discriminate_value hol_ctxt x t1 ::
+ map3 (sel_eq x t1) (index_seq 0 n) arg_Ts args
+ |> foldr1 s_conj) t1
else
raise SAME ()
end
@@ -572,7 +561,7 @@
map Bound (rev js))
val abs_t = Abs (abs_s, abs_T, aux ss Ts (incrs js) depth polar t)
in
- if null js then betapply (abs_t, sko_t)
+ if null js then s_betapply Ts (abs_t, sko_t)
else Const (@{const_name Let}, abs_T --> quant_T) $ sko_t $ abs_t
end
else
@@ -602,11 +591,9 @@
| Const (s0 as @{const_name Ex}, T0) $ Abs (s1, T1, t1) =>
do_quantifier s0 T0 s1 T1 t1
| @{const "op &"} $ t1 $ t2 =>
- @{const "op &"} $ aux ss Ts js depth polar t1
- $ aux ss Ts js depth polar t2
+ s_conj (pairself (aux ss Ts js depth polar) (t1, t2))
| @{const "op |"} $ t1 $ t2 =>
- @{const "op |"} $ aux ss Ts js depth polar t1
- $ aux ss Ts js depth polar t2
+ s_disj (pairself (aux ss Ts js depth polar) (t1, t2))
| @{const "op -->"} $ t1 $ t2 =>
@{const "op -->"} $ aux ss Ts js depth (flip_polarity polar) t1
$ aux ss Ts js depth polar t2
@@ -617,42 +604,30 @@
not (is_well_founded_inductive_pred hol_ctxt x) then
let
val gfp = (fixpoint_kind_of_const thy def_table x = Gfp)
- val (pref, connective, set_oper) =
- if gfp then
- (lbfp_prefix, @{const "op |"},
- @{const_name semilattice_sup_class.sup})
- else
- (ubfp_prefix, @{const "op &"},
- @{const_name semilattice_inf_class.inf})
+ val (pref, connective) =
+ if gfp then (lbfp_prefix, @{const "op |"})
+ else (ubfp_prefix, @{const "op &"})
fun pos () = unrolled_inductive_pred_const hol_ctxt gfp x
|> aux ss Ts js depth polar
fun neg () = Const (pref ^ s, T)
in
- (case polar |> gfp ? flip_polarity of
- Pos => pos ()
- | Neg => neg ()
- | Neut =>
- if is_fun_type T then
- let
- val ((trunk_arg_Ts, rump_arg_T), body_T) =
- T |> strip_type |>> split_last
- val set_T = rump_arg_T --> body_T
- fun app f =
- list_comb (f (),
- map Bound (length trunk_arg_Ts - 1 downto 0))
- in
- List.foldr absdummy
- (Const (set_oper, set_T --> set_T --> set_T)
- $ app pos $ app neg) trunk_arg_Ts
- end
- else
- connective $ pos () $ neg ())
+ case polar |> gfp ? flip_polarity of
+ Pos => pos ()
+ | Neg => neg ()
+ | Neut =>
+ let
+ val arg_Ts = binder_types T
+ fun app f =
+ list_comb (f (), map Bound (length arg_Ts - 1 downto 0))
+ in
+ List.foldr absdummy (connective $ app pos $ app neg) arg_Ts
+ end
end
else
Const x
| t1 $ t2 =>
- betapply (aux ss Ts [] (skolem_depth + 1) polar t1,
- aux ss Ts [] depth Neut t2)
+ s_betapply Ts (aux ss Ts [] (skolem_depth + 1) polar t1,
+ aux ss Ts [] depth Neut t2)
| Abs (s, T, t1) => Abs (s, T, aux ss Ts (incrs js) depth polar t1)
| _ => t
end
@@ -1064,10 +1039,10 @@
| _ => raise SAME ()
else
raise SAME ())
- handle SAME () => betapplys (t, args))
+ handle SAME () => s_betapplys [] (t, args))
| do_term (Abs (s, T, t')) args =
- betapplys (Abs (s, T, do_term t' []), args)
- | do_term t args = betapplys (t, args)
+ s_betapplys [] (Abs (s, T, do_term t' []), args)
+ | do_term t args = s_betapplys [] (t, args)
in do_term t [] end
(** Quantifier massaging: Distributing quantifiers **)
@@ -1223,15 +1198,20 @@
fun finitize_all_types_of_funs (hol_ctxt as {thy, ...}) binarize finitizes monos
(nondef_ts, def_ts) =
- let
- val Ts = ground_types_in_terms hol_ctxt binarize (nondef_ts @ def_ts)
- |> filter_out (fn Type (@{type_name fun_box}, _) => true
- | @{typ signed_bit} => true
- | @{typ unsigned_bit} => true
- | T => is_small_finite_type hol_ctxt T orelse
- triple_lookup (type_match thy) monos T
- = SOME (SOME false))
- in fold (finitize_funs hol_ctxt binarize finitizes) Ts (nondef_ts, def_ts) end
+ if forall (curry (op =) (SOME false) o snd) finitizes then
+ (nondef_ts, def_ts)
+ else
+ let
+ val Ts = ground_types_in_terms hol_ctxt binarize (nondef_ts @ def_ts)
+ |> filter_out (fn Type (@{type_name fun_box}, _) => true
+ | @{typ signed_bit} => true
+ | @{typ unsigned_bit} => true
+ | T => is_small_finite_type hol_ctxt T orelse
+ triple_lookup (type_match thy) monos T
+ = SOME (SOME false))
+ in
+ fold (finitize_funs hol_ctxt binarize finitizes) Ts (nondef_ts, def_ts)
+ end
(** Preprocessor entry point **)
--- a/src/HOL/Tools/Nitpick/nitpick_tests.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_tests.ML Mon Jun 21 17:41:57 2010 +0200
@@ -298,7 +298,6 @@
val peephole_optim = true
val nat_card = 4
val int_card = 9
- val bits = 8
val j0 = 0
val constrs = kodkod_constrs peephole_optim nat_card int_card j0
val (free_rels, pool, table) =
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Mon Jun 21 17:41:57 2010 +0200
@@ -265,12 +265,12 @@
fun replace_ho_args mode hoargs ts =
let
fun replace (Fun _, _) (arg' :: hoargs') = (arg', hoargs')
- | replace (Pair (m1, m2), Const ("Pair", T) $ t1 $ t2) hoargs =
+ | replace (Pair (m1, m2), Const (@{const_name Pair}, T) $ t1 $ t2) hoargs =
let
val (t1', hoargs') = replace (m1, t1) hoargs
val (t2', hoargs'') = replace (m2, t2) hoargs'
in
- (Const ("Pair", T) $ t1' $ t2', hoargs'')
+ (Const (@{const_name Pair}, T) $ t1' $ t2', hoargs'')
end
| replace (_, t) hoargs = (t, hoargs)
in
@@ -290,7 +290,7 @@
fun split_map_mode f mode ts =
let
fun split_arg_mode' (m as Fun _) t = f m t
- | split_arg_mode' (Pair (m1, m2)) (Const ("Pair", _) $ t1 $ t2) =
+ | split_arg_mode' (Pair (m1, m2)) (Const (@{const_name Pair}, _) $ t1 $ t2) =
let
val (i1, o1) = split_arg_mode' m1 t1
val (i2, o2) = split_arg_mode' m2 t2
@@ -334,7 +334,7 @@
end
| fold_map_aterms_prodT comb f T s = f T s
-fun map_filter_prod f (Const ("Pair", _) $ t1 $ t2) =
+fun map_filter_prod f (Const (@{const_name Pair}, _) $ t1 $ t2) =
comb_option HOLogic.mk_prod (map_filter_prod f t1, map_filter_prod f t2)
| map_filter_prod f t = f t
@@ -561,8 +561,8 @@
(* combinators to apply a function to all basic parts of nested products *)
-fun map_products f (Const ("Pair", T) $ t1 $ t2) =
- Const ("Pair", T) $ map_products f t1 $ map_products f t2
+fun map_products f (Const (@{const_name Pair}, T) $ t1 $ t2) =
+ Const (@{const_name Pair}, T) $ map_products f t1 $ map_products f t2
| map_products f t = f t
(* split theorems of case expressions *)
@@ -756,7 +756,7 @@
(case HOLogic.strip_tupleT (fastype_of arg) of
(Ts as _ :: _ :: _) =>
let
- fun rewrite_arg' (Const (@{const_name "Pair"}, _) $ _ $ t2, Type (@{type_name "*"}, [_, T2]))
+ fun rewrite_arg' (Const (@{const_name Pair}, _) $ _ $ t2, Type (@{type_name "*"}, [_, T2]))
(args, (pats, intro_t, ctxt)) = rewrite_arg' (t2, T2) (args, (pats, intro_t, ctxt))
| rewrite_arg' (t, Type (@{type_name "*"}, [T1, T2])) (args, (pats, intro_t, ctxt)) =
let
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Mon Jun 21 17:41:57 2010 +0200
@@ -117,7 +117,7 @@
end;
fun dest_randomT (Type ("fun", [@{typ Random.seed},
- Type ("*", [Type ("*", [T, @{typ "unit => Code_Evaluation.term"}]) ,@{typ Random.seed}])])) = T
+ Type (@{type_name "*"}, [Type (@{type_name "*"}, [T, @{typ "unit => Code_Evaluation.term"}]) ,@{typ Random.seed}])])) = T
| dest_randomT T = raise TYPE ("dest_randomT", [T], [])
fun mk_tracing s t =
@@ -467,7 +467,7 @@
(* generation of case rules from user-given introduction rules *)
-fun mk_args2 (Type ("*", [T1, T2])) st =
+fun mk_args2 (Type (@{type_name "*"}, [T1, T2])) st =
let
val (t1, st') = mk_args2 T1 st
val (t2, st'') = mk_args2 T2 st'
@@ -1099,7 +1099,7 @@
fun all_input_of T =
let
val (Ts, U) = strip_type T
- fun input_of (Type ("*", [T1, T2])) = Pair (input_of T1, input_of T2)
+ fun input_of (Type (@{type_name "*"}, [T1, T2])) = Pair (input_of T1, input_of T2)
| input_of _ = Input
in
if U = HOLogic.boolT then
@@ -1190,7 +1190,7 @@
fun missing_vars vs t = subtract (op =) vs (term_vs t)
-fun output_terms (Const ("Pair", _) $ t1 $ t2, Mode_Pair (d1, d2)) =
+fun output_terms (Const (@{const_name Pair}, _) $ t1 $ t2, Mode_Pair (d1, d2)) =
output_terms (t1, d1) @ output_terms (t2, d2)
| output_terms (t1 $ t2, Mode_App (d1, d2)) =
output_terms (t1, d1) @ output_terms (t2, d2)
@@ -1206,7 +1206,7 @@
SOME ms => SOME (map (fn m => (Context m , [])) ms)
| NONE => NONE)
-fun derivations_of (ctxt : Proof.context) modes vs (Const ("Pair", _) $ t1 $ t2) (Pair (m1, m2)) =
+fun derivations_of (ctxt : Proof.context) modes vs (Const (@{const_name Pair}, _) $ t1 $ t2) (Pair (m1, m2)) =
map_product
(fn (m1, mvars1) => fn (m2, mvars2) => (Mode_Pair (m1, m2), union (op =) mvars1 mvars2))
(derivations_of ctxt modes vs t1 m1) (derivations_of ctxt modes vs t2 m2)
@@ -1236,7 +1236,7 @@
else if eq_mode (m, Output) then
(if is_possible_output ctxt vs t then [(Term Output, [])] else [])
else []
-and all_derivations_of ctxt modes vs (Const ("Pair", _) $ t1 $ t2) =
+and all_derivations_of ctxt modes vs (Const (@{const_name Pair}, _) $ t1 $ t2) =
let
val derivs1 = all_derivations_of ctxt modes vs t1
val derivs2 = all_derivations_of ctxt modes vs t2
@@ -1665,7 +1665,7 @@
(case (t, deriv) of
(t1 $ t2, Mode_App (deriv1, deriv2)) =>
string_of_tderiv ctxt (t1, deriv1) ^ " $ " ^ string_of_tderiv ctxt (t2, deriv2)
- | (Const ("Pair", _) $ t1 $ t2, Mode_Pair (deriv1, deriv2)) =>
+ | (Const (@{const_name Pair}, _) $ t1 $ t2, Mode_Pair (deriv1, deriv2)) =>
"(" ^ string_of_tderiv ctxt (t1, deriv1) ^ ", " ^ string_of_tderiv ctxt (t2, deriv2) ^ ")"
| (t, Term Input) => Syntax.string_of_term ctxt t ^ "[Input]"
| (t, Term Output) => Syntax.string_of_term ctxt t ^ "[Output]"
@@ -1692,7 +1692,7 @@
val bs = map (pair "x") (binder_types (fastype_of t))
val bounds = map Bound (rev (0 upto (length bs) - 1))
in SOME (list_abs (bs, mk_if compfuns (list_comb (t, bounds)))) end
- | (Const ("Pair", _) $ t1 $ t2, Mode_Pair (d1, d2)) =>
+ | (Const (@{const_name Pair}, _) $ t1 $ t2, Mode_Pair (d1, d2)) =>
(case (expr_of (t1, d1), expr_of (t2, d2)) of
(NONE, NONE) => NONE
| (NONE, SOME t) => SOME t
@@ -2010,7 +2010,7 @@
(ks @ [SOME k]))) arities));
fun split_lambda (x as Free _) t = lambda x t
- | split_lambda (Const ("Pair", _) $ t1 $ t2) t =
+ | split_lambda (Const (@{const_name Pair}, _) $ t1 $ t2) t =
HOLogic.mk_split (split_lambda t1 (split_lambda t2 t))
| split_lambda (Const ("Product_Type.Unity", _)) t = Abs ("x", HOLogic.unitT, t)
| split_lambda t _ = raise (TERM ("split_lambda", [t]))
@@ -2019,7 +2019,7 @@
| strip_split_abs (Abs (_, _, t)) = strip_split_abs t
| strip_split_abs t = t
-fun mk_args is_eval (m as Pair (m1, m2), T as Type ("*", [T1, T2])) names =
+fun mk_args is_eval (m as Pair (m1, m2), T as Type (@{type_name "*"}, [T1, T2])) names =
if eq_mode (m, Input) orelse eq_mode (m, Output) then
let
val x = Name.variant names "x"
@@ -3112,7 +3112,7 @@
in
if defined_functions compilation ctxt name then
let
- fun extract_mode (Const ("Pair", _) $ t1 $ t2) = Pair (extract_mode t1, extract_mode t2)
+ fun extract_mode (Const (@{const_name Pair}, _) $ t1 $ t2) = Pair (extract_mode t1, extract_mode t2)
| extract_mode (Free (x, _)) = if member (op =) output_names x then Output else Input
| extract_mode _ = Input
val user_mode = fold_rev (curry Fun) (map extract_mode all_args) Bool
--- a/src/HOL/Tools/Qelim/cooper.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Qelim/cooper.ML Mon Jun 21 17:41:57 2010 +0200
@@ -37,7 +37,7 @@
@{term "max :: int => _"}, @{term "max :: nat => _"},
@{term "min :: int => _"}, @{term "min :: nat => _"},
@{term "uminus :: int => _"}, (*@ {term "uminus :: nat => _"},*)
- @{term "Not"}, @{term "Suc"},
+ @{term "Not"}, @{term Suc},
@{term "Ex :: (int => _) => _"}, @{term "Ex :: (nat => _) => _"},
@{term "All :: (int => _) => _"}, @{term "All :: (nat => _) => _"},
@{term "nat"}, @{term "int"},
@@ -726,7 +726,7 @@
fun isnum t = case t of
Const(@{const_name Groups.zero},_) => true
| Const(@{const_name Groups.one},_) => true
- | @{term "Suc"}$s => isnum s
+ | @{term Suc}$s => isnum s
| @{term "nat"}$s => isnum s
| @{term "int"}$s => isnum s
| Const(@{const_name Groups.uminus},_)$s => isnum s
--- a/src/HOL/Tools/Sledgehammer/metis_tactics.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/metis_tactics.ML Mon Jun 21 17:41:57 2010 +0200
@@ -18,6 +18,7 @@
structure Metis_Tactics : METIS_TACTICS =
struct
+open Sledgehammer_Util
open Sledgehammer_FOL_Clause
open Sledgehammer_Fact_Preprocessor
open Sledgehammer_HOL_Clause
@@ -43,14 +44,16 @@
(* Useful Functions *)
(* ------------------------------------------------------------------------- *)
-(* match untyped terms*)
-fun untyped_aconv (Const(a,_)) (Const(b,_)) = (a=b)
- | untyped_aconv (Free(a,_)) (Free(b,_)) = (a=b)
- | untyped_aconv (Var((a,_),_)) (Var((b,_),_)) = (a=b) (*the index is ignored!*)
- | untyped_aconv (Bound i) (Bound j) = (i=j)
- | untyped_aconv (Abs(a,_,t)) (Abs(b,_,u)) = (a=b) andalso untyped_aconv t u
- | untyped_aconv (t1$t2) (u1$u2) = untyped_aconv t1 u1 andalso untyped_aconv t2 u2
- | untyped_aconv _ _ = false;
+(* Match untyped terms. *)
+fun untyped_aconv (Const (a, _)) (Const(b, _)) = (a = b)
+ | untyped_aconv (Free (a, _)) (Free (b, _)) = (a = b)
+ | untyped_aconv (Var ((a, _), _)) (Var ((b, _), _)) =
+ (a = b) (* The index is ignored, for some reason. *)
+ | untyped_aconv (Bound i) (Bound j) = (i = j)
+ | untyped_aconv (Abs (_, _, t)) (Abs (_, _, u)) = untyped_aconv t u
+ | untyped_aconv (t1 $ t2) (u1 $ u2) =
+ untyped_aconv t1 u1 andalso untyped_aconv t2 u2
+ | untyped_aconv _ _ = false
(* Finding the relative location of an untyped term within a list of terms *)
fun get_index lit =
@@ -84,7 +87,7 @@
val args = map hol_term_to_fol_FO tms
in Metis.Term.Fn (c, tyargs @ args) end
| (CombVar ((v, _), _), []) => Metis.Term.Var v
- | _ => error "hol_term_to_fol_FO";
+ | _ => raise Fail "hol_term_to_fol_FO";
fun hol_term_to_fol_HO (CombVar ((s, _), _)) = Metis.Term.Var s
| hol_term_to_fol_HO (CombConst ((a, _), _, tylist)) =
@@ -118,7 +121,7 @@
metis_lit pol "=" (map hol_term_to_fol_FT tms)
| _ => metis_lit pol "{}" [hol_term_to_fol_FT tm]) (*hBOOL*);
-fun literals_of_hol_thm thy mode t =
+fun literals_of_hol_term thy mode t =
let val (lits, types_sorts) = literals_of_term thy t
in (map (hol_literal_to_fol mode) lits, types_sorts) end;
@@ -134,21 +137,24 @@
fun metis_of_tfree tf =
Metis.Thm.axiom (Metis.LiteralSet.singleton (metis_of_type_literals true tf));
-fun hol_thm_to_fol is_conjecture ctxt mode th =
- let val thy = ProofContext.theory_of ctxt
- val (mlits, types_sorts) =
- (literals_of_hol_thm thy mode o HOLogic.dest_Trueprop o prop_of) th
+fun hol_thm_to_fol is_conjecture ctxt mode j skolem_somes th =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val (skolem_somes, (mlits, types_sorts)) =
+ th |> prop_of |> conceal_skolem_somes j skolem_somes
+ ||> (HOLogic.dest_Trueprop #> literals_of_hol_term thy mode)
in
if is_conjecture then
(Metis.Thm.axiom (Metis.LiteralSet.fromList mlits),
- type_literals_for_types types_sorts)
+ type_literals_for_types types_sorts, skolem_somes)
else
let val tylits = filter_out (default_sort ctxt) types_sorts
|> type_literals_for_types
val mtylits = if Config.get ctxt type_lits
then map (metis_of_type_literals false) tylits else []
in
- (Metis.Thm.axiom (Metis.LiteralSet.fromList(mtylits @ mlits)), [])
+ (Metis.Thm.axiom (Metis.LiteralSet.fromList(mtylits @ mlits)), [],
+ skolem_somes)
end
end;
@@ -194,7 +200,7 @@
fun apply_list rator nargs rands =
let val trands = terms_of rands
in if length trands = nargs then Term (list_comb(rator, trands))
- else error
+ else raise Fail
("apply_list: wrong number of arguments: " ^ Syntax.string_of_term_global Pure.thy rator ^
" expected " ^ Int.toString nargs ^
" received " ^ commas (map (Syntax.string_of_term_global Pure.thy) trands))
@@ -229,12 +235,22 @@
| NONE =>
case strip_prefix tfree_prefix x of
SOME tf => mk_tfree ctxt tf
- | NONE => error ("fol_type_to_isa: " ^ x));
+ | NONE => raise Fail ("fol_type_to_isa: " ^ x));
+
+fun reveal_skolem_somes skolem_somes =
+ map_aterms (fn t as Const (s, T) =>
+ if String.isPrefix skolem_theory_name s then
+ AList.lookup (op =) skolem_somes s
+ |> the |> map_types Type_Infer.paramify_vars
+ else
+ t
+ | t => t)
(*Maps metis terms to isabelle terms*)
-fun fol_term_to_hol_RAW ctxt fol_tm =
+fun hol_term_from_fol_PT ctxt fol_tm =
let val thy = ProofContext.theory_of ctxt
- val _ = trace_msg (fn () => "fol_term_to_hol: " ^ Metis.Term.toString fol_tm)
+ val _ = trace_msg (fn () => "hol_term_from_fol_PT: " ^
+ Metis.Term.toString fol_tm)
fun tm_to_tt (Metis.Term.Var v) =
(case strip_prefix tvar_prefix v of
SOME w => Type (make_tvar w)
@@ -250,7 +266,7 @@
Metis.Term.Fn(fname,ts) => applic_to_tt (fname, ts @ rands)
| _ => case tm_to_tt rator of
Term t => Term (list_comb(t, terms_of (map tm_to_tt rands)))
- | _ => error "tm_to_tt: HO application"
+ | _ => raise Fail "tm_to_tt: HO application"
end
| tm_to_tt (Metis.Term.Fn (fname, args)) = applic_to_tt (fname,args)
and applic_to_tt ("=",ts) =
@@ -265,12 +281,13 @@
val tys = types_of (List.take(tts,ntypes))
in if length tys = ntypes then
apply_list (Const (c, dummyT)) nterms (List.drop(tts,ntypes))
- else error ("Constant " ^ c ^ " expects " ^ Int.toString ntypes ^
- " but gets " ^ Int.toString (length tys) ^
- " type arguments\n" ^
- cat_lines (map (Syntax.string_of_typ ctxt) tys) ^
- " the terms are \n" ^
- cat_lines (map (Syntax.string_of_term ctxt) (terms_of tts)))
+ else
+ raise Fail ("Constant " ^ c ^ " expects " ^ Int.toString ntypes ^
+ " but gets " ^ Int.toString (length tys) ^
+ " type arguments\n" ^
+ cat_lines (map (Syntax.string_of_typ ctxt) tys) ^
+ " the terms are \n" ^
+ cat_lines (map (Syntax.string_of_term ctxt) (terms_of tts)))
end
| NONE => (*Not a constant. Is it a type constructor?*)
case strip_prefix tconst_prefix a of
@@ -284,12 +301,17 @@
SOME b =>
let val opr = Term.Free(b, HOLogic.typeT)
in apply_list opr (length ts) (map tm_to_tt ts) end
- | NONE => error ("unexpected metis function: " ^ a)
- in case tm_to_tt fol_tm of Term t => t | _ => error "fol_tm_to_tt: Term expected" end;
+ | NONE => raise Fail ("unexpected metis function: " ^ a)
+ in
+ case tm_to_tt fol_tm of
+ Term t => t
+ | _ => raise Fail "fol_tm_to_tt: Term expected"
+ end
(*Maps fully-typed metis terms to isabelle terms*)
-fun fol_term_to_hol_FT ctxt fol_tm =
- let val _ = trace_msg (fn () => "fol_term_to_hol_FT: " ^ Metis.Term.toString fol_tm)
+fun hol_term_from_fol_FT ctxt fol_tm =
+ let val _ = trace_msg (fn () => "hol_term_from_fol_FT: " ^
+ Metis.Term.toString fol_tm)
fun cvt (Metis.Term.Fn ("ti", [Metis.Term.Var v, _])) =
(case strip_prefix schematic_var_prefix v of
SOME w => mk_var(w, dummyT)
@@ -302,7 +324,7 @@
| NONE => (*Not a constant. Is it a fixed variable??*)
case strip_prefix fixed_var_prefix x of
SOME v => Free (v, fol_type_to_isa ctxt ty)
- | NONE => error ("fol_term_to_hol_FT bad constant: " ^ x))
+ | NONE => raise Fail ("hol_term_from_fol_FT bad constant: " ^ x))
| cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn (".",[tm1,tm2]), _])) =
cvt tm1 $ cvt tm2
| cvt (Metis.Term.Fn (".",[tm1,tm2])) = (*untyped application*)
@@ -316,21 +338,21 @@
| NONE => (*Not a constant. Is it a fixed variable??*)
case strip_prefix fixed_var_prefix x of
SOME v => Free (v, dummyT)
- | NONE => (trace_msg (fn () => "fol_term_to_hol_FT bad const: " ^ x);
- fol_term_to_hol_RAW ctxt t))
- | cvt t = (trace_msg (fn () => "fol_term_to_hol_FT bad term: " ^ Metis.Term.toString t);
- fol_term_to_hol_RAW ctxt t)
- in cvt fol_tm end;
+ | NONE => (trace_msg (fn () => "hol_term_from_fol_FT bad const: " ^ x);
+ hol_term_from_fol_PT ctxt t))
+ | cvt t = (trace_msg (fn () => "hol_term_from_fol_FT bad term: " ^ Metis.Term.toString t);
+ hol_term_from_fol_PT ctxt t)
+ in fol_tm |> cvt end
-fun fol_term_to_hol ctxt FO = fol_term_to_hol_RAW ctxt
- | fol_term_to_hol ctxt HO = fol_term_to_hol_RAW ctxt
- | fol_term_to_hol ctxt FT = fol_term_to_hol_FT ctxt;
+fun hol_term_from_fol FT = hol_term_from_fol_FT
+ | hol_term_from_fol _ = hol_term_from_fol_PT
-fun fol_terms_to_hol ctxt mode fol_tms =
- let val ts = map (fol_term_to_hol ctxt mode) fol_tms
+fun hol_terms_from_fol ctxt mode skolem_somes fol_tms =
+ let val thy = ProofContext.theory_of ctxt
+ val ts = map (hol_term_from_fol mode ctxt) fol_tms
val _ = trace_msg (fn () => " calling type inference:")
val _ = app (fn t => trace_msg (fn () => Syntax.string_of_term ctxt t)) ts
- val ts' = infer_types ctxt ts;
+ val ts' = ts |> map (reveal_skolem_somes skolem_somes) |> infer_types ctxt
val _ = app (fn t => trace_msg
(fn () => " final term: " ^ Syntax.string_of_term ctxt t ^
" of type " ^ Syntax.string_of_typ ctxt (type_of t)))
@@ -354,7 +376,8 @@
fun lookth thpairs (fth : Metis.Thm.thm) =
the (AList.lookup (uncurry Metis.Thm.equal) thpairs fth)
- handle Option => error ("Failed to find a Metis theorem " ^ Metis.Thm.toString fth);
+ handle Option =>
+ raise Fail ("Failed to find a Metis theorem " ^ Metis.Thm.toString fth);
fun is_TrueI th = Thm.eq_thm(TrueI,th);
@@ -371,22 +394,23 @@
(*This causes variables to have an index of 1 by default. SEE ALSO mk_var above.*)
(* INFERENCE RULE: ASSUME *)
-fun assume_inf ctxt mode atm =
+fun assume_inf ctxt mode skolem_somes atm =
inst_excluded_middle
(ProofContext.theory_of ctxt)
- (singleton (fol_terms_to_hol ctxt mode) (Metis.Term.Fn atm));
+ (singleton (hol_terms_from_fol ctxt mode skolem_somes) (Metis.Term.Fn atm))
(* INFERENCE RULE: INSTANTIATE (Subst). Type instantiations are ignored. Trying to reconstruct
them admits new possibilities of errors, e.g. concerning sorts. Instead we try to arrange
that new TVars are distinct and that types can be inferred from terms.*)
-fun inst_inf ctxt mode thpairs fsubst th =
+fun inst_inf ctxt mode skolem_somes thpairs fsubst th =
let val thy = ProofContext.theory_of ctxt
val i_th = lookth thpairs th
val i_th_vars = Term.add_vars (prop_of i_th) []
fun find_var x = the (List.find (fn ((a,_),_) => a=x) i_th_vars)
fun subst_translation (x,y) =
let val v = find_var x
- val t = fol_term_to_hol ctxt mode y (*we call infer_types below*)
+ (* We call "reveal_skolem_somes" and "infer_types" below. *)
+ val t = hol_term_from_fol mode ctxt y
in SOME (cterm_of thy (Var v), t) end
handle Option =>
(trace_msg (fn() => "List.find failed for the variable " ^ x ^
@@ -401,7 +425,8 @@
val _ = trace_msg (fn () => " isa th: " ^ Display.string_of_thm ctxt i_th)
val substs = map_filter remove_typeinst (Metis.Subst.toList fsubst)
val (vars,rawtms) = ListPair.unzip (map_filter subst_translation substs)
- val tms = infer_types ctxt rawtms;
+ val tms = rawtms |> map (reveal_skolem_somes skolem_somes)
+ |> infer_types ctxt
val ctm_of = cterm_incr_types thy (1 + Thm.maxidx_of i_th)
val substs' = ListPair.zip (vars, map ctm_of tms)
val _ = trace_msg (fn () =>
@@ -409,10 +434,11 @@
(substs' |> map (fn (x, y) =>
Syntax.string_of_term ctxt (term_of x) ^ " |-> " ^
Syntax.string_of_term ctxt (term_of y)))));
- in cterm_instantiate substs' i_th end
+ in cterm_instantiate substs' i_th end
handle THM (msg, _, _) => error ("metis error (inst_inf):\n" ^ msg)
| ERROR msg => error ("metis error (inst_inf):\n" ^ msg ^
- "\n(Perhaps you want to try \"metisFT\".)")
+ "\n(Perhaps you want to try \"metisFT\" if you \
+ \haven't done so already.)")
(* INFERENCE RULE: RESOLVE *)
@@ -428,7 +454,7 @@
| _ => raise THM ("resolve_inc_tyvars: unique result expected", i, [tha,thb])
end;
-fun resolve_inf ctxt mode thpairs atm th1 th2 =
+fun resolve_inf ctxt mode skolem_somes thpairs atm th1 th2 =
let
val i_th1 = lookth thpairs th1 and i_th2 = lookth thpairs th2
val _ = trace_msg (fn () => " isa th1 (pos): " ^ Display.string_of_thm ctxt i_th1)
@@ -438,15 +464,16 @@
else if is_TrueI i_th2 then i_th1
else
let
- val i_atm = singleton (fol_terms_to_hol ctxt mode) (Metis.Term.Fn atm)
+ val i_atm = singleton (hol_terms_from_fol ctxt mode skolem_somes)
+ (Metis.Term.Fn atm)
val _ = trace_msg (fn () => " atom: " ^ Syntax.string_of_term ctxt i_atm)
val prems_th1 = prems_of i_th1
val prems_th2 = prems_of i_th2
val index_th1 = get_index (mk_not i_atm) prems_th1
- handle Empty => error "Failed to find literal in th1"
+ handle Empty => raise Fail "Failed to find literal in th1"
val _ = trace_msg (fn () => " index_th1: " ^ Int.toString index_th1)
val index_th2 = get_index i_atm prems_th2
- handle Empty => error "Failed to find literal in th2"
+ handle Empty => raise Fail "Failed to find literal in th2"
val _ = trace_msg (fn () => " index_th2: " ^ Int.toString index_th2)
in resolve_inc_tyvars (Meson.select_literal index_th1 i_th1, index_th2, i_th2) end
end;
@@ -455,9 +482,9 @@
val refl_x = cterm_of @{theory} (Var (hd (Term.add_vars (prop_of REFL_THM) [])));
val refl_idx = 1 + Thm.maxidx_of REFL_THM;
-fun refl_inf ctxt mode t =
+fun refl_inf ctxt mode skolem_somes t =
let val thy = ProofContext.theory_of ctxt
- val i_t = singleton (fol_terms_to_hol ctxt mode) t
+ val i_t = singleton (hol_terms_from_fol ctxt mode skolem_somes) t
val _ = trace_msg (fn () => " term: " ^ Syntax.string_of_term ctxt i_t)
val c_t = cterm_incr_types thy refl_idx i_t
in cterm_instantiate [(refl_x, c_t)] REFL_THM end;
@@ -467,10 +494,10 @@
| get_ty_arg_size _ _ = 0;
(* INFERENCE RULE: EQUALITY *)
-fun equality_inf ctxt mode (pos, atm) fp fr =
+fun equality_inf ctxt mode skolem_somes (pos, atm) fp fr =
let val thy = ProofContext.theory_of ctxt
val m_tm = Metis.Term.Fn atm
- val [i_atm,i_tm] = fol_terms_to_hol ctxt mode [m_tm, fr]
+ val [i_atm,i_tm] = hol_terms_from_fol ctxt mode skolem_somes [m_tm, fr]
val _ = trace_msg (fn () => "sign of the literal: " ^ Bool.toString pos)
fun replace_item_list lx 0 (_::ls) = lx::ls
| replace_item_list lx i (l::ls) = l :: replace_item_list lx (i-1) ls
@@ -491,6 +518,10 @@
fun path_finder_HO tm [] = (tm, Term.Bound 0)
| path_finder_HO (t$u) (0::ps) = (fn(x,y) => (x, y$u)) (path_finder_HO t ps)
| path_finder_HO (t$u) (_::ps) = (fn(x,y) => (x, t$y)) (path_finder_HO u ps)
+ | path_finder_HO tm ps =
+ raise Fail ("equality_inf, path_finder_HO: path = " ^
+ space_implode " " (map Int.toString ps) ^
+ " isa-term: " ^ Syntax.string_of_term ctxt tm)
fun path_finder_FT tm [] _ = (tm, Term.Bound 0)
| path_finder_FT tm (0::ps) (Metis.Term.Fn ("ti", [t1, _])) =
path_finder_FT tm ps t1
@@ -498,10 +529,11 @@
(fn(x,y) => (x, y$u)) (path_finder_FT t ps t1)
| path_finder_FT (t$u) (1::ps) (Metis.Term.Fn (".", [_, t2])) =
(fn(x,y) => (x, t$y)) (path_finder_FT u ps t2)
- | path_finder_FT tm ps t = error ("equality_inf, path_finder_FT: path = " ^
- space_implode " " (map Int.toString ps) ^
- " isa-term: " ^ Syntax.string_of_term ctxt tm ^
- " fol-term: " ^ Metis.Term.toString t)
+ | path_finder_FT tm ps t =
+ raise Fail ("equality_inf, path_finder_FT: path = " ^
+ space_implode " " (map Int.toString ps) ^
+ " isa-term: " ^ Syntax.string_of_term ctxt tm ^
+ " fol-term: " ^ Metis.Term.toString t)
fun path_finder FO tm ps _ = path_finder_FO tm ps
| path_finder HO (tm as Const(@{const_name "op ="},_) $ _ $ _) (p::ps) _ =
(*equality: not curried, as other predicates are*)
@@ -539,24 +571,28 @@
val factor = Seq.hd o distinct_subgoals_tac;
-fun step _ _ thpairs (fol_th, Metis.Proof.Axiom _) = factor (axiom_inf thpairs fol_th)
- | step ctxt mode _ (_, Metis.Proof.Assume f_atm) = assume_inf ctxt mode f_atm
- | step ctxt mode thpairs (_, Metis.Proof.Subst (f_subst, f_th1)) =
- factor (inst_inf ctxt mode thpairs f_subst f_th1)
- | step ctxt mode thpairs (_, Metis.Proof.Resolve(f_atm, f_th1, f_th2)) =
- factor (resolve_inf ctxt mode thpairs f_atm f_th1 f_th2)
- | step ctxt mode _ (_, Metis.Proof.Refl f_tm) = refl_inf ctxt mode f_tm
- | step ctxt mode _ (_, Metis.Proof.Equality (f_lit, f_p, f_r)) =
- equality_inf ctxt mode f_lit f_p f_r;
+fun step ctxt mode skolem_somes thpairs p =
+ case p of
+ (fol_th, Metis.Proof.Axiom _) => factor (axiom_inf thpairs fol_th)
+ | (_, Metis.Proof.Assume f_atm) => assume_inf ctxt mode skolem_somes f_atm
+ | (_, Metis.Proof.Subst (f_subst, f_th1)) =>
+ factor (inst_inf ctxt mode skolem_somes thpairs f_subst f_th1)
+ | (_, Metis.Proof.Resolve(f_atm, f_th1, f_th2)) =>
+ factor (resolve_inf ctxt mode skolem_somes thpairs f_atm f_th1 f_th2)
+ | (_, Metis.Proof.Refl f_tm) => refl_inf ctxt mode skolem_somes f_tm
+ | (_, Metis.Proof.Equality (f_lit, f_p, f_r)) =>
+ equality_inf ctxt mode skolem_somes f_lit f_p f_r
fun real_literal (_, (c, _)) = not (String.isPrefix class_prefix c);
-fun translate _ _ thpairs [] = thpairs
- | translate mode ctxt thpairs ((fol_th, inf) :: infpairs) =
+(* FIXME: use "fold" instead *)
+fun translate _ _ _ thpairs [] = thpairs
+ | translate ctxt mode skolem_somes thpairs ((fol_th, inf) :: infpairs) =
let val _ = trace_msg (fn () => "=============================================")
val _ = trace_msg (fn () => "METIS THM: " ^ Metis.Thm.toString fol_th)
val _ = trace_msg (fn () => "INFERENCE: " ^ Metis.Proof.inferenceToString inf)
- val th = Meson.flexflex_first_order (step ctxt mode thpairs (fol_th, inf))
+ val th = Meson.flexflex_first_order (step ctxt mode skolem_somes
+ thpairs (fol_th, inf))
val _ = trace_msg (fn () => "ISABELLE THM: " ^ Display.string_of_thm ctxt th)
val _ = trace_msg (fn () => "=============================================")
val n_metis_lits =
@@ -564,7 +600,7 @@
in
if nprems_of th = n_metis_lits then ()
else error "Metis: proof reconstruction has gone wrong";
- translate mode ctxt ((fol_th, th) :: thpairs) infpairs
+ translate ctxt mode skolem_somes ((fol_th, th) :: thpairs) infpairs
end;
(*Determining which axiom clauses are actually used*)
@@ -575,7 +611,8 @@
(* Translation of HO Clauses *)
(* ------------------------------------------------------------------------- *)
-fun cnf_th thy th = hd (cnf_axiom thy th);
+fun cnf_helper_theorem thy raw th =
+ if raw then th else the_single (cnf_axiom thy th)
fun type_ext thy tms =
let val subs = tfree_classes_of_terms tms
@@ -592,7 +629,8 @@
type logic_map =
{axioms: (Metis.Thm.thm * thm) list,
- tfrees: type_literal list};
+ tfrees: type_literal list,
+ skolem_somes: (string * term) list}
fun const_in_metis c (pred, tm_list) =
let
@@ -610,50 +648,74 @@
(*transform isabelle type / arity clause to metis clause *)
fun add_type_thm [] lmap = lmap
- | add_type_thm ((ith, mth) :: cls) {axioms, tfrees} =
- add_type_thm cls {axioms = (mth, ith) :: axioms,
- tfrees = tfrees}
+ | add_type_thm ((ith, mth) :: cls) {axioms, tfrees, skolem_somes} =
+ add_type_thm cls {axioms = (mth, ith) :: axioms, tfrees = tfrees,
+ skolem_somes = skolem_somes}
(*Insert non-logical axioms corresponding to all accumulated TFrees*)
-fun add_tfrees {axioms, tfrees} : logic_map =
- {axioms = (map (fn tf => (metis_of_tfree tf, TrueI)) (distinct op= tfrees)) @ axioms,
- tfrees = tfrees};
+fun add_tfrees {axioms, tfrees, skolem_somes} : logic_map =
+ {axioms = map (rpair TrueI o metis_of_tfree) (distinct (op =) tfrees) @
+ axioms,
+ tfrees = tfrees, skolem_somes = skolem_somes}
fun string_of_mode FO = "FO"
| string_of_mode HO = "HO"
| string_of_mode FT = "FT"
+val helpers =
+ [("c_COMBI", (false, @{thms COMBI_def})),
+ ("c_COMBK", (false, @{thms COMBK_def})),
+ ("c_COMBB", (false, @{thms COMBB_def})),
+ ("c_COMBC", (false, @{thms COMBC_def})),
+ ("c_COMBS", (false, @{thms COMBS_def})),
+ ("c_fequal", (false, @{thms fequal_imp_equal equal_imp_fequal})),
+ ("c_True", (true, @{thms True_or_False})),
+ ("c_False", (true, @{thms True_or_False})),
+ ("c_If", (true, @{thms if_True if_False True_or_False}))]
+
(* Function to generate metis clauses, including comb and type clauses *)
fun build_map mode0 ctxt cls ths =
let val thy = ProofContext.theory_of ctxt
(*The modes FO and FT are sticky. HO can be downgraded to FO.*)
fun set_mode FO = FO
- | set_mode HO = if forall (Meson.is_fol_term thy o prop_of) (cls@ths) then FO else HO
+ | set_mode HO =
+ if forall (is_quasi_fol_term thy o prop_of) (cls @ ths) then FO
+ else HO
| set_mode FT = FT
val mode = set_mode mode0
(*transform isabelle clause to metis clause *)
- fun add_thm is_conjecture ith {axioms, tfrees} : logic_map =
- let val (mth, tfree_lits) = hol_thm_to_fol is_conjecture ctxt mode ith
+ fun add_thm is_conjecture ith {axioms, tfrees, skolem_somes} : logic_map =
+ let
+ val (mth, tfree_lits, skolem_somes) =
+ hol_thm_to_fol is_conjecture ctxt mode (length axioms) skolem_somes
+ ith
in
{axioms = (mth, Meson.make_meta_clause ith) :: axioms,
- tfrees = union (op =) tfree_lits tfrees}
+ tfrees = union (op =) tfree_lits tfrees,
+ skolem_somes = skolem_somes}
end;
- val lmap0 = fold (add_thm true) cls {axioms = [], tfrees = init_tfrees ctxt}
- val lmap = fold (add_thm false) ths (add_tfrees lmap0)
+ val lmap as {skolem_somes, ...} =
+ {axioms = [], tfrees = init_tfrees ctxt, skolem_somes = []}
+ |> fold (add_thm true) cls
+ |> add_tfrees
+ |> fold (add_thm false) ths
val clause_lists = map (Metis.Thm.clause o #1) (#axioms lmap)
- fun used c = exists (Metis.LiteralSet.exists (const_in_metis c o #2)) clause_lists
- (*Now check for the existence of certain combinators*)
- val thI = if used "c_COMBI" then [cnf_th @{theory} @{thm COMBI_def}] else []
- val thK = if used "c_COMBK" then [cnf_th @{theory} @{thm COMBK_def}] else []
- val thB = if used "c_COMBB" then [cnf_th @{theory} @{thm COMBB_def}] else []
- val thC = if used "c_COMBC" then [cnf_th @{theory} @{thm COMBC_def}] else []
- val thS = if used "c_COMBS" then [cnf_th @{theory} @{thm COMBS_def}] else []
- val thEQ = if used "c_fequal" then [cnf_th @{theory} @{thm fequal_imp_equal}, cnf_th @{theory} @{thm equal_imp_fequal}] else []
- val lmap' = if mode=FO then lmap
- else fold (add_thm false) (thEQ @ thS @ thC @ thB @ thK @ thI) lmap
- in
- (mode, add_type_thm (type_ext thy (map prop_of (cls @ ths))) lmap')
- end;
+ fun is_used c =
+ exists (Metis.LiteralSet.exists (const_in_metis c o #2)) clause_lists
+ val lmap =
+ if mode = FO then
+ lmap
+ else
+ let
+ val helper_ths =
+ helpers |> filter (is_used o fst)
+ |> maps (fn (_, (raw, thms)) =>
+ if mode = FT orelse not raw then
+ map (cnf_helper_theorem @{theory} raw) thms
+ else
+ [])
+ in lmap |> fold (add_thm false) helper_ths end
+ in (mode, add_type_thm (type_ext thy (map prop_of (cls @ ths))) lmap) end
fun refute cls =
Metis.Resolution.loop (Metis.Resolution.new Metis.Resolution.default cls);
@@ -664,7 +726,7 @@
exception METIS of string;
-(* Main function to start metis prove and reconstruction *)
+(* Main function to start metis proof and reconstruction *)
fun FOL_SOLVE mode ctxt cls ths0 =
let val thy = ProofContext.theory_of ctxt
val th_cls_pairs =
@@ -674,7 +736,7 @@
val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) cls
val _ = trace_msg (fn () => "THEOREM CLAUSES")
val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) ths
- val (mode, {axioms,tfrees}) = build_map mode ctxt cls ths
+ val (mode, {axioms, tfrees, skolem_somes}) = build_map mode ctxt cls ths
val _ = if null tfrees then ()
else (trace_msg (fn () => "TFREE CLAUSES");
app (fn tf => trace_msg (fn _ => tptp_of_type_literal true tf NONE |> fst)) tfrees)
@@ -694,7 +756,7 @@
val ctxt' = fold Variable.declare_constraints (map prop_of cls) ctxt
(*add constraints arising from converting goal to clause form*)
val proof = Metis.Proof.proof mth
- val result = translate mode ctxt' axioms proof
+ val result = translate ctxt' mode skolem_somes axioms proof
and used = map_filter (used_axioms axioms) proof
val _ = trace_msg (fn () => "METIS COMPLETED...clauses actually used:")
val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) used
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML Mon Jun 21 17:41:57 2010 +0200
@@ -21,12 +21,14 @@
val tvar_classes_of_terms : term list -> string list
val tfree_classes_of_terms : term list -> string list
val type_consts_of_terms : theory -> term list -> string list
+ val is_quasi_fol_term : theory -> term -> bool
val relevant_facts :
bool -> bool -> real -> real -> bool -> int -> bool -> relevance_override
-> Proof.context * (thm list * 'a) -> thm list
-> (thm * (string * int)) list
val prepare_clauses :
- bool -> thm list -> thm list -> (thm * (axiom_name * hol_clause_id)) list
+ bool -> bool -> thm list -> thm list
+ -> (thm * (axiom_name * hol_clause_id)) list
-> (thm * (axiom_name * hol_clause_id)) list -> theory
-> axiom_name vector
* (hol_clause list * hol_clause list * hol_clause list *
@@ -116,6 +118,8 @@
add_const_typ_table (const_with_typ thy (c,typ), tab)
| add_term_consts_typs_rm thy (Free(c, typ), tab) =
add_const_typ_table (const_with_typ thy (c,typ), tab)
+ | add_term_consts_typs_rm _ (Const (@{const_name skolem_id}, _) $ _, tab) =
+ tab
| add_term_consts_typs_rm thy (t $ u, tab) =
add_term_consts_typs_rm thy (t, add_term_consts_typs_rm thy (u, tab))
| add_term_consts_typs_rm thy (Abs(_,_,t), tab) = add_term_consts_typs_rm thy (t, tab)
@@ -368,38 +372,35 @@
fun valid_facts facts =
(facts, []) |-> Facts.fold_static (fn (name, ths0) =>
- let
- fun check_thms a =
- (case try (ProofContext.get_thms ctxt) a of
- NONE => false
- | SOME ths1 => Thm.eq_thms (ths0, ths1));
+ if Facts.is_concealed facts name orelse
+ (respect_no_atp andalso is_package_def name) orelse
+ member (op =) multi_base_blacklist (Long_Name.base_name name) then
+ I
+ else
+ let
+ fun check_thms a =
+ (case try (ProofContext.get_thms ctxt) a of
+ NONE => false
+ | SOME ths1 => Thm.eq_thms (ths0, ths1));
- val name1 = Facts.extern facts name;
- val name2 = Name_Space.extern full_space name;
- val ths = filter_out is_theorem_bad_for_atps ths0;
- in
- if Facts.is_concealed facts name orelse
- (respect_no_atp andalso is_package_def name) then
- I
- else case find_first check_thms [name1, name2, name] of
- NONE => I
- | SOME name' =>
- cons (name' |> forall (member Thm.eq_thm chained_ths) ths
- ? prefix chained_prefix, ths)
- end);
+ val name1 = Facts.extern facts name;
+ val name2 = Name_Space.extern full_space name;
+ val ths = filter_out is_theorem_bad_for_atps ths0;
+ in
+ case find_first check_thms [name1, name2, name] of
+ NONE => I
+ | SOME name' =>
+ cons (name' |> forall (member Thm.eq_thm chained_ths) ths
+ ? prefix chained_prefix, ths)
+ end)
in valid_facts global_facts @ valid_facts local_facts end;
fun multi_name a th (n, pairs) =
(n + 1, (a ^ "(" ^ Int.toString n ^ ")", th) :: pairs);
-fun add_single_names (a, []) pairs = pairs
- | add_single_names (a, [th]) pairs = (a, th) :: pairs
- | add_single_names (a, ths) pairs = #2 (fold (multi_name a) ths (1, pairs));
-
-(*Ignore blacklisted basenames*)
-fun add_multi_names (a, ths) pairs =
- if member (op =) multi_base_blacklist (Long_Name.base_name a) then pairs
- else add_single_names (a, ths) pairs;
+fun add_names (a, []) pairs = pairs
+ | add_names (a, [th]) pairs = (a, th) :: pairs
+ | add_names (a, ths) pairs = #2 (fold (multi_name a) ths (1, pairs))
fun is_multi (a, ths) = length ths > 1 orelse String.isSuffix ".axioms" a;
@@ -408,8 +409,7 @@
fun name_thm_pairs respect_no_atp ctxt name_thms_pairs =
let
val (mults, singles) = List.partition is_multi name_thms_pairs
- val ps = [] |> fold add_single_names singles
- |> fold add_multi_names mults
+ val ps = [] |> fold add_names singles |> fold add_names mults
in ps |> respect_no_atp ? filter_out (No_ATPs.member ctxt o snd) end;
fun is_named ("", th) =
@@ -472,8 +472,12 @@
(* ATP invocation methods setup *)
(***************************************************************)
+fun is_quasi_fol_term thy =
+ Meson.is_fol_term thy o snd o conceal_skolem_somes ~1 []
+
(*Ensures that no higher-order theorems "leak out"*)
-fun restrict_to_logic thy true cls = filter (Meson.is_fol_term thy o prop_of o fst) cls
+fun restrict_to_logic thy true cls =
+ filter (is_quasi_fol_term thy o prop_of o fst) cls
| restrict_to_logic thy false cls = cls;
(**** Predicates to detect unwanted clauses (prolific or likely to cause
@@ -531,7 +535,7 @@
fun remove_dangerous_clauses full_types add_thms =
filter_out (is_dangerous_theorem full_types add_thms o fst)
-fun is_first_order thy = forall (Meson.is_fol_term thy) o map prop_of
+fun is_fol_goal thy = forall (Meson.is_fol_term thy) o map prop_of
fun relevant_facts full_types respect_no_atp relevance_threshold
relevance_convergence defs_relevant max_new theory_relevant
@@ -542,17 +546,17 @@
else
let
val thy = ProofContext.theory_of ctxt
- val add_thms = cnf_for_facts ctxt add
val has_override = not (null add) orelse not (null del)
- val is_FO = is_first_order thy goal_cls
+ val is_FO = is_fol_goal thy goal_cls
val axioms =
- checked_name_thm_pairs respect_no_atp ctxt
+ checked_name_thm_pairs (respect_no_atp andalso not only) ctxt
(if only then map (name_thms_pair_from_ref ctxt chained_ths) add
else all_name_thms_pairs respect_no_atp ctxt chained_ths)
|> cnf_rules_pairs thy
|> not has_override ? make_unique
- |> restrict_to_logic thy is_FO
- |> not only ? remove_dangerous_clauses full_types add_thms
+ |> not only ? restrict_to_logic thy is_FO
+ |> (if only then I
+ else remove_dangerous_clauses full_types (cnf_for_facts ctxt add))
in
relevance_filter ctxt relevance_threshold relevance_convergence
defs_relevant max_new theory_relevant relevance_override
@@ -562,9 +566,9 @@
(* prepare for passing to writer,
create additional clauses based on the information from extra_cls *)
-fun prepare_clauses dfg goal_cls chained_ths axcls extra_cls thy =
+fun prepare_clauses dfg full_types goal_cls chained_ths axcls extra_cls thy =
let
- val is_FO = is_first_order thy goal_cls
+ val is_FO = is_fol_goal thy goal_cls
val ccls = subtract_cls extra_cls goal_cls
val _ = app (fn th => trace_msg (fn _ => Display.string_of_thm_global thy th)) ccls
val ccltms = map prop_of ccls
@@ -576,7 +580,8 @@
val conjectures = make_conjecture_clauses dfg thy ccls
val (_, extra_clauses) = ListPair.unzip (make_axiom_clauses dfg thy extra_cls)
val (clnames, axiom_clauses) = ListPair.unzip (make_axiom_clauses dfg thy axcls)
- val helper_clauses = get_helper_clauses dfg thy is_FO (conjectures, extra_cls, [])
+ val helper_clauses =
+ get_helper_clauses dfg thy is_FO full_types conjectures extra_cls
val (supers', arity_clauses) = make_arity_clauses_dfg dfg thy tycons supers
val classrel_clauses = make_classrel_clauses thy subs supers'
in
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact_minimizer.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact_minimizer.ML Mon Jun 21 17:41:57 2010 +0200
@@ -38,6 +38,7 @@
(* wrapper for calling external prover *)
fun string_for_failure Unprovable = "Unprovable."
+ | string_for_failure IncompleteUnprovable = "Failed."
| string_for_failure TimedOut = "Timed out."
| string_for_failure OutOfResources = "Failed."
| string_for_failure OldSpass = "Error."
@@ -110,9 +111,8 @@
in
(SOME min_thms,
proof_text isar_proof
- (pool, debug, full_types, isar_shrink_factor, ctxt,
- conjecture_shape)
- (K "", proof, internal_thm_names, goal, i) |> fst)
+ (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
+ (full_types, K "", proof, internal_thm_names, goal, i) |> fst)
end
| {outcome = SOME TimedOut, ...} =>
(NONE, "Timeout: You can increase the time limit using the \"timeout\" \
@@ -125,7 +125,8 @@
\option (e.g., \"timeout = " ^
string_of_int (10 + msecs div 1000) ^ " s\").")
| {message, ...} => (NONE, "ATP error: " ^ message))
- handle Sledgehammer_HOL_Clause.TRIVIAL => (SOME [], metis_line i n [])
+ handle Sledgehammer_HOL_Clause.TRIVIAL =>
+ (SOME [], metis_line full_types i n [])
| ERROR msg => (NONE, "Error: " ^ msg)
end
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact_preprocessor.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact_preprocessor.ML Mon Jun 21 17:41:57 2010 +0200
@@ -10,8 +10,10 @@
val chained_prefix: string
val trace: bool Unsynchronized.ref
val trace_msg: (unit -> string) -> unit
+ val skolem_theory_name: string
val skolem_prefix: string
val skolem_infix: string
+ val is_skolem_const_name: string -> bool
val cnf_axiom: theory -> thm -> thm list
val multi_base_blacklist: string list
val is_theorem_bad_for_atps: thm -> bool
@@ -39,6 +41,7 @@
val trace = Unsynchronized.ref false;
fun trace_msg msg = if !trace then tracing (msg ()) else ();
+val skolem_theory_name = "Sledgehammer.Sko"
val skolem_prefix = "sko_"
val skolem_infix = "$"
@@ -75,10 +78,16 @@
(*Keep the full complexity of the original name*)
fun flatten_name s = space_implode "_X" (Long_Name.explode s);
-fun skolem_name thm_name nref var_name =
- skolem_prefix ^ thm_name ^ "_" ^ Int.toString (Unsynchronized.inc nref) ^
+fun skolem_name thm_name j var_name =
+ skolem_prefix ^ thm_name ^ "_" ^ Int.toString j ^
skolem_infix ^ (if var_name = "" then "g" else flatten_name var_name)
+(* Hack: Could return false positives (e.g., a user happens to declare a
+ constant called "SomeTheory.sko_means_shoe_in_$wedish". *)
+val is_skolem_const_name =
+ Long_Name.base_name
+ #> String.isPrefix skolem_prefix andf String.isSubstring skolem_infix
+
fun rhs_extra_types lhsT rhs =
let val lhs_vars = Term.add_tfreesT lhsT []
fun add_new_TFrees (TFree v) =
@@ -87,64 +96,78 @@
val rhs_consts = fold_aterms (fn Const c => insert (op =) c | _ => I) rhs []
in fold (#2 #> Term.fold_atyps add_new_TFrees) rhs_consts [] end;
+fun skolem_type_and_args bound_T body =
+ let
+ val args1 = OldTerm.term_frees body
+ val Ts1 = map type_of args1
+ val Ts2 = rhs_extra_types (Ts1 ---> bound_T) body
+ val args2 = map (fn T => Free (gensym "vsk", T)) Ts2
+ in (Ts2 ---> Ts1 ---> bound_T, args2 @ args1) end
+
(* Traverse a theorem, declaring Skolem function definitions. String "s" is the
suggested prefix for the Skolem constants. *)
fun declare_skolem_funs s th thy =
let
- val nref = Unsynchronized.ref 0 (* FIXME ??? *)
- fun dec_sko (Const (@{const_name Ex}, _) $ (xtp as Abs (s', T, p))) (axs, thy) =
- (*Existential: declare a Skolem function, then insert into body and continue*)
- let
- val cname = skolem_name s nref s'
- val args0 = OldTerm.term_frees xtp (*get the formal parameter list*)
- val Ts = map type_of args0
- val extraTs = rhs_extra_types (Ts ---> T) xtp
- val argsx = map (fn T => Free (gensym "vsk", T)) extraTs
- val args = argsx @ args0
- val cT = extraTs ---> Ts ---> T
- val rhs = list_abs_free (map dest_Free args, HOLogic.choice_const T $ xtp)
- (*Forms a lambda-abstraction over the formal parameters*)
- val (c, thy) =
- Sign.declare_const ((Binding.conceal (Binding.name cname), cT), NoSyn) thy
- val cdef = cname ^ "_def"
- val ((_, ax), thy) =
- Thm.add_def true false (Binding.name cdef, Logic.mk_equals (c, rhs)) thy
- val ax' = Drule.export_without_context ax
- in dec_sko (subst_bound (list_comb (c, args), p)) (ax' :: axs, thy) end
+ val skolem_count = Unsynchronized.ref 0 (* FIXME ??? *)
+ fun dec_sko (Const (@{const_name Ex}, _) $ (body as Abs (s', T, p)))
+ (axs, thy) =
+ (*Existential: declare a Skolem function, then insert into body and continue*)
+ let
+ val id = skolem_name s (Unsynchronized.inc skolem_count) s'
+ val (cT, args) = skolem_type_and_args T body
+ val rhs = list_abs_free (map dest_Free args,
+ HOLogic.choice_const T $ body)
+ (*Forms a lambda-abstraction over the formal parameters*)
+ val (c, thy) =
+ Sign.declare_const ((Binding.conceal (Binding.name id), cT), NoSyn) thy
+ val cdef = id ^ "_def"
+ val ((_, ax), thy) =
+ Thm.add_def true false (Binding.name cdef, Logic.mk_equals (c, rhs)) thy
+ val ax' = Drule.export_without_context ax
+ in dec_sko (subst_bound (list_comb (c, args), p)) (ax' :: axs, thy) end
| dec_sko (Const (@{const_name All}, _) $ (Abs (a, T, p))) thx =
- (*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)) thx end
+ (*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)) thx end
| dec_sko (@{const "op &"} $ p $ q) thx = dec_sko q (dec_sko p thx)
| dec_sko (@{const "op |"} $ p $ q) thx = dec_sko q (dec_sko p thx)
| dec_sko (@{const Trueprop} $ p) thx = dec_sko p thx
| dec_sko t thx = thx
in dec_sko (prop_of th) ([], thy) end
+fun mk_skolem_id t =
+ let val T = fastype_of t in
+ Const (@{const_name skolem_id}, T --> T) $ Envir.beta_eta_contract t
+ end
+
(*Traverse a theorem, accumulating Skolem function definitions.*)
-fun assume_skolem_funs s th =
- let val sko_count = Unsynchronized.ref 0 (* FIXME ??? *)
- fun dec_sko (Const (@{const_name Ex}, _) $ (xtp as Abs (s', T, p))) defs =
- (*Existential: declare a Skolem function, then insert into body and continue*)
- let val skos = map (#1 o Logic.dest_equals) defs (*existing sko fns*)
- val args = subtract (op =) skos (OldTerm.term_frees xtp) (*the formal parameters*)
- val Ts = map type_of args
- val cT = Ts ---> T
- val id = skolem_name s sko_count s'
- val c = Free (id, cT)
- val rhs = list_abs_free (map dest_Free args,
- HOLogic.choice_const T $ xtp)
- (*Forms a lambda-abstraction over the formal parameters*)
- val def = Logic.mk_equals (c, rhs)
- in dec_sko (subst_bound (list_comb(c,args), p)) (def :: defs) end
- | dec_sko (Const (@{const_name All},_) $ Abs (a, T, p)) defs =
- (*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)) defs end
- | dec_sko (@{const "op &"} $ p $ q) defs = dec_sko q (dec_sko p defs)
- | dec_sko (@{const "op |"} $ p $ q) defs = dec_sko q (dec_sko p defs)
- | dec_sko (@{const Trueprop} $ p) defs = dec_sko p defs
- | dec_sko t defs = defs (*Do nothing otherwise*)
+fun assume_skolem_funs inline s th =
+ let
+ val skolem_count = Unsynchronized.ref 0 (* FIXME ??? *)
+ fun dec_sko (Const (@{const_name Ex}, _) $ (body as Abs (s', T, p))) defs =
+ (*Existential: declare a Skolem function, then insert into body and continue*)
+ let
+ val skos = map (#1 o Logic.dest_equals) defs (*existing sko fns*)
+ val args = subtract (op =) skos (OldTerm.term_frees body) (*the formal parameters*)
+ val Ts = map type_of args
+ val cT = Ts ---> T (* FIXME: use "skolem_type_and_args" *)
+ val id = skolem_name s (Unsynchronized.inc skolem_count) s'
+ val c = Free (id, cT)
+ val rhs = list_abs_free (map dest_Free args,
+ HOLogic.choice_const T $ body)
+ |> inline ? mk_skolem_id
+ (*Forms a lambda-abstraction over the formal parameters*)
+ val def = Logic.mk_equals (c, rhs)
+ val comb = list_comb (if inline then rhs else c, args)
+ in dec_sko (subst_bound (comb, p)) (def :: defs) end
+ | dec_sko (Const (@{const_name All},_) $ Abs (a, T, p)) defs =
+ (*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)) defs end
+ | dec_sko (@{const "op &"} $ p $ q) defs = dec_sko q (dec_sko p defs)
+ | dec_sko (@{const "op |"} $ p $ q) defs = dec_sko q (dec_sko p defs)
+ | dec_sko (@{const Trueprop} $ p) defs = dec_sko p defs
+ | dec_sko t defs = defs (*Do nothing otherwise*)
in dec_sko (prop_of th) [] end;
@@ -174,7 +197,11 @@
strip_lambdas (n-1) (freeze_thm (th RS xfun_cong x))
| _ => th;
-val lambda_free = not o Term.has_abs;
+fun is_quasi_lambda_free (Const (@{const_name skolem_id}, _) $ _) = true
+ | is_quasi_lambda_free (t1 $ t2) =
+ is_quasi_lambda_free t1 andalso is_quasi_lambda_free t2
+ | is_quasi_lambda_free (Abs _) = false
+ | is_quasi_lambda_free _ = true
val [f_B,g_B] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_B}));
val [g_C,f_C] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_C}));
@@ -225,7 +252,7 @@
(* Traverse a theorem, remplacing lambda-abstractions with combinators. *)
fun do_introduce_combinators ct =
- if lambda_free (term_of ct) then
+ if is_quasi_lambda_free (term_of ct) then
Thm.reflexive ct
else case term_of ct of
Abs _ =>
@@ -243,7 +270,7 @@
end
fun introduce_combinators th =
- if lambda_free (prop_of th) then
+ if is_quasi_lambda_free (prop_of th) then
th
else
let
@@ -273,19 +300,28 @@
(*Given the definition of a Skolem function, return a theorem to replace
an existential formula by a use of that function.
Example: "EX x. x : A & x ~: B ==> sko A B : A & sko A B ~: B" [.] *)
-fun skolem_of_def def =
- let val (c,rhs) = Thm.dest_equals (cprop_of (freeze_thm def))
- val (ch, frees) = c_variant_abs_multi (rhs, [])
- val (chilbert,cabs) = Thm.dest_comb ch
+fun skolem_theorem_of_def inline def =
+ let
+ val (c, rhs) = Thm.dest_equals (cprop_of (freeze_thm def))
+ val rhs' = rhs |> inline ? (snd o Thm.dest_comb)
+ val (ch, frees) = c_variant_abs_multi (rhs', [])
+ val (chilbert, cabs) = ch |> Thm.dest_comb
val thy = Thm.theory_of_cterm chilbert
val t = Thm.term_of chilbert
- val T = case t of
- Const (@{const_name Eps}, Type (@{type_name fun},[_,T])) => T
- | _ => raise THM ("skolem_of_def: expected Eps", 0, [def])
+ val T =
+ case t of
+ Const (@{const_name Eps}, Type (@{type_name fun}, [_, T])) => T
+ | _ => raise TERM ("skolem_theorem_of_def: expected \"Eps\"", [t])
val cex = Thm.cterm_of thy (HOLogic.exists_const T)
val ex_tm = c_mkTrueprop (Thm.capply cex cabs)
- and conc = c_mkTrueprop (Drule.beta_conv cabs (Drule.list_comb(c,frees)));
- fun tacf [prem] = rewrite_goals_tac [def] THEN rtac (prem RS @{thm someI_ex}) 1
+ and conc =
+ Drule.list_comb (if inline then rhs else c, frees)
+ |> Drule.beta_conv cabs |> c_mkTrueprop
+ fun tacf [prem] =
+ (if inline then rewrite_goals_tac @{thms skolem_id_def_raw}
+ else rewrite_goals_tac [def])
+ THEN rtac ((prem |> inline ? rewrite_rule @{thms skolem_id_def_raw})
+ RS @{thm someI_ex}) 1
in Goal.prove_internal [ex_tm] conc tacf
|> forall_intr_list frees
|> Thm.forall_elim_vars 0 (*Introduce Vars, but don't discharge defs.*)
@@ -303,9 +339,9 @@
in (th3, ctxt) end;
(*Generate Skolem functions for a theorem supplied in nnf*)
-fun assume_skolem_of_def s th =
- map (skolem_of_def o Thm.assume o cterm_of (theory_of_thm th))
- (assume_skolem_funs s th)
+fun skolem_theorems_of_assume inline s th =
+ map (skolem_theorem_of_def inline o Thm.assume o cterm_of (theory_of_thm th))
+ (assume_skolem_funs inline s th)
(*** Blacklisting (more in "Sledgehammer_Fact_Filter") ***)
@@ -364,7 +400,7 @@
else gensym "unknown_thm_";
(*Skolemize a named theorem, with Skolem functions as additional premises.*)
-fun skolem_thm (s, th) =
+fun skolemize_theorem s th =
if member (op =) multi_base_blacklist (Long_Name.base_name s) orelse
is_theorem_bad_for_atps th then
[]
@@ -372,7 +408,8 @@
let
val ctxt0 = Variable.global_thm_context th
val (nnfth, ctxt) = to_nnf th ctxt0
- val defs = assume_skolem_of_def s nnfth
+ val inline = exists_type (exists_subtype (can dest_TFree)) (prop_of nnfth)
+ val defs = skolem_theorems_of_assume inline s nnfth
val (cnfs, ctxt) = Meson.make_cnf defs nnfth ctxt
in
cnfs |> map introduce_combinators
@@ -402,8 +439,8 @@
fun cnf_axiom thy th0 =
let val th = Thm.transfer thy th0 in
case lookup_cache thy th of
- NONE => map Thm.close_derivation (skolem_thm (fake_name th, th))
- | SOME cls => cls
+ SOME cls => cls
+ | NONE => map Thm.close_derivation (skolemize_theorem (fake_name th) th)
end;
@@ -412,15 +449,17 @@
fun pair_name_cls k (n, []) = []
| pair_name_cls k (n, cls::clss) = (cls, (n,k)) :: pair_name_cls (k+1) (n, clss)
-fun cnf_rules_pairs_aux _ pairs [] = pairs
- | cnf_rules_pairs_aux thy pairs ((name,th)::ths) =
- let val pairs' = (pair_name_cls 0 (name, cnf_axiom thy th)) @ pairs
- handle THM _ => pairs |
- CLAUSE _ => pairs
- in cnf_rules_pairs_aux thy pairs' ths end;
-
(*The combination of rev and tail recursion preserves the original order*)
-fun cnf_rules_pairs thy l = cnf_rules_pairs_aux thy [] (rev l);
+fun cnf_rules_pairs thy =
+ let
+ fun aux pairs [] = pairs
+ | aux pairs ((name, th) :: ths) =
+ let
+ val new_pairs = pair_name_cls 0 (name, cnf_axiom thy th)
+ handle THM _ => [] |
+ CLAUSE _ => []
+ in aux (new_pairs @ pairs) ths end
+ in aux [] o rev end
(**** Convert all facts of the theory into FOL or HOL clauses ****)
@@ -438,7 +477,8 @@
fun skolem_cnfs (th, ctxt0, ctxt, nnfth, defs) =
let
- val (cnfs, ctxt) = Meson.make_cnf (map skolem_of_def defs) nnfth ctxt;
+ val (cnfs, ctxt) =
+ Meson.make_cnf (map (skolem_theorem_of_def false) defs) nnfth ctxt
val cnfs' = cnfs
|> map introduce_combinators
|> Variable.export ctxt ctxt0
@@ -455,14 +495,18 @@
if Facts.is_concealed facts name orelse already_seen thy name then I
else cons (name, ths));
val new_thms = (new_facts, []) |-> fold (fn (name, ths) =>
- if member (op =) multi_base_blacklist (Long_Name.base_name name) then I
- else fold_index (fn (i, th) =>
- if is_theorem_bad_for_atps th orelse is_some (lookup_cache thy th) then
- I
- else
- cons (name ^ "_" ^ string_of_int (i + 1), Thm.transfer thy th)) ths)
+ if member (op =) multi_base_blacklist (Long_Name.base_name name) then
+ I
+ else
+ fold_index (fn (i, th) =>
+ if is_theorem_bad_for_atps th orelse
+ is_some (lookup_cache thy th) then
+ I
+ else
+ cons (name ^ "_" ^ string_of_int (i + 1), Thm.transfer thy th)) ths)
in
- if null new_facts then NONE
+ if null new_facts then
+ NONE
else
let
val (defs, thy') = thy
@@ -481,9 +525,9 @@
if !use_skolem_cache then saturate_skolem_cache thy else NONE
-(*The cache can be kept smaller by inspecting the prop of each thm. Can ignore all that are
- lambda_free, but then the individual theory caches become much bigger.*)
-
+(* The cache can be kept smaller by inspecting the prop of each thm. Can ignore
+ all that are quasi-lambda-free, but then the individual theory caches become
+ much bigger. *)
fun strip_subgoal goal i =
let
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fol_clause.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fol_clause.ML Mon Jun 21 17:41:57 2010 +0200
@@ -111,7 +111,7 @@
fun union_all xss = fold (union (op =)) xss []
(* Readable names for the more common symbolic functions. Do not mess with the
- last six entries of the table unless you know what you are doing. *)
+ last nine entries of the table unless you know what you are doing. *)
val const_trans_table =
Symtab.make [(@{const_name "op ="}, "equal"),
(@{const_name "op &"}, "and"),
@@ -123,7 +123,10 @@
(@{const_name COMBK}, "COMBK"),
(@{const_name COMBB}, "COMBB"),
(@{const_name COMBC}, "COMBC"),
- (@{const_name COMBS}, "COMBS")]
+ (@{const_name COMBS}, "COMBS"),
+ (@{const_name True}, "True"),
+ (@{const_name False}, "False"),
+ (@{const_name If}, "If")]
val type_const_trans_table =
Symtab.make [(@{type_name "*"}, "prod"),
@@ -193,8 +196,7 @@
tvar_prefix ^ (ascii_of_indexname (trim_type_var x,i));
fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (trim_type_var x));
-val max_dfg_symbol_length =
- if is_new_spass_version then 1000000 (* arbitrary large number *) else 63
+val max_dfg_symbol_length = 63
(* HACK because SPASS 3.0 truncates identifiers to 63 characters. *)
fun controlled_length dfg s =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_hol_clause.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_hol_clause.ML Mon Jun 21 17:41:57 2010 +0200
@@ -29,14 +29,16 @@
val type_of_combterm : combterm -> fol_type
val strip_combterm_comb : combterm -> combterm * combterm list
val literals_of_term : theory -> term -> literal list * typ list
+ val conceal_skolem_somes :
+ int -> (string * term) list -> term -> (string * term) list * term
exception TRIVIAL
val make_conjecture_clauses : bool -> theory -> thm list -> hol_clause list
val make_axiom_clauses : bool -> theory ->
(thm * (axiom_name * hol_clause_id)) list -> (axiom_name * hol_clause) list
val type_wrapper_name : string
- val get_helper_clauses : bool -> theory -> bool ->
- hol_clause list * (thm * (axiom_name * hol_clause_id)) list * string list ->
- hol_clause list
+ val get_helper_clauses :
+ bool -> theory -> bool -> bool -> hol_clause list
+ -> (thm * (axiom_name * hol_clause_id)) list -> hol_clause list
val write_tptp_file : bool -> bool -> bool -> Path.T ->
hol_clause list * hol_clause list * hol_clause list * hol_clause list *
classrel_clause list * arity_clause list -> name_pool option * int
@@ -52,13 +54,13 @@
open Sledgehammer_FOL_Clause
open Sledgehammer_Fact_Preprocessor
-fun min_arity_of const_min_arity c = the_default 0 (Symtab.lookup const_min_arity c);
+fun min_arity_of const_min_arity c =
+ the_default 0 (Symtab.lookup const_min_arity c)
(*True if the constant ever appears outside of the top-level position in literals.
If false, the constant always receives all of its arguments and is used as a predicate.*)
fun needs_hBOOL explicit_apply const_needs_hBOOL c =
- explicit_apply orelse
- the_default false (Symtab.lookup const_needs_hBOOL c);
+ explicit_apply orelse the_default false (Symtab.lookup const_needs_hBOOL c)
(******************************************************)
@@ -115,28 +117,31 @@
TyVar (make_schematic_type_var x, string_of_indexname x);
-fun const_type_of dfg thy (c,t) =
- let val (tp,ts) = type_of dfg t
- in (tp, ts, map (simp_type_of dfg) (Sign.const_typargs thy (c,t))) end;
-
(* convert a Term.term (with combinators) into a combterm, also accummulate sort info *)
-fun combterm_of dfg thy (Const(c,t)) =
- let val (tp,ts,tvar_list) = const_type_of dfg thy (c,t)
- val c' = CombConst (`(make_fixed_const dfg) c, tp, tvar_list)
+fun combterm_of dfg thy (Const (c, T)) =
+ let
+ val (tp, ts) = type_of dfg T
+ val tvar_list =
+ (if String.isPrefix skolem_theory_name c then
+ [] |> Term.add_tvarsT T |> map TVar
+ else
+ (c, T) |> Sign.const_typargs thy)
+ |> map (simp_type_of dfg)
+ val c' = CombConst (`(make_fixed_const dfg) c, tp, tvar_list)
in (c',ts) end
- | combterm_of dfg _ (Free(v,t)) =
- let val (tp,ts) = type_of dfg t
+ | combterm_of dfg _ (Free(v, T)) =
+ let val (tp,ts) = type_of dfg T
val v' = CombConst (`make_fixed_var v, tp, [])
in (v',ts) end
- | combterm_of dfg _ (Var(v,t)) =
- let val (tp,ts) = type_of dfg t
+ | combterm_of dfg _ (Var(v, T)) =
+ let val (tp,ts) = type_of dfg T
val v' = CombVar ((make_schematic_var v, string_of_indexname v), tp)
in (v',ts) end
| combterm_of dfg thy (P $ Q) =
let val (P',tsP) = combterm_of dfg thy P
val (Q',tsQ) = combterm_of dfg thy Q
in (CombApp(P',Q'), union (op =) tsP tsQ) end
- | combterm_of _ _ (t as Abs _) = raise CLAUSE ("HOL CLAUSE", t);
+ | combterm_of _ _ (t as Abs _) = raise CLAUSE ("HOL clause", t);
fun predicate_of dfg thy ((@{const Not} $ P), polarity) = predicate_of dfg thy (P, not polarity)
| predicate_of dfg thy (t,polarity) = (combterm_of dfg thy (Envir.eta_contract t), polarity);
@@ -153,36 +158,78 @@
fun literals_of_term_dfg dfg thy P = literals_of_term1 dfg thy ([],[]) P;
val literals_of_term = literals_of_term_dfg false;
+fun skolem_name i j num_T_args =
+ skolem_prefix ^ (space_implode "_" (map Int.toString [i, j, num_T_args])) ^
+ skolem_infix ^ "g"
+
+fun conceal_skolem_somes i skolem_somes t =
+ if exists_Const (curry (op =) @{const_name skolem_id} o fst) t then
+ let
+ fun aux skolem_somes
+ (t as (Const (@{const_name skolem_id}, Type (_, [_, T])) $ _)) =
+ let
+ val (skolem_somes, s) =
+ if i = ~1 then
+ (skolem_somes, @{const_name undefined})
+ else case AList.find (op aconv) skolem_somes t of
+ s :: _ => (skolem_somes, s)
+ | [] =>
+ let
+ val s = skolem_theory_name ^ "." ^
+ skolem_name i (length skolem_somes)
+ (length (Term.add_tvarsT T []))
+ in ((s, t) :: skolem_somes, s) end
+ in (skolem_somes, Const (s, T)) end
+ | aux skolem_somes (t1 $ t2) =
+ let
+ val (skolem_somes, t1) = aux skolem_somes t1
+ val (skolem_somes, t2) = aux skolem_somes t2
+ in (skolem_somes, t1 $ t2) end
+ | aux skolem_somes (Abs (s, T, t')) =
+ let val (skolem_somes, t') = aux skolem_somes t' in
+ (skolem_somes, Abs (s, T, t'))
+ end
+ | aux skolem_somes t = (skolem_somes, t)
+ in aux skolem_somes t end
+ else
+ (skolem_somes, t)
+
(* Trivial problem, which resolution cannot handle (empty clause) *)
exception TRIVIAL;
(* making axiom and conjecture clauses *)
-fun make_clause dfg thy (clause_id, axiom_name, kind, th) =
- let val (lits,ctypes_sorts) = literals_of_term_dfg dfg thy (prop_of th)
- in
- if forall isFalse lits then
- raise TRIVIAL
- else
- HOLClause {clause_id = clause_id, axiom_name = axiom_name, th = th,
- kind = kind, literals = lits, ctypes_sorts = ctypes_sorts}
- end;
-
-
-fun add_axiom_clause dfg thy (th, (name, id)) =
- let val cls = make_clause dfg thy (id, name, Axiom, th) in
- not (isTaut cls) ? cons (name, cls)
+fun make_clause dfg thy (clause_id, axiom_name, kind, th) skolem_somes =
+ let
+ val (skolem_somes, t) =
+ th |> prop_of |> conceal_skolem_somes clause_id skolem_somes
+ val (lits, ctypes_sorts) = literals_of_term_dfg dfg thy t
+ in
+ if forall isFalse lits then
+ raise TRIVIAL
+ else
+ (skolem_somes,
+ HOLClause {clause_id = clause_id, axiom_name = axiom_name, th = th,
+ kind = kind, literals = lits, ctypes_sorts = ctypes_sorts})
end
+fun add_axiom_clause dfg thy (th, (name, id)) (skolem_somes, clss) =
+ let
+ val (skolem_somes, cls) =
+ make_clause dfg thy (id, name, Axiom, th) skolem_somes
+ in (skolem_somes, clss |> not (isTaut cls) ? cons (name, cls)) end
+
fun make_axiom_clauses dfg thy clauses =
- fold (add_axiom_clause dfg thy) clauses []
+ ([], []) |> fold (add_axiom_clause dfg thy) clauses |> snd
-fun make_conjecture_clauses_aux _ _ _ [] = []
- | make_conjecture_clauses_aux dfg thy n (th::ths) =
- make_clause dfg thy (n,"conjecture", Conjecture, th) ::
- make_conjecture_clauses_aux dfg thy (n+1) ths;
-
-fun make_conjecture_clauses dfg thy = make_conjecture_clauses_aux dfg thy 0;
-
+fun make_conjecture_clauses dfg thy =
+ let
+ fun aux _ _ [] = []
+ | aux n skolem_somes (th :: ths) =
+ let
+ val (skolem_somes, cls) =
+ make_clause dfg thy (n, "conjecture", Conjecture, th) skolem_somes
+ in cls :: aux (n + 1) skolem_somes ths end
+ in aux 0 [] end
(**********************************************************************)
(* convert clause into ATP specific formats: *)
@@ -376,10 +423,13 @@
fold (add_literal_decls params) literals decls
handle Symtab.DUP a => error ("function " ^ a ^ " has multiple arities")
-fun decls_of_clauses params clauses arity_clauses =
- let val functab =
- init_functab |> fold Symtab.update [(type_wrapper_name, 2), ("hAPP", 2),
- ("tc_bool", 0)]
+fun decls_of_clauses (params as (full_types, explicit_apply, _, _)) clauses
+ arity_clauses =
+ let
+ val functab =
+ init_functab
+ |> fold Symtab.update [(type_wrapper_name, 2), ("hAPP", 2),
+ ("tc_bool", 0)]
val predtab = Symtab.empty |> Symtab.update ("hBOOL", 1)
val (functab, predtab) =
(functab, predtab) |> fold (add_clause_decls params) clauses
@@ -402,10 +452,6 @@
(* write clauses to files *)
(**********************************************************************)
-val init_counters =
- Symtab.make [("c_COMBI", 0), ("c_COMBK", 0), ("c_COMBB", 0), ("c_COMBC", 0),
- ("c_COMBS", 0)];
-
fun count_combterm (CombConst ((c, _), _, _)) =
Symtab.map_entry c (Integer.add 1)
| count_combterm (CombVar _) = I
@@ -415,33 +461,37 @@
fun count_clause (HOLClause {literals, ...}) = fold count_literal literals
-fun count_user_clause user_lemmas (HOLClause {axiom_name, literals, ...}) =
- member (op =) user_lemmas axiom_name ? fold count_literal literals
+val raw_cnf_rules_pairs = map (fn (name, thm) => (thm, (name, 0)))
+fun cnf_helper_thms thy raw =
+ map (`Thm.get_name_hint)
+ #> (if raw then raw_cnf_rules_pairs else cnf_rules_pairs thy)
-fun cnf_helper_thms thy = cnf_rules_pairs thy o map (`Thm.get_name_hint)
+val optional_helpers =
+ [(["c_COMBI", "c_COMBK"], (false, @{thms COMBI_def COMBK_def})),
+ (["c_COMBB", "c_COMBC"], (false, @{thms COMBB_def COMBC_def})),
+ (["c_COMBS"], (false, @{thms COMBS_def}))]
+val optional_typed_helpers =
+ [(["c_True", "c_False"], (true, @{thms True_or_False})),
+ (["c_If"], (true, @{thms if_True if_False True_or_False}))]
+val mandatory_helpers = @{thms fequal_imp_equal equal_imp_fequal}
-fun get_helper_clauses dfg thy isFO (conjectures, axcls, user_lemmas) =
- if isFO then
- []
- else
- let
- val axclauses = map #2 (make_axiom_clauses dfg thy axcls)
- val ct = init_counters |> fold count_clause conjectures
- |> fold (count_user_clause user_lemmas) axclauses
- fun needed c = the (Symtab.lookup ct c) > 0
- val IK = if needed "c_COMBI" orelse needed "c_COMBK"
- then cnf_helper_thms thy [@{thm COMBI_def}, @{thm COMBK_def}]
- else []
- val BC = if needed "c_COMBB" orelse needed "c_COMBC"
- then cnf_helper_thms thy [@{thm COMBB_def}, @{thm COMBC_def}]
- else []
- val S = if needed "c_COMBS" then cnf_helper_thms thy [@{thm COMBS_def}]
- else []
- val other = cnf_helper_thms thy [@{thm fequal_imp_equal},
- @{thm equal_imp_fequal}]
- in
- map #2 (make_axiom_clauses dfg thy (other @ IK @ BC @ S))
- end;
+val init_counters =
+ Symtab.make (maps (maps (map (rpair 0) o fst))
+ [optional_helpers, optional_typed_helpers])
+
+fun get_helper_clauses dfg thy is_FO full_types conjectures axcls =
+ let
+ val axclauses = map snd (make_axiom_clauses dfg thy axcls)
+ val ct = fold (fold count_clause) [conjectures, axclauses] init_counters
+ fun is_needed c = the (Symtab.lookup ct c) > 0
+ val cnfs =
+ (optional_helpers
+ |> full_types ? append optional_typed_helpers
+ |> maps (fn (ss, (raw, ths)) =>
+ if exists is_needed ss then cnf_helper_thms thy raw ths
+ else []))
+ @ (if is_FO then [] else cnf_helper_thms thy false mandatory_helpers)
+ in map snd (make_axiom_clauses dfg thy cnfs) end
(*Find the minimal arity of each function mentioned in the term. Also, note which uses
are not at top level, to see if hBOOL is needed.*)
@@ -553,7 +603,8 @@
val tfree_preds = map dfg_tfree_predicate tfree_lits
val (helper_clauses_strs, pool) =
pool_map (apfst fst oo dfg_clause params) helper_clauses pool
- val (funcs, cl_preds) = decls_of_clauses params (helper_clauses @ conjectures @ axclauses) arity_clauses
+ val (funcs, cl_preds) =
+ decls_of_clauses params (helper_clauses @ conjectures @ axclauses) arity_clauses
val ty_preds = preds_of_clauses axclauses classrel_clauses arity_clauses
val preds = tfree_preds @ cl_preds @ ty_preds
val conjecture_offset =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML Mon Jun 21 17:41:57 2010 +0200
@@ -345,8 +345,8 @@
>> merge_relevance_overrides))
(add_to_relevance_override [])
val parse_sledgehammer_command =
- (Scan.optional Parse.short_ident runN -- parse_params -- parse_relevance_override
- -- Scan.option Parse.nat) #>> sledgehammer_trans
+ (Scan.optional Parse.short_ident runN -- parse_params
+ -- parse_relevance_override -- Scan.option Parse.nat) #>> sledgehammer_trans
val parse_sledgehammer_params_command =
parse_params #>> sledgehammer_params_trans
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML Mon Jun 21 17:41:57 2010 +0200
@@ -14,17 +14,17 @@
val invert_type_const: string -> string
val num_type_args: theory -> string -> int
val strip_prefix: string -> string -> string option
- val metis_line: int -> int -> string list -> string
+ val metis_line: bool -> int -> int -> string list -> string
val metis_proof_text:
- minimize_command * string * string vector * thm * int
+ bool * minimize_command * string * string vector * thm * int
-> string * string list
val isar_proof_text:
- name_pool option * bool * bool * int * Proof.context * int list list
- -> minimize_command * string * string vector * thm * int
+ name_pool option * bool * int * Proof.context * int list list
+ -> bool * minimize_command * string * string vector * thm * int
-> string * string list
val proof_text:
- bool -> name_pool option * bool * bool * int * Proof.context * int list list
- -> minimize_command * string * string vector * thm * int
+ bool -> name_pool option * bool * int * Proof.context * int list list
+ -> bool * minimize_command * string * string vector * thm * int
-> string * string list
end;
@@ -41,12 +41,6 @@
fun is_ident_char c = Char.isAlphaNum c orelse c = #"_"
fun is_head_digit s = Char.isDigit (String.sub (s, 0))
-(* Hack: Could return false positives (e.g., a user happens to declare a
- constant called "SomeTheory.sko_means_shoe_in_$wedish". *)
-val is_skolem_const_name =
- Long_Name.base_name
- #> String.isPrefix skolem_prefix andf String.isSubstring skolem_infix
-
val index_in_shape : int -> int list list -> int =
find_index o exists o curry (op =)
fun is_axiom_clause_number thm_names num = num <= Vector.length thm_names
@@ -263,9 +257,16 @@
fun invert_const c = c |> Symtab.lookup const_trans_table_inv |> the_default c
-(*The number of type arguments of a constant, zero if it's monomorphic*)
+(* The number of type arguments of a constant, zero if it's monomorphic. For
+ (instances of) Skolem pseudoconstants, this information is encoded in the
+ constant name. *)
fun num_type_args thy s =
- length (Sign.const_typargs thy (s, Sign.the_const_type thy s))
+ if String.isPrefix skolem_theory_name s then
+ s |> unprefix skolem_theory_name
+ |> space_explode skolem_infix |> hd
+ |> space_explode "_" |> List.last |> Int.fromString |> the
+ else
+ (s, Sign.the_const_type thy s) |> Sign.const_typargs thy |> length
fun fix_atp_variable_name s =
let
@@ -323,9 +324,12 @@
else
(* Extra args from "hAPP" come after any arguments
given directly to the constant. *)
- Sign.const_instance thy (c,
- map (type_from_node tfrees)
- (drop num_term_args us)))
+ if String.isPrefix skolem_theory_name c then
+ map fastype_of ts ---> HOLogic.typeT
+ else
+ Sign.const_instance thy (c,
+ map (type_from_node tfrees)
+ (drop num_term_args us)))
in list_comb (t, ts) end
| NONE => (* a free or schematic variable *)
let
@@ -580,16 +584,31 @@
| extract_num (num :: "0" :: "Inp" :: _) = Int.fromString num
| extract_num _ = NONE
in atp_proof |> split_lines |> map_filter (extract_num o tokens_of) end
-
-fun apply_command _ 1 = "by "
- | apply_command 1 _ = "apply "
- | apply_command i _ = "prefer " ^ string_of_int i ^ " apply "
-fun metis_command i n [] = apply_command i n ^ "metis"
- | metis_command i n ss =
- apply_command i n ^ "(metis " ^ space_implode " " ss ^ ")"
-fun metis_line i n xs =
+
+val indent_size = 2
+val no_label = ("", ~1)
+
+val raw_prefix = "X"
+val assum_prefix = "A"
+val fact_prefix = "F"
+
+fun string_for_label (s, num) = s ^ string_of_int num
+
+fun metis_using [] = ""
+ | metis_using ls =
+ "using " ^ space_implode " " (map string_for_label ls) ^ " "
+fun metis_apply _ 1 = "by "
+ | metis_apply 1 _ = "apply "
+ | metis_apply i _ = "prefer " ^ string_of_int i ^ " apply "
+fun metis_name full_types = if full_types then "metisFT" else "metis"
+fun metis_call full_types [] = metis_name full_types
+ | metis_call full_types ss =
+ "(" ^ metis_name full_types ^ " " ^ space_implode " " ss ^ ")"
+fun metis_command full_types i n (ls, ss) =
+ metis_using ls ^ metis_apply i n ^ metis_call full_types ss
+fun metis_line full_types i n ss =
"Try this command: " ^
- Markup.markup Markup.sendback (metis_command i n xs) ^ ".\n"
+ Markup.markup Markup.sendback (metis_command full_types i n ([], ss)) ^ ".\n"
fun minimize_line _ [] = ""
| minimize_line minimize_command facts =
case minimize_command facts of
@@ -600,7 +619,8 @@
val unprefix_chained = perhaps (try (unprefix chained_prefix))
-fun metis_proof_text (minimize_command, atp_proof, thm_names, goal, i) =
+fun metis_proof_text (full_types, minimize_command, atp_proof, thm_names, goal,
+ i) =
let
val raw_lemmas =
atp_proof |> extract_clause_numbers_in_atp_proof
@@ -613,8 +633,8 @@
val lemmas = other_lemmas @ chained_lemmas
val n = Logic.count_prems (prop_of goal)
in
- (metis_line i n other_lemmas ^ minimize_line minimize_command lemmas,
- lemmas)
+ (metis_line full_types i n other_lemmas ^
+ minimize_line minimize_command lemmas, lemmas)
end
(** Isar proof construction and manipulation **)
@@ -639,15 +659,6 @@
fun smart_case_split [] facts = ByMetis facts
| smart_case_split proofs facts = CaseSplit (proofs, facts)
-val indent_size = 2
-val no_label = ("", ~1)
-
-val raw_prefix = "X"
-val assum_prefix = "A"
-val fact_prefix = "F"
-
-fun string_for_label (s, num) = s ^ string_of_int num
-
fun add_fact_from_dep thm_names num =
if is_axiom_clause_number thm_names num then
apsnd (insert (op =) (Vector.sub (thm_names, num - 1)))
@@ -924,7 +935,7 @@
step :: aux subst depth nextp proof
in aux [] 0 (1, 1) end
-fun string_for_proof ctxt i n =
+fun string_for_proof ctxt full_types i n =
let
fun fix_print_mode f x =
setmp_CRITICAL show_no_free_types true
@@ -948,7 +959,7 @@
let
val ls = ls |> sort_distinct (prod_ord string_ord int_ord)
val ss = ss |> map unprefix_chained |> sort_distinct string_ord
- in metis_command 1 1 (map string_for_label ls @ ss) end
+ in metis_command full_types 1 1 (ls, ss) end
and do_step ind (Fix xs) =
do_indent ind ^ "fix " ^ space_implode " and " (map do_free xs) ^ "\n"
| do_step ind (Let (t1, t2)) =
@@ -981,17 +992,16 @@
do_indent 0 ^ (if n <> 1 then "next" else "qed") ^ "\n"
in do_proof end
-fun isar_proof_text (pool, debug, full_types, isar_shrink_factor, ctxt,
- conjecture_shape)
- (minimize_command, atp_proof, thm_names, goal, i) =
+fun isar_proof_text (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
+ (other_params as (full_types, _, atp_proof, thm_names, goal,
+ i)) =
let
val thy = ProofContext.theory_of ctxt
val (params, hyp_ts, concl_t) = strip_subgoal goal i
val frees = fold Term.add_frees (concl_t :: hyp_ts) []
val tfrees = fold Term.add_tfrees (concl_t :: hyp_ts) []
val n = Logic.count_prems (prop_of goal)
- val (one_line_proof, lemma_names) =
- metis_proof_text (minimize_command, atp_proof, thm_names, goal, i)
+ val (one_line_proof, lemma_names) = metis_proof_text other_params
fun isar_proof_for () =
case proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor
atp_proof conjecture_shape thm_names params
@@ -1001,7 +1011,7 @@
|> then_chain_proof
|> kill_useless_labels_in_proof
|> relabel_proof
- |> string_for_proof ctxt i n of
+ |> string_for_proof ctxt full_types i n of
"" => ""
| proof => "\nStructured proof:\n" ^ Markup.markup Markup.sendback proof
val isar_proof =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML Mon Jun 21 17:41:57 2010 +0200
@@ -6,7 +6,6 @@
signature SLEDGEHAMMER_UTIL =
sig
- val is_new_spass_version : bool
val pairf : ('a -> 'b) -> ('a -> 'c) -> 'a -> 'b * 'c
val plural_s : int -> string
val serial_commas : string -> string list -> string list
@@ -25,18 +24,6 @@
structure Sledgehammer_Util : SLEDGEHAMMER_UTIL =
struct
-val is_new_spass_version =
- case getenv "SPASS_VERSION" of
- "" => (case getenv "SPASS_HOME" of
- "" => false
- | s =>
- (* Hack: Preliminary versions of the SPASS 3.7 package don't set
- "SPASS_VERSION". *)
- String.isSubstring "/spass-3.7/" s)
- | s => (case s |> space_explode "." |> map Int.fromString of
- SOME m :: SOME n :: _ => m > 3 orelse (m = 3 andalso n >= 5)
- | _ => false)
-
fun pairf f g x = (f x, g x)
fun plural_s n = if n = 1 then "" else "s"
--- a/src/HOL/Tools/TFL/dcterm.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/TFL/dcterm.ML Mon Jun 21 17:41:57 2010 +0200
@@ -129,7 +129,7 @@
val dest_neg = dest_monop "not"
-val dest_pair = dest_binop "Pair";
+val dest_pair = dest_binop @{const_name Pair};
val dest_eq = dest_binop "op ="
val dest_imp = dest_binop "op -->"
val dest_conj = dest_binop "op &"
--- a/src/HOL/Tools/TFL/rules.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/TFL/rules.ML Mon Jun 21 17:41:57 2010 +0200
@@ -591,11 +591,11 @@
local fun dest_pair M = let val {fst,snd} = S.dest_pair M in (fst,snd) end
fun mk_fst tm =
- let val ty as Type("*", [fty,sty]) = type_of tm
- in Const ("fst", ty --> fty) $ tm end
+ let val ty as Type("Product_Type.*", [fty,sty]) = type_of tm
+ in Const ("Product_Type.fst", ty --> fty) $ tm end
fun mk_snd tm =
- let val ty as Type("*", [fty,sty]) = type_of tm
- in Const ("snd", ty --> sty) $ tm end
+ let val ty as Type("Product_Type.*", [fty,sty]) = type_of tm
+ in Const ("Product_Type.snd", ty --> sty) $ tm end
in
fun XFILL tych x vstruct =
let fun traverse p xocc L =
--- a/src/HOL/Tools/TFL/usyntax.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/TFL/usyntax.ML Mon Jun 21 17:41:57 2010 +0200
@@ -197,7 +197,7 @@
local
fun mk_uncurry (xt, yt, zt) =
Const(@{const_name split}, (xt --> yt --> zt) --> prod_ty xt yt --> zt)
-fun dest_pair(Const("Pair",_) $ M $ N) = {fst=M, snd=N}
+fun dest_pair(Const(@{const_name Pair},_) $ M $ N) = {fst=M, snd=N}
| dest_pair _ = raise USYN_ERR "dest_pair" "not a pair"
fun is_var (Var _) = true | is_var (Free _) = true | is_var _ = false
in
@@ -268,11 +268,11 @@
fun mk_pair{fst,snd} =
let val ty1 = type_of fst
val ty2 = type_of snd
- val c = Const("Pair",ty1 --> ty2 --> prod_ty ty1 ty2)
+ val c = Const(@{const_name Pair},ty1 --> ty2 --> prod_ty ty1 ty2)
in list_comb(c,[fst,snd])
end;
-fun dest_pair(Const("Pair",_) $ M $ N) = {fst=M, snd=N}
+fun dest_pair(Const(@{const_name Pair},_) $ M $ N) = {fst=M, snd=N}
| dest_pair _ = raise USYN_ERR "dest_pair" "not a pair";
@@ -372,7 +372,7 @@
(* Miscellaneous *)
fun mk_vstruct ty V =
- let fun follow_prod_type (Type("*",[ty1,ty2])) vs =
+ let fun follow_prod_type (Type(@{type_name "*"},[ty1,ty2])) vs =
let val (ltm,vs1) = follow_prod_type ty1 vs
val (rtm,vs2) = follow_prod_type ty2 vs1
in (mk_pair{fst=ltm, snd=rtm}, vs2) end
@@ -393,7 +393,7 @@
fun dest_relation tm =
if (type_of tm = HOLogic.boolT)
- then let val (Const("op :",_) $ (Const("Pair",_)$y$x) $ R) = tm
+ then let val (Const("op :",_) $ (Const(@{const_name Pair},_)$y$x) $ R) = tm
in (R,y,x)
end handle Bind => raise USYN_ERR "dest_relation" "unexpected term structure"
else raise USYN_ERR "dest_relation" "not a boolean term";
--- a/src/HOL/Tools/float_arith.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/float_arith.ML Mon Jun 21 17:41:57 2010 +0200
@@ -206,7 +206,7 @@
fun mk_float (a, b) = @{term "float"} $
HOLogic.mk_prod (pairself (HOLogic.mk_number HOLogic.intT) (a, b));
-fun dest_float (Const ("Float.float", _) $ (Const ("Pair", _) $ a $ b)) =
+fun dest_float (Const ("Float.float", _) $ (Const (@{const_name Pair}, _) $ a $ b)) =
pairself (snd o HOLogic.dest_number) (a, b)
| dest_float t = ((snd o HOLogic.dest_number) t, 0);
--- a/src/HOL/Tools/groebner.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/groebner.ML Mon Jun 21 17:41:57 2010 +0200
@@ -698,7 +698,7 @@
val holify_polynomial =
let fun holify_varpow (v,n) =
- if n = 1 then v else ring_mk_pow v (Numeral.mk_cnumber @{ctyp "nat"} n) (* FIXME *)
+ if n = 1 then v else ring_mk_pow v (Numeral.mk_cnumber @{ctyp nat} n) (* FIXME *)
fun holify_monomial vars (c,m) =
let val xps = map holify_varpow (filter (fn (_,n) => n <> 0) (vars ~~ m))
in end_itlist ring_mk_mul (mk_const c :: xps)
--- a/src/HOL/Tools/hologic.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/hologic.ML Mon Jun 21 17:41:57 2010 +0200
@@ -289,42 +289,42 @@
(* prod *)
-fun mk_prodT (T1, T2) = Type ("*", [T1, T2]);
+fun mk_prodT (T1, T2) = Type ("Product_Type.*", [T1, T2]);
-fun dest_prodT (Type ("*", [T1, T2])) = (T1, T2)
+fun dest_prodT (Type ("Product_Type.*", [T1, T2])) = (T1, T2)
| dest_prodT T = raise TYPE ("dest_prodT", [T], []);
-fun pair_const T1 T2 = Const ("Pair", [T1, T2] ---> mk_prodT (T1, T2));
+fun pair_const T1 T2 = Const ("Product_Type.Pair", [T1, T2] ---> mk_prodT (T1, T2));
fun mk_prod (t1, t2) =
let val T1 = fastype_of t1 and T2 = fastype_of t2 in
pair_const T1 T2 $ t1 $ t2
end;
-fun dest_prod (Const ("Pair", _) $ t1 $ t2) = (t1, t2)
+fun dest_prod (Const ("Product_Type.Pair", _) $ t1 $ t2) = (t1, t2)
| dest_prod t = raise TERM ("dest_prod", [t]);
fun mk_fst p =
let val pT = fastype_of p in
- Const ("fst", pT --> fst (dest_prodT pT)) $ p
+ Const ("Product_Type.fst", pT --> fst (dest_prodT pT)) $ p
end;
fun mk_snd p =
let val pT = fastype_of p in
- Const ("snd", pT --> snd (dest_prodT pT)) $ p
+ Const ("Product_Type.snd", pT --> snd (dest_prodT pT)) $ p
end;
fun split_const (A, B, C) =
- Const ("split", (A --> B --> C) --> mk_prodT (A, B) --> C);
+ Const ("Product_Type.split", (A --> B --> C) --> mk_prodT (A, B) --> C);
fun mk_split t =
(case Term.fastype_of t of
T as (Type ("fun", [A, Type ("fun", [B, C])])) =>
- Const ("split", T --> mk_prodT (A, B) --> C) $ t
+ Const ("Product_Type.split", T --> mk_prodT (A, B) --> C) $ t
| _ => raise TERM ("mk_split: bad body type", [t]));
(*Maps the type T1 * ... * Tn to [T1, ..., Tn], however nested*)
-fun flatten_tupleT (Type ("*", [T1, T2])) = flatten_tupleT T1 @ flatten_tupleT T2
+fun flatten_tupleT (Type ("Product_Type.*", [T1, T2])) = flatten_tupleT T1 @ flatten_tupleT T2
| flatten_tupleT T = [T];
@@ -334,14 +334,14 @@
| mk_tupleT Ts = foldr1 mk_prodT Ts;
fun strip_tupleT (Type ("Product_Type.unit", [])) = []
- | strip_tupleT (Type ("*", [T1, T2])) = T1 :: strip_tupleT T2
+ | strip_tupleT (Type ("Product_Type.*", [T1, T2])) = T1 :: strip_tupleT T2
| strip_tupleT T = [T];
fun mk_tuple [] = unit
| mk_tuple ts = foldr1 mk_prod ts;
fun strip_tuple (Const ("Product_Type.Unity", _)) = []
- | strip_tuple (Const ("Pair", _) $ t1 $ t2) = t1 :: strip_tuple t2
+ | strip_tuple (Const ("Product_Type.Pair", _) $ t1 $ t2) = t1 :: strip_tuple t2
| strip_tuple t = [t];
@@ -367,14 +367,14 @@
fun strip_ptupleT ps =
let
fun factors p T = if member (op =) ps p then (case T of
- Type ("*", [T1, T2]) =>
+ Type ("Product_Type.*", [T1, T2]) =>
factors (1::p) T1 @ factors (2::p) T2
| _ => ptuple_err "strip_ptupleT") else [T]
in factors [] end;
val flat_tupleT_paths =
let
- fun factors p (Type ("*", [T1, T2])) =
+ fun factors p (Type ("Product_Type.*", [T1, T2])) =
p :: factors (1::p) T1 @ factors (2::p) T2
| factors p _ = []
in factors [] end;
@@ -383,7 +383,7 @@
let
fun mk p T ts =
if member (op =) ps p then (case T of
- Type ("*", [T1, T2]) =>
+ Type ("Product_Type.*", [T1, T2]) =>
let
val (t, ts') = mk (1::p) T1 ts;
val (u, ts'') = mk (2::p) T2 ts'
@@ -395,14 +395,14 @@
fun strip_ptuple ps =
let
fun dest p t = if member (op =) ps p then (case t of
- Const ("Pair", _) $ t $ u =>
+ Const ("Product_Type.Pair", _) $ t $ u =>
dest (1::p) t @ dest (2::p) u
| _ => ptuple_err "strip_ptuple") else [t]
in dest [] end;
val flat_tuple_paths =
let
- fun factors p (Const ("Pair", _) $ t $ u) =
+ fun factors p (Const ("Product_Type.Pair", _) $ t $ u) =
p :: factors (1::p) t @ factors (2::p) u
| factors p _ = []
in factors [] end;
@@ -414,7 +414,7 @@
let
fun ap ((p, T) :: pTs) =
if member (op =) ps p then (case T of
- Type ("*", [T1, T2]) =>
+ Type ("Product_Type.*", [T1, T2]) =>
split_const (T1, T2, map snd pTs ---> T3) $
ap ((1::p, T1) :: (2::p, T2) :: pTs)
| _ => ptuple_err "mk_psplits")
@@ -427,7 +427,7 @@
val strip_psplits =
let
fun strip [] qs Ts t = (t, rev Ts, qs)
- | strip (p :: ps) qs Ts (Const ("split", _) $ t) =
+ | strip (p :: ps) qs Ts (Const ("Product_Type.split", _) $ t) =
strip ((1 :: p) :: (2 :: p) :: ps) (p :: qs) Ts t
| strip (p :: ps) qs Ts (Abs (s, T, t)) = strip ps qs (T :: Ts) t
| strip (p :: ps) qs Ts t = strip ps qs
@@ -438,16 +438,16 @@
(* nat *)
-val natT = Type ("nat", []);
+val natT = Type ("Nat.nat", []);
val zero = Const ("Groups.zero_class.zero", natT);
fun is_zero (Const ("Groups.zero_class.zero", _)) = true
| is_zero _ = false;
-fun mk_Suc t = Const ("Suc", natT --> natT) $ t;
+fun mk_Suc t = Const ("Nat.Suc", natT --> natT) $ t;
-fun dest_Suc (Const ("Suc", _) $ t) = t
+fun dest_Suc (Const ("Nat.Suc", _) $ t) = t
| dest_Suc t = raise TERM ("dest_Suc", [t]);
val Suc_zero = mk_Suc zero;
@@ -459,7 +459,7 @@
in if n < 0 then raise TERM ("mk_nat: negative number", []) else mk n end;
fun dest_nat (Const ("Groups.zero_class.zero", _)) = 0
- | dest_nat (Const ("Suc", _) $ t) = dest_nat t + 1
+ | dest_nat (Const ("Nat.Suc", _) $ t) = dest_nat t + 1
| dest_nat t = raise TERM ("dest_nat", [t]);
val class_size = "Nat.size";
--- a/src/HOL/Tools/inductive_codegen.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/inductive_codegen.ML Mon Jun 21 17:41:57 2010 +0200
@@ -14,7 +14,7 @@
val quickcheck_setup : theory -> theory
end;
-structure InductiveCodegen : INDUCTIVE_CODEGEN =
+structure Inductive_Codegen : INDUCTIVE_CODEGEN =
struct
open Codegen;
@@ -41,7 +41,7 @@
);
-fun warn thm = warning ("InductiveCodegen: Not a proper clause:\n" ^
+fun warn thm = warning ("Inductive_Codegen: Not a proper clause:\n" ^
Display.string_of_thm_without_context thm);
fun add_node x g = Graph.new_node (x, ()) g handle Graph.DUP _ => g;
@@ -324,7 +324,7 @@
| distinct_v t nvs = (t, nvs);
fun is_exhaustive (Var _) = true
- | is_exhaustive (Const ("Pair", _) $ t $ u) =
+ | is_exhaustive (Const (@{const_name Pair}, _) $ t $ u) =
is_exhaustive t andalso is_exhaustive u
| is_exhaustive _ = false;
@@ -569,7 +569,7 @@
and mk_ind_def thy defs gr dep names module modecs intrs nparms =
add_edge_acyclic (hd names, dep) gr handle
Graph.CYCLES (xs :: _) =>
- error ("InductiveCodegen: illegal cyclic dependencies:\n" ^ commas xs)
+ error ("Inductive_Codegen: illegal cyclic dependencies:\n" ^ commas xs)
| Graph.UNDEF _ =>
let
val _ $ u = Logic.strip_imp_concl (hd intrs);
@@ -825,7 +825,7 @@
val s = "structure TestTerm =\nstruct\n\n" ^
cat_lines (map snd code) ^
"\nopen Generated;\n\n" ^ string_of
- (Pretty.block [str "val () = InductiveCodegen.test_fn :=",
+ (Pretty.block [str "val () = Inductive_Codegen.test_fn :=",
Pretty.brk 1, str "(fn p =>", Pretty.brk 1,
str "case Seq.pull (testf p) of", Pretty.brk 1,
str "SOME ", mk_tuple [mk_tuple (map (str o fst) args'), str "_"],
--- a/src/HOL/Tools/inductive_set.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/inductive_set.ML Mon Jun 21 17:41:57 2010 +0200
@@ -401,7 +401,7 @@
else thm
in map preproc end;
-fun code_ind_att optmod = to_pred_att [] #> InductiveCodegen.add optmod NONE;
+fun code_ind_att optmod = to_pred_att [] #> Inductive_Codegen.add optmod NONE;
(**** definition of inductive sets ****)
--- a/src/HOL/Tools/lin_arith.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/lin_arith.ML Mon Jun 21 17:41:57 2010 +0200
@@ -461,7 +461,7 @@
(* ?P ((?n::nat) mod (number_of ?k)) =
((number_of ?k = 0 --> ?P ?n) & (~ (number_of ?k = 0) -->
(ALL i j. j < number_of ?k --> ?n = number_of ?k * i + j --> ?P j))) *)
- | (Const ("Divides.div_class.mod", Type ("fun", [Type ("nat", []), _])), [t1, t2]) =>
+ | (Const ("Divides.div_class.mod", Type ("fun", [@{typ nat}, _])), [t1, t2]) =>
let
val rev_terms = rev terms
val zero = Const (@{const_name Groups.zero}, split_type)
@@ -493,7 +493,7 @@
(* ?P ((?n::nat) div (number_of ?k)) =
((number_of ?k = 0 --> ?P 0) & (~ (number_of ?k = 0) -->
(ALL i j. j < number_of ?k --> ?n = number_of ?k * i + j --> ?P i))) *)
- | (Const ("Divides.div_class.div", Type ("fun", [Type ("nat", []), _])), [t1, t2]) =>
+ | (Const ("Divides.div_class.div", Type ("fun", [@{typ nat}, _])), [t1, t2]) =>
let
val rev_terms = rev terms
val zero = Const (@{const_name Groups.zero}, split_type)
--- a/src/HOL/Tools/meson.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/meson.ML Mon Jun 21 17:41:57 2010 +0200
@@ -9,7 +9,6 @@
sig
val trace: bool Unsynchronized.ref
val term_pair_of: indexname * (typ * 'a) -> term * 'a
- val first_order_resolve: thm -> thm -> thm
val flexflex_first_order: thm -> thm
val size_of_subgoals: thm -> int
val too_many_clauses: Proof.context option -> term -> bool
@@ -98,11 +97,13 @@
let val thy = theory_of_thm thA
val tmA = concl_of thA
val Const("==>",_) $ tmB $ _ = prop_of thB
- val (tyenv, tenv) = Pattern.first_order_match thy (tmB, tmA) (Vartab.empty, Vartab.empty)
+ val tenv =
+ Pattern.first_order_match thy (tmB, tmA)
+ (Vartab.empty, Vartab.empty) |> snd
val ct_pairs = map (pairself (cterm_of thy) o term_pair_of) (Vartab.dest tenv)
in thA RS (cterm_instantiate ct_pairs thB) end) () of
SOME th => th
- | NONE => raise THM ("first_order_resolve", 0, [thA, thB]));
+ | NONE => raise THM ("first_order_resolve", 0, [thA, thB]))
fun flexflex_first_order th =
case (tpairs_of th) of
@@ -315,15 +316,17 @@
val has_meta_conn =
exists_Const (member (op =) ["==", "==>", "=simp=>", "all", "prop"] o #1);
-fun apply_skolem_ths (th, rls) =
- let fun tryall [] = raise THM("apply_skolem_ths", 0, th::rls)
- | tryall (rl::rls) = (first_order_resolve th rl handle THM _ => tryall rls)
- in tryall rls end;
+fun apply_skolem_theorem (th, rls) =
+ let
+ fun tryall [] = raise THM ("apply_skolem_theorem", 0, th::rls)
+ | tryall (rl :: rls) =
+ first_order_resolve th rl handle THM _ => tryall rls
+ in tryall rls end
-(*Conjunctive normal form, adding clauses from th in front of ths (for foldr).
- Strips universal quantifiers and breaks up conjunctions.
- Eliminates existential quantifiers using skoths: Skolemization theorems.*)
-fun cnf skoths ctxt (th,ths) =
+(* Conjunctive normal form, adding clauses from th in front of ths (for foldr).
+ Strips universal quantifiers and breaks up conjunctions.
+ Eliminates existential quantifiers using Skolemization theorems. *)
+fun cnf skolem_ths ctxt (th, ths) =
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*)
@@ -337,7 +340,7 @@
in ctxtr := ctxt'; cnf_aux (th', ths) end
| Const ("Ex", _) =>
(*existential quantifier: Insert Skolem functions*)
- cnf_aux (apply_skolem_ths (th,skoths), ths)
+ cnf_aux (apply_skolem_theorem (th, skolem_ths), ths)
| Const ("op |", _) =>
(*Disjunction of P, Q: Create new goal of proving ?P | ?Q and solve it using
all combinations of converting P, Q to CNF.*)
@@ -353,7 +356,7 @@
else cnf_aux (th,ths)
in (cls, !ctxtr) end;
-fun make_cnf skoths th ctxt = cnf skoths ctxt (th, []);
+fun make_cnf skolem_ths th ctxt = cnf skolem_ths ctxt (th, []);
(*Generalization, removal of redundant equalities, removal of tautologies.*)
fun finish_cnf ths = filter (not o is_taut) (map refl_clause ths);
@@ -395,7 +398,7 @@
since this code expects to be called on a clause form.*)
val is_conn = member (op =)
["Trueprop", "op &", "op |", "op -->", "Not",
- "All", "Ex", "Ball", "Bex"];
+ "All", "Ex", @{const_name Ball}, @{const_name Bex}];
(*True if the term contains a function--not a logical connective--where the type
of any argument contains bool.*)
--- a/src/HOL/Tools/nat_numeral_simprocs.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/nat_numeral_simprocs.ML Mon Jun 21 17:41:57 2010 +0200
@@ -97,7 +97,7 @@
(*Split up a sum into the list of its constituent terms, on the way removing any
Sucs and counting them.*)
-fun dest_Suc_sum (Const ("Suc", _) $ t, (k,ts)) = dest_Suc_sum (t, (k+1,ts))
+fun dest_Suc_sum (Const (@{const_name Suc}, _) $ t, (k,ts)) = dest_Suc_sum (t, (k+1,ts))
| dest_Suc_sum (t, (k,ts)) =
let val (t1,t2) = dest_plus t
in dest_Suc_sum (t1, dest_Suc_sum (t2, (k,ts))) end
--- a/src/HOL/Tools/record.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/record.ML Mon Jun 21 17:41:57 2010 +0200
@@ -64,7 +64,7 @@
signature ISO_TUPLE_SUPPORT =
sig
- val add_iso_tuple_type: bstring * (string * sort) list ->
+ val add_iso_tuple_type: binding * (string * sort) list ->
typ * typ -> theory -> (term * term) * theory
val mk_cons_tuple: term * term -> term
val dest_cons_tuple: term -> term * term
@@ -80,7 +80,7 @@
val iso_tuple_intro = @{thm isomorphic_tuple_intro};
val iso_tuple_intros = Tactic.build_net @{thms isomorphic_tuple.intros};
-val tuple_iso_tuple = (@{const_name tuple_iso_tuple}, @{thm tuple_iso_tuple});
+val tuple_iso_tuple = (@{const_name Record.tuple_iso_tuple}, @{thm tuple_iso_tuple});
fun named_cterm_instantiate values thm =
let
@@ -101,21 +101,21 @@
fun merge data = Symtab.merge Thm.eq_thm_prop data; (* FIXME handle Symtab.DUP ?? *)
);
-fun do_typedef name repT alphas thy =
+fun do_typedef b repT alphas thy =
let
- fun get_thms thy name =
+ val (b_constr, b_proj) = (Binding.prefix_name "make_" b, Binding.prefix_name "dest_" b);
+ fun get_thms thy tyco =
let
- val [({Abs_name, abs_type = absT, ...}, {Rep_inject, Abs_inverse, ...})] =
- Typedef.get_info_global thy name;
- val rewrite_rule =
- MetaSimplifier.rewrite_rule [@{thm iso_tuple_UNIV_I}, @{thm iso_tuple_True_simp}];
+ val SOME { vs, constr, typ = repT, proj_inject, proj_constr, ... } =
+ Typecopy.get_info thy tyco;
+ val absT = Type (tyco, map TFree vs);
in
- (map rewrite_rule [Rep_inject, Abs_inverse], Const (Abs_name, repT --> absT), absT)
+ ((proj_inject, proj_constr), Const (constr, repT --> absT), absT)
end;
in
thy
- |> Typecopy.typecopy (Binding.name name, alphas) repT NONE
- |-> (fn (name, _) => `(fn thy => get_thms thy name))
+ |> Typecopy.typecopy (b, alphas) repT b_constr b_proj
+ |-> (fn (tyco, _) => `(fn thy => get_thms thy tyco))
end;
fun mk_cons_tuple (left, right) =
@@ -124,21 +124,21 @@
val prodT = HOLogic.mk_prodT (leftT, rightT);
val isomT = Type (@{type_name tuple_isomorphism}, [prodT, leftT, rightT]);
in
- Const (@{const_name iso_tuple_cons}, isomT --> leftT --> rightT --> prodT) $
+ Const (@{const_name Record.iso_tuple_cons}, isomT --> leftT --> rightT --> prodT) $
Const (fst tuple_iso_tuple, isomT) $ left $ right
end;
-fun dest_cons_tuple (Const (@{const_name iso_tuple_cons}, _) $ Const _ $ t $ u) = (t, u)
+fun dest_cons_tuple (Const (@{const_name Record.iso_tuple_cons}, _) $ Const _ $ t $ u) = (t, u)
| dest_cons_tuple t = raise TERM ("dest_cons_tuple", [t]);
-fun add_iso_tuple_type (name, alphas) (leftT, rightT) thy =
+fun add_iso_tuple_type (b, alphas) (leftT, rightT) thy =
let
val repT = HOLogic.mk_prodT (leftT, rightT);
- val (([rep_inject, abs_inverse], absC, absT), typ_thy) =
+ val (((rep_inject, abs_inverse), absC, absT), typ_thy) =
thy
- |> do_typedef name repT alphas
- ||> Sign.add_path name;
+ |> do_typedef b repT alphas
+ ||> Sign.add_path (Binding.name_of b); (*FIXME proper prefixing instead*)
(*construct a type and body for the isomorphism constant by
instantiating the theorem to which the definition will be applied*)
@@ -146,7 +146,7 @@
rep_inject RS named_cterm_instantiate [("abst", cterm_of typ_thy absC)] iso_tuple_intro;
val (_, body) = Logic.dest_equals (List.last (prems_of intro_inst));
val isomT = fastype_of body;
- val isom_binding = Binding.name (name ^ isoN);
+ val isom_binding = Binding.suffix_name isoN b;
val isom_name = Sign.full_name typ_thy isom_binding;
val isom = Const (isom_name, isomT);
@@ -157,7 +157,7 @@
[((Binding.conceal (Thm.def_binding isom_binding), Logic.mk_equals (isom, body)), [])];
val iso_tuple = isom_def RS (abs_inverse RS (rep_inject RS iso_tuple_intro));
- val cons = Const (@{const_name iso_tuple_cons}, isomT --> leftT --> rightT --> absT);
+ val cons = Const (@{const_name Record.iso_tuple_cons}, isomT --> leftT --> rightT --> absT);
val thm_thy =
cdef_thy
@@ -1654,7 +1654,7 @@
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
- (suffix suff 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;
@@ -1846,8 +1846,9 @@
fun record_definition (alphas, binding) parent (parents: parent_info list) raw_fields thy =
let
+ val prefix = Binding.name_of binding;
val name = Sign.full_name thy binding;
- val full = Sign.full_name_path thy (Binding.name_of binding); (* FIXME Binding.qualified (!?) *)
+ val full = Sign.full_name_path thy prefix;
val bfields = map (fn (x, T, _) => (x, T)) raw_fields;
val field_syntax = map #3 raw_fields;
@@ -1859,7 +1860,7 @@
val parent_fields_len = length parent_fields;
val parent_variants =
Name.variant_list [moreN, rN, rN ^ "'", wN] (map Long_Name.base_name parent_names);
- val parent_vars = ListPair.map Free (parent_variants, parent_types);
+ val parent_vars = map2 (curry Free) parent_variants parent_types;
val parent_len = length parents;
val fields = map (apfst full) bfields;
@@ -1871,7 +1872,7 @@
val variants =
Name.variant_list (moreN :: rN :: (rN ^ "'") :: wN :: parent_variants)
(map (Binding.name_of o fst) bfields);
- val vars = ListPair.map Free (variants, types);
+ val vars = map2 (curry Free) variants types;
val named_vars = names ~~ vars;
val idxms = 0 upto len;
@@ -2090,13 +2091,13 @@
map (fn (c, x as Free (_, T)) => mk_sel r_rec0 (c, T) === x) named_vars_more;
(*updates*)
- fun mk_upd_prop (i, (c, T)) =
+ fun mk_upd_prop i (c, T) =
let
val x' = Free (Name.variant all_variants (Long_Name.base_name c ^ "'"), T --> T);
val n = parent_fields_len + i;
val args' = nth_map n (K (x' $ nth all_vars_more n)) all_vars_more;
in mk_upd updateN c x' r_rec0 === mk_rec args' 0 end;
- val upd_conv_props = ListPair.map mk_upd_prop (idxms, fields_more);
+ val upd_conv_props = map2 mk_upd_prop idxms fields_more;
(*induct*)
val induct_scheme_prop =
--- a/src/HOL/Tools/refute.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/refute.ML Mon Jun 21 17:41:57 2010 +0200
@@ -517,7 +517,7 @@
end
(* ------------------------------------------------------------------------- *)
-(* get_def: looks up the definition of a constant, as created by "constdefs" *)
+(* get_def: looks up the definition of a constant *)
(* ------------------------------------------------------------------------- *)
(* theory -> string * Term.typ -> (string * Term.term) option *)
@@ -657,14 +657,14 @@
(* other optimizations *)
| Const (@{const_name Finite_Set.card}, _) => t
| Const (@{const_name Finite_Set.finite}, _) => t
- | Const (@{const_name Orderings.less}, Type ("fun", [Type ("nat", []),
- Type ("fun", [Type ("nat", []), Type ("bool", [])])])) => t
- | Const (@{const_name Groups.plus}, Type ("fun", [Type ("nat", []),
- Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => t
- | Const (@{const_name Groups.minus}, Type ("fun", [Type ("nat", []),
- Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => t
- | Const (@{const_name Groups.times}, Type ("fun", [Type ("nat", []),
- Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => t
+ | Const (@{const_name Orderings.less}, Type ("fun", [@{typ nat},
+ Type ("fun", [@{typ nat}, Type ("bool", [])])])) => t
+ | Const (@{const_name Groups.plus}, Type ("fun", [@{typ nat},
+ Type ("fun", [@{typ nat}, @{typ nat}])])) => t
+ | Const (@{const_name Groups.minus}, Type ("fun", [@{typ nat},
+ Type ("fun", [@{typ nat}, @{typ nat}])])) => t
+ | Const (@{const_name Groups.times}, Type ("fun", [@{typ nat},
+ Type ("fun", [@{typ nat}, @{typ nat}])])) => t
| Const (@{const_name List.append}, _) => t
(* UNSOUND
| Const (@{const_name lfp}, _) => t
@@ -711,7 +711,7 @@
(* Note: currently we use "inverse" functions to the definitional *)
(* mechanisms provided by Isabelle/HOL, e.g. for "axclass", *)
- (* "typedef", "constdefs". A more general approach could consider *)
+ (* "typedef", "definition". A more general approach could consider *)
(* *every* axiom of the theory and collect it if it has a constant/ *)
(* type/typeclass in common with the term 't'. *)
@@ -834,17 +834,17 @@
| Const (@{const_name Finite_Set.card}, T) => collect_type_axioms T axs
| Const (@{const_name Finite_Set.finite}, T) =>
collect_type_axioms T axs
- | Const (@{const_name Orderings.less}, T as Type ("fun", [Type ("nat", []),
- Type ("fun", [Type ("nat", []), Type ("bool", [])])])) =>
+ | Const (@{const_name Orderings.less}, T as Type ("fun", [@{typ nat},
+ Type ("fun", [@{typ nat}, Type ("bool", [])])])) =>
collect_type_axioms T axs
- | Const (@{const_name Groups.plus}, T as Type ("fun", [Type ("nat", []),
- Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
+ | Const (@{const_name Groups.plus}, T as Type ("fun", [@{typ nat},
+ Type ("fun", [@{typ nat}, @{typ nat}])])) =>
collect_type_axioms T axs
- | Const (@{const_name Groups.minus}, T as Type ("fun", [Type ("nat", []),
- Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
+ | Const (@{const_name Groups.minus}, T as Type ("fun", [@{typ nat},
+ Type ("fun", [@{typ nat}, @{typ nat}])])) =>
collect_type_axioms T axs
- | Const (@{const_name Groups.times}, T as Type ("fun", [Type ("nat", []),
- Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
+ | Const (@{const_name Groups.times}, T as Type ("fun", [@{typ nat},
+ Type ("fun", [@{typ nat}, @{typ nat}])])) =>
collect_type_axioms T axs
| Const (@{const_name List.append}, T) => collect_type_axioms T axs
(* UNSOUND
@@ -2654,7 +2654,7 @@
case t of
Const (@{const_name Finite_Set.card},
Type ("fun", [Type ("fun", [T, Type ("bool", [])]),
- Type ("nat", [])])) =>
+ @{typ nat}])) =>
let
(* interpretation -> int *)
fun number_of_elements (Node xs) =
@@ -2670,7 +2670,7 @@
| number_of_elements (Leaf _) =
raise REFUTE ("Finite_Set_card_interpreter",
"interpretation for set type is a leaf")
- val size_of_nat = size_of_type thy model (Type ("nat", []))
+ val size_of_nat = size_of_type thy model (@{typ nat})
(* takes an interpretation for a set and returns an interpretation *)
(* for a 'nat' denoting the set's cardinality *)
(* interpretation -> interpretation *)
@@ -2730,10 +2730,10 @@
fun Nat_less_interpreter thy model args t =
case t of
- Const (@{const_name Orderings.less}, Type ("fun", [Type ("nat", []),
- Type ("fun", [Type ("nat", []), Type ("bool", [])])])) =>
+ Const (@{const_name Orderings.less}, Type ("fun", [@{typ nat},
+ Type ("fun", [@{typ nat}, Type ("bool", [])])])) =>
let
- val size_of_nat = size_of_type thy model (Type ("nat", []))
+ val size_of_nat = size_of_type thy model (@{typ nat})
(* the 'n'-th nat is not less than the first 'n' nats, while it *)
(* is less than the remaining 'size_of_nat - n' nats *)
(* int -> interpretation *)
@@ -2753,10 +2753,10 @@
fun Nat_plus_interpreter thy model args t =
case t of
- Const (@{const_name Groups.plus}, Type ("fun", [Type ("nat", []),
- Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
+ Const (@{const_name Groups.plus}, Type ("fun", [@{typ nat},
+ Type ("fun", [@{typ nat}, @{typ nat}])])) =>
let
- val size_of_nat = size_of_type thy model (Type ("nat", []))
+ val size_of_nat = size_of_type thy model (@{typ nat})
(* int -> int -> interpretation *)
fun plus m n =
let
@@ -2784,10 +2784,10 @@
fun Nat_minus_interpreter thy model args t =
case t of
- Const (@{const_name Groups.minus}, Type ("fun", [Type ("nat", []),
- Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
+ Const (@{const_name Groups.minus}, Type ("fun", [@{typ nat},
+ Type ("fun", [@{typ nat}, @{typ nat}])])) =>
let
- val size_of_nat = size_of_type thy model (Type ("nat", []))
+ val size_of_nat = size_of_type thy model (@{typ nat})
(* int -> int -> interpretation *)
fun minus m n =
let
@@ -2812,10 +2812,10 @@
fun Nat_times_interpreter thy model args t =
case t of
- Const (@{const_name Groups.times}, Type ("fun", [Type ("nat", []),
- Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
+ Const (@{const_name Groups.times}, Type ("fun", [@{typ nat},
+ Type ("fun", [@{typ nat}, @{typ nat}])])) =>
let
- val size_of_nat = size_of_type thy model (Type ("nat", []))
+ val size_of_nat = size_of_type thy model (@{typ nat})
(* nat -> nat -> interpretation *)
fun mult m n =
let
@@ -3050,7 +3050,7 @@
fun Product_Type_fst_interpreter thy model args t =
case t of
- Const (@{const_name fst}, Type ("fun", [Type ("*", [T, U]), _])) =>
+ Const (@{const_name fst}, Type ("fun", [Type (@{type_name "*"}, [T, U]), _])) =>
let
val constants_T = make_constants thy model T
val size_U = size_of_type thy model U
@@ -3069,7 +3069,7 @@
fun Product_Type_snd_interpreter thy model args t =
case t of
- Const (@{const_name snd}, Type ("fun", [Type ("*", [T, U]), _])) =>
+ Const (@{const_name snd}, Type ("fun", [Type (@{type_name "*"}, [T, U]), _])) =>
let
val size_T = size_of_type thy model T
val constants_U = make_constants thy model U
@@ -3279,8 +3279,8 @@
add_interpreter "lfp" lfp_interpreter #>
add_interpreter "gfp" gfp_interpreter #>
*)
- add_interpreter "fst" Product_Type_fst_interpreter #>
- add_interpreter "snd" Product_Type_snd_interpreter #>
+ add_interpreter "Product_Type.fst" Product_Type_fst_interpreter #>
+ add_interpreter "Product_Type.snd" Product_Type_snd_interpreter #>
add_printer "stlc" stlc_printer #>
add_printer "IDT" IDT_printer;
--- a/src/HOL/Tools/typecopy.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Tools/typecopy.ML Mon Jun 21 17:41:57 2010 +0200
@@ -6,9 +6,9 @@
signature TYPECOPY =
sig
- type info = { vs: (string * sort) list, constr: string, typ: typ,
- inject: thm, proj: string * typ, proj_def: thm }
- val typecopy: binding * (string * sort) list -> typ -> (binding * binding) option
+ type info = { vs: (string * sort) list, typ: typ, constr: string, proj: string,
+ constr_inject: thm, proj_inject: thm, constr_proj: thm, proj_constr: thm }
+ val typecopy: binding * (string * sort) list -> typ -> binding -> binding
-> theory -> (string * info) * theory
val get_info: theory -> string -> info option
val interpretation: (string -> theory -> theory) -> theory -> theory
@@ -23,14 +23,16 @@
type info = {
vs: (string * sort) list,
+ typ: typ,
constr: string,
- typ: typ,
- inject: thm,
- proj: string * typ,
- proj_def: thm
+ proj: string,
+ constr_inject: thm,
+ proj_inject: thm,
+ constr_proj: thm,
+ proj_constr: thm
};
-structure TypecopyData = Theory_Data
+structure Typecopy_Data = Theory_Data
(
type T = info Symtab.table;
val empty = Symtab.empty;
@@ -38,7 +40,7 @@
fun merge data = Symtab.merge (K true) data;
);
-val get_info = Symtab.lookup o TypecopyData.get;
+val get_info = Symtab.lookup o Typecopy_Data.get;
(* interpretation of type copies *)
@@ -49,40 +51,42 @@
(* introducing typecopies *)
-fun typecopy (raw_tyco, raw_vs) raw_ty constr_proj thy =
+fun add_info tyco vs (({ rep_type, Abs_name, Rep_name, ...},
+ { Abs_inject, Rep_inject, Abs_inverse, Rep_inverse, ... }) : Typedef.info) thy =
+ let
+ val exists_thm =
+ UNIV_I
+ |> Drule.instantiate' [SOME (ctyp_of thy (Logic.varifyT_global rep_type))] [];
+ val constr_inject = Abs_inject OF [exists_thm, exists_thm];
+ val proj_constr = Abs_inverse OF [exists_thm];
+ val info = {
+ vs = vs,
+ typ = rep_type,
+ constr = Abs_name,
+ proj = Rep_name,
+ constr_inject = constr_inject,
+ proj_inject = Rep_inject,
+ constr_proj = Rep_inverse,
+ proj_constr = proj_constr
+ };
+ in
+ thy
+ |> (Typecopy_Data.map o Symtab.update_new) (tyco, info)
+ |> Typecopy_Interpretation.data tyco
+ |> pair (tyco, info)
+ end
+
+fun typecopy (raw_tyco, raw_vs) raw_ty constr proj thy =
let
val ty = Sign.certify_typ thy raw_ty;
val ctxt = ProofContext.init_global thy |> Variable.declare_typ ty;
val vs = map (ProofContext.check_tfree ctxt) raw_vs;
val tac = Tactic.rtac UNIV_witness 1;
- fun add_info tyco (({ abs_type = ty_abs, rep_type = ty_rep, Abs_name = c_abs,
- Rep_name = c_rep, ...}, { Abs_inject = inject, Abs_inverse = inverse, ... })
- : Typedef.info) thy =
- let
- val exists_thm =
- UNIV_I
- |> Drule.instantiate' [SOME (ctyp_of thy (Logic.varifyT_global ty_rep))] [];
- val inject' = inject OF [exists_thm, exists_thm];
- val proj_def = inverse OF [exists_thm];
- val info = {
- vs = vs,
- constr = c_abs,
- typ = ty_rep,
- inject = inject',
- proj = (c_rep, ty_abs --> ty_rep),
- proj_def = proj_def
- };
- in
- thy
- |> (TypecopyData.map o Symtab.update_new) (tyco, info)
- |> Typecopy_Interpretation.data tyco
- |> pair (tyco, info)
- end
in
thy
|> Typedef.add_typedef_global false (SOME raw_tyco) (raw_tyco, vs, NoSyn)
- (HOLogic.mk_UNIV ty) (Option.map swap constr_proj) tac
- |-> (fn (tyco, info) => add_info tyco info)
+ (HOLogic.mk_UNIV ty) (SOME (proj, constr)) tac
+ |-> (fn (tyco, info) => add_info tyco vs info)
end;
@@ -90,10 +94,8 @@
fun add_default_code tyco thy =
let
- val SOME { constr = c, proj = (proj, _), proj_def = proj_eq, vs = vs,
- typ = ty_rep, ... } = get_info thy tyco;
- (* FIXME handle multiple typedef interpretations (!??) *)
- val [(_, { Rep_inject = proj_inject, ... })] = Typedef.get_info_global thy tyco;
+ val SOME { vs, typ = ty_rep, constr = c, proj, proj_constr, proj_inject, ... } =
+ get_info thy tyco;
val constr = (c, Logic.unvarifyT_global (Sign.the_const_type thy c));
val ty = Type (tyco, map TFree vs);
val proj = Const (proj, ty --> ty_rep);
@@ -113,7 +115,7 @@
in
thy
|> Code.add_datatype [constr]
- |> Code.add_eqn proj_eq
+ |> Code.add_eqn proj_constr
|> Theory_Target.instantiation ([tyco], vs, [HOLogic.class_eq])
|> `(fn lthy => Syntax.check_term lthy eq)
|-> (fn eq => Specification.definition
--- a/src/HOL/Transitive_Closure.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Transitive_Closure.thy Mon Jun 21 17:41:57 2010 +0200
@@ -858,7 +858,7 @@
val rtrancl_trans = @{thm rtrancl_trans};
fun decomp (@{const Trueprop} $ t) =
- let fun dec (Const ("op :", _) $ (Const ("Pair", _) $ a $ b) $ rel ) =
+ let fun dec (Const ("op :", _) $ (Const (@{const_name Pair}, _) $ a $ b) $ rel ) =
let fun decr (Const ("Transitive_Closure.rtrancl", _ ) $ r) = (r,"r*")
| decr (Const ("Transitive_Closure.trancl", _ ) $ r) = (r,"r+")
| decr r = (r,"r");
--- a/src/HOL/Wellfounded.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/Wellfounded.thy Mon Jun 21 17:41:57 2010 +0200
@@ -683,10 +683,8 @@
text{* Lexicographic combinations *}
definition
- lex_prod :: "[('a*'a)set, ('b*'b)set] => (('a*'b)*('a*'b))set"
- (infixr "<*lex*>" 80)
-where
- "ra <*lex*> rb == {((a,b),(a',b')). (a,a') : ra | a=a' & (b,b') : rb}"
+ lex_prod :: "('a \<times>'a) set \<Rightarrow> ('b \<times> 'b) set \<Rightarrow> (('a \<times> 'b) \<times> ('a \<times> 'b)) set" (infixr "<*lex*>" 80) where
+ [code del]: "ra <*lex*> rb = {((a, b), (a', b')). (a, a') \<in> ra \<or> a = a' \<and> (b, b') \<in> rb}"
lemma wf_lex_prod [intro!]: "[| wf(ra); wf(rb) |] ==> wf(ra <*lex*> rb)"
apply (unfold wf_def lex_prod_def)
--- a/src/HOL/ex/Codegenerator_Pretty_Test.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/ex/Codegenerator_Pretty_Test.thy Mon Jun 21 17:41:57 2010 +0200
@@ -10,6 +10,6 @@
export_code * in SML module_name CodegenTest
in OCaml module_name CodegenTest file -
in Haskell file -
-(*in Scala file -*)
+ in Scala file -
end
--- a/src/HOL/ex/Dedekind_Real.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/ex/Dedekind_Real.thy Mon Jun 21 17:41:57 2010 +0200
@@ -2006,7 +2006,7 @@
qed
text {*
- There must be other proofs, e.g. @{text "Suc"} of the largest
+ There must be other proofs, e.g. @{text Suc} of the largest
integer in the cut representing @{text "x"}.
*}
--- a/src/HOL/ex/Meson_Test.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/ex/Meson_Test.thy Mon Jun 21 17:41:57 2010 +0200
@@ -16,7 +16,7 @@
below and constants declared in HOL!
*}
-hide_const (open) subset member quotient union inter
+hide_const (open) subset member quotient union inter sum
text {*
Test data for the MESON proof procedure
--- a/src/HOL/ex/Recdefs.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/ex/Recdefs.thy Mon Jun 21 17:41:57 2010 +0200
@@ -120,11 +120,13 @@
*}
consts mapf :: "nat => nat list"
-recdef (permissive) mapf "measure (\<lambda>m. m)"
+recdef mapf "measure (\<lambda>m. m)"
"mapf 0 = []"
"mapf (Suc n) = concat (map (\<lambda>x. mapf x) (replicate n n))"
(hints cong: map_cong)
+(* This used to be an example where recdef_tc etc is needed,
+ but now termination is proved outright
recdef_tc mapf_tc: mapf
apply (rule allI)
apply (case_tac "n = 0")
@@ -138,5 +140,6 @@
done
lemmas mapf_induct = mapf.induct [OF mapf_tc]
+*)
end
--- a/src/HOL/ex/Refute_Examples.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOL/ex/Refute_Examples.thy Mon Jun 21 17:41:57 2010 +0200
@@ -774,7 +774,7 @@
oops
lemma "P Suc"
- refute -- {* @{term "Suc"} is a partial function (regardless of the size
+ refute -- {* @{term Suc} is a partial function (regardless of the size
of the model), hence @{term "P Suc"} is undefined, hence no
model will be found *}
oops
--- a/src/HOLCF/IOA/Modelcheck/MuIOA.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOLCF/IOA/Modelcheck/MuIOA.thy Mon Jun 21 17:41:57 2010 +0200
@@ -34,9 +34,9 @@
exception malformed;
-fun fst_type (Type("*",[a,_])) = a |
+fun fst_type (Type(@{type_name "*"},[a,_])) = a |
fst_type _ = raise malformed;
-fun snd_type (Type("*",[_,a])) = a |
+fun snd_type (Type(@{type_name "*"},[_,a])) = a |
snd_type _ = raise malformed;
fun element_type (Type("set",[a])) = a |
@@ -121,10 +121,10 @@
fun delete_ul_string s = implode(delete_ul (explode s));
-fun type_list_of sign (Type("*",a::b::_)) = (type_list_of sign a) @ (type_list_of sign b) |
+fun type_list_of sign (Type(@{type_name "*"},a::b::_)) = (type_list_of sign a) @ (type_list_of sign b) |
type_list_of sign a = [(Syntax.string_of_typ_global sign a)];
-fun structured_tuple l (Type("*",s::t::_)) =
+fun structured_tuple l (Type(@{type_name "*"},s::t::_)) =
let
val (r,str) = structured_tuple l s;
in
--- a/src/HOLCF/IOA/meta_theory/automaton.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOLCF/IOA/meta_theory/automaton.ML Mon Jun 21 17:41:57 2010 +0200
@@ -298,7 +298,7 @@
(string_of_typ thy t));
fun comp_st_type_of thy [a] = st_type_of thy (aut_type_of thy a) |
-comp_st_type_of thy (a::r) = Type("*",[st_type_of thy (aut_type_of thy a), comp_st_type_of thy r]) |
+comp_st_type_of thy (a::r) = Type(@{type_name "*"},[st_type_of thy (aut_type_of thy a), comp_st_type_of thy r]) |
comp_st_type_of _ _ = error "empty automaton list";
(* checking consistency of action types (for composition) *)
@@ -387,11 +387,11 @@
thy
|> Sign.add_consts_i
[(Binding.name automaton_name,
-Type("*",
-[Type("*",[Type("set",[acttyp]),Type("*",[Type("set",[acttyp]),Type("set",[acttyp])])]),
- Type("*",[Type("set",[st_typ]),
- Type("*",[Type("set",[Type("*",[st_typ,Type("*",[acttyp,st_typ])])]),
- Type("*",[Type("set",[Type("set",[acttyp])]),Type("set",[Type("set",[acttyp])])])])])])
+Type(@{type_name "*"},
+[Type(@{type_name "*"},[Type("set",[acttyp]),Type(@{type_name "*"},[Type("set",[acttyp]),Type("set",[acttyp])])]),
+ Type(@{type_name "*"},[Type("set",[st_typ]),
+ Type(@{type_name "*"},[Type("set",[Type(@{type_name "*"},[st_typ,Type(@{type_name "*"},[acttyp,st_typ])])]),
+ Type(@{type_name "*"},[Type("set",[Type("set",[acttyp])]),Type("set",[Type("set",[acttyp])])])])])])
,NoSyn)]
|> add_defs
[(automaton_name ^ "_def",
@@ -442,11 +442,11 @@
thy
|> Sign.add_consts_i
[(Binding.name automaton_name,
-Type("*",
-[Type("*",[Type("set",[acttyp]),Type("*",[Type("set",[acttyp]),Type("set",[acttyp])])]),
- Type("*",[Type("set",[st_typ]),
- Type("*",[Type("set",[Type("*",[st_typ,Type("*",[acttyp,st_typ])])]),
- Type("*",[Type("set",[Type("set",[acttyp])]),Type("set",[Type("set",[acttyp])])])])])])
+Type(@{type_name "*"},
+[Type(@{type_name "*"},[Type("set",[acttyp]),Type(@{type_name "*"},[Type("set",[acttyp]),Type("set",[acttyp])])]),
+ Type(@{type_name "*"},[Type("set",[st_typ]),
+ Type(@{type_name "*"},[Type("set",[Type(@{type_name "*"},[st_typ,Type(@{type_name "*"},[acttyp,st_typ])])]),
+ Type(@{type_name "*"},[Type("set",[Type("set",[acttyp])]),Type("set",[Type("set",[acttyp])])])])])])
,NoSyn)]
|> add_defs
[(automaton_name ^ "_def",
--- a/src/HOLCF/Tools/Domain/domain_library.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOLCF/Tools/Domain/domain_library.ML Mon Jun 21 17:41:57 2010 +0200
@@ -215,7 +215,7 @@
| prj _ _ _ [] _ = raise Fail "Domain_Library.prj: empty list"
| prj f1 _ x (_::y::ys) 0 = f1 x y
| prj f1 f2 x (y:: ys) j = prj f1 f2 (f2 x y) ys (j-1);
-fun proj x = prj (fn S => K(%%:"fst" $S)) (fn S => K(%%:"snd" $S)) x;
+fun proj x = prj (fn S => K (%%: @{const_name fst} $ S)) (fn S => K (%%: @{const_name snd} $ S)) x;
fun lift tfn = Library.foldr (fn (x,t)=> (mk_trp(tfn x) ===> t));
val UU = %%: @{const_name UU};
--- a/src/HOLCF/Tools/Domain/domain_theorems.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/HOLCF/Tools/Domain/domain_theorems.ML Mon Jun 21 17:41:57 2010 +0200
@@ -179,7 +179,7 @@
then copy_of_dtyp map_tab
mk_take (dtyp_of arg) ` (%# arg)
else (%# arg);
- val lhs = (dc_take dname $ (%%:"Suc" $ %:"n"))`(con_app con args);
+ val lhs = (dc_take dname $ (%%: @{const_name Suc} $ %:"n")) ` (con_app con args);
val rhs = con_app2 con one_rhs args;
val goal = mk_trp (lhs === rhs);
val rules =
--- a/src/Pure/Isar/class_target.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/Pure/Isar/class_target.ML Mon Jun 21 17:41:57 2010 +0200
@@ -465,8 +465,8 @@
explode #> scan_valids #> implode
end;
-fun type_name "*" = "prod"
- | type_name "+" = "sum"
+fun type_name "Product_Type.*" = "prod"
+ | type_name "Sum_Type.+" = "sum"
| type_name s = sanitize_name (Long_Name.base_name s);
fun resort_terms pp algebra consts constraints ts =
--- a/src/Pure/Isar/code.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/Pure/Isar/code.ML Mon Jun 21 17:41:57 2010 +0200
@@ -59,16 +59,20 @@
val add_default_eqn: thm -> theory -> theory
val add_default_eqn_attribute: attribute
val add_default_eqn_attrib: Attrib.src
+ val add_nbe_default_eqn: thm -> theory -> theory
+ val add_nbe_default_eqn_attribute: attribute
+ val add_nbe_default_eqn_attrib: Attrib.src
val del_eqn: thm -> theory -> theory
val del_eqns: string -> theory -> theory
val add_case: thm -> theory -> theory
val add_undefined: string -> theory -> theory
- val get_type: theory -> string -> ((string * sort) list * (string * typ list) list)
+ val get_type: theory -> string -> ((string * sort) list * ((string * string list) * typ list) list)
val get_type_of_constr_or_abstr: theory -> string -> (string * bool) option
val is_constr: theory -> string -> bool
val is_abstr: theory -> string -> bool
val get_cert: theory -> ((thm * bool) list -> (thm * bool) list) -> string -> cert
val get_case_scheme: theory -> string -> (int * (int * string list)) option
+ val get_case_cong: theory -> string -> thm option
val undefineds: theory -> string list
val print_codesetup: theory -> unit
@@ -142,12 +146,12 @@
(* functions *)
-datatype fun_spec = Default of (thm * bool) list
+datatype fun_spec = Default of (thm * bool) list * (thm * bool) list lazy
| Eqns of (thm * bool) list
| Proj of term * string
| Abstr of thm * string;
-val empty_fun_spec = Default [];
+val empty_fun_spec = Default ([], Lazy.value []);
fun is_default (Default _) = true
| is_default _ = false;
@@ -165,7 +169,7 @@
(*with explicit history*),
types: ((serial * ((string * sort) list * typ_spec)) list) Symtab.table
(*with explicit history*),
- cases: (int * (int * string list)) Symtab.table * unit Symtab.table
+ cases: ((int * (int * string list)) * thm) Symtab.table * unit Symtab.table
};
fun make_spec (history_concluded, ((signatures, functions), (types, cases))) =
@@ -425,7 +429,13 @@
of SOME (vs, Abstractor spec) => (vs, spec)
| _ => error ("Not an abstract type: " ^ tyco);
-fun get_type thy = fst o get_type_spec thy;
+fun get_type thy tyco =
+ let
+ val ((vs, cos), _) = get_type_spec thy tyco;
+ fun args_of c tys = map (fst o dest_TFree)
+ (Sign.const_typargs thy (c, tys ---> Type (tyco, map TFree vs)));
+ fun add_typargs (c, tys) = ((c, args_of c tys), tys);
+ in (vs, map add_typargs cos) end;
fun get_type_of_constr_or_abstr thy c =
case (snd o strip_type o const_typ thy) c
@@ -869,10 +879,10 @@
fun retrieve_raw thy c =
Symtab.lookup ((the_functions o the_exec) thy) c
|> Option.map (snd o fst)
- |> the_default (Default [])
+ |> the_default empty_fun_spec
fun get_cert thy f c = case retrieve_raw thy c
- of Default eqns => eqns
+ of Default (_, eqns_lazy) => Lazy.force eqns_lazy
|> (map o apfst) (Thm.transfer thy)
|> f
|> (map o apfst) (AxClass.unoverload thy)
@@ -932,7 +942,8 @@
handle Bind => error "bad case certificate"
| TERM _ => error "bad case certificate";
-fun get_case_scheme thy = Symtab.lookup ((fst o the_cases o the_exec) thy);
+fun get_case_scheme thy = Option.map fst o Symtab.lookup ((fst o the_cases o the_exec) thy);
+fun get_case_cong thy = Option.map snd o Symtab.lookup ((fst o the_cases o the_exec) thy);
val undefineds = Symtab.keys o snd o the_cases o the_exec;
@@ -947,7 +958,7 @@
(Pretty.block o Pretty.fbreaks) (
Pretty.str (string_of_const thy const) :: map (Display.pretty_thm ctxt) thms
);
- fun pretty_function (const, Default eqns) = pretty_equations const (map fst eqns)
+ fun pretty_function (const, Default (_, eqns_lazy)) = pretty_equations const (map fst (Lazy.force eqns_lazy))
| pretty_function (const, Eqns eqns) = pretty_equations const (map fst eqns)
| pretty_function (const, Proj (proj, _)) = Pretty.block
[Pretty.str (string_of_const thy const), Pretty.fbrk, Syntax.pretty_term ctxt proj]
@@ -967,8 +978,8 @@
:: Pretty.str "of"
:: map (Pretty.quote o Syntax.pretty_typ_global thy) tys)) cos)
);
- fun pretty_case (const, (_, (_, []))) = Pretty.str (string_of_const thy const)
- | pretty_case (const, (_, (_, cos))) = (Pretty.block o Pretty.breaks) [
+ fun pretty_case (const, ((_, (_, [])), _)) = Pretty.str (string_of_const thy const)
+ | pretty_case (const, ((_, (_, cos)), _)) = (Pretty.block o Pretty.breaks) [
Pretty.str (string_of_const thy const), Pretty.str "with",
(Pretty.block o Pretty.commas o map (Pretty.str o string_of_const thy)) cos];
val functions = the_functions exec
@@ -1040,21 +1051,26 @@
let
val thm = Thm.close_derivation raw_thm;
val c = const_eqn thy thm;
- fun add_eqn' true (Default eqns) = Default (eqns @ [(thm, proper)])
- | add_eqn' _ (Eqns eqns) =
- let
- val args_of = snd o strip_comb o map_types Type.strip_sorts
- o fst o Logic.dest_equals o Thm.plain_prop_of;
- val args = args_of thm;
- val incr_idx = Logic.incr_indexes ([], Thm.maxidx_of thm + 1);
- fun matches_args args' = length args <= length args' andalso
- Pattern.matchess thy (args, (map incr_idx o take (length args)) args');
- fun drop (thm', proper') = if (proper orelse not proper')
- andalso matches_args (args_of thm') then
- (warning ("Code generator: dropping redundant code equation\n" ^
- Display.string_of_thm_global thy thm'); true)
- else false;
- in Eqns ((thm, proper) :: filter_out drop eqns) end
+ fun update_subsume thy (thm, proper) eqns =
+ let
+ val args_of = snd o strip_comb o map_types Type.strip_sorts
+ o fst o Logic.dest_equals o Thm.plain_prop_of;
+ val args = args_of thm;
+ val incr_idx = Logic.incr_indexes ([], Thm.maxidx_of thm + 1);
+ fun matches_args args' = length args <= length args' andalso
+ Pattern.matchess thy (args, (map incr_idx o take (length args)) args');
+ fun drop (thm', proper') = if (proper orelse not proper')
+ andalso matches_args (args_of thm') then
+ (warning ("Code generator: dropping subsumed code equation\n" ^
+ Display.string_of_thm_global thy thm'); true)
+ else false;
+ in (thm, proper) :: filter_out drop eqns end;
+ fun natural_order thy_ref eqns =
+ (eqns, Lazy.lazy (fn () => fold (update_subsume (Theory.deref thy_ref)) eqns []))
+ fun add_eqn' true (Default (eqns, _)) =
+ Default (natural_order (Theory.check_thy thy) ((thm, proper) :: eqns))
+ (*this restores the natural order and drops syntactic redundancies*)
+ | add_eqn' _ (Eqns eqns) = Eqns (update_subsume thy (thm, proper) eqns)
| add_eqn' false _ = Eqns [(thm, proper)];
in change_fun_spec false c (add_eqn' default) thy end;
@@ -1066,18 +1082,25 @@
of SOME eqn => gen_add_eqn false eqn thy
| NONE => thy;
+fun add_nbe_eqn thm thy =
+ gen_add_eqn false (mk_eqn thy (thm, false)) thy;
+
fun add_default_eqn thm thy =
case mk_eqn_liberal thy thm
of SOME eqn => gen_add_eqn true eqn thy
| NONE => thy;
-fun add_nbe_eqn thm thy =
- gen_add_eqn false (mk_eqn thy (thm, false)) thy;
-
val add_default_eqn_attribute = Thm.declaration_attribute
(fn thm => Context.mapping (add_default_eqn thm) I);
val add_default_eqn_attrib = Attrib.internal (K add_default_eqn_attribute);
+fun add_nbe_default_eqn thm thy =
+ gen_add_eqn true (mk_eqn thy (thm, false)) thy;
+
+val add_nbe_default_eqn_attribute = Thm.declaration_attribute
+ (fn thm => Context.mapping (add_nbe_default_eqn thm) I);
+val add_nbe_default_eqn_attrib = Attrib.internal (K add_nbe_default_eqn_attribute);
+
fun add_abs_eqn raw_thm thy =
let
val (abs_thm, tyco) = (apfst Thm.close_derivation o mk_abs_eqn thy) raw_thm;
@@ -1098,14 +1121,34 @@
(* cases *)
+fun case_cong thy case_const (num_args, (pos, constrs)) =
+ let
+ val ([x, y], ctxt) = Name.variants ["A", "A'"] Name.context;
+ val (zs, _) = Name.variants (replicate (num_args - 1) "") ctxt;
+ val (ws, vs) = chop pos zs;
+ val T = Logic.unvarifyT_global (Sign.the_const_type thy case_const);
+ val Ts = (fst o strip_type) T;
+ val T_cong = nth Ts pos;
+ fun mk_prem z = Free (z, T_cong);
+ fun mk_concl z = list_comb (Const (case_const, T), map2 (curry Free) (ws @ z :: vs) Ts);
+ val (prem, concl) = pairself Logic.mk_equals (pairself mk_prem (x, y), pairself mk_concl (x, y));
+ fun tac { prems, ... } = Simplifier.rewrite_goals_tac prems
+ THEN ALLGOALS (ProofContext.fact_tac [Drule.reflexive_thm]);
+ in Skip_Proof.prove_global thy (x :: y :: zs) [prem] concl tac end;
+
fun add_case thm thy =
let
- val (c, (k, case_pats)) = case_cert thm;
+ val (case_const, (k, case_pats)) = case_cert thm;
val _ = case filter_out (is_constr thy) case_pats
of [] => ()
| cs => error ("Non-constructor(s) in case certificate: " ^ commas (map quote cs));
- val entry = (1 + Int.max (1, length case_pats), (k, case_pats))
- in (map_exec_purge o map_cases o apfst) (Symtab.update (c, entry)) thy end;
+ val entry = (1 + Int.max (1, length case_pats), (k, case_pats));
+ in
+ thy
+ |> Theory.checkpoint
+ |> `(fn thy => case_cong thy case_const entry)
+ |-> (fn cong => (map_exec_purge o map_cases o apfst) (Symtab.update (case_const, (entry, cong))))
+ end;
fun add_undefined c thy =
(map_exec_purge o map_cases o apsnd) (Symtab.update (c, ())) thy;
@@ -1128,7 +1171,7 @@
then insert (op =) c else I)
((the_functions o the_exec) thy) (old_proj :: old_constrs);
fun drop_outdated_cases cases = fold Symtab.delete_safe
- (Symtab.fold (fn (c, (_, (_, cos))) =>
+ (Symtab.fold (fn (c, ((_, (_, cos)), _)) =>
if exists (member (op =) old_constrs) cos
then insert (op =) c else I) cases []) cases;
in
@@ -1145,7 +1188,7 @@
Interpretation(type T = string * serial val eq = eq_snd (op =) : T * T -> bool);
fun datatype_interpretation f = Datatype_Interpretation.interpretation
- (fn (tyco, _) => fn thy => f (tyco, get_type thy tyco) thy);
+ (fn (tyco, _) => fn thy => f (tyco, fst (get_type_spec thy tyco)) thy);
fun add_datatype proto_constrs thy =
let
--- a/src/Pure/Isar/locale.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/Pure/Isar/locale.ML Mon Jun 21 17:41:57 2010 +0200
@@ -73,6 +73,7 @@
val add_dependency: string -> string * morphism -> morphism -> theory -> theory
(* Diagnostic *)
+ val all_locales: theory -> string list
val print_locales: theory -> unit
val print_locale: theory -> bool -> xstring -> unit
val print_registrations: theory -> string -> unit
@@ -149,10 +150,6 @@
fun change_locale name =
Locales.map o apsnd o Symtab.map_entry name o map_locale o apsnd;
-fun print_locales thy =
- Pretty.strs ("locales:" :: map #1 (Name_Space.extern_table (Locales.get thy)))
- |> Pretty.writeln;
-
(*** Primitive operations ***)
@@ -343,20 +340,6 @@
activate_all name thy Element.init (Element.transfer_morphism o Context.theory_of)
([], Context.Proof (ProofContext.init_global thy)) |-> put_idents |> Context.proof_of;
-fun print_locale thy show_facts raw_name =
- let
- val name = intern thy raw_name;
- val ctxt = init name thy;
- fun cons_elem (elem as Notes _) = show_facts ? cons elem
- | cons_elem elem = cons elem;
- val elems =
- activate_all name thy cons_elem (K (Element.transfer_morphism thy)) ([], [])
- |> snd |> rev;
- in
- Pretty.big_list "locale elements:" (map (Pretty.chunks o Element.pretty_ctxt ctxt) elems)
- |> Pretty.writeln
- end;
-
(*** Registrations: interpretations in theories ***)
@@ -458,15 +441,6 @@
in roundup thy activate export dep (get_idents context, context) |-> put_idents end;
-(* Diagnostic *)
-
-fun print_registrations thy raw_name =
- (case these_registrations thy (intern thy raw_name) of
- [] => Pretty.str ("no interpretations")
- | regs => Pretty.big_list "interpretations:" (map (pretty_reg thy) (rev regs)))
- |> Pretty.writeln;
-
-
(* Add and extend registrations *)
fun amend_registration (name, morph) mixin export thy =
@@ -595,4 +569,33 @@
Method.setup (Binding.name "unfold_locales") (Scan.succeed (METHOD o try_intro_locales_tac true))
"back-chain all introduction rules of locales"));
+
+(*** diagnostic commands and interfaces ***)
+
+fun all_locales thy = map #1 (Name_Space.extern_table (Locales.get thy));
+
+fun print_locales thy =
+ Pretty.strs ("locales:" :: all_locales thy)
+ |> Pretty.writeln;
+
+fun print_locale thy show_facts raw_name =
+ let
+ val name = intern thy raw_name;
+ val ctxt = init name thy;
+ fun cons_elem (elem as Notes _) = show_facts ? cons elem
+ | cons_elem elem = cons elem;
+ val elems =
+ activate_all name thy cons_elem (K (Element.transfer_morphism thy)) ([], [])
+ |> snd |> rev;
+ in
+ Pretty.big_list "locale elements:" (map (Pretty.chunks o Element.pretty_ctxt ctxt) elems)
+ |> Pretty.writeln
+ end;
+
+fun print_registrations thy raw_name =
+ (case these_registrations thy (intern thy raw_name) of
+ [] => Pretty.str ("no interpretations")
+ | regs => Pretty.big_list "interpretations:" (map (pretty_reg thy) (rev regs)))
+ |> Pretty.writeln;
+
end;
--- a/src/Pure/simplifier.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/Pure/simplifier.ML Mon Jun 21 17:41:57 2010 +0200
@@ -36,7 +36,7 @@
val inherit_context: simpset -> simpset -> simpset
val the_context: simpset -> Proof.context
val context: Proof.context -> simpset -> simpset
- val global_context: theory -> simpset -> simpset
+ val global_context: theory -> simpset -> simpset
val with_context: Proof.context -> (simpset -> simpset) -> simpset -> simpset
val simproc_i: theory -> string -> term list
-> (theory -> simpset -> term -> thm option) -> simproc
--- a/src/Tools/Code/code_eval.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/Tools/Code/code_eval.ML Mon Jun 21 17:41:57 2010 +0200
@@ -55,7 +55,7 @@
val value_name = "Value.VALUE.value"
val program' = program
|> Graph.new_node (value_name,
- Code_Thingol.Fun (Term.dummy_patternN, (([], ty), [(([], t), (NONE, true))])))
+ Code_Thingol.Fun (Term.dummy_patternN, ((([], ty), [(([], t), (NONE, true))]), NONE)))
|> fold (curry Graph.add_edge value_name) deps;
val (value_code, [SOME value_name']) = Code_ML.evaluation_code_of thy
(the_default target some_target) "" naming program' [value_name];
@@ -122,7 +122,7 @@
fun check_datatype thy tyco consts =
let
- val constrs = (map fst o snd o Code.get_type thy) tyco;
+ val constrs = (map (fst o fst) o snd o Code.get_type thy) tyco;
val missing_constrs = subtract (op =) consts constrs;
val _ = if null missing_constrs then []
else error ("Missing constructor(s) " ^ commas (map quote missing_constrs)
--- a/src/Tools/Code/code_haskell.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/Tools/Code/code_haskell.ML Mon Jun 21 17:41:57 2010 +0200
@@ -32,9 +32,9 @@
| SOME class => class;
fun print_typcontext tyvars vs = case maps (fn (v, sort) => map (pair v) sort) vs
of [] => []
- | classbinds => enum "," "(" ")" (
+ | constraints => enum "," "(" ")" (
map (fn (v, class) =>
- str (class_name class ^ " " ^ lookup_var tyvars v)) classbinds)
+ str (class_name class ^ " " ^ lookup_var tyvars v)) constraints)
@@ str " => ";
fun print_typforall tyvars vs = case map fst vs
of [] => []
@@ -75,7 +75,7 @@
then print_case tyvars some_thm vars fxy cases
else print_app tyvars some_thm vars fxy c_ts
| NONE => print_case tyvars some_thm vars fxy cases)
- and print_app_expr tyvars some_thm vars ((c, (_, tys)), ts) = case contr_classparam_typs c
+ and print_app_expr tyvars some_thm vars ((c, (_, function_typs)), ts) = case contr_classparam_typs c
of [] => (str o deresolve) c :: map (print_term tyvars some_thm vars BR) ts
| fingerprint => let
val ts_fingerprint = ts ~~ take (length ts) fingerprint;
@@ -86,7 +86,7 @@
brackets [print_term tyvars some_thm vars NOBR t, str "::", print_typ tyvars NOBR ty];
in
if needs_annotation then
- (str o deresolve) c :: map2 print_term_anno ts_fingerprint (take (length ts) tys)
+ (str o deresolve) c :: map2 print_term_anno ts_fingerprint (take (length ts) function_typs)
else (str o deresolve) c :: map (print_term tyvars some_thm vars BR) ts
end
and print_app tyvars = gen_print_app (print_app_expr tyvars) (print_term tyvars) syntax_const
@@ -115,7 +115,7 @@
end
| print_case tyvars some_thm vars fxy ((_, []), _) =
(brackify fxy o Pretty.breaks o map str) ["error", "\"empty case\""];
- fun print_stmt (name, Code_Thingol.Fun (_, ((vs, ty), raw_eqs))) =
+ fun print_stmt (name, Code_Thingol.Fun (_, (((vs, ty), raw_eqs), _))) =
let
val tyvars = intro_vars (map fst vs) reserved;
fun print_err n =
@@ -163,7 +163,7 @@
print_typdecl tyvars (vs, (deresolve_base name, map (ITyVar o fst) vs))
]
end
- | print_stmt (name, Code_Thingol.Datatype (_, (vs, [(co, [ty])]))) =
+ | print_stmt (name, Code_Thingol.Datatype (_, (vs, [((co, _), [ty])]))) =
let
val tyvars = intro_vars (map fst vs) reserved;
in
@@ -179,7 +179,7 @@
| print_stmt (name, Code_Thingol.Datatype (_, (vs, co :: cos))) =
let
val tyvars = intro_vars (map fst vs) reserved;
- fun print_co (co, tys) =
+ fun print_co ((co, _), tys) =
concat (
(str o deresolve_base) co
:: map (print_typ tyvars BR) tys
@@ -194,7 +194,7 @@
@ (if deriving_show name then [str "deriving (Read, Show)"] else [])
)
end
- | print_stmt (name, Code_Thingol.Class (_, (v, (superclasses, classparams)))) =
+ | print_stmt (name, Code_Thingol.Class (_, (v, (super_classes, classparams)))) =
let
val tyvars = intro_vars [v] reserved;
fun print_classparam (classparam, ty) =
@@ -207,32 +207,31 @@
Pretty.block_enclose (
Pretty.block [
str "class ",
- Pretty.block (print_typcontext tyvars [(v, map fst superclasses)]),
+ Pretty.block (print_typcontext tyvars [(v, map fst super_classes)]),
str (deresolve_base name ^ " " ^ lookup_var tyvars v),
str " where {"
],
str "};"
) (map print_classparam classparams)
end
- | print_stmt (_, Code_Thingol.Classinst ((class, (tyco, vs)), (_, classparam_insts))) =
+ | print_stmt (_, Code_Thingol.Classinst ((class, (tyco, vs)), (_, (classparam_instances, _)))) =
let
val tyvars = intro_vars (map fst vs) reserved;
- fun print_instdef ((classparam, c_inst), (thm, _)) = case syntax_const classparam
+ fun print_classparam_instance ((classparam, const), (thm, _)) = case syntax_const classparam
of NONE => semicolon [
(str o deresolve_base) classparam,
str "=",
- print_app tyvars (SOME thm) reserved NOBR (c_inst, [])
+ print_app tyvars (SOME thm) reserved NOBR (const, [])
]
| SOME (k, pr) =>
let
- val (c_inst_name, (_, tys)) = c_inst;
- val const = if (is_some o syntax_const) c_inst_name
- then NONE else (SOME o Long_Name.base_name o deresolve) c_inst_name;
- val proto_rhs = Code_Thingol.eta_expand k (c_inst, []);
- val (vs, rhs) = (apfst o map) fst (Code_Thingol.unfold_abs proto_rhs);
+ val (c, (_, tys)) = const;
+ val (vs, rhs) = (apfst o map) fst
+ (Code_Thingol.unfold_abs (Code_Thingol.eta_expand k (const, [])));
+ val s = if (is_some o syntax_const) c
+ then NONE else (SOME o Long_Name.base_name o deresolve) c;
val vars = reserved
- |> intro_vars (the_list const)
- |> intro_vars (map_filter I vs);
+ |> intro_vars (map_filter I (s :: vs));
val lhs = IConst (classparam, (([], []), tys)) `$$ map IVar vs;
(*dictionaries are not relevant at this late stage*)
in
@@ -252,7 +251,7 @@
str " where {"
],
str "};"
- ) (map print_instdef classparam_insts)
+ ) (map print_classparam_instance classparam_instances)
end;
in print_stmt end;
@@ -276,7 +275,8 @@
val (base', nsp_typ') = mk_name_stmt (first_upper base) nsp_typ
in (base', (nsp_fun, nsp_typ')) end;
val add_name = case stmt
- of Code_Thingol.Fun _ => add_fun false
+ of Code_Thingol.Fun (_, (_, SOME _)) => pair base
+ | Code_Thingol.Fun _ => add_fun false
| Code_Thingol.Datatype _ => add_typ
| Code_Thingol.Datatypecons _ => add_fun true
| Code_Thingol.Class _ => add_typ
@@ -284,7 +284,9 @@
| Code_Thingol.Classparam _ => add_fun false
| Code_Thingol.Classinst _ => pair base;
fun add_stmt' base' = case stmt
- of Code_Thingol.Datatypecons _ =>
+ of Code_Thingol.Fun (_, (_, SOME _)) =>
+ I
+ | Code_Thingol.Datatypecons _ =>
cons (name, (Long_Name.append module_name' base', NONE))
| Code_Thingol.Classrel _ => I
| Code_Thingol.Classparam _ =>
--- a/src/Tools/Code/code_ml.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/Tools/Code/code_ml.ML Mon Jun 21 17:41:57 2010 +0200
@@ -33,13 +33,13 @@
ML_Function of string * (typscheme * ((iterm list * iterm) * (thm option * bool)) list)
| ML_Instance of string * ((class * (string * (vname * sort) list))
* ((class * (string * (string * dict list list))) list
- * ((string * const) * (thm * bool)) list));
+ * (((string * const) * (thm * bool)) list * ((string * const) * (thm * bool)) list)));
datatype ml_stmt =
ML_Exc of string * (typscheme * int)
| ML_Val of ml_binding
| ML_Funs of ml_binding list * string list
- | ML_Datas of (string * ((vname * sort) list * (string * itype list) list)) list
+ | ML_Datas of (string * ((vname * sort) list * ((string * vname list) * itype list) list)) list
| ML_Class of string * (vname * ((class * string) list * (string * itype) list));
fun stmt_name_of_binding (ML_Function (name, _)) = name
@@ -121,9 +121,9 @@
then print_case is_pseudo_fun some_thm vars fxy cases
else print_app is_pseudo_fun some_thm vars fxy c_ts
| NONE => print_case is_pseudo_fun some_thm vars fxy cases)
- and print_app_expr is_pseudo_fun some_thm vars (app as ((c, ((_, iss), tys)), ts)) =
+ and print_app_expr is_pseudo_fun some_thm vars (app as ((c, ((_, iss), function_typs)), ts)) =
if is_cons c then
- let val k = length tys in
+ let val k = length function_typs in
if k < 2 orelse length ts = k
then (str o deresolve) c
:: the_list (print_tuple (print_term is_pseudo_fun some_thm vars) BR ts)
@@ -174,8 +174,8 @@
[str "val", str (deresolve name), str ":", print_typscheme typscheme];
fun print_datatype_decl definer (tyco, (vs, cos)) =
let
- fun print_co (co, []) = str (deresolve co)
- | print_co (co, tys) = concat [str (deresolve co), str "of",
+ fun print_co ((co, _), []) = str (deresolve co)
+ | print_co ((co, _), tys) = concat [str (deresolve co), str "of",
enum " *" "" "" (map (print_typ (INFX (2, X))) tys)];
in
concat (
@@ -219,19 +219,19 @@
))
end
| print_def is_pseudo_fun _ definer
- (ML_Instance (inst, ((class, (tyco, vs)), (superinsts, classparams)))) =
+ (ML_Instance (inst, ((class, (tyco, vs)), (super_instances, (classparam_instances, _))))) =
let
- fun print_superinst (_, (classrel, dss)) =
+ fun print_super_instance (_, (classrel, dss)) =
concat [
(str o Long_Name.base_name o deresolve) classrel,
str "=",
print_dict is_pseudo_fun NOBR (DictConst dss)
];
- fun print_classparam ((classparam, c_inst), (thm, _)) =
+ fun print_classparam_instance ((classparam, const), (thm, _)) =
concat [
(str o Long_Name.base_name o deresolve) classparam,
str "=",
- print_app (K false) (SOME thm) reserved NOBR (c_inst, [])
+ print_app (K false) (SOME thm) reserved NOBR (const, [])
];
in pair
(print_val_decl print_dicttypscheme
@@ -243,7 +243,8 @@
else print_dict_args vs)
@ str "="
:: enum "," "{" "}"
- (map print_superinst superinsts @ map print_classparam classparams)
+ (map print_super_instance super_instances
+ @ map print_classparam_instance classparam_instances)
:: str ":"
@@ print_tyco_expr NOBR (class, [tyco `%% map (ITyVar o fst) vs])
))
@@ -300,19 +301,19 @@
sig_ps
(Pretty.chunks (ps @| semicolon [p]))
end
- | print_stmt (ML_Class (class, (v, (superclasses, classparams)))) =
+ | print_stmt (ML_Class (class, (v, (super_classes, classparams)))) =
let
fun print_field s p = concat [str s, str ":", p];
fun print_proj s p = semicolon
(map str ["val", s, "=", "#" ^ s, ":"] @| p);
- fun print_superclass_decl (superclass, classrel) =
+ fun print_super_class_decl (super_class, classrel) =
print_val_decl print_dicttypscheme
- (classrel, ([(v, [class])], (superclass, ITyVar v)));
- fun print_superclass_field (superclass, classrel) =
- print_field (deresolve classrel) (print_dicttyp (superclass, ITyVar v));
- fun print_superclass_proj (superclass, classrel) =
+ (classrel, ([(v, [class])], (super_class, ITyVar v)));
+ fun print_super_class_field (super_class, classrel) =
+ print_field (deresolve classrel) (print_dicttyp (super_class, ITyVar v));
+ fun print_super_class_proj (super_class, classrel) =
print_proj (deresolve classrel)
- (print_dicttypscheme ([(v, [class])], (superclass, ITyVar v)));
+ (print_dicttypscheme ([(v, [class])], (super_class, ITyVar v)));
fun print_classparam_decl (classparam, ty) =
print_val_decl print_typscheme
(classparam, ([(v, [class])], ty));
@@ -323,7 +324,7 @@
(print_typscheme ([(v, [class])], ty));
in pair
(concat [str "type", print_dicttyp (class, ITyVar v)]
- :: map print_superclass_decl superclasses
+ :: map print_super_class_decl super_classes
@ map print_classparam_decl classparams)
(Pretty.chunks (
concat [
@@ -331,11 +332,11 @@
(str o deresolve) class,
str "=",
enum "," "{" "};" (
- map print_superclass_field superclasses
+ map print_super_class_field super_classes
@ map print_classparam_field classparams
)
]
- :: map print_superclass_proj superclasses
+ :: map print_super_class_proj super_classes
@ map print_classparam_proj classparams
))
end;
@@ -465,8 +466,8 @@
[str "val", str (deresolve name), str ":", print_typscheme typscheme];
fun print_datatype_decl definer (tyco, (vs, cos)) =
let
- fun print_co (co, []) = str (deresolve co)
- | print_co (co, tys) = concat [str (deresolve co), str "of",
+ fun print_co ((co, _), []) = str (deresolve co)
+ | print_co ((co, _), tys) = concat [str (deresolve co), str "of",
enum " *" "" "" (map (print_typ (INFX (2, X))) tys)];
in
concat (
@@ -552,19 +553,20 @@
@| print_eqns (is_pseudo_fun name) eqs
))
end
- | print_def is_pseudo_fun _ definer (ML_Instance (inst, ((class, (tyco, vs)), (superinsts, classparams)))) =
+ | print_def is_pseudo_fun _ definer
+ (ML_Instance (inst, ((class, (tyco, vs)), (super_instances, (classparam_instances, _))))) =
let
- fun print_superinst (_, (classrel, dss)) =
+ fun print_super_instance (_, (classrel, dss)) =
concat [
(str o deresolve) classrel,
str "=",
print_dict is_pseudo_fun NOBR (DictConst dss)
];
- fun print_classparam ((classparam, c_inst), (thm, _)) =
+ fun print_classparam_instance ((classparam, const), (thm, _)) =
concat [
(str o deresolve) classparam,
str "=",
- print_app (K false) (SOME thm) reserved NOBR (c_inst, [])
+ print_app (K false) (SOME thm) reserved NOBR (const, [])
];
in pair
(print_val_decl print_dicttypscheme
@@ -575,8 +577,8 @@
:: print_dict_args vs
@ str "="
@@ brackets [
- enum_default "()" ";" "{" "}" (map print_superinst superinsts
- @ map print_classparam classparams),
+ enum_default "()" ";" "{" "}" (map print_super_instance super_instances
+ @ map print_classparam_instance classparam_instances),
str ":",
print_tyco_expr NOBR (class, [tyco `%% map (ITyVar o fst) vs])
]
@@ -633,11 +635,11 @@
sig_ps
(Pretty.chunks (ps @| doublesemicolon [p]))
end
- | print_stmt (ML_Class (class, (v, (superclasses, classparams)))) =
+ | print_stmt (ML_Class (class, (v, (super_classes, classparams)))) =
let
fun print_field s p = concat [str s, str ":", p];
- fun print_superclass_field (superclass, classrel) =
- print_field (deresolve classrel) (print_dicttyp (superclass, ITyVar v));
+ fun print_super_class_field (super_class, classrel) =
+ print_field (deresolve classrel) (print_dicttyp (super_class, ITyVar v));
fun print_classparam_decl (classparam, ty) =
print_val_decl print_typscheme
(classparam, ([(v, [class])], ty));
@@ -652,7 +654,7 @@
(str o deresolve) class,
str "=",
enum_default "unit" ";" "{" "}" (
- map print_superclass_field superclasses
+ map print_super_class_field super_classes
@ map print_classparam_field classparams
)
];
@@ -756,7 +758,12 @@
val base' = if upper then first_upper base else base;
val ([base''], nsp') = Name.variants [base'] nsp;
in (base'', nsp') end;
- fun ml_binding_of_stmt (name, Code_Thingol.Fun (_, (tysm as (vs, ty), raw_eqs))) =
+ fun deps_of names =
+ []
+ |> fold (fold (insert (op =)) o Graph.imm_succs program) names
+ |> subtract (op =) names
+ |> filter_out (Code_Thingol.is_case o Graph.get_node program);
+ fun ml_binding_of_stmt (name, Code_Thingol.Fun (_, ((tysm as (vs, ty), raw_eqs), _))) =
let
val eqs = filter (snd o snd) raw_eqs;
val (eqs', is_value) = if null (filter_out (null o snd) vs) then case eqs
@@ -843,10 +850,7 @@
fun add_stmts' stmts nsp_nodes =
let
val names as (name :: names') = map fst stmts;
- val deps =
- []
- |> fold (fold (insert (op =)) o Graph.imm_succs program) names
- |> subtract (op =) names;
+ val deps = deps_of names;
val (module_names, _) = (split_list o map dest_name) names;
val module_name = (the_single o distinct (op =) o map mk_name_module) module_names
handle Empty =>
@@ -880,9 +884,9 @@
|> apsnd (fold (fn name => fold (add_dep name) deps) names)
|> apsnd (fold_product (curry (map_node module_name_path o Graph.add_edge)) names names)
end;
- val (_, nodes) = empty_module
- |> fold add_stmts' (map (AList.make (Graph.get_node program))
- (rev (Graph.strong_conn program)));
+ val stmts = map (AList.make (Graph.get_node program)) (rev (Graph.strong_conn program))
+ |> filter_out (fn [(_, stmt)] => Code_Thingol.is_case stmt | _ => false);
+ val (_, nodes) = fold add_stmts' stmts empty_module;
fun deresolver prefix name =
let
val module_name = (fst o dest_name) name;
--- a/src/Tools/Code/code_preproc.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/Tools/Code/code_preproc.ML Mon Jun 21 17:41:57 2010 +0200
@@ -28,6 +28,7 @@
-> (code_algebra -> code_graph -> (string * sort) list -> term -> cterm -> thm) -> cterm -> thm
val eval: theory -> ((term -> term) -> 'a -> 'a)
-> (code_algebra -> code_graph -> (string * sort) list -> term -> 'a) -> term -> 'a
+ val pre_post_conv: theory -> (cterm -> thm) -> cterm -> thm
val setup: theory -> theory
end
@@ -262,11 +263,11 @@
of SOME classess => (classess, ([], []))
| NONE => let
val all_classes = complete_proper_sort thy [class];
- val superclasses = remove (op =) class all_classes;
+ val super_classes = remove (op =) class all_classes;
val classess = map (complete_proper_sort thy)
(Sign.arity_sorts thy tyco [class]);
val inst_params = inst_params thy tyco all_classes;
- in (classess, (superclasses, inst_params)) end;
+ in (classess, (super_classes, inst_params)) end;
(* computing instantiations *)
@@ -284,14 +285,14 @@
|> fold (fn styp => fold (ensure_typmatch_inst thy arities eqngr styp) new_classes) styps
|> fold (fn c_k => add_classes thy arities eqngr c_k diff_classes) c_ks
end end
-and add_styp thy arities eqngr c_k tyco_styps vardeps_data =
+and add_styp thy arities eqngr c_k new_tyco_styps vardeps_data =
let
- val (old_styps, classes) = Vargraph.get_node (fst vardeps_data) c_k;
- in if member (op =) old_styps tyco_styps then vardeps_data
+ val (old_tyco_stypss, classes) = Vargraph.get_node (fst vardeps_data) c_k;
+ in if member (op =) old_tyco_stypss new_tyco_styps then vardeps_data
else
vardeps_data
- |> (apfst o Vargraph.map_node c_k o apfst) (cons tyco_styps)
- |> fold (ensure_typmatch_inst thy arities eqngr tyco_styps) classes
+ |> (apfst o Vargraph.map_node c_k o apfst) (cons new_tyco_styps)
+ |> fold (ensure_typmatch_inst thy arities eqngr new_tyco_styps) classes
end
and add_dep thy arities eqngr c_k c_k' vardeps_data =
let
@@ -311,20 +312,20 @@
and ensure_inst thy arities eqngr (inst as (class, tyco)) (vardeps_data as (_, (_, insts))) =
if member (op =) insts inst then vardeps_data
else let
- val (classess, (superclasses, inst_params)) =
+ val (classess, (super_classes, inst_params)) =
obtain_instance thy arities inst;
in
vardeps_data
|> (apsnd o apsnd) (insert (op =) inst)
|> fold_index (fn (k, _) =>
apfst (Vargraph.new_node ((Inst (class, tyco), k), ([] ,[])))) classess
- |> fold (fn superclass => ensure_inst thy arities eqngr (superclass, tyco)) superclasses
+ |> fold (fn super_class => ensure_inst thy arities eqngr (super_class, tyco)) super_classes
|> fold (ensure_fun thy arities eqngr) inst_params
|> fold_index (fn (k, classes) =>
add_classes thy arities eqngr (Inst (class, tyco), k) classes
- #> fold (fn superclass =>
- add_dep thy arities eqngr (Inst (superclass, tyco), k)
- (Inst (class, tyco), k)) superclasses
+ #> fold (fn super_class =>
+ add_dep thy arities eqngr (Inst (super_class, tyco), k)
+ (Inst (class, tyco), k)) super_classes
#> fold (fn inst_param =>
add_dep thy arities eqngr (Fun inst_param, k)
(Inst (class, tyco), k)
@@ -457,6 +458,8 @@
fun eval thy postproc evaluator = gen_eval thy (Thm.cterm_of thy)
(K o postproc (postprocess_term thy)) (simple_evaluator evaluator);
+fun pre_post_conv thy conv = (preprocess_conv thy) then_conv conv then_conv (postprocess_conv thy);
+
(** setup **)
--- a/src/Tools/Code/code_printer.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/Tools/Code/code_printer.ML Mon Jun 21 17:41:57 2010 +0200
@@ -256,12 +256,12 @@
fold_map (Code_Thingol.ensure_declared_const thy) cs naming
|-> (fn cs' => pair (n, f literals cs'));
-fun gen_print_app print_app_expr print_term syntax_const thm vars fxy (app as ((c, (_, tys)), ts)) =
+fun gen_print_app print_app_expr print_term syntax_const thm vars fxy (app as ((c, (_, function_typs)), ts)) =
case syntax_const c
of NONE => brackify fxy (print_app_expr thm vars app)
| SOME (k, print) =>
let
- fun print' fxy ts = print (print_term thm) thm vars fxy (ts ~~ take k tys);
+ fun print' fxy ts = print (print_term thm) thm vars fxy (ts ~~ take k function_typs);
in if k = length ts
then print' fxy ts
else if k < length ts
--- a/src/Tools/Code/code_scala.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/Tools/Code/code_scala.ML Mon Jun 21 17:41:57 2010 +0200
@@ -22,7 +22,8 @@
(** Scala serializer **)
-fun print_scala_stmt labelled_name syntax_tyco syntax_const reserved args_num is_singleton deresolve =
+fun print_scala_stmt labelled_name syntax_tyco syntax_const reserved
+ args_num is_singleton deresolve =
let
val deresolve_base = Long_Name.base_name o deresolve;
val lookup_tyvar = first_upper oo lookup_var;
@@ -61,18 +62,21 @@
then print_case tyvars some_thm vars fxy cases
else print_app tyvars is_pat some_thm vars fxy c_ts
| NONE => print_case tyvars some_thm vars fxy cases)
- and print_app tyvars is_pat some_thm vars fxy (app as ((c, ((tys, _), tys_args)), ts)) =
+ and print_app tyvars is_pat some_thm vars fxy
+ (app as ((c, ((arg_typs, _), function_typs)), ts)) =
let
val k = length ts;
val l = case syntax_const c of NONE => args_num c | SOME (l, _) => l;
- val tys' = if is_pat orelse
- (is_none (syntax_const c) andalso is_singleton c) then [] else tys;
+ val arg_typs' = if is_pat orelse
+ (is_none (syntax_const c) andalso is_singleton c) then [] else arg_typs;
val (no_syntax, print') = case syntax_const c
of NONE => (true, fn ts => applify "(" ")" fxy
- (applify "[" "]" NOBR ((str o deresolve) c) (map (print_typ tyvars NOBR) tys'))
- (map (print_term tyvars is_pat some_thm vars NOBR) ts))
+ (applify "[" "]" NOBR ((str o deresolve) c)
+ (map (print_typ tyvars NOBR) arg_typs'))
+ (map (print_term tyvars is_pat some_thm vars NOBR) ts))
| SOME (_, print) => (false, fn ts =>
- print (print_term tyvars is_pat some_thm) some_thm vars fxy (ts ~~ take l tys_args));
+ print (print_term tyvars is_pat some_thm) some_thm vars fxy
+ (ts ~~ take l function_typs));
in if k = l then print' ts
else if k < l then
print_term tyvars is_pat some_thm vars fxy (Code_Thingol.eta_expand l app)
@@ -82,7 +86,8 @@
Pretty.block (print' ts1 :: map (fn t => Pretty.block
[str ".apply(", print_term tyvars is_pat some_thm vars NOBR t, str ")"]) ts23)
end end
- and print_bind tyvars some_thm fxy p = gen_print_bind (print_term tyvars true) some_thm fxy p
+ and print_bind tyvars some_thm fxy p =
+ gen_print_bind (print_term tyvars true) some_thm fxy p
and print_case tyvars some_thm vars fxy (cases as ((_, [_]), _)) =
let
val (binds, body) = Code_Thingol.unfold_let (ICase cases);
@@ -103,8 +108,9 @@
let
fun print_select (pat, body) =
let
- val (p, vars') = print_bind tyvars some_thm NOBR pat vars;
- in concat [str "case", p, str "=>", print_term tyvars false some_thm vars' NOBR body] end;
+ val (p_pat, vars') = print_bind tyvars some_thm NOBR pat vars;
+ val p_body = print_term tyvars false some_thm vars' NOBR body
+ in concat [str "case", p_pat, str "=>", p_body] end;
in brackify_block fxy
(concat [print_term tyvars false some_thm vars NOBR t, str "match", str "{"])
(map print_select clauses)
@@ -112,24 +118,17 @@
end
| print_case tyvars some_thm vars fxy ((_, []), _) =
(brackify fxy o Pretty.breaks o map str) ["error(\"empty case\")"];
- fun implicit_arguments tyvars vs vars =
- let
- val implicit_typ_ps = maps (fn (v, sort) => map (fn class => Pretty.block
- [(str o deresolve) class, str "[", (str o lookup_tyvar tyvars) v, str "]"]) sort) vs;
- val implicit_names = Name.variant_list [] (maps (fn (v, sort) => map (fn class =>
- lookup_tyvar tyvars v ^ "_" ^ (Long_Name.base_name o deresolve) class) sort) vs);
- val vars' = intro_vars implicit_names vars;
- val implicit_ps = map2 (fn v => fn p => concat [str (v ^ ":"), p])
- implicit_names implicit_typ_ps;
- in ((implicit_names, implicit_ps), vars') end;
- fun print_defhead tyvars vars p vs params tys implicits ty =
+ fun print_defhead tyvars vars p vs params tys implicits (*FIXME simplify*) ty =
Pretty.block [str "def ", print_typed tyvars (applify "(implicit " ")" NOBR
(applify "(" ")" NOBR
- (applify "[" "]" NOBR p (map (str o lookup_tyvar tyvars o fst) vs))
+ (applify "[" "]" NOBR p (map (fn (v, sort) =>
+ (str o implode)
+ (lookup_tyvar tyvars v :: map (prefix ": " o deresolve) sort)) vs))
(map2 (fn param => fn ty => print_typed tyvars
((str o lookup_var vars) param) ty)
params tys)) implicits) ty, str " ="]
- fun print_stmt (name, Code_Thingol.Fun (_, ((vs, ty), raw_eqs))) = (case filter (snd o snd) raw_eqs
+ fun print_stmt (name, Code_Thingol.Fun (_, (((vs, ty), raw_eqs), _))) =
+ (case filter (snd o snd) raw_eqs
of [] =>
let
val (tys, ty') = Code_Thingol.unfold_fun ty;
@@ -156,70 +155,78 @@
val vars1 = reserved
|> intro_base_names
(is_none o syntax_const) deresolve consts
- val ((_, implicit_ps), vars2) = implicit_arguments tyvars vs vars1;
- val params = if simple then (map (fn IVar (SOME x) => x) o fst o fst o hd) eqs
- else aux_params vars2 (map (fst o fst) eqs);
- val vars3 = intro_vars params vars2;
+ val params = if simple
+ then (map (fn IVar (SOME x) => x) o fst o fst o hd) eqs
+ else aux_params vars1 (map (fst o fst) eqs);
+ val vars2 = intro_vars params vars1;
val (tys, ty1) = Code_Thingol.unfold_fun ty;
val (tys1, tys2) = chop (length params) tys;
val ty2 = Library.foldr
(fn (ty1, ty2) => Code_Thingol.fun_tyco `%% [ty1, ty2]) (tys2, ty1);
fun print_tuple [p] = p
| print_tuple ps = enum "," "(" ")" ps;
- fun print_rhs vars' ((_, t), (some_thm, _)) = print_term tyvars false some_thm vars' NOBR t;
+ fun print_rhs vars' ((_, t), (some_thm, _)) =
+ print_term tyvars false some_thm vars' NOBR t;
fun print_clause (eq as ((ts, _), (some_thm, _))) =
let
- val vars' = intro_vars ((fold o Code_Thingol.fold_varnames) (insert (op =)) ts []) vars2;
+ val vars' = intro_vars ((fold o Code_Thingol.fold_varnames)
+ (insert (op =)) ts []) vars1;
in
- concat [str "case", print_tuple (map (print_term tyvars true some_thm vars' NOBR) ts),
+ concat [str "case",
+ print_tuple (map (print_term tyvars true some_thm vars' NOBR) ts),
str "=>", print_rhs vars' eq]
end;
- val head = print_defhead tyvars vars3 ((str o deresolve) name) vs params tys1 implicit_ps ty2;
+ val head = print_defhead tyvars vars2 ((str o deresolve) name) vs params tys1 [] ty2;
in if simple then
- concat [head, print_rhs vars3 (hd eqs)]
+ concat [head, print_rhs vars2 (hd eqs)]
else
Pretty.block_enclose
- (concat [head, print_tuple (map (str o lookup_var vars3) params),
+ (concat [head, print_tuple (map (str o lookup_var vars2) params),
str "match", str "{"], str "}")
(map print_clause eqs)
end)
| print_stmt (name, Code_Thingol.Datatype (_, (vs, cos))) =
let
val tyvars = intro_vars (map fst vs) reserved;
- fun print_co (co, []) =
+ fun print_co ((co, _), []) =
concat [str "final", str "case", str "object", (str o deresolve_base) co,
str "extends", applify "[" "]" NOBR ((str o deresolve_base) name)
(replicate (length vs) (str "Nothing"))]
- | print_co (co, tys) =
+ | print_co ((co, vs_args), tys) =
let
val fields = Name.names (snd reserved) "a" tys;
val vars = intro_vars (map fst fields) reserved;
- fun add_typargs p = applify "[" "]" NOBR p
- (map (str o lookup_tyvar tyvars o fst) vs);
+ fun add_typargs1 p = applify "[" "]" NOBR p
+ (map (str o lookup_tyvar tyvars o fst) vs); (*FIXME*)
+ fun add_typargs2 p = applify "[" "]" NOBR p
+ (map (str o lookup_tyvar tyvars) vs_args); (*FIXME*)
in
concat [
applify "(" ")" NOBR
- (add_typargs ((concat o map str) ["final", "case", "class", deresolve_base co]))
+ (add_typargs2 ((concat o map str)
+ ["final", "case", "class", deresolve_base co]))
(map (uncurry (print_typed tyvars) o apfst str) fields),
str "extends",
- add_typargs ((str o deresolve_base) name)
+ add_typargs1 ((str o deresolve_base) name)
]
end
in
Pretty.chunks (
- applify "[" "]" NOBR ((concat o map str) ["sealed", "class", deresolve_base name])
+ applify "[" "]" NOBR ((concat o map str)
+ ["sealed", "class", deresolve_base name])
(map (str o prefix "+" o lookup_tyvar tyvars o fst) vs)
:: map print_co cos
)
end
- | print_stmt (name, Code_Thingol.Class (_, (v, (superclasses, classparams)))) =
+ | print_stmt (name, Code_Thingol.Class (_, (v, (super_classes, classparams)))) =
let
val tyvars = intro_vars [v] reserved;
val vs = [(v, [name])];
fun add_typarg p = applify "[" "]" NOBR p [(str o lookup_tyvar tyvars) v];
- fun print_superclasses [] = NONE
- | print_superclasses classes = SOME (concat (str "extends"
- :: separate (str "with") (map (add_typarg o str o deresolve o fst) classes)));
+ fun print_super_classes [] = NONE
+ | print_super_classes classes = SOME (concat (str "extends"
+ :: separate (str "with")
+ (map (add_typarg o str o deresolve o fst) classes)));
fun print_tupled_typ ([], ty) =
print_typ tyvars NOBR ty
| print_tupled_typ ([ty1], ty2) =
@@ -228,39 +235,62 @@
concat [enum "," "(" ")" (map (print_typ tyvars NOBR) tys),
str "=>", print_typ tyvars NOBR ty];
fun print_classparam_val (classparam, ty) =
- concat [str "val", (str o Library.enclose "`" "+`:" o deresolve_base) classparam,
+ concat [str "val", (str o Library.enclose "`" "+`:" o deresolve_base)
+ classparam,
(print_tupled_typ o Code_Thingol.unfold_fun) ty]
+ fun implicit_arguments tyvars vs vars = (*FIXME simplifiy*)
+ let
+ val implicit_typ_ps = maps (fn (v, sort) => map (fn class => Pretty.block
+ [(str o deresolve) class, str "[",
+ (str o lookup_tyvar tyvars) v, str "]"]) sort) vs;
+ val implicit_names = Name.variant_list [] (maps (fn (v, sort) =>
+ map (fn class => lookup_tyvar tyvars v ^ "_" ^
+ (Long_Name.base_name o deresolve) class) sort) vs);
+ val vars' = intro_vars implicit_names vars;
+ val implicit_ps = map2 (fn v => fn p => concat [str (v ^ ":"), p])
+ implicit_names implicit_typ_ps;
+ in ((implicit_names, implicit_ps), vars') end;
fun print_classparam_def (classparam, ty) =
let
val (tys, ty) = Code_Thingol.unfold_fun ty;
val params = Name.invents (snd reserved) "a" (length tys);
val vars = intro_vars params reserved;
val (([implicit], [p_implicit]), vars') = implicit_arguments tyvars vs vars;
- val head = print_defhead tyvars vars' ((str o deresolve) classparam) vs params tys [p_implicit] ty;
+ val head = print_defhead tyvars vars' ((str o deresolve) classparam)
+ ((map o apsnd) (K []) vs) params tys [p_implicit] ty;
in
concat [head, applify "(" ")" NOBR
- (Pretty.block [str implicit, str ".", (str o Library.enclose "`" "+`" o deresolve_base) classparam])
+ (Pretty.block [str implicit, str ".",
+ (str o Library.enclose "`" "+`" o deresolve_base) classparam])
(map (str o lookup_var vars') params)]
end;
in
Pretty.chunks (
(Pretty.block_enclose
(concat ([str "trait", add_typarg ((str o deresolve_base) name)]
- @ the_list (print_superclasses superclasses) @ [str "{"]), str "}")
+ @ the_list (print_super_classes super_classes) @ [str "{"]), str "}")
(map print_classparam_val classparams))
:: map print_classparam_def classparams
)
end
| print_stmt (name, Code_Thingol.Classinst ((class, (tyco, vs)),
- (super_instances, classparam_insts))) =
+ (super_instances, (classparam_instances, further_classparam_instances)))) =
let
val tyvars = intro_vars (map fst vs) reserved;
- val insttyp = tyco `%% map (ITyVar o fst) vs;
- val p_inst_typ = print_typ tyvars NOBR insttyp;
- fun add_typ_params p = applify "[" "]" NOBR p (map (str o lookup_tyvar tyvars o fst) vs);
- fun add_inst_typ p = Pretty.block [p, str "[", p_inst_typ, str "]"];
- val ((implicit_names, implicit_ps), vars) = implicit_arguments tyvars vs reserved;
- fun print_classparam_inst ((classparam, const as (_, (_, tys))), (thm, _)) =
+ val classtyp = (class, (tyco, map fst vs));
+ fun print_classtyp' (class, (tyco, vs)) = (*FIXME already exists?*)
+ applify "[" "]" NOBR(*? FIXME*) ((str o deresolve) class)
+ [print_typ tyvars NOBR (tyco `%% map ITyVar vs)]; (*FIXME a mess...*)
+ fun print_typed' tyvars p classtyp = (*FIXME unify with existing print_typed*)
+ Pretty.block [p, str ":", Pretty.brk 1, print_classtyp' classtyp];
+ fun print_defhead' tyvars vars p vs params tys classtyp = (*FIXME unify with existing def_head*)
+ Pretty.block [str "def ", print_typed' tyvars
+ (applify "(" ")" NOBR
+ (applify "[" "]" NOBR p (map (fn (v, sort) =>
+ (str o implode) (lookup_tyvar tyvars v :: map (prefix ": " o deresolve) sort)) vs))
+ (map2 (fn param => fn ty => print_typed tyvars ((str o lookup_var vars) param) ty)
+ params tys)) classtyp, str " ="];
+ fun print_classparam_instance ((classparam, const as (_, (_, tys))), (thm, _)) =
let
val auxs = Name.invents (snd reserved) "a" (length tys);
val vars = intro_vars auxs reserved;
@@ -271,27 +301,10 @@
concat ([str "val", (str o Library.enclose "`" "+`" o deresolve_base) classparam,
str "="] @ args @ [print_app tyvars false (SOME thm) vars NOBR (const, map (IVar o SOME) auxs)])
end;
- in
- Pretty.chunks [
- Pretty.block_enclose
- (concat ([str "trait",
- add_typ_params ((str o deresolve_base) name),
- str "extends",
- Pretty.block [(str o deresolve) class, str "[", p_inst_typ, str "]"]]
- @ map (fn (_, (_, (superinst, _))) => add_typ_params (str ("with " ^ deresolve superinst)))
- super_instances @| str "{"), str "}")
- (map (fn p => Pretty.block [str "implicit val arg$", p]) implicit_ps
- @ map print_classparam_inst classparam_insts),
- concat [str "implicit", str (if null vs then "val" else "def"),
- applify "(implicit " ")" NOBR (applify "[" "]" NOBR
- ((str o deresolve_base) name) (map (str o lookup_tyvar tyvars o fst) vs))
- implicit_ps,
- str "=", str "new", applify "[" "]" NOBR ((str o deresolve_base) name)
- (map (str o lookup_tyvar tyvars o fst) vs),
- Pretty.enum ";" "{ " " }" (map (str o (fn v => "val arg$" ^ v ^ " = " ^ v) o lookup_var vars)
- implicit_names)]
- ]
- end;
+ val head = print_defhead' tyvars reserved ((str o deresolve) name) vs [] [] classtyp;
+ val body = [str "new", print_classtyp' classtyp,
+ Pretty.enum ";" "{ " " }" (map print_classparam_instance (classparam_instances @ further_classparam_instances))];
+ in concat (str "implicit" :: head :: body) end;
in print_stmt end;
fun scala_program_of_program labelled_name module_name reserved raw_module_alias program =
@@ -315,7 +328,10 @@
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;
+ 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
@@ -334,15 +350,17 @@
|> add_name
|-> (fn base' => rpair (add_stmt base' stmts))
end;
- val (_, sca_program) = fold prepare_stmt (AList.make (fn name => Graph.get_node program name)
- (Graph.strong_conn program |> flat)) (((reserved, reserved), reserved), []);
+ 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;
fun serialize_scala raw_module_name labelled_name
raw_reserved includes raw_module_alias
- _ syntax_tyco syntax_const (code_of_pretty, code_writeln) program stmt_names destination =
+ _ syntax_tyco syntax_const (code_of_pretty, code_writeln)
+ program 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";
@@ -351,18 +369,21 @@
module_name reserved raw_module_alias program;
val reserved = make_vars reserved;
fun args_num c = case Graph.get_node program c
- of Code_Thingol.Fun (_, ((_, ty), [])) => (length o fst o Code_Thingol.unfold_fun) ty
- | Code_Thingol.Fun (_, (_, ((ts, _), _) :: _)) => length ts
+ of Code_Thingol.Fun (_, (((_, ty), []), _)) =>
+ (length o fst o Code_Thingol.unfold_fun) ty
+ | Code_Thingol.Fun (_, ((_, ((ts, _), _) :: _), _)) => length ts
| Code_Thingol.Datatypecons (_, tyco) =>
let val Code_Thingol.Datatype (_, (_, constrs)) = Graph.get_node program tyco
- in (length o the o AList.lookup (op =) constrs) c end
+ in (length o the o AList.lookup (eq_fst op =) constrs) (c, []) end (*FIXME simplify lookup*)
| Code_Thingol.Classparam (_, class) =>
- let val Code_Thingol.Class (_, (_, (_, classparams))) = Graph.get_node program class
+ let
+ val Code_Thingol.Class (_, (_, (_, classparams))) =
+ Graph.get_node program class
in (length o fst o Code_Thingol.unfold_fun o the o AList.lookup (op =) classparams) c end;
fun is_singleton c = case Graph.get_node program c
of Code_Thingol.Datatypecons (_, tyco) =>
let val Code_Thingol.Datatype (_, (_, constrs)) = Graph.get_node program tyco
- in (null o the o AList.lookup (op =) constrs) c end
+ in (null o the o AList.lookup (eq_fst op =) constrs) (c, []) end (*FIXME simplify lookup*)
| _ => false;
val print_stmt = print_scala_stmt labelled_name syntax_tyco syntax_const
reserved args_num is_singleton deresolver;
@@ -432,12 +453,13 @@
val setup =
Code_Target.add_target (target, (isar_seri_scala, literals))
- #> Code_Target.add_syntax_tyco target "fun" (SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] =>
- brackify_infix (1, R) fxy (
- print_typ BR ty1 (*product type vs. tupled arguments!*),
- str "=>",
- print_typ (INFX (1, R)) ty2
- )))
+ #> Code_Target.add_syntax_tyco target "fun"
+ (SOME (2, fn print_typ => fn fxy => fn [ty1, ty2] =>
+ brackify_infix (1, R) fxy (
+ print_typ BR ty1 (*product type vs. tupled arguments!*),
+ str "=>",
+ print_typ (INFX (1, R)) ty2
+ )))
#> fold (Code_Target.add_reserved target) [
"abstract", "case", "catch", "class", "def", "do", "else", "extends", "false",
"final", "finally", "for", "forSome", "if", "implicit", "import", "lazy",
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Code/code_simp.ML Mon Jun 21 17:41:57 2010 +0200
@@ -0,0 +1,91 @@
+(* Title: Tools/code/code_simp.ML
+ Author: Florian Haftmann, TU Muenchen
+
+Connecting the simplifier and the code generator.
+*)
+
+signature CODE_SIMP =
+sig
+ val no_frees_conv: conv -> conv
+ val map_ss: (simpset -> simpset) -> theory -> theory
+ val current_conv: theory -> conv
+ val current_tac: theory -> int -> tactic
+ val current_value: theory -> term -> term
+ val make_conv: theory -> simpset option -> string list -> conv
+ val make_tac: theory -> simpset option -> string list -> int -> tactic
+ val setup: theory -> theory
+end;
+
+structure Code_Simp : CODE_SIMP =
+struct
+
+(* avoid free variables during conversion *)
+
+fun no_frees_conv conv ct =
+ let
+ val frees = Thm.add_cterm_frees ct [];
+ fun apply_beta free thm = Thm.combination thm (Thm.reflexive free)
+ |> Conv.fconv_rule (Conv.arg_conv (Conv.try_conv (Thm.beta_conversion false)))
+ |> Conv.fconv_rule (Conv.arg1_conv (Thm.beta_conversion false));
+ in
+ ct
+ |> fold_rev Thm.cabs frees
+ |> conv
+ |> fold apply_beta frees
+ end;
+
+
+(* dedicated simpset *)
+
+structure Simpset = Theory_Data (
+ type T = simpset;
+ val empty = empty_ss;
+ fun extend ss = MetaSimplifier.inherit_context empty_ss ss;
+ val merge = merge_ss;
+);
+
+val map_ss = Simpset.map;
+
+fun simpset_default thy = Simplifier.global_context thy o the_default (Simpset.get thy);
+
+
+(* build simpset and conversion from program *)
+
+fun add_stmt (Code_Thingol.Fun (_, ((_, eqs), some_cong))) ss =
+ ss addsimps (map_filter (fst o snd)) eqs addcongs (the_list some_cong)
+ | add_stmt (Code_Thingol.Classinst (_, (_, (classparam_instances, _)))) ss =
+ ss addsimps (map (fst o snd) classparam_instances)
+ | add_stmt _ ss = ss;
+
+val add_program = Graph.fold (add_stmt o fst o snd)
+
+fun rewrite_modulo thy some_ss program = Simplifier.full_rewrite
+ (add_program program (simpset_default thy some_ss));
+
+fun conclude_tac thy some_ss = Simplifier.full_simp_tac (simpset_default thy some_ss);
+
+
+(* evaluation with current code context *)
+
+fun current_conv thy = no_frees_conv (Code_Thingol.eval_conv thy
+ (fn naming => fn program => fn t => fn deps => rewrite_modulo thy NONE program));
+
+fun current_tac thy = CONVERSION (current_conv thy) THEN' conclude_tac thy NONE;
+
+fun current_value thy = snd o Logic.dest_equals o Thm.prop_of o current_conv thy o Thm.cterm_of thy;
+
+val setup = Method.setup (Binding.name "code_simp")
+ (Scan.succeed (SIMPLE_METHOD' o (CHANGED_PROP oo current_tac o ProofContext.theory_of)))
+ "simplification with code equations"
+ #> Value.add_evaluator ("simp", current_value o ProofContext.theory_of);
+
+
+(* evaluation with freezed code context *)
+
+fun make_conv thy some_ss consts = no_frees_conv (Code_Preproc.pre_post_conv thy
+ ((rewrite_modulo thy some_ss o snd o snd o Code_Thingol.consts_program thy false) consts));
+
+fun make_tac thy some_ss consts = CONVERSION (make_conv thy some_ss consts)
+ THEN' conclude_tac thy some_ss;
+
+end;
--- a/src/Tools/Code/code_thingol.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/Tools/Code/code_thingol.ML Mon Jun 21 17:41:57 2010 +0200
@@ -66,20 +66,23 @@
datatype stmt =
NoStmt
- | Fun of string * (typscheme * ((iterm list * iterm) * (thm option * bool)) list)
- | Datatype of string * ((vname * sort) list * (string * itype list) list)
+ | Fun of string * ((typscheme * ((iterm list * iterm) * (thm option * bool)) list) * thm option)
+ | Datatype of string * ((vname * sort) list *
+ ((string * vname list (*type argument wrt. canonical order*)) * itype list) list)
| Datatypecons of string * string
| Class of class * (vname * ((class * string) list * (string * itype) list))
| Classrel of class * class
| Classparam of string * class
- | Classinst of (class * (string * (vname * sort) list))
- * ((class * (string * (string * dict list list))) list
- * ((string * const) * (thm * bool)) list)
+ | Classinst of (class * (string * (vname * sort) list) (*class and arity*))
+ * ((class * (string * (string * dict list list))) list (*super instances*)
+ * (((string * const) * (thm * bool)) list (*class parameter instances*)
+ * ((string * const) * (thm * bool)) list (*super class parameter instances*)))
type program = stmt Graph.T
val empty_funs: program -> string list
val map_terms_bottom_up: (iterm -> iterm) -> iterm -> iterm
val map_terms_stmt: (iterm -> iterm) -> stmt -> stmt
val is_cons: program -> string -> bool
+ val is_case: stmt -> bool
val contr_classparam_typs: program -> string -> itype option list
val labelled_name: theory -> program -> string -> string
val group_stmts: theory -> program
@@ -276,8 +279,8 @@
in
fun namify_class thy = namify thy Long_Name.base_name thyname_of_class;
-fun namify_classrel thy = namify thy (fn (class1, class2) =>
- Long_Name.base_name class2 ^ "_" ^ Long_Name.base_name class1)
+fun namify_classrel thy = namify thy (fn (sub_class, super_class) =>
+ Long_Name.base_name super_class ^ "_" ^ Long_Name.base_name sub_class)
(fn thy => thyname_of_class thy o fst);
(*order fits nicely with composed projections*)
fun namify_tyco thy "fun" = "Pure.fun"
@@ -386,11 +389,12 @@
of SOME const' => (const', naming)
| NONE => declare_const thy const naming;
-val fun_tyco = "Pure.fun.tyco" (*depends on suffix_tyco and namify_tyco!*);
+val fun_tyco = Long_Name.append (namify_tyco Pure.thy "fun") suffix_tyco
+ (*depends on add_suffix*);
val unfold_fun = unfoldr
(fn tyco `%% [ty1, ty2] => if tyco = fun_tyco then SOME (ty1, ty2) else NONE
- | _ => NONE); (*depends on suffix_tyco and namify_tyco!*)
+ | _ => NONE);
end; (* local *)
@@ -400,20 +404,22 @@
type typscheme = (vname * sort) list * itype;
datatype stmt =
NoStmt
- | Fun of string * (typscheme * ((iterm list * iterm) * (thm option * bool)) list)
- | Datatype of string * ((vname * sort) list * (string * itype list) list)
+ | Fun of string * ((typscheme * ((iterm list * iterm) * (thm option * bool)) list) * thm option)
+ | Datatype of string * ((vname * sort) list * ((string * vname list) * itype list) list)
| Datatypecons of string * string
| Class of class * (vname * ((class * string) list * (string * itype) list))
| Classrel of class * class
| Classparam of string * class
| Classinst of (class * (string * (vname * sort) list))
* ((class * (string * (string * dict list list))) list
- * ((string * const) * (thm * bool)) list);
+ * (((string * const) * (thm * bool)) list
+ * ((string * const) * (thm * bool)) list))
+ (*see also signature*);
type program = stmt Graph.T;
fun empty_funs program =
- Graph.fold (fn (name, (Fun (c, (_, [])), _)) => cons c
+ Graph.fold (fn (name, (Fun (c, ((_, []), _)), _)) => cons c
| _ => I) program [];
fun map_terms_bottom_up f (t as IConst _) = f t
@@ -426,22 +432,27 @@
(ICase (((map_terms_bottom_up f t, ty), (map o pairself)
(map_terms_bottom_up f) ps), map_terms_bottom_up f t0));
+fun map_classparam_instances_as_term f =
+ (map o apfst o apsnd) (fn const => case f (IConst const) of IConst const' => const')
+
fun map_terms_stmt f NoStmt = NoStmt
- | map_terms_stmt f (Fun (c, (tysm, eqs))) = Fun (c, (tysm, (map o apfst)
- (fn (ts, t) => (map f ts, f t)) eqs))
+ | map_terms_stmt f (Fun (c, ((tysm, eqs), case_cong))) = Fun (c, ((tysm, (map o apfst)
+ (fn (ts, t) => (map f ts, f t)) eqs), case_cong))
| map_terms_stmt f (stmt as Datatype _) = stmt
| map_terms_stmt f (stmt as Datatypecons _) = stmt
| map_terms_stmt f (stmt as Class _) = stmt
| map_terms_stmt f (stmt as Classrel _) = stmt
| map_terms_stmt f (stmt as Classparam _) = stmt
- | map_terms_stmt f (Classinst (arity, (superinsts, classparams))) =
- Classinst (arity, (superinsts, (map o apfst o apsnd) (fn const =>
- case f (IConst const) of IConst const' => const') classparams));
+ | map_terms_stmt f (Classinst (arity, (super_instances, classparam_instances))) =
+ Classinst (arity, (super_instances, (pairself o map_classparam_instances_as_term) f classparam_instances));
fun is_cons program name = case Graph.get_node program name
of Datatypecons _ => true
| _ => false;
+fun is_case (Fun (_, (_, SOME _))) = true
+ | is_case _ = false;
+
fun contr_classparam_typs program name = case Graph.get_node program name
of Classparam (_, class) => let
val Class (_, (_, (_, params))) = Graph.get_node program class;
@@ -470,6 +481,10 @@
val Datatype (tyco, _) = Graph.get_node program tyco
in quote (Sign.extern_type thy tyco ^ " :: " ^ Sign.extern_class thy class) end
+fun linear_stmts program =
+ rev (Graph.strong_conn program)
+ |> map (AList.make (Graph.get_node program));
+
fun group_stmts thy program =
let
fun is_fun (_, Fun _) = true | is_fun _ = false;
@@ -489,8 +504,7 @@
else error ("Illegal mutual dependencies: " ^
(commas o map (labelled_name thy program o fst)) stmts)
in
- rev (Graph.strong_conn program)
- |> map (AList.make (Graph.get_node program))
+ linear_stmts program
|> map group
end;
@@ -549,8 +563,9 @@
val (vs, cos) = Code.get_type thy tyco;
val stmt_datatype =
fold_map (translate_tyvar_sort thy algbr eqngr permissive) vs
- ##>> fold_map (fn (c, tys) =>
+ ##>> fold_map (fn ((c, vs), tys) =>
ensure_const thy algbr eqngr permissive c
+ ##>> pair (map (unprefix "'") vs)
##>> fold_map (translate_typ thy algbr eqngr permissive) tys) cos
#>> (fn info => Datatype (tyco, info));
in ensure_stmt lookup_tyco (declare_tyco thy) stmt_datatype tyco end
@@ -565,12 +580,12 @@
fun stmt_fun cert =
let
val ((vs, ty), eqns) = Code.equations_of_cert thy cert;
+ val some_case_cong = Code.get_case_cong thy c;
in
fold_map (translate_tyvar_sort thy algbr eqngr permissive) vs
##>> translate_typ thy algbr eqngr permissive ty
- ##>> (fold_map (translate_equation thy algbr eqngr permissive) eqns
- #>> map_filter I)
- #>> (fn info => Fun (c, info))
+ ##>> translate_eqns thy algbr eqngr permissive eqns
+ #>> (fn info => Fun (c, (info, some_case_cong)))
end;
val stmt_const = case Code.get_type_of_constr_or_abstr thy c
of SOME (tyco, _) => stmt_datatypecons tyco
@@ -580,57 +595,63 @@
in ensure_stmt lookup_const (declare_const thy) stmt_const c end
and ensure_class thy (algbr as (_, algebra)) eqngr permissive class =
let
- val superclasses = (Sorts.minimize_sort algebra o Sorts.super_classes algebra) class;
+ val super_classes = (Sorts.minimize_sort algebra o Sorts.super_classes algebra) class;
val cs = #params (AxClass.get_info thy class);
val stmt_class =
- fold_map (fn superclass => ensure_class thy algbr eqngr permissive superclass
- ##>> ensure_classrel thy algbr eqngr permissive (class, superclass)) superclasses
+ fold_map (fn super_class => ensure_class thy algbr eqngr permissive super_class
+ ##>> ensure_classrel thy algbr eqngr permissive (class, super_class)) super_classes
##>> fold_map (fn (c, ty) => ensure_const thy algbr eqngr permissive c
##>> translate_typ thy algbr eqngr permissive ty) cs
#>> (fn info => Class (class, (unprefix "'" Name.aT, info)))
in ensure_stmt lookup_class (declare_class thy) stmt_class class end
-and ensure_classrel thy algbr eqngr permissive (subclass, superclass) =
+and ensure_classrel thy algbr eqngr permissive (sub_class, super_class) =
let
val stmt_classrel =
- ensure_class thy algbr eqngr permissive subclass
- ##>> ensure_class thy algbr eqngr permissive superclass
+ ensure_class thy algbr eqngr permissive sub_class
+ ##>> ensure_class thy algbr eqngr permissive super_class
#>> Classrel;
- in ensure_stmt lookup_classrel (declare_classrel thy) stmt_classrel (subclass, superclass) end
+ in ensure_stmt lookup_classrel (declare_classrel thy) stmt_classrel (sub_class, super_class) end
and ensure_inst thy (algbr as (_, algebra)) eqngr permissive (class, tyco) =
let
- val superclasses = (Sorts.minimize_sort algebra o Sorts.super_classes algebra) class;
- val classparams = these (try (#params o AxClass.get_info thy) class);
+ val super_classes = (Sorts.minimize_sort algebra o Sorts.super_classes algebra) class;
+ val these_classparams = these o try (#params o AxClass.get_info thy);
+ val classparams = these_classparams class;
+ val further_classparams = maps these_classparams
+ ((Sorts.complete_sort algebra o Sorts.super_classes algebra) class);
val vs = Name.names Name.context "'a" (Sorts.mg_domain algebra tyco [class]);
val sorts' = Sorts.mg_domain (Sign.classes_of thy) tyco [class];
val vs' = map2 (fn (v, sort1) => fn sort2 => (v,
Sorts.inter_sort (Sign.classes_of thy) (sort1, sort2))) vs sorts';
val arity_typ = Type (tyco, map TFree vs);
val arity_typ' = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort)) vs');
- fun translate_superarity superclass =
- ensure_class thy algbr eqngr permissive superclass
- ##>> ensure_classrel thy algbr eqngr permissive (class, superclass)
- ##>> translate_dicts thy algbr eqngr permissive NONE (arity_typ, [superclass])
- #>> (fn ((superclass, classrel), [DictConst (inst, dss)]) =>
- (superclass, (classrel, (inst, dss))));
- fun translate_classparam_inst (c, ty) =
+ fun translate_super_instance super_class =
+ ensure_class thy algbr eqngr permissive super_class
+ ##>> ensure_classrel thy algbr eqngr permissive (class, super_class)
+ ##>> translate_dicts thy algbr eqngr permissive NONE (arity_typ, [super_class])
+ #>> (fn ((super_class, classrel), [DictConst (inst, dss)]) =>
+ (super_class, (classrel, (inst, dss))));
+ fun translate_classparam_instance (c, ty) =
let
- val c_inst = Const (c, map_type_tfree (K arity_typ') ty);
- val thm = AxClass.unoverload_conv thy (Thm.cterm_of thy c_inst);
- val c_ty = (apsnd Logic.unvarifyT_global o dest_Const o snd
+ val raw_const = Const (c, map_type_tfree (K arity_typ') ty);
+ val thm = AxClass.unoverload_conv thy (Thm.cterm_of thy raw_const);
+ val const = (apsnd Logic.unvarifyT_global o dest_Const o snd
o Logic.dest_equals o Thm.prop_of) thm;
in
ensure_const thy algbr eqngr permissive c
- ##>> translate_const thy algbr eqngr permissive (SOME thm) (c_ty, NONE)
- #>> (fn (c, IConst c_inst) => ((c, c_inst), (thm, true)))
+ ##>> translate_const thy algbr eqngr permissive (SOME thm) (const, NONE)
+ #>> (fn (c, IConst const') => ((c, const'), (thm, true)))
end;
val stmt_inst =
ensure_class thy algbr eqngr permissive class
##>> ensure_tyco thy algbr eqngr permissive tyco
##>> fold_map (translate_tyvar_sort thy algbr eqngr permissive) vs
- ##>> fold_map translate_superarity superclasses
- ##>> fold_map translate_classparam_inst classparams
- #>> (fn ((((class, tyco), arity), superinsts), classparams) =>
- Classinst ((class, (tyco, arity)), (superinsts, classparams)));
+ ##>> fold_map translate_super_instance super_classes
+ ##>> fold_map translate_classparam_instance classparams
+ ##>> fold_map translate_classparam_instance further_classparams
+ #>> (fn (((((class, tyco), arity_args), super_instances),
+ classparam_instances), further_classparam_instances) =>
+ Classinst ((class, (tyco, arity_args)), (super_instances,
+ (classparam_instances, further_classparam_instances))));
in ensure_stmt lookup_instance (declare_instance thy) stmt_inst (class, tyco) end
and translate_typ thy algbr eqngr permissive (TFree (v, _)) =
pair (ITyVar (unprefix "'" v))
@@ -664,9 +685,9 @@
fold_map (translate_term thy algbr eqngr permissive some_thm) args
##>> translate_term thy algbr eqngr permissive some_thm (rhs, some_abs)
#>> rpair (some_thm, proper)
-and translate_equation thy algbr eqngr permissive eqn prgrm =
- prgrm |> translate_eqn thy algbr eqngr permissive eqn |>> SOME
- handle PERMISSIVE () => (NONE, prgrm)
+and translate_eqns thy algbr eqngr permissive eqns prgrm =
+ prgrm |> fold_map (translate_eqn thy algbr eqngr permissive) eqns
+ handle PERMISSIVE () => ([], prgrm)
and translate_const thy algbr eqngr permissive some_thm ((c, ty), some_abs) =
let
val _ = if (case some_abs of NONE => true | SOME abs => not (c = abs))
@@ -674,15 +695,15 @@
then translation_error thy permissive some_thm
"Abstraction violation" ("constant " ^ Code.string_of_const thy c)
else ()
- val tys = Sign.const_typargs thy (c, ty);
+ val arg_typs = Sign.const_typargs thy (c, ty);
val sorts = Code_Preproc.sortargs eqngr c;
- val tys_args = (fst o Term.strip_type) ty;
+ val function_typs = (fst o Term.strip_type) ty;
in
ensure_const thy algbr eqngr permissive c
- ##>> fold_map (translate_typ thy algbr eqngr permissive) tys
- ##>> fold_map (translate_dicts thy algbr eqngr permissive some_thm) (tys ~~ sorts)
- ##>> fold_map (translate_typ thy algbr eqngr permissive) tys_args
- #>> (fn (((c, tys), iss), tys_args) => IConst (c, ((tys, iss), tys_args)))
+ ##>> fold_map (translate_typ thy algbr eqngr permissive) arg_typs
+ ##>> fold_map (translate_dicts thy algbr eqngr permissive some_thm) (arg_typs ~~ sorts)
+ ##>> fold_map (translate_typ thy algbr eqngr permissive) function_typs
+ #>> (fn (((c, arg_typs), dss), function_typs) => IConst (c, ((arg_typs, dss), function_typs)))
end
and translate_app_const thy algbr eqngr permissive some_thm ((c_ty, ts), some_abs) =
translate_const thy algbr eqngr permissive some_thm (c_ty, some_abs)
@@ -772,8 +793,8 @@
| Local of (class * class) list * (string * (int * sort));
fun class_relation (Global ((_, tyco), yss), _) class =
Global ((class, tyco), yss)
- | class_relation (Local (classrels, v), subclass) superclass =
- Local ((subclass, superclass) :: classrels, v);
+ | class_relation (Local (classrels, v), sub_class) super_class =
+ Local ((sub_class, super_class) :: classrels, v);
fun type_constructor (tyco, _) yss class =
Global ((class, tyco), (map o map) fst yss);
fun type_variable (TFree (v, sort)) =
@@ -844,10 +865,10 @@
##>> translate_typ thy algbr eqngr false ty
##>> translate_term thy algbr eqngr false NONE (Code.subst_signatures thy t, NONE)
#>> (fn ((vs, ty), t) => Fun
- (Term.dummy_patternN, ((vs, ty), [(([], t), (NONE, true))])));
+ (Term.dummy_patternN, (((vs, ty), [(([], t), (NONE, true))]), NONE)));
fun term_value (dep, (naming, program1)) =
let
- val Fun (_, (vs_ty, [(([], t), _)])) =
+ val Fun (_, ((vs_ty, [(([], t), _)]), _)) =
Graph.get_node program1 Term.dummy_patternN;
val deps = Graph.imm_succs program1 Term.dummy_patternN;
val program2 = Graph.del_nodes [Term.dummy_patternN] program1;
--- a/src/Tools/Code_Generator.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/Tools/Code_Generator.thy Mon Jun 21 17:41:57 2010 +0200
@@ -13,6 +13,7 @@
"~~/src/Tools/value.ML"
"~~/src/Tools/Code/code_preproc.ML"
"~~/src/Tools/Code/code_thingol.ML"
+ "~~/src/Tools/Code/code_simp.ML"
"~~/src/Tools/Code/code_printer.ML"
"~~/src/Tools/Code/code_target.ML"
"~~/src/Tools/Code/code_ml.ML"
@@ -24,6 +25,7 @@
setup {*
Code_Preproc.setup
+ #> Code_Simp.setup
#> Code_ML.setup
#> Code_Eval.setup
#> Code_Haskell.setup
--- a/src/Tools/nbe.ML Mon Jun 21 11:37:25 2010 +0200
+++ b/src/Tools/nbe.ML Mon Jun 21 17:41:57 2010 +0200
@@ -396,17 +396,17 @@
(* preparing function equations *)
-fun eqns_of_stmt (_, Code_Thingol.Fun (_, (_, []))) =
+fun eqns_of_stmt (_, Code_Thingol.Fun (_, ((_, []), _))) =
[]
- | eqns_of_stmt (const, Code_Thingol.Fun (_, ((vs, _), eqns))) =
+ | eqns_of_stmt (const, Code_Thingol.Fun (_, (((vs, _), eqns), _))) =
[(const, (vs, map fst eqns))]
| eqns_of_stmt (_, Code_Thingol.Datatypecons _) =
[]
| eqns_of_stmt (_, Code_Thingol.Datatype _) =
[]
- | eqns_of_stmt (class, Code_Thingol.Class (_, (v, (superclasses, classops)))) =
+ | eqns_of_stmt (class, Code_Thingol.Class (_, (v, (super_classes, classparams)))) =
let
- val names = map snd superclasses @ map fst classops;
+ val names = map snd super_classes @ map fst classparams;
val params = Name.invent_list [] "d" (length names);
fun mk (k, name) =
(name, ([(v, [])],
@@ -417,10 +417,10 @@
[]
| eqns_of_stmt (_, Code_Thingol.Classparam _) =
[]
- | eqns_of_stmt (inst, Code_Thingol.Classinst ((class, (_, arities)), (superinsts, instops))) =
- [(inst, (arities, [([], IConst (class, (([], []), [])) `$$
- map (fn (_, (_, (inst, dicts))) => IConst (inst, (([], dicts), []))) superinsts
- @ map (IConst o snd o fst) instops)]))];
+ | eqns_of_stmt (inst, Code_Thingol.Classinst ((class, (_, arity_args)), (super_instances, (classparam_instances, _)))) =
+ [(inst, (arity_args, [([], IConst (class, (([], []), [])) `$$
+ map (fn (_, (_, (inst, dss))) => IConst (inst, (([], dss), []))) super_instances
+ @ map (IConst o snd o fst) classparam_instances)]))];
fun compile_stmts ctxt stmts_deps =
let
@@ -581,19 +581,6 @@
fun norm_oracle thy program vsp_ty_t deps ct =
raw_norm_oracle (thy, program, vsp_ty_t, deps, ct);
-fun no_frees_conv conv ct =
- let
- val frees = Thm.add_cterm_frees ct [];
- fun apply_beta free thm = Thm.combination thm (Thm.reflexive free)
- |> Conv.fconv_rule (Conv.arg_conv (Conv.try_conv (Thm.beta_conversion false)))
- |> Conv.fconv_rule (Conv.arg1_conv (Thm.beta_conversion false));
- in
- ct
- |> fold_rev Thm.cabs frees
- |> conv
- |> fold apply_beta frees
- end;
-
fun no_frees_rew rew t =
let
val frees = map Free (Term.add_frees t []);
@@ -604,7 +591,7 @@
|> (fn t' => Term.betapplys (t', frees))
end;
-val norm_conv = no_frees_conv (fn ct =>
+val norm_conv = Code_Simp.no_frees_conv (fn ct =>
let
val thy = Thm.theory_of_cterm ct;
in lift_triv_classes_conv thy (Code_Thingol.eval_conv thy (K (norm_oracle thy))) ct end);
--- a/src/ZF/ZF.thy Mon Jun 21 11:37:25 2010 +0200
+++ b/src/ZF/ZF.thy Mon Jun 21 17:41:57 2010 +0200
@@ -199,10 +199,7 @@
finalconsts
0 Pow Inf Union PrimReplace mem
-defs
-(*don't try to use constdefs: the declaration order is tightly constrained*)
-
- (* Bounded Quantifiers *)
+defs (* Bounded Quantifiers *)
Ball_def: "Ball(A, P) == \<forall>x. x\<in>A --> P(x)"
Bex_def: "Bex(A, P) == \<exists>x. x\<in>A & P(x)"