# HG changeset patch # User huffman # Date 1240842377 25200 # Node ID 7c871a9cf6f4fc7aabec476ad66b46ee524ece4a # Parent 644d18da3c7785a66eb74cce7b78743d632d4256# Parent ed7364584aa7119a71c2d39ab02916df169a1d3f merged diff -r 644d18da3c77 -r 7c871a9cf6f4 CONTRIBUTORS --- a/CONTRIBUTORS Wed Apr 22 11:00:25 2009 -0700 +++ b/CONTRIBUTORS Mon Apr 27 07:26:17 2009 -0700 @@ -7,6 +7,10 @@ Contributions to this Isabelle version -------------------------------------- + +Contributions to Isabelle2009 +----------------------------- + * March 2009: Robert Himmelmann, TUM and Amine Chaieb, University of Cambridge Elementary topology in Euclidean space. diff -r 644d18da3c77 -r 7c871a9cf6f4 NEWS --- a/NEWS Wed Apr 22 11:00:25 2009 -0700 +++ b/NEWS Mon Apr 27 07:26:17 2009 -0700 @@ -4,6 +4,26 @@ New in this Isabelle version ---------------------------- +*** Pure *** + +* On instantiation of classes, remaining undefined class parameters are +formally declared. INCOMPATIBILITY. + + +*** HOL *** + +* Class semiring_div requires superclass no_zero_divisors and proof of div_mult_mult1; +theorems div_mult_mult1, div_mult_mult2, div_mult_mult1_if, div_mult_mult1 and +div_mult_mult2 have been generalized to class semiring_div, subsuming former +theorems zdiv_zmult_zmult1, zdiv_zmult_zmult1_if, zdiv_zmult_zmult1 and zdiv_zmult_zmult2. +div_mult_mult1 is now [simp] by default. INCOMPATIBILITY. + +* Power operations on relations and functions are now one dedicate constant compow with +infix syntax "^^". Power operations on multiplicative monoids retains syntax "^" +and is now defined generic in class power. INCOMPATIBILITY. + +* ML antiquotation @{code_datatype} inserts definition of a datatype generated +by the code generator; see Predicate.thy for an example. New in Isabelle2009 (April 2009) @@ -187,7 +207,7 @@ * Keyword 'code_exception' now named 'code_abort'. INCOMPATIBILITY. -* Unified theorem tables for both code code generators. Thus [code +* Unified theorem tables for both code generators. Thus [code func] has disappeared and only [code] remains. INCOMPATIBILITY. * Command 'find_consts' searches for constants based on type and name diff -r 644d18da3c77 -r 7c871a9cf6f4 contrib/SystemOnTPTP/remote --- a/contrib/SystemOnTPTP/remote Wed Apr 22 11:00:25 2009 -0700 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,120 +0,0 @@ -#!/usr/bin/env perl -# -# Wrapper for custom remote provers on SystemOnTPTP -# Author: Fabian Immler, TU Muenchen -# - -use warnings; -use strict; -use Getopt::Std; -use HTTP::Request::Common; -use LWP; - -my $SystemOnTPTPFormReplyURL = "http://www.cs.miami.edu/~tptp/cgi-bin/SystemOnTPTPFormReply"; - -# default parameters -my %URLParameters = ( - "NoHTML" => 1, - "QuietFlag" => "-q01", - "X2TPTP" => "-S", - "SubmitButton" => "RunSelectedSystems", - "ProblemSource" => "UPLOAD", - ); - -#----Get format and transform options if specified -my %Options; -getopts("hws:t:c:",\%Options); - -#----Usage -sub usage() { - print("Usage: remote [] \n"); - print(" are ...\n"); - print(" -h - print this help\n"); - print(" -w - list available ATP systems\n"); - print(" -s - specified system to use\n"); - print(" -t - CPU time limit for system\n"); - print(" -c - custom command for system\n"); - print(" - TPTP problem file\n"); - exit(0); -} -if (exists($Options{'h'})) { - usage(); -} -#----What systems flag -if (exists($Options{'w'})) { - $URLParameters{"SubmitButton"} = "ListSystems"; - delete($URLParameters{"ProblemSource"}); -} -#----Selected system -my $System; -if (exists($Options{'s'})) { - $System = $Options{'s'}; -} else { - # use Vampire as default - $System = "Vampire---9.0"; -} -$URLParameters{"System___$System"} = $System; - -#----Time limit -if (exists($Options{'t'})) { - $URLParameters{"TimeLimit___$System"} = $Options{'t'}; -} -#----Custom command -if (exists($Options{'c'})) { - $URLParameters{"Command___$System"} = $Options{'c'}; -} - -#----Get single file name -if (exists($URLParameters{"ProblemSource"})) { - if (scalar(@ARGV) >= 1) { - $URLParameters{"UPLOADProblem"} = [shift(@ARGV)]; - } else { - print("Missing problem file\n"); - usage(); - die; - } -} - -# Query Server -my $Agent = LWP::UserAgent->new; -if (exists($Options{'t'})) { - # give server more time to respond - $Agent->timeout($Options{'t'} + 10); -} -my $Request = POST($SystemOnTPTPFormReplyURL, - Content_Type => 'form-data',Content => \%URLParameters); -my $Response = $Agent->request($Request); - -#catch errors / failure -if(! $Response->is_success){ - print "HTTP-Error: " . $Response->message . "\n"; - exit(-1); -} elsif (exists($Options{'w'})) { - print $Response->content; - exit (0); -} elsif ($Response->content =~ /WARNING: (\S*) does not exist/) { - print "Specified System $1 does not exist\n"; - exit(-1); -} elsif ($Response->content =~ /%\s*Result\s*:\s*Unsatisfiable.*\n%\s*Output\s*:\s*(CNF)?Refutation.*\n%/) { - my @lines = split( /\n/, $Response->content); - my $extract = ""; - foreach my $line (@lines){ - #ignore comments - if ($line !~ /^%/ && !($line eq "")) { - $extract .= "$line"; - } - } - # insert newlines after ').' - $extract =~ s/\s//g; - $extract =~ s/\)\.cnf/\)\.\ncnf/g; - - # orientation for res_reconstruct.ML - print "# SZS output start CNFRefutation.\n"; - print "$extract\n"; - print "# SZS output end CNFRefutation.\n"; - exit(0); -} else { - print "Remote-script could not extract proof:\n".$Response->content; - exit(-1); -} - diff -r 644d18da3c77 -r 7c871a9cf6f4 doc-src/Codegen/Thy/Program.thy --- a/doc-src/Codegen/Thy/Program.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/doc-src/Codegen/Thy/Program.thy Mon Apr 27 07:26:17 2009 -0700 @@ -323,7 +323,7 @@ *} -subsection {* Equality and wellsortedness *} +subsection {* Equality *} text {* Surely you have already noticed how equality is treated @@ -358,60 +358,7 @@ manually like any other type class. Though this @{text eq} class is designed to get rarely in - the way, a subtlety - enters the stage when definitions of overloaded constants - are dependent on operational equality. For example, let - us define a lexicographic ordering on tuples - (also see theory @{theory Product_ord}): -*} - -instantiation %quote "*" :: (order, order) order -begin - -definition %quote [code del]: - "x \ y \ fst x < fst y \ fst x = fst y \ snd x \ snd y" - -definition %quote [code del]: - "x < y \ fst x < fst y \ fst x = fst y \ snd x < snd y" - -instance %quote proof -qed (auto simp: less_eq_prod_def less_prod_def intro: order_less_trans) - -end %quote - -lemma %quote order_prod [code]: - "(x1 \ 'a\order, y1 \ 'b\order) < (x2, y2) \ - x1 < x2 \ x1 = x2 \ y1 < y2" - "(x1 \ 'a\order, y1 \ 'b\order) \ (x2, y2) \ - x1 < x2 \ x1 = x2 \ y1 \ y2" - by (simp_all add: less_prod_def less_eq_prod_def) - -text {* - \noindent Then code generation will fail. Why? The definition - of @{term "op \"} depends on equality on both arguments, - which are polymorphic and impose an additional @{class eq} - class constraint, which the preprocessor does not propagate - (for technical reasons). - - The solution is to add @{class eq} explicitly to the first sort arguments in the - code theorems: -*} - -lemma %quote order_prod_code [code]: - "(x1 \ 'a\{order, eq}, y1 \ 'b\order) < (x2, y2) \ - x1 < x2 \ x1 = x2 \ y1 < y2" - "(x1 \ 'a\{order, eq}, y1 \ 'b\order) \ (x2, y2) \ - x1 < x2 \ x1 = x2 \ y1 \ y2" - by (simp_all add: less_prod_def less_eq_prod_def) - -text {* - \noindent Then code generation succeeds: -*} - -text %quote {*@{code_stmts "op \ \ _ \ _ \ _ \ _ \ bool" (SML)}*} - -text {* - In some cases, the automatically derived code equations + the way, in some cases the automatically derived code equations for equality on a particular type may not be appropriate. As example, watch the following datatype representing monomorphic parametric types (where type constructors diff -r 644d18da3c77 -r 7c871a9cf6f4 doc-src/Codegen/Thy/document/Program.tex --- a/doc-src/Codegen/Thy/document/Program.tex Wed Apr 22 11:00:25 2009 -0700 +++ b/doc-src/Codegen/Thy/document/Program.tex Mon Apr 27 07:26:17 2009 -0700 @@ -714,7 +714,7 @@ \end{isamarkuptext}% \isamarkuptrue% % -\isamarkupsubsection{Equality and wellsortedness% +\isamarkupsubsection{Equality% } \isamarkuptrue% % @@ -801,141 +801,7 @@ manually like any other type class. Though this \isa{eq} class is designed to get rarely in - the way, a subtlety - enters the stage when definitions of overloaded constants - are dependent on operational equality. For example, let - us define a lexicographic ordering on tuples - (also see theory \hyperlink{theory.Product-ord}{\mbox{\isa{Product{\isacharunderscore}ord}}}):% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{instantiation}\isamarkupfalse% -\ {\isachardoublequoteopen}{\isacharasterisk}{\isachardoublequoteclose}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}order{\isacharcomma}\ order{\isacharparenright}\ order\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ {\isacharbrackleft}code\ del{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}x\ {\isasymle}\ y\ {\isasymlongleftrightarrow}\ fst\ x\ {\isacharless}\ fst\ y\ {\isasymor}\ fst\ x\ {\isacharequal}\ fst\ y\ {\isasymand}\ snd\ x\ {\isasymle}\ snd\ y{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\ {\isacharbrackleft}code\ del{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}x\ {\isacharless}\ y\ {\isasymlongleftrightarrow}\ fst\ x\ {\isacharless}\ fst\ y\ {\isasymor}\ fst\ x\ {\isacharequal}\ fst\ y\ {\isasymand}\ snd\ x\ {\isacharless}\ snd\ y{\isachardoublequoteclose}\isanewline -\isanewline -\isacommand{instance}\isamarkupfalse% -\ \isacommand{proof}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -\ {\isacharparenleft}auto\ simp{\isacharcolon}\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}prod{\isacharunderscore}def\ intro{\isacharcolon}\ order{\isacharunderscore}less{\isacharunderscore}trans{\isacharparenright}\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ order{\isacharunderscore}prod\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}order{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}order{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ less{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def{\isacharparenright}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Then code generation will fail. Why? The definition - of \isa{op\ {\isasymle}} depends on equality on both arguments, - which are polymorphic and impose an additional \isa{eq} - class constraint, which the preprocessor does not propagate - (for technical reasons). - - The solution is to add \isa{eq} explicitly to the first sort arguments in the - code theorems:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -\isacommand{lemma}\isamarkupfalse% -\ order{\isacharunderscore}prod{\isacharunderscore}code\ {\isacharbrackleft}code{\isacharbrackright}{\isacharcolon}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}order{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isacharless}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isacharless}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline -\ \ {\isachardoublequoteopen}{\isacharparenleft}x{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}a{\isasymColon}{\isacharbraceleft}order{\isacharcomma}\ eq{\isacharbraceright}{\isacharcomma}\ y{\isadigit{1}}\ {\isasymColon}\ {\isacharprime}b{\isasymColon}order{\isacharparenright}\ {\isasymle}\ {\isacharparenleft}x{\isadigit{2}}{\isacharcomma}\ y{\isadigit{2}}{\isacharparenright}\ {\isasymlongleftrightarrow}\isanewline -\ \ \ \ \ x{\isadigit{1}}\ {\isacharless}\ x{\isadigit{2}}\ {\isasymor}\ x{\isadigit{1}}\ {\isacharequal}\ x{\isadigit{2}}\ {\isasymand}\ y{\isadigit{1}}\ {\isasymle}\ y{\isadigit{2}}{\isachardoublequoteclose}\isanewline -\ \ \isacommand{by}\isamarkupfalse% -\ {\isacharparenleft}simp{\isacharunderscore}all\ add{\isacharcolon}\ less{\isacharunderscore}prod{\isacharunderscore}def\ less{\isacharunderscore}eq{\isacharunderscore}prod{\isacharunderscore}def{\isacharparenright}% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -\noindent Then code generation succeeds:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimquote -% -\endisadelimquote -% -\isatagquote -% -\begin{isamarkuptext}% -\isatypewriter% -\noindent% -\hspace*{0pt}structure Example = \\ -\hspace*{0pt}struct\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a eq = {\char123}eq :~'a -> 'a -> bool{\char125};\\ -\hspace*{0pt}fun eq (A{\char95}:'a eq) = {\char35}eq A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a ord = {\char123}less{\char95}eq :~'a -> 'a -> bool,~less :~'a -> 'a -> bool{\char125};\\ -\hspace*{0pt}fun less{\char95}eq (A{\char95}:'a ord) = {\char35}less{\char95}eq A{\char95};\\ -\hspace*{0pt}fun less (A{\char95}:'a ord) = {\char35}less A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun eqop A{\char95}~a b = eq A{\char95}~a b;\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a preorder = {\char123}Orderings{\char95}{\char95}ord{\char95}preorder :~'a ord{\char125};\\ -\hspace*{0pt}fun ord{\char95}preorder (A{\char95}:'a preorder) = {\char35}Orderings{\char95}{\char95}ord{\char95}preorder A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}type 'a order = {\char123}Orderings{\char95}{\char95}preorder{\char95}order :~'a preorder{\char125};\\ -\hspace*{0pt}fun preorder{\char95}order (A{\char95}:'a order) = {\char35}Orderings{\char95}{\char95}preorder{\char95}order A{\char95};\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun less{\char95}eqa (A1{\char95},~A2{\char95}) B{\char95}~(x1,~y1) (x2,~y2) =\\ -\hspace*{0pt} ~less ((ord{\char95}preorder o preorder{\char95}order) A2{\char95}) x1 x2 orelse\\ -\hspace*{0pt} ~~~eqop A1{\char95}~x1 x2 andalso\\ -\hspace*{0pt} ~~~~~less{\char95}eq ((ord{\char95}preorder o preorder{\char95}order) B{\char95}) y1 y2\\ -\hspace*{0pt} ~| less{\char95}eqa (A1{\char95},~A2{\char95}) B{\char95}~(x1,~y1) (x2,~y2) =\\ -\hspace*{0pt} ~~~less ((ord{\char95}preorder o preorder{\char95}order) A2{\char95}) x1 x2 orelse\\ -\hspace*{0pt} ~~~~~eqop A1{\char95}~x1 x2 andalso\\ -\hspace*{0pt} ~~~~~~~less{\char95}eq ((ord{\char95}preorder o preorder{\char95}order) B{\char95}) y1 y2;\\ -\hspace*{0pt}\\ -\hspace*{0pt}end;~(*struct Example*)% -\end{isamarkuptext}% -\isamarkuptrue% -% -\endisatagquote -{\isafoldquote}% -% -\isadelimquote -% -\endisadelimquote -% -\begin{isamarkuptext}% -In some cases, the automatically derived code equations + the way, in some cases the automatically derived code equations for equality on a particular type may not be appropriate. As example, watch the following datatype representing monomorphic parametric types (where type constructors diff -r 644d18da3c77 -r 7c871a9cf6f4 doc-src/Main/Docs/Main_Doc.thy --- a/doc-src/Main/Docs/Main_Doc.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/doc-src/Main/Docs/Main_Doc.thy Mon Apr 27 07:26:17 2009 -0700 @@ -268,6 +268,7 @@ @{const Transitive_Closure.rtrancl} & @{term_type_only Transitive_Closure.rtrancl "('a*'a)set\('a*'a)set"}\\ @{const Transitive_Closure.trancl} & @{term_type_only Transitive_Closure.trancl "('a*'a)set\('a*'a)set"}\\ @{const Transitive_Closure.reflcl} & @{term_type_only Transitive_Closure.reflcl "('a*'a)set\('a*'a)set"}\\ +@{const compower} & @{term_type_only "op ^^ :: ('a*'a)set\nat\('a*'a)set" "('a*'a)set\nat\('a*'a)set"}\\ \end{tabular} \subsubsection*{Syntax} @@ -318,7 +319,6 @@ @{term "op + :: nat \ nat \ nat"} & @{term "op - :: nat \ nat \ nat"} & @{term "op * :: nat \ nat \ nat"} & -@{term "op ^ :: nat \ nat \ nat"} & @{term "op div :: nat \ nat \ nat"}& @{term "op mod :: nat \ nat \ nat"}& @{term "op dvd :: nat \ nat \ bool"}\\ @@ -331,7 +331,9 @@ \end{tabular} \begin{tabular}{@ {} l @ {~::~} l @ {}} -@{const Nat.of_nat} & @{typeof Nat.of_nat} +@{const Nat.of_nat} & @{typeof Nat.of_nat}\\ +@{term "op ^^ :: ('a \ 'a) \ nat \ 'a \ 'a"} & + @{term_type_only "op ^^ :: ('a \ 'a) \ nat \ 'a \ 'a" "('a \ 'a) \ nat \ 'a \ 'a"} \end{tabular} \section{Int} @@ -450,14 +452,6 @@ \end{tabular} -\section{Iterated Functions and Relations} - -Theory: @{theory Relation_Power} - -Iterated functions \ @{term[source]"(f::'a\'a) ^ n"} \ -and relations \ @{term[source]"(r::('a\'a)set) ^ n"}. - - \section{Option} @{datatype option} diff -r 644d18da3c77 -r 7c871a9cf6f4 doc-src/Main/Docs/document/Main_Doc.tex --- a/doc-src/Main/Docs/document/Main_Doc.tex Wed Apr 22 11:00:25 2009 -0700 +++ b/doc-src/Main/Docs/document/Main_Doc.tex Mon Apr 27 07:26:17 2009 -0700 @@ -279,6 +279,7 @@ \isa{rtrancl} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set}\\ \isa{trancl} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set}\\ \isa{reflcl} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set}\\ +\isa{op\ {\isacharcircum}{\isacharcircum}} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set}\\ \end{tabular} \subsubsection*{Syntax} @@ -328,7 +329,6 @@ \isa{op\ {\isacharplus}} & \isa{op\ {\isacharminus}} & \isa{op\ {\isacharasterisk}} & -\isa{op\ {\isacharcircum}} & \isa{op\ div}& \isa{op\ mod}& \isa{op\ dvd}\\ @@ -341,7 +341,9 @@ \end{tabular} \begin{tabular}{@ {} l @ {~::~} l @ {}} -\isa{of{\isacharunderscore}nat} & \isa{nat\ {\isasymRightarrow}\ {\isacharprime}a} +\isa{of{\isacharunderscore}nat} & \isa{nat\ {\isasymRightarrow}\ {\isacharprime}a}\\ +\isa{op\ {\isacharcircum}{\isacharcircum}} & + \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a} \end{tabular} \section{Int} @@ -460,14 +462,6 @@ \end{tabular} -\section{Iterated Functions and Relations} - -Theory: \isa{Relation{\isacharunderscore}Power} - -Iterated functions \ \isa{{\isachardoublequote}{\isacharparenleft}f{\isacharcolon}{\isacharcolon}{\isacharprime}a{\isasymRightarrow}{\isacharprime}a{\isacharparenright}\ {\isacharcircum}\ n{\isachardoublequote}} \ -and relations \ \isa{{\isachardoublequote}{\isacharparenleft}r{\isacharcolon}{\isacharcolon}{\isacharparenleft}{\isacharprime}a{\isasymtimes}{\isacharprime}a{\isacharparenright}set{\isacharparenright}\ {\isacharcircum}\ n{\isachardoublequote}}. - - \section{Option} \isa{\isacommand{datatype}\ {\isacharprime}a\ option\ {\isacharequal}\ None\ {\isacharbar}\ Some\ {\isacharprime}a} diff -r 644d18da3c77 -r 7c871a9cf6f4 doc-src/TutorialI/tutorial.tex --- a/doc-src/TutorialI/tutorial.tex Wed Apr 22 11:00:25 2009 -0700 +++ b/doc-src/TutorialI/tutorial.tex Mon Apr 27 07:26:17 2009 -0700 @@ -39,10 +39,11 @@ %University of Cambridge\\ %Computer Laboratory } +\pagenumbering{roman} \maketitle +\newpage -\pagenumbering{roman} -\setcounter{page}{5} +%\setcounter{page}{5} %\vspace*{\fill} %\begin{center} %\LARGE In memoriam \\[1ex] @@ -52,6 +53,7 @@ %\vspace*{\fill} %\vspace*{\fill} %\newpage + \include{preface} \tableofcontents diff -r 644d18da3c77 -r 7c871a9cf6f4 doc-src/more_antiquote.ML --- a/doc-src/more_antiquote.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/doc-src/more_antiquote.ML Mon Apr 27 07:26:17 2009 -0700 @@ -88,7 +88,7 @@ let val thy = ProofContext.theory_of ctxt; val const = Code_Unit.check_const thy raw_const; - val (_, funcgr) = Code_Wellsorted.make thy [const]; + val (_, funcgr) = Code_Wellsorted.obtain thy [const] []; fun holize thm = @{thm meta_eq_to_obj_eq} OF [thm]; val thms = Code_Wellsorted.eqns funcgr const |> map_filter (fn (thm, linear) => if linear then SOME thm else NONE) diff -r 644d18da3c77 -r 7c871a9cf6f4 lib/scripts/SystemOnTPTP --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/lib/scripts/SystemOnTPTP Mon Apr 27 07:26:17 2009 -0700 @@ -0,0 +1,120 @@ +#!/usr/bin/env perl +# +# Wrapper for custom remote provers on SystemOnTPTP +# Author: Fabian Immler, TU Muenchen +# + +use warnings; +use strict; +use Getopt::Std; +use HTTP::Request::Common; +use LWP; + +my $SystemOnTPTPFormReplyURL = "http://www.cs.miami.edu/~tptp/cgi-bin/SystemOnTPTPFormReply"; + +# default parameters +my %URLParameters = ( + "NoHTML" => 1, + "QuietFlag" => "-q01", + "X2TPTP" => "-S", + "SubmitButton" => "RunSelectedSystems", + "ProblemSource" => "UPLOAD", + ); + +#----Get format and transform options if specified +my %Options; +getopts("hws:t:c:",\%Options); + +#----Usage +sub usage() { + print("Usage: remote [] \n"); + print(" are ...\n"); + print(" -h - print this help\n"); + print(" -w - list available ATP systems\n"); + print(" -s - specified system to use\n"); + print(" -t - CPU time limit for system\n"); + print(" -c - custom command for system\n"); + print(" - TPTP problem file\n"); + exit(0); +} +if (exists($Options{'h'})) { + usage(); +} +#----What systems flag +if (exists($Options{'w'})) { + $URLParameters{"SubmitButton"} = "ListSystems"; + delete($URLParameters{"ProblemSource"}); +} +#----Selected system +my $System; +if (exists($Options{'s'})) { + $System = $Options{'s'}; +} else { + # use Vampire as default + $System = "Vampire---9.0"; +} +$URLParameters{"System___$System"} = $System; + +#----Time limit +if (exists($Options{'t'})) { + $URLParameters{"TimeLimit___$System"} = $Options{'t'}; +} +#----Custom command +if (exists($Options{'c'})) { + $URLParameters{"Command___$System"} = $Options{'c'}; +} + +#----Get single file name +if (exists($URLParameters{"ProblemSource"})) { + if (scalar(@ARGV) >= 1) { + $URLParameters{"UPLOADProblem"} = [shift(@ARGV)]; + } else { + print("Missing problem file\n"); + usage(); + die; + } +} + +# Query Server +my $Agent = LWP::UserAgent->new; +if (exists($Options{'t'})) { + # give server more time to respond + $Agent->timeout($Options{'t'} + 10); +} +my $Request = POST($SystemOnTPTPFormReplyURL, + Content_Type => 'form-data',Content => \%URLParameters); +my $Response = $Agent->request($Request); + +#catch errors / failure +if(! $Response->is_success){ + print "HTTP-Error: " . $Response->message . "\n"; + exit(-1); +} elsif (exists($Options{'w'})) { + print $Response->content; + exit (0); +} elsif ($Response->content =~ /WARNING: (\S*) does not exist/) { + print "Specified System $1 does not exist\n"; + exit(-1); +} elsif ($Response->content =~ /%\s*Result\s*:\s*Unsatisfiable.*\n%\s*Output\s*:\s*(CNF)?Refutation.*\n%/) { + my @lines = split( /\n/, $Response->content); + my $extract = ""; + foreach my $line (@lines){ + #ignore comments + if ($line !~ /^%/ && !($line eq "")) { + $extract .= "$line"; + } + } + # insert newlines after ').' + $extract =~ s/\s//g; + $extract =~ s/\)\.cnf/\)\.\ncnf/g; + + # orientation for res_reconstruct.ML + print "# SZS output start CNFRefutation.\n"; + print "$extract\n"; + print "# SZS output end CNFRefutation.\n"; + exit(0); +} else { + print "Remote-script could not extract proof:\n".$Response->content; + exit(-1); +} + diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Algebra/abstract/Ring2.thy --- a/src/HOL/Algebra/abstract/Ring2.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Algebra/abstract/Ring2.thy Mon Apr 27 07:26:17 2009 -0700 @@ -12,7 +12,7 @@ subsection {* Ring axioms *} -class ring = zero + one + plus + minus + uminus + times + inverse + power + Ring_and_Field.dvd + +class ring = zero + one + plus + minus + uminus + times + inverse + power + dvd + assumes a_assoc: "(a + b) + c = a + (b + c)" and l_zero: "0 + a = a" and l_neg: "(-a) + a = 0" @@ -28,8 +28,6 @@ assumes minus_def: "a - b = a + (-b)" and inverse_def: "inverse a = (if a dvd 1 then THE x. a*x = 1 else 0)" and divide_def: "a / b = a * inverse b" - and power_0 [simp]: "a ^ 0 = 1" - and power_Suc [simp]: "a ^ Suc n = a ^ n * a" begin definition assoc :: "'a \ 'a \ bool" (infixl "assoc" 50) where diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Algebra/poly/LongDiv.thy --- a/src/HOL/Algebra/poly/LongDiv.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Algebra/poly/LongDiv.thy Mon Apr 27 07:26:17 2009 -0700 @@ -1,6 +1,5 @@ (* Experimental theory: long division of polynomials - $Id$ Author: Clemens Ballarin, started 23 June 1999 *) @@ -133,9 +132,9 @@ delsimprocs [ring_simproc]) 1 *}) apply (tactic {* asm_simp_tac (@{simpset} delsimprocs [ring_simproc]) 1 *}) apply (tactic {* simp_tac (@{simpset} addsimps [thm "minus_def", thm "smult_r_distr", - thm "smult_r_minus", thm "monom_mult_smult", thm "smult_assoc1", thm "smult_assoc2"] + thm "smult_r_minus", thm "monom_mult_smult", thm "smult_assoc2"] delsimprocs [ring_simproc]) 1 *}) - apply simp + apply (simp add: smult_assoc1 [symmetric]) done ML {* diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Algebra/poly/UnivPoly2.thy --- a/src/HOL/Algebra/poly/UnivPoly2.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Algebra/poly/UnivPoly2.thy Mon Apr 27 07:26:17 2009 -0700 @@ -155,16 +155,6 @@ end -instantiation up :: ("{times, one, comm_monoid_add}") power -begin - -primrec power_up where - "(a \ 'a up) ^ 0 = 1" - | "(a \ 'a up) ^ Suc n = a ^ n * a" - -instance .. - -end subsection {* Effect of operations on coefficients *} @@ -328,8 +318,9 @@ qed show "(p + q) * r = p * r + q * r" by (rule up_eqI) simp - show "p * q = q * p" + show "\q. p * q = q * p" proof (rule up_eqI) + fix q fix n { fix k @@ -354,11 +345,11 @@ by (simp add: up_inverse_def) show "p / q = p * inverse q" by (simp add: up_divide_def) - fix n - show "p ^ 0 = 1" by simp - show "p ^ Suc n = p ^ n * p" by simp qed +instance up :: (ring) recpower proof +qed simp_all + (* Further properties of monom *) lemma monom_zero [simp]: diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Bali/Trans.thy --- a/src/HOL/Bali/Trans.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Bali/Trans.thy Mon Apr 27 07:26:17 2009 -0700 @@ -359,7 +359,7 @@ abbreviation stepn:: "[prog, term \ state,nat,term \ state] \ bool" ("_\_ \_ _"[61,82,82] 81) - where "G\p \n p' \ (p,p') \ {(x, y). step G x y}^n" + where "G\p \n p' \ (p,p') \ {(x, y). step G x y}^^n" abbreviation steptr:: "[prog,term \ state,term \ state] \ bool" ("_\_ \* _"[61,82,82] 81) @@ -370,25 +370,6 @@ Smallstep zu Bigstep, nur wenn nicht die Ausdrücke Callee, FinA ,\ *) -lemma rtrancl_imp_rel_pow: "p \ R^* \ \n. p \ R^n" -proof - - assume "p \ R\<^sup>*" - moreover obtain x y where p: "p = (x,y)" by (cases p) - ultimately have "(x,y) \ R\<^sup>*" by hypsubst - hence "\n. (x,y) \ R^n" - proof induct - fix a have "(a,a) \ R^0" by simp - thus "\n. (a,a) \ R ^ n" .. - next - fix a b c assume "\n. (a,b) \ R ^ n" - then obtain n where "(a,b) \ R^n" .. - moreover assume "(b,c) \ R" - ultimately have "(a,c) \ R^(Suc n)" by auto - thus "\n. (a,c) \ R^n" .. - qed - with p show ?thesis by hypsubst -qed - (* lemma imp_eval_trans: assumes eval: "G\s0 \t\\ (v,s1)" diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Code_Eval.thy --- a/src/HOL/Code_Eval.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Code_Eval.thy Mon Apr 27 07:26:17 2009 -0700 @@ -161,6 +161,7 @@ signature EVAL = sig val mk_term: ((string * typ) -> term) -> (typ -> term) -> term -> term + val mk_term_of: typ -> term -> term val eval_ref: (unit -> term) option ref val eval_term: theory -> term -> term end; @@ -175,8 +176,7 @@ fun eval_term thy t = t |> Eval.mk_term_of (fastype_of t) - |> (fn t => Code_ML.eval_term ("Eval.eval_ref", eval_ref) thy t []) - |> Code.postprocess_term thy; + |> (fn t => Code_ML.eval NONE ("Eval.eval_ref", eval_ref) I thy t []); end; *} diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Code_Setup.thy --- a/src/HOL/Code_Setup.thy Wed Apr 22 11:00:25 2009 -0700 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,253 +0,0 @@ -(* Title: HOL/Code_Setup.thy - ID: $Id$ - Author: Florian Haftmann -*) - -header {* Setup of code generators and related tools *} - -theory Code_Setup -imports HOL -begin - -subsection {* Generic code generator foundation *} - -text {* Datatypes *} - -code_datatype True False - -code_datatype "TYPE('a\{})" - -code_datatype Trueprop "prop" - -text {* Code equations *} - -lemma [code]: - shows "(True \ PROP P) \ PROP P" - and "(False \ Q) \ Trueprop True" - and "(PROP P \ True) \ Trueprop True" - and "(Q \ False) \ Trueprop (\ Q)" by (auto intro!: equal_intr_rule) - -lemma [code]: - shows "False \ x \ False" - and "True \ x \ x" - and "x \ False \ False" - and "x \ True \ x" by simp_all - -lemma [code]: - shows "False \ x \ x" - and "True \ x \ True" - and "x \ False \ x" - and "x \ True \ True" by simp_all - -lemma [code]: - shows "\ True \ False" - and "\ False \ True" by (rule HOL.simp_thms)+ - -lemmas [code] = Let_def if_True if_False - -lemmas [code, code unfold, symmetric, code post] = imp_conv_disj - -text {* Equality *} - -context eq -begin - -lemma equals_eq [code inline, code]: "op = \ eq" - by (rule eq_reflection) (rule ext, rule ext, rule sym, rule eq_equals) - -declare eq [code unfold, code inline del] - -declare equals_eq [symmetric, code post] - -end - -declare simp_thms(6) [code nbe] - -hide (open) const eq -hide const eq - -setup {* - Code_Unit.add_const_alias @{thm equals_eq} -*} - -text {* Cases *} - -lemma Let_case_cert: - assumes "CASE \ (\x. Let x f)" - shows "CASE x \ f x" - using assms by simp_all - -lemma If_case_cert: - assumes "CASE \ (\b. If b f g)" - shows "(CASE True \ f) &&& (CASE False \ g)" - using assms by simp_all - -setup {* - Code.add_case @{thm Let_case_cert} - #> Code.add_case @{thm If_case_cert} - #> Code.add_undefined @{const_name undefined} -*} - -code_abort undefined - - -subsection {* Generic code generator preprocessor *} - -setup {* - Code.map_pre (K HOL_basic_ss) - #> Code.map_post (K HOL_basic_ss) -*} - - -subsection {* Generic code generator target languages *} - -text {* type bool *} - -code_type bool - (SML "bool") - (OCaml "bool") - (Haskell "Bool") - -code_const True and False and Not and "op &" and "op |" and If - (SML "true" and "false" and "not" - and infixl 1 "andalso" and infixl 0 "orelse" - and "!(if (_)/ then (_)/ else (_))") - (OCaml "true" and "false" and "not" - and infixl 4 "&&" and infixl 2 "||" - and "!(if (_)/ then (_)/ else (_))") - (Haskell "True" and "False" and "not" - and infixl 3 "&&" and infixl 2 "||" - and "!(if (_)/ then (_)/ else (_))") - -code_reserved SML - bool true false not - -code_reserved OCaml - bool not - -text {* using built-in Haskell equality *} - -code_class eq - (Haskell "Eq") - -code_const "eq_class.eq" - (Haskell infixl 4 "==") - -code_const "op =" - (Haskell infixl 4 "==") - -text {* undefined *} - -code_const undefined - (SML "!(raise/ Fail/ \"undefined\")") - (OCaml "failwith/ \"undefined\"") - (Haskell "error/ \"undefined\"") - - -subsection {* SML code generator setup *} - -types_code - "bool" ("bool") -attach (term_of) {* -fun term_of_bool b = if b then HOLogic.true_const else HOLogic.false_const; -*} -attach (test) {* -fun gen_bool i = - let val b = one_of [false, true] - in (b, fn () => term_of_bool b) end; -*} - "prop" ("bool") -attach (term_of) {* -fun term_of_prop b = - HOLogic.mk_Trueprop (if b then HOLogic.true_const else HOLogic.false_const); -*} - -consts_code - "Trueprop" ("(_)") - "True" ("true") - "False" ("false") - "Not" ("Bool.not") - "op |" ("(_ orelse/ _)") - "op &" ("(_ andalso/ _)") - "If" ("(if _/ then _/ else _)") - -setup {* -let - -fun eq_codegen thy defs dep thyname b t gr = - (case strip_comb t of - (Const ("op =", Type (_, [Type ("fun", _), _])), _) => NONE - | (Const ("op =", _), [t, u]) => - let - val (pt, gr') = Codegen.invoke_codegen thy defs dep thyname false t gr; - val (pu, gr'') = Codegen.invoke_codegen thy defs dep thyname false u gr'; - val (_, gr''') = Codegen.invoke_tycodegen thy defs dep thyname false HOLogic.boolT gr''; - in - SOME (Codegen.parens - (Pretty.block [pt, Codegen.str " =", Pretty.brk 1, pu]), gr''') - end - | (t as Const ("op =", _), ts) => SOME (Codegen.invoke_codegen - thy defs dep thyname b (Codegen.eta_expand t ts 2) gr) - | _ => NONE); - -in - Codegen.add_codegen "eq_codegen" eq_codegen -end -*} - - -subsection {* Evaluation and normalization by evaluation *} - -setup {* - Value.add_evaluator ("SML", Codegen.eval_term o ProofContext.theory_of) -*} - -ML {* -structure Eval_Method = -struct - -val eval_ref : (unit -> bool) option ref = ref NONE; - -end; -*} - -oracle eval_oracle = {* fn ct => - let - val thy = Thm.theory_of_cterm ct; - val t = Thm.term_of ct; - val dummy = @{cprop True}; - in case try HOLogic.dest_Trueprop t - of SOME t' => if Code_ML.eval_term - ("Eval_Method.eval_ref", Eval_Method.eval_ref) thy t' [] - then Thm.capply (Thm.capply @{cterm "op \ \ prop \ prop \ prop"} ct) dummy - else dummy - | NONE => dummy - end -*} - -ML {* -fun gen_eval_method conv ctxt = SIMPLE_METHOD' - (CONVERSION (Conv.params_conv (~1) (K (Conv.concl_conv (~1) conv)) ctxt) - THEN' rtac TrueI) -*} - -method_setup eval = {* Scan.succeed (gen_eval_method eval_oracle) *} - "solve goal by evaluation" - -method_setup evaluation = {* Scan.succeed (gen_eval_method Codegen.evaluation_conv) *} - "solve goal by evaluation" - -method_setup normalization = {* - Scan.succeed (K (SIMPLE_METHOD' (CONVERSION Nbe.norm_conv THEN' (fn k => TRY (rtac TrueI k))))) -*} "solve goal by normalization" - - -subsection {* Quickcheck *} - -setup {* - Quickcheck.add_generator ("SML", Codegen.test_term) -*} - -quickcheck_params [size = 5, iterations = 50] - -end diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Complex.thy --- a/src/HOL/Complex.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Complex.thy Mon Apr 27 07:26:17 2009 -0700 @@ -159,19 +159,7 @@ subsection {* Exponentiation *} -instantiation complex :: recpower -begin - -primrec power_complex where - "z ^ 0 = (1\complex)" -| "z ^ Suc n = (z\complex) * z ^ n" - -instance proof -qed simp_all - -declare power_complex.simps [simp del] - -end +instance complex :: recpower .. subsection {* Numerals and Arithmetic *} diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Decision_Procs/Approximation.thy --- a/src/HOL/Decision_Procs/Approximation.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Decision_Procs/Approximation.thy Mon Apr 27 07:26:17 2009 -0700 @@ -23,8 +23,8 @@ qed lemma horner_schema: fixes f :: "nat \ nat" and G :: "nat \ nat \ nat" and F :: "nat \ nat" - assumes f_Suc: "\n. f (Suc n) = G ((F^n) s) (f n)" - shows "horner F G n ((F^j') s) (f j') x = (\ j = 0..< n. -1^j * (1 / real (f (j' + j))) * x^j)" + assumes f_Suc: "\n. f (Suc n) = G ((F ^^ n) s) (f n)" + shows "horner F G n ((F ^^ j') s) (f j') x = (\ j = 0..< n. -1 ^ j * (1 / real (f (j' + j))) * x ^ j)" proof (induct n arbitrary: i k j') case (Suc n) @@ -33,13 +33,13 @@ qed auto lemma horner_bounds': - assumes "0 \ Ifloat x" and f_Suc: "\n. f (Suc n) = G ((F^n) s) (f n)" + assumes "0 \ Ifloat x" and f_Suc: "\n. f (Suc n) = G ((F ^^ n) s) (f n)" and lb_0: "\ i k x. lb 0 i k x = 0" and lb_Suc: "\ n i k x. lb (Suc n) i k x = lapprox_rat prec 1 (int k) - x * (ub n (F i) (G i k) x)" and ub_0: "\ i k x. ub 0 i k x = 0" and ub_Suc: "\ n i k x. ub (Suc n) i k x = rapprox_rat prec 1 (int k) - x * (lb n (F i) (G i k) x)" - shows "Ifloat (lb n ((F^j') s) (f j') x) \ horner F G n ((F^j') s) (f j') (Ifloat x) \ - horner F G n ((F^j') s) (f j') (Ifloat x) \ Ifloat (ub n ((F^j') s) (f j') x)" + shows "Ifloat (lb n ((F ^^ j') s) (f j') x) \ horner F G n ((F ^^ j') s) (f j') (Ifloat x) \ + horner F G n ((F ^^ j') s) (f j') (Ifloat x) \ Ifloat (ub n ((F ^^ j') s) (f j') x)" (is "?lb n j' \ ?horner n j' \ ?horner n j' \ ?ub n j'") proof (induct n arbitrary: j') case 0 thus ?case unfolding lb_0 ub_0 horner.simps by auto @@ -49,15 +49,15 @@ proof (rule add_mono) show "Ifloat (lapprox_rat prec 1 (int (f j'))) \ 1 / real (f j')" using lapprox_rat[of prec 1 "int (f j')"] by auto from Suc[where j'="Suc j'", unfolded funpow.simps comp_def f_Suc, THEN conjunct2] `0 \ Ifloat x` - show "- Ifloat (x * ub n (F ((F ^ j') s)) (G ((F ^ j') s) (f j')) x) \ - (Ifloat x * horner F G n (F ((F ^ j') s)) (G ((F ^ j') s) (f j')) (Ifloat x))" + show "- Ifloat (x * ub n (F ((F ^^ j') s)) (G ((F ^^ j') s) (f j')) x) \ - (Ifloat x * horner F G n (F ((F ^^ j') s)) (G ((F ^^ j') s) (f j')) (Ifloat x))" unfolding Ifloat_mult neg_le_iff_le by (rule mult_left_mono) qed moreover have "?horner (Suc n) j' \ ?ub (Suc n) j'" unfolding ub_Suc ub_Suc horner.simps Ifloat_sub diff_def proof (rule add_mono) show "1 / real (f j') \ Ifloat (rapprox_rat prec 1 (int (f j')))" using rapprox_rat[of 1 "int (f j')" prec] by auto from Suc[where j'="Suc j'", unfolded funpow.simps comp_def f_Suc, THEN conjunct1] `0 \ Ifloat x` - show "- (Ifloat x * horner F G n (F ((F ^ j') s)) (G ((F ^ j') s) (f j')) (Ifloat x)) \ - - Ifloat (x * lb n (F ((F ^ j') s)) (G ((F ^ j') s) (f j')) x)" + show "- (Ifloat x * horner F G n (F ((F ^^ j') s)) (G ((F ^^ j') s) (f j')) (Ifloat x)) \ + - Ifloat (x * lb n (F ((F ^^ j') s)) (G ((F ^^ j') s) (f j')) x)" unfolding Ifloat_mult neg_le_iff_le by (rule mult_left_mono) qed ultimately show ?case by blast @@ -73,13 +73,13 @@ *} lemma horner_bounds: fixes F :: "nat \ nat" and G :: "nat \ nat \ nat" - assumes "0 \ Ifloat x" and f_Suc: "\n. f (Suc n) = G ((F^n) s) (f n)" + assumes "0 \ Ifloat x" and f_Suc: "\n. f (Suc n) = G ((F ^^ n) s) (f n)" and lb_0: "\ i k x. lb 0 i k x = 0" and lb_Suc: "\ n i k x. lb (Suc n) i k x = lapprox_rat prec 1 (int k) - x * (ub n (F i) (G i k) x)" and ub_0: "\ i k x. ub 0 i k x = 0" and ub_Suc: "\ n i k x. ub (Suc n) i k x = rapprox_rat prec 1 (int k) - x * (lb n (F i) (G i k) x)" - shows "Ifloat (lb n ((F^j') s) (f j') x) \ (\j=0..j=0.. Ifloat (ub n ((F^j') s) (f j') x)" (is "?ub") + shows "Ifloat (lb n ((F ^^ j') s) (f j') x) \ (\j=0..j=0.. Ifloat (ub n ((F ^^ j') s) (f j') x)" (is "?ub") proof - have "?lb \ ?ub" using horner_bounds'[where lb=lb, OF `0 \ Ifloat x` f_Suc lb_0 lb_Suc ub_0 ub_Suc] @@ -88,29 +88,29 @@ qed lemma horner_bounds_nonpos: fixes F :: "nat \ nat" and G :: "nat \ nat \ nat" - assumes "Ifloat x \ 0" and f_Suc: "\n. f (Suc n) = G ((F^n) s) (f n)" + assumes "Ifloat x \ 0" and f_Suc: "\n. f (Suc n) = G ((F ^^ n) s) (f n)" and lb_0: "\ i k x. lb 0 i k x = 0" and lb_Suc: "\ n i k x. lb (Suc n) i k x = lapprox_rat prec 1 (int k) + x * (ub n (F i) (G i k) x)" and ub_0: "\ i k x. ub 0 i k x = 0" and ub_Suc: "\ n i k x. ub (Suc n) i k x = rapprox_rat prec 1 (int k) + x * (lb n (F i) (G i k) x)" - shows "Ifloat (lb n ((F^j') s) (f j') x) \ (\j=0..j=0.. Ifloat (ub n ((F^j') s) (f j') x)" (is "?ub") + shows "Ifloat (lb n ((F ^^ j') s) (f j') x) \ (\j=0..j=0.. Ifloat (ub n ((F ^^ j') s) (f j') x)" (is "?ub") proof - { fix x y z :: float have "x - y * z = x + - y * z" - by (cases x, cases y, cases z, simp add: plus_float.simps minus_float.simps uminus_float.simps times_float.simps algebra_simps) + by (cases x, cases y, cases z, simp add: plus_float.simps minus_float_def uminus_float.simps times_float.simps algebra_simps) } note diff_mult_minus = this { fix x :: float have "- (- x) = x" by (cases x, auto simp add: uminus_float.simps) } note minus_minus = this have move_minus: "Ifloat (-x) = -1 * Ifloat x" by auto - have sum_eq: "(\j=0..j=0..j = 0.. {0 ..< n}" show "1 / real (f (j' + j)) * Ifloat x ^ j = -1 ^ j * (1 / real (f (j' + j))) * Ifloat (- x) ^ j" unfolding move_minus power_mult_distrib real_mult_assoc[symmetric] - unfolding real_mult_commute unfolding real_mult_assoc[of "-1^j", symmetric] power_mult_distrib[symmetric] + unfolding real_mult_commute unfolding real_mult_assoc[of "-1 ^ j", symmetric] power_mult_distrib[symmetric] by auto qed @@ -160,21 +160,21 @@ else (0, (max (-l) u) ^ n))" lemma float_power_bnds: assumes "(l1, u1) = float_power_bnds n l u" and "x \ {Ifloat l .. Ifloat u}" - shows "x^n \ {Ifloat l1..Ifloat u1}" + shows "x ^ n \ {Ifloat l1..Ifloat u1}" proof (cases "even n") case True show ?thesis proof (cases "0 < l") case True hence "odd n \ 0 < l" and "0 \ Ifloat l" unfolding less_float_def by auto have u1: "u1 = u ^ n" and l1: "l1 = l ^ n" using assms unfolding float_power_bnds_def if_P[OF `odd n \ 0 < l`] by auto - have "Ifloat l^n \ x^n" and "x^n \ Ifloat u^n " using `0 \ Ifloat l` and assms unfolding atLeastAtMost_iff using power_mono[of "Ifloat l" x] power_mono[of x "Ifloat u"] by auto + have "Ifloat l ^ n \ x ^ n" and "x ^ n \ Ifloat u ^ n " using `0 \ Ifloat l` and assms unfolding atLeastAtMost_iff using power_mono[of "Ifloat l" x] power_mono[of x "Ifloat u"] by auto thus ?thesis using assms `0 < l` unfolding atLeastAtMost_iff l1 u1 float_power less_float_def by auto next case False hence P: "\ (odd n \ 0 < l)" using `even n` by auto show ?thesis proof (cases "u < 0") case True hence "0 \ - Ifloat u" and "- Ifloat u \ - x" and "0 \ - x" and "-x \ - Ifloat l" using assms unfolding less_float_def by auto - hence "Ifloat u^n \ x^n" and "x^n \ Ifloat l^n" using power_mono[of "-x" "-Ifloat l" n] power_mono[of "-Ifloat u" "-x" n] + hence "Ifloat u ^ n \ x ^ n" and "x ^ n \ Ifloat l ^ n" using power_mono[of "-x" "-Ifloat l" n] power_mono[of "-Ifloat u" "-x" n] unfolding power_minus_even[OF `even n`] by auto moreover have u1: "u1 = l ^ n" and l1: "l1 = u ^ n" using assms unfolding float_power_bnds_def if_not_P[OF P] if_P[OF True] by auto ultimately show ?thesis using float_power by auto @@ -194,11 +194,11 @@ next case False hence "odd n \ 0 < l" by auto have u1: "u1 = u ^ n" and l1: "l1 = l ^ n" using assms unfolding float_power_bnds_def if_P[OF `odd n \ 0 < l`] by auto - have "Ifloat l^n \ x^n" and "x^n \ Ifloat u^n " using assms unfolding atLeastAtMost_iff using power_mono_odd[OF False] by auto + have "Ifloat l ^ n \ x ^ n" and "x ^ n \ Ifloat u ^ n " using assms unfolding atLeastAtMost_iff using power_mono_odd[OF False] by auto thus ?thesis unfolding atLeastAtMost_iff l1 u1 float_power less_float_def by auto qed -lemma bnds_power: "\ x l u. (l1, u1) = float_power_bnds n l u \ x \ {Ifloat l .. Ifloat u} \ Ifloat l1 \ x^n \ x^n \ Ifloat u1" +lemma bnds_power: "\ x l u. (l1, u1) = float_power_bnds n l u \ x \ {Ifloat l .. Ifloat u} \ Ifloat l1 \ x ^ n \ x ^ n \ Ifloat u1" using float_power_bnds by auto section "Square root" @@ -794,8 +794,8 @@ let "?f n" = "fact (2 * n)" { fix n - have F: "\m. ((\i. i + 2) ^ n) m = m + 2 * n" by (induct n arbitrary: m, auto) - have "?f (Suc n) = ?f n * ((\i. i + 2) ^ n) 1 * (((\i. i + 2) ^ n) 1 + 1)" + have F: "\m. ((\i. i + 2) ^^ n) m = m + 2 * n" by (induct n arbitrary: m, auto) + have "?f (Suc n) = ?f n * ((\i. i + 2) ^^ n) 1 * (((\i. i + 2) ^^ n) 1 + 1)" unfolding F by auto } note f_eq = this from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0, @@ -811,7 +811,7 @@ have "0 < x * x" using `0 < x` unfolding less_float_def Ifloat_mult Ifloat_0 using mult_pos_pos[where a="Ifloat x" and b="Ifloat x"] by auto - { fix x n have "(\ i=0.. i=0.. i = 0 ..< 2 * n. (if even(i) then (-1 ^ (i div 2))/(real (fact i)) else 0) * x ^ i)" (is "?sum = ?ifsum") proof - have "?sum = ?sum + (\ j = 0 ..< n. 0)" by auto @@ -905,8 +905,8 @@ let "?f n" = "fact (2 * n + 1)" { fix n - have F: "\m. ((\i. i + 2) ^ n) m = m + 2 * n" by (induct n arbitrary: m, auto) - have "?f (Suc n) = ?f n * ((\i. i + 2) ^ n) 2 * (((\i. i + 2) ^ n) 2 + 1)" + have F: "\m. ((\i. i + 2) ^^ n) m = m + 2 * n" by (induct n arbitrary: m, auto) + have "?f (Suc n) = ?f n * ((\i. i + 2) ^^ n) 2 * (((\i. i + 2) ^^ n) 2 + 1)" unfolding F by auto } note f_eq = this from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0, @@ -1382,8 +1382,8 @@ shows "exp (Ifloat x) \ { Ifloat (lb_exp_horner prec (get_even n) 1 1 x) .. Ifloat (ub_exp_horner prec (get_odd n) 1 1 x) }" proof - { fix n - have F: "\ m. ((\i. i + 1) ^ n) m = n + m" by (induct n, auto) - have "fact (Suc n) = fact n * ((\i. i + 1) ^ n) 1" unfolding F by auto } note f_eq = this + have F: "\ m. ((\i. i + 1) ^^ n) m = n + m" by (induct n, auto) + have "fact (Suc n) = fact n * ((\i. i + 1) ^^ n) 1" unfolding F by auto } note f_eq = this note bounds = horner_bounds_nonpos[where f="fact" and lb="lb_exp_horner prec" and ub="ub_exp_horner prec" and j'=0 and s=1, OF assms f_eq lb_exp_horner.simps ub_exp_horner.simps] @@ -1462,7 +1462,8 @@ finally have "0 < Ifloat ((?horner x) ^ num)" . } ultimately show ?thesis - unfolding lb_exp.simps if_not_P[OF `\ 0 < x`] Let_def by (cases "floor_fl x", cases "x < - 1", auto simp add: le_float_def less_float_def normfloat) + unfolding lb_exp.simps if_not_P[OF `\ 0 < x`] Let_def + by (cases "floor_fl x", cases "x < - 1", auto simp add: float_power le_float_def less_float_def) qed lemma exp_boundaries': assumes "x \ 0" @@ -1631,10 +1632,10 @@ lemma ln_bounds: assumes "0 \ x" and "x < 1" - shows "(\i=0..<2*n. -1^i * (1 / real (i + 1)) * x^(Suc i)) \ ln (x + 1)" (is "?lb") - and "ln (x + 1) \ (\i=0..<2*n + 1. -1^i * (1 / real (i + 1)) * x^(Suc i))" (is "?ub") + shows "(\i=0..<2*n. -1^i * (1 / real (i + 1)) * x ^ (Suc i)) \ ln (x + 1)" (is "?lb") + and "ln (x + 1) \ (\i=0..<2*n + 1. -1^i * (1 / real (i + 1)) * x ^ (Suc i))" (is "?ub") proof - - let "?a n" = "(1/real (n +1)) * x^(Suc n)" + let "?a n" = "(1/real (n +1)) * x ^ (Suc n)" have ln_eq: "(\ i. -1^i * ?a i) = ln (x + 1)" using ln_series[of "x + 1"] `0 \ x` `x < 1` by auto @@ -2479,7 +2480,7 @@ fun lift_var (Free (varname, _)) = (case AList.lookup (op =) bound_eqs varname of SOME bound => bound | NONE => raise TERM ("No bound equations found for " ^ varname, [])) - | lift_var t = raise TERM ("Can not convert expression " ^ + | lift_var t = raise TERM ("Can not convert expression " ^ (Syntax.string_of_term ctxt t), [t]) val _ $ vs = HOLogic.dest_Trueprop (Logic.strip_imp_concl goal') diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Decision_Procs/cooper_tac.ML --- a/src/HOL/Decision_Procs/cooper_tac.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Decision_Procs/cooper_tac.ML Mon Apr 27 07:26:17 2009 -0700 @@ -76,7 +76,7 @@ @{thm "div_by_1"}, @{thm "mod_by_1"}, @{thm "div_1"}, @{thm "mod_1"}, Suc_plus1] addsimps @{thms add_ac} - addsimprocs [cancel_div_mod_proc] + addsimprocs [cancel_div_mod_nat_proc, cancel_div_mod_int_proc] val simpset0 = HOL_basic_ss addsimps [mod_div_equality', Suc_plus1] addsimps comp_arith diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Decision_Procs/mir_tac.ML --- a/src/HOL/Decision_Procs/mir_tac.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Decision_Procs/mir_tac.ML Mon Apr 27 07:26:17 2009 -0700 @@ -99,7 +99,7 @@ @{thm "div_by_1"}, @{thm "mod_by_1"}, @{thm "div_1"}, @{thm "mod_1"}, @{thm "Suc_plus1"}] addsimps @{thms add_ac} - addsimprocs [cancel_div_mod_proc] + addsimprocs [cancel_div_mod_nat_proc, cancel_div_mod_int_proc] val simpset0 = HOL_basic_ss addsimps [mod_div_equality', Suc_plus1] addsimps comp_ths diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Divides.thy --- a/src/HOL/Divides.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Divides.thy Mon Apr 27 07:26:17 2009 -0700 @@ -1,5 +1,4 @@ (* Title: HOL/Divides.thy - ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1999 University of Cambridge *) @@ -20,11 +19,12 @@ subsection {* Abstract division in commutative semirings. *} -class semiring_div = comm_semiring_1_cancel + div + +class semiring_div = comm_semiring_1_cancel + no_zero_divisors + div + assumes mod_div_equality: "a div b * b + a mod b = a" and div_by_0 [simp]: "a div 0 = 0" and div_0 [simp]: "0 div a = 0" and div_mult_self1 [simp]: "b \ 0 \ (a + c * b) div b = c + a div b" + and div_mult_mult1 [simp]: "c \ 0 \ (c * a) div (c * b) = a div b" begin text {* @{const div} and @{const mod} *} @@ -38,16 +38,16 @@ by (simp only: add_ac) lemma div_mod_equality: "((a div b) * b + a mod b) + c = a + c" -by (simp add: mod_div_equality) + by (simp add: mod_div_equality) lemma div_mod_equality2: "(b * (a div b) + a mod b) + c = a + c" -by (simp add: mod_div_equality2) + by (simp add: mod_div_equality2) lemma mod_by_0 [simp]: "a mod 0 = a" -using mod_div_equality [of a zero] by simp + using mod_div_equality [of a zero] by simp lemma mod_0 [simp]: "0 mod a = 0" -using mod_div_equality [of zero a] div_0 by simp + using mod_div_equality [of zero a] div_0 by simp lemma div_mult_self2 [simp]: assumes "b \ 0" @@ -72,7 +72,7 @@ qed lemma mod_mult_self2 [simp]: "(a + b * c) mod b = a mod b" -by (simp add: mult_commute [of b]) + by (simp add: mult_commute [of b]) lemma div_mult_self1_is_id [simp]: "b \ 0 \ b * a div b = a" using div_mult_self2 [of b 0 a] by simp @@ -238,9 +238,9 @@ by (simp only: mod_add_eq [symmetric]) qed -lemma div_add[simp]: "z dvd x \ z dvd y +lemma div_add [simp]: "z dvd x \ z dvd y \ (x + y) div z = x div z + y div z" -by(cases "z=0", simp, unfold dvd_def, auto simp add: algebra_simps) +by (cases "z = 0", simp, unfold dvd_def, auto simp add: algebra_simps) text {* Multiplication respects modular equivalence. *} @@ -297,21 +297,41 @@ finally show ?thesis . qed -end - -lemma div_mult_div_if_dvd: "(y::'a::{semiring_div,no_zero_divisors}) dvd x \ - z dvd w \ (x div y) * (w div z) = (x * w) div (y * z)" -unfolding dvd_def - apply clarify - apply (case_tac "y = 0") - apply simp - apply (case_tac "z = 0") - apply simp - apply (simp add: algebra_simps) +lemma div_mult_div_if_dvd: + "y dvd x \ z dvd w \ (x div y) * (w div z) = (x * w) div (y * z)" + apply (cases "y = 0", simp) + apply (cases "z = 0", simp) + apply (auto elim!: dvdE simp add: algebra_simps) apply (subst mult_assoc [symmetric]) apply (simp add: no_zero_divisors) -done + done + +lemma div_mult_mult2 [simp]: + "c \ 0 \ (a * c) div (b * c) = a div b" + by (drule div_mult_mult1) (simp add: mult_commute) + +lemma div_mult_mult1_if [simp]: + "(c * a) div (c * b) = (if c = 0 then 0 else a div b)" + by simp_all +lemma mod_mult_mult1: + "(c * a) mod (c * b) = c * (a mod b)" +proof (cases "c = 0") + case True then show ?thesis by simp +next + case False + from mod_div_equality + have "((c * a) div (c * b)) * (c * b) + (c * a) mod (c * b) = c * a" . + with False have "c * ((a div b) * b + a mod b) + (c * a) mod (c * b) + = c * a + c * (a mod b)" by (simp add: algebra_simps) + with mod_div_equality show ?thesis by simp +qed + +lemma mod_mult_mult2: + "(a * c) mod (b * c) = (a mod b) * c" + using mod_mult_mult1 [of c a b] by (simp add: mult_commute) + +end lemma div_power: "(y::'a::{semiring_div,no_zero_divisors,recpower}) dvd x \ (x div y)^n = x^n div y^n" @@ -398,15 +418,17 @@ @{term "q\nat"}(uotient) and @{term "r\nat"}(emainder). *} -definition divmod_rel :: "nat \ nat \ nat \ nat \ bool" where - "divmod_rel m n q r \ m = q * n + r \ (if n > 0 then 0 \ r \ r < n else q = 0)" +definition divmod_rel :: "nat \ nat \ nat \ nat \ bool" where + "divmod_rel m n qr \ + m = fst qr * n + snd qr \ + (if n = 0 then fst qr = 0 else if n > 0 then 0 \ snd qr \ snd qr < n else n < snd qr \ snd qr \ 0)" text {* @{const divmod_rel} is total: *} lemma divmod_rel_ex: - obtains q r where "divmod_rel m n q r" + obtains q r where "divmod_rel m n (q, r)" proof (cases "n = 0") - case True with that show thesis + case True with that show thesis by (auto simp add: divmod_rel_def) next case False @@ -436,13 +458,14 @@ text {* @{const divmod_rel} is injective: *} -lemma divmod_rel_unique_div: - assumes "divmod_rel m n q r" - and "divmod_rel m n q' r'" - shows "q = q'" +lemma divmod_rel_unique: + assumes "divmod_rel m n qr" + and "divmod_rel m n qr'" + shows "qr = qr'" proof (cases "n = 0") case True with assms show ?thesis - by (simp add: divmod_rel_def) + by (cases qr, cases qr') + (simp add: divmod_rel_def) next case False have aux: "\q r q' r'. q' * n + r' = q * n + r \ r < n \ q' \ (q\nat)" @@ -450,18 +473,11 @@ apply (subst less_iff_Suc_add) apply (auto simp add: add_mult_distrib) done - from `n \ 0` assms show ?thesis - by (auto simp add: divmod_rel_def - intro: order_antisym dest: aux sym) -qed - -lemma divmod_rel_unique_mod: - assumes "divmod_rel m n q r" - and "divmod_rel m n q' r'" - shows "r = r'" -proof - - from assms have "q = q'" by (rule divmod_rel_unique_div) - with assms show ?thesis by (simp add: divmod_rel_def) + from `n \ 0` assms have "fst qr = fst qr'" + by (auto simp add: divmod_rel_def intro: order_antisym dest: aux sym) + moreover from this assms have "snd qr = snd qr'" + by (simp add: divmod_rel_def) + ultimately show ?thesis by (cases qr, cases qr') simp qed text {* @@ -473,7 +489,21 @@ begin definition divmod :: "nat \ nat \ nat \ nat" where - [code del]: "divmod m n = (THE (q, r). divmod_rel m n q r)" + [code del]: "divmod m n = (THE qr. divmod_rel m n qr)" + +lemma divmod_rel_divmod: + "divmod_rel m n (divmod m n)" +proof - + from divmod_rel_ex + obtain qr where rel: "divmod_rel m n qr" . + then show ?thesis + by (auto simp add: divmod_def intro: theI elim: divmod_rel_unique) +qed + +lemma divmod_eq: + assumes "divmod_rel m n qr" + shows "divmod m n = qr" + using assms by (auto intro: divmod_rel_unique divmod_rel_divmod) definition div_nat where "m div n = fst (divmod m n)" @@ -485,30 +515,18 @@ "divmod m n = (m div n, m mod n)" unfolding div_nat_def mod_nat_def by simp -lemma divmod_eq: - assumes "divmod_rel m n q r" - shows "divmod m n = (q, r)" - using assms by (auto simp add: divmod_def - dest: divmod_rel_unique_div divmod_rel_unique_mod) - lemma div_eq: - assumes "divmod_rel m n q r" + assumes "divmod_rel m n (q, r)" shows "m div n = q" - using assms by (auto dest: divmod_eq simp add: div_nat_def) + using assms by (auto dest: divmod_eq simp add: divmod_div_mod) lemma mod_eq: - assumes "divmod_rel m n q r" + assumes "divmod_rel m n (q, r)" shows "m mod n = r" - using assms by (auto dest: divmod_eq simp add: mod_nat_def) + using assms by (auto dest: divmod_eq simp add: divmod_div_mod) -lemma divmod_rel: "divmod_rel m n (m div n) (m mod n)" -proof - - from divmod_rel_ex - obtain q r where rel: "divmod_rel m n q r" . - moreover with div_eq mod_eq have "m div n = q" and "m mod n = r" - by simp_all - ultimately show ?thesis by simp -qed +lemma divmod_rel: "divmod_rel m n (m div n, m mod n)" + by (simp add: div_nat_def mod_nat_def divmod_rel_divmod) lemma divmod_zero: "divmod m 0 = (0, m)" @@ -531,10 +549,10 @@ assumes "0 < n" and "n \ m" shows "divmod m n = (Suc ((m - n) div n), (m - n) mod n)" proof - - from divmod_rel have divmod_m_n: "divmod_rel m n (m div n) (m mod n)" . + from divmod_rel have divmod_m_n: "divmod_rel m n (m div n, m mod n)" . with assms have m_div_n: "m div n \ 1" by (cases "m div n") (auto simp add: divmod_rel_def) - from assms divmod_m_n have "divmod_rel (m - n) n (m div n - Suc 0) (m mod n)" + from assms divmod_m_n have "divmod_rel (m - n) n (m div n - Suc 0, m mod n)" by (cases "m div n") (auto simp add: divmod_rel_def) with divmod_eq have "divmod (m - n) n = (m div n - Suc 0, m mod n)" by simp moreover from divmod_div_mod have "divmod (m - n) n = ((m - n) div n, (m - n) mod n)" . @@ -569,55 +587,74 @@ shows "m mod n = (m - n) mod n" using assms divmod_step divmod_div_mod by (cases "n = 0") simp_all -instance proof - fix m n :: nat show "m div n * n + m mod n = m" - using divmod_rel [of m n] by (simp add: divmod_rel_def) -next - fix n :: nat show "n div 0 = 0" - using divmod_zero divmod_div_mod [of n 0] by simp -next - fix n :: nat show "0 div n = 0" - using divmod_rel [of 0 n] by (cases n) (simp_all add: divmod_rel_def) -next - fix m n q :: nat assume "n \ 0" then show "(q + m * n) div n = m + q div n" - by (induct m) (simp_all add: le_div_geq) +instance proof - + have [simp]: "\n::nat. n div 0 = 0" + by (simp add: div_nat_def divmod_zero) + have [simp]: "\n::nat. 0 div n = 0" + proof - + fix n :: nat + show "0 div n = 0" + by (cases "n = 0") simp_all + qed + show "OFCLASS(nat, semiring_div_class)" proof + fix m n :: nat + show "m div n * n + m mod n = m" + using divmod_rel [of m n] by (simp add: divmod_rel_def) + next + fix m n q :: nat + assume "n \ 0" + then show "(q + m * n) div n = m + q div n" + by (induct m) (simp_all add: le_div_geq) + next + fix m n q :: nat + assume "m \ 0" + then show "(m * n) div (m * q) = n div q" + proof (cases "n \ 0 \ q \ 0") + case False then show ?thesis by auto + next + case True with `m \ 0` + have "m > 0" and "n > 0" and "q > 0" by auto + then have "\a b. divmod_rel n q (a, b) \ divmod_rel (m * n) (m * q) (a, m * b)" + by (auto simp add: divmod_rel_def) (simp_all add: algebra_simps) + moreover from divmod_rel have "divmod_rel n q (n div q, n mod q)" . + ultimately have "divmod_rel (m * n) (m * q) (n div q, m * (n mod q))" . + then show ?thesis by (simp add: div_eq) + qed + qed simp_all qed end text {* Simproc for cancelling @{const div} and @{const mod} *} -(*lemmas mod_div_equality_nat = semiring_div_class.times_div_mod_plus_zero_one.mod_div_equality [of "m\nat" n, standard] -lemmas mod_div_equality2_nat = mod_div_equality2 [of "n\nat" m, standard*) +ML {* +local + +structure CancelDivMod = CancelDivModFun(struct -ML {* -structure CancelDivModData = -struct - -val div_name = @{const_name div}; -val mod_name = @{const_name mod}; -val mk_binop = HOLogic.mk_binop; -val mk_sum = Nat_Arith.mk_sum; -val dest_sum = Nat_Arith.dest_sum; + val div_name = @{const_name div}; + val mod_name = @{const_name mod}; + val mk_binop = HOLogic.mk_binop; + val mk_sum = Nat_Arith.mk_sum; + val dest_sum = Nat_Arith.dest_sum; -(*logic*) + val div_mod_eqs = map mk_meta_eq [@{thm div_mod_equality}, @{thm div_mod_equality2}]; -val div_mod_eqs = map mk_meta_eq [@{thm div_mod_equality}, @{thm div_mod_equality2}] - -val trans = trans + val trans = trans; -val prove_eq_sums = - let val simps = @{thm add_0} :: @{thm add_0_right} :: @{thms add_ac} - in Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac simps) end; + val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac + (@{thm monoid_add_class.add_0_left} :: @{thm monoid_add_class.add_0_right} :: @{thms add_ac})) -end; +end) -structure CancelDivMod = CancelDivModFun(CancelDivModData); +in -val cancel_div_mod_proc = Simplifier.simproc (the_context ()) +val cancel_div_mod_nat_proc = Simplifier.simproc (the_context ()) "cancel_div_mod" ["(m::nat) + n"] (K CancelDivMod.proc); -Addsimprocs[cancel_div_mod_proc]; +val _ = Addsimprocs [cancel_div_mod_nat_proc]; + +end *} text {* code generator setup *} @@ -658,7 +695,7 @@ fixes m n :: nat assumes "n > 0" shows "m mod n < (n::nat)" - using assms divmod_rel unfolding divmod_rel_def by auto + using assms divmod_rel [of m n] unfolding divmod_rel_def by auto lemma mod_less_eq_dividend [simp]: fixes m n :: nat @@ -700,18 +737,19 @@ subsubsection {* Quotient and Remainder *} lemma divmod_rel_mult1_eq: - "[| divmod_rel b c q r; c > 0 |] - ==> divmod_rel (a*b) c (a*q + a*r div c) (a*r mod c)" + "divmod_rel b c (q, r) \ c > 0 + \ divmod_rel (a * b) c (a * q + a * r div c, a * r mod c)" by (auto simp add: split_ifs divmod_rel_def algebra_simps) -lemma div_mult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::nat)" +lemma div_mult1_eq: + "(a * b) div c = a * (b div c) + a * (b mod c) div (c::nat)" apply (cases "c = 0", simp) apply (blast intro: divmod_rel [THEN divmod_rel_mult1_eq, THEN div_eq]) done lemma divmod_rel_add1_eq: - "[| divmod_rel a c aq ar; divmod_rel b c bq br; c > 0 |] - ==> divmod_rel (a + b) c (aq + bq + (ar+br) div c) ((ar + br) mod c)" + "divmod_rel a c (aq, ar) \ divmod_rel b c (bq, br) \ c > 0 + \ divmod_rel (a + b) c (aq + bq + (ar + br) div c, (ar + br) mod c)" by (auto simp add: split_ifs divmod_rel_def algebra_simps) (*NOT suitable for rewriting: the RHS has an instance of the LHS*) @@ -728,8 +766,9 @@ apply (simp add: add_mult_distrib2) done -lemma divmod_rel_mult2_eq: "[| divmod_rel a b q r; 0 < b; 0 < c |] - ==> divmod_rel a (b*c) (q div c) (b*(q mod c) + r)" +lemma divmod_rel_mult2_eq: + "divmod_rel a b (q, r) \ 0 < b \ 0 < c + \ divmod_rel a (b * c) (q div c, b *(q mod c) + r)" by (auto simp add: mult_ac divmod_rel_def add_mult_distrib2 [symmetric] mod_lemma) lemma div_mult2_eq: "a div (b*c) = (a div b) div (c::nat)" @@ -745,23 +784,6 @@ done -subsubsection{*Cancellation of Common Factors in Division*} - -lemma div_mult_mult_lemma: - "[| (0::nat) < b; 0 < c |] ==> (c*a) div (c*b) = a div b" -by (auto simp add: div_mult2_eq) - -lemma div_mult_mult1 [simp]: "(0::nat) < c ==> (c*a) div (c*b) = a div b" - apply (cases "b = 0") - apply (auto simp add: linorder_neq_iff [of b] div_mult_mult_lemma) - done - -lemma div_mult_mult2 [simp]: "(0::nat) < c ==> (a*c) div (b*c) = a div b" - apply (drule div_mult_mult1) - apply (auto simp add: mult_commute) - done - - subsubsection{*Further Facts about Quotient and Remainder*} lemma div_1 [simp]: "m div Suc 0 = m" @@ -769,7 +791,7 @@ (* Monotonicity of div in first argument *) -lemma div_le_mono [rule_format]: +lemma div_le_mono [rule_format (no_asm)]: "\m::nat. m \ n --> (m div k) \ (n div k)" apply (case_tac "k=0", simp) apply (induct "n" rule: nat_less_induct, clarify) @@ -824,12 +846,6 @@ apply (simp_all) done -lemma nat_div_eq_0 [simp]: "(n::nat) > 0 ==> ((m div n) = 0) = (m < n)" -by(auto, subst mod_div_equality [of m n, symmetric], auto) - -lemma nat_div_gt_0 [simp]: "(n::nat) > 0 ==> ((m div n) > 0) = (m >= n)" -by (subst neq0_conv [symmetric], auto) - declare div_less_dividend [simp] text{*A fact for the mutilated chess board*} @@ -915,16 +931,10 @@ done lemma dvd_imp_le: "[| k dvd n; 0 < n |] ==> k \ (n::nat)" -by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc) - -lemma nat_dvd_not_less: "(0::nat) < m \ m < n \ \ n dvd m" -by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc) + by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc) lemma dvd_mult_div_cancel: "n dvd m ==> n * (m div n) = (m::nat)" - apply (subgoal_tac "m mod n = 0") - apply (simp add: mult_div_cancel) - apply (simp only: dvd_eq_mod_eq_0) - done + by (simp add: dvd_eq_mod_eq_0 mult_div_cancel) lemma nat_zero_less_power_iff [simp]: "(x^n > 0) = (x > (0::nat) | n=0)" by (induct n) auto @@ -1001,9 +1011,11 @@ from A B show ?lhs .. next assume P: ?lhs - then have "divmod_rel m n q (m - n * q)" + then have "divmod_rel m n (q, m - n * q)" unfolding divmod_rel_def by (auto simp add: mult_ac) - then show ?rhs using divmod_rel by (rule divmod_rel_unique_div) + with divmod_rel_unique divmod_rel [of m n] + have "(q, m - n * q) = (m div n, m mod n)" by auto + then show ?rhs by simp qed theorem split_div': @@ -1155,4 +1167,9 @@ with j show ?thesis by blast qed +lemma nat_dvd_not_less: + fixes m n :: nat + shows "0 < m \ m < n \ \ n dvd m" +by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc) + end diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Groebner_Basis.thy --- a/src/HOL/Groebner_Basis.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Groebner_Basis.thy Mon Apr 27 07:26:17 2009 -0700 @@ -5,7 +5,7 @@ header {* Semiring normalization and Groebner Bases *} theory Groebner_Basis -imports NatBin +imports Nat_Numeral uses "Tools/Groebner_Basis/misc.ML" "Tools/Groebner_Basis/normalizer_data.ML" diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/HOL.thy --- a/src/HOL/HOL.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/HOL.thy Mon Apr 27 07:26:17 2009 -0700 @@ -5,9 +5,10 @@ header {* The basis of Higher-Order Logic *} theory HOL -imports Pure +imports Pure "~~/src/Tools/Code_Generator" uses ("Tools/hologic.ML") + "~~/src/Tools/auto_solve.ML" "~~/src/Tools/IsaPlanner/zipper.ML" "~~/src/Tools/IsaPlanner/isand.ML" "~~/src/Tools/IsaPlanner/rw_tools.ML" @@ -27,16 +28,6 @@ "~~/src/Tools/atomize_elim.ML" "~~/src/Tools/induct.ML" ("~~/src/Tools/induct_tacs.ML") - "~~/src/Tools/value.ML" - "~~/src/Tools/code/code_name.ML" - "~~/src/Tools/code/code_funcgr.ML" (*formal dependency*) - "~~/src/Tools/code/code_wellsorted.ML" - "~~/src/Tools/code/code_thingol.ML" - "~~/src/Tools/code/code_printer.ML" - "~~/src/Tools/code/code_target.ML" - "~~/src/Tools/code/code_ml.ML" - "~~/src/Tools/code/code_haskell.ML" - "~~/src/Tools/nbe.ML" ("Tools/recfun_codegen.ML") begin @@ -1674,37 +1665,264 @@ *} -subsection {* Code generator basics -- see further theory @{text "Code_Setup"} *} +subsection {* Code generator setup *} + +subsubsection {* SML code generator setup *} + +use "Tools/recfun_codegen.ML" + +setup {* + Codegen.setup + #> RecfunCodegen.setup +*} + +types_code + "bool" ("bool") +attach (term_of) {* +fun term_of_bool b = if b then HOLogic.true_const else HOLogic.false_const; +*} +attach (test) {* +fun gen_bool i = + let val b = one_of [false, true] + in (b, fn () => term_of_bool b) end; +*} + "prop" ("bool") +attach (term_of) {* +fun term_of_prop b = + HOLogic.mk_Trueprop (if b then HOLogic.true_const else HOLogic.false_const); +*} -text {* Equality *} +consts_code + "Trueprop" ("(_)") + "True" ("true") + "False" ("false") + "Not" ("Bool.not") + "op |" ("(_ orelse/ _)") + "op &" ("(_ andalso/ _)") + "If" ("(if _/ then _/ else _)") + +setup {* +let + +fun eq_codegen thy defs dep thyname b t gr = + (case strip_comb t of + (Const ("op =", Type (_, [Type ("fun", _), _])), _) => NONE + | (Const ("op =", _), [t, u]) => + let + val (pt, gr') = Codegen.invoke_codegen thy defs dep thyname false t gr; + val (pu, gr'') = Codegen.invoke_codegen thy defs dep thyname false u gr'; + val (_, gr''') = Codegen.invoke_tycodegen thy defs dep thyname false HOLogic.boolT gr''; + in + SOME (Codegen.parens + (Pretty.block [pt, Codegen.str " =", Pretty.brk 1, pu]), gr''') + end + | (t as Const ("op =", _), ts) => SOME (Codegen.invoke_codegen + thy defs dep thyname b (Codegen.eta_expand t ts 2) gr) + | _ => NONE); + +in + Codegen.add_codegen "eq_codegen" eq_codegen +end +*} + +subsubsection {* Equality *} class eq = fixes eq :: "'a \ 'a \ bool" assumes eq_equals: "eq x y \ x = y" begin -lemma eq: "eq = (op =)" +lemma eq [code unfold, code inline del]: "eq = (op =)" by (rule ext eq_equals)+ lemma eq_refl: "eq x x \ True" unfolding eq by rule+ +lemma equals_eq [code inline]: "(op =) \ eq" + by (rule eq_reflection) (rule ext, rule ext, rule sym, rule eq_equals) + +declare equals_eq [symmetric, code post] + end -text {* Module setup *} +declare equals_eq [code] + + +subsubsection {* Generic code generator foundation *} + +text {* Datatypes *} + +code_datatype True False + +code_datatype "TYPE('a\{})" + +code_datatype Trueprop "prop" + +text {* Code equations *} + +lemma [code]: + shows "(True \ PROP P) \ PROP P" + and "(False \ Q) \ Trueprop True" + and "(PROP P \ True) \ Trueprop True" + and "(Q \ False) \ Trueprop (\ Q)" by (auto intro!: equal_intr_rule) + +lemma [code]: + shows "False \ x \ False" + and "True \ x \ x" + and "x \ False \ False" + and "x \ True \ x" by simp_all + +lemma [code]: + shows "False \ x \ x" + and "True \ x \ True" + and "x \ False \ x" + and "x \ True \ True" by simp_all + +lemma [code]: + shows "\ True \ False" + and "\ False \ True" by (rule HOL.simp_thms)+ -use "Tools/recfun_codegen.ML" +lemmas [code] = Let_def if_True if_False + +lemmas [code, code unfold, symmetric, code post] = imp_conv_disj + +text {* Equality *} + +declare simp_thms(6) [code nbe] + +hide (open) const eq +hide const eq + +setup {* + Code_Unit.add_const_alias @{thm equals_eq} +*} + +text {* Cases *} + +lemma Let_case_cert: + assumes "CASE \ (\x. Let x f)" + shows "CASE x \ f x" + using assms by simp_all + +lemma If_case_cert: + assumes "CASE \ (\b. If b f g)" + shows "(CASE True \ f) &&& (CASE False \ g)" + using assms by simp_all + +setup {* + Code.add_case @{thm Let_case_cert} + #> Code.add_case @{thm If_case_cert} + #> Code.add_undefined @{const_name undefined} +*} + +code_abort undefined + +subsubsection {* Generic code generator preprocessor *} setup {* - Code_ML.setup - #> Code_Haskell.setup - #> Nbe.setup - #> Codegen.setup - #> RecfunCodegen.setup + Code.map_pre (K HOL_basic_ss) + #> Code.map_post (K HOL_basic_ss) *} +subsubsection {* Generic code generator target languages *} -subsection {* Nitpick hooks *} +text {* type bool *} + +code_type bool + (SML "bool") + (OCaml "bool") + (Haskell "Bool") + +code_const True and False and Not and "op &" and "op |" and If + (SML "true" and "false" and "not" + and infixl 1 "andalso" and infixl 0 "orelse" + and "!(if (_)/ then (_)/ else (_))") + (OCaml "true" and "false" and "not" + and infixl 4 "&&" and infixl 2 "||" + and "!(if (_)/ then (_)/ else (_))") + (Haskell "True" and "False" and "not" + and infixl 3 "&&" and infixl 2 "||" + and "!(if (_)/ then (_)/ else (_))") + +code_reserved SML + bool true false not + +code_reserved OCaml + bool not + +text {* using built-in Haskell equality *} + +code_class eq + (Haskell "Eq") + +code_const "eq_class.eq" + (Haskell infixl 4 "==") + +code_const "op =" + (Haskell infixl 4 "==") + +text {* undefined *} + +code_const undefined + (SML "!(raise/ Fail/ \"undefined\")") + (OCaml "failwith/ \"undefined\"") + (Haskell "error/ \"undefined\"") + +subsubsection {* Evaluation and normalization by evaluation *} + +setup {* + Value.add_evaluator ("SML", Codegen.eval_term o ProofContext.theory_of) +*} + +ML {* +structure Eval_Method = +struct + +val eval_ref : (unit -> bool) option ref = ref NONE; + +end; +*} + +oracle eval_oracle = {* fn ct => + let + val thy = Thm.theory_of_cterm ct; + val t = Thm.term_of ct; + val dummy = @{cprop True}; + in case try HOLogic.dest_Trueprop t + of SOME t' => if Code_ML.eval NONE + ("Eval_Method.eval_ref", Eval_Method.eval_ref) (K I) thy t' [] + then Thm.capply (Thm.capply @{cterm "op \ \ prop \ prop \ prop"} ct) dummy + else dummy + | NONE => dummy + end +*} + +ML {* +fun gen_eval_method conv ctxt = SIMPLE_METHOD' + (CONVERSION (Conv.params_conv (~1) (K (Conv.concl_conv (~1) conv)) ctxt) + THEN' rtac TrueI) +*} + +method_setup eval = {* Scan.succeed (gen_eval_method eval_oracle) *} + "solve goal by evaluation" + +method_setup evaluation = {* Scan.succeed (gen_eval_method Codegen.evaluation_conv) *} + "solve goal by evaluation" + +method_setup normalization = {* + Scan.succeed (K (SIMPLE_METHOD' (CONVERSION Nbe.norm_conv THEN' (fn k => TRY (rtac TrueI k))))) +*} "solve goal by normalization" + +subsubsection {* Quickcheck *} + +setup {* + Quickcheck.add_generator ("SML", Codegen.test_term) +*} + +quickcheck_params [size = 5, iterations = 50] + + +subsection {* Nitpick setup *} text {* This will be relocated once Nitpick is moved to HOL. *} @@ -1730,10 +1948,14 @@ val description = "introduction rules for (co)inductive predicates as needed by Nitpick" ) *} -setup {* Nitpick_Const_Def_Thms.setup - #> Nitpick_Const_Simp_Thms.setup - #> Nitpick_Const_Psimp_Thms.setup - #> Nitpick_Ind_Intro_Thms.setup *} + +setup {* + Nitpick_Const_Def_Thms.setup + #> Nitpick_Const_Simp_Thms.setup + #> Nitpick_Const_Psimp_Thms.setup + #> Nitpick_Ind_Intro_Thms.setup +*} + subsection {* Legacy tactics and ML bindings *} diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/HoareParallel/OG_Tran.thy --- a/src/HOL/HoareParallel/OG_Tran.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/HoareParallel/OG_Tran.thy Mon Apr 27 07:26:17 2009 -0700 @@ -74,7 +74,7 @@ abbreviation ann_transition_n :: "('a ann_com_op \ 'a) \ nat \ ('a ann_com_op \ 'a) \ bool" ("_ -_\ _"[81,81] 100) where - "con_0 -n\ con_1 \ (con_0, con_1) \ ann_transition^n" + "con_0 -n\ con_1 \ (con_0, con_1) \ ann_transition ^^ n" abbreviation ann_transitions :: "('a ann_com_op \ 'a) \ ('a ann_com_op \ 'a) \ bool" @@ -84,7 +84,7 @@ abbreviation transition_n :: "('a com \ 'a) \ nat \ ('a com \ 'a) \ bool" ("_ -P_\ _"[81,81,81] 100) where - "con_0 -Pn\ con_1 \ (con_0, con_1) \ transition^n" + "con_0 -Pn\ con_1 \ (con_0, con_1) \ transition ^^ n" subsection {* Definition of Semantics *} diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/IMP/Compiler0.thy --- a/src/HOL/IMP/Compiler0.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/IMP/Compiler0.thy Mon Apr 27 07:26:17 2009 -0700 @@ -45,7 +45,7 @@ abbreviation stepan :: "[instr list,state,nat,nat,state,nat] \ bool" ("_ \/ (3\_,_\/ -(_)\ \_,_\)" [50,0,0,0,0,0] 50) where - "P \ \s,m\ -(i)\ \t,n\ == ((s,m),t,n) : ((stepa1 P)^i)" + "P \ \s,m\ -(i)\ \t,n\ == ((s,m),t,n) : (stepa1 P ^^ i)" subsection "The compiler" diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/IMP/Machines.thy --- a/src/HOL/IMP/Machines.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/IMP/Machines.thy Mon Apr 27 07:26:17 2009 -0700 @@ -1,7 +1,6 @@ - -(* $Id$ *) - -theory Machines imports Natural begin +theory Machines +imports Natural +begin lemma rtrancl_eq: "R^* = Id \ (R O R^*)" by (fast intro: rtrancl_into_rtrancl elim: rtranclE) @@ -11,20 +10,22 @@ lemmas converse_rel_powE = rel_pow_E2 -lemma R_O_Rn_commute: "R O R^n = R^n O R" +lemma R_O_Rn_commute: "R O R ^^ n = R ^^ n O R" by (induct n) (simp, simp add: O_assoc [symmetric]) lemma converse_in_rel_pow_eq: - "((x,z) \ R^n) = (n=0 \ z=x \ (\m y. n = Suc m \ (x,y) \ R \ (y,z) \ R^m))" + "((x,z) \ R ^^ n) = (n=0 \ z=x \ (\m y. n = Suc m \ (x,y) \ R \ (y,z) \ R ^^ m))" apply(rule iffI) apply(blast elim:converse_rel_powE) apply (fastsimp simp add:gr0_conv_Suc R_O_Rn_commute) done -lemma rel_pow_plus: "R^(m+n) = R^n O R^m" +lemma rel_pow_plus: + "R ^^ (m+n) = R ^^ n O R ^^ m" by (induct n) (simp, simp add: O_assoc) -lemma rel_pow_plusI: "\ (x,y) \ R^m; (y,z) \ R^n \ \ (x,z) \ R^(m+n)" +lemma rel_pow_plusI: + "\ (x,y) \ R ^^ m; (y,z) \ R ^^ n \ \ (x,z) \ R ^^ (m+n)" by (simp add: rel_pow_plus rel_compI) subsection "Instructions" @@ -57,7 +58,7 @@ abbreviation exec0n :: "[instrs, nat,state, nat, nat,state] \ bool" ("(_/ \ (1\_,/_\)/ -_\ (1\_,/_\))" [50,0,0,0,0] 50) where - "p \ \i,s\ -n\ \j,t\ == ((i,s),j,t) : (exec01 p)^n" + "p \ \i,s\ -n\ \j,t\ == ((i,s),j,t) : (exec01 p)^^n" subsection "M0 with lists" @@ -89,7 +90,7 @@ abbreviation stepan :: "[instrs,instrs,state, nat, instrs,instrs,state] \ bool" ("((1\_,/_,/_\)/ -_\ (1\_,/_,/_\))" 50) where - "\p,q,s\ -i\ \p',q',t\ == ((p,q,s),p',q',t) : (stepa1^i)" + "\p,q,s\ -i\ \p',q',t\ == ((p,q,s),p',q',t) : (stepa1^^i)" inductive_cases execE: "((i#is,p,s), (is',p',s')) : stepa1" diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/IMP/Transition.thy --- a/src/HOL/IMP/Transition.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/IMP/Transition.thy Mon Apr 27 07:26:17 2009 -0700 @@ -1,5 +1,4 @@ (* Title: HOL/IMP/Transition.thy - ID: $Id$ Author: Tobias Nipkow & Robert Sandner, TUM Isar Version: Gerwin Klein, 2001 Copyright 1996 TUM @@ -69,7 +68,7 @@ abbreviation evalcn :: "[(com option\state),nat,(com option\state)] \ bool" ("_ -_\\<^sub>1 _" [60,60,60] 60) where - "cs -n\\<^sub>1 cs' == (cs,cs') \ evalc1^n" + "cs -n\\<^sub>1 cs' == (cs,cs') \ evalc1^^n" abbreviation evalc' :: "[(com option\state),(com option\state)] \ bool" @@ -77,28 +76,9 @@ "cs \\<^sub>1\<^sup>* cs' == (cs,cs') \ evalc1^*" (*<*) -(* fixme: move to Relation_Power.thy *) -lemma rel_pow_Suc_E2 [elim!]: - "[| (x, z) \ R ^ Suc n; !!y. [| (x, y) \ R; (y, z) \ R ^ n |] ==> P |] ==> P" - by (blast dest: rel_pow_Suc_D2) +declare rel_pow_Suc_E2 [elim!] +(*>*) -lemma rtrancl_imp_rel_pow: "p \ R^* \ \n. p \ R^n" -proof (induct p) - fix x y - assume "(x, y) \ R\<^sup>*" - thus "\n. (x, y) \ R^n" - proof induct - fix a have "(a, a) \ R^0" by simp - thus "\n. (a, a) \ R ^ n" .. - next - fix a b c assume "\n. (a, b) \ R ^ n" - then obtain n where "(a, b) \ R^n" .. - moreover assume "(b, c) \ R" - ultimately have "(a, c) \ R^(Suc n)" by auto - thus "\n. (a, c) \ R^n" .. - qed -qed -(*>*) text {* As for the big step semantics you can read these rules in a syntax directed way: @@ -189,8 +169,8 @@ (*<*) (* FIXME: relpow.simps don't work *) lemmas [simp del] = relpow.simps -lemma rel_pow_0 [simp]: "!!R::('a*'a) set. R^0 = Id" by (simp add: relpow.simps) -lemma rel_pow_Suc_0 [simp]: "!!R::('a*'a) set. R^(Suc 0) = R" by (simp add: relpow.simps) +lemma rel_pow_0 [simp]: "!!R::('a*'a) set. R ^^ 0 = Id" by (simp add: relpow.simps) +lemma rel_pow_Suc_0 [simp]: "!!R::('a*'a) set. R ^^ Suc 0 = R" by (simp add: relpow.simps) (*>*) lemma evalc1_None_0 [simp]: "\s\ -n\\<^sub>1 y = (n = 0 \ y = \s\)" diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Import/HOL/HOL4Base.thy --- a/src/HOL/Import/HOL/HOL4Base.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Import/HOL/HOL4Base.thy Mon Apr 27 07:26:17 2009 -0700 @@ -2794,8 +2794,8 @@ by (import numeral numeral_fact) lemma numeral_funpow: "ALL n::nat. - ((f::'a::type => 'a::type) ^ n) (x::'a::type) = - (if n = 0 then x else (f ^ (n - 1)) (f x))" + ((f::'a::type => 'a::type) ^^ n) (x::'a::type) = + (if n = 0 then x else (f ^^ (n - 1)) (f x))" by (import numeral numeral_funpow) ;end_setup diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Import/HOL/HOL4Word32.thy --- a/src/HOL/Import/HOL/HOL4Word32.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Import/HOL/HOL4Word32.thy Mon Apr 27 07:26:17 2009 -0700 @@ -434,15 +434,15 @@ by (import word32 EQUIV_QT) lemma FUNPOW_THM: "ALL (f::'a::type => 'a::type) (n::nat) x::'a::type. - (f ^ n) (f x) = f ((f ^ n) x)" + (f ^^ n) (f x) = f ((f ^^ n) x)" by (import word32 FUNPOW_THM) lemma FUNPOW_THM2: "ALL (f::'a::type => 'a::type) (n::nat) x::'a::type. - (f ^ Suc n) x = f ((f ^ n) x)" + (f ^^ Suc n) x = f ((f ^^ n) x)" by (import word32 FUNPOW_THM2) lemma FUNPOW_COMP: "ALL (f::'a::type => 'a::type) (m::nat) (n::nat) a::'a::type. - (f ^ m) ((f ^ n) a) = (f ^ (m + n)) a" + (f ^^ m) ((f ^^ n) a) = (f ^^ (m + n)) a" by (import word32 FUNPOW_COMP) lemma INw_MODw: "ALL n::nat. INw (MODw n)" @@ -1170,23 +1170,23 @@ constdefs word_lsr :: "word32 => nat => word32" - "word_lsr == %(a::word32) n::nat. (word_lsr1 ^ n) a" + "word_lsr == %(a::word32) n::nat. (word_lsr1 ^^ n) a" -lemma word_lsr: "ALL (a::word32) n::nat. word_lsr a n = (word_lsr1 ^ n) a" +lemma word_lsr: "ALL (a::word32) n::nat. word_lsr a n = (word_lsr1 ^^ n) a" by (import word32 word_lsr) constdefs word_asr :: "word32 => nat => word32" - "word_asr == %(a::word32) n::nat. (word_asr1 ^ n) a" + "word_asr == %(a::word32) n::nat. (word_asr1 ^^ n) a" -lemma word_asr: "ALL (a::word32) n::nat. word_asr a n = (word_asr1 ^ n) a" +lemma word_asr: "ALL (a::word32) n::nat. word_asr a n = (word_asr1 ^^ n) a" by (import word32 word_asr) constdefs word_ror :: "word32 => nat => word32" - "word_ror == %(a::word32) n::nat. (word_ror1 ^ n) a" + "word_ror == %(a::word32) n::nat. (word_ror1 ^^ n) a" -lemma word_ror: "ALL (a::word32) n::nat. word_ror a n = (word_ror1 ^ n) a" +lemma word_ror: "ALL (a::word32) n::nat. word_ror a n = (word_ror1 ^^ n) a" by (import word32 word_ror) consts @@ -1583,4 +1583,3 @@ ;end_setup end - diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Import/HOL/arithmetic.imp --- a/src/HOL/Import/HOL/arithmetic.imp Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Import/HOL/arithmetic.imp Mon Apr 27 07:26:17 2009 -0700 @@ -43,7 +43,7 @@ "TWO" > "HOL4Base.arithmetic.TWO" "TIMES2" > "NatSimprocs.nat_mult_2" "SUC_SUB1" > "HOL4Base.arithmetic.SUC_SUB1" - "SUC_ONE_ADD" > "NatBin.Suc_eq_add_numeral_1_left" + "SUC_ONE_ADD" > "Nat_Numeral.Suc_eq_add_numeral_1_left" "SUC_NOT" > "Nat.nat.simps_2" "SUC_ELIM_THM" > "HOL4Base.arithmetic.SUC_ELIM_THM" "SUC_ADD_SYM" > "HOL4Base.arithmetic.SUC_ADD_SYM" @@ -233,7 +233,7 @@ "EVEN_AND_ODD" > "HOL4Base.arithmetic.EVEN_AND_ODD" "EVEN_ADD" > "HOL4Base.arithmetic.EVEN_ADD" "EVEN" > "HOL4Base.arithmetic.EVEN" - "EQ_MULT_LCANCEL" > "NatBin.nat_mult_eq_cancel_disj" + "EQ_MULT_LCANCEL" > "Nat_Numeral.nat_mult_eq_cancel_disj" "EQ_MONO_ADD_EQ" > "Nat.nat_add_right_cancel" "EQ_LESS_EQ" > "Orderings.order_eq_iff" "EQ_ADD_RCANCEL" > "Nat.nat_add_right_cancel" diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Import/HOL/real.imp --- a/src/HOL/Import/HOL/real.imp Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Import/HOL/real.imp Mon Apr 27 07:26:17 2009 -0700 @@ -99,7 +99,7 @@ "REAL_POW_INV" > "Power.power_inverse" "REAL_POW_DIV" > "Power.power_divide" "REAL_POW_ADD" > "Power.power_add" - "REAL_POW2_ABS" > "NatBin.power2_abs" + "REAL_POW2_ABS" > "Nat_Numeral.power2_abs" "REAL_POS_NZ" > "HOL4Real.real.REAL_POS_NZ" "REAL_POS" > "RealDef.real_of_nat_ge_zero" "REAL_POASQ" > "HOL4Real.real.REAL_POASQ" @@ -210,7 +210,7 @@ "REAL_LE_RDIV_EQ" > "Ring_and_Field.pos_le_divide_eq" "REAL_LE_RDIV" > "Ring_and_Field.mult_imp_le_div_pos" "REAL_LE_RADD" > "OrderedGroup.add_le_cancel_right" - "REAL_LE_POW2" > "NatBin.zero_compare_simps_12" + "REAL_LE_POW2" > "Nat_Numeral.zero_compare_simps_12" "REAL_LE_NEGTOTAL" > "HOL4Real.real.REAL_LE_NEGTOTAL" "REAL_LE_NEGR" > "OrderedGroup.le_minus_self_iff" "REAL_LE_NEGL" > "OrderedGroup.minus_le_self_iff" @@ -313,7 +313,7 @@ "POW_ONE" > "Power.power_one" "POW_NZ" > "Power.field_power_not_zero" "POW_MUL" > "Power.power_mult_distrib" - "POW_MINUS1" > "NatBin.power_minus1_even" + "POW_MINUS1" > "Nat_Numeral.power_minus1_even" "POW_M1" > "HOL4Real.real.POW_M1" "POW_LT" > "HOL4Real.real.POW_LT" "POW_LE" > "Power.power_mono" @@ -323,7 +323,7 @@ "POW_ABS" > "Power.power_abs" "POW_2_LT" > "RealPow.two_realpow_gt" "POW_2_LE1" > "RealPow.two_realpow_ge_one" - "POW_2" > "NatBin.power2_eq_square" + "POW_2" > "Nat_Numeral.power2_eq_square" "POW_1" > "Power.power_one_right" "POW_0" > "Power.power_0_Suc" "ABS_ZERO" > "OrderedGroup.abs_eq_0" @@ -335,7 +335,7 @@ "ABS_SIGN2" > "HOL4Real.real.ABS_SIGN2" "ABS_SIGN" > "HOL4Real.real.ABS_SIGN" "ABS_REFL" > "HOL4Real.real.ABS_REFL" - "ABS_POW2" > "NatBin.abs_power2" + "ABS_POW2" > "Nat_Numeral.abs_power2" "ABS_POS" > "OrderedGroup.abs_ge_zero" "ABS_NZ" > "OrderedGroup.zero_less_abs_iff" "ABS_NEG" > "OrderedGroup.abs_minus_cancel" diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Import/HOL4Compat.thy --- a/src/HOL/Import/HOL4Compat.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Import/HOL4Compat.thy Mon Apr 27 07:26:17 2009 -0700 @@ -202,19 +202,13 @@ constdefs FUNPOW :: "('a => 'a) => nat => 'a => 'a" - "FUNPOW f n == f ^ n" + "FUNPOW f n == f ^^ n" -lemma FUNPOW: "(ALL f x. (f ^ 0) x = x) & - (ALL f n x. (f ^ Suc n) x = (f ^ n) (f x))" -proof auto - fix f n x - have "ALL x. f ((f ^ n) x) = (f ^ n) (f x)" - by (induct n,auto) - thus "f ((f ^ n) x) = (f ^ n) (f x)" - .. -qed +lemma FUNPOW: "(ALL f x. (f ^^ 0) x = x) & + (ALL f n x. (f ^^ Suc n) x = (f ^^ n) (f x))" + by (simp add: funpow_swap1) -lemma [hol4rew]: "FUNPOW f n = f ^ n" +lemma [hol4rew]: "FUNPOW f n = f ^^ n" by (simp add: FUNPOW_def) lemma ADD: "(!n. (0::nat) + n = n) & (!m n. Suc m + n = Suc (m + n))" @@ -224,7 +218,7 @@ by simp lemma SUB: "(!m. (0::nat) - m = 0) & (!m n. (Suc m) - n = (if m < n then 0 else Suc (m - n)))" - by (simp, arith) + by (simp) arith lemma MAX_DEF: "max (m::nat) n = (if m < n then n else m)" by (simp add: max_def) diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Import/HOLLight/hollight.imp --- a/src/HOL/Import/HOLLight/hollight.imp Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Import/HOLLight/hollight.imp Mon Apr 27 07:26:17 2009 -0700 @@ -1515,7 +1515,7 @@ "EQ_REFL_T" > "HOL.simp_thms_6" "EQ_REFL" > "Presburger.fm_modd_pinf" "EQ_MULT_RCANCEL" > "Nat.mult_cancel2" - "EQ_MULT_LCANCEL" > "NatBin.nat_mult_eq_cancel_disj" + "EQ_MULT_LCANCEL" > "Nat_Numeral.nat_mult_eq_cancel_disj" "EQ_IMP_LE" > "HOLLight.hollight.EQ_IMP_LE" "EQ_EXT" > "HOL.meta_eq_to_obj_eq" "EQ_CLAUSES" > "HOLLight.hollight.EQ_CLAUSES" diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Int.thy --- a/src/HOL/Int.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Int.thy Mon Apr 27 07:26:17 2009 -0700 @@ -1536,7 +1536,7 @@ by (simp add: abs_if) lemma abs_power_minus_one [simp]: - "abs(-1 ^ n) = (1::'a::{ordered_idom,number_ring,recpower})" + "abs(-1 ^ n) = (1::'a::{ordered_idom,number_ring})" by (simp add: power_abs) lemma of_int_number_of_eq [simp]: @@ -1848,42 +1848,21 @@ subsection {* Integer Powers *} -instantiation int :: recpower +context ring_1 begin -primrec power_int where - "p ^ 0 = (1\int)" - | "p ^ (Suc n) = (p\int) * (p ^ n)" - -instance proof - fix z :: int - fix n :: nat - show "z ^ 0 = 1" by simp - show "z ^ Suc n = z * (z ^ n)" by simp -qed - -declare power_int.simps [simp del] +lemma of_int_power: + "of_int (z ^ n) = of_int z ^ n" + by (induct n) simp_all end -lemma zpower_zadd_distrib: "x ^ (y + z) = ((x ^ y) * (x ^ z)::int)" - by (rule Power.power_add) - -lemma zpower_zpower: "(x ^ y) ^ z = (x ^ (y * z)::int)" - by (rule Power.power_mult [symmetric]) - -lemma zero_less_zpower_abs_iff [simp]: - "(0 < abs x ^ n) \ (x \ (0::int) | n = 0)" - by (induct n) (auto simp add: zero_less_mult_iff) - -lemma zero_le_zpower_abs [simp]: "(0::int) \ abs x ^ n" - by (induct n) (auto simp add: zero_le_mult_iff) - -lemma of_int_power: - "of_int (z ^ n) = (of_int z ^ n :: 'a::{recpower, ring_1})" - by (induct n) simp_all - -lemma int_power: "int (m^n) = (int m) ^ n" +lemma zpower_zpower: + "(x ^ y) ^ z = (x ^ (y * z)::int)" + by (rule power_mult [symmetric]) + +lemma int_power: + "int (m ^ n) = int m ^ n" by (rule of_nat_power) lemmas zpower_int = int_power [symmetric] @@ -2224,6 +2203,8 @@ subsection {* Legacy theorems *} +instance int :: recpower .. + lemmas zminus_zminus = minus_minus [of "z::int", standard] lemmas zminus_0 = minus_zero [where 'a=int] lemmas zminus_zadd_distrib = minus_add_distrib [of "z::int" "w", standard] @@ -2278,4 +2259,15 @@ lemmas zless_le = less_int_def lemmas int_eq_of_nat = TrueI +lemma zpower_zadd_distrib: + "x ^ (y + z) = ((x ^ y) * (x ^ z)::int)" + by (rule power_add) + +lemma zero_less_zpower_abs_iff: + "(0 < abs x ^ n) \ (x \ (0::int) | n = 0)" + by (rule zero_less_power_abs_iff) + +lemma zero_le_zpower_abs: "(0::int) \ abs x ^ n" + by (rule zero_le_power_abs) + end diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/IntDiv.thy --- a/src/HOL/IntDiv.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/IntDiv.thy Mon Apr 27 07:26:17 2009 -0700 @@ -249,33 +249,33 @@ text {* Tool setup *} ML {* -local +local -structure CancelDivMod = CancelDivModFun( -struct - val div_name = @{const_name Divides.div}; - val mod_name = @{const_name Divides.mod}; +structure CancelDivMod = CancelDivModFun(struct + + val div_name = @{const_name div}; + val mod_name = @{const_name mod}; val mk_binop = HOLogic.mk_binop; val mk_sum = Int_Numeral_Simprocs.mk_sum HOLogic.intT; val dest_sum = Int_Numeral_Simprocs.dest_sum; - val div_mod_eqs = - map mk_meta_eq [@{thm zdiv_zmod_equality}, - @{thm zdiv_zmod_equality2}]; + + val div_mod_eqs = map mk_meta_eq [@{thm zdiv_zmod_equality}, @{thm zdiv_zmod_equality2}]; + val trans = trans; - val prove_eq_sums = - let - val simps = @{thm diff_int_def} :: Int_Numeral_Simprocs.add_0s @ @{thms zadd_ac} - in Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac simps) end; + + val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac + (@{thm diff_minus} :: @{thms add_0s} @ @{thms add_ac})) + end) in -val cancel_zdiv_zmod_proc = Simplifier.simproc (the_context ()) - "cancel_zdiv_zmod" ["(m::int) + n"] (K CancelDivMod.proc) +val cancel_div_mod_int_proc = Simplifier.simproc (the_context ()) + "cancel_zdiv_zmod" ["(k::int) + l"] (K CancelDivMod.proc); -end; +val _ = Addsimprocs [cancel_div_mod_int_proc]; -Addsimprocs [cancel_zdiv_zmod_proc] +end *} lemma pos_mod_conj : "(0::int) < b ==> 0 \ a mod b & a mod b < b" @@ -711,6 +711,26 @@ show "(a + c * b) div b = c + a div b" unfolding zdiv_zadd1_eq [of a "c * b"] using not0 by (simp add: zmod_zmult1_eq zmod_zdiv_trivial zdiv_zmult1_eq) +next + fix a b c :: int + assume "a \ 0" + then show "(a * b) div (a * c) = b div c" + proof (cases "b \ 0 \ c \ 0") + case False then show ?thesis by auto + next + case True then have "b \ 0" and "c \ 0" by auto + with `a \ 0` + have "\q r. divmod_rel b c (q, r) \ divmod_rel (a * b) (a * c) (q, a * r)" + apply (auto simp add: divmod_rel_def) + apply (auto simp add: algebra_simps) + apply (auto simp add: zero_less_mult_iff zero_le_mult_iff mult_le_0_iff) + apply (simp_all add: mult_less_cancel_left_disj mult_commute [of _ a]) + done + moreover with `c \ 0` divmod_rel_div_mod have "divmod_rel b c (b div c, b mod c)" by auto + ultimately have "divmod_rel (a * b) (a * c) (b div c, a * (b mod c))" . + moreover from `a \ 0` `c \ 0` have "a * c \ 0" by simp + ultimately show ?thesis by (rule divmod_rel_div) + qed qed auto lemma posDivAlg_div_mod: @@ -808,52 +828,6 @@ done -subsection{*Cancellation of Common Factors in div*} - -lemma zdiv_zmult_zmult1_aux1: - "[| (0::int) < b; c \ 0 |] ==> (c*a) div (c*b) = a div b" -by (subst zdiv_zmult2_eq, auto) - -lemma zdiv_zmult_zmult1_aux2: - "[| b < (0::int); c \ 0 |] ==> (c*a) div (c*b) = a div b" -apply (subgoal_tac " (c * (-a)) div (c * (-b)) = (-a) div (-b) ") -apply (rule_tac [2] zdiv_zmult_zmult1_aux1, auto) -done - -lemma zdiv_zmult_zmult1: "c \ (0::int) ==> (c*a) div (c*b) = a div b" -apply (case_tac "b = 0", simp) -apply (auto simp add: linorder_neq_iff zdiv_zmult_zmult1_aux1 zdiv_zmult_zmult1_aux2) -done - -lemma zdiv_zmult_zmult1_if[simp]: - "(k*m) div (k*n) = (if k = (0::int) then 0 else m div n)" -by (simp add:zdiv_zmult_zmult1) - - -subsection{*Distribution of Factors over mod*} - -lemma zmod_zmult_zmult1_aux1: - "[| (0::int) < b; c \ 0 |] ==> (c*a) mod (c*b) = c * (a mod b)" -by (subst zmod_zmult2_eq, auto) - -lemma zmod_zmult_zmult1_aux2: - "[| b < (0::int); c \ 0 |] ==> (c*a) mod (c*b) = c * (a mod b)" -apply (subgoal_tac " (c * (-a)) mod (c * (-b)) = c * ((-a) mod (-b))") -apply (rule_tac [2] zmod_zmult_zmult1_aux1, auto) -done - -lemma zmod_zmult_zmult1: "(c*a) mod (c*b) = (c::int) * (a mod b)" -apply (case_tac "b = 0", simp) -apply (case_tac "c = 0", simp) -apply (auto simp add: linorder_neq_iff zmod_zmult_zmult1_aux1 zmod_zmult_zmult1_aux2) -done - -lemma zmod_zmult_zmult2: "(a*c) mod (b*c) = (a mod b) * (c::int)" -apply (cut_tac c = c in zmod_zmult_zmult1) -apply (auto simp add: mult_commute) -done - - subsection {*Splitting Rules for div and mod*} text{*The proofs of the two lemmas below are essentially identical*} @@ -937,7 +911,7 @@ right_distrib) thus ?thesis by (subst zdiv_zadd1_eq, - simp add: zdiv_zmult_zmult1 zmod_zmult_zmult1 one_less_a2 + simp add: mod_mult_mult1 one_less_a2 div_pos_pos_trivial) qed @@ -961,7 +935,7 @@ then number_of v div (number_of w) else (number_of v + (1::int)) div (number_of w))" apply (simp only: number_of_eq numeral_simps UNIV_I split: split_if) -apply (simp add: zdiv_zmult_zmult1 pos_zdiv_mult_2 neg_zdiv_mult_2 add_ac) +apply (simp add: pos_zdiv_mult_2 neg_zdiv_mult_2 add_ac) done @@ -977,7 +951,7 @@ apply (auto simp add: add_commute [of 1] mult_commute add1_zle_eq pos_mod_bound) apply (subst mod_add_eq) -apply (simp add: zmod_zmult_zmult2 mod_pos_pos_trivial) +apply (simp add: mod_mult_mult2 mod_pos_pos_trivial) apply (rule mod_pos_pos_trivial) apply (auto simp add: mod_pos_pos_trivial ring_distribs) apply (subgoal_tac "0 \ b mod a", arith, simp) @@ -998,7 +972,7 @@ "number_of (Int.Bit0 v) mod number_of (Int.Bit0 w) = (2::int) * (number_of v mod number_of w)" apply (simp only: number_of_eq numeral_simps) -apply (simp add: zmod_zmult_zmult1 pos_zmod_mult_2 +apply (simp add: mod_mult_mult1 pos_zmod_mult_2 neg_zmod_mult_2 add_ac) done @@ -1008,7 +982,7 @@ then 2 * (number_of v mod number_of w) + 1 else 2 * ((number_of v + (1::int)) mod number_of w) - 1)" apply (simp only: number_of_eq numeral_simps) -apply (simp add: zmod_zmult_zmult1 pos_zmod_mult_2 +apply (simp add: mod_mult_mult1 pos_zmod_mult_2 neg_zmod_mult_2 add_ac) done @@ -1090,9 +1064,7 @@ done lemma zdvd_zmod: "f dvd m ==> f dvd (n::int) ==> f dvd m mod n" - apply (simp add: dvd_def) - apply (auto simp add: zmod_zmult_zmult1) - done + by (auto elim!: dvdE simp add: mod_mult_mult1) lemma zdvd_zmod_imp_zdvd: "k dvd m mod n ==> k dvd n ==> k dvd (m::int)" apply (subgoal_tac "k dvd n * (m div n) + m mod n") @@ -1247,9 +1219,9 @@ lemmas zmod_simps = mod_add_left_eq [symmetric] mod_add_right_eq [symmetric] - IntDiv.zmod_zmult1_eq [symmetric] - mod_mult_left_eq [symmetric] - IntDiv.zpower_zmod + zmod_zmult1_eq [symmetric] + mod_mult_left_eq [symmetric] + zpower_zmod zminus_zmod zdiff_zmod_left zdiff_zmod_right text {* Distributive laws for function @{text nat}. *} diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/IsaMakefile Mon Apr 27 07:26:17 2009 -0700 @@ -89,7 +89,7 @@ $(SRC)/Tools/IsaPlanner/rw_tools.ML \ $(SRC)/Tools/IsaPlanner/zipper.ML \ $(SRC)/Tools/atomize_elim.ML \ - $(SRC)/Tools/code/code_funcgr.ML \ + $(SRC)/Tools/auto_solve.ML \ $(SRC)/Tools/code/code_haskell.ML \ $(SRC)/Tools/code/code_ml.ML \ $(SRC)/Tools/code/code_name.ML \ @@ -103,10 +103,11 @@ $(SRC)/Tools/intuitionistic.ML \ $(SRC)/Tools/induct_tacs.ML \ $(SRC)/Tools/nbe.ML \ + $(SRC)/Tools/quickcheck.ML \ $(SRC)/Tools/project_rule.ML \ $(SRC)/Tools/random_word.ML \ $(SRC)/Tools/value.ML \ - Code_Setup.thy \ + $(SRC)/Tools/Code_Generator.thy \ HOL.thy \ Tools/hologic.ML \ Tools/recfun_codegen.ML \ @@ -216,10 +217,9 @@ List.thy \ Main.thy \ Map.thy \ - NatBin.thy \ + Nat_Numeral.thy \ Presburger.thy \ Recdef.thy \ - Relation_Power.thy \ SetInterval.thy \ $(SRC)/Provers/Arith/assoc_fold.ML \ $(SRC)/Provers/Arith/cancel_numeral_factor.ML \ diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Library/Code_Index.thy --- a/src/HOL/Library/Code_Index.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Library/Code_Index.thy Mon Apr 27 07:26:17 2009 -0700 @@ -144,7 +144,7 @@ subsection {* Basic arithmetic *} -instantiation index :: "{minus, ordered_semidom, Divides.div, linorder}" +instantiation index :: "{minus, ordered_semidom, semiring_div, linorder}" begin definition [simp, code del]: @@ -172,7 +172,7 @@ "n < m \ nat_of n < nat_of m" instance proof -qed (auto simp add: left_distrib) +qed (auto simp add: index left_distrib div_mult_self1) end diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Library/Coinductive_List.thy --- a/src/HOL/Library/Coinductive_List.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Library/Coinductive_List.thy Mon Apr 27 07:26:17 2009 -0700 @@ -786,7 +786,7 @@ lemma funpow_lmap: fixes f :: "'a \ 'a" - shows "(lmap f ^ n) (LCons b l) = LCons ((f ^ n) b) ((lmap f ^ n) l)" + shows "(lmap f ^^ n) (LCons b l) = LCons ((f ^^ n) b) ((lmap f ^^ n) l)" by (induct n) simp_all @@ -796,35 +796,35 @@ proof fix x have "(h x, iterates f x) \ - {((lmap f ^ n) (h u), (lmap f ^ n) (iterates f u)) | u n. True}" + {((lmap f ^^ n) (h u), (lmap f ^^ n) (iterates f u)) | u n. True}" proof - - have "(h x, iterates f x) = ((lmap f ^ 0) (h x), (lmap f ^ 0) (iterates f x))" + have "(h x, iterates f x) = ((lmap f ^^ 0) (h x), (lmap f ^^ 0) (iterates f x))" by simp then show ?thesis by blast qed then show "h x = iterates f x" proof (coinduct rule: llist_equalityI) case (Eqllist q) - then obtain u n where "q = ((lmap f ^ n) (h u), (lmap f ^ n) (iterates f u))" + then obtain u n where "q = ((lmap f ^^ n) (h u), (lmap f ^^ n) (iterates f u))" (is "_ = (?q1, ?q2)") by auto - also have "?q1 = LCons ((f ^ n) u) ((lmap f ^ Suc n) (h u))" + also have "?q1 = LCons ((f ^^ n) u) ((lmap f ^^ Suc n) (h u))" proof - - have "?q1 = (lmap f ^ n) (LCons u (lmap f (h u)))" + have "?q1 = (lmap f ^^ n) (LCons u (lmap f (h u)))" by (subst h) rule - also have "\ = LCons ((f ^ n) u) ((lmap f ^ n) (lmap f (h u)))" + also have "\ = LCons ((f ^^ n) u) ((lmap f ^^ n) (lmap f (h u)))" by (rule funpow_lmap) - also have "(lmap f ^ n) (lmap f (h u)) = (lmap f ^ Suc n) (h u)" + also have "(lmap f ^^ n) (lmap f (h u)) = (lmap f ^^ Suc n) (h u)" by (simp add: funpow_swap1) finally show ?thesis . qed - also have "?q2 = LCons ((f ^ n) u) ((lmap f ^ Suc n) (iterates f u))" + also have "?q2 = LCons ((f ^^ n) u) ((lmap f ^^ Suc n) (iterates f u))" proof - - have "?q2 = (lmap f ^ n) (LCons u (iterates f (f u)))" + have "?q2 = (lmap f ^^ n) (LCons u (iterates f (f u)))" by (subst iterates) rule - also have "\ = LCons ((f ^ n) u) ((lmap f ^ n) (iterates f (f u)))" + also have "\ = LCons ((f ^^ n) u) ((lmap f ^^ n) (iterates f (f u)))" by (rule funpow_lmap) - also have "(lmap f ^ n) (iterates f (f u)) = (lmap f ^ Suc n) (iterates f u)" + also have "(lmap f ^^ n) (iterates f (f u)) = (lmap f ^^ Suc n) (iterates f u)" by (simp add: lmap_iterates funpow_swap1) finally show ?thesis . qed diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Library/Continuity.thy --- a/src/HOL/Library/Continuity.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Library/Continuity.thy Mon Apr 27 07:26:17 2009 -0700 @@ -5,7 +5,7 @@ header {* Continuity and iterations (of set transformers) *} theory Continuity -imports Relation_Power Main +imports Transitive_Closure Main begin subsection {* Continuity for complete lattices *} @@ -48,25 +48,25 @@ qed lemma continuous_lfp: - assumes "continuous F" shows "lfp F = (SUP i. (F^i) bot)" + assumes "continuous F" shows "lfp F = (SUP i. (F ^^ i) bot)" proof - note mono = continuous_mono[OF `continuous F`] - { fix i have "(F^i) bot \ lfp F" + { fix i have "(F ^^ i) bot \ lfp F" proof (induct i) - show "(F^0) bot \ lfp F" by simp + show "(F ^^ 0) bot \ lfp F" by simp next case (Suc i) - have "(F^(Suc i)) bot = F((F^i) bot)" by simp + have "(F ^^ Suc i) bot = F((F ^^ i) bot)" by simp also have "\ \ F(lfp F)" by(rule monoD[OF mono Suc]) also have "\ = lfp F" by(simp add:lfp_unfold[OF mono, symmetric]) finally show ?case . qed } - hence "(SUP i. (F^i) bot) \ lfp F" by (blast intro!:SUP_leI) - moreover have "lfp F \ (SUP i. (F^i) bot)" (is "_ \ ?U") + hence "(SUP i. (F ^^ i) bot) \ lfp F" by (blast intro!:SUP_leI) + moreover have "lfp F \ (SUP i. (F ^^ i) bot)" (is "_ \ ?U") proof (rule lfp_lowerbound) - have "chain(%i. (F^i) bot)" + have "chain(%i. (F ^^ i) bot)" proof - - { fix i have "(F^i) bot \ (F^(Suc i)) bot" + { fix i have "(F ^^ i) bot \ (F ^^ (Suc i)) bot" proof (induct i) case 0 show ?case by simp next @@ -74,7 +74,7 @@ qed } thus ?thesis by(auto simp add:chain_def) qed - hence "F ?U = (SUP i. (F^(i+1)) bot)" using `continuous F` by (simp add:continuous_def) + hence "F ?U = (SUP i. (F ^^ (i+1)) bot)" using `continuous F` by (simp add:continuous_def) also have "\ \ ?U" by(fast intro:SUP_leI le_SUPI) finally show "F ?U \ ?U" . qed @@ -193,7 +193,7 @@ definition up_iterate :: "('a set => 'a set) => nat => 'a set" where - "up_iterate f n = (f^n) {}" + "up_iterate f n = (f ^^ n) {}" lemma up_iterate_0 [simp]: "up_iterate f 0 = {}" by (simp add: up_iterate_def) @@ -245,7 +245,7 @@ definition down_iterate :: "('a set => 'a set) => nat => 'a set" where - "down_iterate f n = (f^n) UNIV" + "down_iterate f n = (f ^^ n) UNIV" lemma down_iterate_0 [simp]: "down_iterate f 0 = UNIV" by (simp add: down_iterate_def) diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Library/Euclidean_Space.thy --- a/src/HOL/Library/Euclidean_Space.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Library/Euclidean_Space.thy Mon Apr 27 07:26:17 2009 -0700 @@ -253,12 +253,7 @@ "vector_power x 0 = 1" | "vector_power x (Suc n) = x * vector_power x n" -instantiation "^" :: (recpower,type) recpower -begin - definition vec_power_def: "op ^ \ vector_power" - instance - apply (intro_classes) by (simp_all add: vec_power_def) -end +instance "^" :: (recpower,type) recpower .. instance "^" :: (semiring,type) semiring apply (intro_classes) by (vector ring_simps)+ diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Library/Eval_Witness.thy --- a/src/HOL/Library/Eval_Witness.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Library/Eval_Witness.thy Mon Apr 27 07:26:17 2009 -0700 @@ -68,7 +68,7 @@ | dest_exs _ _ = sys_error "dest_exs"; val t = dest_exs (length ws) (HOLogic.dest_Trueprop goal); in - if Code_ML.eval_term ("Eval_Witness_Method.eval_ref", Eval_Witness_Method.eval_ref) thy t ws + if Code_ML.eval NONE ("Eval_Witness_Method.eval_ref", Eval_Witness_Method.eval_ref) (K I) thy t ws then Thm.cterm_of thy goal else @{cprop True} (*dummy*) end diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Library/Float.thy --- a/src/HOL/Library/Float.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Library/Float.thy Mon Apr 27 07:26:17 2009 -0700 @@ -15,8 +15,8 @@ datatype float = Float int int -fun Ifloat :: "float \ real" where -"Ifloat (Float a b) = real a * pow2 b" +primrec Ifloat :: "float \ real" where + "Ifloat (Float a b) = real a * pow2 b" instantiation float :: zero begin definition zero_float where "0 = Float 0 0" @@ -33,11 +33,11 @@ instance .. end -fun mantissa :: "float \ int" where -"mantissa (Float a b) = a" +primrec mantissa :: "float \ int" where + "mantissa (Float a b) = a" -fun scale :: "float \ int" where -"scale (Float a b) = b" +primrec scale :: "float \ int" where + "scale (Float a b) = b" lemma Ifloat_neg_exp: "e < 0 \ Ifloat (Float m e) = real m * inverse (2^nat (-e))" by auto lemma Ifloat_nge0_exp: "\ 0 \ e \ Ifloat (Float m e) = real m * inverse (2^nat (-e))" by auto @@ -320,12 +320,12 @@ end instantiation float :: uminus begin -fun uminus_float where [simp del]: "uminus_float (Float m e) = Float (-m) e" +primrec uminus_float where [simp del]: "uminus_float (Float m e) = Float (-m) e" instance .. end instantiation float :: minus begin -fun minus_float where [simp del]: "(z::float) - w = z + (- w)" +definition minus_float where [simp del]: "(z::float) - w = z + (- w)" instance .. end @@ -334,11 +334,11 @@ instance .. end -fun float_pprt :: "float \ float" where -"float_pprt (Float a e) = (if 0 <= a then (Float a e) else 0)" +primrec float_pprt :: "float \ float" where + "float_pprt (Float a e) = (if 0 <= a then (Float a e) else 0)" -fun float_nprt :: "float \ float" where -"float_nprt (Float a e) = (if 0 <= a then 0 else (Float a e))" +primrec float_nprt :: "float \ float" where + "float_nprt (Float a e) = (if 0 <= a then 0 else (Float a e))" instantiation float :: ord begin definition le_float_def: "z \ w \ Ifloat z \ Ifloat w" @@ -354,7 +354,7 @@ by (cases a, simp add: uminus_float.simps) lemma Ifloat_sub[simp]: "Ifloat (a - b) = Ifloat a - Ifloat b" - by (cases a, cases b, simp add: minus_float.simps) + by (cases a, cases b, simp add: minus_float_def) lemma Ifloat_mult[simp]: "Ifloat (a*b) = Ifloat a * Ifloat b" by (cases a, cases b, simp add: times_float.simps pow2_add) @@ -443,37 +443,10 @@ lemma Ifloat_min: "Ifloat (min x y) = min (Ifloat x) (Ifloat y)" unfolding min_def le_float_def by auto lemma Ifloat_max: "Ifloat (max a b) = max (Ifloat a) (Ifloat b)" unfolding max_def le_float_def by auto -instantiation float :: power begin -fun power_float where [simp del]: "(Float m e) ^ n = Float (m ^ n) (e * int n)" -instance .. -end - -instance float :: recpower -proof (intro_classes) - fix a :: float show "a ^ 0 = 1" by (cases a, auto simp add: power_float.simps one_float_def) -next - fix a :: float and n :: nat show "a ^ (Suc n) = a * a ^ n" - by (cases a, auto simp add: power_float.simps times_float.simps algebra_simps) -qed +instance float :: recpower .. -lemma float_power: "Ifloat (x ^ n) = (Ifloat x) ^ n" -proof (cases x) - case (Float m e) - - have "pow2 e ^ n = pow2 (e * int n)" - proof (cases "e >= 0") - case True hence e_nat: "e = int (nat e)" by auto - hence "pow2 e ^ n = (2 ^ nat e) ^ n" using pow2_int[of "nat e"] by auto - thus ?thesis unfolding power_mult[symmetric] unfolding pow2_int[symmetric] int_mult e_nat[symmetric] . - next - case False hence e_minus: "-e = int (nat (-e))" by auto - hence "pow2 (-e) ^ n = (2 ^ nat (-e)) ^ n" using pow2_int[of "nat (-e)"] by auto - hence "pow2 (-e) ^ n = pow2 ((-e) * int n)" unfolding power_mult[symmetric] unfolding pow2_int[symmetric] int_mult e_minus[symmetric] zmult_zminus . - thus ?thesis unfolding pow2_neg[of "-e"] pow2_neg[of "-e * int n"] unfolding zmult_zminus zminus_zminus nonzero_power_inverse[OF pow2_neq_zero, symmetric] - using nonzero_inverse_eq_imp_eq[OF _ pow2_neq_zero pow2_neq_zero] by auto - qed - thus ?thesis by (auto simp add: Float power_mult_distrib Ifloat.simps power_float.simps) -qed +lemma float_power: "Ifloat (x ^ n) = Ifloat x ^ n" + by (induct n) simp_all lemma zero_le_pow2[simp]: "0 \ pow2 s" apply (subgoal_tac "0 < pow2 s") @@ -1182,12 +1155,12 @@ unfolding x_eq y_eq float_divr.simps Let_def le_float_def Ifloat_0 Ifloat_mult by (auto intro!: mult_nonneg_nonpos) qed -fun round_down :: "nat \ float \ float" where +primrec round_down :: "nat \ float \ float" where "round_down prec (Float m e) = (let d = bitlen m - int prec in if 0 < d then let P = 2^nat d ; n = m div P in Float n (e + d) else Float m e)" -fun round_up :: "nat \ float \ float" where +primrec round_up :: "nat \ float \ float" where "round_up prec (Float m e) = (let d = bitlen m - int prec in if 0 < d then let P = 2^nat d ; n = m div P ; r = m mod P in Float (n + (if r = 0 then 0 else 1)) (e + d) else Float m e)" @@ -1314,8 +1287,8 @@ finally show ?thesis . qed -fun float_abs :: "float \ float" where -"float_abs (Float m e) = Float \m\ e" +primrec float_abs :: "float \ float" where + "float_abs (Float m e) = Float \m\ e" instantiation float :: abs begin definition abs_float_def: "\x\ = float_abs x" @@ -1329,8 +1302,8 @@ thus ?thesis unfolding Float abs_float_def float_abs.simps Ifloat.simps by auto qed -fun floor_fl :: "float \ float" where -"floor_fl (Float m e) = (if 0 \ e then Float m e +primrec floor_fl :: "float \ float" where + "floor_fl (Float m e) = (if 0 \ e then Float m e else Float (m div (2 ^ (nat (-e)))) 0)" lemma floor_fl: "Ifloat (floor_fl x) \ Ifloat x" @@ -1358,8 +1331,8 @@ declare floor_fl.simps[simp del] -fun ceiling_fl :: "float \ float" where -"ceiling_fl (Float m e) = (if 0 \ e then Float m e +primrec ceiling_fl :: "float \ float" where + "ceiling_fl (Float m e) = (if 0 \ e then Float m e else Float (m div (2 ^ (nat (-e))) + 1) 0)" lemma ceiling_fl: "Ifloat x \ Ifloat (ceiling_fl x)" diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Library/Formal_Power_Series.thy --- a/src/HOL/Library/Formal_Power_Series.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Library/Formal_Power_Series.thy Mon Apr 27 07:26:17 2009 -0700 @@ -680,30 +680,14 @@ subsection {* Powers*} -instantiation fps :: (semiring_1) power -begin - -fun fps_pow :: "nat \ 'a fps \ 'a fps" where - "fps_pow 0 f = 1" -| "fps_pow (Suc n) f = f * fps_pow n f" - -definition fps_power_def: "power (f::'a fps) n = fps_pow n f" -instance .. -end - -instantiation fps :: (comm_ring_1) recpower -begin -instance - apply (intro_classes) - by (simp_all add: fps_power_def) -end +instance fps :: (semiring_1) recpower .. lemma fps_power_zeroth_eq_one: "a$0 =1 \ a^n $ 0 = (1::'a::semiring_1)" - by (induct n, auto simp add: fps_power_def expand_fps_eq fps_mult_nth) + by (induct n, auto simp add: expand_fps_eq fps_mult_nth) lemma fps_power_first_eq: "(a:: 'a::comm_ring_1 fps)$0 =1 \ a^n $ 1 = of_nat n * a$1" proof(induct n) - case 0 thus ?case by (simp add: fps_power_def) + case 0 thus ?case by simp next case (Suc n) note h = Suc.hyps[OF `a$0 = 1`] @@ -712,13 +696,13 @@ qed lemma startsby_one_power:"a $ 0 = (1::'a::comm_ring_1) \ a^n $ 0 = 1" - by (induct n, auto simp add: fps_power_def fps_mult_nth) + by (induct n, auto simp add: fps_mult_nth) lemma startsby_zero_power:"a $0 = (0::'a::comm_ring_1) \ n > 0 \ a^n $0 = 0" - by (induct n, auto simp add: fps_power_def fps_mult_nth) + by (induct n, auto simp add: fps_mult_nth) lemma startsby_power:"a $0 = (v::'a::{comm_ring_1, recpower}) \ a^n $0 = v^n" - by (induct n, auto simp add: fps_power_def fps_mult_nth power_Suc) + by (induct n, auto simp add: fps_mult_nth power_Suc) lemma startsby_zero_power_iff[simp]: "a^n $0 = (0::'a::{idom, recpower}) \ (n \ 0 \ a$0 = 0)" @@ -901,7 +885,7 @@ lemma X_power_iff: "X^k = Abs_fps (\n. if n = k then (1::'a::comm_ring_1) else 0)" proof(induct k) - case 0 thus ?case by (simp add: X_def fps_power_def fps_eq_iff) + case 0 thus ?case by (simp add: X_def fps_eq_iff) next case (Suc k) {fix m @@ -979,7 +963,7 @@ (* {a_{n+k}}_0^infty Corresponds to (f - setsum (\i. a_i * x^i))/x^h, for h>0*) lemma fps_power_mult_eq_shift: - "X^Suc k * Abs_fps (\n. a (n + Suc k)) = Abs_fps a - setsum (\i. fps_const (a i :: 'a:: field) * X^i) {0 .. k}" (is "?lhs = ?rhs") + "X^Suc k * Abs_fps (\n. a (n + Suc k)) = Abs_fps a - setsum (\i. fps_const (a i :: 'a:: comm_ring_1) * X^i) {0 .. k}" (is "?lhs = ?rhs") proof- {fix n:: nat have "?lhs $ n = (if n < Suc k then 0 else a n)" @@ -990,7 +974,7 @@ next case (Suc k) note th = Suc.hyps[symmetric] - have "(Abs_fps a - setsum (\i. fps_const (a i :: 'a:: field) * X^i) {0 .. Suc k})$n = (Abs_fps a - setsum (\i. fps_const (a i :: 'a:: field) * X^i) {0 .. k} - fps_const (a (Suc k)) * X^ Suc k) $ n" by (simp add: ring_simps) + have "(Abs_fps a - setsum (\i. fps_const (a i :: 'a) * X^i) {0 .. Suc k})$n = (Abs_fps a - setsum (\i. fps_const (a i :: 'a) * X^i) {0 .. k} - fps_const (a (Suc k)) * X^ Suc k) $ n" by (simp add: ring_simps) also have "\ = (if n < Suc k then 0 else a n) - (fps_const (a (Suc k)) * X^ Suc k)$n" using th unfolding fps_sub_nth by simp @@ -1022,13 +1006,16 @@ lemma XD_linear[simp]: "XD (fps_const c * a + fps_const d * b) = fps_const c * XD a + fps_const d * XD (b :: ('a::comm_ring_1) fps)" by simp -lemma XDN_linear: "(XD^n) (fps_const c * a + fps_const d * b) = fps_const c * (XD^n) a + fps_const d * (XD^n) (b :: ('a::comm_ring_1) fps)" +lemma XDN_linear: + "(XD ^^ n) (fps_const c * a + fps_const d * b) = fps_const c * (XD ^^ n) a + fps_const d * (XD ^^ n) (b :: ('a::comm_ring_1) fps)" by (induct n, simp_all) lemma fps_mult_X_deriv_shift: "X* fps_deriv a = Abs_fps (\n. of_nat n* a$n)" by (simp add: fps_eq_iff) -lemma fps_mult_XD_shift: "(XD ^k) (a:: ('a::{comm_ring_1, recpower, ring_char_0}) fps) = Abs_fps (\n. (of_nat n ^ k) * a$n)" -by (induct k arbitrary: a) (simp_all add: power_Suc XD_def fps_eq_iff ring_simps del: One_nat_def) + +lemma fps_mult_XD_shift: + "(XD ^^ k) (a:: ('a::{comm_ring_1, recpower}) fps) = Abs_fps (\n. (of_nat n ^ k) * a$n)" + by (induct k arbitrary: a) (simp_all add: power_Suc XD_def fps_eq_iff ring_simps del: One_nat_def) subsubsection{* Rule 3 is trivial and is given by @{text fps_times_def}*} subsubsection{* Rule 5 --- summation and "division" by (1 - X)*} @@ -1901,19 +1888,16 @@ done lemma fps_compose_1[simp]: "1 oo a = 1" - by (simp add: fps_eq_iff fps_compose_nth fps_power_def mult_delta_left setsum_delta) + by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum_delta) lemma fps_compose_0[simp]: "0 oo a = 0" by (simp add: fps_eq_iff fps_compose_nth) -lemma fps_pow_0: "fps_pow n 0 = (if n = 0 then 1 else 0)" - by (induct n, simp_all) - lemma fps_compose_0_right[simp]: "a oo 0 = fps_const (a$0)" - by (auto simp add: fps_eq_iff fps_compose_nth fps_power_def fps_pow_0 setsum_0') + by (auto simp add: fps_eq_iff fps_compose_nth power_0_left setsum_0') lemma fps_compose_add_distrib: "(a + b) oo c = (a oo c) + (b oo c)" - by (simp add: fps_eq_iff fps_compose_nth ring_simps setsum_addf) + by (simp add: fps_eq_iff fps_compose_nth ring_simps setsum_addf) lemma fps_compose_setsum_distrib: "(setsum f S) oo a = setsum (\i. f i oo a) S" proof- @@ -2341,11 +2325,11 @@ proof- have "fps_deriv ?lhs = 0" apply (simp add: fps_deriv_power fps_sin_deriv fps_cos_deriv power_Suc) - by (simp add: fps_power_def ring_simps fps_const_neg[symmetric] del: fps_const_neg) + by (simp add: ring_simps fps_const_neg[symmetric] del: fps_const_neg) then have "?lhs = fps_const (?lhs $ 0)" unfolding fps_deriv_eq_0_iff . also have "\ = 1" - by (auto simp add: fps_eq_iff fps_power_def numeral_2_eq_2 fps_mult_nth fps_cos_def fps_sin_def) + by (auto simp add: fps_eq_iff numeral_2_eq_2 fps_mult_nth fps_cos_def fps_sin_def) finally show ?thesis . qed diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Library/Numeral_Type.thy --- a/src/HOL/Library/Numeral_Type.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Library/Numeral_Type.thy Mon Apr 27 07:26:17 2009 -0700 @@ -154,8 +154,8 @@ locale mod_type = fixes n :: int - and Rep :: "'a::{zero,one,plus,times,uminus,minus,power} \ int" - and Abs :: "int \ 'a::{zero,one,plus,times,uminus,minus,power}" + and Rep :: "'a::{zero,one,plus,times,uminus,minus} \ int" + and Abs :: "int \ 'a::{zero,one,plus,times,uminus,minus}" assumes type: "type_definition Rep Abs {0.. int" - and Abs :: "int \ 'a::{number_ring,power}" + and Rep :: "'a::{number_ring} \ int" + and Abs :: "int \ 'a::{number_ring}" begin lemma of_nat_eq: "of_nat k = Abs (int k mod n)" @@ -272,7 +271,7 @@ @{typ num1}, since 0 and 1 are not distinct. *} -instantiation num1 :: "{comm_ring,comm_monoid_mult,number,recpower}" +instantiation num1 :: "{comm_ring,comm_monoid_mult,number}" begin lemma num1_eq_iff: "(x::num1) = (y::num1) \ True" @@ -284,7 +283,7 @@ end instantiation - bit0 and bit1 :: (finite) "{zero,one,plus,times,uminus,minus,power}" + bit0 and bit1 :: (finite) "{zero,one,plus,times,uminus,minus}" begin definition Abs_bit0' :: "int \ 'a bit0" where @@ -299,7 +298,6 @@ definition "x * y = Abs_bit0' (Rep_bit0 x * Rep_bit0 y)" definition "x - y = Abs_bit0' (Rep_bit0 x - Rep_bit0 y)" definition "- x = Abs_bit0' (- Rep_bit0 x)" -definition "x ^ k = Abs_bit0' (Rep_bit0 x ^ k)" definition "0 = Abs_bit1 0" definition "1 = Abs_bit1 1" @@ -307,7 +305,6 @@ definition "x * y = Abs_bit1' (Rep_bit1 x * Rep_bit1 y)" definition "x - y = Abs_bit1' (Rep_bit1 x - Rep_bit1 y)" definition "- x = Abs_bit1' (- Rep_bit1 x)" -definition "x ^ k = Abs_bit1' (Rep_bit1 x ^ k)" instance .. @@ -326,7 +323,6 @@ apply (rule times_bit0_def [unfolded Abs_bit0'_def]) apply (rule minus_bit0_def [unfolded Abs_bit0'_def]) apply (rule uminus_bit0_def [unfolded Abs_bit0'_def]) -apply (rule power_bit0_def [unfolded Abs_bit0'_def]) done interpretation bit1: @@ -342,7 +338,6 @@ apply (rule times_bit1_def [unfolded Abs_bit1'_def]) apply (rule minus_bit1_def [unfolded Abs_bit1'_def]) apply (rule uminus_bit1_def [unfolded Abs_bit1'_def]) -apply (rule power_bit1_def [unfolded Abs_bit1'_def]) done instance bit0 :: (finite) "{comm_ring_1,recpower}" @@ -386,9 +381,6 @@ lemmas bit0_iszero_number_of [simp] = bit0.iszero_number_of lemmas bit1_iszero_number_of [simp] = bit1.iszero_number_of -declare power_Suc [where ?'a="'a::finite bit0", standard, simp] -declare power_Suc [where ?'a="'a::finite bit1", standard, simp] - subsection {* Syntax *} diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Library/Polynomial.thy --- a/src/HOL/Library/Polynomial.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Library/Polynomial.thy Mon Apr 27 07:26:17 2009 -0700 @@ -632,19 +632,7 @@ shows "a \ 0 \ p dvd smult a q \ p dvd q" by (safe elim!: dvd_smult dvd_smult_cancel) -instantiation poly :: (comm_semiring_1) recpower -begin - -primrec power_poly where - "(p::'a poly) ^ 0 = 1" -| "(p::'a poly) ^ (Suc n) = p * p ^ n" - -instance - by default simp_all - -declare power_poly.simps [simp del] - -end +instance poly :: (comm_semiring_1) recpower .. lemma degree_power_le: "degree (p ^ n) \ degree p * n" by (induct n, simp, auto intro: order_trans degree_mult_le) @@ -987,6 +975,30 @@ by (simp add: pdivmod_rel_def left_distrib) thus "(x + z * y) div y = z + x div y" by (rule div_poly_eq) +next + fix x y z :: "'a poly" + assume "x \ 0" + show "(x * y) div (x * z) = y div z" + proof (cases "y \ 0 \ z \ 0") + have "\x::'a poly. pdivmod_rel x 0 0 x" + by (rule pdivmod_rel_by_0) + then have [simp]: "\x::'a poly. x div 0 = 0" + by (rule div_poly_eq) + have "\x::'a poly. pdivmod_rel 0 x 0 0" + by (rule pdivmod_rel_0) + then have [simp]: "\x::'a poly. 0 div x = 0" + by (rule div_poly_eq) + case False then show ?thesis by auto + next + case True then have "y \ 0" and "z \ 0" by auto + with `x \ 0` + have "\q r. pdivmod_rel y z q r \ pdivmod_rel (x * y) (x * z) q (x * r)" + by (auto simp add: pdivmod_rel_def algebra_simps) + (rule classical, simp add: degree_mult_eq) + moreover from pdivmod_rel have "pdivmod_rel y z (y div z) (y mod z)" . + ultimately have "pdivmod_rel (x * y) (x * z) (y div z) (x * (y mod z))" . + then show ?thesis by (simp add: div_poly_eq) + qed qed end diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Library/Quickcheck.thy --- a/src/HOL/Library/Quickcheck.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Library/Quickcheck.thy Mon Apr 27 07:26:17 2009 -0700 @@ -47,6 +47,8 @@ val eval_ref : (unit -> int -> int * int -> term list option * (int * int)) option ref = ref NONE; +val target = "Quickcheck"; + fun mk_generator_expr thy prop tys = let val bound_max = length tys - 1; @@ -72,14 +74,75 @@ let val tys = (map snd o fst o strip_abs) t; val t' = mk_generator_expr thy t tys; - val f = Code_ML.eval_term ("Quickcheck.eval_ref", eval_ref) thy t' []; - in f #> Random_Engine.run #> (Option.map o map) (Code.postprocess_term thy) end; + val f = Code_ML.eval (SOME target) ("Quickcheck.eval_ref", eval_ref) + (fn proc => fn g => fn s => g s #>> (Option.map o map) proc) thy t' []; + in f #> Random_Engine.run end; end *} setup {* - Quickcheck.add_generator ("code", Quickcheck.compile_generator_expr o ProofContext.theory_of) + Code_Target.extend_target (Quickcheck.target, (Code_ML.target_Eval, K I)) + #> Quickcheck.add_generator ("code", Quickcheck.compile_generator_expr o ProofContext.theory_of) *} + +subsection {* Type @{typ "'a \ 'b"} *} + +ML {* +structure Random_Engine = +struct + +open Random_Engine; + +fun random_fun (T1 : typ) (T2 : typ) (eq : 'a -> 'a -> bool) (term_of : 'a -> term) + (random : Random_Engine.seed -> ('b * (unit -> term)) * Random_Engine.seed) + (random_split : Random_Engine.seed -> Random_Engine.seed * Random_Engine.seed) + (seed : Random_Engine.seed) = + let + val (seed', seed'') = random_split seed; + val state = ref (seed', [], Const (@{const_name undefined}, T1 --> T2)); + val fun_upd = Const (@{const_name fun_upd}, + (T1 --> T2) --> T1 --> T2 --> T1 --> T2); + fun random_fun' x = + let + val (seed, fun_map, f_t) = ! state; + in case AList.lookup (uncurry eq) fun_map x + of SOME y => y + | NONE => let + val t1 = term_of x; + val ((y, t2), seed') = random seed; + val fun_map' = (x, y) :: fun_map; + val f_t' = fun_upd $ f_t $ t1 $ t2 (); + val _ = state := (seed', fun_map', f_t'); + in y end + end; + fun term_fun' () = #3 (! state); + in ((random_fun', term_fun'), seed'') end; + end +*} + +axiomatization + random_fun_aux :: "typerep \ typerep \ ('a \ 'a \ bool) \ ('a \ term) + \ (seed \ ('b \ (unit \ term)) \ seed) \ (seed \ seed \ seed) + \ seed \ (('a \ 'b) \ (unit \ term)) \ seed" + +code_const random_fun_aux (Quickcheck "Random'_Engine.random'_fun") + -- {* With enough criminal energy this can be abused to derive @{prop False}; + for this reason we use a distinguished target @{text Quickcheck} + not spoiling the regular trusted code generation *} + +instantiation "fun" :: ("{eq, term_of}", "{type, random}") random +begin + +definition random_fun :: "index \ seed \ (('a \ 'b) \ (unit \ term)) \ seed" where + "random n = random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Eval.term_of (random n) split_seed" + +instance .. + +end + +code_reserved Quickcheck Random_Engine + +end diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Library/Topology_Euclidean_Space.thy --- a/src/HOL/Library/Topology_Euclidean_Space.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Library/Topology_Euclidean_Space.thy Mon Apr 27 07:26:17 2009 -0700 @@ -5441,7 +5441,7 @@ have "1 - c > 0" using c by auto from s(2) obtain z0 where "z0 \ s" by auto - def z \ "\ n::nat. fun_pow n f z0" + def z \ "\n. (f ^^ n) z0" { fix n::nat have "z n \ s" unfolding z_def proof(induct n) case 0 thus ?case using `z0 \s` by auto @@ -5580,7 +5580,7 @@ using dist[THEN bspec[where x=x], THEN bspec[where x=y]] by auto } note dist' = this def y \ "g x" have [simp]:"y\s" unfolding y_def using gs[unfolded image_subset_iff] and `x\s` by blast - def f \ "\ n. fun_pow n g" + def f \ "\n. g ^^ n" have [simp]:"\n z. g (f n z) = f (Suc n) z" unfolding f_def by auto have [simp]:"\z. f 0 z = z" unfolding f_def by auto { fix n::nat and z assume "z\s" diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Library/Word.thy --- a/src/HOL/Library/Word.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Library/Word.thy Mon Apr 27 07:26:17 2009 -0700 @@ -1,5 +1,4 @@ (* Title: HOL/Library/Word.thy - ID: $Id$ Author: Sebastian Skalberg (TU Muenchen) *) @@ -40,10 +39,8 @@ Zero ("\") | One ("\") -primrec - bitval :: "bit => nat" -where - "bitval \ = 0" +primrec bitval :: "bit => nat" where + "bitval \ = 0" | "bitval \ = 1" consts @@ -1531,7 +1528,7 @@ show ?thesis apply simp apply (subst power_Suc [symmetric]) - apply (simp del: power_int.simps) + apply simp done qed finally show ?thesis . diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Library/reflection.ML --- a/src/HOL/Library/reflection.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Library/reflection.ML Mon Apr 27 07:26:17 2009 -0700 @@ -314,5 +314,6 @@ in (rtac th i THEN TRY(rtac TrueI i)) st end); fun reflection_tac ctxt = gen_reflection_tac ctxt Codegen.evaluation_conv; + (*FIXME why Codegen.evaluation_conv? very specific...*) end diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/List.thy --- a/src/HOL/List.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/List.thy Mon Apr 27 07:26:17 2009 -0700 @@ -5,7 +5,7 @@ header {* The datatype of finite lists *} theory List -imports Plain Relation_Power Presburger Recdef ATP_Linkup +imports Plain Presburger Recdef ATP_Linkup uses "Tools/string_syntax.ML" begin @@ -198,7 +198,7 @@ definition rotate :: "nat \ 'a list \ 'a list" where - "rotate n = rotate1 ^ n" + "rotate n = rotate1 ^^ n" definition list_all2 :: "('a => 'b => bool) => 'a list => 'b list => bool" where diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Map.thy --- a/src/HOL/Map.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Map.thy Mon Apr 27 07:26:17 2009 -0700 @@ -11,7 +11,7 @@ imports List begin -types ('a,'b) "~=>" = "'a => 'b option" (infixr 0) +types ('a,'b) "~=>" = "'a => 'b option" (infixr "~=>" 0) translations (type) "a ~=> b " <= (type) "a => b option" syntax (xsymbols) diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/NSA/HyperDef.thy --- a/src/HOL/NSA/HyperDef.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/NSA/HyperDef.thy Mon Apr 27 07:26:17 2009 -0700 @@ -459,7 +459,7 @@ by transfer (rule power_add) lemma hyperpow_one [simp]: - "\r. (r::'a::recpower star) pow (1::hypnat) = r" + "\r. (r::'a::monoid_mult star) pow (1::hypnat) = r" by transfer (rule power_one_right) lemma hyperpow_two: diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/NSA/StarDef.thy --- a/src/HOL/NSA/StarDef.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/NSA/StarDef.thy Mon Apr 27 07:26:17 2009 -0700 @@ -1,5 +1,4 @@ (* Title : HOL/Hyperreal/StarDef.thy - ID : $Id$ Author : Jacques D. Fleuriot and Brian Huffman *) @@ -546,16 +545,6 @@ end -instantiation star :: (power) power -begin - -definition - star_power_def: "(op ^) \ \x n. ( *f* (\x. x ^ n)) x" - -instance .. - -end - instantiation star :: (ord) ord begin @@ -574,7 +563,7 @@ star_add_def star_diff_def star_minus_def star_mult_def star_divide_def star_inverse_def star_le_def star_less_def star_abs_def star_sgn_def - star_div_def star_mod_def star_power_def + star_div_def star_mod_def text {* Class operations preserve standard elements *} @@ -614,15 +603,11 @@ lemma Standard_mod: "\x \ Standard; y \ Standard\ \ x mod y \ Standard" by (simp add: star_mod_def) -lemma Standard_power: "x \ Standard \ x ^ n \ Standard" -by (simp add: star_power_def) - lemmas Standard_simps [simp] = Standard_zero Standard_one Standard_number_of Standard_add Standard_diff Standard_minus Standard_mult Standard_divide Standard_inverse Standard_abs Standard_div Standard_mod - Standard_power text {* @{term star_of} preserves class operations *} @@ -650,9 +635,6 @@ lemma star_of_mod: "star_of (x mod y) = star_of x mod star_of y" by transfer (rule refl) -lemma star_of_power: "star_of (x ^ n) = star_of x ^ n" -by transfer (rule refl) - lemma star_of_abs: "star_of (abs x) = abs (star_of x)" by transfer (rule refl) @@ -717,8 +699,7 @@ lemmas star_of_simps [simp] = star_of_add star_of_diff star_of_minus star_of_mult star_of_divide star_of_inverse - star_of_div star_of_mod - star_of_power star_of_abs + star_of_div star_of_mod star_of_abs star_of_zero star_of_one star_of_number_of star_of_less star_of_le star_of_eq star_of_0_less star_of_0_le star_of_0_eq @@ -970,25 +951,35 @@ instance star :: (ordered_idom) ordered_idom .. instance star :: (ordered_field) ordered_field .. -subsection {* Power classes *} + +subsection {* Power *} + +instance star :: (recpower) recpower .. -text {* - Proving the class axiom @{thm [source] power_Suc} for type - @{typ "'a star"} is a little tricky, because it quantifies - over values of type @{typ nat}. The transfer principle does - not handle quantification over non-star types in general, - but we can work around this by fixing an arbitrary @{typ nat} - value, and then applying the transfer principle. -*} +lemma star_power_def [transfer_unfold]: + "(op ^) \ \x n. ( *f* (\x. x ^ n)) x" +proof (rule eq_reflection, rule ext, rule ext) + fix n :: nat + show "\x::'a star. x ^ n = ( *f* (\x. x ^ n)) x" + proof (induct n) + case 0 + have "\x::'a star. ( *f* (\x. 1)) x = 1" + by transfer simp + then show ?case by simp + next + case (Suc n) + have "\x::'a star. x * ( *f* (\x\'a. x ^ n)) x = ( *f* (\x\'a. x * x ^ n)) x" + by transfer simp + with Suc show ?case by simp + qed +qed -instance star :: (recpower) recpower -proof - show "\a::'a star. a ^ 0 = 1" - by transfer (rule power_0) -next - fix n show "\a::'a star. a ^ Suc n = a * a ^ n" - by transfer (rule power_Suc) -qed +lemma Standard_power [simp]: "x \ Standard \ x ^ n \ Standard" + by (simp add: star_power_def) + +lemma star_of_power [simp]: "star_of (x ^ n) = star_of x ^ n" + by transfer (rule refl) + subsection {* Number classes *} diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Nat.thy --- a/src/HOL/Nat.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Nat.thy Mon Apr 27 07:26:17 2009 -0700 @@ -1164,6 +1164,64 @@ end +subsection {* Natural operation of natural numbers on functions *} + +text {* + We use the same logical constant for the power operations on + functions and relations, in order to share the same syntax. +*} + +consts compow :: "nat \ ('a \ 'b) \ ('a \ 'b)" + +abbreviation compower :: "('a \ 'b) \ nat \ 'a \ 'b" (infixr "^^" 80) where + "f ^^ n \ compow n f" + +notation (latex output) + compower ("(_\<^bsup>_\<^esup>)" [1000] 1000) + +notation (HTML output) + compower ("(_\<^bsup>_\<^esup>)" [1000] 1000) + +text {* @{text "f ^^ n = f o ... o f"}, the n-fold composition of @{text f} *} + +overloading + funpow == "compow :: nat \ ('a \ 'a) \ ('a \ 'a)" +begin + +primrec funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" where + "funpow 0 f = id" + | "funpow (Suc n) f = f o funpow n f" + +end + +text {* for code generation *} + +definition funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" where + funpow_code_def [code post]: "funpow = compow" + +lemmas [code inline] = funpow_code_def [symmetric] + +lemma [code]: + "funpow 0 f = id" + "funpow (Suc n) f = f o funpow n f" + unfolding funpow_code_def by simp_all + +hide (open) const funpow + +lemma funpow_add: + "f ^^ (m + n) = f ^^ m \ f ^^ n" + by (induct m) simp_all + +lemma funpow_swap1: + "f ((f ^^ n) x) = (f ^^ n) (f x)" +proof - + have "f ((f ^^ n) x) = (f ^^ (n + 1)) x" by simp + also have "\ = (f ^^ n o f ^^ 1) x" by (simp only: funpow_add) + also have "\ = (f ^^ n) (f x)" by simp + finally show ?thesis . +qed + + subsection {* Embedding of the Naturals into any @{text semiring_1}: @{term of_nat} *} @@ -1189,7 +1247,7 @@ "of_nat_aux inc 0 i = i" | "of_nat_aux inc (Suc n) i = of_nat_aux inc n (inc i)" -- {* tail recursive *} -lemma of_nat_code [code, code unfold, code inline del]: +lemma of_nat_code: "of_nat n = of_nat_aux (\i. i + 1) n 0" proof (induct n) case 0 then show ?case by simp @@ -1201,9 +1259,11 @@ by simp with Suc show ?case by (simp add: add_commute) qed - + end +declare of_nat_code [code, code unfold, code inline del] + text{*Class for unital semirings with characteristic zero. Includes non-ordered rings like the complex numbers.*} diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/NatBin.thy --- a/src/HOL/NatBin.thy Wed Apr 22 11:00:25 2009 -0700 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,975 +0,0 @@ -(* Title: HOL/NatBin.thy - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1999 University of Cambridge -*) - -header {* Binary arithmetic for the natural numbers *} - -theory NatBin -imports IntDiv -uses ("Tools/nat_simprocs.ML") -begin - -text {* - Arithmetic for naturals is reduced to that for the non-negative integers. -*} - -instantiation nat :: number -begin - -definition - nat_number_of_def [code inline, code del]: "number_of v = nat (number_of v)" - -instance .. - -end - -lemma [code post]: - "nat (number_of v) = number_of v" - unfolding nat_number_of_def .. - -abbreviation (xsymbols) - power2 :: "'a::power => 'a" ("(_\)" [1000] 999) where - "x\ == x^2" - -notation (latex output) - power2 ("(_\)" [1000] 999) - -notation (HTML output) - power2 ("(_\)" [1000] 999) - - -subsection {* Predicate for negative binary numbers *} - -definition neg :: "int \ bool" where - "neg Z \ Z < 0" - -lemma not_neg_int [simp]: "~ neg (of_nat n)" -by (simp add: neg_def) - -lemma neg_zminus_int [simp]: "neg (- (of_nat (Suc n)))" -by (simp add: neg_def neg_less_0_iff_less del: of_nat_Suc) - -lemmas neg_eq_less_0 = neg_def - -lemma not_neg_eq_ge_0: "(~neg x) = (0 \ x)" -by (simp add: neg_def linorder_not_less) - -text{*To simplify inequalities when Numeral1 can get simplified to 1*} - -lemma not_neg_0: "~ neg 0" -by (simp add: One_int_def neg_def) - -lemma not_neg_1: "~ neg 1" -by (simp add: neg_def linorder_not_less zero_le_one) - -lemma neg_nat: "neg z ==> nat z = 0" -by (simp add: neg_def order_less_imp_le) - -lemma not_neg_nat: "~ neg z ==> of_nat (nat z) = z" -by (simp add: linorder_not_less neg_def) - -text {* - If @{term Numeral0} is rewritten to 0 then this rule can't be applied: - @{term Numeral0} IS @{term "number_of Pls"} -*} - -lemma not_neg_number_of_Pls: "~ neg (number_of Int.Pls)" - by (simp add: neg_def) - -lemma neg_number_of_Min: "neg (number_of Int.Min)" - by (simp add: neg_def) - -lemma neg_number_of_Bit0: - "neg (number_of (Int.Bit0 w)) = neg (number_of w)" - by (simp add: neg_def) - -lemma neg_number_of_Bit1: - "neg (number_of (Int.Bit1 w)) = neg (number_of w)" - by (simp add: neg_def) - -lemmas neg_simps [simp] = - not_neg_0 not_neg_1 - not_neg_number_of_Pls neg_number_of_Min - neg_number_of_Bit0 neg_number_of_Bit1 - - -subsection{*Function @{term nat}: Coercion from Type @{typ int} to @{typ nat}*} - -declare nat_0 [simp] nat_1 [simp] - -lemma nat_number_of [simp]: "nat (number_of w) = number_of w" -by (simp add: nat_number_of_def) - -lemma nat_numeral_0_eq_0 [simp]: "Numeral0 = (0::nat)" -by (simp add: nat_number_of_def) - -lemma nat_numeral_1_eq_1 [simp]: "Numeral1 = (1::nat)" -by (simp add: nat_1 nat_number_of_def) - -lemma numeral_1_eq_Suc_0: "Numeral1 = Suc 0" -by (simp add: nat_numeral_1_eq_1) - -lemma numeral_2_eq_2: "2 = Suc (Suc 0)" -apply (unfold nat_number_of_def) -apply (rule nat_2) -done - - -subsection{*Function @{term int}: Coercion from Type @{typ nat} to @{typ int}*} - -lemma int_nat_number_of [simp]: - "int (number_of v) = - (if neg (number_of v :: int) then 0 - else (number_of v :: int))" - unfolding nat_number_of_def number_of_is_id neg_def - by simp - - -subsubsection{*Successor *} - -lemma Suc_nat_eq_nat_zadd1: "(0::int) <= z ==> Suc (nat z) = nat (1 + z)" -apply (rule sym) -apply (simp add: nat_eq_iff int_Suc) -done - -lemma Suc_nat_number_of_add: - "Suc (number_of v + n) = - (if neg (number_of v :: int) then 1+n else number_of (Int.succ v) + n)" - unfolding nat_number_of_def number_of_is_id neg_def numeral_simps - by (simp add: Suc_nat_eq_nat_zadd1 add_ac) - -lemma Suc_nat_number_of [simp]: - "Suc (number_of v) = - (if neg (number_of v :: int) then 1 else number_of (Int.succ v))" -apply (cut_tac n = 0 in Suc_nat_number_of_add) -apply (simp cong del: if_weak_cong) -done - - -subsubsection{*Addition *} - -lemma add_nat_number_of [simp]: - "(number_of v :: nat) + number_of v' = - (if v < Int.Pls then number_of v' - else if v' < Int.Pls then number_of v - else number_of (v + v'))" - unfolding nat_number_of_def number_of_is_id numeral_simps - by (simp add: nat_add_distrib) - -lemma nat_number_of_add_1 [simp]: - "number_of v + (1::nat) = - (if v < Int.Pls then 1 else number_of (Int.succ v))" - unfolding nat_number_of_def number_of_is_id numeral_simps - by (simp add: nat_add_distrib) - -lemma nat_1_add_number_of [simp]: - "(1::nat) + number_of v = - (if v < Int.Pls then 1 else number_of (Int.succ v))" - unfolding nat_number_of_def number_of_is_id numeral_simps - by (simp add: nat_add_distrib) - -lemma nat_1_add_1 [simp]: "1 + 1 = (2::nat)" - by (rule int_int_eq [THEN iffD1]) simp - - -subsubsection{*Subtraction *} - -lemma diff_nat_eq_if: - "nat z - nat z' = - (if neg z' then nat z - else let d = z-z' in - if neg d then 0 else nat d)" -by (simp add: Let_def nat_diff_distrib [symmetric] neg_eq_less_0 not_neg_eq_ge_0) - - -lemma diff_nat_number_of [simp]: - "(number_of v :: nat) - number_of v' = - (if v' < Int.Pls then number_of v - else let d = number_of (v + uminus v') in - if neg d then 0 else nat d)" - unfolding nat_number_of_def number_of_is_id numeral_simps neg_def - by auto - -lemma nat_number_of_diff_1 [simp]: - "number_of v - (1::nat) = - (if v \ Int.Pls then 0 else number_of (Int.pred v))" - unfolding nat_number_of_def number_of_is_id numeral_simps - by auto - - -subsubsection{*Multiplication *} - -lemma mult_nat_number_of [simp]: - "(number_of v :: nat) * number_of v' = - (if v < Int.Pls then 0 else number_of (v * v'))" - unfolding nat_number_of_def number_of_is_id numeral_simps - by (simp add: nat_mult_distrib) - - -subsubsection{*Quotient *} - -lemma div_nat_number_of [simp]: - "(number_of v :: nat) div number_of v' = - (if neg (number_of v :: int) then 0 - else nat (number_of v div number_of v'))" - unfolding nat_number_of_def number_of_is_id neg_def - by (simp add: nat_div_distrib) - -lemma one_div_nat_number_of [simp]: - "Suc 0 div number_of v' = nat (1 div number_of v')" -by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) - - -subsubsection{*Remainder *} - -lemma mod_nat_number_of [simp]: - "(number_of v :: nat) mod number_of v' = - (if neg (number_of v :: int) then 0 - else if neg (number_of v' :: int) then number_of v - else nat (number_of v mod number_of v'))" - unfolding nat_number_of_def number_of_is_id neg_def - by (simp add: nat_mod_distrib) - -lemma one_mod_nat_number_of [simp]: - "Suc 0 mod number_of v' = - (if neg (number_of v' :: int) then Suc 0 - else nat (1 mod number_of v'))" -by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) - - -subsubsection{* Divisibility *} - -lemmas dvd_eq_mod_eq_0_number_of = - dvd_eq_mod_eq_0 [of "number_of x" "number_of y", standard] - -declare dvd_eq_mod_eq_0_number_of [simp] - -ML -{* -val nat_number_of_def = thm"nat_number_of_def"; - -val nat_number_of = thm"nat_number_of"; -val nat_numeral_0_eq_0 = thm"nat_numeral_0_eq_0"; -val nat_numeral_1_eq_1 = thm"nat_numeral_1_eq_1"; -val numeral_1_eq_Suc_0 = thm"numeral_1_eq_Suc_0"; -val numeral_2_eq_2 = thm"numeral_2_eq_2"; -val nat_div_distrib = thm"nat_div_distrib"; -val nat_mod_distrib = thm"nat_mod_distrib"; -val int_nat_number_of = thm"int_nat_number_of"; -val Suc_nat_eq_nat_zadd1 = thm"Suc_nat_eq_nat_zadd1"; -val Suc_nat_number_of_add = thm"Suc_nat_number_of_add"; -val Suc_nat_number_of = thm"Suc_nat_number_of"; -val add_nat_number_of = thm"add_nat_number_of"; -val diff_nat_eq_if = thm"diff_nat_eq_if"; -val diff_nat_number_of = thm"diff_nat_number_of"; -val mult_nat_number_of = thm"mult_nat_number_of"; -val div_nat_number_of = thm"div_nat_number_of"; -val mod_nat_number_of = thm"mod_nat_number_of"; -*} - - -subsection{*Comparisons*} - -subsubsection{*Equals (=) *} - -lemma eq_nat_nat_iff: - "[| (0::int) <= z; 0 <= z' |] ==> (nat z = nat z') = (z=z')" -by (auto elim!: nonneg_eq_int) - -lemma eq_nat_number_of [simp]: - "((number_of v :: nat) = number_of v') = - (if neg (number_of v :: int) then (number_of v' :: int) \ 0 - else if neg (number_of v' :: int) then (number_of v :: int) = 0 - else v = v')" - unfolding nat_number_of_def number_of_is_id neg_def - by auto - - -subsubsection{*Less-than (<) *} - -lemma less_nat_number_of [simp]: - "(number_of v :: nat) < number_of v' \ - (if v < v' then Int.Pls < v' else False)" - unfolding nat_number_of_def number_of_is_id numeral_simps - by auto - - -subsubsection{*Less-than-or-equal *} - -lemma le_nat_number_of [simp]: - "(number_of v :: nat) \ number_of v' \ - (if v \ v' then True else v \ Int.Pls)" - unfolding nat_number_of_def number_of_is_id numeral_simps - by auto - -(*Maps #n to n for n = 0, 1, 2*) -lemmas numerals = nat_numeral_0_eq_0 nat_numeral_1_eq_1 numeral_2_eq_2 - - -subsection{*Powers with Numeric Exponents*} - -text{*We cannot refer to the number @{term 2} in @{text Ring_and_Field.thy}. -We cannot prove general results about the numeral @{term "-1"}, so we have to -use @{term "- 1"} instead.*} - -lemma power2_eq_square: "(a::'a::recpower)\ = a * a" - by (simp add: numeral_2_eq_2 Power.power_Suc) - -lemma zero_power2 [simp]: "(0::'a::{semiring_1,recpower})\ = 0" - by (simp add: power2_eq_square) - -lemma one_power2 [simp]: "(1::'a::{semiring_1,recpower})\ = 1" - by (simp add: power2_eq_square) - -lemma power3_eq_cube: "(x::'a::recpower) ^ 3 = x * x * x" - apply (subgoal_tac "3 = Suc (Suc (Suc 0))") - apply (erule ssubst) - apply (simp add: power_Suc mult_ac) - apply (unfold nat_number_of_def) - apply (subst nat_eq_iff) - apply simp -done - -text{*Squares of literal numerals will be evaluated.*} -lemmas power2_eq_square_number_of = - power2_eq_square [of "number_of w", standard] -declare power2_eq_square_number_of [simp] - - -lemma zero_le_power2[simp]: "0 \ (a\::'a::{ordered_idom,recpower})" - by (simp add: power2_eq_square) - -lemma zero_less_power2[simp]: - "(0 < a\) = (a \ (0::'a::{ordered_idom,recpower}))" - by (force simp add: power2_eq_square zero_less_mult_iff linorder_neq_iff) - -lemma power2_less_0[simp]: - fixes a :: "'a::{ordered_idom,recpower}" - shows "~ (a\ < 0)" -by (force simp add: power2_eq_square mult_less_0_iff) - -lemma zero_eq_power2[simp]: - "(a\ = 0) = (a = (0::'a::{ordered_idom,recpower}))" - by (force simp add: power2_eq_square mult_eq_0_iff) - -lemma abs_power2[simp]: - "abs(a\) = (a\::'a::{ordered_idom,recpower})" - by (simp add: power2_eq_square abs_mult abs_mult_self) - -lemma power2_abs[simp]: - "(abs a)\ = (a\::'a::{ordered_idom,recpower})" - by (simp add: power2_eq_square abs_mult_self) - -lemma power2_minus[simp]: - "(- a)\ = (a\::'a::{comm_ring_1,recpower})" - by (simp add: power2_eq_square) - -lemma power2_le_imp_le: - fixes x y :: "'a::{ordered_semidom,recpower}" - shows "\x\ \ y\; 0 \ y\ \ x \ y" -unfolding numeral_2_eq_2 by (rule power_le_imp_le_base) - -lemma power2_less_imp_less: - fixes x y :: "'a::{ordered_semidom,recpower}" - shows "\x\ < y\; 0 \ y\ \ x < y" -by (rule power_less_imp_less_base) - -lemma power2_eq_imp_eq: - fixes x y :: "'a::{ordered_semidom,recpower}" - shows "\x\ = y\; 0 \ x; 0 \ y\ \ x = y" -unfolding numeral_2_eq_2 by (erule (2) power_eq_imp_eq_base, simp) - -lemma power_minus1_even[simp]: "(- 1) ^ (2*n) = (1::'a::{comm_ring_1,recpower})" -proof (induct n) - case 0 show ?case by simp -next - case (Suc n) then show ?case by (simp add: power_Suc power_add) -qed - -lemma power_minus1_odd: "(- 1) ^ Suc(2*n) = -(1::'a::{comm_ring_1,recpower})" - by (simp add: power_Suc) - -lemma power_even_eq: "(a::'a::recpower) ^ (2*n) = (a^n)^2" -by (subst mult_commute) (simp add: power_mult) - -lemma power_odd_eq: "(a::int) ^ Suc(2*n) = a * (a^n)^2" -by (simp add: power_even_eq) - -lemma power_minus_even [simp]: - "(-a) ^ (2*n) = (a::'a::{comm_ring_1,recpower}) ^ (2*n)" -by (simp add: power_minus1_even power_minus [of a]) - -lemma zero_le_even_power'[simp]: - "0 \ (a::'a::{ordered_idom,recpower}) ^ (2*n)" -proof (induct "n") - case 0 - show ?case by (simp add: zero_le_one) -next - case (Suc n) - have "a ^ (2 * Suc n) = (a*a) * a ^ (2*n)" - by (simp add: mult_ac power_add power2_eq_square) - thus ?case - by (simp add: prems zero_le_mult_iff) -qed - -lemma odd_power_less_zero: - "(a::'a::{ordered_idom,recpower}) < 0 ==> a ^ Suc(2*n) < 0" -proof (induct "n") - case 0 - then show ?case by simp -next - case (Suc n) - have "a ^ Suc (2 * Suc n) = (a*a) * a ^ Suc(2*n)" - by (simp add: mult_ac power_add power2_eq_square) - thus ?case - by (simp del: power_Suc add: prems mult_less_0_iff mult_neg_neg) -qed - -lemma odd_0_le_power_imp_0_le: - "0 \ a ^ Suc(2*n) ==> 0 \ (a::'a::{ordered_idom,recpower})" -apply (insert odd_power_less_zero [of a n]) -apply (force simp add: linorder_not_less [symmetric]) -done - -text{*Simprules for comparisons where common factors can be cancelled.*} -lemmas zero_compare_simps = - add_strict_increasing add_strict_increasing2 add_increasing - zero_le_mult_iff zero_le_divide_iff - zero_less_mult_iff zero_less_divide_iff - mult_le_0_iff divide_le_0_iff - mult_less_0_iff divide_less_0_iff - zero_le_power2 power2_less_0 - -subsubsection{*Nat *} - -lemma Suc_pred': "0 < n ==> n = Suc(n - 1)" -by (simp add: numerals) - -(*Expresses a natural number constant as the Suc of another one. - NOT suitable for rewriting because n recurs in the condition.*) -lemmas expand_Suc = Suc_pred' [of "number_of v", standard] - -subsubsection{*Arith *} - -lemma Suc_eq_add_numeral_1: "Suc n = n + 1" -by (simp add: numerals) - -lemma Suc_eq_add_numeral_1_left: "Suc n = 1 + n" -by (simp add: numerals) - -(* These two can be useful when m = number_of... *) - -lemma add_eq_if: "(m::nat) + n = (if m=0 then n else Suc ((m - 1) + n))" - unfolding One_nat_def by (cases m) simp_all - -lemma mult_eq_if: "(m::nat) * n = (if m=0 then 0 else n + ((m - 1) * n))" - unfolding One_nat_def by (cases m) simp_all - -lemma power_eq_if: "(p ^ m :: nat) = (if m=0 then 1 else p * (p ^ (m - 1)))" - unfolding One_nat_def by (cases m) simp_all - - -subsection{*Comparisons involving (0::nat) *} - -text{*Simplification already does @{term "n<0"}, @{term "n\0"} and @{term "0\n"}.*} - -lemma eq_number_of_0 [simp]: - "number_of v = (0::nat) \ v \ Int.Pls" - unfolding nat_number_of_def number_of_is_id numeral_simps - by auto - -lemma eq_0_number_of [simp]: - "(0::nat) = number_of v \ v \ Int.Pls" -by (rule trans [OF eq_sym_conv eq_number_of_0]) - -lemma less_0_number_of [simp]: - "(0::nat) < number_of v \ Int.Pls < v" - unfolding nat_number_of_def number_of_is_id numeral_simps - by simp - -lemma neg_imp_number_of_eq_0: "neg (number_of v :: int) ==> number_of v = (0::nat)" -by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric]) - - - -subsection{*Comparisons involving @{term Suc} *} - -lemma eq_number_of_Suc [simp]: - "(number_of v = Suc n) = - (let pv = number_of (Int.pred v) in - if neg pv then False else nat pv = n)" -apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less - number_of_pred nat_number_of_def - split add: split_if) -apply (rule_tac x = "number_of v" in spec) -apply (auto simp add: nat_eq_iff) -done - -lemma Suc_eq_number_of [simp]: - "(Suc n = number_of v) = - (let pv = number_of (Int.pred v) in - if neg pv then False else nat pv = n)" -by (rule trans [OF eq_sym_conv eq_number_of_Suc]) - -lemma less_number_of_Suc [simp]: - "(number_of v < Suc n) = - (let pv = number_of (Int.pred v) in - if neg pv then True else nat pv < n)" -apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less - number_of_pred nat_number_of_def - split add: split_if) -apply (rule_tac x = "number_of v" in spec) -apply (auto simp add: nat_less_iff) -done - -lemma less_Suc_number_of [simp]: - "(Suc n < number_of v) = - (let pv = number_of (Int.pred v) in - if neg pv then False else n < nat pv)" -apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less - number_of_pred nat_number_of_def - split add: split_if) -apply (rule_tac x = "number_of v" in spec) -apply (auto simp add: zless_nat_eq_int_zless) -done - -lemma le_number_of_Suc [simp]: - "(number_of v <= Suc n) = - (let pv = number_of (Int.pred v) in - if neg pv then True else nat pv <= n)" -by (simp add: Let_def less_Suc_number_of linorder_not_less [symmetric]) - -lemma le_Suc_number_of [simp]: - "(Suc n <= number_of v) = - (let pv = number_of (Int.pred v) in - if neg pv then False else n <= nat pv)" -by (simp add: Let_def less_number_of_Suc linorder_not_less [symmetric]) - - -lemma eq_number_of_Pls_Min: "(Numeral0 ::int) ~= number_of Int.Min" -by auto - - - -subsection{*Max and Min Combined with @{term Suc} *} - -lemma max_number_of_Suc [simp]: - "max (Suc n) (number_of v) = - (let pv = number_of (Int.pred v) in - if neg pv then Suc n else Suc(max n (nat pv)))" -apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def - split add: split_if nat.split) -apply (rule_tac x = "number_of v" in spec) -apply auto -done - -lemma max_Suc_number_of [simp]: - "max (number_of v) (Suc n) = - (let pv = number_of (Int.pred v) in - if neg pv then Suc n else Suc(max (nat pv) n))" -apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def - split add: split_if nat.split) -apply (rule_tac x = "number_of v" in spec) -apply auto -done - -lemma min_number_of_Suc [simp]: - "min (Suc n) (number_of v) = - (let pv = number_of (Int.pred v) in - if neg pv then 0 else Suc(min n (nat pv)))" -apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def - split add: split_if nat.split) -apply (rule_tac x = "number_of v" in spec) -apply auto -done - -lemma min_Suc_number_of [simp]: - "min (number_of v) (Suc n) = - (let pv = number_of (Int.pred v) in - if neg pv then 0 else Suc(min (nat pv) n))" -apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def - split add: split_if nat.split) -apply (rule_tac x = "number_of v" in spec) -apply auto -done - -subsection{*Literal arithmetic involving powers*} - -lemma nat_power_eq: "(0::int) <= z ==> nat (z^n) = nat z ^ n" -apply (induct "n") -apply (simp_all (no_asm_simp) add: nat_mult_distrib) -done - -lemma power_nat_number_of: - "(number_of v :: nat) ^ n = - (if neg (number_of v :: int) then 0^n else nat ((number_of v :: int) ^ n))" -by (simp only: simp_thms neg_nat not_neg_eq_ge_0 nat_number_of_def nat_power_eq - split add: split_if cong: imp_cong) - - -lemmas power_nat_number_of_number_of = power_nat_number_of [of _ "number_of w", standard] -declare power_nat_number_of_number_of [simp] - - - -text{*For arbitrary rings*} - -lemma power_number_of_even: - fixes z :: "'a::{number_ring,recpower}" - shows "z ^ number_of (Int.Bit0 w) = (let w = z ^ (number_of w) in w * w)" -unfolding Let_def nat_number_of_def number_of_Bit0 -apply (rule_tac x = "number_of w" in spec, clarify) -apply (case_tac " (0::int) <= x") -apply (auto simp add: nat_mult_distrib power_even_eq power2_eq_square) -done - -lemma power_number_of_odd: - fixes z :: "'a::{number_ring,recpower}" - shows "z ^ number_of (Int.Bit1 w) = (if (0::int) <= number_of w - then (let w = z ^ (number_of w) in z * w * w) else 1)" -unfolding Let_def nat_number_of_def number_of_Bit1 -apply (rule_tac x = "number_of w" in spec, auto) -apply (simp only: nat_add_distrib nat_mult_distrib) -apply simp -apply (auto simp add: nat_add_distrib nat_mult_distrib power_even_eq power2_eq_square neg_nat power_Suc) -done - -lemmas zpower_number_of_even = power_number_of_even [where 'a=int] -lemmas zpower_number_of_odd = power_number_of_odd [where 'a=int] - -lemmas power_number_of_even_number_of [simp] = - power_number_of_even [of "number_of v", standard] - -lemmas power_number_of_odd_number_of [simp] = - power_number_of_odd [of "number_of v", standard] - - - -ML -{* -val numeral_ss = @{simpset} addsimps @{thms numerals}; - -val nat_bin_arith_setup = - Lin_Arith.map_data - (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} => - {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, - inj_thms = inj_thms, - lessD = lessD, neqE = neqE, - simpset = simpset addsimps @{thms neg_simps} @ - [@{thm Suc_nat_number_of}, @{thm int_nat_number_of}]}) -*} - -declaration {* K nat_bin_arith_setup *} - -(* Enable arith to deal with div/mod k where k is a numeral: *) -declare split_div[of _ _ "number_of k", standard, arith_split] -declare split_mod[of _ _ "number_of k", standard, arith_split] - -lemma nat_number_of_Pls: "Numeral0 = (0::nat)" - by (simp add: number_of_Pls nat_number_of_def) - -lemma nat_number_of_Min: "number_of Int.Min = (0::nat)" - apply (simp only: number_of_Min nat_number_of_def nat_zminus_int) - done - -lemma nat_number_of_Bit0: - "number_of (Int.Bit0 w) = (let n::nat = number_of w in n + n)" - unfolding nat_number_of_def number_of_is_id numeral_simps Let_def - by auto - -lemma nat_number_of_Bit1: - "number_of (Int.Bit1 w) = - (if neg (number_of w :: int) then 0 - else let n = number_of w in Suc (n + n))" - unfolding nat_number_of_def number_of_is_id numeral_simps neg_def Let_def - by auto - -lemmas nat_number = - nat_number_of_Pls nat_number_of_Min - nat_number_of_Bit0 nat_number_of_Bit1 - -lemma Let_Suc [simp]: "Let (Suc n) f == f (Suc n)" - by (simp add: Let_def) - -lemma power_m1_even: "(-1) ^ (2*n) = (1::'a::{number_ring,recpower})" -by (simp add: power_mult power_Suc); - -lemma power_m1_odd: "(-1) ^ Suc(2*n) = (-1::'a::{number_ring,recpower})" -by (simp add: power_mult power_Suc); - - -subsection{*Literal arithmetic and @{term of_nat}*} - -lemma of_nat_double: - "0 \ x ==> of_nat (nat (2 * x)) = of_nat (nat x) + of_nat (nat x)" -by (simp only: mult_2 nat_add_distrib of_nat_add) - -lemma nat_numeral_m1_eq_0: "-1 = (0::nat)" -by (simp only: nat_number_of_def) - -lemma of_nat_number_of_lemma: - "of_nat (number_of v :: nat) = - (if 0 \ (number_of v :: int) - then (number_of v :: 'a :: number_ring) - else 0)" -by (simp add: int_number_of_def nat_number_of_def number_of_eq of_nat_nat); - -lemma of_nat_number_of_eq [simp]: - "of_nat (number_of v :: nat) = - (if neg (number_of v :: int) then 0 - else (number_of v :: 'a :: number_ring))" -by (simp only: of_nat_number_of_lemma neg_def, simp) - - -subsection {*Lemmas for the Combination and Cancellation Simprocs*} - -lemma nat_number_of_add_left: - "number_of v + (number_of v' + (k::nat)) = - (if neg (number_of v :: int) then number_of v' + k - else if neg (number_of v' :: int) then number_of v + k - else number_of (v + v') + k)" - unfolding nat_number_of_def number_of_is_id neg_def - by auto - -lemma nat_number_of_mult_left: - "number_of v * (number_of v' * (k::nat)) = - (if v < Int.Pls then 0 - else number_of (v * v') * k)" -by simp - - -subsubsection{*For @{text combine_numerals}*} - -lemma left_add_mult_distrib: "i*u + (j*u + k) = (i+j)*u + (k::nat)" -by (simp add: add_mult_distrib) - - -subsubsection{*For @{text cancel_numerals}*} - -lemma nat_diff_add_eq1: - "j <= (i::nat) ==> ((i*u + m) - (j*u + n)) = (((i-j)*u + m) - n)" -by (simp split add: nat_diff_split add: add_mult_distrib) - -lemma nat_diff_add_eq2: - "i <= (j::nat) ==> ((i*u + m) - (j*u + n)) = (m - ((j-i)*u + n))" -by (simp split add: nat_diff_split add: add_mult_distrib) - -lemma nat_eq_add_iff1: - "j <= (i::nat) ==> (i*u + m = j*u + n) = ((i-j)*u + m = n)" -by (auto split add: nat_diff_split simp add: add_mult_distrib) - -lemma nat_eq_add_iff2: - "i <= (j::nat) ==> (i*u + m = j*u + n) = (m = (j-i)*u + n)" -by (auto split add: nat_diff_split simp add: add_mult_distrib) - -lemma nat_less_add_iff1: - "j <= (i::nat) ==> (i*u + m < j*u + n) = ((i-j)*u + m < n)" -by (auto split add: nat_diff_split simp add: add_mult_distrib) - -lemma nat_less_add_iff2: - "i <= (j::nat) ==> (i*u + m < j*u + n) = (m < (j-i)*u + n)" -by (auto split add: nat_diff_split simp add: add_mult_distrib) - -lemma nat_le_add_iff1: - "j <= (i::nat) ==> (i*u + m <= j*u + n) = ((i-j)*u + m <= n)" -by (auto split add: nat_diff_split simp add: add_mult_distrib) - -lemma nat_le_add_iff2: - "i <= (j::nat) ==> (i*u + m <= j*u + n) = (m <= (j-i)*u + n)" -by (auto split add: nat_diff_split simp add: add_mult_distrib) - - -subsubsection{*For @{text cancel_numeral_factors} *} - -lemma nat_mult_le_cancel1: "(0::nat) < k ==> (k*m <= k*n) = (m<=n)" -by auto - -lemma nat_mult_less_cancel1: "(0::nat) < k ==> (k*m < k*n) = (m (k*m = k*n) = (m=n)" -by auto - -lemma nat_mult_div_cancel1: "(0::nat) < k ==> (k*m) div (k*n) = (m div n)" -by auto - -lemma nat_mult_dvd_cancel_disj[simp]: - "(k*m) dvd (k*n) = (k=0 | m dvd (n::nat))" -by(auto simp: dvd_eq_mod_eq_0 mod_mult_distrib2[symmetric]) - -lemma nat_mult_dvd_cancel1: "0 < k \ (k*m) dvd (k*n::nat) = (m dvd n)" -by(auto) - - -subsubsection{*For @{text cancel_factor} *} - -lemma nat_mult_le_cancel_disj: "(k*m <= k*n) = ((0::nat) < k --> m<=n)" -by auto - -lemma nat_mult_less_cancel_disj: "(k*m < k*n) = ((0::nat) < k & m Suc m - n = m - (n - Numeral1)" -by (simp add: numeral_0_eq_0 numeral_1_eq_1 split add: nat_diff_split) - -text {*Now just instantiating @{text n} to @{text "number_of v"} does - the right simplification, but with some redundant inequality - tests.*} -lemma neg_number_of_pred_iff_0: - "neg (number_of (Int.pred v)::int) = (number_of v = (0::nat))" -apply (subgoal_tac "neg (number_of (Int.pred v)) = (number_of v < Suc 0) ") -apply (simp only: less_Suc_eq_le le_0_eq) -apply (subst less_number_of_Suc, simp) -done - -text{*No longer required as a simprule because of the @{text inverse_fold} - simproc*} -lemma Suc_diff_number_of: - "Int.Pls < v ==> - Suc m - (number_of v) = m - (number_of (Int.pred v))" -apply (subst Suc_diff_eq_diff_pred) -apply simp -apply (simp del: nat_numeral_1_eq_1) -apply (auto simp only: diff_nat_number_of less_0_number_of [symmetric] - neg_number_of_pred_iff_0) -done - -lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n" -by (simp add: numerals split add: nat_diff_split) - - -subsubsection{*For @{term nat_case} and @{term nat_rec}*} - -lemma nat_case_number_of [simp]: - "nat_case a f (number_of v) = - (let pv = number_of (Int.pred v) in - if neg pv then a else f (nat pv))" -by (simp split add: nat.split add: Let_def neg_number_of_pred_iff_0) - -lemma nat_case_add_eq_if [simp]: - "nat_case a f ((number_of v) + n) = - (let pv = number_of (Int.pred v) in - if neg pv then nat_case a f n else f (nat pv + n))" -apply (subst add_eq_if) -apply (simp split add: nat.split - del: nat_numeral_1_eq_1 - add: nat_numeral_1_eq_1 [symmetric] - numeral_1_eq_Suc_0 [symmetric] - neg_number_of_pred_iff_0) -done - -lemma nat_rec_number_of [simp]: - "nat_rec a f (number_of v) = - (let pv = number_of (Int.pred v) in - if neg pv then a else f (nat pv) (nat_rec a f (nat pv)))" -apply (case_tac " (number_of v) ::nat") -apply (simp_all (no_asm_simp) add: Let_def neg_number_of_pred_iff_0) -apply (simp split add: split_if_asm) -done - -lemma nat_rec_add_eq_if [simp]: - "nat_rec a f (number_of v + n) = - (let pv = number_of (Int.pred v) in - if neg pv then nat_rec a f n - else f (nat pv + n) (nat_rec a f (nat pv + n)))" -apply (subst add_eq_if) -apply (simp split add: nat.split - del: nat_numeral_1_eq_1 - add: nat_numeral_1_eq_1 [symmetric] - numeral_1_eq_Suc_0 [symmetric] - neg_number_of_pred_iff_0) -done - - -subsubsection{*Various Other Lemmas*} - -text {*Evens and Odds, for Mutilated Chess Board*} - -text{*Lemmas for specialist use, NOT as default simprules*} -lemma nat_mult_2: "2 * z = (z+z::nat)" -proof - - have "2*z = (1 + 1)*z" by simp - also have "... = z+z" by (simp add: left_distrib) - finally show ?thesis . -qed - -lemma nat_mult_2_right: "z * 2 = (z+z::nat)" -by (subst mult_commute, rule nat_mult_2) - -text{*Case analysis on @{term "n<2"}*} -lemma less_2_cases: "(n::nat) < 2 ==> n = 0 | n = Suc 0" -by arith - -lemma div2_Suc_Suc [simp]: "Suc(Suc m) div 2 = Suc (m div 2)" -by arith - -lemma add_self_div_2 [simp]: "(m + m) div 2 = (m::nat)" -by (simp add: nat_mult_2 [symmetric]) - -lemma mod2_Suc_Suc [simp]: "Suc(Suc(m)) mod 2 = m mod 2" -apply (subgoal_tac "m mod 2 < 2") -apply (erule less_2_cases [THEN disjE]) -apply (simp_all (no_asm_simp) add: Let_def mod_Suc nat_1) -done - -lemma mod2_gr_0 [simp]: "!!m::nat. (0 < m mod 2) = (m mod 2 = 1)" -apply (subgoal_tac "m mod 2 < 2") -apply (force simp del: mod_less_divisor, simp) -done - -text{*Removal of Small Numerals: 0, 1 and (in additive positions) 2*} - -lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)" -by simp - -lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)" -by simp - -text{*Can be used to eliminate long strings of Sucs, but not by default*} -lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n" -by simp - - -text{*These lemmas collapse some needless occurrences of Suc: - at least three Sucs, since two and fewer are rewritten back to Suc again! - We already have some rules to simplify operands smaller than 3.*} - -lemma div_Suc_eq_div_add3 [simp]: "m div (Suc (Suc (Suc n))) = m div (3+n)" -by (simp add: Suc3_eq_add_3) - -lemma mod_Suc_eq_mod_add3 [simp]: "m mod (Suc (Suc (Suc n))) = m mod (3+n)" -by (simp add: Suc3_eq_add_3) - -lemma Suc_div_eq_add3_div: "(Suc (Suc (Suc m))) div n = (3+m) div n" -by (simp add: Suc3_eq_add_3) - -lemma Suc_mod_eq_add3_mod: "(Suc (Suc (Suc m))) mod n = (3+m) mod n" -by (simp add: Suc3_eq_add_3) - -lemmas Suc_div_eq_add3_div_number_of = - Suc_div_eq_add3_div [of _ "number_of v", standard] -declare Suc_div_eq_add3_div_number_of [simp] - -lemmas Suc_mod_eq_add3_mod_number_of = - Suc_mod_eq_add3_mod [of _ "number_of v", standard] -declare Suc_mod_eq_add3_mod_number_of [simp] - -end diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Nat_Numeral.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nat_Numeral.thy Mon Apr 27 07:26:17 2009 -0700 @@ -0,0 +1,980 @@ +(* Title: HOL/Nat_Numeral.thy + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1999 University of Cambridge +*) + +header {* Binary numerals for the natural numbers *} + +theory Nat_Numeral +imports IntDiv +uses ("Tools/nat_simprocs.ML") +begin + +text {* + Arithmetic for naturals is reduced to that for the non-negative integers. +*} + +instantiation nat :: number +begin + +definition + nat_number_of_def [code inline, code del]: "number_of v = nat (number_of v)" + +instance .. + +end + +lemma [code post]: + "nat (number_of v) = number_of v" + unfolding nat_number_of_def .. + +context recpower +begin + +abbreviation (xsymbols) + power2 :: "'a \ 'a" ("(_\)" [1000] 999) where + "x\ \ x ^ 2" + +notation (latex output) + power2 ("(_\)" [1000] 999) + +notation (HTML output) + power2 ("(_\)" [1000] 999) + +end + + +subsection {* Predicate for negative binary numbers *} + +definition neg :: "int \ bool" where + "neg Z \ Z < 0" + +lemma not_neg_int [simp]: "~ neg (of_nat n)" +by (simp add: neg_def) + +lemma neg_zminus_int [simp]: "neg (- (of_nat (Suc n)))" +by (simp add: neg_def neg_less_0_iff_less del: of_nat_Suc) + +lemmas neg_eq_less_0 = neg_def + +lemma not_neg_eq_ge_0: "(~neg x) = (0 \ x)" +by (simp add: neg_def linorder_not_less) + +text{*To simplify inequalities when Numeral1 can get simplified to 1*} + +lemma not_neg_0: "~ neg 0" +by (simp add: One_int_def neg_def) + +lemma not_neg_1: "~ neg 1" +by (simp add: neg_def linorder_not_less zero_le_one) + +lemma neg_nat: "neg z ==> nat z = 0" +by (simp add: neg_def order_less_imp_le) + +lemma not_neg_nat: "~ neg z ==> of_nat (nat z) = z" +by (simp add: linorder_not_less neg_def) + +text {* + If @{term Numeral0} is rewritten to 0 then this rule can't be applied: + @{term Numeral0} IS @{term "number_of Pls"} +*} + +lemma not_neg_number_of_Pls: "~ neg (number_of Int.Pls)" + by (simp add: neg_def) + +lemma neg_number_of_Min: "neg (number_of Int.Min)" + by (simp add: neg_def) + +lemma neg_number_of_Bit0: + "neg (number_of (Int.Bit0 w)) = neg (number_of w)" + by (simp add: neg_def) + +lemma neg_number_of_Bit1: + "neg (number_of (Int.Bit1 w)) = neg (number_of w)" + by (simp add: neg_def) + +lemmas neg_simps [simp] = + not_neg_0 not_neg_1 + not_neg_number_of_Pls neg_number_of_Min + neg_number_of_Bit0 neg_number_of_Bit1 + + +subsection{*Function @{term nat}: Coercion from Type @{typ int} to @{typ nat}*} + +declare nat_0 [simp] nat_1 [simp] + +lemma nat_number_of [simp]: "nat (number_of w) = number_of w" +by (simp add: nat_number_of_def) + +lemma nat_numeral_0_eq_0 [simp]: "Numeral0 = (0::nat)" +by (simp add: nat_number_of_def) + +lemma nat_numeral_1_eq_1 [simp]: "Numeral1 = (1::nat)" +by (simp add: nat_1 nat_number_of_def) + +lemma numeral_1_eq_Suc_0: "Numeral1 = Suc 0" +by (simp add: nat_numeral_1_eq_1) + +lemma numeral_2_eq_2: "2 = Suc (Suc 0)" +apply (unfold nat_number_of_def) +apply (rule nat_2) +done + + +subsection{*Function @{term int}: Coercion from Type @{typ nat} to @{typ int}*} + +lemma int_nat_number_of [simp]: + "int (number_of v) = + (if neg (number_of v :: int) then 0 + else (number_of v :: int))" + unfolding nat_number_of_def number_of_is_id neg_def + by simp + + +subsubsection{*Successor *} + +lemma Suc_nat_eq_nat_zadd1: "(0::int) <= z ==> Suc (nat z) = nat (1 + z)" +apply (rule sym) +apply (simp add: nat_eq_iff int_Suc) +done + +lemma Suc_nat_number_of_add: + "Suc (number_of v + n) = + (if neg (number_of v :: int) then 1+n else number_of (Int.succ v) + n)" + unfolding nat_number_of_def number_of_is_id neg_def numeral_simps + by (simp add: Suc_nat_eq_nat_zadd1 add_ac) + +lemma Suc_nat_number_of [simp]: + "Suc (number_of v) = + (if neg (number_of v :: int) then 1 else number_of (Int.succ v))" +apply (cut_tac n = 0 in Suc_nat_number_of_add) +apply (simp cong del: if_weak_cong) +done + + +subsubsection{*Addition *} + +lemma add_nat_number_of [simp]: + "(number_of v :: nat) + number_of v' = + (if v < Int.Pls then number_of v' + else if v' < Int.Pls then number_of v + else number_of (v + v'))" + unfolding nat_number_of_def number_of_is_id numeral_simps + by (simp add: nat_add_distrib) + +lemma nat_number_of_add_1 [simp]: + "number_of v + (1::nat) = + (if v < Int.Pls then 1 else number_of (Int.succ v))" + unfolding nat_number_of_def number_of_is_id numeral_simps + by (simp add: nat_add_distrib) + +lemma nat_1_add_number_of [simp]: + "(1::nat) + number_of v = + (if v < Int.Pls then 1 else number_of (Int.succ v))" + unfolding nat_number_of_def number_of_is_id numeral_simps + by (simp add: nat_add_distrib) + +lemma nat_1_add_1 [simp]: "1 + 1 = (2::nat)" + by (rule int_int_eq [THEN iffD1]) simp + + +subsubsection{*Subtraction *} + +lemma diff_nat_eq_if: + "nat z - nat z' = + (if neg z' then nat z + else let d = z-z' in + if neg d then 0 else nat d)" +by (simp add: Let_def nat_diff_distrib [symmetric] neg_eq_less_0 not_neg_eq_ge_0) + + +lemma diff_nat_number_of [simp]: + "(number_of v :: nat) - number_of v' = + (if v' < Int.Pls then number_of v + else let d = number_of (v + uminus v') in + if neg d then 0 else nat d)" + unfolding nat_number_of_def number_of_is_id numeral_simps neg_def + by auto + +lemma nat_number_of_diff_1 [simp]: + "number_of v - (1::nat) = + (if v \ Int.Pls then 0 else number_of (Int.pred v))" + unfolding nat_number_of_def number_of_is_id numeral_simps + by auto + + +subsubsection{*Multiplication *} + +lemma mult_nat_number_of [simp]: + "(number_of v :: nat) * number_of v' = + (if v < Int.Pls then 0 else number_of (v * v'))" + unfolding nat_number_of_def number_of_is_id numeral_simps + by (simp add: nat_mult_distrib) + + +subsubsection{*Quotient *} + +lemma div_nat_number_of [simp]: + "(number_of v :: nat) div number_of v' = + (if neg (number_of v :: int) then 0 + else nat (number_of v div number_of v'))" + unfolding nat_number_of_def number_of_is_id neg_def + by (simp add: nat_div_distrib) + +lemma one_div_nat_number_of [simp]: + "Suc 0 div number_of v' = nat (1 div number_of v')" +by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) + + +subsubsection{*Remainder *} + +lemma mod_nat_number_of [simp]: + "(number_of v :: nat) mod number_of v' = + (if neg (number_of v :: int) then 0 + else if neg (number_of v' :: int) then number_of v + else nat (number_of v mod number_of v'))" + unfolding nat_number_of_def number_of_is_id neg_def + by (simp add: nat_mod_distrib) + +lemma one_mod_nat_number_of [simp]: + "Suc 0 mod number_of v' = + (if neg (number_of v' :: int) then Suc 0 + else nat (1 mod number_of v'))" +by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) + + +subsubsection{* Divisibility *} + +lemmas dvd_eq_mod_eq_0_number_of = + dvd_eq_mod_eq_0 [of "number_of x" "number_of y", standard] + +declare dvd_eq_mod_eq_0_number_of [simp] + +ML +{* +val nat_number_of_def = thm"nat_number_of_def"; + +val nat_number_of = thm"nat_number_of"; +val nat_numeral_0_eq_0 = thm"nat_numeral_0_eq_0"; +val nat_numeral_1_eq_1 = thm"nat_numeral_1_eq_1"; +val numeral_1_eq_Suc_0 = thm"numeral_1_eq_Suc_0"; +val numeral_2_eq_2 = thm"numeral_2_eq_2"; +val nat_div_distrib = thm"nat_div_distrib"; +val nat_mod_distrib = thm"nat_mod_distrib"; +val int_nat_number_of = thm"int_nat_number_of"; +val Suc_nat_eq_nat_zadd1 = thm"Suc_nat_eq_nat_zadd1"; +val Suc_nat_number_of_add = thm"Suc_nat_number_of_add"; +val Suc_nat_number_of = thm"Suc_nat_number_of"; +val add_nat_number_of = thm"add_nat_number_of"; +val diff_nat_eq_if = thm"diff_nat_eq_if"; +val diff_nat_number_of = thm"diff_nat_number_of"; +val mult_nat_number_of = thm"mult_nat_number_of"; +val div_nat_number_of = thm"div_nat_number_of"; +val mod_nat_number_of = thm"mod_nat_number_of"; +*} + + +subsection{*Comparisons*} + +subsubsection{*Equals (=) *} + +lemma eq_nat_nat_iff: + "[| (0::int) <= z; 0 <= z' |] ==> (nat z = nat z') = (z=z')" +by (auto elim!: nonneg_eq_int) + +lemma eq_nat_number_of [simp]: + "((number_of v :: nat) = number_of v') = + (if neg (number_of v :: int) then (number_of v' :: int) \ 0 + else if neg (number_of v' :: int) then (number_of v :: int) = 0 + else v = v')" + unfolding nat_number_of_def number_of_is_id neg_def + by auto + + +subsubsection{*Less-than (<) *} + +lemma less_nat_number_of [simp]: + "(number_of v :: nat) < number_of v' \ + (if v < v' then Int.Pls < v' else False)" + unfolding nat_number_of_def number_of_is_id numeral_simps + by auto + + +subsubsection{*Less-than-or-equal *} + +lemma le_nat_number_of [simp]: + "(number_of v :: nat) \ number_of v' \ + (if v \ v' then True else v \ Int.Pls)" + unfolding nat_number_of_def number_of_is_id numeral_simps + by auto + +(*Maps #n to n for n = 0, 1, 2*) +lemmas numerals = nat_numeral_0_eq_0 nat_numeral_1_eq_1 numeral_2_eq_2 + + +subsection{*Powers with Numeric Exponents*} + +text{*We cannot refer to the number @{term 2} in @{text Ring_and_Field.thy}. +We cannot prove general results about the numeral @{term "-1"}, so we have to +use @{term "- 1"} instead.*} + +lemma power2_eq_square: "(a::'a::recpower)\ = a * a" + by (simp add: numeral_2_eq_2 Power.power_Suc) + +lemma zero_power2 [simp]: "(0::'a::{semiring_1,recpower})\ = 0" + by (simp add: power2_eq_square) + +lemma one_power2 [simp]: "(1::'a::{semiring_1,recpower})\ = 1" + by (simp add: power2_eq_square) + +lemma power3_eq_cube: "(x::'a::recpower) ^ 3 = x * x * x" + apply (subgoal_tac "3 = Suc (Suc (Suc 0))") + apply (erule ssubst) + apply (simp add: power_Suc mult_ac) + apply (unfold nat_number_of_def) + apply (subst nat_eq_iff) + apply simp +done + +text{*Squares of literal numerals will be evaluated.*} +lemmas power2_eq_square_number_of = + power2_eq_square [of "number_of w", standard] +declare power2_eq_square_number_of [simp] + + +lemma zero_le_power2[simp]: "0 \ (a\::'a::{ordered_idom,recpower})" + by (simp add: power2_eq_square) + +lemma zero_less_power2[simp]: + "(0 < a\) = (a \ (0::'a::{ordered_idom,recpower}))" + by (force simp add: power2_eq_square zero_less_mult_iff linorder_neq_iff) + +lemma power2_less_0[simp]: + fixes a :: "'a::{ordered_idom,recpower}" + shows "~ (a\ < 0)" +by (force simp add: power2_eq_square mult_less_0_iff) + +lemma zero_eq_power2[simp]: + "(a\ = 0) = (a = (0::'a::{ordered_idom,recpower}))" + by (force simp add: power2_eq_square mult_eq_0_iff) + +lemma abs_power2[simp]: + "abs(a\) = (a\::'a::{ordered_idom,recpower})" + by (simp add: power2_eq_square abs_mult abs_mult_self) + +lemma power2_abs[simp]: + "(abs a)\ = (a\::'a::{ordered_idom,recpower})" + by (simp add: power2_eq_square abs_mult_self) + +lemma power2_minus[simp]: + "(- a)\ = (a\::'a::{comm_ring_1,recpower})" + by (simp add: power2_eq_square) + +lemma power2_le_imp_le: + fixes x y :: "'a::{ordered_semidom,recpower}" + shows "\x\ \ y\; 0 \ y\ \ x \ y" +unfolding numeral_2_eq_2 by (rule power_le_imp_le_base) + +lemma power2_less_imp_less: + fixes x y :: "'a::{ordered_semidom,recpower}" + shows "\x\ < y\; 0 \ y\ \ x < y" +by (rule power_less_imp_less_base) + +lemma power2_eq_imp_eq: + fixes x y :: "'a::{ordered_semidom,recpower}" + shows "\x\ = y\; 0 \ x; 0 \ y\ \ x = y" +unfolding numeral_2_eq_2 by (erule (2) power_eq_imp_eq_base, simp) + +lemma power_minus1_even[simp]: "(- 1) ^ (2*n) = (1::'a::{comm_ring_1,recpower})" +proof (induct n) + case 0 show ?case by simp +next + case (Suc n) then show ?case by (simp add: power_Suc power_add) +qed + +lemma power_minus1_odd: "(- 1) ^ Suc(2*n) = -(1::'a::{comm_ring_1,recpower})" + by (simp add: power_Suc) + +lemma power_even_eq: "(a::'a::recpower) ^ (2*n) = (a^n)^2" + by (subst mult_commute) (simp add: power_mult) + +lemma power_odd_eq: "(a::int) ^ Suc(2*n) = a * (a^n)^2" + by (simp add: power_even_eq) + +lemma power_minus_even [simp]: + "(-a) ^ (2*n) = (a::'a::{comm_ring_1,recpower}) ^ (2*n)" + by (simp add: power_minus [of a]) + +lemma zero_le_even_power'[simp]: + "0 \ (a::'a::{ordered_idom,recpower}) ^ (2*n)" +proof (induct "n") + case 0 + show ?case by (simp add: zero_le_one) +next + case (Suc n) + have "a ^ (2 * Suc n) = (a*a) * a ^ (2*n)" + by (simp add: mult_ac power_add power2_eq_square) + thus ?case + by (simp add: prems zero_le_mult_iff) +qed + +lemma odd_power_less_zero: + "(a::'a::{ordered_idom,recpower}) < 0 ==> a ^ Suc(2*n) < 0" +proof (induct "n") + case 0 + then show ?case by simp +next + case (Suc n) + have "a ^ Suc (2 * Suc n) = (a*a) * a ^ Suc(2*n)" + by (simp add: mult_ac power_add power2_eq_square) + thus ?case + by (simp del: power_Suc add: prems mult_less_0_iff mult_neg_neg) +qed + +lemma odd_0_le_power_imp_0_le: + "0 \ a ^ Suc(2*n) ==> 0 \ (a::'a::{ordered_idom,recpower})" +apply (insert odd_power_less_zero [of a n]) +apply (force simp add: linorder_not_less [symmetric]) +done + +text{*Simprules for comparisons where common factors can be cancelled.*} +lemmas zero_compare_simps = + add_strict_increasing add_strict_increasing2 add_increasing + zero_le_mult_iff zero_le_divide_iff + zero_less_mult_iff zero_less_divide_iff + mult_le_0_iff divide_le_0_iff + mult_less_0_iff divide_less_0_iff + zero_le_power2 power2_less_0 + +subsubsection{*Nat *} + +lemma Suc_pred': "0 < n ==> n = Suc(n - 1)" +by (simp add: numerals) + +(*Expresses a natural number constant as the Suc of another one. + NOT suitable for rewriting because n recurs in the condition.*) +lemmas expand_Suc = Suc_pred' [of "number_of v", standard] + +subsubsection{*Arith *} + +lemma Suc_eq_add_numeral_1: "Suc n = n + 1" +by (simp add: numerals) + +lemma Suc_eq_add_numeral_1_left: "Suc n = 1 + n" +by (simp add: numerals) + +(* These two can be useful when m = number_of... *) + +lemma add_eq_if: "(m::nat) + n = (if m=0 then n else Suc ((m - 1) + n))" + unfolding One_nat_def by (cases m) simp_all + +lemma mult_eq_if: "(m::nat) * n = (if m=0 then 0 else n + ((m - 1) * n))" + unfolding One_nat_def by (cases m) simp_all + +lemma power_eq_if: "(p ^ m :: nat) = (if m=0 then 1 else p * (p ^ (m - 1)))" + unfolding One_nat_def by (cases m) simp_all + + +subsection{*Comparisons involving (0::nat) *} + +text{*Simplification already does @{term "n<0"}, @{term "n\0"} and @{term "0\n"}.*} + +lemma eq_number_of_0 [simp]: + "number_of v = (0::nat) \ v \ Int.Pls" + unfolding nat_number_of_def number_of_is_id numeral_simps + by auto + +lemma eq_0_number_of [simp]: + "(0::nat) = number_of v \ v \ Int.Pls" +by (rule trans [OF eq_sym_conv eq_number_of_0]) + +lemma less_0_number_of [simp]: + "(0::nat) < number_of v \ Int.Pls < v" + unfolding nat_number_of_def number_of_is_id numeral_simps + by simp + +lemma neg_imp_number_of_eq_0: "neg (number_of v :: int) ==> number_of v = (0::nat)" +by (simp del: nat_numeral_0_eq_0 add: nat_numeral_0_eq_0 [symmetric]) + + + +subsection{*Comparisons involving @{term Suc} *} + +lemma eq_number_of_Suc [simp]: + "(number_of v = Suc n) = + (let pv = number_of (Int.pred v) in + if neg pv then False else nat pv = n)" +apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less + number_of_pred nat_number_of_def + split add: split_if) +apply (rule_tac x = "number_of v" in spec) +apply (auto simp add: nat_eq_iff) +done + +lemma Suc_eq_number_of [simp]: + "(Suc n = number_of v) = + (let pv = number_of (Int.pred v) in + if neg pv then False else nat pv = n)" +by (rule trans [OF eq_sym_conv eq_number_of_Suc]) + +lemma less_number_of_Suc [simp]: + "(number_of v < Suc n) = + (let pv = number_of (Int.pred v) in + if neg pv then True else nat pv < n)" +apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less + number_of_pred nat_number_of_def + split add: split_if) +apply (rule_tac x = "number_of v" in spec) +apply (auto simp add: nat_less_iff) +done + +lemma less_Suc_number_of [simp]: + "(Suc n < number_of v) = + (let pv = number_of (Int.pred v) in + if neg pv then False else n < nat pv)" +apply (simp only: simp_thms Let_def neg_eq_less_0 linorder_not_less + number_of_pred nat_number_of_def + split add: split_if) +apply (rule_tac x = "number_of v" in spec) +apply (auto simp add: zless_nat_eq_int_zless) +done + +lemma le_number_of_Suc [simp]: + "(number_of v <= Suc n) = + (let pv = number_of (Int.pred v) in + if neg pv then True else nat pv <= n)" +by (simp add: Let_def less_Suc_number_of linorder_not_less [symmetric]) + +lemma le_Suc_number_of [simp]: + "(Suc n <= number_of v) = + (let pv = number_of (Int.pred v) in + if neg pv then False else n <= nat pv)" +by (simp add: Let_def less_number_of_Suc linorder_not_less [symmetric]) + + +lemma eq_number_of_Pls_Min: "(Numeral0 ::int) ~= number_of Int.Min" +by auto + + + +subsection{*Max and Min Combined with @{term Suc} *} + +lemma max_number_of_Suc [simp]: + "max (Suc n) (number_of v) = + (let pv = number_of (Int.pred v) in + if neg pv then Suc n else Suc(max n (nat pv)))" +apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def + split add: split_if nat.split) +apply (rule_tac x = "number_of v" in spec) +apply auto +done + +lemma max_Suc_number_of [simp]: + "max (number_of v) (Suc n) = + (let pv = number_of (Int.pred v) in + if neg pv then Suc n else Suc(max (nat pv) n))" +apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def + split add: split_if nat.split) +apply (rule_tac x = "number_of v" in spec) +apply auto +done + +lemma min_number_of_Suc [simp]: + "min (Suc n) (number_of v) = + (let pv = number_of (Int.pred v) in + if neg pv then 0 else Suc(min n (nat pv)))" +apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def + split add: split_if nat.split) +apply (rule_tac x = "number_of v" in spec) +apply auto +done + +lemma min_Suc_number_of [simp]: + "min (number_of v) (Suc n) = + (let pv = number_of (Int.pred v) in + if neg pv then 0 else Suc(min (nat pv) n))" +apply (simp only: Let_def neg_eq_less_0 number_of_pred nat_number_of_def + split add: split_if nat.split) +apply (rule_tac x = "number_of v" in spec) +apply auto +done + +subsection{*Literal arithmetic involving powers*} + +lemma nat_power_eq: "(0::int) <= z ==> nat (z^n) = nat z ^ n" +apply (induct "n") +apply (simp_all (no_asm_simp) add: nat_mult_distrib) +done + +lemma power_nat_number_of: + "(number_of v :: nat) ^ n = + (if neg (number_of v :: int) then 0^n else nat ((number_of v :: int) ^ n))" +by (simp only: simp_thms neg_nat not_neg_eq_ge_0 nat_number_of_def nat_power_eq + split add: split_if cong: imp_cong) + + +lemmas power_nat_number_of_number_of = power_nat_number_of [of _ "number_of w", standard] +declare power_nat_number_of_number_of [simp] + + + +text{*For arbitrary rings*} + +lemma power_number_of_even: + fixes z :: "'a::{number_ring,recpower}" + shows "z ^ number_of (Int.Bit0 w) = (let w = z ^ (number_of w) in w * w)" +unfolding Let_def nat_number_of_def number_of_Bit0 +apply (rule_tac x = "number_of w" in spec, clarify) +apply (case_tac " (0::int) <= x") +apply (auto simp add: nat_mult_distrib power_even_eq power2_eq_square) +done + +lemma power_number_of_odd: + fixes z :: "'a::{number_ring,recpower}" + shows "z ^ number_of (Int.Bit1 w) = (if (0::int) <= number_of w + then (let w = z ^ (number_of w) in z * w * w) else 1)" +unfolding Let_def nat_number_of_def number_of_Bit1 +apply (rule_tac x = "number_of w" in spec, auto) +apply (simp only: nat_add_distrib nat_mult_distrib) +apply simp +apply (auto simp add: nat_add_distrib nat_mult_distrib power_even_eq power2_eq_square neg_nat power_Suc) +done + +lemmas zpower_number_of_even = power_number_of_even [where 'a=int] +lemmas zpower_number_of_odd = power_number_of_odd [where 'a=int] + +lemmas power_number_of_even_number_of [simp] = + power_number_of_even [of "number_of v", standard] + +lemmas power_number_of_odd_number_of [simp] = + power_number_of_odd [of "number_of v", standard] + + + +ML +{* +val numeral_ss = @{simpset} addsimps @{thms numerals}; + +val nat_bin_arith_setup = + Lin_Arith.map_data + (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} => + {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, + inj_thms = inj_thms, + lessD = lessD, neqE = neqE, + simpset = simpset addsimps @{thms neg_simps} @ + [@{thm Suc_nat_number_of}, @{thm int_nat_number_of}]}) +*} + +declaration {* K nat_bin_arith_setup *} + +(* Enable arith to deal with div/mod k where k is a numeral: *) +declare split_div[of _ _ "number_of k", standard, arith_split] +declare split_mod[of _ _ "number_of k", standard, arith_split] + +lemma nat_number_of_Pls: "Numeral0 = (0::nat)" + by (simp add: number_of_Pls nat_number_of_def) + +lemma nat_number_of_Min: "number_of Int.Min = (0::nat)" + apply (simp only: number_of_Min nat_number_of_def nat_zminus_int) + done + +lemma nat_number_of_Bit0: + "number_of (Int.Bit0 w) = (let n::nat = number_of w in n + n)" + unfolding nat_number_of_def number_of_is_id numeral_simps Let_def + by auto + +lemma nat_number_of_Bit1: + "number_of (Int.Bit1 w) = + (if neg (number_of w :: int) then 0 + else let n = number_of w in Suc (n + n))" + unfolding nat_number_of_def number_of_is_id numeral_simps neg_def Let_def + by auto + +lemmas nat_number = + nat_number_of_Pls nat_number_of_Min + nat_number_of_Bit0 nat_number_of_Bit1 + +lemma Let_Suc [simp]: "Let (Suc n) f == f (Suc n)" + by (simp add: Let_def) + +lemma power_m1_even: "(-1) ^ (2*n) = (1::'a::{number_ring,recpower})" +by (simp add: power_mult power_Suc); + +lemma power_m1_odd: "(-1) ^ Suc(2*n) = (-1::'a::{number_ring,recpower})" +by (simp add: power_mult power_Suc); + + +subsection{*Literal arithmetic and @{term of_nat}*} + +lemma of_nat_double: + "0 \ x ==> of_nat (nat (2 * x)) = of_nat (nat x) + of_nat (nat x)" +by (simp only: mult_2 nat_add_distrib of_nat_add) + +lemma nat_numeral_m1_eq_0: "-1 = (0::nat)" +by (simp only: nat_number_of_def) + +lemma of_nat_number_of_lemma: + "of_nat (number_of v :: nat) = + (if 0 \ (number_of v :: int) + then (number_of v :: 'a :: number_ring) + else 0)" +by (simp add: int_number_of_def nat_number_of_def number_of_eq of_nat_nat); + +lemma of_nat_number_of_eq [simp]: + "of_nat (number_of v :: nat) = + (if neg (number_of v :: int) then 0 + else (number_of v :: 'a :: number_ring))" +by (simp only: of_nat_number_of_lemma neg_def, simp) + + +subsection {*Lemmas for the Combination and Cancellation Simprocs*} + +lemma nat_number_of_add_left: + "number_of v + (number_of v' + (k::nat)) = + (if neg (number_of v :: int) then number_of v' + k + else if neg (number_of v' :: int) then number_of v + k + else number_of (v + v') + k)" + unfolding nat_number_of_def number_of_is_id neg_def + by auto + +lemma nat_number_of_mult_left: + "number_of v * (number_of v' * (k::nat)) = + (if v < Int.Pls then 0 + else number_of (v * v') * k)" +by simp + + +subsubsection{*For @{text combine_numerals}*} + +lemma left_add_mult_distrib: "i*u + (j*u + k) = (i+j)*u + (k::nat)" +by (simp add: add_mult_distrib) + + +subsubsection{*For @{text cancel_numerals}*} + +lemma nat_diff_add_eq1: + "j <= (i::nat) ==> ((i*u + m) - (j*u + n)) = (((i-j)*u + m) - n)" +by (simp split add: nat_diff_split add: add_mult_distrib) + +lemma nat_diff_add_eq2: + "i <= (j::nat) ==> ((i*u + m) - (j*u + n)) = (m - ((j-i)*u + n))" +by (simp split add: nat_diff_split add: add_mult_distrib) + +lemma nat_eq_add_iff1: + "j <= (i::nat) ==> (i*u + m = j*u + n) = ((i-j)*u + m = n)" +by (auto split add: nat_diff_split simp add: add_mult_distrib) + +lemma nat_eq_add_iff2: + "i <= (j::nat) ==> (i*u + m = j*u + n) = (m = (j-i)*u + n)" +by (auto split add: nat_diff_split simp add: add_mult_distrib) + +lemma nat_less_add_iff1: + "j <= (i::nat) ==> (i*u + m < j*u + n) = ((i-j)*u + m < n)" +by (auto split add: nat_diff_split simp add: add_mult_distrib) + +lemma nat_less_add_iff2: + "i <= (j::nat) ==> (i*u + m < j*u + n) = (m < (j-i)*u + n)" +by (auto split add: nat_diff_split simp add: add_mult_distrib) + +lemma nat_le_add_iff1: + "j <= (i::nat) ==> (i*u + m <= j*u + n) = ((i-j)*u + m <= n)" +by (auto split add: nat_diff_split simp add: add_mult_distrib) + +lemma nat_le_add_iff2: + "i <= (j::nat) ==> (i*u + m <= j*u + n) = (m <= (j-i)*u + n)" +by (auto split add: nat_diff_split simp add: add_mult_distrib) + + +subsubsection{*For @{text cancel_numeral_factors} *} + +lemma nat_mult_le_cancel1: "(0::nat) < k ==> (k*m <= k*n) = (m<=n)" +by auto + +lemma nat_mult_less_cancel1: "(0::nat) < k ==> (k*m < k*n) = (m (k*m = k*n) = (m=n)" +by auto + +lemma nat_mult_div_cancel1: "(0::nat) < k ==> (k*m) div (k*n) = (m div n)" +by auto + +lemma nat_mult_dvd_cancel_disj[simp]: + "(k*m) dvd (k*n) = (k=0 | m dvd (n::nat))" +by(auto simp: dvd_eq_mod_eq_0 mod_mult_distrib2[symmetric]) + +lemma nat_mult_dvd_cancel1: "0 < k \ (k*m) dvd (k*n::nat) = (m dvd n)" +by(auto) + + +subsubsection{*For @{text cancel_factor} *} + +lemma nat_mult_le_cancel_disj: "(k*m <= k*n) = ((0::nat) < k --> m<=n)" +by auto + +lemma nat_mult_less_cancel_disj: "(k*m < k*n) = ((0::nat) < k & m Suc m - n = m - (n - Numeral1)" +by (simp add: numeral_0_eq_0 numeral_1_eq_1 split add: nat_diff_split) + +text {*Now just instantiating @{text n} to @{text "number_of v"} does + the right simplification, but with some redundant inequality + tests.*} +lemma neg_number_of_pred_iff_0: + "neg (number_of (Int.pred v)::int) = (number_of v = (0::nat))" +apply (subgoal_tac "neg (number_of (Int.pred v)) = (number_of v < Suc 0) ") +apply (simp only: less_Suc_eq_le le_0_eq) +apply (subst less_number_of_Suc, simp) +done + +text{*No longer required as a simprule because of the @{text inverse_fold} + simproc*} +lemma Suc_diff_number_of: + "Int.Pls < v ==> + Suc m - (number_of v) = m - (number_of (Int.pred v))" +apply (subst Suc_diff_eq_diff_pred) +apply simp +apply (simp del: nat_numeral_1_eq_1) +apply (auto simp only: diff_nat_number_of less_0_number_of [symmetric] + neg_number_of_pred_iff_0) +done + +lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n" +by (simp add: numerals split add: nat_diff_split) + + +subsubsection{*For @{term nat_case} and @{term nat_rec}*} + +lemma nat_case_number_of [simp]: + "nat_case a f (number_of v) = + (let pv = number_of (Int.pred v) in + if neg pv then a else f (nat pv))" +by (simp split add: nat.split add: Let_def neg_number_of_pred_iff_0) + +lemma nat_case_add_eq_if [simp]: + "nat_case a f ((number_of v) + n) = + (let pv = number_of (Int.pred v) in + if neg pv then nat_case a f n else f (nat pv + n))" +apply (subst add_eq_if) +apply (simp split add: nat.split + del: nat_numeral_1_eq_1 + add: nat_numeral_1_eq_1 [symmetric] + numeral_1_eq_Suc_0 [symmetric] + neg_number_of_pred_iff_0) +done + +lemma nat_rec_number_of [simp]: + "nat_rec a f (number_of v) = + (let pv = number_of (Int.pred v) in + if neg pv then a else f (nat pv) (nat_rec a f (nat pv)))" +apply (case_tac " (number_of v) ::nat") +apply (simp_all (no_asm_simp) add: Let_def neg_number_of_pred_iff_0) +apply (simp split add: split_if_asm) +done + +lemma nat_rec_add_eq_if [simp]: + "nat_rec a f (number_of v + n) = + (let pv = number_of (Int.pred v) in + if neg pv then nat_rec a f n + else f (nat pv + n) (nat_rec a f (nat pv + n)))" +apply (subst add_eq_if) +apply (simp split add: nat.split + del: nat_numeral_1_eq_1 + add: nat_numeral_1_eq_1 [symmetric] + numeral_1_eq_Suc_0 [symmetric] + neg_number_of_pred_iff_0) +done + + +subsubsection{*Various Other Lemmas*} + +text {*Evens and Odds, for Mutilated Chess Board*} + +text{*Lemmas for specialist use, NOT as default simprules*} +lemma nat_mult_2: "2 * z = (z+z::nat)" +proof - + have "2*z = (1 + 1)*z" by simp + also have "... = z+z" by (simp add: left_distrib) + finally show ?thesis . +qed + +lemma nat_mult_2_right: "z * 2 = (z+z::nat)" +by (subst mult_commute, rule nat_mult_2) + +text{*Case analysis on @{term "n<2"}*} +lemma less_2_cases: "(n::nat) < 2 ==> n = 0 | n = Suc 0" +by arith + +lemma div2_Suc_Suc [simp]: "Suc(Suc m) div 2 = Suc (m div 2)" +by arith + +lemma add_self_div_2 [simp]: "(m + m) div 2 = (m::nat)" +by (simp add: nat_mult_2 [symmetric]) + +lemma mod2_Suc_Suc [simp]: "Suc(Suc(m)) mod 2 = m mod 2" +apply (subgoal_tac "m mod 2 < 2") +apply (erule less_2_cases [THEN disjE]) +apply (simp_all (no_asm_simp) add: Let_def mod_Suc nat_1) +done + +lemma mod2_gr_0 [simp]: "!!m::nat. (0 < m mod 2) = (m mod 2 = 1)" +apply (subgoal_tac "m mod 2 < 2") +apply (force simp del: mod_less_divisor, simp) +done + +text{*Removal of Small Numerals: 0, 1 and (in additive positions) 2*} + +lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)" +by simp + +lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)" +by simp + +text{*Can be used to eliminate long strings of Sucs, but not by default*} +lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n" +by simp + + +text{*These lemmas collapse some needless occurrences of Suc: + at least three Sucs, since two and fewer are rewritten back to Suc again! + We already have some rules to simplify operands smaller than 3.*} + +lemma div_Suc_eq_div_add3 [simp]: "m div (Suc (Suc (Suc n))) = m div (3+n)" +by (simp add: Suc3_eq_add_3) + +lemma mod_Suc_eq_mod_add3 [simp]: "m mod (Suc (Suc (Suc n))) = m mod (3+n)" +by (simp add: Suc3_eq_add_3) + +lemma Suc_div_eq_add3_div: "(Suc (Suc (Suc m))) div n = (3+m) div n" +by (simp add: Suc3_eq_add_3) + +lemma Suc_mod_eq_add3_mod: "(Suc (Suc (Suc m))) mod n = (3+m) mod n" +by (simp add: Suc3_eq_add_3) + +lemmas Suc_div_eq_add3_div_number_of = + Suc_div_eq_add3_div [of _ "number_of v", standard] +declare Suc_div_eq_add3_div_number_of [simp] + +lemmas Suc_mod_eq_add3_mod_number_of = + Suc_mod_eq_add3_mod [of _ "number_of v", standard] +declare Suc_mod_eq_add3_mod_number_of [simp] + +end \ No newline at end of file diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Nominal/Examples/Fsub.thy --- a/src/HOL/Nominal/Examples/Fsub.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Nominal/Examples/Fsub.thy Mon Apr 27 07:26:17 2009 -0700 @@ -245,7 +245,7 @@ apply (simp add: dj_perm_forget[OF dj_tyvrs_vrs]) done -lemma ty_vrs_fresh[fresh]: +lemma ty_vrs_fresh: fixes x::"vrs" and T::"ty" shows "x \ T" @@ -422,7 +422,7 @@ by (nominal_induct T avoiding: X T' rule: ty.strong_induct) (perm_simp add: fresh_left)+ -lemma type_subst_fresh[fresh]: +lemma type_subst_fresh: fixes X::"tyvrs" assumes "X \ T" and "X \ P" shows "X \ T[Y \ P]\<^sub>\" @@ -430,7 +430,7 @@ by (nominal_induct T avoiding: X Y P rule:ty.strong_induct) (auto simp add: abs_fresh) -lemma fresh_type_subst_fresh[fresh]: +lemma fresh_type_subst_fresh: assumes "X\T'" shows "X\T[X \ T']\<^sub>\" using assms @@ -458,18 +458,19 @@ | "(VarB X U)[Y \ T]\<^sub>b = VarB X (U[Y \ T]\<^sub>\)" by auto -lemma binding_subst_fresh[fresh]: +lemma binding_subst_fresh: fixes X::"tyvrs" assumes "X \ a" and "X \ P" shows "X \ a[Y \ P]\<^sub>b" using assms -by (nominal_induct a rule:binding.strong_induct) - (auto simp add: freshs) +by (nominal_induct a rule: binding.strong_induct) + (auto simp add: type_subst_fresh) -lemma binding_subst_identity: "X \ B \ B[X \ U]\<^sub>b = B" - by (induct B rule: binding.induct) - (simp_all add: fresh_atm type_subst_identity) +lemma binding_subst_identity: + shows "X \ B \ B[X \ U]\<^sub>b = B" +by (induct B rule: binding.induct) + (simp_all add: fresh_atm type_subst_identity) consts subst_tyc :: "env \ tyvrs \ ty \ env" ("_[_ \ _]\<^sub>e" [100,100,100] 100) @@ -478,14 +479,14 @@ "([])[Y \ T]\<^sub>e= []" "(B#\)[Y \ T]\<^sub>e = (B[Y \ T]\<^sub>b)#(\[Y \ T]\<^sub>e)" -lemma ctxt_subst_fresh'[fresh]: +lemma ctxt_subst_fresh': fixes X::"tyvrs" assumes "X \ \" and "X \ P" shows "X \ \[Y \ P]\<^sub>e" using assms by (induct \) - (auto simp add: fresh_list_cons freshs) + (auto simp add: fresh_list_cons binding_subst_fresh) lemma ctxt_subst_mem_TVarB: "TVarB X T \ set \ \ TVarB X (T[Y \ U]\<^sub>\) \ set (\[Y \ U]\<^sub>e)" by (induct \) auto @@ -1188,8 +1189,8 @@ using assms by (induct, auto) nominal_inductive typing - by (auto dest!: typing_ok intro: closed_in_fresh fresh_domain - simp: abs_fresh fresh_prod fresh_atm freshs valid_ty_domain_fresh fresh_trm_domain) +by (auto dest!: typing_ok intro: closed_in_fresh fresh_domain type_subst_fresh + simp: abs_fresh fresh_type_subst_fresh ty_vrs_fresh valid_ty_domain_fresh fresh_trm_domain) lemma ok_imp_VarB_closed_in: assumes ok: "\ \ ok" diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Nominal/Nominal.thy --- a/src/HOL/Nominal/Nominal.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Nominal/Nominal.thy Mon Apr 27 07:26:17 2009 -0700 @@ -18,25 +18,98 @@ types 'x prm = "('x \ 'x) list" -(* polymorphic operations for permutation and swapping *) +(* polymorphic constants for permutation and swapping *) consts perm :: "'x prm \ 'a \ 'a" (infixr "\" 80) swap :: "('x \ 'x) \ 'x \ 'x" +(* a "private" copy of the option type used in the abstraction function *) +datatype 'a noption = nSome 'a | nNone + +(* a "private" copy of the product type used in the nominal induct method *) +datatype ('a,'b) nprod = nPair 'a 'b + (* an auxiliary constant for the decision procedure involving *) -(* permutations (to avoid loops when using perm-composition) *) +(* permutations (to avoid loops when using perm-compositions) *) constdefs "perm_aux pi x \ pi\x" -(* permutation on functions *) -defs (unchecked overloaded) - perm_fun_def: "pi\(f::'a\'b) \ (\x. pi\f((rev pi)\x))" - -(* permutation on bools *) -primrec (unchecked perm_bool) - true_eqvt: "pi\True = True" - false_eqvt: "pi\False = False" - +(* overloaded permutation operations *) +overloading + perm_fun \ "perm :: 'x prm \ ('a\'b) \ ('a\'b)" (unchecked) + perm_bool \ "perm :: 'x prm \ bool \ bool" (unchecked) + perm_unit \ "perm :: 'x prm \ unit \ unit" (unchecked) + perm_prod \ "perm :: 'x prm \ ('a\'b) \ ('a\'b)" (unchecked) + perm_list \ "perm :: 'x prm \ 'a list \ 'a list" (unchecked) + perm_option \ "perm :: 'x prm \ 'a option \ 'a option" (unchecked) + perm_char \ "perm :: 'x prm \ char \ char" (unchecked) + perm_nat \ "perm :: 'x prm \ nat \ nat" (unchecked) + perm_int \ "perm :: 'x prm \ int \ int" (unchecked) + + perm_noption \ "perm :: 'x prm \ 'a noption \ 'a noption" (unchecked) + perm_nprod \ "perm :: 'x prm \ ('a, 'b) nprod \ ('a, 'b) nprod" (unchecked) +begin + +definition + perm_fun_def: "perm_fun pi (f::'a\'b) \ (\x. pi\f((rev pi)\x))" + +fun + perm_bool :: "'x prm \ bool \ bool" +where + true_eqvt: "perm_bool pi True = True" +| false_eqvt: "perm_bool pi False = False" + +fun + perm_unit :: "'x prm \ unit \ unit" +where + "perm_unit pi () = ()" + +fun + perm_prod :: "'x prm \ ('a\'b) \ ('a\'b)" +where + "perm_prod pi (x,y) = (pi\x,pi\y)" + +fun + perm_list :: "'x prm \ 'a list \ 'a list" +where + nil_eqvt: "perm_list pi [] = []" +| cons_eqvt: "perm_list pi (x#xs) = (pi\x)#(pi\xs)" + +fun + perm_option :: "'x prm \ 'a option \ 'a option" +where + some_eqvt: "perm_option pi (Some x) = Some (pi\x)" +| none_eqvt: "perm_option pi None = None" + +definition + perm_char :: "'x prm \ char \ char" +where + perm_char_def: "perm_char pi c \ c" + +definition + perm_nat :: "'x prm \ nat \ nat" +where + perm_nat_def: "perm_nat pi i \ i" + +definition + perm_int :: "'x prm \ int \ int" +where + perm_int_def: "perm_int pi i \ i" + +fun + perm_noption :: "'x prm \ 'a noption \ 'a noption" +where + nsome_eqvt: "perm_noption pi (nSome x) = nSome (pi\x)" +| nnone_eqvt: "perm_noption pi nNone = nNone" + +fun + perm_nprod :: "'x prm \ ('a, 'b) nprod \ ('a, 'b) nprod" +where + "perm_nprod pi (nPair x y) = nPair (pi\x) (pi\y)" +end + + +(* permutations on booleans *) lemma perm_bool: shows "pi\(b::bool) = b" by (cases b) auto @@ -54,8 +127,7 @@ lemma if_eqvt: fixes pi::"'a prm" shows "pi\(if b then c1 else c2) = (if (pi\b) then (pi\c1) else (pi\c2))" -apply(simp add: perm_fun_def) -done + by (simp add: perm_fun_def) lemma imp_eqvt: shows "pi\(A\B) = ((pi\A)\(pi\B))" @@ -82,13 +154,7 @@ shows "(pi\(X\Y)) = (pi\X) \ (pi\Y)" by (simp add: perm_fun_def perm_bool Un_iff [unfolded mem_def] expand_fun_eq) -(* permutation on units and products *) -primrec (unchecked perm_unit) - "pi\() = ()" - -primrec (unchecked perm_prod) - "pi\(x,y) = (pi\x,pi\y)" - +(* permutations on products *) lemma fst_eqvt: "pi\(fst x) = fst (pi\x)" by (cases x) simp @@ -98,10 +164,6 @@ by (cases x) simp (* permutation on lists *) -primrec (unchecked perm_list) - nil_eqvt: "pi\[] = []" - cons_eqvt: "pi\(x#xs) = (pi\x)#(pi\xs)" - lemma append_eqvt: fixes pi :: "'x prm" and l1 :: "'a list" @@ -115,41 +177,12 @@ shows "pi\(rev l) = rev (pi\l)" by (induct l) (simp_all add: append_eqvt) -(* permutation on options *) - -primrec (unchecked perm_option) - some_eqvt: "pi\Some(x) = Some(pi\x)" - none_eqvt: "pi\None = None" - -(* a "private" copy of the option type used in the abstraction function *) -datatype 'a noption = nSome 'a | nNone - -primrec (unchecked perm_noption) - nSome_eqvt: "pi\nSome(x) = nSome(pi\x)" - nNone_eqvt: "pi\nNone = nNone" - -(* a "private" copy of the product type used in the nominal induct method *) -datatype ('a,'b) nprod = nPair 'a 'b - -primrec (unchecked perm_nprod) - perm_nProd_def: "pi\(nPair x1 x2) = nPair (pi\x1) (pi\x2)" - -(* permutation on characters (used in strings) *) -defs (unchecked overloaded) - perm_char_def: "pi\(c::char) \ c" - +(* permutation on characters and strings *) lemma perm_string: fixes s::"string" shows "pi\s = s" -by (induct s)(auto simp add: perm_char_def) - -(* permutation on ints *) -defs (unchecked overloaded) - perm_int_def: "pi\(i::int) \ i" - -(* permutation on nats *) -defs (unchecked overloaded) - perm_nat_def: "pi\(i::nat) \ i" + by (induct s)(auto simp add: perm_char_def) + section {* permutation equality *} (*==============================*) @@ -170,11 +203,12 @@ supports :: "'x set \ 'a \ bool" (infixl "supports" 80) "S supports x \ \a b. (a\S \ b\S \ [(a,b)]\x=x)" +(* lemmas about supp *) lemma supp_fresh_iff: fixes x :: "'a" shows "(supp x) = {a::'x. \a\x}" -apply(simp add: fresh_def) -done + by (simp add: fresh_def) + lemma supp_unit: shows "supp () = {}" @@ -205,14 +239,13 @@ fixes x :: "'a" and xs :: "'a list" shows "supp (x#xs) = (supp x)\(supp xs)" -apply(auto simp add: supp_def Collect_imp_eq Collect_neg_eq) -done + by (auto simp add: supp_def Collect_imp_eq Collect_neg_eq) lemma supp_list_append: fixes xs :: "'a list" and ys :: "'a list" shows "supp (xs@ys) = (supp xs)\(supp ys)" - by (induct xs, auto simp add: supp_list_nil supp_list_cons) + by (induct xs) (auto simp add: supp_list_nil supp_list_cons) lemma supp_list_rev: fixes xs :: "'a list" @@ -221,47 +254,40 @@ lemma supp_bool: fixes x :: "bool" - shows "supp (x) = {}" - apply(case_tac "x") - apply(simp_all add: supp_def) -done + shows "supp x = {}" + by (cases "x") (simp_all add: supp_def) lemma supp_some: fixes x :: "'a" shows "supp (Some x) = (supp x)" - apply(simp add: supp_def) - done + by (simp add: supp_def) lemma supp_none: fixes x :: "'a" shows "supp (None) = {}" - apply(simp add: supp_def) - done + by (simp add: supp_def) lemma supp_int: fixes i::"int" shows "supp (i) = {}" - apply(simp add: supp_def perm_int_def) - done + by (simp add: supp_def perm_int_def) lemma supp_nat: fixes n::"nat" - shows "supp (n) = {}" - apply(simp add: supp_def perm_nat_def) - done + shows "(supp n) = {}" + by (simp add: supp_def perm_nat_def) lemma supp_char: fixes c::"char" - shows "supp (c) = {}" - apply(simp add: supp_def perm_char_def) - done + shows "(supp c) = {}" + by (simp add: supp_def perm_char_def) lemma supp_string: fixes s::"string" - shows "supp (s) = {}" -apply(simp add: supp_def perm_string) -done - + shows "(supp s) = {}" + by (simp add: supp_def perm_string) + +(* lemmas about freshness *) lemma fresh_set_empty: shows "a\{}" by (simp add: fresh_def supp_set_empty) @@ -344,7 +370,6 @@ by (simp add: fresh_def supp_bool) text {* Normalization of freshness results; cf.\ @{text nominal_induct} *} - lemma fresh_unit_elim: shows "(a\() \ PROP C) \ PROP C" by (simp add: fresh_def supp_unit) @@ -371,63 +396,6 @@ Simplifier.map_ss (fn ss => ss setmksimps (mksimps mksimps_pairs)) *} -section {* generalisation of freshness to lists and sets of atoms *} -(*================================================================*) - -consts - fresh_star :: "'b \ 'a \ bool" ("_ \* _" [100,100] 100) - -defs (overloaded) - fresh_star_set: "xs\*c \ \x\xs. x\c" - -defs (overloaded) - fresh_star_list: "xs\*c \ \x\set xs. x\c" - -lemmas fresh_star_def = fresh_star_list fresh_star_set - -lemma fresh_star_prod_set: - fixes xs::"'a set" - shows "xs\*(a,b) = (xs\*a \ xs\*b)" -by (auto simp add: fresh_star_def fresh_prod) - -lemma fresh_star_prod_list: - fixes xs::"'a list" - shows "xs\*(a,b) = (xs\*a \ xs\*b)" -by (auto simp add: fresh_star_def fresh_prod) - -lemmas fresh_star_prod = fresh_star_prod_list fresh_star_prod_set - -lemma fresh_star_set_eq: "set xs \* c = xs \* c" - by (simp add: fresh_star_def) - -lemma fresh_star_Un_elim: - "((S \ T) \* c \ PROP C) \ (S \* c \ T \* c \ PROP C)" - apply rule - apply (simp_all add: fresh_star_def) - apply (erule meta_mp) - apply blast - done - -lemma fresh_star_insert_elim: - "(insert x S \* c \ PROP C) \ (x \ c \ S \* c \ PROP C)" - by rule (simp_all add: fresh_star_def) - -lemma fresh_star_empty_elim: - "({} \* c \ PROP C) \ PROP C" - by (simp add: fresh_star_def) - -text {* Normalization of freshness results; cf.\ @{text nominal_induct} *} - -lemma fresh_star_unit_elim: - shows "((a::'a set)\*() \ PROP C) \ PROP C" - and "((b::'a list)\*() \ PROP C) \ PROP C" - by (simp_all add: fresh_star_def fresh_def supp_unit) - -lemma fresh_star_prod_elim: - shows "((a::'a set)\*(x,y) \ PROP C) \ (a\*x \ a\*y \ PROP C)" - and "((b::'a list)\*(x,y) \ PROP C) \ (b\*x \ b\*y \ PROP C)" - by (rule, simp_all add: fresh_star_prod)+ - section {* Abstract Properties for Permutations and Atoms *} (*=========================================================*) @@ -487,7 +455,7 @@ shows "swap (a,b) c = (if a=c then b else (if b=c then a else c))" using a by (simp only: at_def) -(* rules to calculate simple premutations *) +(* rules to calculate simple permutations *) lemmas at_calc = at2 at1 at3 lemma at_swap_simps: @@ -682,7 +650,6 @@ shows "pi1 \ pi2 \ (rev pi1) \ (rev pi2)" by (simp add: at_prm_rev_eq[OF at]) - lemma at_ds1: fixes a :: "'x" assumes at: "at TYPE('x)" @@ -838,15 +805,18 @@ by (auto intro: ex_in_inf[OF at, OF fs] simp add: fresh_def) lemma at_finite_select: - shows "at (TYPE('a)) \ finite (S::'a set) \ \x. x \ S" - apply (drule Diff_infinite_finite) - apply (simp add: at_def) - apply blast - apply (subgoal_tac "UNIV - S \ {}") - apply (simp only: ex_in_conv [symmetric]) - apply blast - apply (rule notI) - apply simp + fixes S::"'a set" + assumes a: "at TYPE('a)" + and b: "finite S" + shows "\x. x \ S" + using a b + apply(drule_tac S="UNIV::'a set" in Diff_infinite_finite) + apply(simp add: at_def) + apply(subgoal_tac "UNIV - S \ {}") + apply(simp only: ex_in_conv [symmetric]) + apply(blast) + apply(rule notI) + apply(simp) done lemma at_different: @@ -1222,8 +1192,8 @@ assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" shows "pi\(x=y) = (pi\x = pi\y)" -using assms -by (auto simp add: pt_bij perm_bool) + using pt at + by (auto simp add: pt_bij perm_bool) lemma pt_bij3: fixes pi :: "'x prm" @@ -1231,7 +1201,7 @@ and y :: "'a" assumes a: "x=y" shows "(pi\x = pi\y)" -using a by simp + using a by simp lemma pt_bij4: fixes pi :: "'x prm" @@ -1241,7 +1211,7 @@ and at: "at TYPE('x)" and a: "pi\x = pi\y" shows "x = y" -using a by (simp add: pt_bij[OF pt, OF at]) + using a by (simp add: pt_bij[OF pt, OF at]) lemma pt_swap_bij: fixes a :: "'x" @@ -1574,35 +1544,6 @@ apply(simp add: pt_rev_pi[OF ptb, OF at]) done -lemma pt_fresh_star_bij_ineq: - fixes pi :: "'x prm" - and x :: "'a" - and a :: "'y set" - and b :: "'y list" - assumes pta: "pt TYPE('a) TYPE('x)" - and ptb: "pt TYPE('y) TYPE('x)" - and at: "at TYPE('x)" - and cp: "cp TYPE('a) TYPE('x) TYPE('y)" - shows "(pi\a)\*(pi\x) = a\*x" - and "(pi\b)\*(pi\x) = b\*x" -apply(unfold fresh_star_def) -apply(auto) -apply(drule_tac x="pi\xa" in bspec) -apply(rule pt_set_bij2[OF ptb, OF at]) -apply(assumption) -apply(simp add: fresh_star_def pt_fresh_bij_ineq[OF pta, OF ptb, OF at, OF cp]) -apply(drule_tac x="(rev pi)\xa" in bspec) -apply(simp add: pt_set_bij1[OF ptb, OF at]) -apply(simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp]) -apply(drule_tac x="pi\xa" in bspec) -apply(simp add: pt_set_bij1[OF ptb, OF at]) -apply(simp add: pt_set_eqvt [OF ptb at] pt_rev_pi[OF pt_list_inst[OF ptb], OF at]) -apply(simp add: pt_fresh_bij_ineq[OF pta, OF ptb, OF at, OF cp]) -apply(drule_tac x="(rev pi)\xa" in bspec) -apply(simp add: pt_set_bij1[OF ptb, OF at] pt_set_eqvt [OF ptb at]) -apply(simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp]) -done - lemma pt_fresh_left: fixes pi :: "'x prm" and x :: "'a" @@ -1651,56 +1592,6 @@ apply(rule at) done -lemma pt_fresh_star_bij: - fixes pi :: "'x prm" - and x :: "'a" - and a :: "'x set" - and b :: "'x list" - assumes pt: "pt TYPE('a) TYPE('x)" - and at: "at TYPE('x)" - shows "(pi\a)\*(pi\x) = a\*x" - and "(pi\b)\*(pi\x) = b\*x" -apply(rule pt_fresh_star_bij_ineq(1)) -apply(rule pt) -apply(rule at_pt_inst) -apply(rule at)+ -apply(rule cp_pt_inst) -apply(rule pt) -apply(rule at) -apply(rule pt_fresh_star_bij_ineq(2)) -apply(rule pt) -apply(rule at_pt_inst) -apply(rule at)+ -apply(rule cp_pt_inst) -apply(rule pt) -apply(rule at) -done - -lemma pt_fresh_star_eqvt: - fixes pi :: "'x prm" - and x :: "'a" - and a :: "'x set" - and b :: "'x list" - assumes pt: "pt TYPE('a) TYPE('x)" - and at: "at TYPE('x)" - shows "pi\(a\*x) = (pi\a)\*(pi\x)" - and "pi\(b\*x) = (pi\b)\*(pi\x)" - by (simp_all add: perm_bool pt_fresh_star_bij[OF pt, OF at]) - -lemma pt_fresh_star_eqvt_ineq: - fixes pi::"'x prm" - and a::"'y set" - and b::"'y list" - and x::"'a" - assumes pta: "pt TYPE('a) TYPE('x)" - and ptb: "pt TYPE('y) TYPE('x)" - and at: "at TYPE('x)" - and cp: "cp TYPE('a) TYPE('x) TYPE('y)" - and dj: "disjoint TYPE('y) TYPE('x)" - shows "pi\(a\*x) = (pi\a)\*(pi\x)" - and "pi\(b\*x) = (pi\b)\*(pi\x)" - by (simp_all add: pt_fresh_star_bij_ineq[OF pta, OF ptb, OF at, OF cp] dj_perm_forget[OF dj] perm_bool) - lemma pt_fresh_bij1: fixes pi :: "'x prm" and x :: "'a" @@ -1753,7 +1644,6 @@ (* the next two lemmas are needed in the proof *) (* of the structural induction principle *) - lemma pt_fresh_aux: fixes a::"'x" and b::"'x" @@ -1857,27 +1747,6 @@ thus ?thesis using eq3 by simp qed -lemma pt_freshs_freshs: - assumes pt: "pt TYPE('a) TYPE('x)" - and at: "at TYPE ('x)" - and pi: "set (pi::'x prm) \ Xs \ Ys" - and Xs: "Xs \* (x::'a)" - and Ys: "Ys \* x" - shows "pi \ x = x" - using pi -proof (induct pi) - case Nil - show ?case by (simp add: pt1 [OF pt]) -next - case (Cons p pi) - obtain a b where p: "p = (a, b)" by (cases p) - with Cons Xs Ys have "a \ x" "b \ x" - by (simp_all add: fresh_star_def) - with Cons p show ?case - by (simp add: pt_fresh_fresh [OF pt at] - pt2 [OF pt, of "[(a, b)]" pi, simplified]) -qed - lemma pt_pi_fresh_fresh: fixes x :: "'a" and pi :: "'x prm" @@ -1943,8 +1812,7 @@ thus ?thesis by (simp add: pt2[OF pt]) qed -section {* equivaraince for some connectives *} - +section {* equivariance for some connectives *} lemma pt_all_eqvt: fixes pi :: "'x prm" and x :: "'a" @@ -1990,8 +1858,6 @@ apply(rule theI'[OF unique]) done - - section {* facts about supports *} (*==============================*) @@ -2160,6 +2026,7 @@ shows "(x \ X) = (x \ X)" by (simp add: at_fin_set_supp fresh_def at fs) + section {* Permutations acting on Functions *} (*==========================================*) @@ -2540,9 +2407,8 @@ and a1: "a\x" and a2: "a\X" shows "a\(insert x X)" -using a1 a2 -apply(simp add: fresh_fin_insert[OF pt, OF at, OF fs, OF f]) -done + using a1 a2 + by (simp add: fresh_fin_insert[OF pt, OF at, OF fs, OF f]) lemma pt_list_set_supp: fixes xs :: "'a list" @@ -2571,14 +2437,191 @@ shows "a\(set xs) = a\xs" by (simp add: fresh_def pt_list_set_supp[OF pt, OF at, OF fs]) + +section {* generalisation of freshness to lists and sets of atoms *} +(*================================================================*) + +consts + fresh_star :: "'b \ 'a \ bool" ("_ \* _" [100,100] 100) + +defs (overloaded) + fresh_star_set: "xs\*c \ \x\xs. x\c" + +defs (overloaded) + fresh_star_list: "xs\*c \ \x\set xs. x\c" + +lemmas fresh_star_def = fresh_star_list fresh_star_set + +lemma fresh_star_prod_set: + fixes xs::"'a set" + shows "xs\*(a,b) = (xs\*a \ xs\*b)" +by (auto simp add: fresh_star_def fresh_prod) + +lemma fresh_star_prod_list: + fixes xs::"'a list" + shows "xs\*(a,b) = (xs\*a \ xs\*b)" + by (auto simp add: fresh_star_def fresh_prod) + +lemmas fresh_star_prod = fresh_star_prod_list fresh_star_prod_set + +lemma fresh_star_set_eq: "set xs \* c = xs \* c" + by (simp add: fresh_star_def) + +lemma fresh_star_Un_elim: + "((S \ T) \* c \ PROP C) \ (S \* c \ T \* c \ PROP C)" + apply rule + apply (simp_all add: fresh_star_def) + apply (erule meta_mp) + apply blast + done + +lemma fresh_star_insert_elim: + "(insert x S \* c \ PROP C) \ (x \ c \ S \* c \ PROP C)" + by rule (simp_all add: fresh_star_def) + +lemma fresh_star_empty_elim: + "({} \* c \ PROP C) \ PROP C" + by (simp add: fresh_star_def) + +text {* Normalization of freshness results; see \ @{text nominal_induct} *} + +lemma fresh_star_unit_elim: + shows "((a::'a set)\*() \ PROP C) \ PROP C" + and "((b::'a list)\*() \ PROP C) \ PROP C" + by (simp_all add: fresh_star_def fresh_def supp_unit) + +lemma fresh_star_prod_elim: + shows "((a::'a set)\*(x,y) \ PROP C) \ (a\*x \ a\*y \ PROP C)" + and "((b::'a list)\*(x,y) \ PROP C) \ (b\*x \ b\*y \ PROP C)" + by (rule, simp_all add: fresh_star_prod)+ + + +lemma pt_fresh_star_bij_ineq: + fixes pi :: "'x prm" + and x :: "'a" + and a :: "'y set" + and b :: "'y list" + assumes pta: "pt TYPE('a) TYPE('x)" + and ptb: "pt TYPE('y) TYPE('x)" + and at: "at TYPE('x)" + and cp: "cp TYPE('a) TYPE('x) TYPE('y)" + shows "(pi\a)\*(pi\x) = a\*x" + and "(pi\b)\*(pi\x) = b\*x" +apply(unfold fresh_star_def) +apply(auto) +apply(drule_tac x="pi\xa" in bspec) +apply(erule pt_set_bij2[OF ptb, OF at]) +apply(simp add: fresh_star_def pt_fresh_bij_ineq[OF pta, OF ptb, OF at, OF cp]) +apply(drule_tac x="(rev pi)\xa" in bspec) +apply(simp add: pt_set_bij1[OF ptb, OF at]) +apply(simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp]) +apply(drule_tac x="pi\xa" in bspec) +apply(simp add: pt_set_bij1[OF ptb, OF at]) +apply(simp add: pt_set_eqvt [OF ptb at] pt_rev_pi[OF pt_list_inst[OF ptb], OF at]) +apply(simp add: pt_fresh_bij_ineq[OF pta, OF ptb, OF at, OF cp]) +apply(drule_tac x="(rev pi)\xa" in bspec) +apply(simp add: pt_set_bij1[OF ptb, OF at] pt_set_eqvt [OF ptb at]) +apply(simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp]) +done + +lemma pt_fresh_star_bij: + fixes pi :: "'x prm" + and x :: "'a" + and a :: "'x set" + and b :: "'x list" + assumes pt: "pt TYPE('a) TYPE('x)" + and at: "at TYPE('x)" + shows "(pi\a)\*(pi\x) = a\*x" + and "(pi\b)\*(pi\x) = b\*x" +apply(rule pt_fresh_star_bij_ineq(1)) +apply(rule pt) +apply(rule at_pt_inst) +apply(rule at)+ +apply(rule cp_pt_inst) +apply(rule pt) +apply(rule at) +apply(rule pt_fresh_star_bij_ineq(2)) +apply(rule pt) +apply(rule at_pt_inst) +apply(rule at)+ +apply(rule cp_pt_inst) +apply(rule pt) +apply(rule at) +done + +lemma pt_fresh_star_eqvt: + fixes pi :: "'x prm" + and x :: "'a" + and a :: "'x set" + and b :: "'x list" + assumes pt: "pt TYPE('a) TYPE('x)" + and at: "at TYPE('x)" + shows "pi\(a\*x) = (pi\a)\*(pi\x)" + and "pi\(b\*x) = (pi\b)\*(pi\x)" + by (simp_all add: perm_bool pt_fresh_star_bij[OF pt, OF at]) + +lemma pt_fresh_star_eqvt_ineq: + fixes pi::"'x prm" + and a::"'y set" + and b::"'y list" + and x::"'a" + assumes pta: "pt TYPE('a) TYPE('x)" + and ptb: "pt TYPE('y) TYPE('x)" + and at: "at TYPE('x)" + and cp: "cp TYPE('a) TYPE('x) TYPE('y)" + and dj: "disjoint TYPE('y) TYPE('x)" + shows "pi\(a\*x) = (pi\a)\*(pi\x)" + and "pi\(b\*x) = (pi\b)\*(pi\x)" + by (simp_all add: pt_fresh_star_bij_ineq[OF pta, OF ptb, OF at, OF cp] dj_perm_forget[OF dj] perm_bool) + +lemma pt_freshs_freshs: + assumes pt: "pt TYPE('a) TYPE('x)" + and at: "at TYPE ('x)" + and pi: "set (pi::'x prm) \ Xs \ Ys" + and Xs: "Xs \* (x::'a)" + and Ys: "Ys \* x" + shows "pi\x = x" + using pi +proof (induct pi) + case Nil + show ?case by (simp add: pt1 [OF pt]) +next + case (Cons p pi) + obtain a b where p: "p = (a, b)" by (cases p) + with Cons Xs Ys have "a \ x" "b \ x" + by (simp_all add: fresh_star_def) + with Cons p show ?case + by (simp add: pt_fresh_fresh [OF pt at] + pt2 [OF pt, of "[(a, b)]" pi, simplified]) +qed + +lemma pt_fresh_star_pi: + fixes x::"'a" + and pi::"'x prm" + assumes pt: "pt TYPE('a) TYPE('x)" + and at: "at TYPE('x)" + and a: "((supp x)::'x set)\* pi" + shows "pi\x = x" +using a +apply(induct pi) +apply(auto simp add: fresh_star_def fresh_list_cons fresh_prod pt1[OF pt]) +apply(subgoal_tac "((a,b)#pi)\x = ([(a,b)]@pi)\x") +apply(simp only: pt2[OF pt]) +apply(rule pt_fresh_fresh[OF pt at]) +apply(simp add: fresh_def at_supp[OF at]) +apply(blast) +apply(simp add: fresh_def at_supp[OF at]) +apply(blast) +apply(simp add: pt2[OF pt]) +done + section {* Infrastructure lemmas for strong rule inductions *} (*==========================================================*) - text {* For every set of atoms, there is another set of atoms avoiding a finitely supported c and there is a permutation - which make 'translates' between both sets. + which 'translates' between both sets. *} lemma at_set_avoiding_aux: fixes Xs::"'a set" @@ -3365,7 +3408,6 @@ syntax ABS :: "type \ type \ type" ("\_\_" [1000,1000] 1000) - section {* lemmas for deciding permutation equations *} (*===================================================*) @@ -3526,8 +3568,8 @@ shows "pi\(x div y) = (pi\x) div (pi\y)" by (simp add:perm_int_def) -(*******************************************************************) -(* Setup of the theorem attributes eqvt, eqvt_force, fresh and bij *) +(*******************************************************) +(* Setup of the theorem attributes eqvt and eqvt_force *) use "nominal_thmdecls.ML" setup "NominalThmDecls.setup" diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Nominal/nominal_thmdecls.ML --- a/src/HOL/Nominal/nominal_thmdecls.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Nominal/nominal_thmdecls.ML Mon Apr 27 07:26:17 2009 -0700 @@ -1,12 +1,12 @@ (* Authors: Julien Narboux and Christian Urban This file introduces the infrastructure for the lemma - declaration "eqvts" "bijs" and "freshs". + collection "eqvts". - By attaching [eqvt] [bij] or [fresh] to a lemma, the lemma gets stored - in a data-slot in the context. Possible modifiers - are [attribute add] and [attribute del] for adding and deleting, - respectively the lemma from the data-slot. + By attaching [eqvt] or [eqvt_force] to a lemma, it will get + stored in a data-slot in the context. Possible modifiers + are [... add] and [... del] for adding and deleting, + respectively, the lemma from the data-slot. *) signature NOMINAL_THMDECLS = @@ -17,9 +17,6 @@ val eqvt_force_del: attribute val setup: theory -> theory val get_eqvt_thms: Proof.context -> thm list - val get_fresh_thms: Proof.context -> thm list - val get_bij_thms: Proof.context -> thm list - val NOMINAL_EQVT_DEBUG : bool ref end; @@ -29,13 +26,11 @@ structure Data = GenericDataFun ( - type T = {eqvts:thm list, freshs:thm list, bijs:thm list}; - val empty = ({eqvts=[], freshs=[], bijs=[]}:T); - val extend = I; - fun merge _ (r1:T,r2:T) = {eqvts = Thm.merge_thms (#eqvts r1, #eqvts r2), - freshs = Thm.merge_thms (#freshs r1, #freshs r2), - bijs = Thm.merge_thms (#bijs r1, #bijs r2)} -); + type T = thm list + val empty = []:T + val extend = I + fun merge _ (r1:T, r2:T) = Thm.merge_thms (r1, r2) +) (* Exception for when a theorem does not conform with form of an equivariance lemma. *) (* There are two forms: one is an implication (for relations) and the other is an *) @@ -46,72 +41,68 @@ (* the implicational case it is also checked that the variables and permutation fit *) (* together, i.e. are of the right "pt_class", so that a stronger version of the *) (* equality-lemma can be derived. *) -exception EQVT_FORM of string; +exception EQVT_FORM of string -val get_eqvt_thms = Context.Proof #> Data.get #> #eqvts; -val get_fresh_thms = Context.Proof #> Data.get #> #freshs; -val get_bij_thms = Context.Proof #> Data.get #> #bijs; +val NOMINAL_EQVT_DEBUG = ref false -(* FIXME: should be a function in a library *) -fun mk_permT T = HOLogic.listT (HOLogic.mk_prodT (T, T)); - -val NOMINAL_EQVT_DEBUG = ref false; - -fun tactic (msg,tac) = - if !NOMINAL_EQVT_DEBUG - then tac THEN print_tac ("after "^msg) - else tac +fun tactic (msg, tac) = + if !NOMINAL_EQVT_DEBUG + then tac THEN' (K (print_tac ("after " ^ msg))) + else tac -fun tactic_eqvt ctx orig_thm pi pi' = - let - val mypi = Thm.cterm_of ctx pi - val T = fastype_of pi' - val mypifree = Thm.cterm_of ctx (Const ("List.rev", T --> T) $ pi') - val perm_pi_simp = PureThy.get_thms ctx "perm_pi_simp" - in - EVERY [tactic ("iffI applied",rtac iffI 1), - tactic ("remove pi with perm_boolE", (dtac @{thm perm_boolE} 1)), - tactic ("solve with orig_thm", (etac orig_thm 1)), - tactic ("applies orig_thm instantiated with rev pi", - dtac (Drule.cterm_instantiate [(mypi,mypifree)] orig_thm) 1), - tactic ("getting rid of the pi on the right", - (rtac @{thm perm_boolI} 1)), - tactic ("getting rid of all remaining perms", - full_simp_tac (HOL_basic_ss addsimps perm_pi_simp) 1)] - end; +fun prove_eqvt_tac ctxt orig_thm pi pi' = +let + val mypi = Thm.cterm_of ctxt pi + val T = fastype_of pi' + val mypifree = Thm.cterm_of ctxt (Const (@{const_name "rev"}, T --> T) $ pi') + val perm_pi_simp = PureThy.get_thms ctxt "perm_pi_simp" +in + EVERY1 [tactic ("iffI applied", rtac @{thm iffI}), + tactic ("remove pi with perm_boolE", dtac @{thm perm_boolE}), + tactic ("solve with orig_thm", etac orig_thm), + tactic ("applies orig_thm instantiated with rev pi", + dtac (Drule.cterm_instantiate [(mypi,mypifree)] orig_thm)), + tactic ("getting rid of the pi on the right", rtac @{thm perm_boolI}), + tactic ("getting rid of all remaining perms", + full_simp_tac (HOL_basic_ss addsimps perm_pi_simp))] +end; fun get_derived_thm ctxt hyp concl orig_thm pi typi = let val thy = ProofContext.theory_of ctxt; val pi' = Var (pi, typi); - val lhs = Const ("Nominal.perm", typi --> HOLogic.boolT --> HOLogic.boolT) $ pi' $ hyp; + val lhs = Const (@{const_name "perm"}, typi --> HOLogic.boolT --> HOLogic.boolT) $ pi' $ hyp; val ([goal_term, pi''], ctxt') = Variable.import_terms false [HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, concl)), pi'] ctxt val _ = Display.print_cterm (cterm_of thy goal_term) in Goal.prove ctxt' [] [] goal_term - (fn _ => tactic_eqvt thy orig_thm pi' pi'') |> + (fn _ => prove_eqvt_tac thy orig_thm pi' pi'') |> singleton (ProofContext.export ctxt' ctxt) end -(* replaces every variable x in t with pi o x *) -fun apply_pi trm (pi,typi) = - let - fun only_vars t = - (case t of - Var (n,ty) => (Const ("Nominal.perm",typi --> ty --> ty) $ (Var (pi,typi)) $ (Var (n,ty))) - | _ => t) +(* replaces in t every variable, say x, with pi o x *) +fun apply_pi trm (pi, typi) = +let + fun replace n ty = + let + val c = Const (@{const_name "perm"}, typi --> ty --> ty) + val v1 = Var (pi, typi) + val v2 = Var (n, ty) in - map_aterms only_vars trm - end; + c $ v1 $ v2 + end +in + map_aterms (fn Var (n, ty) => replace n ty | t => t) trm +end (* returns *the* pi which is in front of all variables, provided there *) (* exists such a pi; otherwise raises EQVT_FORM *) fun get_pi t thy = let fun get_pi_aux s = (case s of - (Const ("Nominal.perm",typrm) $ - (Var (pi,typi as Type("List.list",[Type ("*",[Type (tyatm,[]),_])]))) $ + (Const (@{const_name "perm"} ,typrm) $ + (Var (pi,typi as Type(@{type_name "list"}, [Type ("*", [Type (tyatm,[]),_])]))) $ (Var (n,ty))) => let (* FIXME: this should be an operation the library *) @@ -130,7 +121,7 @@ (* to ensure that all pi's must have been the same, i.e. distinct returns *) (* a singleton-list *) (case (distinct (op =) (get_pi_aux t)) of - [(pi,typi)] => (pi,typi) + [(pi,typi)] => (pi, typi) | _ => raise EQVT_FORM "All permutation should be the same") end; @@ -155,8 +146,8 @@ else raise EQVT_FORM "Type Implication" end (* case: eqvt-lemma is of the equational form *) - | (Const ("Trueprop", _) $ (Const ("op =", _) $ - (Const ("Nominal.perm",typrm) $ Var (pi,typi) $ lhs) $ rhs)) => + | (Const (@{const_name "Trueprop"}, _) $ (Const (@{const_name "op ="}, _) $ + (Const (@{const_name "perm"},typrm) $ Var (pi,typi) $ lhs) $ rhs)) => (if (apply_pi lhs (pi,typi)) = rhs then [orig_thm] else raise EQVT_FORM "Type Equality") @@ -165,38 +156,24 @@ fold (fn thm => Data.map (flag thm)) thms_to_be_added context end handle EQVT_FORM s => - error (Display.string_of_thm orig_thm ^ " does not comply with the form of an equivariance lemma ("^s^").") - -(* in cases of bij- and freshness, we just add the lemmas to the *) -(* data-slot *) - -fun eqvt_map f (r:Data.T) = {eqvts = f (#eqvts r), freshs = #freshs r, bijs = #bijs r}; -fun fresh_map f (r:Data.T) = {eqvts = #eqvts r, freshs = f (#freshs r), bijs = #bijs r}; -fun bij_map f (r:Data.T) = {eqvts = #eqvts r, freshs = #freshs r, bijs = f (#bijs r)}; - -val eqvt_add = Thm.declaration_attribute (eqvt_add_del_aux (eqvt_map o Thm.add_thm)); -val eqvt_del = Thm.declaration_attribute (eqvt_add_del_aux (eqvt_map o Thm.del_thm)); - -val eqvt_force_add = Thm.declaration_attribute (Data.map o eqvt_map o Thm.add_thm); -val eqvt_force_del = Thm.declaration_attribute (Data.map o eqvt_map o Thm.del_thm); -val bij_add = Thm.declaration_attribute (Data.map o bij_map o Thm.add_thm); -val bij_del = Thm.declaration_attribute (Data.map o bij_map o Thm.del_thm); -val fresh_add = Thm.declaration_attribute (Data.map o fresh_map o Thm.add_thm); -val fresh_del = Thm.declaration_attribute (Data.map o fresh_map o Thm.del_thm); + error (Display.string_of_thm orig_thm ^ + " does not comply with the form of an equivariance lemma (" ^ s ^").") +val eqvt_add = Thm.declaration_attribute (eqvt_add_del_aux (Thm.add_thm)); +val eqvt_del = Thm.declaration_attribute (eqvt_add_del_aux (Thm.del_thm)); + +val eqvt_force_add = Thm.declaration_attribute (Data.map o Thm.add_thm); +val eqvt_force_del = Thm.declaration_attribute (Data.map o Thm.del_thm); + +val get_eqvt_thms = Context.Proof #> Data.get; val setup = - Attrib.setup @{binding eqvt} (Attrib.add_del eqvt_add eqvt_del) - "equivariance theorem declaration" #> - Attrib.setup @{binding eqvt_force} (Attrib.add_del eqvt_force_add eqvt_force_del) - "equivariance theorem declaration (without checking the form of the lemma)" #> - Attrib.setup @{binding fresh} (Attrib.add_del fresh_add fresh_del) - "freshness theorem declaration" #> - Attrib.setup @{binding "bij"} (Attrib.add_del bij_add bij_del) - "bijection theorem declaration" #> - PureThy.add_thms_dynamic (Binding.name "eqvts", #eqvts o Data.get) #> - PureThy.add_thms_dynamic (Binding.name "freshs", #freshs o Data.get) #> - PureThy.add_thms_dynamic (Binding.name "bijs", #bijs o Data.get); + Attrib.setup @{binding eqvt} (Attrib.add_del eqvt_add eqvt_del) + "equivariance theorem declaration" + #> Attrib.setup @{binding eqvt_force} (Attrib.add_del eqvt_force_add eqvt_force_del) + "equivariance theorem declaration (without checking the form of the lemma)" + #> PureThy.add_thms_dynamic (Binding.name "eqvts", Data.get) + end; diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Orderings.thy --- a/src/HOL/Orderings.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Orderings.thy Mon Apr 27 07:26:17 2009 -0700 @@ -5,7 +5,7 @@ header {* Abstract orderings *} theory Orderings -imports Code_Setup +imports HOL uses "~~/src/Provers/order.ML" begin diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Power.thy --- a/src/HOL/Power.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Power.thy Mon Apr 27 07:26:17 2009 -0700 @@ -1,95 +1,179 @@ (* Title: HOL/Power.thy - ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1997 University of Cambridge - *) -header{*Exponentiation*} +header {* Exponentiation *} theory Power imports Nat begin -class power = - fixes power :: "'a \ nat \ 'a" (infixr "^" 80) +subsection {* Powers for Arbitrary Monoids *} + +class power = one + times +begin -subsection{*Powers for Arbitrary Monoids*} +primrec power :: "'a \ nat \ 'a" (infixr "^" 80) where + power_0: "a ^ 0 = 1" + | power_Suc: "a ^ Suc n = a * a ^ n" + +notation (latex output) + power ("(_\<^bsup>_\<^esup>)" [1000] 1000) -class recpower = monoid_mult + power + - assumes power_0 [simp]: "a ^ 0 = 1" - assumes power_Suc [simp]: "a ^ Suc n = a * (a ^ n)" +notation (HTML output) + power ("(_\<^bsup>_\<^esup>)" [1000] 1000) + +end + +context monoid_mult +begin -lemma power_0_Suc [simp]: "(0::'a::{recpower,semiring_0}) ^ (Suc n) = 0" +subclass power .. + +lemma power_one [simp]: + "1 ^ n = 1" + by (induct n) simp_all + +lemma power_one_right [simp]: + "a ^ 1 = a" by simp -text{*It looks plausible as a simprule, but its effect can be strange.*} -lemma power_0_left: "0^n = (if n=0 then 1 else (0::'a::{recpower,semiring_0}))" - by (induct n) simp_all - -lemma power_one [simp]: "1^n = (1::'a::recpower)" - by (induct n) simp_all - -lemma power_one_right [simp]: "(a::'a::recpower) ^ 1 = a" - unfolding One_nat_def by simp - -lemma power_commutes: "(a::'a::recpower) ^ n * a = a * a ^ n" +lemma power_commutes: + "a ^ n * a = a * a ^ n" by (induct n) (simp_all add: mult_assoc) -lemma power_Suc2: "(a::'a::recpower) ^ Suc n = a ^ n * a" +lemma power_Suc2: + "a ^ Suc n = a ^ n * a" by (simp add: power_commutes) -lemma power_add: "(a::'a::recpower) ^ (m+n) = (a^m) * (a^n)" - by (induct m) (simp_all add: mult_ac) +lemma power_add: + "a ^ (m + n) = a ^ m * a ^ n" + by (induct m) (simp_all add: algebra_simps) -lemma power_mult: "(a::'a::recpower) ^ (m*n) = (a^m) ^ n" +lemma power_mult: + "a ^ (m * n) = (a ^ m) ^ n" by (induct n) (simp_all add: power_add) -lemma power_mult_distrib: "((a::'a::{recpower,comm_monoid_mult}) * b) ^ n = (a^n) * (b^n)" +end + +context comm_monoid_mult +begin + +lemma power_mult_distrib: + "(a * b) ^ n = (a ^ n) * (b ^ n)" by (induct n) (simp_all add: mult_ac) -lemma zero_less_power[simp]: - "0 < (a::'a::{ordered_semidom,recpower}) ==> 0 < a^n" -by (induct n) (simp_all add: mult_pos_pos) +end + +context semiring_1 +begin + +lemma of_nat_power: + "of_nat (m ^ n) = of_nat m ^ n" + by (induct n) (simp_all add: of_nat_mult) + +end + +context comm_semiring_1 +begin + +text {* The divides relation *} + +lemma le_imp_power_dvd: + assumes "m \ n" shows "a ^ m dvd a ^ n" +proof + have "a ^ n = a ^ (m + (n - m))" + using `m \ n` by simp + also have "\ = a ^ m * a ^ (n - m)" + by (rule power_add) + finally show "a ^ n = a ^ m * a ^ (n - m)" . +qed + +lemma power_le_dvd: + "a ^ n dvd b \ m \ n \ a ^ m dvd b" + by (rule dvd_trans [OF le_imp_power_dvd]) + +lemma dvd_power_same: + "x dvd y \ x ^ n dvd y ^ n" + by (induct n) (auto simp add: mult_dvd_mono) + +lemma dvd_power_le: + "x dvd y \ m \ n \ x ^ n dvd y ^ m" + by (rule power_le_dvd [OF dvd_power_same]) -lemma zero_le_power[simp]: - "0 \ (a::'a::{ordered_semidom,recpower}) ==> 0 \ a^n" -by (induct n) (simp_all add: mult_nonneg_nonneg) +lemma dvd_power [simp]: + assumes "n > (0::nat) \ x = 1" + shows "x dvd (x ^ n)" +using assms proof + assume "0 < n" + then have "x ^ n = x ^ Suc (n - 1)" by simp + then show "x dvd (x ^ n)" by simp +next + assume "x = 1" + then show "x dvd (x ^ n)" by simp +qed + +end + +context ring_1 +begin + +lemma power_minus: + "(- a) ^ n = (- 1) ^ n * a ^ n" +proof (induct n) + case 0 show ?case by simp +next + case (Suc n) then show ?case + by (simp del: power_Suc add: power_Suc2 mult_assoc) +qed + +end + +context ordered_semidom +begin + +lemma zero_less_power [simp]: + "0 < a \ 0 < a ^ n" + by (induct n) (simp_all add: mult_pos_pos) + +lemma zero_le_power [simp]: + "0 \ a \ 0 \ a ^ n" + by (induct n) (simp_all add: mult_nonneg_nonneg) lemma one_le_power[simp]: - "1 \ (a::'a::{ordered_semidom,recpower}) ==> 1 \ a^n" -apply (induct "n") -apply simp_all -apply (rule order_trans [OF _ mult_mono [of 1 _ 1]]) -apply (simp_all add: order_trans [OF zero_le_one]) -done - -lemma gt1_imp_ge0: "1 < a ==> 0 \ (a::'a::ordered_semidom)" - by (simp add: order_trans [OF zero_le_one order_less_imp_le]) + "1 \ a \ 1 \ a ^ n" + apply (induct n) + apply simp_all + apply (rule order_trans [OF _ mult_mono [of 1 _ 1]]) + apply (simp_all add: order_trans [OF zero_le_one]) + done lemma power_gt1_lemma: - assumes gt1: "1 < (a::'a::{ordered_semidom,recpower})" - shows "1 < a * a^n" + assumes gt1: "1 < a" + shows "1 < a * a ^ n" proof - - have "1*1 < a*1" using gt1 by simp - also have "\ \ a * a^n" using gt1 - by (simp only: mult_mono gt1_imp_ge0 one_le_power order_less_imp_le + from gt1 have "0 \ a" + by (fact order_trans [OF zero_le_one less_imp_le]) + have "1 * 1 < a * 1" using gt1 by simp + also have "\ \ a * a ^ n" using gt1 + by (simp only: mult_mono `0 \ a` one_le_power order_less_imp_le zero_le_one order_refl) finally show ?thesis by simp qed -lemma one_less_power[simp]: - "\1 < (a::'a::{ordered_semidom,recpower}); 0 < n\ \ 1 < a ^ n" -by (cases n, simp_all add: power_gt1_lemma) +lemma power_gt1: + "1 < a \ 1 < a ^ Suc n" + by (simp add: power_gt1_lemma) -lemma power_gt1: - "1 < (a::'a::{ordered_semidom,recpower}) ==> 1 < a ^ (Suc n)" -by (simp add: power_gt1_lemma) +lemma one_less_power [simp]: + "1 < a \ 0 < n \ 1 < a ^ n" + by (cases n) (simp_all add: power_gt1_lemma) lemma power_le_imp_le_exp: - assumes gt1: "(1::'a::{recpower,ordered_semidom}) < a" - shows "!!n. a^m \ a^n ==> m \ n" -proof (induct m) + assumes gt1: "1 < a" + shows "a ^ m \ a ^ n \ m \ n" +proof (induct m arbitrary: n) case 0 show ?case by simp next @@ -97,212 +181,128 @@ show ?case proof (cases n) case 0 - from prems have "a * a^m \ 1" by simp + with Suc.prems Suc.hyps have "a * a ^ m \ 1" by simp with gt1 show ?thesis by (force simp only: power_gt1_lemma - linorder_not_less [symmetric]) + not_less [symmetric]) next case (Suc n) - from prems show ?thesis + with Suc.prems Suc.hyps show ?thesis by (force dest: mult_left_le_imp_le - simp add: order_less_trans [OF zero_less_one gt1]) + simp add: less_trans [OF zero_less_one gt1]) qed qed text{*Surely we can strengthen this? It holds for @{text "0 (a^m = a^n) = (m=n)" + "1 < a \ a ^ m = a ^ n \ m = n" by (force simp add: order_antisym power_le_imp_le_exp) text{*Can relax the first premise to @{term "0 m < n" -by (simp add: order_less_le [of m n] order_less_le [of "a^m" "a^n"] - power_le_imp_le_exp) - + "1 < a \ a ^ m < a ^ n \ m < n" + by (simp add: order_less_le [of m n] less_le [of "a^m" "a^n"] + power_le_imp_le_exp) lemma power_mono: - "[|a \ b; (0::'a::{recpower,ordered_semidom}) \ a|] ==> a^n \ b^n" -apply (induct "n") -apply simp_all -apply (auto intro: mult_mono order_trans [of 0 a b]) -done + "a \ b \ 0 \ a \ a ^ n \ b ^ n" + by (induct n) + (auto intro: mult_mono order_trans [of 0 a b]) lemma power_strict_mono [rule_format]: - "[|a < b; (0::'a::{recpower,ordered_semidom}) \ a|] - ==> 0 < n --> a^n < b^n" -apply (induct "n") -apply (auto simp add: mult_strict_mono order_le_less_trans [of 0 a b]) -done - -lemma power_eq_0_iff [simp]: - "(a^n = 0) \ - (a = (0::'a::{mult_zero,zero_neq_one,no_zero_divisors,recpower}) & n\0)" -apply (induct "n") -apply (auto simp add: no_zero_divisors) -done - - -lemma field_power_not_zero: - "a \ (0::'a::{ring_1_no_zero_divisors,recpower}) ==> a^n \ 0" -by force - -lemma nonzero_power_inverse: - fixes a :: "'a::{division_ring,recpower}" - shows "a \ 0 ==> inverse (a ^ n) = (inverse a) ^ n" -apply (induct "n") -apply (auto simp add: nonzero_inverse_mult_distrib power_commutes) -done (* TODO: reorient or rename to nonzero_inverse_power *) - -text{*Perhaps these should be simprules.*} -lemma power_inverse: - fixes a :: "'a::{division_ring,division_by_zero,recpower}" - shows "inverse (a ^ n) = (inverse a) ^ n" -apply (cases "a = 0") -apply (simp add: power_0_left) -apply (simp add: nonzero_power_inverse) -done (* TODO: reorient or rename to inverse_power *) - -lemma power_one_over: "1 / (a::'a::{field,division_by_zero,recpower})^n = - (1 / a)^n" -apply (simp add: divide_inverse) -apply (rule power_inverse) -done - -lemma nonzero_power_divide: - "b \ 0 ==> (a/b) ^ n = ((a::'a::{field,recpower}) ^ n) / (b ^ n)" -by (simp add: divide_inverse power_mult_distrib nonzero_power_inverse) - -lemma power_divide: - "(a/b) ^ n = ((a::'a::{field,division_by_zero,recpower}) ^ n / b ^ n)" -apply (case_tac "b=0", simp add: power_0_left) -apply (rule nonzero_power_divide) -apply assumption -done - -lemma power_abs: "abs(a ^ n) = abs(a::'a::{ordered_idom,recpower}) ^ n" -apply (induct "n") -apply (auto simp add: abs_mult) -done - -lemma abs_power_minus [simp]: - fixes a:: "'a::{ordered_idom,recpower}" shows "abs((-a) ^ n) = abs(a ^ n)" - by (simp add: abs_minus_cancel power_abs) - -lemma zero_less_power_abs_iff [simp,noatp]: - "(0 < (abs a)^n) = (a \ (0::'a::{ordered_idom,recpower}) | n=0)" -proof (induct "n") - case 0 - show ?case by simp -next - case (Suc n) - show ?case by (auto simp add: prems zero_less_mult_iff) -qed - -lemma zero_le_power_abs [simp]: - "(0::'a::{ordered_idom,recpower}) \ (abs a)^n" -by (rule zero_le_power [OF abs_ge_zero]) - -lemma power_minus: "(-a) ^ n = (- 1)^n * (a::'a::{ring_1,recpower}) ^ n" -proof (induct n) - case 0 show ?case by simp -next - case (Suc n) then show ?case - by (simp del: power_Suc add: power_Suc2 mult_assoc) -qed + "a < b \ 0 \ a \ 0 < n \ a ^ n < b ^ n" + by (induct n) + (auto simp add: mult_strict_mono le_less_trans [of 0 a b]) text{*Lemma for @{text power_strict_decreasing}*} lemma power_Suc_less: - "[|(0::'a::{ordered_semidom,recpower}) < a; a < 1|] - ==> a * a^n < a^n" -apply (induct n) -apply (auto simp add: mult_strict_left_mono) -done + "0 < a \ a < 1 \ a * a ^ n < a ^ n" + by (induct n) + (auto simp add: mult_strict_left_mono) -lemma power_strict_decreasing: - "[|n < N; 0 < a; a < (1::'a::{ordered_semidom,recpower})|] - ==> a^N < a^n" -apply (erule rev_mp) -apply (induct "N") -apply (auto simp add: power_Suc_less less_Suc_eq) -apply (rename_tac m) -apply (subgoal_tac "a * a^m < 1 * a^n", simp) -apply (rule mult_strict_mono) -apply (auto simp add: order_less_imp_le) -done +lemma power_strict_decreasing [rule_format]: + "n < N \ 0 < a \ a < 1 \ a ^ N < a ^ n" +proof (induct N) + case 0 then show ?case by simp +next + case (Suc N) then show ?case + apply (auto simp add: power_Suc_less less_Suc_eq) + apply (subgoal_tac "a * a^N < 1 * a^n") + apply simp + apply (rule mult_strict_mono) apply auto + done +qed text{*Proof resembles that of @{text power_strict_decreasing}*} -lemma power_decreasing: - "[|n \ N; 0 \ a; a \ (1::'a::{ordered_semidom,recpower})|] - ==> a^N \ a^n" -apply (erule rev_mp) -apply (induct "N") -apply (auto simp add: le_Suc_eq) -apply (rename_tac m) -apply (subgoal_tac "a * a^m \ 1 * a^n", simp) -apply (rule mult_mono) -apply auto -done +lemma power_decreasing [rule_format]: + "n \ N \ 0 \ a \ a \ 1 \ a ^ N \ a ^ n" +proof (induct N) + case 0 then show ?case by simp +next + case (Suc N) then show ?case + apply (auto simp add: le_Suc_eq) + apply (subgoal_tac "a * a^N \ 1 * a^n", simp) + apply (rule mult_mono) apply auto + done +qed lemma power_Suc_less_one: - "[| 0 < a; a < (1::'a::{ordered_semidom,recpower}) |] ==> a ^ Suc n < 1" -apply (insert power_strict_decreasing [of 0 "Suc n" a], simp) -done + "0 < a \ a < 1 \ a ^ Suc n < 1" + using power_strict_decreasing [of 0 "Suc n" a] by simp text{*Proof again resembles that of @{text power_strict_decreasing}*} -lemma power_increasing: - "[|n \ N; (1::'a::{ordered_semidom,recpower}) \ a|] ==> a^n \ a^N" -apply (erule rev_mp) -apply (induct "N") -apply (auto simp add: le_Suc_eq) -apply (rename_tac m) -apply (subgoal_tac "1 * a^n \ a * a^m", simp) -apply (rule mult_mono) -apply (auto simp add: order_trans [OF zero_le_one]) -done +lemma power_increasing [rule_format]: + "n \ N \ 1 \ a \ a ^ n \ a ^ N" +proof (induct N) + case 0 then show ?case by simp +next + case (Suc N) then show ?case + apply (auto simp add: le_Suc_eq) + apply (subgoal_tac "1 * a^n \ a * a^N", simp) + apply (rule mult_mono) apply (auto simp add: order_trans [OF zero_le_one]) + done +qed text{*Lemma for @{text power_strict_increasing}*} lemma power_less_power_Suc: - "(1::'a::{ordered_semidom,recpower}) < a ==> a^n < a * a^n" -apply (induct n) -apply (auto simp add: mult_strict_left_mono order_less_trans [OF zero_less_one]) -done + "1 < a \ a ^ n < a * a ^ n" + by (induct n) (auto simp add: mult_strict_left_mono less_trans [OF zero_less_one]) -lemma power_strict_increasing: - "[|n < N; (1::'a::{ordered_semidom,recpower}) < a|] ==> a^n < a^N" -apply (erule rev_mp) -apply (induct "N") -apply (auto simp add: power_less_power_Suc less_Suc_eq) -apply (rename_tac m) -apply (subgoal_tac "1 * a^n < a * a^m", simp) -apply (rule mult_strict_mono) -apply (auto simp add: order_less_trans [OF zero_less_one] order_less_imp_le) -done +lemma power_strict_increasing [rule_format]: + "n < N \ 1 < a \ a ^ n < a ^ N" +proof (induct N) + case 0 then show ?case by simp +next + case (Suc N) then show ?case + apply (auto simp add: power_less_power_Suc less_Suc_eq) + apply (subgoal_tac "1 * a^n < a * a^N", simp) + apply (rule mult_strict_mono) apply (auto simp add: less_trans [OF zero_less_one] less_imp_le) + done +qed lemma power_increasing_iff [simp]: - "1 < (b::'a::{ordered_semidom,recpower}) ==> (b ^ x \ b ^ y) = (x \ y)" -by (blast intro: power_le_imp_le_exp power_increasing order_less_imp_le) + "1 < b \ b ^ x \ b ^ y \ x \ y" + by (blast intro: power_le_imp_le_exp power_increasing less_imp_le) lemma power_strict_increasing_iff [simp]: - "1 < (b::'a::{ordered_semidom,recpower}) ==> (b ^ x < b ^ y) = (x < y)" + "1 < b \ b ^ x < b ^ y \ x < y" by (blast intro: power_less_imp_less_exp power_strict_increasing) lemma power_le_imp_le_base: -assumes le: "a ^ Suc n \ b ^ Suc n" - and ynonneg: "(0::'a::{ordered_semidom,recpower}) \ b" -shows "a \ b" + assumes le: "a ^ Suc n \ b ^ Suc n" + and ynonneg: "0 \ b" + shows "a \ b" proof (rule ccontr) assume "~ a \ b" then have "b < a" by (simp only: linorder_not_le) then have "b ^ Suc n < a ^ Suc n" by (simp only: prems power_strict_mono) - from le and this show "False" + from le and this show False by (simp add: linorder_not_less [symmetric]) qed lemma power_less_imp_less_base: - fixes a b :: "'a::{ordered_semidom,recpower}" assumes less: "a ^ n < b ^ n" assumes nonneg: "0 \ b" shows "a < b" @@ -310,98 +310,144 @@ assume "~ a < b" hence "b \ a" by (simp only: linorder_not_less) hence "b ^ n \ a ^ n" using nonneg by (rule power_mono) - thus "~ a ^ n < b ^ n" by (simp only: linorder_not_less) + thus "\ a ^ n < b ^ n" by (simp only: linorder_not_less) qed lemma power_inject_base: - "[| a ^ Suc n = b ^ Suc n; 0 \ a; 0 \ b |] - ==> a = (b::'a::{ordered_semidom,recpower})" -by (blast intro: power_le_imp_le_base order_antisym order_eq_refl sym) + "a ^ Suc n = b ^ Suc n \ 0 \ a \ 0 \ b \ a = b" +by (blast intro: power_le_imp_le_base antisym eq_refl sym) lemma power_eq_imp_eq_base: - fixes a b :: "'a::{ordered_semidom,recpower}" - shows "\a ^ n = b ^ n; 0 \ a; 0 \ b; 0 < n\ \ a = b" -by (cases n, simp_all del: power_Suc, rule power_inject_base) + "a ^ n = b ^ n \ 0 \ a \ 0 \ b \ 0 < n \ a = b" + by (cases n) (simp_all del: power_Suc, rule power_inject_base) -text {* The divides relation *} +end + +context ordered_idom +begin -lemma le_imp_power_dvd: - fixes a :: "'a::{comm_semiring_1,recpower}" - assumes "m \ n" shows "a^m dvd a^n" -proof - have "a^n = a^(m + (n - m))" - using `m \ n` by simp - also have "\ = a^m * a^(n - m)" - by (rule power_add) - finally show "a^n = a^m * a^(n - m)" . +lemma power_abs: + "abs (a ^ n) = abs a ^ n" + by (induct n) (auto simp add: abs_mult) + +lemma abs_power_minus [simp]: + "abs ((-a) ^ n) = abs (a ^ n)" + by (simp add: abs_minus_cancel power_abs) + +lemma zero_less_power_abs_iff [simp, noatp]: + "0 < abs a ^ n \ a \ 0 \ n = 0" +proof (induct n) + case 0 show ?case by simp +next + case (Suc n) show ?case by (auto simp add: Suc zero_less_mult_iff) qed -lemma power_le_dvd: - fixes a b :: "'a::{comm_semiring_1,recpower}" - shows "a^n dvd b \ m \ n \ a^m dvd b" - by (rule dvd_trans [OF le_imp_power_dvd]) - - -lemma dvd_power_same: - "(x::'a::{comm_semiring_1,recpower}) dvd y \ x^n dvd y^n" -by (induct n) (auto simp add: mult_dvd_mono) - -lemma dvd_power_le: - "(x::'a::{comm_semiring_1,recpower}) dvd y \ m >= n \ x^n dvd y^m" -by(rule power_le_dvd[OF dvd_power_same]) +lemma zero_le_power_abs [simp]: + "0 \ abs a ^ n" + by (rule zero_le_power [OF abs_ge_zero]) -lemma dvd_power [simp]: - "n > 0 | (x::'a::{comm_semiring_1,recpower}) = 1 \ x dvd x^n" -apply (erule disjE) - apply (subgoal_tac "x ^ n = x^(Suc (n - 1))") - apply (erule ssubst) - apply (subst power_Suc) - apply auto -done +end - -subsection{*Exponentiation for the Natural Numbers*} - -instantiation nat :: recpower +context ring_1_no_zero_divisors begin -primrec power_nat where - "p ^ 0 = (1\nat)" - | "p ^ (Suc n) = (p\nat) * (p ^ n)" +lemma field_power_not_zero: + "a \ 0 \ a ^ n \ 0" + by (induct n) auto + +end + +context division_ring +begin -instance proof - fix z n :: nat - show "z^0 = 1" by simp - show "z^(Suc n) = z * (z^n)" by simp -qed +text {* FIXME reorient or rename to @{text nonzero_inverse_power} *} +lemma nonzero_power_inverse: + "a \ 0 \ inverse (a ^ n) = (inverse a) ^ n" + by (induct n) + (simp_all add: nonzero_inverse_mult_distrib power_commutes field_power_not_zero) -declare power_nat.simps [simp del] +end + +context field +begin + +lemma nonzero_power_divide: + "b \ 0 \ (a / b) ^ n = a ^ n / b ^ n" + by (simp add: divide_inverse power_mult_distrib nonzero_power_inverse) end -lemma of_nat_power: - "of_nat (m ^ n) = (of_nat m::'a::{semiring_1,recpower}) ^ n" -by (induct n, simp_all add: of_nat_mult) +lemma power_0_Suc [simp]: + "(0::'a::{power, semiring_0}) ^ Suc n = 0" + by simp + +text{*It looks plausible as a simprule, but its effect can be strange.*} +lemma power_0_left: + "0 ^ n = (if n = 0 then 1 else (0::'a::{power, semiring_0}))" + by (induct n) simp_all + +lemma power_eq_0_iff [simp]: + "a ^ n = 0 \ + a = (0::'a::{mult_zero,zero_neq_one,no_zero_divisors,power}) \ n \ 0" + by (induct n) + (auto simp add: no_zero_divisors elim: contrapos_pp) + +lemma power_diff: + fixes a :: "'a::field" + assumes nz: "a \ 0" + shows "n \ m \ a ^ (m - n) = a ^ m / a ^ n" + by (induct m n rule: diff_induct) (simp_all add: nz) -lemma nat_one_le_power [simp]: "Suc 0 \ i ==> Suc 0 \ i^n" -by (rule one_le_power [of i n, unfolded One_nat_def]) +text{*Perhaps these should be simprules.*} +lemma power_inverse: + fixes a :: "'a::{division_ring,division_by_zero,power}" + shows "inverse (a ^ n) = (inverse a) ^ n" +apply (cases "a = 0") +apply (simp add: power_0_left) +apply (simp add: nonzero_power_inverse) +done (* TODO: reorient or rename to inverse_power *) + +lemma power_one_over: + "1 / (a::'a::{field,division_by_zero, power}) ^ n = (1 / a) ^ n" + by (simp add: divide_inverse) (rule power_inverse) -lemma nat_zero_less_power_iff [simp]: "(x^n > 0) = (x > (0::nat) | n=0)" -by (induct "n", auto) +lemma power_divide: + "(a / b) ^ n = (a::'a::{field,division_by_zero}) ^ n / b ^ n" +apply (cases "b = 0") +apply (simp add: power_0_left) +apply (rule nonzero_power_divide) +apply assumption +done + +class recpower = monoid_mult + + +subsection {* Exponentiation for the Natural Numbers *} + +instance nat :: recpower .. + +lemma nat_one_le_power [simp]: + "Suc 0 \ i \ Suc 0 \ i ^ n" + by (rule one_le_power [of i n, unfolded One_nat_def]) + +lemma nat_zero_less_power_iff [simp]: + "x ^ n > 0 \ x > (0::nat) \ n = 0" + by (induct n) auto lemma nat_power_eq_Suc_0_iff [simp]: - "((x::nat)^m = Suc 0) = (m = 0 | x = Suc 0)" -by (induct_tac m, auto) + "x ^ m = Suc 0 \ m = 0 \ x = Suc 0" + by (induct m) auto -lemma power_Suc_0[simp]: "(Suc 0)^n = Suc 0" -by simp +lemma power_Suc_0 [simp]: + "Suc 0 ^ n = Suc 0" + by simp text{*Valid for the naturals, but what if @{text"0nat)" - assumes less: "i^m < i^n" + assumes less: "i ^ m < i ^ n" shows "m < n" proof (cases "i = 1") case True with less power_one [where 'a = nat] show ?thesis by simp @@ -410,10 +456,4 @@ from power_strict_increasing_iff [OF this] less show ?thesis .. qed -lemma power_diff: - assumes nz: "a ~= 0" - shows "n <= m ==> (a::'a::{recpower, field}) ^ (m-n) = (a^m) / (a^n)" - by (induct m n rule: diff_induct) - (simp_all add: nonzero_mult_divide_cancel_left nz) - end diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Predicate.thy --- a/src/HOL/Predicate.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Predicate.thy Mon Apr 27 07:26:17 2009 -0700 @@ -622,6 +622,51 @@ "pred_rec f P = f (eval P)" by (cases P) simp +text {* for evaluation of predicate enumerations *} + +ML {* +signature PREDICATE = +sig + datatype 'a pred = Seq of (unit -> 'a seq) + and 'a seq = Empty | Insert of 'a * 'a pred | Join of 'a pred * 'a seq + val yield: 'a pred -> ('a * 'a pred) option + val yieldn: int -> 'a pred -> 'a list * 'a pred +end; + +structure Predicate : PREDICATE = +struct + +@{code_datatype pred = Seq}; +@{code_datatype seq = Empty | Insert | Join}; + +fun yield (Seq f) = next (f ()) +and next @{code Empty} = NONE + | next (@{code Insert} (x, P)) = SOME (x, P) + | next (@{code Join} (P, xq)) = (case yield P + of NONE => next xq + | SOME (x, Q) => SOME (x, @{code Seq} (fn _ => @{code Join} (Q, xq)))) + +fun anamorph f k x = (if k = 0 then ([], x) + else case f x + of NONE => ([], x) + | SOME (v, y) => let + val (vs, z) = anamorph f (k - 1) y + in (v :: vs, z) end) + +fun yieldn P = anamorph yield P; + +end; +*} + +code_reserved Eval Predicate + +code_type pred and seq + (Eval "_/ Predicate.pred" and "_/ Predicate.seq") + +code_const Seq and Empty and Insert and Join + (Eval "Predicate.Seq" and "Predicate.Empty" and "Predicate.Insert/ (_,/ _)" and "Predicate.Join/ (_,/ _)") + + no_notation inf (infixl "\" 70) and sup (infixl "\" 65) and diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Product_Type.thy --- a/src/HOL/Product_Type.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Product_Type.thy Mon Apr 27 07:26:17 2009 -0700 @@ -84,6 +84,14 @@ lemma unit_abs_eta_conv [simp,noatp]: "(%u::unit. f ()) = f" by (rule ext) simp +instantiation unit :: default +begin + +definition "default = ()" + +instance .. + +end text {* code generator setup *} diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Rational.thy --- a/src/HOL/Rational.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Rational.thy Mon Apr 27 07:26:17 2009 -0700 @@ -156,11 +156,6 @@ then show ?thesis by (simp add: mult_rat [symmetric]) qed -primrec power_rat -where - "q ^ 0 = (1\rat)" -| "q ^ Suc n = (q\rat) * (q ^ n)" - instance proof fix q r s :: rat show "(q * r) * s = q * (r * s)" by (cases q, cases r, cases s) (simp add: eq_rat) @@ -193,15 +188,8 @@ next fix q :: rat show "q * 1 = q" by (cases q) (simp add: One_rat_def eq_rat) -next - fix q :: rat - fix n :: nat - show "q ^ 0 = 1" by simp - show "q ^ (Suc n) = q * (q ^ n)" by simp qed -declare power_rat.simps [simp del] - end lemma of_nat_rat: "of_nat k = Fract (of_nat k) 1" @@ -222,7 +210,8 @@ definition rat_number_of_def [code del]: "number_of w = Fract w 1" -instance by intro_classes (simp add: rat_number_of_def of_int_rat) +instance proof +qed (simp add: rat_number_of_def of_int_rat) end diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/RealPow.thy --- a/src/HOL/RealPow.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/RealPow.thy Mon Apr 27 07:26:17 2009 -0700 @@ -12,24 +12,7 @@ declare abs_mult_self [simp] -instantiation real :: recpower -begin - -primrec power_real where - "r ^ 0 = (1\real)" -| "r ^ Suc n = (r\real) * r ^ n" - -instance proof - fix z :: real - fix n :: nat - show "z^0 = 1" by simp - show "z^(Suc n) = z * (z^n)" by simp -qed - -declare power_real.simps [simp del] - -end - +instance real :: recpower .. lemma two_realpow_ge_one [simp]: "(1::real) \ 2 ^ n" by simp @@ -47,7 +30,6 @@ lemma realpow_minus_mult [rule_format]: "0 < n --> (x::real) ^ (n - 1) * x = x ^ n" -unfolding One_nat_def apply (simp split add: nat_diff_split) done diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Relation_Power.thy --- a/src/HOL/Relation_Power.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Relation_Power.thy Mon Apr 27 07:26:17 2009 -0700 @@ -9,132 +9,124 @@ imports Power Transitive_Closure Plain begin -instance - "fun" :: (type, type) power .. - --{* only type @{typ "'a => 'a"} should be in class @{text power}!*} +consts funpower :: "('a \ 'b) \ nat \ 'a \ 'b" (infixr "^^" 80) overloading - relpow \ "power \ ('a \ 'a) set \ nat \ ('a \ 'a) set" (unchecked) + relpow \ "funpower \ ('a \ 'a) set \ nat \ ('a \ 'a) set" begin -text {* @{text "R ^ n = R O ... O R"}, the n-fold composition of @{text R} *} +text {* @{text "R ^^ n = R O ... O R"}, the n-fold composition of @{text R} *} primrec relpow where - "(R \ ('a \ 'a) set) ^ 0 = Id" - | "(R \ ('a \ 'a) set) ^ Suc n = R O (R ^ n)" + "(R \ ('a \ 'a) set) ^^ 0 = Id" + | "(R \ ('a \ 'a) set) ^^ Suc n = R O (R ^^ n)" end overloading - funpow \ "power \ ('a \ 'a) \ nat \ 'a \ 'a" (unchecked) + funpow \ "funpower \ ('a \ 'a) \ nat \ 'a \ 'a" begin -text {* @{text "f ^ n = f o ... o f"}, the n-fold composition of @{text f} *} +text {* @{text "f ^^ n = f o ... o f"}, the n-fold composition of @{text f} *} primrec funpow where - "(f \ 'a \ 'a) ^ 0 = id" - | "(f \ 'a \ 'a) ^ Suc n = f o (f ^ n)" + "(f \ 'a \ 'a) ^^ 0 = id" + | "(f \ 'a \ 'a) ^^ Suc n = f o (f ^^ n)" end -text{*WARNING: due to the limits of Isabelle's type classes, exponentiation on -functions and relations has too general a domain, namely @{typ "('a * 'b)set"} -and @{typ "'a => 'b"}. Explicit type constraints may therefore be necessary. -For example, @{term "range(f^n) = A"} and @{term "Range(R^n) = B"} need -constraints.*} - -text {* - Circumvent this problem for code generation: -*} - -primrec - fun_pow :: "nat \ ('a \ 'a) \ 'a \ 'a" -where - "fun_pow 0 f = id" +primrec fun_pow :: "nat \ ('a \ 'a) \ 'a \ 'a" where + "fun_pow 0 f = id" | "fun_pow (Suc n) f = f o fun_pow n f" -lemma funpow_fun_pow [code unfold]: "f ^ n = fun_pow n f" +lemma funpow_fun_pow [code unfold]: + "f ^^ n = fun_pow n f" unfolding funpow_def fun_pow_def .. -lemma funpow_add: "f ^ (m+n) = f^m o f^n" +lemma funpow_add: + "f ^^ (m + n) = f ^^ m o f ^^ n" by (induct m) simp_all -lemma funpow_swap1: "f((f^n) x) = (f^n)(f x)" +lemma funpow_swap1: + "f ((f ^^ n) x) = (f ^^ n) (f x)" proof - - have "f((f^n) x) = (f^(n+1)) x" unfolding One_nat_def by simp - also have "\ = (f^n o f^1) x" by (simp only: funpow_add) - also have "\ = (f^n)(f x)" unfolding One_nat_def by simp + have "f ((f ^^ n) x) = (f ^^ (n+1)) x" unfolding One_nat_def by simp + also have "\ = (f ^^ n o f ^^ 1) x" by (simp only: funpow_add) + also have "\ = (f ^^ n) (f x)" unfolding One_nat_def by simp finally show ?thesis . qed lemma rel_pow_1 [simp]: - fixes R :: "('a*'a)set" - shows "R^1 = R" - unfolding One_nat_def by simp - -lemma rel_pow_0_I: "(x,x) : R^0" + fixes R :: "('a * 'a) set" + shows "R ^^ 1 = R" by simp -lemma rel_pow_Suc_I: "[| (x,y) : R^n; (y,z):R |] ==> (x,z):R^(Suc n)" +lemma rel_pow_0_I: + "(x, x) \ R ^^ 0" + by simp + +lemma rel_pow_Suc_I: + "(x, y) \ R ^^ n \ (y, z) \ R \ (x, z) \ R ^^ Suc n" by auto lemma rel_pow_Suc_I2: - "(x, y) : R \ (y, z) : R^n \ (x,z) : R^(Suc n)" - apply (induct n arbitrary: z) - apply simp - apply fastsimp - done + "(x, y) \ R \ (y, z) \ R ^^ n \ (x, z) \ R ^^ Suc n" + by (induct n arbitrary: z) (simp, fastsimp) -lemma rel_pow_0_E: "[| (x,y) : R^0; x=y ==> P |] ==> P" +lemma rel_pow_0_E: + "(x, y) \ R ^^ 0 \ (x = y \ P) \ P" by simp lemma rel_pow_Suc_E: - "[| (x,z) : R^(Suc n); !!y. [| (x,y) : R^n; (y,z) : R |] ==> P |] ==> P" + "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R ^^ n \ (y, z) \ R \ P) \ P" by auto lemma rel_pow_E: - "[| (x,z) : R^n; [| n=0; x = z |] ==> P; - !!y m. [| n = Suc m; (x,y) : R^m; (y,z) : R |] ==> P - |] ==> P" + "(x, z) \ R ^^ n \ (n = 0 \ x = z \ P) + \ (\y m. n = Suc m \ (x, y) \ R ^^ m \ (y, z) \ R \ P) + \ P" by (cases n) auto lemma rel_pow_Suc_D2: - "(x, z) : R^(Suc n) \ (\y. (x,y) : R & (y,z) : R^n)" + "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R \ (y, z) \ R ^^ n)" apply (induct n arbitrary: x z) apply (blast intro: rel_pow_0_I elim: rel_pow_0_E rel_pow_Suc_E) apply (blast intro: rel_pow_Suc_I elim: rel_pow_0_E rel_pow_Suc_E) done lemma rel_pow_Suc_D2': - "\x y z. (x,y) : R^n & (y,z) : R --> (\w. (x,w) : R & (w,z) : R^n)" + "\x y z. (x, y) \ R ^^ n \ (y, z) \ R \ (\w. (x, w) \ R \ (w, z) \ R ^^ n)" by (induct n) (simp_all, blast) lemma rel_pow_E2: - "[| (x,z) : R^n; [| n=0; x = z |] ==> P; - !!y m. [| n = Suc m; (x,y) : R; (y,z) : R^m |] ==> P - |] ==> P" - apply (case_tac n, simp) + "(x, z) \ R ^^ n \ (n = 0 \ x = z \ P) + \ (\y m. n = Suc m \ (x, y) \ R \ (y, z) \ R ^^ m \ P) + \ P" + apply (cases n, simp) apply (cut_tac n=nat and R=R in rel_pow_Suc_D2', simp, blast) done -lemma rtrancl_imp_UN_rel_pow: "!!p. p:R^* ==> p : (UN n. R^n)" - apply (simp only: split_tupled_all) +lemma rtrancl_imp_UN_rel_pow: + "p \ R^* \ p \ (\n. R ^^ n)" + apply (cases p) apply (simp only:) apply (erule rtrancl_induct) apply (blast intro: rel_pow_0_I rel_pow_Suc_I)+ done -lemma rel_pow_imp_rtrancl: "!!p. p:R^n ==> p:R^*" - apply (simp only: split_tupled_all) - apply (induct n) +lemma rel_pow_imp_rtrancl: + "p \ R ^^ n \ p \ R^*" + apply (induct n arbitrary: p) + apply (simp_all only: split_tupled_all) apply (blast intro: rtrancl_refl elim: rel_pow_0_E) apply (blast elim: rel_pow_Suc_E intro: rtrancl_into_rtrancl) done -lemma rtrancl_is_UN_rel_pow: "R^* = (UN n. R^n)" +lemma rtrancl_is_UN_rel_pow: + "R^* = (UN n. R ^^ n)" by (blast intro: rtrancl_imp_UN_rel_pow rel_pow_imp_rtrancl) lemma trancl_power: - "x \ r^+ = (\n > 0. x \ r^n)" + "x \ r^+ = (\n > 0. x \ r ^^ n)" apply (cases x) apply simp apply (rule iffI) @@ -151,30 +143,12 @@ done lemma single_valued_rel_pow: - "!!r::('a * 'a)set. single_valued r ==> single_valued (r^n)" + fixes R :: "('a * 'a) set" + shows "single_valued R \ single_valued (R ^^ n)" + apply (induct n arbitrary: R) + apply simp_all apply (rule single_valuedI) - apply (induct n) - apply simp apply (fast dest: single_valuedD elim: rel_pow_Suc_E) done -ML -{* -val funpow_add = thm "funpow_add"; -val rel_pow_1 = thm "rel_pow_1"; -val rel_pow_0_I = thm "rel_pow_0_I"; -val rel_pow_Suc_I = thm "rel_pow_Suc_I"; -val rel_pow_Suc_I2 = thm "rel_pow_Suc_I2"; -val rel_pow_0_E = thm "rel_pow_0_E"; -val rel_pow_Suc_E = thm "rel_pow_Suc_E"; -val rel_pow_E = thm "rel_pow_E"; -val rel_pow_Suc_D2 = thm "rel_pow_Suc_D2"; -val rel_pow_Suc_D2 = thm "rel_pow_Suc_D2"; -val rel_pow_E2 = thm "rel_pow_E2"; -val rtrancl_imp_UN_rel_pow = thm "rtrancl_imp_UN_rel_pow"; -val rel_pow_imp_rtrancl = thm "rel_pow_imp_rtrancl"; -val rtrancl_is_UN_rel_pow = thm "rtrancl_is_UN_rel_pow"; -val single_valued_rel_pow = thm "single_valued_rel_pow"; -*} - end diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Ring_and_Field.thy --- a/src/HOL/Ring_and_Field.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Ring_and_Field.thy Mon Apr 27 07:26:17 2009 -0700 @@ -2226,15 +2226,21 @@ qed qed -instance ordered_idom \ pordered_ring_abs -by default (auto simp add: abs_if not_less - equal_neg_zero neg_equal_zero mult_less_0_iff) - -lemma abs_mult: "abs (a * b) = abs a * abs (b::'a::ordered_idom)" -by (simp add: abs_eq_mult linorder_linear) - -lemma abs_mult_self: "abs a * abs a = a * (a::'a::ordered_idom)" -by (simp add: abs_if) +context ordered_idom +begin + +subclass pordered_ring_abs proof +qed (auto simp add: abs_if not_less equal_neg_zero neg_equal_zero mult_less_0_iff) + +lemma abs_mult: + "abs (a * b) = abs a * abs b" + by (rule abs_eq_mult) auto + +lemma abs_mult_self: + "abs a * abs a = a * a" + by (simp add: abs_if) + +end lemma nonzero_abs_inverse: "a \ 0 ==> abs (inverse (a::'a::ordered_field)) = inverse (abs a)" diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/SizeChange/Graphs.thy --- a/src/HOL/SizeChange/Graphs.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/SizeChange/Graphs.thy Mon Apr 27 07:26:17 2009 -0700 @@ -228,18 +228,8 @@ qed qed -instantiation graph :: (type, monoid_mult) "{semiring_1, idem_add, recpower, star}" -begin - -primrec power_graph :: "('a\type, 'b\monoid_mult) graph \ nat => ('a, 'b) graph" -where - "(A \ ('a, 'b) graph) ^ 0 = 1" -| "(A \ ('a, 'b) graph) ^ Suc n = A * (A ^ n)" - -definition - graph_star_def: "star (G \ ('a, 'b) graph) = (SUP n. G ^ n)" - -instance proof +instance graph :: (type, monoid_mult) "{semiring_1, idem_add, recpower}" +proof fix a b c :: "('a, 'b) graph" show "1 * a = a" @@ -258,10 +248,16 @@ show "a + a = a" unfolding graph_plus_def by simp - show "a ^ 0 = 1" "\n. a ^ (Suc n) = a * a ^ n" - by simp_all qed +instantiation graph :: (type, monoid_mult) star +begin + +definition + graph_star_def: "star (G \ ('a, 'b) graph) = (SUP n. G ^ n)" + +instance .. + end lemma graph_leqI: @@ -351,7 +347,7 @@ lemma in_tcl: "has_edge (tcl G) a x b = (\n>0. has_edge (G ^ n) a x b)" - apply (auto simp: tcl_is_SUP in_SUP simp del: power_graph.simps power_Suc) + apply (auto simp: tcl_is_SUP in_SUP simp del: power.simps power_Suc) apply (rule_tac x = "n - 1" in exI, auto) done diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/SizeChange/Interpretation.thy --- a/src/HOL/SizeChange/Interpretation.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/SizeChange/Interpretation.thy Mon Apr 27 07:26:17 2009 -0700 @@ -35,7 +35,7 @@ and nia: "\x. \accp R x \ \accp R (f x)" by blast - let ?s = "\i. (f ^ i) x" + let ?s = "\i. (f ^^ i) x" { fix i diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/SizeChange/Size_Change_Termination.thy --- a/src/HOL/SizeChange/Size_Change_Termination.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/SizeChange/Size_Change_Termination.thy Mon Apr 27 07:26:17 2009 -0700 @@ -1,5 +1,4 @@ (* Title: HOL/Library/Size_Change_Termination.thy - ID: $Id$ Author: Alexander Krauss, TU Muenchen *) diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Tools/Qelim/presburger.ML --- a/src/HOL/Tools/Qelim/presburger.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Tools/Qelim/presburger.ML Mon Apr 27 07:26:17 2009 -0700 @@ -131,7 +131,7 @@ @{thm "div_0"}, @{thm "mod_0"}, @{thm "div_by_1"}, @{thm "mod_by_1"}, @{thm "div_1"}, @{thm "mod_1"}, @{thm "Suc_plus1"}] @ @{thms add_ac} - addsimprocs [cancel_div_mod_proc] + addsimprocs [cancel_div_mod_nat_proc, cancel_div_mod_int_proc] val splits_ss = comp_ss addsimps [@{thm "mod_div_equality'"}] addsplits [@{thm "split_zdiv"}, @{thm "split_zmod"}, @{thm "split_div'"}, @{thm "split_min"}, @{thm "split_max"}, @{thm "abs_split"}] diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Tools/atp_manager.ML --- a/src/HOL/Tools/atp_manager.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Tools/atp_manager.ML Mon Apr 27 07:26:17 2009 -0700 @@ -51,15 +51,17 @@ fun set_timeout time = CRITICAL (fn () => timeout := time); val _ = - ProofGeneralPgip.add_preference "Proof" + ProofGeneralPgip.add_preference Preferences.category_proof (Preferences.string_pref atps "ATP: provers" "Default automatic provers (separated by whitespace)"); -val _ = ProofGeneralPgip.add_preference "Proof" +val _ = + ProofGeneralPgip.add_preference Preferences.category_proof (Preferences.int_pref max_atps "ATP: maximum number" "How many provers may run in parallel"); -val _ = ProofGeneralPgip.add_preference "Proof" +val _ = + ProofGeneralPgip.add_preference Preferences.category_proof (Preferences.int_pref timeout "ATP: timeout" "ATPs will be interrupted after this time (in seconds)"); diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Tools/atp_wrapper.ML --- a/src/HOL/Tools/atp_wrapper.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Tools/atp_wrapper.ML Mon Apr 27 07:26:17 2009 -0700 @@ -1,5 +1,4 @@ (* Title: HOL/Tools/atp_wrapper.ML - ID: $Id$ Author: Fabian Immler, TU Muenchen Wrapper functions for external ATPs. @@ -179,7 +178,7 @@ fun remote_prover_opts max_new theory_const args timeout = tptp_prover_opts max_new theory_const - (Path.explode "$ISABELLE_HOME/contrib/SystemOnTPTP/remote", args ^ " -t " ^ string_of_int timeout) + (Path.explode "$ISABELLE_HOME/lib/scripts/SystemOnTPTP", args ^ " -t " ^ string_of_int timeout) timeout; val remote_prover = remote_prover_opts 60 false; diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Tools/int_factor_simprocs.ML --- a/src/HOL/Tools/int_factor_simprocs.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Tools/int_factor_simprocs.ML Mon Apr 27 07:26:17 2009 -0700 @@ -1,5 +1,4 @@ (* Title: HOL/int_factor_simprocs.ML - ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 2000 University of Cambridge @@ -46,13 +45,13 @@ @{thm mult_zero_right}, @{thm mult_Bit1}, @{thm mult_1_right}]; end -(*Version for integer division*) -structure IntDivCancelNumeralFactor = CancelNumeralFactorFun +(*Version for semiring_div*) +structure DivCancelNumeralFactor = CancelNumeralFactorFun (open CancelNumeralFactorCommon val prove_conv = Arith_Data.prove_conv val mk_bal = HOLogic.mk_binop @{const_name Divides.div} - val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.intT - val cancel = @{thm zdiv_zmult_zmult1} RS trans + val dest_bal = HOLogic.dest_bin @{const_name Divides.div} Term.dummyT + val cancel = @{thm div_mult_mult1} RS trans val neg_exchanges = false ) @@ -108,8 +107,9 @@ "(l::'a::{ordered_idom,number_ring}) <= m * n"], K LeCancelNumeralFactor.proc), ("int_div_cancel_numeral_factors", - ["((l::int) * m) div n", "(l::int) div (m * n)"], - K IntDivCancelNumeralFactor.proc), + ["((l::'a::{semiring_div,number_ring}) * m) div n", + "(l::'a::{semiring_div,number_ring}) div (m * n)"], + K DivCancelNumeralFactor.proc), ("divide_cancel_numeral_factor", ["((l::'a::{division_by_zero,field,number_ring}) * m) / n", "(l::'a::{division_by_zero,field,number_ring}) / (m * n)", @@ -284,24 +284,25 @@ @{thm mult_less_cancel_left_pos} @{thm mult_less_cancel_left_neg} ); -(*zdiv_zmult_zmult1_if is for integer division (div).*) -structure IntDivCancelFactor = ExtractCommonTermFun +(*for semirings with division*) +structure DivCancelFactor = ExtractCommonTermFun (open CancelFactorCommon val prove_conv = Arith_Data.prove_conv val mk_bal = HOLogic.mk_binop @{const_name Divides.div} - val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.intT - val simp_conv = K (K (SOME @{thm zdiv_zmult_zmult1_if})) + val dest_bal = HOLogic.dest_bin @{const_name Divides.div} Term.dummyT + val simp_conv = K (K (SOME @{thm div_mult_mult1_if})) ); -structure IntModCancelFactor = ExtractCommonTermFun +structure ModCancelFactor = ExtractCommonTermFun (open CancelFactorCommon val prove_conv = Arith_Data.prove_conv val mk_bal = HOLogic.mk_binop @{const_name Divides.mod} val dest_bal = HOLogic.dest_bin @{const_name Divides.mod} HOLogic.intT - val simp_conv = K (K (SOME @{thm zmod_zmult_zmult1})) + val simp_conv = K (K (SOME @{thm mod_mult_mult1})) ); -structure IntDvdCancelFactor = ExtractCommonTermFun +(*for idoms*) +structure DvdCancelFactor = ExtractCommonTermFun (open CancelFactorCommon val prove_conv = Arith_Data.prove_conv val mk_bal = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd} @@ -321,8 +322,8 @@ val cancel_factors = map Arith_Data.prep_simproc [("ring_eq_cancel_factor", - ["(l::'a::{idom}) * m = n", - "(l::'a::{idom}) = m * n"], + ["(l::'a::idom) * m = n", + "(l::'a::idom) = m * n"], K EqCancelFactor.proc), ("ordered_ring_le_cancel_factor", ["(l::'a::ordered_ring) * m <= n", @@ -333,14 +334,14 @@ "(l::'a::ordered_ring) < m * n"], K LessCancelFactor.proc), ("int_div_cancel_factor", - ["((l::int) * m) div n", "(l::int) div (m * n)"], - K IntDivCancelFactor.proc), + ["((l::'a::semiring_div) * m) div n", "(l::'a::semiring_div) div (m * n)"], + K DivCancelFactor.proc), ("int_mod_cancel_factor", - ["((l::int) * m) mod n", "(l::int) mod (m * n)"], - K IntModCancelFactor.proc), + ["((l::'a::semiring_div) * m) mod n", "(l::'a::semiring_div) mod (m * n)"], + K ModCancelFactor.proc), ("dvd_cancel_factor", ["((l::'a::idom) * m) dvd n", "(l::'a::idom) dvd (m * n)"], - K IntDvdCancelFactor.proc), + K DvdCancelFactor.proc), ("divide_cancel_factor", ["((l::'a::{division_by_zero,field}) * m) / n", "(l::'a::{division_by_zero,field}) / (m * n)"], diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Transitive_Closure.thy --- a/src/HOL/Transitive_Closure.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Transitive_Closure.thy Mon Apr 27 07:26:17 2009 -0700 @@ -630,6 +630,140 @@ declare trancl_into_rtrancl [elim] +subsection {* The power operation on relations *} + +text {* @{text "R ^^ n = R O ... O R"}, the n-fold composition of @{text R} *} + +overloading + relpow == "compow :: nat \ ('a \ 'a) set \ ('a \ 'a) set" +begin + +primrec relpow :: "nat \ ('a \ 'a) set \ ('a \ 'a) set" where + "relpow 0 R = Id" + | "relpow (Suc n) R = R O (R ^^ n)" + +end + +lemma rel_pow_1 [simp]: + fixes R :: "('a \ 'a) set" + shows "R ^^ 1 = R" + by simp + +lemma rel_pow_0_I: + "(x, x) \ R ^^ 0" + by simp + +lemma rel_pow_Suc_I: + "(x, y) \ R ^^ n \ (y, z) \ R \ (x, z) \ R ^^ Suc n" + by auto + +lemma rel_pow_Suc_I2: + "(x, y) \ R \ (y, z) \ R ^^ n \ (x, z) \ R ^^ Suc n" + by (induct n arbitrary: z) (simp, fastsimp) + +lemma rel_pow_0_E: + "(x, y) \ R ^^ 0 \ (x = y \ P) \ P" + by simp + +lemma rel_pow_Suc_E: + "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R ^^ n \ (y, z) \ R \ P) \ P" + by auto + +lemma rel_pow_E: + "(x, z) \ R ^^ n \ (n = 0 \ x = z \ P) + \ (\y m. n = Suc m \ (x, y) \ R ^^ m \ (y, z) \ R \ P) + \ P" + by (cases n) auto + +lemma rel_pow_Suc_D2: + "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R \ (y, z) \ R ^^ n)" + apply (induct n arbitrary: x z) + apply (blast intro: rel_pow_0_I elim: rel_pow_0_E rel_pow_Suc_E) + apply (blast intro: rel_pow_Suc_I elim: rel_pow_0_E rel_pow_Suc_E) + done + +lemma rel_pow_Suc_E2: + "(x, z) \ R ^^ Suc n \ (\y. (x, y) \ R \ (y, z) \ R ^^ n \ P) \ P" + by (blast dest: rel_pow_Suc_D2) + +lemma rel_pow_Suc_D2': + "\x y z. (x, y) \ R ^^ n \ (y, z) \ R \ (\w. (x, w) \ R \ (w, z) \ R ^^ n)" + by (induct n) (simp_all, blast) + +lemma rel_pow_E2: + "(x, z) \ R ^^ n \ (n = 0 \ x = z \ P) + \ (\y m. n = Suc m \ (x, y) \ R \ (y, z) \ R ^^ m \ P) + \ P" + apply (cases n, simp) + apply (cut_tac n=nat and R=R in rel_pow_Suc_D2', simp, blast) + done + +lemma rtrancl_imp_UN_rel_pow: + assumes "p \ R^*" + shows "p \ (\n. R ^^ n)" +proof (cases p) + case (Pair x y) + with assms have "(x, y) \ R^*" by simp + then have "(x, y) \ (\n. R ^^ n)" proof induct + case base show ?case by (blast intro: rel_pow_0_I) + next + case step then show ?case by (blast intro: rel_pow_Suc_I) + qed + with Pair show ?thesis by simp +qed + +lemma rel_pow_imp_rtrancl: + assumes "p \ R ^^ n" + shows "p \ R^*" +proof (cases p) + case (Pair x y) + with assms have "(x, y) \ R ^^ n" by simp + then have "(x, y) \ R^*" proof (induct n arbitrary: x y) + case 0 then show ?case by simp + next + case Suc then show ?case + by (blast elim: rel_pow_Suc_E intro: rtrancl_into_rtrancl) + qed + with Pair show ?thesis by simp +qed + +lemma rtrancl_is_UN_rel_pow: + "R^* = (\n. R ^^ n)" + by (blast intro: rtrancl_imp_UN_rel_pow rel_pow_imp_rtrancl) + +lemma rtrancl_power: + "p \ R^* \ (\n. p \ R ^^ n)" + by (simp add: rtrancl_is_UN_rel_pow) + +lemma trancl_power: + "p \ R^+ \ (\n > 0. p \ R ^^ n)" + apply (cases p) + apply simp + apply (rule iffI) + apply (drule tranclD2) + apply (clarsimp simp: rtrancl_is_UN_rel_pow) + apply (rule_tac x="Suc n" in exI) + apply (clarsimp simp: rel_comp_def) + apply fastsimp + apply clarsimp + apply (case_tac n, simp) + apply clarsimp + apply (drule rel_pow_imp_rtrancl) + apply (drule rtrancl_into_trancl1) apply auto + done + +lemma rtrancl_imp_rel_pow: + "p \ R^* \ \n. p \ R ^^ n" + by (auto dest: rtrancl_imp_UN_rel_pow) + +lemma single_valued_rel_pow: + fixes R :: "('a * 'a) set" + shows "single_valued R \ single_valued (R ^^ n)" + apply (induct n arbitrary: R) + apply simp_all + apply (rule single_valuedI) + apply (fast dest: single_valuedD elim: rel_pow_Suc_E) + done subsection {* Setup of transitivity reasoner *} diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/UNITY/Comp.thy --- a/src/HOL/UNITY/Comp.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/UNITY/Comp.thy Mon Apr 27 07:26:17 2009 -0700 @@ -15,14 +15,22 @@ header{*Composition: Basic Primitives*} -theory Comp imports Union begin +theory Comp +imports Union +begin -instance program :: (type) ord .. +instantiation program :: (type) ord +begin -defs - component_def: "F \ H == \G. F\G = H" - strict_component_def: "(F < (H::'a program)) == (F \ H & F \ H)" +definition + component_def: "F \ H <-> (\G. F\G = H)" +definition + strict_component_def: "F < (H::'a program) <-> (F \ H & F \ H)" + +instance .. + +end constdefs component_of :: "'a program =>'a program=> bool" @@ -114,7 +122,7 @@ by (auto simp add: stable_def component_constrains) (*Used in Guar.thy to show that programs are partially ordered*) -lemmas program_less_le = strict_component_def [THEN meta_eq_to_obj_eq] +lemmas program_less_le = strict_component_def subsection{*The preserves property*} @@ -229,8 +237,7 @@ apply (blast intro: Join_assoc [symmetric]) done -lemmas strict_component_of_eq = - strict_component_of_def [THEN meta_eq_to_obj_eq, standard] +lemmas strict_component_of_eq = strict_component_of_def (** localize **) lemma localize_Init_eq [simp]: "Init (localize v F) = Init F" diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/UNITY/Transformers.thy --- a/src/HOL/UNITY/Transformers.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/UNITY/Transformers.thy Mon Apr 27 07:26:17 2009 -0700 @@ -338,10 +338,10 @@ constdefs wens_single_finite :: "[('a*'a) set, 'a set, nat] => 'a set" - "wens_single_finite act B k == \i \ atMost k. ((wp act)^i) B" + "wens_single_finite act B k == \i \ atMost k. (wp act ^^ i) B" wens_single :: "[('a*'a) set, 'a set] => 'a set" - "wens_single act B == \i. ((wp act)^i) B" + "wens_single act B == \i. (wp act ^^ i) B" lemma wens_single_Un_eq: "single_valued act diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Wellfounded.thy --- a/src/HOL/Wellfounded.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Wellfounded.thy Mon Apr 27 07:26:17 2009 -0700 @@ -7,7 +7,7 @@ header {*Well-founded Recursion*} theory Wellfounded -imports Finite_Set Transitive_Closure Nat +imports Finite_Set Transitive_Closure uses ("Tools/function_package/size.ML") begin diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Word/BinBoolList.thy --- a/src/HOL/Word/BinBoolList.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Word/BinBoolList.thy Mon Apr 27 07:26:17 2009 -0700 @@ -38,7 +38,7 @@ if y then rbl_add ws x else ws)" lemma butlast_power: - "(butlast ^ n) bl = take (length bl - n) bl" + "(butlast ^^ n) bl = take (length bl - n) bl" by (induct n) (auto simp: butlast_take) lemma bin_to_bl_aux_Pls_minus_simp [simp]: @@ -370,14 +370,14 @@ done lemma nth_rest_power_bin [rule_format] : - "ALL n. bin_nth ((bin_rest ^ k) w) n = bin_nth w (n + k)" + "ALL n. bin_nth ((bin_rest ^^ k) w) n = bin_nth w (n + k)" apply (induct k, clarsimp) apply clarsimp apply (simp only: bin_nth.Suc [symmetric] add_Suc) done lemma take_rest_power_bin: - "m <= n ==> take m (bin_to_bl n w) = bin_to_bl m ((bin_rest ^ (n - m)) w)" + "m <= n ==> take m (bin_to_bl n w) = bin_to_bl m ((bin_rest ^^ (n - m)) w)" apply (rule nth_equalityI) apply simp apply (clarsimp simp add: nth_bin_to_bl nth_rest_power_bin) diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Word/BinGeneral.thy --- a/src/HOL/Word/BinGeneral.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Word/BinGeneral.thy Mon Apr 27 07:26:17 2009 -0700 @@ -439,7 +439,7 @@ apply clarsimp apply (simp add: bin_last_mod bin_rest_div Bit_def cong: number_of_False_cong) - apply (clarsimp simp: zmod_zmult_zmult1 [symmetric] + apply (clarsimp simp: mod_mult_mult1 [symmetric] zmod_zdiv_equality [THEN diff_eq_eq [THEN iffD2 [THEN sym]]]) apply (rule trans [symmetric, OF _ emep1]) apply auto @@ -822,8 +822,8 @@ by (induct n) auto lemma bin_rest_power_trunc [rule_format] : - "(bin_rest ^ k) (bintrunc n bin) = - bintrunc (n - k) ((bin_rest ^ k) bin)" + "(bin_rest ^^ k) (bintrunc n bin) = + bintrunc (n - k) ((bin_rest ^^ k) bin)" by (induct k) (auto simp: bin_rest_trunc) lemma bin_rest_trunc_i: @@ -857,7 +857,7 @@ by (rule ext) auto lemma rco_lem: - "f o g o f = g o f ==> f o (g o f) ^ n = g ^ n o f" + "f o g o f = g o f ==> f o (g o f) ^^ n = g ^^ n o f" apply (rule ext) apply (induct_tac n) apply (simp_all (no_asm)) @@ -867,7 +867,7 @@ apply simp done -lemma rco_alt: "(f o g) ^ n o f = f o (g o f) ^ n" +lemma rco_alt: "(f o g) ^^ n o f = f o (g o f) ^^ n" apply (rule ext) apply (induct n) apply (simp_all add: o_def) @@ -891,8 +891,9 @@ subsection {* Miscellaneous lemmas *} -lemmas funpow_minus_simp = - trans [OF gen_minus [where f = "power f"] funpow_Suc, standard] +lemma funpow_minus_simp: + "0 < n \ f ^^ n = f \ f ^^ (n - 1)" + by (cases n) simp_all lemmas funpow_pred_simp [simp] = funpow_minus_simp [of "number_of bin", simplified nobm1, standard] diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Word/BinOperations.thy --- a/src/HOL/Word/BinOperations.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Word/BinOperations.thy Mon Apr 27 07:26:17 2009 -0700 @@ -641,7 +641,7 @@ apply (simp add: bin_rest_div zdiv_zmult2_eq) apply (case_tac b rule: bin_exhaust) apply simp - apply (simp add: Bit_def zmod_zmult_zmult1 p1mod22k + apply (simp add: Bit_def mod_mult_mult1 p1mod22k split: bit.split cong: number_of_False_cong) done diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Word/Num_Lemmas.thy --- a/src/HOL/Word/Num_Lemmas.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Word/Num_Lemmas.thy Mon Apr 27 07:26:17 2009 -0700 @@ -66,7 +66,7 @@ apply (safe dest!: even_equiv_def [THEN iffD1]) apply (subst pos_zmod_mult_2) apply arith - apply (simp add: zmod_zmult_zmult1) + apply (simp add: mod_mult_mult1) done lemmas eme1p = emep1 [simplified add_commute] diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Word/TdThs.thy --- a/src/HOL/Word/TdThs.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Word/TdThs.thy Mon Apr 27 07:26:17 2009 -0700 @@ -110,7 +110,7 @@ done lemma fn_comm_power: - "fa o tr = tr o fr ==> fa ^ n o tr = tr o fr ^ n" + "fa o tr = tr o fr ==> fa ^^ n o tr = tr o fr ^^ n" apply (rule ext) apply (induct n) apply (auto dest: fun_cong) diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Word/WordArith.thy --- a/src/HOL/Word/WordArith.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Word/WordArith.thy Mon Apr 27 07:26:17 2009 -0700 @@ -778,6 +778,8 @@ apply (simp add: word_mult_1) done +instance word :: (len0) recpower .. + instance word :: (len0) comm_semiring by (intro_classes) (simp add : word_left_distrib) diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Word/WordBitwise.thy --- a/src/HOL/Word/WordBitwise.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Word/WordBitwise.thy Mon Apr 27 07:26:17 2009 -0700 @@ -443,8 +443,10 @@ lemmas test_bit_2p = refl [THEN test_bit_2p', unfolded word_size] -lemmas nth_w2p = test_bit_2p [unfolded of_int_number_of_eq - word_of_int [symmetric] Int.of_int_power] +lemma nth_w2p: + "((2\'a\len word) ^ n) !! m \ m = n \ m < len_of TYPE('a\len)" + unfolding test_bit_2p [symmetric] word_of_int [symmetric] + by (simp add: of_int_power) lemma uint_2p: "(0::'a::len word) < 2 ^ n ==> uint (2 ^ n::'a::len word) = 2 ^ n" diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Word/WordDefinition.thy --- a/src/HOL/Word/WordDefinition.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Word/WordDefinition.thy Mon Apr 27 07:26:17 2009 -0700 @@ -99,7 +99,7 @@ subsection "Arithmetic operations" -instantiation word :: (len0) "{number, uminus, minus, plus, one, zero, times, Divides.div, power, ord, bit}" +instantiation word :: (len0) "{number, uminus, minus, plus, one, zero, times, Divides.div, ord, bit}" begin definition @@ -126,10 +126,6 @@ definition word_mod_def: "a mod b = word_of_int (uint a mod uint b)" -primrec power_word where - "(a\'a word) ^ 0 = 1" - | "(a\'a word) ^ Suc n = a * a ^ n" - definition word_number_of_def: "number_of w = word_of_int w" @@ -157,7 +153,7 @@ instance .. -end +end definition word_succ :: "'a :: len0 word => 'a word" @@ -207,10 +203,10 @@ "shiftr1 w = word_of_int (bin_rest (uint w))" definition - shiftl_def: "w << n = (shiftl1 ^ n) w" + shiftl_def: "w << n = (shiftl1 ^^ n) w" definition - shiftr_def: "w >> n = (shiftr1 ^ n) w" + shiftr_def: "w >> n = (shiftr1 ^^ n) w" instance .. @@ -245,7 +241,7 @@ "bshiftr1 b w == of_bl (b # butlast (to_bl w))" sshiftr :: "'a :: len word => nat => 'a word" (infixl ">>>" 55) - "w >>> n == (sshiftr1 ^ n) w" + "w >>> n == (sshiftr1 ^^ n) w" mask :: "nat => 'a::len word" "mask n == (1 << n) - 1" @@ -268,7 +264,7 @@ case ys of [] => [] | x # xs => last ys # butlast ys" rotater :: "nat => 'a list => 'a list" - "rotater n == rotater1 ^ n" + "rotater n == rotater1 ^^ n" word_rotr :: "nat => 'a :: len0 word => 'a :: len0 word" "word_rotr n w == of_bl (rotater n (to_bl w))" @@ -303,7 +299,7 @@ constdefs -- "Largest representable machine integer." max_word :: "'a::len word" - "max_word \ word_of_int (2^len_of TYPE('a) - 1)" + "max_word \ word_of_int (2 ^ len_of TYPE('a) - 1)" consts of_bool :: "bool \ 'a::len word" diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/Word/WordShift.thy --- a/src/HOL/Word/WordShift.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/Word/WordShift.thy Mon Apr 27 07:26:17 2009 -0700 @@ -361,14 +361,14 @@ lemma shiftr_no': "w = number_of bin ==> - (w::'a::len0 word) >> n = number_of ((bin_rest ^ n) (bintrunc (size w) bin))" + (w::'a::len0 word) >> n = number_of ((bin_rest ^^ n) (bintrunc (size w) bin))" apply clarsimp apply (rule word_eqI) apply (auto simp: nth_shiftr nth_rest_power_bin nth_bintr word_size) done lemma sshiftr_no': - "w = number_of bin ==> w >>> n = number_of ((bin_rest ^ n) + "w = number_of bin ==> w >>> n = number_of ((bin_rest ^^ n) (sbintrunc (size w - 1) bin))" apply clarsimp apply (rule word_eqI) diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/base.ML --- a/src/HOL/base.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/base.ML Mon Apr 27 07:26:17 2009 -0700 @@ -1,2 +1,2 @@ (*side-entry for HOL-Base*) -use_thy "Code_Setup"; +use_thy "HOL"; diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/ex/NormalForm.thy --- a/src/HOL/ex/NormalForm.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/ex/NormalForm.thy Mon Apr 27 07:26:17 2009 -0700 @@ -1,7 +1,6 @@ -(* Authors: Klaus Aehlig, Tobias Nipkow -*) +(* Authors: Klaus Aehlig, Tobias Nipkow *) -header {* Test of normalization function *} +header {* Testing implementation of normalization by evaluation *} theory NormalForm imports Main Rational @@ -19,18 +18,13 @@ datatype n = Z | S n -consts - add :: "n \ n \ n" - add2 :: "n \ n \ n" - mul :: "n \ n \ n" - mul2 :: "n \ n \ n" - exp :: "n \ n \ n" -primrec - "add Z = id" - "add (S m) = S o add m" -primrec - "add2 Z n = n" - "add2 (S m) n = S(add2 m n)" +primrec add :: "n \ n \ n" where + "add Z = id" + | "add (S m) = S o add m" + +primrec add2 :: "n \ n \ n" where + "add2 Z n = n" + | "add2 (S m) n = S(add2 m n)" declare add2.simps [code] lemma [code nbe]: "add2 (add2 n m) k = add2 n (add2 m k)" @@ -44,15 +38,17 @@ lemma "add2 (add2 (S n) (S m)) (S k) = S(S(S(add2 n (add2 m k))))" by normalization lemma "add2 (add2 (S n) (add2 (S m) Z)) (S k) = S(S(S(add2 n (add2 m k))))" by normalization -primrec - "mul Z = (%n. Z)" - "mul (S m) = (%n. add (mul m n) n)" -primrec - "mul2 Z n = Z" - "mul2 (S m) n = add2 n (mul2 m n)" -primrec - "exp m Z = S Z" - "exp m (S n) = mul (exp m n) m" +primrec mul :: "n \ n \ n" where + "mul Z = (%n. Z)" + | "mul (S m) = (%n. add (mul m n) n)" + +primrec mul2 :: "n \ n \ n" where + "mul2 Z n = Z" + | "mul2 (S m) n = add2 n (mul2 m n)" + +primrec exp :: "n \ n \ n" where + "exp m Z = S Z" + | "exp m (S n) = mul (exp m n) m" lemma "mul2 (S(S(S(S(S Z))))) (S(S(S Z))) = S(S(S(S(S(S(S(S(S(S(S(S(S(S(S Z))))))))))))))" by normalization lemma "mul (S(S(S(S(S Z))))) (S(S(S Z))) = S(S(S(S(S(S(S(S(S(S(S(S(S(S(S Z))))))))))))))" by normalization diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/ex/Predicate_Compile.thy --- a/src/HOL/ex/Predicate_Compile.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/ex/Predicate_Compile.thy Mon Apr 27 07:26:17 2009 -0700 @@ -1,37 +1,86 @@ theory Predicate_Compile -imports Complex_Main Code_Index Lattice_Syntax +imports Complex_Main Lattice_Syntax Code_Eval uses "predicate_compile.ML" begin +text {* Package setup *} + setup {* Predicate_Compile.setup *} -primrec "next" :: "('a Predicate.pred \ ('a \ 'a Predicate.pred) option) - \ 'a Predicate.seq \ ('a \ 'a Predicate.pred) option" where - "next yield Predicate.Empty = None" - | "next yield (Predicate.Insert x P) = Some (x, P)" - | "next yield (Predicate.Join P xq) = (case yield P - of None \ next yield xq | Some (x, Q) \ Some (x, Predicate.Seq (\_. Predicate.Join Q xq)))" + +text {* Experimental code *} + +definition pred_map :: "('a \ 'b) \ 'a Predicate.pred \ 'b Predicate.pred" where + "pred_map f P = Predicate.bind P (Predicate.single o f)" ML {* -let - fun yield (@{code Predicate.Seq} f) = @{code next} yield (f ()) -in - yield @{code "\ :: 'a Predicate.pred"} (*replace bottom with sequence to evaluate*) -end +structure Predicate = +struct + +open Predicate; + +val pred_ref = ref (NONE : (unit -> term Predicate.pred) option); + +fun eval_pred thy t = + t + |> Eval.mk_term_of (fastype_of t) + |> (fn t => Code_ML.eval NONE ("Predicate.pred_ref", pred_ref) @{code pred_map} thy t []); + +fun eval_pred_elems thy t T length = + t |> eval_pred thy |> yieldn length |> fst |> HOLogic.mk_list T; + +fun analyze_compr thy t = + let + val split = case t of (Const (@{const_name Collect}, _) $ t') => t' + | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term_global thy t); + val (body, Ts, fp) = HOLogic.strip_split split; + val (t_pred, args) = strip_comb body; + val pred = case t_pred of Const (pred, _) => pred + | _ => error ("Not a constant: " ^ Syntax.string_of_term_global thy t_pred); + val mode = map is_Bound args; (*FIXME what about higher-order modes?*) + val args' = filter_out is_Bound args; + val T = HOLogic.mk_tupleT fp Ts; + val mk = HOLogic.mk_tuple' fp T; + in (((pred, mode), args), (mk, T)) end; + +end; *} -fun anamorph :: "('b \ ('a \ 'b) option) \ index \ 'b \ 'a list \ 'b" where - "anamorph f k x = (if k = 0 then ([], x) - else case f x of None \ ([], x) | Some (v, y) \ let (vs, z) = anamorph f (k - 1) y in (v # vs, z))" + +text {* Example(s) *} + +inductive even :: "nat \ bool" and odd :: "nat \ bool" where + "even 0" + | "even n \ odd (Suc n)" + | "odd n \ even (Suc n)" + +setup {* pred_compile "even" *} +thm even_codegen + + +inductive append :: "'a list \ 'a list \ 'a list \ bool" where + append_Nil: "append [] xs xs" + | append_Cons: "append xs ys zs \ append (x # xs) ys (x # zs)" -ML {* -let - fun yield (@{code Predicate.Seq} f) = @{code next} yield (f ()) - fun yieldn k = @{code anamorph} yield k -in - yieldn 0 (*replace with number of elements to retrieve*) - @{code "\ :: 'a Predicate.pred"} (*replace bottom with sequence to evaluate*) -end -*} +setup {* pred_compile "append" *} +thm append_codegen + + +inductive partition :: "('a \ bool) \ 'a list \ 'a list \ 'a list \ bool" + for f where + "partition f [] [] []" + | "f x \ partition f xs ys zs \ partition f (x # xs) (x # ys) zs" + | "\ f x \ partition f xs ys zs \ partition f (x # xs) ys (x # zs)" + +setup {* pred_compile "partition" *} +thm partition_codegen + +setup {* pred_compile "tranclp" *} +thm tranclp_codegen + +ML_val {* Predicate_Compile.modes_of @{theory} @{const_name partition} *} +ML_val {* Predicate_Compile.modes_of @{theory} @{const_name tranclp} *} + +ML_val {* Predicate.analyze_compr @{theory} @{term "{n. odd n}"} *} end \ No newline at end of file diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/ex/Quickcheck_Generators.thy --- a/src/HOL/ex/Quickcheck_Generators.thy Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/ex/Quickcheck_Generators.thy Mon Apr 27 07:26:17 2009 -0700 @@ -6,62 +6,6 @@ imports Quickcheck State_Monad begin -subsection {* Type @{typ "'a \ 'b"} *} - -ML {* -structure Random_Engine = -struct - -open Random_Engine; - -fun random_fun (T1 : typ) (T2 : typ) (eq : 'a -> 'a -> bool) (term_of : 'a -> term) - (random : Random_Engine.seed -> ('b * (unit -> term)) * Random_Engine.seed) - (random_split : Random_Engine.seed -> Random_Engine.seed * Random_Engine.seed) - (seed : Random_Engine.seed) = - let - val (seed', seed'') = random_split seed; - val state = ref (seed', [], Const (@{const_name undefined}, T1 --> T2)); - val fun_upd = Const (@{const_name fun_upd}, - (T1 --> T2) --> T1 --> T2 --> T1 --> T2); - fun random_fun' x = - let - val (seed, fun_map, f_t) = ! state; - in case AList.lookup (uncurry eq) fun_map x - of SOME y => y - | NONE => let - val t1 = term_of x; - val ((y, t2), seed') = random seed; - val fun_map' = (x, y) :: fun_map; - val f_t' = fun_upd $ f_t $ t1 $ t2 (); - val _ = state := (seed', fun_map', f_t'); - in y end - end; - fun term_fun' () = #3 (! state); - in ((random_fun', term_fun'), seed'') end; - -end -*} - -axiomatization - random_fun_aux :: "typerep \ typerep \ ('a \ 'a \ bool) \ ('a \ term) - \ (seed \ ('b \ (unit \ term)) \ seed) \ (seed \ seed \ seed) - \ seed \ (('a \ 'b) \ (unit \ term)) \ seed" - -code_const random_fun_aux (SML "Random'_Engine.random'_fun") - -instantiation "fun" :: ("{eq, term_of}", "{type, random}") random -begin - -definition random_fun :: "index \ seed \ (('a \ 'b) \ (unit \ term)) \ seed" where - "random n = random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Eval.term_of (random n) split_seed" - -instance .. - -end - -code_reserved SML Random_Engine - - subsection {* Datatypes *} definition diff -r 644d18da3c77 -r 7c871a9cf6f4 src/HOL/ex/predicate_compile.ML --- a/src/HOL/ex/predicate_compile.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/HOL/ex/predicate_compile.ML Mon Apr 27 07:26:17 2009 -0700 @@ -6,38 +6,119 @@ signature PREDICATE_COMPILE = sig - val create_def_equation': string -> (int list option list * int list) option -> theory -> theory + type mode = int list option list * int list + val create_def_equation': string -> mode option -> theory -> theory val create_def_equation: string -> theory -> theory - val intro_rule: theory -> string -> (int list option list * int list) -> thm - val elim_rule: theory -> string -> (int list option list * int list) -> thm + val intro_rule: theory -> string -> mode -> thm + val elim_rule: theory -> string -> mode -> thm val strip_intro_concl : term -> int -> (term * (term list * term list)) val code_ind_intros_attrib : attribute val code_ind_cases_attrib : attribute + val print_alternative_rules : theory -> theory + val modename_of: theory -> string -> mode -> string + val modes_of: theory -> string -> mode list val setup : theory -> theory - val print_alternative_rules : theory -> theory val do_proofs: bool ref end; structure Predicate_Compile: PREDICATE_COMPILE = struct +(** auxiliary **) + +(* debug stuff *) + +fun tracing s = (if ! Toplevel.debug then Output.tracing s else ()); + +fun print_tac s = (if ! Toplevel.debug then Tactical.print_tac s else Seq.single); +fun debug_tac msg = (fn st => (tracing msg; Seq.single st)); + +val do_proofs = ref true; + + +(** fundamentals **) + +(* syntactic operations *) + +fun mk_eq (x, xs) = + let fun mk_eqs _ [] = [] + | mk_eqs a (b::cs) = + HOLogic.mk_eq (Free (a, fastype_of b), b) :: mk_eqs a cs + in mk_eqs x xs end; + +fun mk_tupleT [] = HOLogic.unitT + | mk_tupleT Ts = foldr1 HOLogic.mk_prodT Ts; + +fun mk_tuple [] = HOLogic.unit + | mk_tuple ts = foldr1 HOLogic.mk_prod ts; + +fun dest_tuple (Const (@{const_name Product_Type.Unity}, _)) = [] + | dest_tuple (Const (@{const_name Pair}, _) $ t1 $ t2) = t1 :: (dest_tuple t2) + | dest_tuple t = [t] + +fun mk_pred_enumT T = Type ("Predicate.pred", [T]) + +fun dest_pred_enumT (Type ("Predicate.pred", [T])) = T + | dest_pred_enumT T = raise TYPE ("dest_pred_enumT", [T], []); + +fun mk_Enum f = + let val T as Type ("fun", [T', _]) = fastype_of f + in + Const (@{const_name Predicate.Pred}, T --> mk_pred_enumT T') $ f + end; + +fun mk_Eval (f, x) = + let val T = fastype_of x + in + Const (@{const_name Predicate.eval}, mk_pred_enumT T --> T --> HOLogic.boolT) $ f $ x + end; + +fun mk_empty T = Const (@{const_name Orderings.bot}, mk_pred_enumT T); + +fun mk_single t = + let val T = fastype_of t + in Const(@{const_name Predicate.single}, T --> mk_pred_enumT T) $ t end; + +fun mk_bind (x, f) = + let val T as Type ("fun", [_, U]) = fastype_of f + in + Const (@{const_name Predicate.bind}, fastype_of x --> T --> U) $ x $ f + end; + +val mk_sup = HOLogic.mk_binop @{const_name sup}; + +fun mk_if_predenum cond = Const (@{const_name Predicate.if_pred}, + HOLogic.boolT --> mk_pred_enumT HOLogic.unitT) $ cond; + +fun mk_not_pred t = let val T = mk_pred_enumT HOLogic.unitT + in Const (@{const_name Predicate.not_pred}, T --> T) $ t end + + +(* data structures *) + +type mode = int list option list * int list; + +val mode_ord = prod_ord (list_ord (option_ord (list_ord int_ord))) (list_ord int_ord); + structure PredModetab = TableFun( - type key = (string * (int list option list * int list)) - val ord = prod_ord fast_string_ord (prod_ord - (list_ord (option_ord (list_ord int_ord))) (list_ord int_ord))) + type key = string * mode + val ord = prod_ord fast_string_ord mode_ord +); +(*FIXME scrap boilerplate*) + structure IndCodegenData = TheoryDataFun ( type T = {names : string PredModetab.table, - modes : ((int list option list * int list) list) Symtab.table, + modes : mode list Symtab.table, function_defs : Thm.thm Symtab.table, function_intros : Thm.thm Symtab.table, function_elims : Thm.thm Symtab.table, - intro_rules : (Thm.thm list) Symtab.table, + intro_rules : Thm.thm list Symtab.table, elim_rules : Thm.thm Symtab.table, nparams : int Symtab.table - }; + }; (*FIXME: better group tables according to key*) (* names: map from inductive predicate and mode to function name (string). modes: map from inductive predicates to modes function_defs: map from function name to definition @@ -115,26 +196,12 @@ intro_rules = #intro_rules x, elim_rules = #elim_rules x, nparams = f (#nparams x)}) thy -(* Debug stuff and tactics ***********************************************************) - -fun tracing s = (if ! Toplevel.debug then Output.tracing s else ()); -fun print_tac s = (if ! Toplevel.debug then Tactical.print_tac s else Seq.single); - -fun debug_tac msg = (fn st => - (tracing msg; Seq.single st)); - (* removes first subgoal *) fun mycheat_tac thy i st = (Tactic.rtac (SkipProof.make_thm thy (Var (("A", 0), propT))) i) st -val (do_proofs : bool ref) = ref true; - (* Lightweight mode analysis **********************************************) -(* Hack for message from old code generator *) -val message = tracing; - - (**************************************************************************) (* source code from old code generator ************************************) @@ -153,7 +220,8 @@ | _ => false) in check end; -(**** check if a type is an equality type (i.e. doesn't contain fun) ****) +(**** check if a type is an equality type (i.e. doesn't contain fun) + FIXME this is only an approximation ****) fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts | is_eqT _ = true; @@ -165,7 +233,7 @@ | SOME js => enclose "[" "]" (commas (map string_of_int js))) (iss @ [SOME is])); -fun print_modes modes = message ("Inferred modes:\n" ^ +fun print_modes modes = tracing ("Inferred modes:\n" ^ cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map string_of_mode ms)) modes)); @@ -182,6 +250,7 @@ (get_args' is (i+1) ts) in get_args' is 1 ts end +(*FIXME this function should not be named merge... make it local instead*) fun merge xs [] = xs | merge [] ys = ys | merge (x::xs) (y::ys) = if length x >= length y then x::merge xs (y::ys) @@ -197,7 +266,8 @@ fun cprods xss = foldr (map op :: o cprod) [[]] xss; -datatype mode = Mode of (int list option list * int list) * int list * mode option list; +datatype hmode = Mode of mode * int list * hmode option list; (*FIXME don't understand + why there is another mode type!?*) fun modes_of modes t = let @@ -285,11 +355,11 @@ in (p, List.filter (fn m => case find_index (not o check_mode_clause thy param_vs modes m) rs of ~1 => true - | i => (message ("Clause " ^ string_of_int (i+1) ^ " of " ^ + | i => (tracing ("Clause " ^ string_of_int (i+1) ^ " of " ^ p ^ " violates mode " ^ string_of_mode m); false)) ms) end; -fun fixp f (x : (string * (int list option list * int list) list) list) = +fun fixp f (x : (string * mode list) list) = let val y = f x in if x = y then x else fixp f y end; @@ -306,66 +376,6 @@ (*****************************************************************************************) (**** term construction ****) -fun mk_eq (x, xs) = - let fun mk_eqs _ [] = [] - | mk_eqs a (b::cs) = - HOLogic.mk_eq (Free (a, fastype_of b), b) :: mk_eqs a cs - in mk_eqs x xs end; - -fun mk_tuple [] = HOLogic.unit - | mk_tuple ts = foldr1 HOLogic.mk_prod ts; - -fun dest_tuple (Const (@{const_name Product_Type.Unity}, _)) = [] - | dest_tuple (Const (@{const_name Pair}, _) $ t1 $ t2) = t1 :: (dest_tuple t2) - | dest_tuple t = [t] - -fun mk_tupleT [] = HOLogic.unitT - | mk_tupleT Ts = foldr1 HOLogic.mk_prodT Ts; - -fun mk_pred_enumT T = Type ("Predicate.pred", [T]) - -fun dest_pred_enumT (Type ("Predicate.pred", [T])) = T - | dest_pred_enumT T = raise TYPE ("dest_pred_enumT", [T], []); - -fun mk_single t = - let val T = fastype_of t - in Const(@{const_name Predicate.single}, T --> mk_pred_enumT T) $ t end; - -fun mk_empty T = Const (@{const_name Orderings.bot}, mk_pred_enumT T); - -fun mk_if_predenum cond = Const (@{const_name Predicate.if_pred}, - HOLogic.boolT --> mk_pred_enumT HOLogic.unitT) - $ cond - -fun mk_not_pred t = let val T = mk_pred_enumT HOLogic.unitT - in Const (@{const_name Predicate.not_pred}, T --> T) $ t end - -fun mk_bind (x, f) = - let val T as Type ("fun", [_, U]) = fastype_of f - in - Const (@{const_name Predicate.bind}, fastype_of x --> T --> U) $ x $ f - end; - -fun mk_Enum f = - let val T as Type ("fun", [T', _]) = fastype_of f - in - Const (@{const_name Predicate.Pred}, T --> mk_pred_enumT T') $ f - end; - -fun mk_Eval (f, x) = - let val T = fastype_of x - in - Const (@{const_name Predicate.eval}, mk_pred_enumT T --> T --> HOLogic.boolT) $ f $ x - end; - -fun mk_Eval' f = - let val T = fastype_of f - in - Const (@{const_name Predicate.eval}, T --> dest_pred_enumT T --> HOLogic.boolT) $ f - end; - -val mk_sup = HOLogic.mk_binop @{const_name sup}; - (* for simple modes (e.g. parameters) only: better call it param_funT *) (* or even better: remove it and only use funT'_of - some modifications to funT'_of necessary *) fun funT_of T NONE = T @@ -424,13 +434,16 @@ (v', mk_empty U')])) end; -fun modename thy name mode = let +fun modename_of thy name mode = let val v = (PredModetab.lookup (#names (IndCodegenData.get thy)) (name, mode)) - in if (is_some v) then the v - else error ("fun modename - definition not found: name: " ^ name ^ " mode: " ^ (makestring mode)) + in if (is_some v) then the v (*FIXME use case here*) + else error ("fun modename_of - definition not found: name: " ^ name ^ " mode: " ^ (makestring mode)) end -(* function can be removed *) +fun modes_of thy = + these o Symtab.lookup ((#modes o IndCodegenData.get) thy); + +(*FIXME function can be removed*) fun mk_funcomp f t = let val names = Term.add_free_names t []; @@ -449,7 +462,7 @@ val f' = case f of Const (name, T) => if AList.defined op = modes name then - Const (modename thy name (iss, is'), funT'_of (iss, is') T) + Const (modename_of thy name (iss, is'), funT'_of (iss, is') T) else error "compile param: Not an inductive predicate with correct mode" | Free (name, T) => Free (name, funT_of T (SOME is')) in list_comb (f', params' @ args') end @@ -463,7 +476,7 @@ val (Ts, Us) = get_args is (curry Library.drop (length ms) (fst (strip_type T))) val params' = map (compile_param thy modes) (ms ~~ params) - val mode_id = modename thy name mode + val mode_id = modename_of thy name mode in list_comb (Const (mode_id, ((map fastype_of params') @ Ts) ---> mk_pred_enumT (mk_tupleT Us)), params') end @@ -556,7 +569,7 @@ val cl_ts = map (fn cl => compile_clause thy all_vs param_vs modes mode cl (mk_tuple xs)) cls; - val mode_id = modename thy s mode + val mode_id = modename_of thy s mode in HOLogic.mk_Trueprop (HOLogic.mk_eq (list_comb (Const (mode_id, (Ts1' @ Us1) ---> @@ -591,7 +604,7 @@ fold Term.add_consts intrs [] |> map fst |> filter_out (member (op =) preds) |> filter (is_ind_pred thy) -fun print_arities arities = message ("Arities:\n" ^ +fun print_arities arities = tracing ("Arities:\n" ^ cat_lines (map (fn (s, (ks, k)) => s ^ ": " ^ space_implode " -> " (map (fn NONE => "X" | SOME k' => string_of_int k') @@ -691,10 +704,10 @@ (* Proving equivalence of term *) -fun intro_rule thy pred mode = modename thy pred mode +fun intro_rule thy pred mode = modename_of thy pred mode |> Symtab.lookup (#function_intros (IndCodegenData.get thy)) |> the -fun elim_rule thy pred mode = modename thy pred mode +fun elim_rule thy pred mode = modename_of thy pred mode |> Symtab.lookup (#function_elims (IndCodegenData.get thy)) |> the fun pred_intros thy predname = let @@ -711,7 +724,7 @@ end fun function_definition thy pred mode = - modename thy pred mode |> Symtab.lookup (#function_defs (IndCodegenData.get thy)) |> the + modename_of thy pred mode |> Symtab.lookup (#function_defs (IndCodegenData.get thy)) |> the fun is_Type (Type _) = true | is_Type _ = false @@ -973,7 +986,7 @@ in nth (#elims (snd ind_result)) index end) fun prove_one_direction thy all_vs param_vs modes clauses ((pred, T), mode) = let - val elim_rule = the (Symtab.lookup (#function_elims (IndCodegenData.get thy)) (modename thy pred mode)) + val elim_rule = the (Symtab.lookup (#function_elims (IndCodegenData.get thy)) (modename_of thy pred mode)) (* val ind_result = InductivePackage.the_inductive (ProofContext.init thy) pred val index = find_index (fn s => s = pred) (#names (fst ind_result)) val (_, T) = dest_Const (nth (#preds (snd ind_result)) index) *) @@ -1225,7 +1238,7 @@ (* main function *********************************************************************) (*************************************************************************************) -fun create_def_equation' ind_name (mode : (int list option list * int list) option) thy = +fun create_def_equation' ind_name (mode : mode option) thy = let val _ = tracing ("starting create_def_equation' with " ^ ind_name) val (prednames, preds) = @@ -1249,6 +1262,7 @@ val _ = tracing ("calling preds: " ^ makestring name_of_calls) val _ = tracing "starting recursive compilations" fun rec_call name thy = + (*FIXME use member instead of infix mem*) if not (name mem (Symtab.keys (#modes (IndCodegenData.get thy)))) then create_def_equation name thy else thy val thy'' = fold rec_call name_of_calls thy' diff -r 644d18da3c77 -r 7c871a9cf6f4 src/Provers/Arith/cancel_div_mod.ML --- a/src/Provers/Arith/cancel_div_mod.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/Provers/Arith/cancel_div_mod.ML Mon Apr 27 07:26:17 2009 -0700 @@ -69,7 +69,7 @@ fun cancel ss t pq = let val teqt' = Data.prove_eq_sums ss (t, rearrange t pq) - in hd(Data.div_mod_eqs RL [teqt' RS transitive_thm]) end; + in hd (Data.div_mod_eqs RL [teqt' RS transitive_thm]) end; fun proc ss t = let val (divs,mods) = coll_div_mod t ([],[]) diff -r 644d18da3c77 -r 7c871a9cf6f4 src/Pure/IsaMakefile --- a/src/Pure/IsaMakefile Wed Apr 22 11:00:25 2009 -0700 +++ b/src/Pure/IsaMakefile Mon Apr 27 07:26:17 2009 -0700 @@ -40,9 +40,8 @@ Pure: $(OUT)/Pure -$(OUT)/Pure: $(BOOTSTRAP_FILES) ../Tools/auto_solve.ML \ - ../Tools/quickcheck.ML Concurrent/ROOT.ML Concurrent/future.ML \ - Concurrent/mailbox.ML Concurrent/par_list.ML \ +$(OUT)/Pure: $(BOOTSTRAP_FILES) Concurrent/ROOT.ML \ + Concurrent/future.ML Concurrent/mailbox.ML Concurrent/par_list.ML \ Concurrent/par_list_dummy.ML Concurrent/simple_thread.ML \ Concurrent/synchronized.ML Concurrent/task_queue.ML General/ROOT.ML \ General/alist.ML General/antiquote.ML General/balanced_tree.ML \ diff -r 644d18da3c77 -r 7c871a9cf6f4 src/Pure/Isar/code.ML --- a/src/Pure/Isar/code.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/Pure/Isar/code.ML Mon Apr 27 07:26:17 2009 -0700 @@ -29,8 +29,6 @@ val add_undefined: string -> theory -> theory val purge_data: theory -> theory - val coregular_algebra: theory -> Sorts.algebra - val operational_algebra: theory -> (sort -> sort) * Sorts.algebra val these_eqns: theory -> string -> (thm * bool) list val these_raw_eqns: theory -> string -> (thm * bool) list val get_datatype: theory -> string -> ((string * sort) list * (string * typ list) list) @@ -469,39 +467,6 @@ fun mk_default_eqn thy = Code_Unit.try_thm (check_linear o Code_Unit.mk_eqn thy); -(** operational sort algebra and class discipline **) - -local - -fun arity_constraints thy algebra (class, tyco) = - let - val base_constraints = Sorts.mg_domain algebra tyco [class]; - val classparam_constraints = Sorts.complete_sort algebra [class] - |> maps (map fst o these o try (#params o AxClass.get_info thy)) - |> map_filter (fn c => try (AxClass.param_of_inst thy) (c, tyco)) - |> maps (map fst o get_eqns thy) - |> map (map (snd o dest_TVar) o Sign.const_typargs thy o Code_Unit.const_typ_eqn); - val inter_sorts = map2 (curry (Sorts.inter_sort algebra)); - in fold inter_sorts classparam_constraints base_constraints end; - -fun retrieve_algebra thy operational = - Sorts.subalgebra (Syntax.pp_global thy) operational - (SOME o arity_constraints thy (Sign.classes_of thy)) - (Sign.classes_of thy); - -in - -fun coregular_algebra thy = retrieve_algebra thy (K true) |> snd; -fun operational_algebra thy = - let - fun add_iff_operational class = - can (AxClass.get_info thy) class ? cons class; - val operational_classes = fold add_iff_operational (Sign.all_classes thy) [] - in retrieve_algebra thy (member (op =) operational_classes) end; - -end; (*local*) - - (** interfaces and attributes **) fun delete_force msg key xs = diff -r 644d18da3c77 -r 7c871a9cf6f4 src/Pure/ProofGeneral/ROOT.ML --- a/src/Pure/ProofGeneral/ROOT.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/Pure/ProofGeneral/ROOT.ML Mon Apr 27 07:26:17 2009 -0700 @@ -14,11 +14,7 @@ use "pgip_isabelle.ML"; -(use - |> setmp Proofterm.proofs 1 - |> setmp quick_and_dirty true - |> setmp auto_quickcheck true - |> setmp auto_solve true) "preferences.ML"; +use "preferences.ML"; use "pgip_parser.ML"; diff -r 644d18da3c77 -r 7c871a9cf6f4 src/Pure/ProofGeneral/preferences.ML --- a/src/Pure/ProofGeneral/preferences.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/Pure/ProofGeneral/preferences.ML Mon Apr 27 07:26:17 2009 -0700 @@ -6,6 +6,10 @@ signature PREFERENCES = sig + val category_display: string + val category_advanced_display: string + val category_tracing: string + val category_proof: string type preference = {name: string, descr: string, @@ -29,6 +33,14 @@ structure Preferences: PREFERENCES = struct +(* categories *) + +val category_display = "Display"; +val category_advanced_display = "Advanced Display"; +val category_tracing = "Tracing"; +val category_proof = "Proof" + + (* preferences and preference tables *) type preference = @@ -66,11 +78,11 @@ (* preferences of Pure *) -val proof_pref = +val proof_pref = setmp Proofterm.proofs 1 (fn () => let fun get () = PgipTypes.bool_to_pgstring (! Proofterm.proofs >= 2); fun set s = Proofterm.proofs := (if PgipTypes.read_pgipbool s then 2 else 1); - in mkpref get set PgipTypes.Pgipbool "full-proofs" "Record full proof objects internally" end; + in mkpref get set PgipTypes.Pgipbool "full-proofs" "Record full proof objects internally" end) (); val thm_depsN = "thm_deps"; val thm_deps_pref = @@ -145,24 +157,13 @@ bool_pref Toplevel.debug "debugging" "Whether to enable debugging.", - bool_pref Quickcheck.auto - "auto-quickcheck" - "Whether to enable quickcheck automatically.", - nat_pref Quickcheck.auto_time_limit - "auto-quickcheck-time-limit" - "Time limit for automatic quickcheck (in milliseconds).", - bool_pref AutoSolve.auto - "auto-solve" - "Try to solve newly declared lemmas with existing theorems.", - nat_pref AutoSolve.auto_time_limit - "auto-solve-time-limit" - "Time limit for seeking automatic solutions (in milliseconds).", thm_deps_pref]; val proof_preferences = - [bool_pref quick_and_dirty - "quick-and-dirty" - "Take a few short cuts", + [setmp quick_and_dirty true (fn () => + bool_pref quick_and_dirty + "quick-and-dirty" + "Take a few short cuts") (), bool_pref Toplevel.skip_proofs "skip-proofs" "Skip over proofs (interactive-only)", @@ -175,10 +176,10 @@ "Check proofs in parallel"]; val pure_preferences = - [("Display", display_preferences), - ("Advanced Display", advanced_display_preferences), - ("Tracing", tracing_preferences), - ("Proof", proof_preferences)]; + [(category_display, display_preferences), + (category_advanced_display, advanced_display_preferences), + (category_tracing, tracing_preferences), + (category_proof, proof_preferences)]; (* table of categories and preferences; names must be unique *) @@ -203,6 +204,6 @@ else if exists (fn {name, ...} => name = #name pref) prefs then (warning ("Preference already exists: " ^ quote (#name pref)); (cat, prefs)) - else (cat, pref :: prefs)); + else (cat, prefs @ [pref])); end; diff -r 644d18da3c77 -r 7c871a9cf6f4 src/Pure/Tools/ROOT.ML --- a/src/Pure/Tools/ROOT.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/Pure/Tools/ROOT.ML Mon Apr 27 07:26:17 2009 -0700 @@ -1,16 +1,9 @@ -(* Title: Pure/Tools/ROOT.ML - -Miscellaneous tools and packages for Pure Isabelle. -*) +(* Miscellaneous tools and packages for Pure Isabelle *) use "named_thms.ML"; -(*basic XML support*) use "xml_syntax.ML"; use "find_theorems.ML"; use "find_consts.ML"; -(*quickcheck/autosolve needed here because of pg preferences*) -use "../../Tools/quickcheck.ML"; -use "../../Tools/auto_solve.ML"; diff -r 644d18da3c77 -r 7c871a9cf6f4 src/Pure/axclass.ML --- a/src/Pure/axclass.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/Pure/axclass.ML Mon Apr 27 07:26:17 2009 -0700 @@ -286,74 +286,6 @@ handle TYPE (msg, _, _) => error msg; -(* primitive rules *) - -fun add_classrel th thy = - let - fun err () = raise THM ("add_classrel: malformed class relation", 0, [th]); - val prop = Thm.plain_prop_of (Thm.transfer thy th); - val rel = Logic.dest_classrel prop handle TERM _ => err (); - val (c1, c2) = cert_classrel thy rel handle TYPE _ => err (); - in - thy - |> Sign.primitive_classrel (c1, c2) - |> put_classrel ((c1, c2), Thm.close_derivation (Drule.unconstrainTs th)) - |> perhaps complete_arities - end; - -fun add_arity th thy = - let - fun err () = raise THM ("add_arity: malformed type arity", 0, [th]); - val prop = Thm.plain_prop_of (Thm.transfer thy th); - val (t, Ss, c) = Logic.dest_arity prop handle TERM _ => err (); - val _ = map (Sign.certify_sort thy) Ss = Ss orelse err (); - val _ = case filter_out (fn c => can (get_inst_param thy) (c, t)) (params_of thy c) - of [] => () - | cs => Output.legacy_feature - ("Missing specifications for overloaded parameters " ^ commas_quote cs) - val th' = Drule.unconstrainTs th; - in - thy - |> Sign.primitive_arity (t, Ss, [c]) - |> put_arity ((t, Ss, c), th') - end; - - -(* tactical proofs *) - -fun prove_classrel raw_rel tac thy = - let - val ctxt = ProofContext.init thy; - val (c1, c2) = cert_classrel thy raw_rel; - val th = Goal.prove ctxt [] [] (Logic.mk_classrel (c1, c2)) (K tac) handle ERROR msg => - cat_error msg ("The error(s) above occurred while trying to prove class relation " ^ - quote (Syntax.string_of_classrel ctxt [c1, c2])); - in - thy - |> PureThy.add_thms [((Binding.name - (prefix classrel_prefix (Logic.name_classrel (c1, c2))), th), [])] - |-> (fn [th'] => add_classrel th') - end; - -fun prove_arity raw_arity tac thy = - let - val ctxt = ProofContext.init thy; - val arity = Sign.cert_arity thy raw_arity; - val names = map (prefix arity_prefix) (Logic.name_arities arity); - val props = Logic.mk_arities arity; - val ths = Goal.prove_multi ctxt [] [] props - (fn _ => Goal.precise_conjunction_tac (length props) 1 THEN tac) handle ERROR msg => - cat_error msg ("The error(s) above occurred while trying to prove type arity " ^ - quote (Syntax.string_of_arity ctxt arity)); - in - thy - |> PureThy.add_thms (map (rpair []) (map Binding.name names ~~ ths)) - |-> fold add_arity - end; - - -(* instance parameters and overloaded definitions *) - (* declaration and definition of instances of overloaded constants *) fun declare_overloaded (c, T) thy = @@ -398,6 +330,74 @@ end; +(* primitive rules *) + +fun add_classrel th thy = + let + fun err () = raise THM ("add_classrel: malformed class relation", 0, [th]); + val prop = Thm.plain_prop_of (Thm.transfer thy th); + val rel = Logic.dest_classrel prop handle TERM _ => err (); + val (c1, c2) = cert_classrel thy rel handle TYPE _ => err (); + in + thy + |> Sign.primitive_classrel (c1, c2) + |> put_classrel ((c1, c2), Thm.close_derivation (Drule.unconstrainTs th)) + |> perhaps complete_arities + end; + +fun add_arity th thy = + let + fun err () = raise THM ("add_arity: malformed type arity", 0, [th]); + val prop = Thm.plain_prop_of (Thm.transfer thy th); + val (t, Ss, c) = Logic.dest_arity prop handle TERM _ => err (); + val T = Type (t, map TFree (Name.names Name.context Name.aT Ss)); + val missing_params = Sign.complete_sort thy [c] + |> maps (these o Option.map #params o try (get_info thy)) + |> filter_out (fn (const, _) => can (get_inst_param thy) (const, t)) + |> (map o apsnd o map_atyps) (K T); + val _ = map (Sign.certify_sort thy) Ss = Ss orelse err (); + val th' = Drule.unconstrainTs th; + in + thy + |> fold (snd oo declare_overloaded) missing_params + |> Sign.primitive_arity (t, Ss, [c]) + |> put_arity ((t, Ss, c), th') + end; + + +(* tactical proofs *) + +fun prove_classrel raw_rel tac thy = + let + val ctxt = ProofContext.init thy; + val (c1, c2) = cert_classrel thy raw_rel; + val th = Goal.prove ctxt [] [] (Logic.mk_classrel (c1, c2)) (K tac) handle ERROR msg => + cat_error msg ("The error(s) above occurred while trying to prove class relation " ^ + quote (Syntax.string_of_classrel ctxt [c1, c2])); + in + thy + |> PureThy.add_thms [((Binding.name + (prefix classrel_prefix (Logic.name_classrel (c1, c2))), th), [])] + |-> (fn [th'] => add_classrel th') + end; + +fun prove_arity raw_arity tac thy = + let + val ctxt = ProofContext.init thy; + val arity = Sign.cert_arity thy raw_arity; + val names = map (prefix arity_prefix) (Logic.name_arities arity); + val props = Logic.mk_arities arity; + val ths = Goal.prove_multi ctxt [] [] props + (fn _ => Goal.precise_conjunction_tac (length props) 1 THEN tac) handle ERROR msg => + cat_error msg ("The error(s) above occurred while trying to prove type arity " ^ + quote (Syntax.string_of_arity ctxt arity)); + in + thy + |> PureThy.add_thms (map (rpair []) (map Binding.name names ~~ ths)) + |-> fold add_arity + end; + + (** class definitions **) diff -r 644d18da3c77 -r 7c871a9cf6f4 src/Tools/Code_Generator.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Tools/Code_Generator.thy Mon Apr 27 07:26:17 2009 -0700 @@ -0,0 +1,28 @@ +(* Title: Tools/Code_Generator.thy + Author: Florian Haftmann, TU Muenchen +*) + +header {* Loading the code generator modules *} + +theory Code_Generator +imports Pure +uses + "~~/src/Tools/value.ML" + "~~/src/Tools/quickcheck.ML" + "~~/src/Tools/code/code_name.ML" + "~~/src/Tools/code/code_wellsorted.ML" + "~~/src/Tools/code/code_thingol.ML" + "~~/src/Tools/code/code_printer.ML" + "~~/src/Tools/code/code_target.ML" + "~~/src/Tools/code/code_ml.ML" + "~~/src/Tools/code/code_haskell.ML" + "~~/src/Tools/nbe.ML" +begin + +setup {* + Code_ML.setup + #> Code_Haskell.setup + #> Nbe.setup +*} + +end \ No newline at end of file diff -r 644d18da3c77 -r 7c871a9cf6f4 src/Tools/auto_solve.ML --- a/src/Tools/auto_solve.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/Tools/auto_solve.ML Mon Apr 27 07:26:17 2009 -0700 @@ -14,18 +14,34 @@ val auto : bool ref val auto_time_limit : int ref val limit : int ref - - val seek_solution : bool -> Proof.state -> Proof.state end; structure AutoSolve : AUTO_SOLVE = struct +(* preferences *) + val auto = ref false; val auto_time_limit = ref 2500; val limit = ref 5; -fun seek_solution int state = +val _ = + ProofGeneralPgip.add_preference Preferences.category_tracing + (setmp auto true (fn () => + Preferences.bool_pref auto + "auto-solve" + "Try to solve newly declared lemmas with existing theorems.") ()); + +val _ = + ProofGeneralPgip.add_preference Preferences.category_tracing + (Preferences.nat_pref auto_time_limit + "auto-solve-time-limit" + "Time limit for seeking automatic solutions (in milliseconds)."); + + +(* hook *) + +val _ = Context.>> (Specification.add_theorem_hook (fn int => fn state => let val ctxt = Proof.context_of state; @@ -76,12 +92,10 @@ if int andalso ! auto andalso not (! Toplevel.quiet) then go () else state - end; + end)); end; -val _ = Context.>> (Specification.add_theorem_hook AutoSolve.seek_solution); - val auto_solve = AutoSolve.auto; val auto_solve_time_limit = AutoSolve.auto_time_limit; diff -r 644d18da3c77 -r 7c871a9cf6f4 src/Tools/code/code_funcgr.ML --- a/src/Tools/code/code_funcgr.ML Wed Apr 22 11:00:25 2009 -0700 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,335 +0,0 @@ -(* Title: Tools/code/code_funcgr.ML - Author: Florian Haftmann, TU Muenchen - -Retrieving, normalizing and structuring code equations in graph -with explicit dependencies. - -Legacy. To be replaced by Tools/code/code_wellsorted.ML -*) - -signature CODE_WELLSORTED = -sig - type T - val eqns: T -> string -> (thm * bool) list - val typ: T -> string -> (string * sort) list * typ - val all: T -> string list - val pretty: theory -> T -> Pretty.T - val make: theory -> string list - -> ((sort -> sort) * Sorts.algebra) * T - val eval_conv: theory - -> (term -> term * (((sort -> sort) * Sorts.algebra) -> T -> thm)) -> cterm -> thm - val eval_term: theory - -> (term -> term * (((sort -> sort) * Sorts.algebra) -> T -> 'a)) -> term -> 'a - val timing: bool ref -end - -structure Code_Wellsorted : CODE_WELLSORTED = -struct - -(** the graph type **) - -type T = (((string * sort) list * typ) * (thm * bool) list) Graph.T; - -fun eqns funcgr = - these o Option.map snd o try (Graph.get_node funcgr); - -fun typ funcgr = - fst o Graph.get_node funcgr; - -fun all funcgr = Graph.keys funcgr; - -fun pretty thy funcgr = - AList.make (snd o Graph.get_node funcgr) (Graph.keys funcgr) - |> (map o apfst) (Code_Unit.string_of_const thy) - |> sort (string_ord o pairself fst) - |> map (fn (s, thms) => - (Pretty.block o Pretty.fbreaks) ( - Pretty.str s - :: map (Display.pretty_thm o fst) thms - )) - |> Pretty.chunks; - - -(** generic combinators **) - -fun fold_consts f thms = - thms - |> maps (op :: o swap o apfst (snd o strip_comb) o Logic.dest_equals o Thm.plain_prop_of) - |> (fold o fold_aterms) (fn Const c => f c | _ => I); - -fun consts_of (const, []) = [] - | consts_of (const, thms as _ :: _) = - let - fun the_const (c, _) = if c = const then I else insert (op =) c - in fold_consts the_const (map fst thms) [] end; - -fun insts_of thy algebra tys sorts = - let - fun class_relation (x, _) _ = x; - fun type_constructor tyco xs class = - (tyco, class) :: (maps o maps) fst xs; - fun type_variable (TVar (_, sort)) = map (pair []) sort - | type_variable (TFree (_, sort)) = map (pair []) sort; - fun of_sort_deriv ty sort = - Sorts.of_sort_derivation (Syntax.pp_global thy) algebra - { class_relation = class_relation, type_constructor = type_constructor, - type_variable = type_variable } - (ty, sort) handle Sorts.CLASS_ERROR _ => [] (*permissive!*) - in (flat o flat) (map2 of_sort_deriv tys sorts) end; - -fun meets_of thy algebra = - let - fun meet_of ty sort tab = - Sorts.meet_sort algebra (ty, sort) tab - handle Sorts.CLASS_ERROR _ => tab (*permissive!*); - in fold2 meet_of end; - - -(** graph algorithm **) - -val timing = ref false; - -local - -fun resort_thms thy algebra typ_of thms = - let - val cs = fold_consts (insert (op =)) thms []; - fun meets (c, ty) = case typ_of c - of SOME (vs, _) => - meets_of thy algebra (Sign.const_typargs thy (c, ty)) (map snd vs) - | NONE => I; - val tab = fold meets cs Vartab.empty; - in map (Code_Unit.inst_thm thy tab) thms end; - -fun resort_eqnss thy algebra funcgr = - let - val typ_funcgr = try (fst o Graph.get_node funcgr); - val resort_dep = (apsnd o burrow_fst) (resort_thms thy algebra typ_funcgr); - fun resort_rec typ_of (c, []) = (true, (c, [])) - | resort_rec typ_of (c, thms as (thm, _) :: _) = if is_some (AxClass.inst_of_param thy c) - then (true, (c, thms)) - else let - val (_, (vs, ty)) = Code_Unit.head_eqn thy thm; - val thms' as (thm', _) :: _ = burrow_fst (resort_thms thy algebra typ_of) thms - val (_, (vs', ty')) = Code_Unit.head_eqn thy thm'; (*FIXME simplify check*) - in (Sign.typ_equiv thy (ty, ty'), (c, thms')) end; - fun resort_recs eqnss = - let - fun typ_of c = case these (AList.lookup (op =) eqnss c) - of (thm, _) :: _ => (SOME o snd o Code_Unit.head_eqn thy) thm - | [] => NONE; - val (unchangeds, eqnss') = split_list (map (resort_rec typ_of) eqnss); - val unchanged = fold (fn x => fn y => x andalso y) unchangeds true; - in (unchanged, eqnss') end; - fun resort_rec_until eqnss = - let - val (unchanged, eqnss') = resort_recs eqnss; - in if unchanged then eqnss' else resort_rec_until eqnss' end; - in map resort_dep #> resort_rec_until end; - -fun instances_of thy algebra insts = - let - val thy_classes = (#classes o Sorts.rep_algebra o Sign.classes_of) thy; - fun all_classparams tyco class = - these (try (#params o AxClass.get_info thy) class) - |> map_filter (fn (c, _) => try (AxClass.param_of_inst thy) (c, tyco)) - in - Symtab.empty - |> fold (fn (tyco, class) => - Symtab.map_default (tyco, []) (insert (op =) class)) insts - |> (fn tab => Symtab.fold (fn (tyco, classes) => append (maps (all_classparams tyco) - (Graph.all_succs thy_classes classes))) tab []) - end; - -fun instances_of_consts thy algebra funcgr consts = - let - fun inst (cexpr as (c, ty)) = insts_of thy algebra - (Sign.const_typargs thy (c, ty)) ((map snd o fst) (typ funcgr c)); - in - [] - |> fold (fold (insert (op =)) o inst) consts - |> instances_of thy algebra - end; - -fun ensure_const' thy algebra funcgr const auxgr = - if can (Graph.get_node funcgr) const - then (NONE, auxgr) - else if can (Graph.get_node auxgr) const - then (SOME const, auxgr) - else if is_some (Code.get_datatype_of_constr thy const) then - auxgr - |> Graph.new_node (const, []) - |> pair (SOME const) - else let - val thms = Code.these_eqns thy const - |> burrow_fst (Code_Unit.norm_args thy) - |> burrow_fst (Code_Unit.norm_varnames thy Code_Name.purify_tvar Code_Name.purify_var); - val rhs = consts_of (const, thms); - in - auxgr - |> Graph.new_node (const, thms) - |> fold_map (ensure_const thy algebra funcgr) rhs - |-> (fn rhs' => fold (fn SOME const' => Graph.add_edge (const, const') - | NONE => I) rhs') - |> pair (SOME const) - end -and ensure_const thy algebra funcgr const = - let - val timeap = if !timing - then Output.timeap_msg ("time for " ^ Code_Unit.string_of_const thy const) - else I; - in timeap (ensure_const' thy algebra funcgr const) end; - -fun merge_eqnss thy algebra raw_eqnss funcgr = - let - val eqnss = raw_eqnss - |> resort_eqnss thy algebra funcgr - |> filter_out (can (Graph.get_node funcgr) o fst); - fun typ_eqn c [] = Code.default_typscheme thy c - | typ_eqn c (thms as (thm, _) :: _) = (snd o Code_Unit.head_eqn thy) thm; - fun add_eqns (const, thms) = - Graph.new_node (const, (typ_eqn const thms, thms)); - fun add_deps (eqns as (const, thms)) funcgr = - let - val deps = consts_of eqns; - val insts = instances_of_consts thy algebra funcgr - (fold_consts (insert (op =)) (map fst thms) []); - in - funcgr - |> ensure_consts thy algebra insts - |> fold (curry Graph.add_edge const) deps - |> fold (curry Graph.add_edge const) insts - end; - in - funcgr - |> fold add_eqns eqnss - |> fold add_deps eqnss - end -and ensure_consts thy algebra cs funcgr = - let - val auxgr = Graph.empty - |> fold (snd oo ensure_const thy algebra funcgr) cs; - in - funcgr - |> fold (merge_eqnss thy algebra) - (map (AList.make (Graph.get_node auxgr)) - (rev (Graph.strong_conn auxgr))) - end; - -in - -(** retrieval interfaces **) - -val ensure_consts = ensure_consts; - -fun proto_eval thy cterm_of evaluator_lift evaluator proto_ct funcgr = - let - val ct = cterm_of proto_ct; - val _ = Sign.no_vars (Syntax.pp_global thy) (Thm.term_of ct); - val _ = Term.fold_types (Type.no_tvars #> K I) (Thm.term_of ct) (); - fun consts_of t = - fold_aterms (fn Const c_ty => cons c_ty | _ => I) t []; - val algebra = Code.coregular_algebra thy; - val thm = Code.preprocess_conv thy ct; - val ct' = Thm.rhs_of thm; - val t' = Thm.term_of ct'; - val consts = map fst (consts_of t'); - val funcgr' = ensure_consts thy algebra consts funcgr; - val (t'', evaluator_funcgr) = evaluator t'; - val consts' = consts_of t''; - val dicts = instances_of_consts thy algebra funcgr' consts'; - val funcgr'' = ensure_consts thy algebra dicts funcgr'; - in (evaluator_lift (evaluator_funcgr (Code.operational_algebra thy)) thm funcgr'', funcgr'') end; - -fun proto_eval_conv thy = - let - fun evaluator_lift evaluator thm1 funcgr = - let - val thm2 = evaluator funcgr; - val thm3 = Code.postprocess_conv thy (Thm.rhs_of thm2); - in - Thm.transitive thm1 (Thm.transitive thm2 thm3) handle THM _ => - error ("could not construct evaluation proof:\n" - ^ (cat_lines o map Display.string_of_thm) [thm1, thm2, thm3]) - end; - in proto_eval thy I evaluator_lift end; - -fun proto_eval_term thy = - let - fun evaluator_lift evaluator _ funcgr = evaluator funcgr; - in proto_eval thy (Thm.cterm_of thy) evaluator_lift end; - -end; (*local*) - -structure Funcgr = CodeDataFun -( - type T = T; - val empty = Graph.empty; - fun purge _ cs funcgr = - Graph.del_nodes ((Graph.all_preds funcgr - o filter (can (Graph.get_node funcgr))) cs) funcgr; -); - -fun make thy = - pair (Code.operational_algebra thy) - o Funcgr.change thy o ensure_consts thy (Code.coregular_algebra thy); - -fun eval_conv thy f = - fst o Funcgr.change_yield thy o proto_eval_conv thy f; - -fun eval_term thy f = - fst o Funcgr.change_yield thy o proto_eval_term thy f; - - -(** diagnostic commands **) - -fun code_depgr thy consts = - let - val (_, gr) = make thy consts; - val select = Graph.all_succs gr consts; - in - gr - |> not (null consts) ? Graph.subgraph (member (op =) select) - |> Graph.map_nodes ((apsnd o map o apfst) (AxClass.overload thy)) - end; - -fun code_thms thy = Pretty.writeln o pretty thy o code_depgr thy; - -fun code_deps thy consts = - let - val gr = code_depgr thy consts; - fun mk_entry (const, (_, (_, parents))) = - let - val name = Code_Unit.string_of_const thy const; - val nameparents = map (Code_Unit.string_of_const thy) parents; - in { name = name, ID = name, dir = "", unfold = true, - path = "", parents = nameparents } - end; - val prgr = Graph.fold ((fn x => fn xs => xs @ [x]) o mk_entry) gr []; - in Present.display_graph prgr end; - -local - -structure P = OuterParse -and K = OuterKeyword - -fun code_thms_cmd thy = code_thms thy o op @ o Code_Name.read_const_exprs thy; -fun code_deps_cmd thy = code_deps thy o op @ o Code_Name.read_const_exprs thy; - -in - -val _ = - OuterSyntax.improper_command "code_thms" "print system of code equations for code" OuterKeyword.diag - (Scan.repeat P.term_group - >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory - o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of))); - -val _ = - OuterSyntax.improper_command "code_deps" "visualize dependencies of code equations for code" OuterKeyword.diag - (Scan.repeat P.term_group - >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory - o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of))); - -end; - -end; (*struct*) diff -r 644d18da3c77 -r 7c871a9cf6f4 src/Tools/code/code_ml.ML --- a/src/Tools/code/code_ml.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/Tools/code/code_ml.ML Mon Apr 27 07:26:17 2009 -0700 @@ -6,8 +6,9 @@ signature CODE_ML = sig - val eval_term: string * (unit -> 'a) option ref - -> theory -> term -> string list -> 'a + val eval: string option -> string * (unit -> 'a) option ref + -> ((term -> term) -> 'a -> 'a) -> theory -> term -> string list -> 'a + val target_Eval: string val setup: theory -> theory end; @@ -22,6 +23,7 @@ val target_SML = "SML"; val target_OCaml = "OCaml"; +val target_Eval = "Eval"; datatype ml_stmt = MLExc of string * int @@ -907,36 +909,38 @@ in (deresolver, nodes) end; fun serialize_ml target compile pr_module pr_stmt raw_module_name labelled_name reserved_names includes raw_module_alias - _ syntax_tyco syntax_const naming program cs destination = + _ syntax_tyco syntax_const naming program stmt_names destination = let val is_cons = Code_Thingol.is_cons program; - val stmt_names = Code_Target.stmt_names_of_destination destination; - val module_name = if null stmt_names then raw_module_name else SOME "Code"; + val present_stmt_names = Code_Target.stmt_names_of_destination destination; + val is_present = not (null present_stmt_names); + val module_name = if is_present then SOME "Code" else raw_module_name; val (deresolver, nodes) = ml_node_of_program labelled_name module_name reserved_names raw_module_alias program; val reserved_names = Code_Printer.make_vars reserved_names; fun pr_node prefix (Dummy _) = NONE - | pr_node prefix (Stmt (_, stmt)) = if null stmt_names orelse - (not o null o filter (member (op =) stmt_names) o stmt_names_of) stmt then SOME + | pr_node prefix (Stmt (_, stmt)) = if is_present andalso + (null o filter (member (op =) present_stmt_names) o stmt_names_of) stmt + then NONE + else SOME (pr_stmt naming labelled_name syntax_tyco syntax_const reserved_names (deresolver prefix) is_cons stmt) - else NONE | pr_node prefix (Module (module_name, (_, nodes))) = separate (str "") ((map_filter (pr_node (prefix @ [module_name]) o Graph.get_node nodes) o rev o flat o Graph.strong_conn) nodes) - |> (if null stmt_names then pr_module module_name else Pretty.chunks) + |> (if is_present then Pretty.chunks else pr_module module_name) |> SOME; - val cs' = (map o try) - (deresolver (if is_some module_name then the_list module_name else [])) cs; + val stmt_names' = (map o try) + (deresolver (if is_some module_name then the_list module_name else [])) stmt_names; val p = Pretty.chunks (separate (str "") (map snd includes @ (map_filter (pr_node [] o Graph.get_node nodes) o rev o flat o Graph.strong_conn) nodes)); in Code_Target.mk_serialization target (case compile of SOME compile => SOME (compile o Code_Target.code_of_pretty) | NONE => NONE) (fn NONE => Code_Target.code_writeln | SOME file => File.write file o Code_Target.code_of_pretty) - (rpair cs' o Code_Target.code_of_pretty) p destination + (rpair stmt_names' o Code_Target.code_of_pretty) p destination end; end; (*local*) @@ -944,20 +948,20 @@ (** ML (system language) code for evaluation and instrumentalization **) -fun ml_code_of thy = Code_Target.serialize_custom thy (target_SML, +fun eval_code_of some_target thy = Code_Target.serialize_custom thy (the_default target_Eval some_target, (fn _ => fn [] => serialize_ml target_SML (SOME (K ())) (K Pretty.chunks) pr_sml_stmt (SOME ""), literals_sml)); (* evaluation *) -fun eval eval'' term_of reff thy ct args = +fun eval some_target reff postproc thy t args = let val ctxt = ProofContext.init thy; - val _ = if null (Term.add_frees (term_of ct) []) then () else error ("Term " - ^ quote (Syntax.string_of_term_global thy (term_of ct)) + val _ = if null (Term.add_frees t []) then () else error ("Term " + ^ quote (Syntax.string_of_term_global thy t) ^ " to be evaluated contains free variables"); - fun eval' naming program ((vs, ty), t) deps = + fun evaluator naming program (((_, (_, ty)), _), t) deps = let val _ = if Code_Thingol.contains_dictvar t then error "Term to be evaluated contains free dictionaries" else (); @@ -966,13 +970,11 @@ |> Graph.new_node (value_name, Code_Thingol.Fun (Term.dummy_patternN, (([], ty), [(([], t), (Drule.dummy_thm, true))]))) |> fold (curry Graph.add_edge value_name) deps; - val (value_code, [SOME value_name']) = ml_code_of thy naming program' [value_name]; + val (value_code, [SOME value_name']) = eval_code_of some_target thy naming program' [value_name]; val sml_code = "let\n" ^ value_code ^ "\nin " ^ value_name' ^ space_implode " " (map (enclose "(" ")") args) ^ " end"; in ML_Context.evaluate ctxt false reff sml_code end; - in eval'' thy (rpair eval') ct end; - -fun eval_term reff = eval Code_Thingol.eval_term I reff; + in Code_Thingol.eval thy I postproc evaluator t end; (* instrumentalization by antiquotation *) @@ -981,42 +983,69 @@ structure CodeAntiqData = ProofDataFun ( - type T = string list * (bool * (string * (string * (string * string) list) lazy)); - fun init _ = ([], (true, ("", Lazy.value ("", [])))); + type T = (string list * string list) * (bool * (string + * (string * ((string * string) list * (string * string) list)) lazy)); + fun init _ = (([], []), (true, ("", Lazy.value ("", ([], []))))); ); val is_first_occ = fst o snd o CodeAntiqData.get; -fun delayed_code thy consts () = +fun delayed_code thy tycos consts () = let val (consts', (naming, program)) = Code_Thingol.consts_program thy consts; - val (ml_code, consts'') = ml_code_of thy naming program consts'; - val const_tab = map2 (fn const => fn NONE => - error ("Constant " ^ (quote o Code_Unit.string_of_const thy) const - ^ "\nhas a user-defined serialization") - | SOME const' => (const, const')) consts consts'' - in (ml_code, const_tab) end; + val tycos' = map (the o Code_Thingol.lookup_tyco naming) tycos; + val (ml_code, target_names) = eval_code_of NONE thy naming program (consts' @ tycos'); + val (consts'', tycos'') = chop (length consts') target_names; + val consts_map = map2 (fn const => fn NONE => + error ("Constant " ^ (quote o Code_Unit.string_of_const thy) const + ^ "\nhas a user-defined serialization") + | SOME const'' => (const, const'')) consts consts'' + val tycos_map = map2 (fn tyco => fn NONE => + error ("Type " ^ (quote o Sign.extern_type thy) tyco + ^ "\nhas a user-defined serialization") + | SOME tyco'' => (tyco, tyco'')) tycos tycos''; + in (ml_code, (tycos_map, consts_map)) end; -fun register_const const ctxt = +fun register_code new_tycos new_consts ctxt = let - val (consts, (_, (struct_name, _))) = CodeAntiqData.get ctxt; - val consts' = insert (op =) const consts; + val ((tycos, consts), (_, (struct_name, _))) = CodeAntiqData.get ctxt; + val tycos' = fold (insert (op =)) new_tycos tycos; + val consts' = fold (insert (op =)) new_consts consts; val (struct_name', ctxt') = if struct_name = "" then ML_Antiquote.variant "Code" ctxt else (struct_name, ctxt); - val acc_code = Lazy.lazy (delayed_code (ProofContext.theory_of ctxt) consts'); - in CodeAntiqData.put (consts', (false, (struct_name', acc_code))) ctxt' end; + val acc_code = Lazy.lazy (delayed_code (ProofContext.theory_of ctxt) tycos' consts'); + in CodeAntiqData.put ((tycos', consts'), (false, (struct_name', acc_code))) ctxt' end; + +fun register_const const = register_code [] [const]; -fun print_code struct_name is_first const ctxt = +fun register_datatype tyco constrs = register_code [tyco] constrs; + +fun print_const const all_struct_name tycos_map consts_map = + (Long_Name.append all_struct_name o the o AList.lookup (op =) consts_map) const; + +fun print_datatype tyco constrs all_struct_name tycos_map consts_map = let - val (consts, (_, (struct_code_name, acc_code))) = CodeAntiqData.get ctxt; - val (raw_ml_code, consts_map) = Lazy.force acc_code; - val const'' = Long_Name.append (Long_Name.append struct_name struct_code_name) - ((the o AList.lookup (op =) consts_map) const); + val upperize = implode o nth_map 0 Symbol.to_ascii_upper o explode; + fun check_base name name'' = + if upperize (Long_Name.base_name name) = upperize name'' + then () else error ("Name as printed " ^ quote name'' + ^ "\ndiffers from logical base name " ^ quote (Long_Name.base_name name) ^ "; sorry."); + val tyco'' = (the o AList.lookup (op =) tycos_map) tyco; + val constrs'' = map (the o AList.lookup (op =) consts_map) constrs; + val _ = check_base tyco tyco''; + val _ = map2 check_base constrs constrs''; + in "datatype " ^ tyco'' ^ " = datatype " ^ Long_Name.append all_struct_name tyco'' end; + +fun print_code struct_name is_first print_it ctxt = + let + val (_, (_, (struct_code_name, acc_code))) = CodeAntiqData.get ctxt; + val (raw_ml_code, (tycos_map, consts_map)) = Lazy.force acc_code; val ml_code = if is_first then "\nstructure " ^ struct_code_name ^ " =\nstruct\n\n" ^ raw_ml_code ^ "\nend;\n\n" else ""; - in (ml_code, const'') end; + val all_struct_name = Long_Name.append struct_name struct_code_name; + in (ml_code, print_it all_struct_name tycos_map consts_map) end; in @@ -1025,7 +1054,19 @@ val const = Code_Unit.check_const (ProofContext.theory_of background) raw_const; val is_first = is_first_occ background; val background' = register_const const background; - in (print_code struct_name is_first const, background') end; + in (print_code struct_name is_first (print_const const), background') end; + +fun ml_code_datatype_antiq (raw_tyco, raw_constrs) {struct_name, background} = + let + val thy = ProofContext.theory_of background; + val tyco = Sign.intern_type thy raw_tyco; + val constrs = map (Code_Unit.check_const thy) raw_constrs; + val constrs' = (map fst o snd o Code.get_datatype thy) tyco; + val _ = if gen_eq_set (op =) (constrs, constrs') then () + else error ("Type " ^ quote tyco ^ ": given constructors diverge from real constructors") + val is_first = is_first_occ background; + val background' = register_datatype tyco constrs background; + in (print_code struct_name is_first (print_datatype tyco constrs), background') end; end; (*local*) @@ -1033,6 +1074,10 @@ (** Isar setup **) val _ = ML_Context.add_antiq "code" (fn _ => Args.term >> ml_code_antiq); +val _ = ML_Context.add_antiq "code_datatype" (fn _ => + (Args.tyname --| Scan.lift (Args.$$$ "=") + -- (Args.term ::: Scan.repeat (Scan.lift (Args.$$$ "|") |-- Args.term))) + >> ml_code_datatype_antiq); fun isar_seri_sml module_name = Code_Target.parse_args (Scan.succeed ()) @@ -1048,6 +1093,7 @@ val setup = Code_Target.add_target (target_SML, (isar_seri_sml, literals_sml)) #> Code_Target.add_target (target_OCaml, (isar_seri_ocaml, literals_ocaml)) + #> Code_Target.extend_target (target_Eval, (target_SML, K I)) #> Code_Target.add_syntax_tyco target_SML "fun" (SOME (2, fn pr_typ => fn fxy => fn [ty1, ty2] => brackify_infix (1, R) fxy [ pr_typ (INFX (1, X)) ty1, diff -r 644d18da3c77 -r 7c871a9cf6f4 src/Tools/code/code_target.ML --- a/src/Tools/code/code_target.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/Tools/code/code_target.ML Mon Apr 27 07:26:17 2009 -0700 @@ -68,7 +68,7 @@ fun compile f = (code_setmp f Compile; ()); fun export f = (code_setmp f Export; ()); fun file p f = (code_setmp f (File p); ()); -fun string cs f = fst (the (code_setmp f (String cs))); +fun string stmts f = fst (the (code_setmp f (String stmts))); fun stmt_names_of_destination (String stmts) = stmts | stmt_names_of_destination _ = []; diff -r 644d18da3c77 -r 7c871a9cf6f4 src/Tools/code/code_thingol.ML --- a/src/Tools/code/code_thingol.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/Tools/code/code_thingol.ML Mon Apr 27 07:26:17 2009 -0700 @@ -83,11 +83,11 @@ val consts_program: theory -> string list -> string list * (naming * program) val cached_program: theory -> naming * program - val eval_conv: theory - -> (term -> term * (naming -> program -> typscheme * iterm -> string list -> thm)) + val eval_conv: theory -> (sort -> sort) + -> (naming -> program -> (((string * sort) list * typscheme) * (string * itype) list) * iterm -> string list -> cterm -> thm) -> cterm -> thm - val eval_term: theory - -> (term -> term * (naming -> program -> typscheme * iterm -> string list -> 'a)) + val eval: theory -> (sort -> sort) -> ((term -> term) -> 'a -> 'a) + -> (naming -> program -> (((string * sort) list * typscheme) * (string * itype) list) * iterm -> string list -> 'a) -> term -> 'a end; @@ -459,7 +459,45 @@ (* translation *) -fun ensure_class thy (algbr as (_, algebra)) funcgr class = +fun ensure_tyco thy algbr funcgr tyco = + let + val stmt_datatype = + let + val (vs, cos) = Code.get_datatype thy tyco; + in + fold_map (translate_tyvar_sort thy algbr funcgr) vs + ##>> fold_map (fn (c, tys) => + ensure_const thy algbr funcgr c + ##>> fold_map (translate_typ thy algbr funcgr) tys) cos + #>> (fn info => Datatype (tyco, info)) + end; + in ensure_stmt lookup_tyco (declare_tyco thy) stmt_datatype tyco end +and ensure_const thy algbr funcgr c = + let + fun stmt_datatypecons tyco = + ensure_tyco thy algbr funcgr tyco + #>> (fn tyco => Datatypecons (c, tyco)); + fun stmt_classparam class = + ensure_class thy algbr funcgr class + #>> (fn class => Classparam (c, class)); + fun stmt_fun ((vs, ty), raw_thms) = + let + val thms = if null (Term.add_tfreesT ty []) orelse (null o fst o strip_type) ty + then raw_thms + else (map o apfst) (Code_Unit.expand_eta thy 1) raw_thms; + in + fold_map (translate_tyvar_sort thy algbr funcgr) vs + ##>> translate_typ thy algbr funcgr ty + ##>> fold_map (translate_eq thy algbr funcgr) thms + #>> (fn info => Fun (c, info)) + end; + val stmt_const = case Code.get_datatype_of_constr thy c + of SOME tyco => stmt_datatypecons tyco + | NONE => (case AxClass.class_of_param thy c + of SOME class => stmt_classparam class + | NONE => stmt_fun (Code_Wellsorted.typ funcgr c, Code_Wellsorted.eqns funcgr c)) + in ensure_stmt lookup_const (declare_const thy) stmt_const c end +and ensure_class thy (algbr as (_, algebra)) funcgr class = let val superclasses = (Sorts.minimize_sort algebra o Sorts.super_classes algebra) class; val cs = #params (AxClass.get_info thy class); @@ -477,65 +515,6 @@ ##>> ensure_class thy algbr funcgr superclass #>> Classrel; in ensure_stmt lookup_classrel (declare_classrel thy) stmt_classrel (subclass, superclass) end -and ensure_tyco thy algbr funcgr tyco = - let - val stmt_datatype = - let - val (vs, cos) = Code.get_datatype thy tyco; - in - fold_map (translate_tyvar_sort thy algbr funcgr) vs - ##>> fold_map (fn (c, tys) => - ensure_const thy algbr funcgr c - ##>> fold_map (translate_typ thy algbr funcgr) tys) cos - #>> (fn info => Datatype (tyco, info)) - end; - in ensure_stmt lookup_tyco (declare_tyco thy) stmt_datatype tyco end -and translate_tyvar_sort thy (algbr as (proj_sort, _)) funcgr (v, sort) = - fold_map (ensure_class thy algbr funcgr) (proj_sort sort) - #>> (fn sort => (unprefix "'" v, sort)) -and translate_typ thy algbr funcgr (TFree (v, _)) = - pair (ITyVar (unprefix "'" v)) - | translate_typ thy algbr funcgr (Type (tyco, tys)) = - ensure_tyco thy algbr funcgr tyco - ##>> fold_map (translate_typ thy algbr funcgr) tys - #>> (fn (tyco, tys) => tyco `%% tys) -and translate_dicts thy (algbr as (proj_sort, algebra)) funcgr thm (ty, sort) = - let - val pp = Syntax.pp_global thy; - datatype typarg = - Global of (class * string) * typarg list list - | 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); - fun type_constructor tyco yss class = - Global ((class, tyco), (map o map) fst yss); - fun type_variable (TFree (v, sort)) = - let - val sort' = proj_sort sort; - in map_index (fn (n, class) => (Local ([], (v, (n, sort'))), class)) sort' end; - val typargs = Sorts.of_sort_derivation pp algebra - {class_relation = class_relation, type_constructor = type_constructor, - type_variable = type_variable} (ty, proj_sort sort) - handle Sorts.CLASS_ERROR e => not_wellsorted thy thm ty sort e; - fun mk_dict (Global (inst, yss)) = - ensure_inst thy algbr funcgr inst - ##>> (fold_map o fold_map) mk_dict yss - #>> (fn (inst, dss) => DictConst (inst, dss)) - | mk_dict (Local (classrels, (v, (k, sort)))) = - fold_map (ensure_classrel thy algbr funcgr) classrels - #>> (fn classrels => DictVar (classrels, (unprefix "'" v, (k, length sort)))) - in fold_map mk_dict typargs end -and translate_eq thy algbr funcgr (thm, linear) = - let - val (args, rhs) = (apfst (snd o strip_comb) o Logic.dest_equals - o Logic.unvarify o prop_of) thm; - in - fold_map (translate_term thy algbr funcgr (SOME thm)) args - ##>> translate_term thy algbr funcgr (SOME thm) rhs - #>> rpair (thm, linear) - end and ensure_inst thy (algbr as (_, algebra)) funcgr (class, tyco) = let val superclasses = (Sorts.minimize_sort algebra o Sorts.super_classes algebra) class; @@ -572,31 +551,12 @@ #>> (fn ((((class, tyco), arity), superarities), classparams) => Classinst ((class, (tyco, arity)), (superarities, classparams))); in ensure_stmt lookup_instance (declare_instance thy) stmt_inst (class, tyco) end -and ensure_const thy algbr funcgr c = - let - fun stmt_datatypecons tyco = +and translate_typ thy algbr funcgr (TFree (v, _)) = + pair (ITyVar (unprefix "'" v)) + | translate_typ thy algbr funcgr (Type (tyco, tys)) = ensure_tyco thy algbr funcgr tyco - #>> (fn tyco => Datatypecons (c, tyco)); - fun stmt_classparam class = - ensure_class thy algbr funcgr class - #>> (fn class => Classparam (c, class)); - fun stmt_fun ((vs, ty), raw_thms) = - let - val thms = if null (Term.add_tfreesT ty []) orelse (null o fst o strip_type) ty - then raw_thms - else (map o apfst) (Code_Unit.expand_eta thy 1) raw_thms; - in - fold_map (translate_tyvar_sort thy algbr funcgr) vs - ##>> translate_typ thy algbr funcgr ty - ##>> fold_map (translate_eq thy algbr funcgr) thms - #>> (fn info => Fun (c, info)) - end; - val stmt_const = case Code.get_datatype_of_constr thy c - of SOME tyco => stmt_datatypecons tyco - | NONE => (case AxClass.class_of_param thy c - of SOME class => stmt_classparam class - | NONE => stmt_fun (Code_Wellsorted.typ funcgr c, Code_Wellsorted.eqns funcgr c)) - in ensure_stmt lookup_const (declare_const thy) stmt_const c end + ##>> fold_map (translate_typ thy algbr funcgr) tys + #>> (fn (tyco, tys) => tyco `%% tys) and translate_term thy algbr funcgr thm (Const (c, ty)) = translate_app thy algbr funcgr thm ((c, ty), []) | translate_term thy algbr funcgr thm (Free (v, _)) = @@ -617,6 +577,15 @@ translate_term thy algbr funcgr thm t' ##>> fold_map (translate_term thy algbr funcgr thm) ts #>> (fn (t, ts) => t `$$ ts) +and translate_eq thy algbr funcgr (thm, linear) = + let + val (args, rhs) = (apfst (snd o strip_comb) o Logic.dest_equals + o Logic.unvarify o prop_of) thm; + in + fold_map (translate_term thy algbr funcgr (SOME thm)) args + ##>> translate_term thy algbr funcgr (SOME thm) rhs + #>> rpair (thm, linear) + end and translate_const thy algbr funcgr thm (c, ty) = let val tys = Sign.const_typargs thy (c, ty); @@ -695,7 +664,38 @@ and translate_app thy algbr funcgr thm (c_ty_ts as ((c, _), _)) = case Code.get_case_scheme thy c of SOME case_scheme => translate_app_case thy algbr funcgr thm case_scheme c_ty_ts - | NONE => translate_app_const thy algbr funcgr thm c_ty_ts; + | NONE => translate_app_const thy algbr funcgr thm c_ty_ts +and translate_tyvar_sort thy (algbr as (proj_sort, _)) funcgr (v, sort) = + fold_map (ensure_class thy algbr funcgr) (proj_sort sort) + #>> (fn sort => (unprefix "'" v, sort)) +and translate_dicts thy (algbr as (proj_sort, algebra)) funcgr thm (ty, sort) = + let + val pp = Syntax.pp_global thy; + datatype typarg = + Global of (class * string) * typarg list list + | 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); + fun type_constructor tyco yss class = + Global ((class, tyco), (map o map) fst yss); + fun type_variable (TFree (v, sort)) = + let + val sort' = proj_sort sort; + in map_index (fn (n, class) => (Local ([], (v, (n, sort'))), class)) sort' end; + val typargs = Sorts.of_sort_derivation pp algebra + {class_relation = class_relation, type_constructor = type_constructor, + type_variable = type_variable} (ty, proj_sort sort) + handle Sorts.CLASS_ERROR e => not_wellsorted thy thm ty sort e; + fun mk_dict (Global (inst, yss)) = + ensure_inst thy algbr funcgr inst + ##>> (fold_map o fold_map) mk_dict yss + #>> (fn (inst, dss) => DictConst (inst, dss)) + | mk_dict (Local (classrels, (v, (k, sort)))) = + fold_map (ensure_classrel thy algbr funcgr) classrels + #>> (fn classrels => DictVar (classrels, (unprefix "'" v, (k, length sort)))) + in fold_map mk_dict typargs end; (* store *) @@ -733,14 +733,14 @@ fun generate_consts thy algebra funcgr = fold_map (ensure_const thy algebra funcgr); in - invoke_generation thy (Code_Wellsorted.make thy cs) generate_consts cs + invoke_generation thy (Code_Wellsorted.obtain thy cs []) generate_consts cs |-> project_consts end; (* value evaluation *) -fun ensure_value thy algbr funcgr t = +fun ensure_value thy algbr funcgr (fs, t) = let val ty = fastype_of t; val vs = fold_term_types (K (fold_atyps (insert (eq_fst op =) @@ -751,33 +751,91 @@ ##>> translate_term thy algbr funcgr NONE t #>> (fn ((vs, ty), t) => Fun (Term.dummy_patternN, ((vs, ty), [(([], t), (Drule.dummy_thm, true))]))); - fun term_value (dep, (naming, program1)) = + fun term_value fs (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; val deps_all = Graph.all_succs program2 deps; val program3 = Graph.subgraph (member (op =) deps_all) program2; - in (((naming, program3), (((vs, ty), t), deps)), (dep, (naming, program2))) end; + in (((naming, program3), (((vs_ty, fs), t), deps)), (dep, (naming, program2))) end; in ensure_stmt ((K o K) NONE) pair stmt_value Term.dummy_patternN #> snd - #> term_value + #> fold_map (fn (v, ty) => translate_typ thy algbr funcgr ty + #-> (fn ty' => pair (v, ty'))) fs + #-> (fn fs' => term_value fs') + end; + +fun base_evaluator thy evaluator algebra funcgr vs t = + let + val fs = Term.add_frees t []; + val (((naming, program), ((((vs', ty'), fs'), t'), deps)), _) = + invoke_generation thy (algebra, funcgr) ensure_value (fs, t); + val vs'' = map (fn (v, _) => (v, (the o AList.lookup (op =) vs o prefix "'") v)) vs'; + in evaluator naming program (((vs'', (vs', ty')), fs'), t') deps end; + +fun eval_conv thy prep_sort = Code_Wellsorted.eval_conv thy prep_sort o base_evaluator thy; +fun eval thy prep_sort postproc = Code_Wellsorted.eval thy prep_sort postproc o base_evaluator thy; + + +(** diagnostic commands **) + +fun code_depgr thy consts = + let + val (_, eqngr) = Code_Wellsorted.obtain thy consts []; + val select = Graph.all_succs eqngr consts; + in + eqngr + |> not (null consts) ? Graph.subgraph (member (op =) select) + |> Graph.map_nodes ((apsnd o map o apfst) (AxClass.overload thy)) end; -fun eval thy evaluator t = +fun code_thms thy = Pretty.writeln o Code_Wellsorted.pretty thy o code_depgr thy; + +fun code_deps thy consts = let - val (t', evaluator'') = evaluator t; - fun evaluator' algebra funcgr = - let - val (((naming, program), (vs_ty_t, deps)), _) = - invoke_generation thy (algebra, funcgr) ensure_value t'; - in evaluator'' naming program vs_ty_t deps end; - in (t', evaluator') end + val eqngr = code_depgr thy consts; + val constss = Graph.strong_conn eqngr; + val mapping = Symtab.empty |> fold (fn consts => fold (fn const => + Symtab.update (const, consts)) consts) constss; + fun succs consts = consts + |> maps (Graph.imm_succs eqngr) + |> subtract (op =) consts + |> map (the o Symtab.lookup mapping) + |> distinct (op =); + val conn = [] |> fold (fn consts => cons (consts, succs consts)) constss; + fun namify consts = map (Code_Unit.string_of_const thy) consts + |> commas; + val prgr = map (fn (consts, constss) => + { name = namify consts, ID = namify consts, dir = "", unfold = true, + path = "", parents = map namify constss }) conn; + in Present.display_graph prgr end; + +local -fun eval_conv thy = Code_Wellsorted.eval_conv thy o eval thy; -fun eval_term thy = Code_Wellsorted.eval_term thy o eval thy; +structure P = OuterParse +and K = OuterKeyword + +fun code_thms_cmd thy = code_thms thy o op @ o Code_Name.read_const_exprs thy; +fun code_deps_cmd thy = code_deps thy o op @ o Code_Name.read_const_exprs thy; + +in + +val _ = + OuterSyntax.improper_command "code_thms" "print system of code equations for code" OuterKeyword.diag + (Scan.repeat P.term_group + >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory + o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of))); + +val _ = + OuterSyntax.improper_command "code_deps" "visualize dependencies of code equations for code" OuterKeyword.diag + (Scan.repeat P.term_group + >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory + o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of))); + +end; end; (*struct*) diff -r 644d18da3c77 -r 7c871a9cf6f4 src/Tools/code/code_wellsorted.ML --- a/src/Tools/code/code_wellsorted.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/Tools/code/code_wellsorted.ML Mon Apr 27 07:26:17 2009 -0700 @@ -7,25 +7,26 @@ signature CODE_WELLSORTED = sig - type T - val eqns: T -> string -> (thm * bool) list - val typ: T -> string -> (string * sort) list * typ - val all: T -> string list - val pretty: theory -> T -> Pretty.T - val make: theory -> string list - -> ((sort -> sort) * Sorts.algebra) * T - val eval_conv: theory - -> (term -> term * (((sort -> sort) * Sorts.algebra) -> T -> thm)) -> cterm -> thm - val eval_term: theory - -> (term -> term * (((sort -> sort) * Sorts.algebra) -> T -> 'a)) -> term -> 'a + type code_algebra + type code_graph + val eqns: code_graph -> string -> (thm * bool) list + val typ: code_graph -> string -> (string * sort) list * typ + val all: code_graph -> string list + val pretty: theory -> code_graph -> Pretty.T + val obtain: theory -> string list -> term list -> code_algebra * code_graph + val eval_conv: theory -> (sort -> sort) + -> (code_algebra -> code_graph -> (string * sort) list -> term -> cterm -> thm) -> cterm -> thm + val eval: theory -> (sort -> sort) -> ((term -> term) -> 'a -> 'a) + -> (code_algebra -> code_graph -> (string * sort) list -> term -> 'a) -> term -> 'a end structure Code_Wellsorted : CODE_WELLSORTED = struct -(** the equation graph type **) +(** the algebra and code equation graph types **) -type T = (((string * sort) list * typ) * (thm * bool) list) Graph.T; +type code_algebra = (sort -> sort) * Sorts.algebra; +type code_graph = (((string * sort) list * typ) * (thm * bool) list) Graph.T; fun eqns eqngr = these o Option.map snd o try (Graph.get_node eqngr); fun typ eqngr = fst o Graph.get_node eqngr; @@ -47,8 +48,10 @@ (* auxiliary *) +fun is_proper_class thy = can (AxClass.get_info thy); + fun complete_proper_sort thy = - Sign.complete_sort thy #> filter (can (AxClass.get_info thy)); + Sign.complete_sort thy #> filter (is_proper_class thy); fun inst_params thy tyco = map (fn (c, _) => AxClass.param_of_inst thy (c, tyco)) @@ -232,8 +235,7 @@ ((class, tyco), map (fn k => (snd o Vargraph.get_node vardeps) (Inst (class, tyco), k)) (0 upto Sign.arity_number thy tyco - 1)); -fun add_eqs thy (proj_sort, algebra) vardeps - (c, (proto_lhs, proto_eqns)) (rhss, eqngr) = +fun add_eqs thy vardeps (c, (proto_lhs, proto_eqns)) (rhss, eqngr) = if can (Graph.get_node eqngr) c then (rhss, eqngr) else let val lhs = map_index (fn (k, (v, _)) => @@ -246,72 +248,30 @@ val eqngr' = Graph.new_node (c, (tyscm, eqns)) eqngr; in (map (pair c) rhss' @ rhss, eqngr') end; -fun extend_arities_eqngr thy cs cs_rhss (arities, eqngr) = +fun extend_arities_eqngr thy cs ts (arities, eqngr) = let - val cs_rhss' = (map o apsnd o map) (styp_of NONE) cs_rhss; + val cs_rhss = (fold o fold_aterms) (fn Const (c_ty as (c, _)) => + insert (op =) (c, (map (styp_of NONE) o Sign.const_typargs thy) c_ty) | _ => I) ts []; val (vardeps, (eqntab, insts)) = empty_vardeps_data |> fold (assert_fun thy arities eqngr) cs - |> fold (assert_rhs thy arities eqngr) cs_rhss'; + |> fold (assert_rhs thy arities eqngr) cs_rhss; val arities' = fold (add_arity thy vardeps) insts arities; val pp = Syntax.pp_global thy; - val is_proper_class = can (AxClass.get_info thy); - val (proj_sort, algebra) = Sorts.subalgebra pp is_proper_class + val algebra = Sorts.subalgebra pp (is_proper_class thy) (AList.lookup (op =) arities') (Sign.classes_of thy); - val (rhss, eqngr') = Symtab.fold - (add_eqs thy (proj_sort, algebra) vardeps) eqntab ([], eqngr); - fun deps_of (c, rhs) = c :: - maps (dicts_of thy (proj_sort, algebra)) - (rhs ~~ (map snd o fst o fst o Graph.get_node eqngr') c); + val (rhss, eqngr') = Symtab.fold (add_eqs thy vardeps) eqntab ([], eqngr); + fun deps_of (c, rhs) = c :: maps (dicts_of thy algebra) + (rhs ~~ (map snd o fst o fst o Graph.get_node eqngr') c); val eqngr'' = fold (fn (c, rhs) => fold (curry Graph.add_edge c) (deps_of rhs)) rhss eqngr'; - in ((proj_sort, algebra), (arities', eqngr'')) end; + in (algebra, (arities', eqngr'')) end; -(** retrieval interfaces **) - -fun proto_eval thy cterm_of evaluator_lift evaluator proto_ct arities_eqngr = - let - val ct = cterm_of proto_ct; - val _ = Sign.no_vars (Syntax.pp_global thy) (Thm.term_of ct); - val _ = Term.fold_types (Type.no_tvars #> K I) (Thm.term_of ct) (); - fun consts_of t = - fold_aterms (fn Const c_ty => cons c_ty | _ => I) t []; - val thm = Code.preprocess_conv thy ct; - val ct' = Thm.rhs_of thm; - val t' = Thm.term_of ct'; - val (t'', evaluator_eqngr) = evaluator t'; - val consts = map fst (consts_of t'); - val consts' = consts_of t''; - val const_matches' = fold (fn (c, ty) => - insert (op =) (c, Sign.const_typargs thy (c, ty))) consts' []; - val (algebra', arities_eqngr') = - extend_arities_eqngr thy consts const_matches' arities_eqngr; - in - (evaluator_lift (evaluator_eqngr algebra') thm (snd arities_eqngr'), - arities_eqngr') - end; - -fun proto_eval_conv thy = - let - fun evaluator_lift evaluator thm1 eqngr = - let - val thm2 = evaluator eqngr; - val thm3 = Code.postprocess_conv thy (Thm.rhs_of thm2); - in - Thm.transitive thm1 (Thm.transitive thm2 thm3) handle THM _ => - error ("could not construct evaluation proof:\n" - ^ (cat_lines o map Display.string_of_thm) [thm1, thm2, thm3]) - end; - in proto_eval thy I evaluator_lift end; - -fun proto_eval_term thy = - let - fun evaluator_lift evaluator _ eqngr = evaluator eqngr; - in proto_eval thy (Thm.cterm_of thy) evaluator_lift end; +(** store **) structure Wellsorted = CodeDataFun ( - type T = ((string * class) * sort list) list * T; + type T = ((string * class) * sort list) list * code_graph; val empty = ([], Graph.empty); fun purge thy cs (arities, eqngr) = let @@ -327,71 +287,52 @@ in (arities', eqngr') end; ); -fun make thy cs = apsnd snd - (Wellsorted.change_yield thy (extend_arities_eqngr thy cs [])); -fun eval_conv thy f = - fst o Wellsorted.change_yield thy o proto_eval_conv thy f; +(** retrieval interfaces **) -fun eval_term thy f = - fst o Wellsorted.change_yield thy o proto_eval_term thy f; - - -(** diagnostic commands **) +fun obtain thy cs ts = apsnd snd + (Wellsorted.change_yield thy (extend_arities_eqngr thy cs ts)); -fun code_depgr thy consts = - let - val (_, eqngr) = make thy consts; - val select = Graph.all_succs eqngr consts; - in - eqngr - |> not (null consts) ? Graph.subgraph (member (op =) select) - |> Graph.map_nodes ((apsnd o map o apfst) (AxClass.overload thy)) - end; +fun prepare_sorts prep_sort (Const (c, ty)) = Const (c, map_type_tfree + (fn (v, sort) => TFree (v, prep_sort sort)) ty) + | prepare_sorts prep_sort (t1 $ t2) = + prepare_sorts prep_sort t1 $ prepare_sorts prep_sort t2 + | prepare_sorts prep_sort (Abs (v, ty, t)) = + Abs (v, Type.strip_sorts ty, prepare_sorts prep_sort t) + | prepare_sorts _ (Term.Free (v, ty)) = Term.Free (v, Type.strip_sorts ty) + | prepare_sorts _ (t as Bound _) = t; -fun code_thms thy = Pretty.writeln o pretty thy o code_depgr thy; - -fun code_deps thy consts = +fun gen_eval thy cterm_of conclude_evaluation prep_sort evaluator proto_ct = let - val eqngr = code_depgr thy consts; - val constss = Graph.strong_conn eqngr; - val mapping = Symtab.empty |> fold (fn consts => fold (fn const => - Symtab.update (const, consts)) consts) constss; - fun succs consts = consts - |> maps (Graph.imm_succs eqngr) - |> subtract (op =) consts - |> map (the o Symtab.lookup mapping) - |> distinct (op =); - val conn = [] |> fold (fn consts => cons (consts, succs consts)) constss; - fun namify consts = map (Code_Unit.string_of_const thy) consts - |> commas; - val prgr = map (fn (consts, constss) => - { name = namify consts, ID = namify consts, dir = "", unfold = true, - path = "", parents = map namify constss }) conn; - in Present.display_graph prgr end; - -local + val ct = cterm_of proto_ct; + val _ = (Term.map_types Type.no_tvars o Sign.no_vars (Syntax.pp_global thy)) + (Thm.term_of ct); + val thm = Code.preprocess_conv thy ct; + val ct' = Thm.rhs_of thm; + val t' = Thm.term_of ct'; + val vs = Term.add_tfrees t' []; + val consts = fold_aterms + (fn Const (c, _) => insert (op =) c | _ => I) t' []; + val t'' = prepare_sorts prep_sort t'; + val (algebra', eqngr') = obtain thy consts [t'']; + in conclude_evaluation (evaluator algebra' eqngr' vs t'' ct') thm end; -structure P = OuterParse -and K = OuterKeyword - -fun code_thms_cmd thy = code_thms thy o op @ o Code_Name.read_const_exprs thy; -fun code_deps_cmd thy = code_deps thy o op @ o Code_Name.read_const_exprs thy; - -in +fun simple_evaluator evaluator algebra eqngr vs t ct = + evaluator algebra eqngr vs t; -val _ = - OuterSyntax.improper_command "code_thms" "print system of code equations for code" OuterKeyword.diag - (Scan.repeat P.term_group - >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory - o Toplevel.keep ((fn thy => code_thms_cmd thy cs) o Toplevel.theory_of))); +fun eval_conv thy = + let + fun conclude_evaluation thm2 thm1 = + let + val thm3 = Code.postprocess_conv thy (Thm.rhs_of thm2); + in + Thm.transitive thm1 (Thm.transitive thm2 thm3) handle THM _ => + error ("could not construct evaluation proof:\n" + ^ (cat_lines o map Display.string_of_thm) [thm1, thm2, thm3]) + end; + in gen_eval thy I conclude_evaluation end; -val _ = - OuterSyntax.improper_command "code_deps" "visualize dependencies of code equations for code" OuterKeyword.diag - (Scan.repeat P.term_group - >> (fn cs => Toplevel.no_timing o Toplevel.unknown_theory - o Toplevel.keep ((fn thy => code_deps_cmd thy cs) o Toplevel.theory_of))); - -end; +fun eval thy prep_sort postproc evaluator = gen_eval thy (Thm.cterm_of thy) + (K o postproc (Code.postprocess_term thy)) prep_sort (simple_evaluator evaluator); end; (*struct*) diff -r 644d18da3c77 -r 7c871a9cf6f4 src/Tools/nbe.ML --- a/src/Tools/nbe.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/Tools/nbe.ML Mon Apr 27 07:26:17 2009 -0700 @@ -7,7 +7,7 @@ signature NBE = sig val norm_conv: cterm -> thm - val norm_term: theory -> term -> term + val norm: theory -> term -> term datatype Univ = Const of int * Univ list (*named (uninterpreted) constants*) @@ -350,7 +350,7 @@ (* term evaluation *) -fun eval_term ctxt gr deps ((vs, ty) : typscheme, t) = +fun eval_term ctxt gr deps (vs : (string * sort) list, t) = let val frees = Code_Thingol.fold_unbound_varnames (insert (op =)) t [] val frees' = map (fn v => Free (v, [])) frees; @@ -364,6 +364,15 @@ (* reification *) +fun typ_of_itype program vs (ityco `%% itys) = + let + val Code_Thingol.Datatype (tyco, _) = Graph.get_node program ityco; + in Type (tyco, map (typ_of_itype program vs) itys) end + | typ_of_itype program vs (ITyVar v) = + let + val sort = (the o AList.lookup (op =) vs) v; + in TFree ("'" ^ v, sort) end; + fun term_of_univ thy program idx_tab t = let fun take_until f [] = [] @@ -418,41 +427,40 @@ (* compilation, evaluation and reification *) -fun compile_eval thy naming program vs_ty_t deps = +fun compile_eval thy naming program vs_t deps = let val ctxt = ProofContext.init thy; val (_, (gr, (_, idx_tab))) = Nbe_Functions.change thy (ensure_stmts ctxt naming program o snd); in - vs_ty_t + vs_t |> eval_term ctxt gr deps |> term_of_univ thy program idx_tab end; (* evaluation with type reconstruction *) -fun eval thy t naming program vs_ty_t deps = +fun normalize thy naming program (((vs0, (vs, ty)), fs), t) deps = let fun subst_const f = map_aterms (fn t as Term.Const (c, ty) => Term.Const (f c, ty) | t => t); - val subst_triv_consts = subst_const (Code_Unit.resubst_alias thy); - val ty = type_of t; - val type_free = AList.lookup (op =) - (map (fn (s, T) => (s, Term.Free (s, T))) (Term.add_frees t [])); - val type_frees = Term.map_aterms - (fn (t as Term.Free (s, _)) => the_default t (type_free s) | t => t); + val resubst_triv_consts = subst_const (Code_Unit.resubst_alias thy); + val ty' = typ_of_itype program vs0 ty; + val fs' = (map o apsnd) (typ_of_itype program vs0) fs; + val type_frees = Term.map_aterms (fn (t as Term.Free (s, _)) => + Term.Free (s, (the o AList.lookup (op =) fs') s) | t => t); fun type_infer t = singleton (TypeInfer.infer_types (Syntax.pp_global thy) (Sign.tsig_of thy) I (try (Type.strip_sorts o Sign.the_const_type thy)) (K NONE) Name.context 0) - (TypeInfer.constrain ty t); + (TypeInfer.constrain ty' t); fun check_tvars t = if null (Term.add_tvars t []) then t else error ("Illegal schematic type variables in normalized term: " ^ setmp show_types true (Syntax.string_of_term_global thy) t); val string_of_term = setmp show_types true (Syntax.string_of_term_global thy); in - compile_eval thy naming program vs_ty_t deps + compile_eval thy naming program (vs, t) deps |> tracing (fn t => "Normalized:\n" ^ string_of_term t) - |> subst_triv_consts + |> resubst_triv_consts |> type_frees |> tracing (fn t => "Vars typed:\n" ^ string_of_term t) |> type_infer @@ -463,39 +471,36 @@ (* evaluation oracle *) -val (_, norm_oracle) = Context.>>> (Context.map_theory_result - (Thm.add_oracle (Binding.name "norm", fn (thy, t, naming, program, vs_ty_t, deps) => - Thm.cterm_of thy (Logic.mk_equals (t, eval thy t naming program vs_ty_t deps))))); +fun add_triv_classes thy = curry (Sorts.inter_sort (Sign.classes_of thy)) + (Code_Unit.triv_classes thy); -fun add_triv_classes thy = +fun mk_equals thy lhs raw_rhs = let - val inters = curry (Sorts.inter_sort (Sign.classes_of thy)) - (Code_Unit.triv_classes thy); - fun map_sorts f = (map_types o map_atyps) - (fn TVar (v, sort) => TVar (v, f sort) - | TFree (v, sort) => TFree (v, f sort)); - in map_sorts inters end; + val ty = Thm.typ_of (Thm.ctyp_of_term lhs); + val eq = Thm.cterm_of thy (Term.Const ("==", ty --> ty --> propT)); + val rhs = Thm.cterm_of thy raw_rhs; + in Thm.mk_binop eq lhs rhs end; + +val (_, raw_norm_oracle) = Context.>>> (Context.map_theory_result + (Thm.add_oracle (Binding.name "norm", fn (thy, naming, program, vsp_ty_fs_t, deps, ct) => + mk_equals thy ct (normalize thy naming program vsp_ty_fs_t deps)))); + +fun norm_oracle thy naming program vsp_ty_fs_t deps ct = + raw_norm_oracle (thy, naming, program, vsp_ty_fs_t, deps, ct); fun norm_conv ct = let val thy = Thm.theory_of_cterm ct; - fun evaluator' t naming program vs_ty_t deps = - norm_oracle (thy, t, naming, program, vs_ty_t, deps); - fun evaluator t = (add_triv_classes thy t, evaluator' t); - in Code_Thingol.eval_conv thy evaluator ct end; + in Code_Thingol.eval_conv thy (add_triv_classes thy) (norm_oracle thy) ct end; -fun norm_term thy t = - let - fun evaluator' t naming program vs_ty_t deps = eval thy t naming program vs_ty_t deps; - fun evaluator t = (add_triv_classes thy t, evaluator' t); - in (Code.postprocess_term thy o Code_Thingol.eval_term thy evaluator) t end; +fun norm thy = Code_Thingol.eval thy (add_triv_classes thy) I (normalize thy); (* evaluation command *) fun norm_print_term ctxt modes t = let val thy = ProofContext.theory_of ctxt; - val t' = norm_term thy t; + val t' = norm thy t; val ty' = Term.type_of t'; val ctxt' = Variable.auto_fixes t ctxt; val p = PrintMode.with_modes modes (fn () => @@ -510,8 +515,7 @@ let val ctxt = Toplevel.context_of state in norm_print_term ctxt modes (Syntax.read_term ctxt s) end; -val setup = - Value.add_evaluator ("nbe", norm_term o ProofContext.theory_of); +val setup = Value.add_evaluator ("nbe", norm o ProofContext.theory_of); local structure P = OuterParse and K = OuterKeyword in diff -r 644d18da3c77 -r 7c871a9cf6f4 src/Tools/quickcheck.ML --- a/src/Tools/quickcheck.ML Wed Apr 22 11:00:25 2009 -0700 +++ b/src/Tools/quickcheck.ML Mon Apr 27 07:26:17 2009 -0700 @@ -6,28 +6,48 @@ signature QUICKCHECK = sig - val test_term: Proof.context -> bool -> string option -> int -> int -> term -> (string * term) list option; - val add_generator: string * (Proof.context -> term -> int -> term list option) -> theory -> theory val auto: bool ref val auto_time_limit: int ref + val test_term: Proof.context -> bool -> string option -> int -> int -> term -> + (string * term) list option + val add_generator: string * (Proof.context -> term -> int -> term list option) -> theory -> theory end; structure Quickcheck : QUICKCHECK = struct +(* preferences *) + +val auto = ref false; +val auto_time_limit = ref 2500; + +val _ = + ProofGeneralPgip.add_preference Preferences.category_tracing + (setmp auto true (fn () => + Preferences.bool_pref auto + "auto-quickcheck" + "Whether to enable quickcheck automatically.") ()); + +val _ = + ProofGeneralPgip.add_preference Preferences.category_tracing + (Preferences.nat_pref auto_time_limit + "auto-quickcheck-time-limit" + "Time limit for automatic quickcheck (in milliseconds)."); + + (* quickcheck configuration -- default parameters, test generators *) datatype test_params = Test_Params of { size: int, iterations: int, default_type: typ option }; -fun dest_test_params (Test_Params { size, iterations, default_type}) = +fun dest_test_params (Test_Params { size, iterations, default_type }) = ((size, iterations), default_type); fun mk_test_params ((size, iterations), default_type) = Test_Params { size = size, iterations = iterations, default_type = default_type }; fun map_test_params f (Test_Params { size, iterations, default_type}) = mk_test_params (f ((size, iterations), default_type)); -fun merge_test_params (Test_Params {size = size1, iterations = iterations1, default_type = default_type1}, - Test_Params {size = size2, iterations = iterations2, default_type = default_type2}) = +fun merge_test_params (Test_Params { size = size1, iterations = iterations1, default_type = default_type1 }, + Test_Params { size = size2, iterations = iterations2, default_type = default_type2 }) = mk_test_params ((Int.max (size1, size2), Int.max (iterations1, iterations2)), case default_type1 of NONE => default_type2 | _ => default_type1); @@ -138,10 +158,7 @@ (* automatic testing *) -val auto = ref false; -val auto_time_limit = ref 5000; - -fun test_goal_auto int state = +val _ = Context.>> (Specification.add_theorem_hook (fn int => fn state => let val ctxt = Proof.context_of state; val assms = map term_of (Assumption.all_assms_of ctxt); @@ -162,12 +179,10 @@ if int andalso !auto andalso not (!Toplevel.quiet) then test () else state - end; - -val _ = Context.>> (Specification.add_theorem_hook test_goal_auto); + end)); -(* Isar interfaces *) +(* Isar commands *) fun read_nat s = case (Library.read_int o Symbol.explode) s of (k, []) => if k >= 0 then k