# HG changeset patch # User chaieb # Date 1248376401 -7200 # Node ID abda97d2deea79fed9a6e37cdbceb9eac8abaead # Parent 53716a67c3b1251a9fde6091bcfa5d7555ff08fa# Parent 63686057cbe8379ec80c4f9e80d43757ea624b68 merged diff -r 63686057cbe8 -r abda97d2deea Admin/isatest/isatest-makedist --- a/Admin/isatest/isatest-makedist Thu Jul 23 21:12:57 2009 +0200 +++ b/Admin/isatest/isatest-makedist Thu Jul 23 21:13:21 2009 +0200 @@ -110,8 +110,8 @@ sleep 15 $SSH macbroy5 "$MAKEALL $HOME/settings/mac-poly" sleep 15 -$SSH macbroy6 "$MAKEALL $HOME/settings/at-mac-poly-5.1-para" -sleep 15 +#$SSH macbroy6 "$MAKEALL $HOME/settings/at-mac-poly-5.1-para" +#sleep 15 $SSH atbroy51 "$HOME/admin/isatest/isatest-annomaly" echo ------------------- spawned tests successfully --- `date` --- $HOSTNAME >> $DISTLOG 2>&1 diff -r 63686057cbe8 -r abda97d2deea Admin/isatest/settings/at-mac-poly-5.1-para --- a/Admin/isatest/settings/at-mac-poly-5.1-para Thu Jul 23 21:12:57 2009 +0200 +++ b/Admin/isatest/settings/at-mac-poly-5.1-para Thu Jul 23 21:13:21 2009 +0200 @@ -25,4 +25,4 @@ ISABELLE_USEDIR_OPTIONS="-i false -d false -M 4" -HOL_USEDIR_OPTIONS="-p 2 -Q false" +HOL_USEDIR_OPTIONS="-p 2 -q 0" diff -r 63686057cbe8 -r abda97d2deea Admin/isatest/settings/at-poly-5.1-para-e --- a/Admin/isatest/settings/at-poly-5.1-para-e Thu Jul 23 21:12:57 2009 +0200 +++ b/Admin/isatest/settings/at-poly-5.1-para-e Thu Jul 23 21:13:21 2009 +0200 @@ -24,4 +24,4 @@ ISABELLE_USEDIR_OPTIONS="-i true -d pdf -v true -M 20" -HOL_USEDIR_OPTIONS="-p 2 -Q false" +HOL_USEDIR_OPTIONS="-p 2 -q 0" diff -r 63686057cbe8 -r abda97d2deea Admin/isatest/settings/at64-poly-5.1-para --- a/Admin/isatest/settings/at64-poly-5.1-para Thu Jul 23 21:12:57 2009 +0200 +++ b/Admin/isatest/settings/at64-poly-5.1-para Thu Jul 23 21:13:21 2009 +0200 @@ -24,4 +24,4 @@ ISABELLE_USEDIR_OPTIONS="-i true -d pdf -v true -M 2" -HOL_USEDIR_OPTIONS="-p 2 -Q false" +HOL_USEDIR_OPTIONS="-p 2 -q 0" diff -r 63686057cbe8 -r abda97d2deea Admin/isatest/settings/mac-poly-M4 --- a/Admin/isatest/settings/mac-poly-M4 Thu Jul 23 21:12:57 2009 +0200 +++ b/Admin/isatest/settings/mac-poly-M4 Thu Jul 23 21:13:21 2009 +0200 @@ -25,4 +25,4 @@ ISABELLE_USEDIR_OPTIONS="-i false -d false -M 4" -HOL_USEDIR_OPTIONS="-p 2 -Q false" +HOL_USEDIR_OPTIONS="-p 2 -q 0" diff -r 63686057cbe8 -r abda97d2deea Admin/isatest/settings/mac-poly-M8 --- a/Admin/isatest/settings/mac-poly-M8 Thu Jul 23 21:12:57 2009 +0200 +++ b/Admin/isatest/settings/mac-poly-M8 Thu Jul 23 21:13:21 2009 +0200 @@ -4,7 +4,7 @@ ML_SYSTEM="polyml-experimental" ML_PLATFORM="x86-darwin" ML_HOME="$POLYML_HOME/$ML_PLATFORM" - ML_OPTIONS="--mutable 600 --immutable 1800" + ML_OPTIONS="--mutable 800 --immutable 2000" ISABELLE_HOME_USER=~/isabelle-mac-poly-M8 @@ -25,4 +25,4 @@ ISABELLE_USEDIR_OPTIONS="-i false -d false -M 8" -HOL_USEDIR_OPTIONS="-p 2 -Q false" +HOL_USEDIR_OPTIONS="-p 2 -q 0" diff -r 63686057cbe8 -r abda97d2deea Admin/isatest/settings/mac-poly64-M4 --- a/Admin/isatest/settings/mac-poly64-M4 Thu Jul 23 21:12:57 2009 +0200 +++ b/Admin/isatest/settings/mac-poly64-M4 Thu Jul 23 21:13:21 2009 +0200 @@ -25,4 +25,4 @@ ISABELLE_USEDIR_OPTIONS="-i false -d false -M 4" -HOL_USEDIR_OPTIONS="-p 2 -Q false" +HOL_USEDIR_OPTIONS="-p 2 -q 1" diff -r 63686057cbe8 -r abda97d2deea NEWS --- a/NEWS Thu Jul 23 21:12:57 2009 +0200 +++ b/NEWS Thu Jul 23 21:13:21 2009 +0200 @@ -18,6 +18,11 @@ *** HOL *** +* More convenient names for set intersection and union. INCOMPATIBILITY: + + Set.Int ~> Set.inter + Set.Un ~> Set.union + * Code generator attributes follow the usual underscore convention: code_unfold replaces code unfold code_post replaces code post @@ -28,6 +33,14 @@ * New type class boolean_algebra. +* Refinements to lattices classes: + - added boolean_algebra type class + - less default intro/elim rules in locale variant, more default + intro/elim rules in class variant: more uniformity + - lemma ge_sup_conv renamed to le_sup_iff, in accordance with le_inf_iff + - dropped lemma alias inf_ACI for inf_aci (same for sup_ACI and sup_aci) + - renamed ACI to inf_sup_aci + * 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 @@ -41,6 +54,9 @@ multiplicative monoids retains syntax "^" and is now defined generic in class power. INCOMPATIBILITY. +* GCD now has functions Gcd/GCD and Lcm/LCM for the gcd and lcm of finite and +infinite sets. It is shown that they form a complete lattice. + * ML antiquotation @{code_datatype} inserts definition of a datatype generated by the code generator; see Predicate.thy for an example. @@ -112,6 +128,15 @@ cominators for "args". INCOMPATIBILITY, need to use simplified Attrib/Method.setup introduced in Isabelle2009. +* Proper context for simpset_of, claset_of, clasimpset_of. May fall +back on global_simpset_of, global_claset_of, global_clasimpset_of as +last resort. INCOMPATIBILITY. + +* Display.pretty_thm now requires a proper context (cf. former +ProofContext.pretty_thm). May fall back on Display.pretty_thm_global +or even Display.pretty_thm_without_context as last resort. +INCOMPATIBILITY. + *** System *** @@ -120,6 +145,9 @@ * Removed "compress" option from isabelle-process and isabelle usedir; this is always enabled. +* More fine-grained control of proof parallelism, cf. +Goal.parallel_proofs in ML and usedir option -q LEVEL. + New in Isabelle2009 (April 2009) diff -r 63686057cbe8 -r abda97d2deea doc-src/IsarRef/showsymbols --- a/doc-src/IsarRef/showsymbols Thu Jul 23 21:12:57 2009 +0200 +++ b/doc-src/IsarRef/showsymbols Thu Jul 23 21:13:21 2009 +0200 @@ -1,6 +1,4 @@ #!/usr/bin/env perl -# -# $Id$ print "\\begin{supertabular}{ll\@{\\qquad}ll}\n"; diff -r 63686057cbe8 -r abda97d2deea doc-src/Main/Docs/Main_Doc.thy --- a/doc-src/Main/Docs/Main_Doc.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/doc-src/Main/Docs/Main_Doc.thy Thu Jul 23 21:13:21 2009 +0200 @@ -107,11 +107,11 @@ \begin{supertabular}{@ {} l @ {~::~} l l @ {}} @{const Set.empty} & @{term_type_only "Set.empty" "'a set"}\\ -@{const insert} & @{term_type_only insert "'a\'a set\'a set"}\\ +@{const Set.insert} & @{term_type_only insert "'a\'a set\'a set"}\\ @{const Collect} & @{term_type_only Collect "('a\bool)\'a set"}\\ @{const "op :"} & @{term_type_only "op :" "'a\'a set\bool"} & (\texttt{:})\\ -@{const Set.Un} & @{term_type_only Set.Un "'a set\'a set \ 'a set"} & (\texttt{Un})\\ -@{const Set.Int} & @{term_type_only Set.Int "'a set\'a set \ 'a set"} & (\texttt{Int})\\ +@{const Set.union} & @{term_type_only Set.Un "'a set\'a set \ 'a set"} & (\texttt{Un})\\ +@{const Set.inter} & @{term_type_only Set.Int "'a set\'a set \ 'a set"} & (\texttt{Int})\\ @{const UNION} & @{term_type_only UNION "'a set\('a \ 'b set) \ 'b set"}\\ @{const INTER} & @{term_type_only INTER "'a set\('a \ 'b set) \ 'b set"}\\ @{const Union} & @{term_type_only Union "'a set set\'a set"}\\ diff -r 63686057cbe8 -r abda97d2deea doc-src/System/Makefile --- a/doc-src/System/Makefile Thu Jul 23 21:12:57 2009 +0200 +++ b/doc-src/System/Makefile Thu Jul 23 21:13:21 2009 +0200 @@ -1,7 +1,3 @@ -# -# $Id$ -# - ## targets default: dvi diff -r 63686057cbe8 -r abda97d2deea doc-src/System/Thy/Basics.thy --- a/doc-src/System/Thy/Basics.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/doc-src/System/Thy/Basics.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1,5 +1,3 @@ -(* $Id$ *) - theory Basics imports Pure begin diff -r 63686057cbe8 -r abda97d2deea doc-src/System/Thy/Interfaces.thy --- a/doc-src/System/Thy/Interfaces.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/doc-src/System/Thy/Interfaces.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1,5 +1,3 @@ -(* $Id$ *) - theory Interfaces imports Pure begin diff -r 63686057cbe8 -r abda97d2deea doc-src/System/Thy/Misc.thy --- a/doc-src/System/Thy/Misc.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/doc-src/System/Thy/Misc.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1,5 +1,3 @@ -(* $Id$ *) - theory Misc imports Pure begin diff -r 63686057cbe8 -r abda97d2deea doc-src/System/Thy/Presentation.thy --- a/doc-src/System/Thy/Presentation.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/doc-src/System/Thy/Presentation.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1,5 +1,3 @@ -(* $Id$ *) - theory Presentation imports Pure begin @@ -442,7 +440,6 @@ -D PATH dump generated document sources into PATH -M MAX multithreading: maximum number of worker threads (default 1) -P PATH set path for remote theory browsing information - -Q BOOL check proofs in parallel (default true) -T LEVEL multithreading: trace level (default 0) -V VERSION declare alternative document VERSION -b build mode (output heap image, using current dir) @@ -451,7 +448,8 @@ -g BOOL generate session graph image for document (default false) -i BOOL generate theory browser information (default false) -m MODE add print mode for output - -p LEVEL set level of detail for proof objects + -p LEVEL set level of detail for proof objects (default 0) + -q LEVEL set level of parallel proof checking (default 1) -r reset session path -s NAME override session NAME -t BOOL internal session timing (default false) @@ -564,6 +562,12 @@ for internal proof objects, see also the \emph{Isabelle Reference Manual}~\cite{isabelle-ref}. + \medskip The @{verbatim "-q"} option specifies the level of parallel + proof checking: @{verbatim 0} no proofs, @{verbatim 1} toplevel + proofs (default), @{verbatim 2} toplevel and nested Isar proofs. + The resulting speedup may vary, depending on the number of worker + threads, granularity of proofs, and whether proof terms are enabled. + \medskip The @{verbatim "-t"} option produces a more detailed internal timing report of the session. @@ -583,13 +587,6 @@ bottle-necks, e.g.\ due to excessive wait states when locking critical code sections. - \medskip The @{verbatim "-Q"} option tells whether individual proofs - should be checked in parallel; the default is @{verbatim true}. - Note that fine-grained proof parallelism requires considerable - amounts of extra memory, since full proof context information is - maintained for each independent derivation thread, until checking is - completed. - \medskip Any @{tool usedir} session is named by some \emph{session identifier}. These accumulate, documenting the way sessions depend on others. For example, consider @{verbatim "Pure/FOL/ex"}, which diff -r 63686057cbe8 -r abda97d2deea doc-src/System/Thy/ROOT.ML --- a/doc-src/System/Thy/ROOT.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/doc-src/System/Thy/ROOT.ML Thu Jul 23 21:13:21 2009 +0200 @@ -1,6 +1,3 @@ - -(* $Id$ *) - set ThyOutput.source; use "../../antiquote_setup.ML"; diff -r 63686057cbe8 -r abda97d2deea doc-src/System/Thy/document/Basics.tex --- a/doc-src/System/Thy/document/Basics.tex Thu Jul 23 21:12:57 2009 +0200 +++ b/doc-src/System/Thy/document/Basics.tex Thu Jul 23 21:13:21 2009 +0200 @@ -3,8 +3,6 @@ \def\isabellecontext{Basics}% % \isadelimtheory -\isanewline -\isanewline % \endisadelimtheory % diff -r 63686057cbe8 -r abda97d2deea doc-src/System/Thy/document/Interfaces.tex --- a/doc-src/System/Thy/document/Interfaces.tex Thu Jul 23 21:12:57 2009 +0200 +++ b/doc-src/System/Thy/document/Interfaces.tex Thu Jul 23 21:13:21 2009 +0200 @@ -3,8 +3,6 @@ \def\isabellecontext{Interfaces}% % \isadelimtheory -\isanewline -\isanewline % \endisadelimtheory % diff -r 63686057cbe8 -r abda97d2deea doc-src/System/Thy/document/Misc.tex --- a/doc-src/System/Thy/document/Misc.tex Thu Jul 23 21:12:57 2009 +0200 +++ b/doc-src/System/Thy/document/Misc.tex Thu Jul 23 21:13:21 2009 +0200 @@ -3,8 +3,6 @@ \def\isabellecontext{Misc}% % \isadelimtheory -\isanewline -\isanewline % \endisadelimtheory % diff -r 63686057cbe8 -r abda97d2deea doc-src/System/Thy/document/Presentation.tex --- a/doc-src/System/Thy/document/Presentation.tex Thu Jul 23 21:12:57 2009 +0200 +++ b/doc-src/System/Thy/document/Presentation.tex Thu Jul 23 21:13:21 2009 +0200 @@ -3,8 +3,6 @@ \def\isabellecontext{Presentation}% % \isadelimtheory -\isanewline -\isanewline % \endisadelimtheory % @@ -468,7 +466,6 @@ -D PATH dump generated document sources into PATH -M MAX multithreading: maximum number of worker threads (default 1) -P PATH set path for remote theory browsing information - -Q BOOL check proofs in parallel (default true) -T LEVEL multithreading: trace level (default 0) -V VERSION declare alternative document VERSION -b build mode (output heap image, using current dir) @@ -477,7 +474,8 @@ -g BOOL generate session graph image for document (default false) -i BOOL generate theory browser information (default false) -m MODE add print mode for output - -p LEVEL set level of detail for proof objects + -p LEVEL set level of detail for proof objects (default 0) + -q LEVEL set level of parallel proof checking (default 1) -r reset session path -s NAME override session NAME -t BOOL internal session timing (default false) @@ -581,6 +579,12 @@ for internal proof objects, see also the \emph{Isabelle Reference Manual}~\cite{isabelle-ref}. + \medskip The \verb|-q| option specifies the level of parallel + proof checking: \verb|0| no proofs, \verb|1| toplevel + proofs (default), \verb|2| toplevel and nested Isar proofs. + The resulting speedup may vary, depending on the number of worker + threads, granularity of proofs, and whether proof terms are enabled. + \medskip The \verb|-t| option produces a more detailed internal timing report of the session. @@ -599,13 +603,6 @@ bottle-necks, e.g.\ due to excessive wait states when locking critical code sections. - \medskip The \verb|-Q| option tells whether individual proofs - should be checked in parallel; the default is \verb|true|. - Note that fine-grained proof parallelism requires considerable - amounts of extra memory, since full proof context information is - maintained for each independent derivation thread, until checking is - completed. - \medskip Any \hyperlink{tool.usedir}{\mbox{\isa{\isatt{usedir}}}} session is named by some \emph{session identifier}. These accumulate, documenting the way sessions depend on others. For example, consider \verb|Pure/FOL/ex|, which diff -r 63686057cbe8 -r abda97d2deea doc-src/System/system.tex --- a/doc-src/System/system.tex Thu Jul 23 21:12:57 2009 +0200 +++ b/doc-src/System/system.tex Thu Jul 23 21:13:21 2009 +0200 @@ -1,6 +1,3 @@ - -%% $Id$ - \documentclass[12pt,a4paper]{report} \usepackage{supertabular} \usepackage{graphicx} diff -r 63686057cbe8 -r abda97d2deea doc-src/TutorialI/Protocol/Message.thy --- a/doc-src/TutorialI/Protocol/Message.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/doc-src/TutorialI/Protocol/Message.thy Thu Jul 23 21:13:21 2009 +0200 @@ -915,15 +915,15 @@ lemmas Fake_parts_sing_imp_Un = Fake_parts_sing [THEN [2] rev_subsetD] method_setup spy_analz = {* - Scan.succeed (SIMPLE_METHOD' o spy_analz_tac o local_clasimpset_of) *} + Scan.succeed (SIMPLE_METHOD' o spy_analz_tac o clasimpset_of) *} "for proving the Fake case when analz is involved" method_setup atomic_spy_analz = {* - Scan.succeed (SIMPLE_METHOD' o atomic_spy_analz_tac o local_clasimpset_of) *} + Scan.succeed (SIMPLE_METHOD' o atomic_spy_analz_tac o clasimpset_of) *} "for debugging spy_analz" method_setup Fake_insert_simp = {* - Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac o local_simpset_of) *} + Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac o simpset_of) *} "for debugging spy_analz" diff -r 63686057cbe8 -r abda97d2deea doc-src/TutorialI/Protocol/Public.thy --- a/doc-src/TutorialI/Protocol/Public.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/doc-src/TutorialI/Protocol/Public.thy Thu Jul 23 21:13:21 2009 +0200 @@ -154,7 +154,7 @@ ML {* fun possibility_tac ctxt = REPEAT (*omit used_Says so that Nonces start from different traces!*) - (ALLGOALS (simp_tac (local_simpset_of ctxt delsimps [used_Says])) + (ALLGOALS (simp_tac (simpset_of ctxt delsimps [used_Says])) THEN REPEAT_FIRST (eq_assume_tac ORELSE' resolve_tac [refl, conjI, @{thm Nonce_supply}])); diff -r 63686057cbe8 -r abda97d2deea doc-src/rail.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/rail.ML Thu Jul 23 21:13:21 2009 +0200 @@ -0,0 +1,468 @@ +datatype token = + Identifier of string | + Special_Identifier of string | + Text of string | + Anot of string | + Symbol of string | + EOF; + +fun is_identifier (Identifier _) = true + | is_identifier (Special_Identifier _ ) = true + | is_identifier _ = false; + +fun is_text (Text _) = true + | is_text _ = false; + +fun is_anot (Anot _) = true + | is_anot _ = false; + +fun is_symbol (Symbol _) = true + | is_symbol _ = false; + +fun is_ str = (fn s => s = Symbol str); + + +local (* copied from antiquote-setup... *) +fun translate f = Symbol.explode #> map f #> implode; + +fun clean_name "\" = "dots" + | clean_name ".." = "ddot" + | clean_name "." = "dot" + | clean_name "_" = "underscore" + | clean_name "{" = "braceleft" + | clean_name "}" = "braceright" + | clean_name s = s |> translate (fn "_" => "-" | "\" => "-" | c => c); + +fun no_check _ _ = true; + +fun false_check _ _ = false; + +fun thy_check intern defined ctxt = + let val thy = ProofContext.theory_of ctxt + in defined thy o intern thy end; + +fun check_tool name = + File.exists (Path.append (Path.explode "~~/lib/Tools") (Path.basic name)); + +val arg = enclose "{" "}"; + +val markup_table = +(* [(kind, (markup, check, style))*) + Symtab.make [ + ("syntax", ("", no_check, true)), + ("command", ("isacommand", K (is_some o OuterKeyword.command_keyword), true)), + ("keyword", ("isakeyword", K OuterKeyword.is_keyword, true)), + ("element", ("isakeyword", K OuterKeyword.is_keyword, true)), + ("method", ("", thy_check Method.intern Method.defined, true)), + ("attribute", ("", thy_check Attrib.intern Attrib.defined, true)), + ("fact", ("", no_check, true)), + ("variable", ("", no_check, true)), + ("case", ("", no_check, true)), + ("antiquotation", ("", K ThyOutput.defined_command, true)), + ("antiquotation option", ("", K ThyOutput.defined_option, true)), (* kann mein scanner nicht erkennen *) + ("setting", ("isatt", (fn _ => fn name => is_some (OS.Process.getEnv name)), true)), + ("inference", ("", no_check, true)), + ("executable", ("isatt", no_check, true)), + ("tool", ("isatt", K check_tool, true)), + ("file", ("isatt", K (File.exists o Path.explode), true)), + ("theory", ("", K ThyInfo.known_thy, true)) + ]; + +in + +fun decode_link ctxt (kind,index,logic,name) = + let + val (markup, check, style) = case Symtab.lookup markup_table kind of + SOME x => x + | NONE => ("", false_check, false); + val hyper_name = + "{" ^ Long_Name.append kind (Long_Name.append logic (clean_name name)) ^ "}"; + val hyper = + enclose ("\\hyperlink" ^ hyper_name ^ "{") "}" #> + index = "def" ? enclose ("\\hypertarget" ^ hyper_name ^ "{") "}"; + val idx = + if index = "" then "" + else "\\index" ^ index ^ arg logic ^ arg kind ^ arg name; + in + if check ctxt name then + (idx ^ + (Output.output name + |> (if markup = "" then I else enclose ("\\" ^ markup ^ "{") "}") + |> (if ! ThyOutput.quotes then quote else I) + |> (if ! ThyOutput.display then enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}" + else hyper o enclose "\\mbox{\\isa{" "}}")), style) + else ("Bad " ^ kind ^ " " ^ name, false) + end; +end; + +val blanks = + Scan.many Symbol.is_blank >> implode; + +val scan_symbol = + $$ ";" || $$ ":"|| $$ "("|| $$ ")"|| $$ "+"|| $$ "|"|| $$ "*"|| $$ "?"|| $$ "\\"; + +(* escaped version *) +val scan_link = (* @{kind{_def|_ref (logic) name} *) + let + fun pseudo_antiquote inner_scan = ($$ "@" ^^ $$ "{") |-- inner_scan --| (blanks -- $$ "}"); + fun parens scan = $$ "(" |-- scan --| $$ ")"; + fun opt_quotes scan = $$ "'" |-- scan --| $$ "'" || scan; + val letters = Scan.many Symbol.is_letter >> implode; + val kind_name = letters; + val opt_kind_type = Scan.optional ( + $$ "_" |-- (Scan.this_string "def" || Scan.this_string "ref")) ""; + val logic_name = letters; + val escaped_singlequote = $$ "\\" |-- $$ "'"; + val text = Scan.repeat (Scan.one Symbol.is_letter || escaped_singlequote) >> implode; + in + pseudo_antiquote ( + kind_name -- opt_kind_type -- + (blanks |-- Scan.optional ( parens logic_name ) "") -- + (blanks |-- opt_quotes text) ) + >> (fn (((kind,index),logic),name) => (kind, index, logic, name)) +end; + +(* escaped version *) +fun scan_identifier ctxt = + let fun is_identifier_start s = + Symbol.is_letter s orelse + s = "_" + fun is_identifier_rest s = + Symbol.is_letter s orelse + Symbol.is_digit s orelse + s = "_" orelse + s = "." + in + (Scan.one is_identifier_start ::: + Scan.repeat (Scan.one is_identifier_rest || ($$ "\\" |-- $$ "'")) + ) >> (Identifier o enclose "\\isa{" "}" o Output.output o implode) || + scan_link >> (decode_link ctxt) >> + (fn (txt, style) => + if style then Special_Identifier(txt) + else Identifier(txt)) +end; + +fun scan_anot ctxt = + let val scan_anot = + Scan.many (fn s => + s <> "\n" andalso + s <> "\t" andalso + s <> "]" andalso + Symbol.is_regular s) >> implode + in + $$ "[" |-- scan_link --| $$ "]" >> (fst o (decode_link ctxt)) || + $$ "[" |-- scan_anot --| $$ "]" >> Output.output +end; + +(* escaped version *) +fun scan_text ctxt = + let + val text_sq = + Scan.repeat ( + Scan.one (fn s => + s <> "\n" andalso + s <> "\t" andalso + s <> "'" andalso + s <> "\\" andalso + Symbol.is_regular s) || + ($$ "\\" |-- $$ "'") + ) >> implode + fun quoted scan = $$ "'" |-- scan --| $$ "'"; + in + quoted scan_link >> (fst o (decode_link ctxt)) || + quoted text_sq >> (enclose "\\isa{" "}" o Output.output) +end; + +fun scan_rail ctxt = + Scan.repeat ( blanks |-- ( + scan_identifier ctxt || + scan_anot ctxt >> Anot || + scan_text ctxt >> Text || + scan_symbol >> Symbol) + ) --| blanks; + +fun lex_rail txt ctxt = (* Symbol_Pos fuer spaeter durchgereicht *) + Symbol.scanner "Malformed rail-declaration" (scan_rail ctxt) (map fst (Symbol_Pos.explode txt)); + +val lex = lex_rail; + +datatype id_kind = UNKNOWN | TOKEN | TERM | NTERM; + +datatype id_type = + Id of string * id_kind | + Null_Id; + +datatype body_kind = + CAT | BAR | PLUS | + CR | EMPTY | ANNOTE | IDENT | STRING | + Null_Kind; + +datatype body_type = + Body of body_kind * string * string * id_type * body_type list | + Body_Pos of body_kind * string * string * id_type * body_type list * int * int | + Empty_Body | + Null_Body; + +datatype rule = + Rule of id_type * body_type; + +fun new_id id kind = Id (id, kind); + +fun is_empty (Body(kind,_,_,_,_)) = kind = EMPTY; + +fun new_body (kind, Null_Body, Null_Body) = Body (kind, "", "", Null_Id, []) + | new_body (kind, body1, body2) = Body (kind, "", "", Null_Id, body1 :: [body2]); + +fun is_kind_of kind (Body(bodyKind,_,_,_,_)) = kind = bodyKind + | is_kind_of _ _ = false; + +fun add_list (Body(kind, text, annot, id, bodies), body) = + Body(kind, text, annot, id, bodies @ [body]); + +fun cat_body_lists (Body(kind, text, annot, id, bodies1), Body(_,_,_,_, bodies2)) = + Body(kind, text, annot, id, bodies1 @ bodies2); + +fun add_body (kind, body1 as Body(kind1,_,_,_,_), body2 as Body(kind2,_,_,_,_)) = + if kind = kind1 andalso kind <> BAR then + if kind = kind2 then + cat_body_lists(body1, body2) + else (* kind <> kind2 *) + add_list(body1, body2) + else (* kind <> kind1 orelse kind = BAR *) + if kind = kind2 then + cat_body_lists(add_list(new_body(kind,Null_Body,Null_Body), body1), body2) + else (* kind <> kind2 *) + add_list(add_list(new_body(kind,Null_Body,Null_Body), body1), body2); + +fun rev_body (body as Body (kind, text, annot, id, [])) = body + | rev_body (Body (CAT, text, annot, id, bodies)) = + Body(CAT, text, annot, id, map rev_body (rev bodies)) + | rev_body (Body (kind, text, annot, id, bodies)) = + Body(kind, text, annot, id, map rev_body bodies) + | rev_body body = body; + +fun set_body_text text (Body(k,_,a,i,b)) = Body(k,text,a,i,b); +fun set_body_anot anot (Body(k,t,_,i,b)) = Body(k,t,anot,i,b); +fun set_body_ident id (Body(k,t,a,_,b)) = Body(k,t,a, new_id id TOKEN,b); +fun set_body_special_ident id (Body(k,t,a,_,b)) = Body(k,t,a, new_id id TERM,b); + + +fun mk_eof _ = EOF; +fun is_eof s = s = EOF; +val stopper = Scan.stopper mk_eof is_eof; + +(* TODO: change this, so the next or next two tokens are printed *) +fun lex_err msg (cs, _) = "rail grammar error: " ^ msg cs; +fun !!! msg scan = Scan.!! (lex_err (K msg)) scan; +fun $$$ tok = Scan.one (is_ tok); + + +local +fun new_bar_body([], body2) = body2 + | new_bar_body(body1::bodies, body2) = + add_body(BAR, body1, new_bar_body(bodies, body2)); + +fun new_cat_body(body::[]) = body + | new_cat_body(body1::bodies) = add_body(CAT, body1, new_cat_body(bodies)); + +fun new_annote_body (Anot anot) = + set_body_text anot (new_body(ANNOTE, Empty_Body, Empty_Body)); + +fun new_text_annote_body (Text text, Anot anot) = + set_body_anot anot (set_body_text text (new_body(STRING, Empty_Body, Empty_Body))); + +fun new_ident_body (Identifier ident, Anot anot) = + set_body_anot anot (set_body_ident ident (new_body(IDENT, Empty_Body, Empty_Body))) + | new_ident_body (Special_Identifier ident, Anot anot) = + set_body_anot anot (set_body_special_ident ident (new_body(IDENT, Empty_Body, Empty_Body))); + +val new_empty_body = new_body(EMPTY, Null_Body, Null_Body); +in + +fun parse_body x = + ( + Scan.repeat1 (parse_body0 --| $$$ "|") -- !!! "body0 expected" (parse_body0) >> + new_bar_body || + parse_body0 + ) x +and parse_body0 x = + ( + Scan.one is_anot -- !!! "body1 expected" (parse_body1) >> + (fn (anot, body) => add_body(CAT, new_annote_body(anot), body)) || + parse_body1 + ) x +and parse_body1 x = + ( + parse_body2 -- ($$$ "*" |-- !!! "body4e expected" (parse_body4e)) >> + (fn (body1, body2) => + if is_empty body2 then + add_body(PLUS, new_empty_body, rev_body body1) + else + add_body(BAR, new_empty_body, add_body (PLUS, body1, rev_body body2)) ) || + parse_body2 -- ($$$ "+" |-- !!! "body4e expected" (parse_body4e)) >> + (fn (body1, body2) => new_body (PLUS, body1, rev_body body2) ) || + parse_body2e + ) x +and parse_body2e x = + ( + parse_body2 || + (fn toks => (new_empty_body, toks)) + ) x +and parse_body2 x = + ( + Scan.repeat1 (parse_body3) >> new_cat_body + ) x +and parse_body3 x = + ( + parse_body4 --| $$$ "?" >> (fn body => new_body (BAR, new_empty_body, body) ) || + parse_body4 + ) x +and parse_body4e x = + ( + parse_body4 || + (fn toks => (new_empty_body, toks)) + ) x +and parse_body4 x = + ( + $$$ "(" |-- !!! "body0 or ')' expected" (parse_body --| $$$ ")") || + Scan.one is_text -- (Scan.optional (Scan.one is_anot) (Anot(""))) >> + (fn (text, anot) => new_text_annote_body (text,anot)) || + Scan.one is_identifier -- (Scan.optional (Scan.one is_anot) (Anot(""))) >> + (fn (id, anot) => new_ident_body (id,anot)) || + $$$ "\\" >> (fn _ => new_body (CR, Null_Body, Null_Body)) + ) x; +end; + +fun new_named_rule (Identifier name, body) = Rule(Id(name, UNKNOWN), body) + | new_named_rule (Special_Identifier name, body) = Rule(Id(name, UNKNOWN), body); +fun new_anonym_rule body = Rule(Null_Id, body); + +val parse_rule = + (Scan.one (is_identifier) -- ($$$ ":" |-- !!! "body expected" (parse_body)) ) >> + new_named_rule || + parse_body >> new_anonym_rule; + +val parse_rules = + Scan.repeat ( parse_rule --| $$$ ";") @@@ Scan.single parse_rule; + +fun parse_rail s = + Scan.read stopper parse_rules s; + +val parse = parse_rail; + +fun getystart (Body_Pos(_,_,_,_,_,ystart,_)) = ystart; +fun getynext (Body_Pos(_,_,_,_,_,_,ynext)) = ynext; + +fun position_body (body as Body(kind, text, annot, id, bodies), ystart) = + let fun max (x,y) = if x > y then x else y + fun set_body_position (Body(kind, text, annot, id, bodies), ystart, ynext) = + Body_Pos(kind, text, annot, id, bodies, ystart, ynext) + fun pos_bodies_cat ([],_,ynext,liste) = (liste, ynext) + | pos_bodies_cat (x::xs, ystart, ynext, liste) = + if is_kind_of CR x then + (case set_body_position(x, ystart, ynext+1) of + body as Body_Pos(_,_,_,_,_,_,ynext1) => + pos_bodies_cat(xs, ynext1, max(ynext,ynext1), liste@[body]) + ) + else + (case position_body(x, ystart) of + body as Body_Pos(_,_,_,_,_,_,ynext1) => + pos_bodies_cat(xs, ystart, max(ynext,ynext1), liste@[body]) + ) + fun pos_bodies_bar_plus ([],_,ynext,liste) = (liste, ynext) + | pos_bodies_bar_plus (x::xs, ystart, ynext, liste) = + (case position_body(x, ystart) of + body as Body_Pos(_,_,_,_,_,_,ynext1) => + pos_bodies_bar_plus(xs, ynext1, max(ynext,ynext1), liste@[body]) + ) + in + (case kind of + CAT => (case pos_bodies_cat(bodies,ystart,ystart+1,[]) of + (bodiesPos, ynext) => + Body_Pos(kind, text, annot, id, bodiesPos, ystart, ynext)) + | BAR => (case pos_bodies_bar_plus(bodies,ystart,ystart+1,[]) of + (bodiesPos, ynext) => + Body_Pos(kind, text, annot, id, bodiesPos, ystart, ynext)) + | PLUS => (case pos_bodies_bar_plus(bodies,ystart,ystart+1,[]) of + (bodiesPos, ynext) => + Body_Pos(kind, text, annot, id, bodiesPos, ystart, ynext)) + | CR => set_body_position(body, ystart, ystart+3) + | EMPTY => set_body_position(body, ystart, ystart+1) + | ANNOTE => set_body_position(body, ystart, ystart+1) + | IDENT => set_body_position(body, ystart, ystart+1) + | STRING => set_body_position(body, ystart, ystart+1) + ) + end; + +fun format_body (Body_Pos(EMPTY,_,_,_,_,_,_), _) = "" + | format_body (Body_Pos(CAT,_,_,_,bodies,_,_), cent) = + let fun format_bodies([]) = "" + | format_bodies(x::xs) = format_body (x, "") ^ format_bodies(xs) + in + format_bodies(bodies) + end + | format_body (Body_Pos(BAR,_,_,_,bodies,_,_),cent) = + let fun format_bodies([]) = "\\rail@endbar\n" + | format_bodies(x::xs) = + "\\rail@nextbar{" ^ string_of_int(getystart(x)) ^"}\n" ^ + format_body(x, "") ^ format_bodies(xs) + in + "\\rail@bar\n" ^ format_body(hd(bodies), "") ^ format_bodies(tl(bodies)) + end + | format_body (Body_Pos(PLUS,_,_,_,x::y::xs,_,_),cent) = + "\\rail@plus\n" ^ format_body(x, cent) ^ + "\\rail@nextplus{" ^ string_of_int(getystart(y)) ^ "}\n" ^ + format_body(y, "c") ^ + "\\rail@endplus\n" + | format_body (Body_Pos(ANNOTE,text,_,_,_,_,_),cent) = + "\\rail@annote[" ^ text ^ "]\n" + | format_body (Body_Pos(IDENT,_,annot,Id(name,TERM),_,_,_),cent) = + "\\rail@" ^ cent ^ "token{" ^ name ^ "}" ^ "[" ^ annot ^ "]\n" + | format_body (Body_Pos(IDENT,_,annot,Id(name,_),_,_,_),cent) = + "\\rail@" ^ cent ^ "nont{" ^ name ^ "}" ^ "[" ^ annot ^ "]\n" + | format_body (Body_Pos(CR,_,_,_,_,_,ynext),cent) = + "\\rail@cr{" ^ string_of_int(ynext) ^ "}\n" + | format_body (Body_Pos(STRING,text,annot,_,_,_,_),cent) = + "\\rail@" ^ cent ^ "term{" ^ text ^ "}[" ^ annot ^ "]\n" + | format_body _ = + "\\rail@unknown\n"; + +fun out_body (Id(name,_), body) = + let val bodyPos as Body_Pos(_,_,_,_,_,_,ynext) = position_body(body,0) + in + "\\rail@begin{" ^ string_of_int(ynext) ^ "}{" ^ name ^ "}\n" ^ + format_body(bodyPos,"") ^ + "\\rail@end\n" + end + | out_body (Null_Id, body) = out_body (Id("", UNKNOWN), body); + +fun out_rule (Rule(id, body)) = + if is_empty body then "" + else out_body (id, body); + +fun out_rules ([]) = "" + | out_rules (rule::rules) = out_rule rule ^ out_rules rules; + +val output_no_rules = + "\\rail@begin{1}{}\n" ^ + "\\rail@setbox{\\bfseries ???}\n" ^ + "\\rail@oval\n" ^ + "\\rail@end\n"; + + +fun print (SOME rules) = + "\\begin{railoutput}\n" ^ + out_rules rules ^ + "\\end{railoutput}\n" + | print (NONE) = + "\\begin{railoutput}\n" ^ + output_no_rules ^ + "\\end{railoutput}\n"; + +fun process txt ctxt = + lex txt ctxt + |> parse + |> print; + +val _ = ThyOutput.antiquotation "rail" (Scan.lift ( OuterParse.position Args.name )) + (fn {context = ctxt,...} => fn txt => process txt ctxt); diff -r 63686057cbe8 -r abda97d2deea etc/settings --- a/etc/settings Thu Jul 23 21:12:57 2009 +0200 +++ b/etc/settings Thu Jul 23 21:13:21 2009 +0200 @@ -96,7 +96,7 @@ # Specifically for the HOL image HOL_USEDIR_OPTIONS="" -#HOL_USEDIR_OPTIONS="-p 2 -Q false" +#HOL_USEDIR_OPTIONS="-p 2 -q 1" #Source file identification (default: full name + date stamp) ISABELLE_FILE_IDENT="" diff -r 63686057cbe8 -r abda97d2deea lib/Tools/usedir --- a/lib/Tools/usedir Thu Jul 23 21:12:57 2009 +0200 +++ b/lib/Tools/usedir Thu Jul 23 21:13:21 2009 +0200 @@ -19,7 +19,6 @@ echo " -D PATH dump generated document sources into PATH" echo " -M MAX multithreading: maximum number of worker threads (default 1)" echo " -P PATH set path for remote theory browsing information" - echo " -Q BOOL check proofs in parallel (default true)" echo " -T LEVEL multithreading: trace level (default 0)" echo " -V VERSION declare alternative document VERSION" echo " -b build mode (output heap image, using current dir)" @@ -28,7 +27,8 @@ echo " -g BOOL generate session graph image for document (default false)" echo " -i BOOL generate HTML and graph browser information (default false)" echo " -m MODE add print mode for output" - echo " -p LEVEL set level of detail for proof objects" + echo " -p LEVEL set level of detail for proof objects (default 0)" + echo " -q LEVEL set level of parallel proof checking (default 1)" echo " -r reset session path" echo " -s NAME override session NAME" echo " -t BOOL internal session timing (default false)" @@ -73,7 +73,6 @@ DUMP="" MAXTHREADS="1" RPATH="" -PARALLEL_PROOFS="true" TRACETHREADS="0" DOCUMENT_VERSIONS="" BUILD="" @@ -84,14 +83,15 @@ MODES="" RESET=false SESSION="" -PROOFS=0 +PROOFS="0" +PARALLEL_PROOFS="1" TIMING=false VERBOSE=false function getoptions() { OPTIND=1 - while getopts "C:D:M:P:Q:T:V:bd:f:g:i:m:p:rs:t:v:" OPT + while getopts "C:D:M:P:T:V:bd:f:g:i:m:p:q:rs:t:v:" OPT do case "$OPT" in C) @@ -112,9 +112,6 @@ P) RPATH="$OPTARG" ;; - Q) - PARALLEL_PROOFS="$OPTARG" - ;; T) check_number "$OPTARG" TRACETHREADS="$OPTARG" @@ -154,6 +151,10 @@ check_number "$OPTARG" PROOFS="$OPTARG" ;; + q) + check_number "$OPTARG" + PARALLEL_PROOFS="$OPTARG" + ;; r) RESET=true ;; diff -r 63686057cbe8 -r abda97d2deea src/CCL/CCL.thy --- a/src/CCL/CCL.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/CCL/CCL.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1,5 +1,4 @@ (* Title: CCL/CCL.thy - ID: $Id$ Author: Martin Coen Copyright 1993 University of Cambridge *) @@ -249,18 +248,18 @@ ML {* -val caseB_lemmas = mk_lemmas (thms "caseBs") +val caseB_lemmas = mk_lemmas @{thms caseBs} val ccl_dstncts = let fun mk_raw_dstnct_thm rls s = - prove_goal (the_context ()) s (fn _=> [rtac notI 1,eresolve_tac rls 1]) + prove_goal @{theory} s (fn _=> [rtac notI 1,eresolve_tac rls 1]) in map (mk_raw_dstnct_thm caseB_lemmas) - (mk_dstnct_rls (the_context ()) ["bot","true","false","pair","lambda"]) end + (mk_dstnct_rls @{theory} ["bot","true","false","pair","lambda"]) end fun mk_dstnct_thms thy defs inj_rls xs = - let fun mk_dstnct_thm rls s = prove_goalw thy defs s - (fn _ => [simp_tac (simpset_of thy addsimps (rls@inj_rls)) 1]) - in map (mk_dstnct_thm ccl_dstncts) (mk_dstnct_rls thy xs) end + let fun mk_dstnct_thm rls s = prove_goalw thy defs s + (fn _ => [simp_tac (global_simpset_of thy addsimps (rls@inj_rls)) 1]) + in map (mk_dstnct_thm ccl_dstncts) (mk_dstnct_rls thy xs) end fun mkall_dstnct_thms thy defs i_rls xss = List.concat (map (mk_dstnct_thms thy defs i_rls) xss) @@ -273,9 +272,9 @@ val XH_to_Ds = map XH_to_D val XH_to_Es = map XH_to_E; -bind_thms ("ccl_rews", thms "caseBs" @ ccl_injs @ ccl_dstncts); +bind_thms ("ccl_rews", @{thms caseBs} @ ccl_injs @ ccl_dstncts); bind_thms ("ccl_dstnctsEs", ccl_dstncts RL [notE]); -bind_thms ("ccl_injDs", XH_to_Ds (thms "ccl_injs")); +bind_thms ("ccl_injDs", XH_to_Ds @{thms ccl_injs}); *} lemmas [simp] = ccl_rews @@ -388,13 +387,13 @@ ML {* local - fun mk_thm s = prove_goal (the_context ()) s (fn _ => - [rtac notI 1,dtac (thm "case_pocong") 1,etac rev_mp 5, + fun mk_thm s = prove_goal @{theory} s (fn _ => + [rtac notI 1, dtac @{thm case_pocong} 1, etac rev_mp 5, ALLGOALS (simp_tac @{simpset}), - REPEAT (resolve_tac [thm "po_refl", thm "npo_lam_bot"] 1)]) + REPEAT (resolve_tac [@{thm po_refl}, @{thm npo_lam_bot}] 1)]) in -val npo_rls = [thm "npo_pair_lam", thm "npo_lam_pair"] @ map mk_thm +val npo_rls = [@{thm npo_pair_lam}, @{thm npo_lam_pair}] @ map mk_thm ["~ true [= false", "~ false [= true", "~ true [= ", "~ [= true", "~ true [= lam x. f(x)","~ lam x. f(x) [= true", diff -r 63686057cbe8 -r abda97d2deea src/CCL/Hered.thy --- a/src/CCL/Hered.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/CCL/Hered.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1,5 +1,4 @@ (* Title: CCL/Hered.thy - ID: $Id$ Author: Martin Coen Copyright 1993 University of Cambridge *) @@ -113,7 +112,7 @@ res_inst_tac ctxt [(("R", 0), s)] @{thm HTT_coinduct3} i val HTTgenIs = - map (mk_genIs (the_context ()) @{thms data_defs} @{thm HTTgenXH} @{thm HTTgen_mono}) + map (mk_genIs @{theory} @{thms data_defs} @{thm HTTgenXH} @{thm HTTgen_mono}) ["true : HTTgen(R)", "false : HTTgen(R)", "[| a : R; b : R |] ==> : HTTgen(R)", diff -r 63686057cbe8 -r abda97d2deea src/CCL/Term.thy --- a/src/CCL/Term.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/CCL/Term.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1,5 +1,4 @@ (* Title: CCL/Term.thy - ID: $Id$ Author: Martin Coen Copyright 1993 University of Cambridge *) @@ -274,8 +273,9 @@ ML {* -bind_thms ("term_injs", map (mk_inj_rl (the_context ()) - [thm "applyB", thm "splitB", thm "whenBinl", thm "whenBinr", thm "ncaseBsucc", thm "lcaseBcons"]) +bind_thms ("term_injs", map (mk_inj_rl @{theory} + [@{thm applyB}, @{thm splitB}, @{thm whenBinl}, @{thm whenBinr}, + @{thm ncaseBsucc}, @{thm lcaseBcons}]) ["(inl(a) = inl(a')) <-> (a=a')", "(inr(a) = inr(a')) <-> (a=a')", "(succ(a) = succ(a')) <-> (a=a')", @@ -287,7 +287,7 @@ ML {* bind_thms ("term_dstncts", - mkall_dstnct_thms (the_context ()) (thms "data_defs") (thms "ccl_injs" @ thms "term_injs") + mkall_dstnct_thms @{theory} @{thms data_defs} (@{thms ccl_injs} @ @{thms term_injs}) [["bot","inl","inr"], ["bot","zero","succ"], ["bot","nil","cons"]]); *} @@ -297,8 +297,8 @@ ML {* local - fun mk_thm s = prove_goalw (the_context ()) (thms "data_defs") s (fn _ => - [simp_tac (@{simpset} addsimps (thms "ccl_porews")) 1]) + fun mk_thm s = prove_goalw @{theory} @{thms data_defs} s (fn _ => + [simp_tac (@{simpset} addsimps @{thms ccl_porews}) 1]) in val term_porews = map mk_thm ["inl(a) [= inl(a') <-> a [= a'", "inr(b) [= inr(b') <-> b [= b'", diff -r 63686057cbe8 -r abda97d2deea src/CCL/Type.thy --- a/src/CCL/Type.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/CCL/Type.thy Thu Jul 23 21:13:21 2009 +0200 @@ -132,10 +132,10 @@ fun mk_ncanT_tac ctxt defs top_crls crls s = prove_goalw (ProofContext.theory_of ctxt) defs s (fn major::prems => [(resolve_tac ([major] RL top_crls) 1), (REPEAT_SOME (eresolve_tac (crls @ [exE,bexE,conjE,disjE]))), - (ALLGOALS (asm_simp_tac (local_simpset_of ctxt))), + (ALLGOALS (asm_simp_tac (simpset_of ctxt))), (ALLGOALS (ares_tac (prems RL [lemma]) ORELSE' etac bspec )), - (safe_tac (local_claset_of ctxt addSIs prems))]) + (safe_tac (claset_of ctxt addSIs prems))]) end *} @@ -408,7 +408,7 @@ fun mk_genIs thy defs genXH gen_mono s = prove_goalw thy defs s (fn prems => [rtac (genXH RS iffD2) 1, - simp_tac (simpset_of thy) 1, + simp_tac (global_simpset_of thy) 1, TRY (fast_tac (@{claset} addIs ([genXH RS iffD2,gen_mono RS coinduct3_mono_lemma RS lfpI] @ prems)) 1)]) @@ -428,7 +428,7 @@ ML {* -val POgenIs = map (mk_genIs (the_context ()) (thms "data_defs") (thm "POgenXH") (thm "POgen_mono")) +val POgenIs = map (mk_genIs @{theory} @{thms data_defs} @{thm POgenXH} @{thm POgen_mono}) [" : POgen(R)", " : POgen(R)", "[| : R; : R |] ==> <,> : POgen(R)", @@ -442,10 +442,10 @@ "[| : lfp(%x. POgen(x) Un R Un PO); : lfp(%x. POgen(x) Un R Un PO) |] ==> : POgen(lfp(%x. POgen(x) Un R Un PO))"]; fun POgen_tac ctxt (rla,rlb) i = - SELECT_GOAL (safe_tac (local_claset_of ctxt)) i THEN - rtac (rlb RS (rla RS (thm "ssubst_pair"))) i THEN - (REPEAT (resolve_tac (POgenIs @ [thm "PO_refl" RS (thm "POgen_mono" RS ci3_AI)] @ - (POgenIs RL [thm "POgen_mono" RS ci3_AgenI]) @ [thm "POgen_mono" RS ci3_RI]) i)); + SELECT_GOAL (safe_tac (claset_of ctxt)) i THEN + rtac (rlb RS (rla RS @{thm ssubst_pair})) i THEN + (REPEAT (resolve_tac (POgenIs @ [@{thm PO_refl} RS (@{thm POgen_mono} RS ci3_AI)] @ + (POgenIs RL [@{thm POgen_mono} RS ci3_AgenI]) @ [@{thm POgen_mono} RS ci3_RI]) i)); *} @@ -458,7 +458,7 @@ ML {* -val EQgenIs = map (mk_genIs (the_context ()) (thms "data_defs") (thm "EQgenXH") (thm "EQgen_mono")) +val EQgenIs = map (mk_genIs @{theory} @{thms data_defs} @{thm EQgenXH} @{thm EQgen_mono}) [" : EQgen(R)", " : EQgen(R)", "[| : R; : R |] ==> <,> : EQgen(R)", @@ -481,9 +481,9 @@ fun EQgen_tac ctxt rews i = SELECT_GOAL - (TRY (safe_tac (local_claset_of ctxt)) THEN + (TRY (safe_tac (claset_of ctxt)) THEN resolve_tac ((rews@[refl]) RL ((rews@[refl]) RL [@{thm ssubst_pair}])) i THEN - ALLGOALS (simp_tac (local_simpset_of ctxt)) THEN + ALLGOALS (simp_tac (simpset_of ctxt)) THEN ALLGOALS EQgen_raw_tac) i *} diff -r 63686057cbe8 -r abda97d2deea src/FOL/cladata.ML --- a/src/FOL/cladata.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/FOL/cladata.ML Thu Jul 23 21:13:21 2009 +0200 @@ -19,7 +19,7 @@ structure Cla = ClassicalFun(Classical_Data); structure BasicClassical: BASIC_CLASSICAL = Cla; open BasicClassical; -ML_Antiquote.value "claset" (Scan.succeed "Cla.local_claset_of (ML_Context.the_local_context ())"); +ML_Antiquote.value "claset" (Scan.succeed "Cla.claset_of (ML_Context.the_local_context ())"); (*Propositional rules*) diff -r 63686057cbe8 -r abda97d2deea src/FOL/simpdata.ML --- a/src/FOL/simpdata.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/FOL/simpdata.ML Thu Jul 23 21:13:21 2009 +0200 @@ -85,11 +85,11 @@ end); val defEX_regroup = - Simplifier.simproc (the_context ()) + Simplifier.simproc @{theory} "defined EX" ["EX x. P(x)"] Quantifier1.rearrange_ex; val defALL_regroup = - Simplifier.simproc (the_context ()) + Simplifier.simproc @{theory} "defined ALL" ["ALL x. P(x)"] Quantifier1.rearrange_all; @@ -164,6 +164,6 @@ open Clasimp; ML_Antiquote.value "clasimpset" - (Scan.succeed "Clasimp.local_clasimpset_of (ML_Context.the_local_context ())"); + (Scan.succeed "Clasimp.clasimpset_of (ML_Context.the_local_context ())"); val FOL_css = (FOL_cs, FOL_ss); diff -r 63686057cbe8 -r abda97d2deea src/FOLP/classical.ML --- a/src/FOLP/classical.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/FOLP/classical.ML Thu Jul 23 21:13:21 2009 +0200 @@ -124,11 +124,11 @@ val empty_cs = make_cs{safeIs=[], safeEs=[], hazIs=[], hazEs=[]}; fun print_cs (CS{safeIs,safeEs,hazIs,hazEs,...}) = - (writeln"Introduction rules"; Display.prths hazIs; - writeln"Safe introduction rules"; Display.prths safeIs; - writeln"Elimination rules"; Display.prths hazEs; - writeln"Safe elimination rules"; Display.prths safeEs; - ()); + writeln (cat_lines + (["Introduction rules"] @ map Display.string_of_thm_without_context hazIs @ + ["Safe introduction rules"] @ map Display.string_of_thm_without_context safeIs @ + ["Elimination rules"] @ map Display.string_of_thm_without_context hazEs @ + ["Safe elimination rules"] @ map Display.string_of_thm_without_context safeEs)); fun (CS{safeIs,safeEs,hazIs,hazEs,...}) addSIs ths = make_cs {safeIs=ths@safeIs, safeEs=safeEs, hazIs=hazIs, hazEs=hazEs}; diff -r 63686057cbe8 -r abda97d2deea src/FOLP/ex/Prolog.ML --- a/src/FOLP/ex/Prolog.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/FOLP/ex/Prolog.ML Thu Jul 23 21:13:21 2009 +0200 @@ -13,7 +13,7 @@ by (resolve_tac [appNil,appCons] 1); by (resolve_tac [appNil,appCons] 1); by (resolve_tac [appNil,appCons] 1); -prth (result()); +result(); Goal "app(?x, c:d:Nil, a:b:c:d:Nil)"; by (REPEAT (resolve_tac [appNil,appCons] 1)); diff -r 63686057cbe8 -r abda97d2deea src/FOLP/simp.ML --- a/src/FOLP/simp.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/FOLP/simp.ML Thu Jul 23 21:13:21 2009 +0200 @@ -113,7 +113,7 @@ let fun norm thm = case lhs_of(concl_of thm) of Const(n,_)$_ => n - | _ => (Display.prths normE_thms; error"No constant in lhs of a norm_thm") + | _ => error "No constant in lhs of a norm_thm" in map norm normE_thms end; fun lhs_is_NORM(thm,i) = case lhs_of_eq i thm of @@ -122,7 +122,7 @@ val refl_tac = resolve_tac refl_thms; fun find_res thms thm = - let fun find [] = (Display.prths thms; error"Check Simp_Data") + let fun find [] = error "Check Simp_Data" | find(th::thms) = thm RS th handle THM _ => find thms in find thms end; @@ -249,8 +249,8 @@ fun insert_thm_warn th net = Net.insert_term Thm.eq_thm_prop (concl_of th, th) net handle Net.INSERT => - (writeln"\nDuplicate rewrite or congruence rule:"; Display.print_thm th; - net); + (writeln ("Duplicate rewrite or congruence rule:\n" ^ + Display.string_of_thm_without_context th); net); val insert_thms = fold_rev insert_thm_warn; @@ -275,8 +275,8 @@ fun delete_thm_warn th net = Net.delete_term Thm.eq_thm_prop (concl_of th, th) net handle Net.DELETE => - (writeln"\nNo such rewrite or congruence rule:"; Display.print_thm th; - net); + (writeln ("No such rewrite or congruence rule:\n" ^ + Display.string_of_thm_without_context th); net); val delete_thms = fold_rev delete_thm_warn; @@ -290,9 +290,9 @@ fun delrew(SS{auto_tac,congs,cong_net,mk_simps,simps,simp_net}, thm) = let fun find((p as (th,ths))::ps',ps) = if Thm.eq_thm_prop(thm,th) then (ths,ps@ps') else find(ps',p::ps) - | find([],simps') = (writeln"\nNo such rewrite or congruence rule:"; - Display.print_thm thm; - ([],simps')) + | find([],simps') = + (writeln ("No such rewrite or congruence rule:\n" ^ + Display.string_of_thm_without_context thm); ([], simps')) val (thms,simps') = find(simps,[]) in SS{auto_tac=auto_tac, congs=congs, cong_net=cong_net, mk_simps=mk_simps, simps = simps', simp_net = delete_thms thms simp_net } @@ -311,8 +311,9 @@ fun dest_ss(SS{congs,simps,...}) = (congs, map #1 simps); fun print_ss(SS{congs,simps,...}) = - (writeln"Congruences:"; Display.prths congs; - writeln"Rewrite Rules:"; Display.prths (map #1 simps); ()); + writeln (cat_lines + (["Congruences:"] @ map Display.string_of_thm_without_context congs @ + ["Rewrite Rules:"] @ map (Display.string_of_thm_without_context o #1) simps)); (* Rewriting with conditionals *) @@ -435,7 +436,8 @@ val rwrls = map mk_trans (List.concat(map mk_rew_rules thms)); val anet' = List.foldr lhs_insert_thm anet rwrls in if !tracing andalso not(null new_rws) - then (writeln"Adding rewrites:"; Display.prths new_rws; ()) + then writeln (cat_lines + ("Adding rewrites:" :: map Display.string_of_thm_without_context new_rws)) else (); (ss,thm,anet',anet::ats,cs) end; diff -r 63686057cbe8 -r abda97d2deea src/HOL/Algebra/abstract/Ring2.thy --- a/src/HOL/Algebra/abstract/Ring2.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Algebra/abstract/Ring2.thy Thu Jul 23 21:13:21 2009 +0200 @@ -286,7 +286,7 @@ else SOME rew end; in - val ring_simproc = Simplifier.simproc (the_context ()) "ring" lhss (K proc); + val ring_simproc = Simplifier.simproc @{theory} "ring" lhss (K proc); end; *} diff -r 63686057cbe8 -r abda97d2deea src/HOL/Auth/Message.thy --- a/src/HOL/Auth/Message.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Auth/Message.thy Thu Jul 23 21:13:21 2009 +0200 @@ -856,6 +856,8 @@ Abstraction over i is ESSENTIAL: it delays the dereferencing of claset DEPENDS UPON "X" REFERRING TO THE FRADULENT MESSAGE *) +fun impOfSubs th = th RSN (2, @{thm rev_subsetD}) + (*Apply rules to break down assumptions of the form Y \ parts(insert X H) and Y \ analz(insert X H) *) @@ -937,15 +939,15 @@ lemmas Fake_parts_sing_imp_Un = Fake_parts_sing [THEN [2] rev_subsetD] method_setup spy_analz = {* - Scan.succeed (SIMPLE_METHOD' o Message.spy_analz_tac o local_clasimpset_of) *} + Scan.succeed (SIMPLE_METHOD' o Message.spy_analz_tac o clasimpset_of) *} "for proving the Fake case when analz is involved" method_setup atomic_spy_analz = {* - Scan.succeed (SIMPLE_METHOD' o Message.atomic_spy_analz_tac o local_clasimpset_of) *} + Scan.succeed (SIMPLE_METHOD' o Message.atomic_spy_analz_tac o clasimpset_of) *} "for debugging spy_analz" method_setup Fake_insert_simp = {* - Scan.succeed (SIMPLE_METHOD' o Message.Fake_insert_simp_tac o local_simpset_of) *} + Scan.succeed (SIMPLE_METHOD' o Message.Fake_insert_simp_tac o simpset_of) *} "for debugging spy_analz" end diff -r 63686057cbe8 -r abda97d2deea src/HOL/Auth/Public.thy --- a/src/HOL/Auth/Public.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Auth/Public.thy Thu Jul 23 21:13:21 2009 +0200 @@ -407,7 +407,7 @@ (*Tactic for possibility theorems*) fun possibility_tac ctxt = REPEAT (*omit used_Says so that Nonces start from different traces!*) - (ALLGOALS (simp_tac (local_simpset_of ctxt delsimps [@{thm used_Says}])) + (ALLGOALS (simp_tac (simpset_of ctxt delsimps [@{thm used_Says}])) THEN REPEAT_FIRST (eq_assume_tac ORELSE' resolve_tac [refl, conjI, @{thm Nonce_supply}])) @@ -416,7 +416,7 @@ nonces and keys initially*) fun basic_possibility_tac ctxt = REPEAT - (ALLGOALS (asm_simp_tac (local_simpset_of ctxt setSolver safe_solver)) + (ALLGOALS (asm_simp_tac (simpset_of ctxt setSolver safe_solver)) THEN REPEAT_FIRST (resolve_tac [refl, conjI])) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Auth/Shared.thy --- a/src/HOL/Auth/Shared.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Auth/Shared.thy Thu Jul 23 21:13:21 2009 +0200 @@ -204,7 +204,7 @@ such as Nonce ?N \ used evs that match Nonce_supply*) fun possibility_tac ctxt = (REPEAT - (ALLGOALS (simp_tac (local_simpset_of ctxt + (ALLGOALS (simp_tac (simpset_of ctxt delsimps [@{thm used_Says}, @{thm used_Notes}, @{thm used_Gets}] setSolver safe_solver)) THEN @@ -215,7 +215,7 @@ nonces and keys initially*) fun basic_possibility_tac ctxt = REPEAT - (ALLGOALS (asm_simp_tac (local_simpset_of ctxt setSolver safe_solver)) + (ALLGOALS (asm_simp_tac (simpset_of ctxt setSolver safe_solver)) THEN REPEAT_FIRST (resolve_tac [refl, conjI])) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Auth/Smartcard/ShoupRubin.thy --- a/src/HOL/Auth/Smartcard/ShoupRubin.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Auth/Smartcard/ShoupRubin.thy Thu Jul 23 21:13:21 2009 +0200 @@ -757,7 +757,7 @@ (*SR9*) dresolve_tac [@{thm Gets_imp_knows_Spy_parts_Snd}] 19 THEN (*Oops1*) dresolve_tac [@{thm Outpts_B_Card_form_7}] 25 THEN (*Oops2*) dresolve_tac [@{thm Outpts_A_Card_form_10}] 27 THEN - (*Base*) (force_tac (local_clasimpset_of ctxt)) 1 + (*Base*) (force_tac (clasimpset_of ctxt)) 1 val analz_prepare_tac = prepare_tac THEN diff -r 63686057cbe8 -r abda97d2deea src/HOL/Auth/Smartcard/ShoupRubinBella.thy --- a/src/HOL/Auth/Smartcard/ShoupRubinBella.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Auth/Smartcard/ShoupRubinBella.thy Thu Jul 23 21:13:21 2009 +0200 @@ -755,7 +755,7 @@ fun prepare_tac ctxt = (*SR_U8*) forward_tac [@{thm Outpts_B_Card_form_7}] 14 THEN - (*SR_U8*) clarify_tac (local_claset_of ctxt) 15 THEN + (*SR_U8*) clarify_tac (claset_of ctxt) 15 THEN (*SR_U9*) forward_tac [@{thm Outpts_A_Card_form_4}] 16 THEN (*SR_U11*) forward_tac [@{thm Outpts_A_Card_form_10}] 21 @@ -765,7 +765,7 @@ (*SR_U9*) dresolve_tac [@{thm Gets_imp_knows_Spy_parts_Snd}] 19 THEN (*Oops1*) dresolve_tac [@{thm Outpts_B_Card_form_7}] 25 THEN (*Oops2*) dresolve_tac [@{thm Outpts_A_Card_form_10}] 27 THEN - (*Base*) (force_tac (local_clasimpset_of ctxt)) 1 + (*Base*) (force_tac (clasimpset_of ctxt)) 1 fun analz_prepare_tac ctxt = prepare_tac ctxt THEN diff -r 63686057cbe8 -r abda97d2deea src/HOL/Auth/Smartcard/Smartcard.thy --- a/src/HOL/Auth/Smartcard/Smartcard.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Auth/Smartcard/Smartcard.thy Thu Jul 23 21:13:21 2009 +0200 @@ -375,7 +375,7 @@ such as Nonce ?N \ used evs that match Nonce_supply*) fun possibility_tac ctxt = (REPEAT - (ALLGOALS (simp_tac (local_simpset_of ctxt + (ALLGOALS (simp_tac (simpset_of ctxt delsimps [@{thm used_Says}, @{thm used_Notes}, @{thm used_Gets}, @{thm used_Inputs}, @{thm used_C_Gets}, @{thm used_Outpts}, @{thm used_A_Gets}] setSolver safe_solver)) @@ -387,7 +387,7 @@ nonces and keys initially*) fun basic_possibility_tac ctxt = REPEAT - (ALLGOALS (asm_simp_tac (local_simpset_of ctxt setSolver safe_solver)) + (ALLGOALS (asm_simp_tac (simpset_of ctxt setSolver safe_solver)) THEN REPEAT_FIRST (resolve_tac [refl, conjI])) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Bali/Basis.thy --- a/src/HOL/Bali/Basis.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Bali/Basis.thy Thu Jul 23 21:13:21 2009 +0200 @@ -229,7 +229,7 @@ ML {* fun sum3_instantiate ctxt thm = map (fn s => - simplify (Simplifier.local_simpset_of ctxt delsimps[@{thm not_None_eq}]) + simplify (simpset_of ctxt delsimps[@{thm not_None_eq}]) (read_instantiate ctxt [(("t", 0), "In" ^ s ^ " ?x")] thm)) ["1l","2","3","1r"] *} (* e.g. lemmas is_stmt_rews = is_stmt_def [of "In1l x", simplified] *) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Bali/DeclConcepts.thy --- a/src/HOL/Bali/DeclConcepts.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Bali/DeclConcepts.thy Thu Jul 23 21:13:21 2009 +0200 @@ -648,7 +648,7 @@ G \Method old member_of superNew \ \ G\new overrides\<^sub>S old" -| Indirect: "\G\new overrides\<^sub>S inter; G\inter overrides\<^sub>S old\ +| Indirect: "\G\new overrides\<^sub>S intr; G\intr overrides\<^sub>S old\ \ G\new overrides\<^sub>S old" text {* Dynamic overriding (used during the typecheck of the compiler) *} @@ -668,7 +668,7 @@ G\resTy new \ resTy old \ \ G\new overrides old" -| Indirect: "\G\new overrides inter; G\inter overrides old\ +| Indirect: "\G\new overrides intr; G\intr overrides old\ \ G\new overrides old" syntax diff -r 63686057cbe8 -r abda97d2deea src/HOL/Code_Eval.thy --- a/src/HOL/Code_Eval.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Code_Eval.thy Thu Jul 23 21:13:21 2009 +0200 @@ -32,7 +32,7 @@ \ 'a \ (unit \ term) \ 'b \ (unit \ term)" where "valapp f x = (fst f (fst x), \u. App (snd f ()) (snd x ()))" -lemma valapp_code [code, code_inline]: +lemma valapp_code [code, code_unfold]: "valapp (f, tf) (x, tx) = (f x, \u. App (tf ()) (tx ()))" by (simp only: valapp_def fst_conv snd_conv) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Code_Numeral.thy --- a/src/HOL/Code_Numeral.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Code_Numeral.thy Thu Jul 23 21:13:21 2009 +0200 @@ -176,13 +176,13 @@ end -lemma zero_code_numeral_code [code_inline, code]: +lemma zero_code_numeral_code [code, code_unfold]: "(0\code_numeral) = Numeral0" by (simp add: number_of_code_numeral_def Pls_def) lemma [code_post]: "Numeral0 = (0\code_numeral)" using zero_code_numeral_code .. -lemma one_code_numeral_code [code_inline, code]: +lemma one_code_numeral_code [code, code_unfold]: "(1\code_numeral) = Numeral1" by (simp add: number_of_code_numeral_def Pls_def Bit1_def) lemma [code_post]: "Numeral1 = (1\code_numeral)" diff -r 63686057cbe8 -r abda97d2deea src/HOL/Complete_Lattice.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complete_Lattice.thy Thu Jul 23 21:13:21 2009 +0200 @@ -0,0 +1,794 @@ +(* Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel; Florian Haftmann, TU Muenchen *) + +header {* Complete lattices, with special focus on sets *} + +theory Complete_Lattice +imports Set +begin + +notation + less_eq (infix "\" 50) and + less (infix "\" 50) and + inf (infixl "\" 70) and + sup (infixl "\" 65) + + +subsection {* Abstract complete lattices *} + +class complete_lattice = lattice + bot + top + + fixes Inf :: "'a set \ 'a" ("\_" [900] 900) + and Sup :: "'a set \ 'a" ("\_" [900] 900) + assumes Inf_lower: "x \ A \ \A \ x" + and Inf_greatest: "(\x. x \ A \ z \ x) \ z \ \A" + assumes Sup_upper: "x \ A \ x \ \A" + and Sup_least: "(\x. x \ A \ x \ z) \ \A \ z" +begin + +lemma Inf_Sup: "\A = \{b. \a \ A. b \ a}" + by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) + +lemma Sup_Inf: "\A = \{b. \a \ A. a \ b}" + by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) + +lemma Inf_Univ: "\UNIV = \{}" + unfolding Sup_Inf by auto + +lemma Sup_Univ: "\UNIV = \{}" + unfolding Inf_Sup by auto + +lemma Inf_insert: "\insert a A = a \ \A" + by (auto intro: le_infI le_infI1 le_infI2 antisym Inf_greatest Inf_lower) + +lemma Sup_insert: "\insert a A = a \ \A" + by (auto intro: le_supI le_supI1 le_supI2 antisym Sup_least Sup_upper) + +lemma Inf_singleton [simp]: + "\{a} = a" + by (auto intro: antisym Inf_lower Inf_greatest) + +lemma Sup_singleton [simp]: + "\{a} = a" + by (auto intro: antisym Sup_upper Sup_least) + +lemma Inf_insert_simp: + "\insert a A = (if A = {} then a else a \ \A)" + by (cases "A = {}") (simp_all, simp add: Inf_insert) + +lemma Sup_insert_simp: + "\insert a A = (if A = {} then a else a \ \A)" + by (cases "A = {}") (simp_all, simp add: Sup_insert) + +lemma Inf_binary: + "\{a, b} = a \ b" + by (auto simp add: Inf_insert_simp) + +lemma Sup_binary: + "\{a, b} = a \ b" + by (auto simp add: Sup_insert_simp) + +lemma bot_def: + "bot = \{}" + by (auto intro: antisym Sup_least) + +lemma top_def: + "top = \{}" + by (auto intro: antisym Inf_greatest) + +lemma sup_bot [simp]: + "x \ bot = x" + using bot_least [of x] by (simp add: le_iff_sup sup_commute) + +lemma inf_top [simp]: + "x \ top = x" + using top_greatest [of x] by (simp add: le_iff_inf inf_commute) + +definition SUPR :: "'b set \ ('b \ 'a) \ 'a" where + "SUPR A f = \ (f ` A)" + +definition INFI :: "'b set \ ('b \ 'a) \ 'a" where + "INFI A f = \ (f ` A)" + +end + +syntax + "_SUP1" :: "pttrns => 'b => 'b" ("(3SUP _./ _)" [0, 10] 10) + "_SUP" :: "pttrn => 'a set => 'b => 'b" ("(3SUP _:_./ _)" [0, 10] 10) + "_INF1" :: "pttrns => 'b => 'b" ("(3INF _./ _)" [0, 10] 10) + "_INF" :: "pttrn => 'a set => 'b => 'b" ("(3INF _:_./ _)" [0, 10] 10) + +translations + "SUP x y. B" == "SUP x. SUP y. B" + "SUP x. B" == "CONST SUPR CONST UNIV (%x. B)" + "SUP x. B" == "SUP x:CONST UNIV. B" + "SUP x:A. B" == "CONST SUPR A (%x. B)" + "INF x y. B" == "INF x. INF y. B" + "INF x. B" == "CONST INFI CONST UNIV (%x. B)" + "INF x. B" == "INF x:CONST UNIV. B" + "INF x:A. B" == "CONST INFI A (%x. B)" + +print_translation {* [ +Syntax.preserve_binder_abs2_tr' @{const_syntax SUPR} "_SUP", +Syntax.preserve_binder_abs2_tr' @{const_syntax INFI} "_INF" +] *} -- {* to avoid eta-contraction of body *} + +context complete_lattice +begin + +lemma le_SUPI: "i : A \ M i \ (SUP i:A. M i)" + by (auto simp add: SUPR_def intro: Sup_upper) + +lemma SUP_leI: "(\i. i : A \ M i \ u) \ (SUP i:A. M i) \ u" + by (auto simp add: SUPR_def intro: Sup_least) + +lemma INF_leI: "i : A \ (INF i:A. M i) \ M i" + by (auto simp add: INFI_def intro: Inf_lower) + +lemma le_INFI: "(\i. i : A \ u \ M i) \ u \ (INF i:A. M i)" + by (auto simp add: INFI_def intro: Inf_greatest) + +lemma SUP_const[simp]: "A \ {} \ (SUP i:A. M) = M" + by (auto intro: antisym SUP_leI le_SUPI) + +lemma INF_const[simp]: "A \ {} \ (INF i:A. M) = M" + by (auto intro: antisym INF_leI le_INFI) + +end + + +subsection {* @{typ bool} and @{typ "_ \ _"} as complete lattice *} + +instantiation bool :: complete_lattice +begin + +definition + Inf_bool_def: "\A \ (\x\A. x)" + +definition + Sup_bool_def: "\A \ (\x\A. x)" + +instance proof +qed (auto simp add: Inf_bool_def Sup_bool_def le_bool_def) + +end + +lemma Inf_empty_bool [simp]: + "\{}" + unfolding Inf_bool_def by auto + +lemma not_Sup_empty_bool [simp]: + "\ \{}" + unfolding Sup_bool_def by auto + +lemma INFI_bool_eq: + "INFI = Ball" +proof (rule ext)+ + fix A :: "'a set" + fix P :: "'a \ bool" + show "(INF x:A. P x) \ (\x \ A. P x)" + by (auto simp add: Ball_def INFI_def Inf_bool_def) +qed + +lemma SUPR_bool_eq: + "SUPR = Bex" +proof (rule ext)+ + fix A :: "'a set" + fix P :: "'a \ bool" + show "(SUP x:A. P x) \ (\x \ A. P x)" + by (auto simp add: Bex_def SUPR_def Sup_bool_def) +qed + +instantiation "fun" :: (type, complete_lattice) complete_lattice +begin + +definition + Inf_fun_def [code del]: "\A = (\x. \{y. \f\A. y = f x})" + +definition + Sup_fun_def [code del]: "\A = (\x. \{y. \f\A. y = f x})" + +instance proof +qed (auto simp add: Inf_fun_def Sup_fun_def le_fun_def + intro: Inf_lower Sup_upper Inf_greatest Sup_least) + +end + +lemma Inf_empty_fun: + "\{} = (\_. \{})" + by (simp add: Inf_fun_def) + +lemma Sup_empty_fun: + "\{} = (\_. \{})" + by (simp add: Sup_fun_def) + + +subsection {* Union *} + +definition Union :: "'a set set \ 'a set" where + Sup_set_eq [symmetric]: "Union S = \S" + +notation (xsymbols) + Union ("\_" [90] 90) + +lemma Union_eq: + "\A = {x. \B \ A. x \ B}" +proof (rule set_ext) + fix x + have "(\Q\{P. \B\A. P \ x \ B}. Q) \ (\B\A. x \ B)" + by auto + then show "x \ \A \ x \ {x. \B\A. x \ B}" + by (simp add: Sup_set_eq [symmetric] Sup_fun_def Sup_bool_def) (simp add: mem_def) +qed + +lemma Union_iff [simp, noatp]: + "A \ \C \ (\X\C. A\X)" + by (unfold Union_eq) blast + +lemma UnionI [intro]: + "X \ C \ A \ X \ A \ \C" + -- {* The order of the premises presupposes that @{term C} is rigid; + @{term A} may be flexible. *} + by auto + +lemma UnionE [elim!]: + "A \ \C \ (\X. A\X \ X\C \ R) \ R" + by auto + +lemma Union_upper: "B \ A ==> B \ Union A" + by (iprover intro: subsetI UnionI) + +lemma Union_least: "(!!X. X \ A ==> X \ C) ==> Union A \ C" + by (iprover intro: subsetI elim: UnionE dest: subsetD) + +lemma Un_eq_Union: "A \ B = \{A, B}" + by blast + +lemma Union_empty [simp]: "Union({}) = {}" + by blast + +lemma Union_UNIV [simp]: "Union UNIV = UNIV" + by blast + +lemma Union_insert [simp]: "Union (insert a B) = a \ \B" + by blast + +lemma Union_Un_distrib [simp]: "\(A Un B) = \A \ \B" + by blast + +lemma Union_Int_subset: "\(A \ B) \ \A \ \B" + by blast + +lemma Union_empty_conv [simp,noatp]: "(\A = {}) = (\x\A. x = {})" + by blast + +lemma empty_Union_conv [simp,noatp]: "({} = \A) = (\x\A. x = {})" + by blast + +lemma Union_disjoint: "(\C \ A = {}) = (\B\C. B \ A = {})" + by blast + +lemma subset_Pow_Union: "A \ Pow (\A)" + by blast + +lemma Union_Pow_eq [simp]: "\(Pow A) = A" + by blast + +lemma Union_mono: "A \ B ==> \A \ \B" + by blast + + +subsection {* Unions of families *} + +definition UNION :: "'a set \ ('a \ 'b set) \ 'b set" where + SUPR_set_eq [symmetric]: "UNION S f = (SUP x:S. f x)" + +syntax + "@UNION1" :: "pttrns => 'b set => 'b set" ("(3UN _./ _)" [0, 10] 10) + "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3UN _:_./ _)" [0, 10] 10) + +syntax (xsymbols) + "@UNION1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) + "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 10] 10) + +syntax (latex output) + "@UNION1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) + "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 10] 10) + +translations + "UN x y. B" == "UN x. UN y. B" + "UN x. B" == "CONST UNION CONST UNIV (%x. B)" + "UN x. B" == "UN x:CONST UNIV. B" + "UN x:A. B" == "CONST UNION A (%x. B)" + +text {* + Note the difference between ordinary xsymbol syntax of indexed + unions and intersections (e.g.\ @{text"\a\<^isub>1\A\<^isub>1. B"}) + and their \LaTeX\ rendition: @{term"\a\<^isub>1\A\<^isub>1. B"}. The + former does not make the index expression a subscript of the + union/intersection symbol because this leads to problems with nested + subscripts in Proof General. +*} + +print_translation {* [ +Syntax.preserve_binder_abs2_tr' @{const_syntax UNION} "@UNION" +] *} -- {* to avoid eta-contraction of body *} + +lemma UNION_eq_Union_image: + "(\x\A. B x) = \(B`A)" + by (simp add: SUPR_def SUPR_set_eq [symmetric] Sup_set_eq) + +lemma Union_def: + "\S = (\x\S. x)" + by (simp add: UNION_eq_Union_image image_def) + +lemma UNION_def [noatp]: + "(\x\A. B x) = {y. \x\A. y \ B x}" + by (auto simp add: UNION_eq_Union_image Union_eq) + +lemma Union_image_eq [simp]: + "\(B`A) = (\x\A. B x)" + by (rule sym) (fact UNION_eq_Union_image) + +lemma UN_iff [simp]: "(b: (UN x:A. B x)) = (EX x:A. b: B x)" + by (unfold UNION_def) blast + +lemma UN_I [intro]: "a:A ==> b: B a ==> b: (UN x:A. B x)" + -- {* The order of the premises presupposes that @{term A} is rigid; + @{term b} may be flexible. *} + by auto + +lemma UN_E [elim!]: "b : (UN x:A. B x) ==> (!!x. x:A ==> b: B x ==> R) ==> R" + by (unfold UNION_def) blast + +lemma UN_cong [cong]: + "A = B ==> (!!x. x:B ==> C x = D x) ==> (UN x:A. C x) = (UN x:B. D x)" + by (simp add: UNION_def) + +lemma strong_UN_cong: + "A = B ==> (!!x. x:B =simp=> C x = D x) ==> (UN x:A. C x) = (UN x:B. D x)" + by (simp add: UNION_def simp_implies_def) + +lemma image_eq_UN: "f`A = (UN x:A. {f x})" + by blast + +lemma UN_upper: "a \ A ==> B a \ (\x\A. B x)" + by blast + +lemma UN_least: "(!!x. x \ A ==> B x \ C) ==> (\x\A. B x) \ C" + by (iprover intro: subsetI elim: UN_E dest: subsetD) + +lemma Collect_bex_eq [noatp]: "{x. \y\A. P x y} = (\y\A. {x. P x y})" + by blast + +lemma UN_insert_distrib: "u \ A ==> (\x\A. insert a (B x)) = insert a (\x\A. B x)" + by blast + +lemma UN_empty [simp,noatp]: "(\x\{}. B x) = {}" + by blast + +lemma UN_empty2 [simp]: "(\x\A. {}) = {}" + by blast + +lemma UN_singleton [simp]: "(\x\A. {x}) = A" + by blast + +lemma UN_absorb: "k \ I ==> A k \ (\i\I. A i) = (\i\I. A i)" + by auto + +lemma UN_insert [simp]: "(\x\insert a A. B x) = B a \ UNION A B" + by blast + +lemma UN_Un[simp]: "(\i \ A \ B. M i) = (\i\A. M i) \ (\i\B. M i)" + by blast + +lemma UN_UN_flatten: "(\x \ (\y\A. B y). C x) = (\y\A. \x\B y. C x)" + by blast + +lemma UN_subset_iff: "((\i\I. A i) \ B) = (\i\I. A i \ B)" + by blast + +lemma image_Union: "f ` \S = (\x\S. f ` x)" + by blast + +lemma UN_constant [simp]: "(\y\A. c) = (if A = {} then {} else c)" + by auto + +lemma UN_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" + by blast + +lemma UNION_empty_conv[simp]: + "({} = (UN x:A. B x)) = (\x\A. B x = {})" + "((UN x:A. B x) = {}) = (\x\A. B x = {})" +by blast+ + +lemma Collect_ex_eq [noatp]: "{x. \y. P x y} = (\y. {x. P x y})" + by blast + +lemma ball_UN: "(\z \ UNION A B. P z) = (\x\A. \z \ B x. P z)" + by blast + +lemma bex_UN: "(\z \ UNION A B. P z) = (\x\A. \z\B x. P z)" + by blast + +lemma Un_eq_UN: "A \ B = (\b. if b then A else B)" + by (auto simp add: split_if_mem2) + +lemma UN_bool_eq: "(\b::bool. A b) = (A True \ A False)" + by (auto intro: bool_contrapos) + +lemma UN_Pow_subset: "(\x\A. Pow (B x)) \ Pow (\x\A. B x)" + by blast + +lemma UN_mono: + "A \ B ==> (!!x. x \ A ==> f x \ g x) ==> + (\x\A. f x) \ (\x\B. g x)" + by (blast dest: subsetD) + +lemma vimage_Union: "f -` (Union A) = (UN X:A. f -` X)" + by blast + +lemma vimage_UN: "f-`(UN x:A. B x) = (UN x:A. f -` B x)" + by blast + +lemma vimage_eq_UN: "f-`B = (UN y: B. f-`{y})" + -- {* NOT suitable for rewriting *} + by blast + +lemma image_UN: "(f ` (UNION A B)) = (UN x:A.(f ` (B x)))" +by blast + + +subsection {* Inter *} + +definition Inter :: "'a set set \ 'a set" where + Inf_set_eq [symmetric]: "Inter S = \S" + +notation (xsymbols) + Inter ("\_" [90] 90) + +lemma Inter_eq [code del]: + "\A = {x. \B \ A. x \ B}" +proof (rule set_ext) + fix x + have "(\Q\{P. \B\A. P \ x \ B}. Q) \ (\B\A. x \ B)" + by auto + then show "x \ \A \ x \ {x. \B \ A. x \ B}" + by (simp add: Inf_fun_def Inf_bool_def Inf_set_eq [symmetric]) (simp add: mem_def) +qed + +lemma Inter_iff [simp,noatp]: "(A : Inter C) = (ALL X:C. A:X)" + by (unfold Inter_eq) blast + +lemma InterI [intro!]: "(!!X. X:C ==> A:X) ==> A : Inter C" + by (simp add: Inter_eq) + +text {* + \medskip A ``destruct'' rule -- every @{term X} in @{term C} + contains @{term A} as an element, but @{prop "A:X"} can hold when + @{prop "X:C"} does not! This rule is analogous to @{text spec}. +*} + +lemma InterD [elim]: "A : Inter C ==> X:C ==> A:X" + by auto + +lemma InterE [elim]: "A : Inter C ==> (X~:C ==> R) ==> (A:X ==> R) ==> R" + -- {* ``Classical'' elimination rule -- does not require proving + @{prop "X:C"}. *} + by (unfold Inter_eq) blast + +lemma Inter_lower: "B \ A ==> Inter A \ B" + by blast + +lemma Inter_subset: + "[| !!X. X \ A ==> X \ B; A ~= {} |] ==> \A \ B" + by blast + +lemma Inter_greatest: "(!!X. X \ A ==> C \ X) ==> C \ Inter A" + by (iprover intro: InterI subsetI dest: subsetD) + +lemma Int_eq_Inter: "A \ B = \{A, B}" + by blast + +lemma Inter_empty [simp]: "\{} = UNIV" + by blast + +lemma Inter_UNIV [simp]: "\UNIV = {}" + by blast + +lemma Inter_insert [simp]: "\(insert a B) = a \ \B" + by blast + +lemma Inter_Un_subset: "\A \ \B \ \(A \ B)" + by blast + +lemma Inter_Un_distrib: "\(A \ B) = \A \ \B" + by blast + +lemma Inter_UNIV_conv [simp,noatp]: + "(\A = UNIV) = (\x\A. x = UNIV)" + "(UNIV = \A) = (\x\A. x = UNIV)" + by blast+ + +lemma Inter_anti_mono: "B \ A ==> \A \ \B" + by blast + + +subsection {* Intersections of families *} + +definition INTER :: "'a set \ ('a \ 'b set) \ 'b set" where + INFI_set_eq [symmetric]: "INTER S f = (INF x:S. f x)" + +syntax + "@INTER1" :: "pttrns => 'b set => 'b set" ("(3INT _./ _)" [0, 10] 10) + "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3INT _:_./ _)" [0, 10] 10) + +syntax (xsymbols) + "@INTER1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) + "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 10] 10) + +syntax (latex output) + "@INTER1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) + "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 10] 10) + +translations + "INT x y. B" == "INT x. INT y. B" + "INT x. B" == "CONST INTER CONST UNIV (%x. B)" + "INT x. B" == "INT x:CONST UNIV. B" + "INT x:A. B" == "CONST INTER A (%x. B)" + +print_translation {* [ +Syntax.preserve_binder_abs2_tr' @{const_syntax INTER} "@INTER" +] *} -- {* to avoid eta-contraction of body *} + +lemma INTER_eq_Inter_image: + "(\x\A. B x) = \(B`A)" + by (simp add: INFI_def INFI_set_eq [symmetric] Inf_set_eq) + +lemma Inter_def: + "\S = (\x\S. x)" + by (simp add: INTER_eq_Inter_image image_def) + +lemma INTER_def: + "(\x\A. B x) = {y. \x\A. y \ B x}" + by (auto simp add: INTER_eq_Inter_image Inter_eq) + +lemma Inter_image_eq [simp]: + "\(B`A) = (\x\A. B x)" + by (rule sym) (fact INTER_eq_Inter_image) + +lemma INT_iff [simp]: "(b: (INT x:A. B x)) = (ALL x:A. b: B x)" + by (unfold INTER_def) blast + +lemma INT_I [intro!]: "(!!x. x:A ==> b: B x) ==> b : (INT x:A. B x)" + by (unfold INTER_def) blast + +lemma INT_D [elim]: "b : (INT x:A. B x) ==> a:A ==> b: B a" + by auto + +lemma INT_E [elim]: "b : (INT x:A. B x) ==> (b: B a ==> R) ==> (a~:A ==> R) ==> R" + -- {* "Classical" elimination -- by the Excluded Middle on @{prop "a:A"}. *} + by (unfold INTER_def) blast + +lemma INT_cong [cong]: + "A = B ==> (!!x. x:B ==> C x = D x) ==> (INT x:A. C x) = (INT x:B. D x)" + by (simp add: INTER_def) + +lemma Collect_ball_eq: "{x. \y\A. P x y} = (\y\A. {x. P x y})" + by blast + +lemma Collect_all_eq: "{x. \y. P x y} = (\y. {x. P x y})" + by blast + +lemma INT_lower: "a \ A ==> (\x\A. B x) \ B a" + by blast + +lemma INT_greatest: "(!!x. x \ A ==> C \ B x) ==> C \ (\x\A. B x)" + by (iprover intro: INT_I subsetI dest: subsetD) + +lemma INT_empty [simp]: "(\x\{}. B x) = UNIV" + by blast + +lemma INT_absorb: "k \ I ==> A k \ (\i\I. A i) = (\i\I. A i)" + by blast + +lemma INT_subset_iff: "(B \ (\i\I. A i)) = (\i\I. B \ A i)" + by blast + +lemma INT_insert [simp]: "(\x \ insert a A. B x) = B a \ INTER A B" + by blast + +lemma INT_Un: "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" + by blast + +lemma INT_insert_distrib: + "u \ A ==> (\x\A. insert a (B x)) = insert a (\x\A. B x)" + by blast + +lemma INT_constant [simp]: "(\y\A. c) = (if A = {} then UNIV else c)" + by auto + +lemma INT_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" + -- {* Look: it has an \emph{existential} quantifier *} + by blast + +lemma INTER_UNIV_conv[simp]: + "(UNIV = (INT x:A. B x)) = (\x\A. B x = UNIV)" + "((INT x:A. B x) = UNIV) = (\x\A. B x = UNIV)" +by blast+ + +lemma INT_bool_eq: "(\b::bool. A b) = (A True \ A False)" + by (auto intro: bool_induct) + +lemma Pow_INT_eq: "Pow (\x\A. B x) = (\x\A. Pow (B x))" + by blast + +lemma INT_anti_mono: + "B \ A ==> (!!x. x \ A ==> f x \ g x) ==> + (\x\A. f x) \ (\x\A. g x)" + -- {* The last inclusion is POSITIVE! *} + by (blast dest: subsetD) + +lemma vimage_INT: "f-`(INT x:A. B x) = (INT x:A. f -` B x)" + by blast + + +subsection {* Distributive laws *} + +lemma Int_Union: "A \ \B = (\C\B. A \ C)" + by blast + +lemma Int_Union2: "\B \ A = (\C\B. C \ A)" + by blast + +lemma Un_Union_image: "(\x\C. A x \ B x) = \(A`C) \ \(B`C)" + -- {* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: *} + -- {* Union of a family of unions *} + by blast + +lemma UN_Un_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" + -- {* Equivalent version *} + by blast + +lemma Un_Inter: "A \ \B = (\C\B. A \ C)" + by blast + +lemma Int_Inter_image: "(\x\C. A x \ B x) = \(A`C) \ \(B`C)" + by blast + +lemma INT_Int_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" + -- {* Equivalent version *} + by blast + +lemma Int_UN_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" + -- {* Halmos, Naive Set Theory, page 35. *} + by blast + +lemma Un_INT_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" + by blast + +lemma Int_UN_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" + by blast + +lemma Un_INT_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" + by blast + + +subsection {* Complement *} + +lemma Compl_UN [simp]: "-(\x\A. B x) = (\x\A. -B x)" + by blast + +lemma Compl_INT [simp]: "-(\x\A. B x) = (\x\A. -B x)" + by blast + + +subsection {* Miniscoping and maxiscoping *} + +text {* \medskip Miniscoping: pushing in quantifiers and big Unions + and Intersections. *} + +lemma UN_simps [simp]: + "!!a B C. (UN x:C. insert a (B x)) = (if C={} then {} else insert a (UN x:C. B x))" + "!!A B C. (UN x:C. A x Un B) = ((if C={} then {} else (UN x:C. A x) Un B))" + "!!A B C. (UN x:C. A Un B x) = ((if C={} then {} else A Un (UN x:C. B x)))" + "!!A B C. (UN x:C. A x Int B) = ((UN x:C. A x) Int B)" + "!!A B C. (UN x:C. A Int B x) = (A Int (UN x:C. B x))" + "!!A B C. (UN x:C. A x - B) = ((UN x:C. A x) - B)" + "!!A B C. (UN x:C. A - B x) = (A - (INT x:C. B x))" + "!!A B. (UN x: Union A. B x) = (UN y:A. UN x:y. B x)" + "!!A B C. (UN z: UNION A B. C z) = (UN x:A. UN z: B(x). C z)" + "!!A B f. (UN x:f`A. B x) = (UN a:A. B (f a))" + by auto + +lemma INT_simps [simp]: + "!!A B C. (INT x:C. A x Int B) = (if C={} then UNIV else (INT x:C. A x) Int B)" + "!!A B C. (INT x:C. A Int B x) = (if C={} then UNIV else A Int (INT x:C. B x))" + "!!A B C. (INT x:C. A x - B) = (if C={} then UNIV else (INT x:C. A x) - B)" + "!!A B C. (INT x:C. A - B x) = (if C={} then UNIV else A - (UN x:C. B x))" + "!!a B C. (INT x:C. insert a (B x)) = insert a (INT x:C. B x)" + "!!A B C. (INT x:C. A x Un B) = ((INT x:C. A x) Un B)" + "!!A B C. (INT x:C. A Un B x) = (A Un (INT x:C. B x))" + "!!A B. (INT x: Union A. B x) = (INT y:A. INT x:y. B x)" + "!!A B C. (INT z: UNION A B. C z) = (INT x:A. INT z: B(x). C z)" + "!!A B f. (INT x:f`A. B x) = (INT a:A. B (f a))" + by auto + +lemma ball_simps [simp,noatp]: + "!!A P Q. (ALL x:A. P x | Q) = ((ALL x:A. P x) | Q)" + "!!A P Q. (ALL x:A. P | Q x) = (P | (ALL x:A. Q x))" + "!!A P Q. (ALL x:A. P --> Q x) = (P --> (ALL x:A. Q x))" + "!!A P Q. (ALL x:A. P x --> Q) = ((EX x:A. P x) --> Q)" + "!!P. (ALL x:{}. P x) = True" + "!!P. (ALL x:UNIV. P x) = (ALL x. P x)" + "!!a B P. (ALL x:insert a B. P x) = (P a & (ALL x:B. P x))" + "!!A P. (ALL x:Union A. P x) = (ALL y:A. ALL x:y. P x)" + "!!A B P. (ALL x: UNION A B. P x) = (ALL a:A. ALL x: B a. P x)" + "!!P Q. (ALL x:Collect Q. P x) = (ALL x. Q x --> P x)" + "!!A P f. (ALL x:f`A. P x) = (ALL x:A. P (f x))" + "!!A P. (~(ALL x:A. P x)) = (EX x:A. ~P x)" + by auto + +lemma bex_simps [simp,noatp]: + "!!A P Q. (EX x:A. P x & Q) = ((EX x:A. P x) & Q)" + "!!A P Q. (EX x:A. P & Q x) = (P & (EX x:A. Q x))" + "!!P. (EX x:{}. P x) = False" + "!!P. (EX x:UNIV. P x) = (EX x. P x)" + "!!a B P. (EX x:insert a B. P x) = (P(a) | (EX x:B. P x))" + "!!A P. (EX x:Union A. P x) = (EX y:A. EX x:y. P x)" + "!!A B P. (EX x: UNION A B. P x) = (EX a:A. EX x:B a. P x)" + "!!P Q. (EX x:Collect Q. P x) = (EX x. Q x & P x)" + "!!A P f. (EX x:f`A. P x) = (EX x:A. P (f x))" + "!!A P. (~(EX x:A. P x)) = (ALL x:A. ~P x)" + by auto + +lemma ball_conj_distrib: + "(ALL x:A. P x & Q x) = ((ALL x:A. P x) & (ALL x:A. Q x))" + by blast + +lemma bex_disj_distrib: + "(EX x:A. P x | Q x) = ((EX x:A. P x) | (EX x:A. Q x))" + by blast + + +text {* \medskip Maxiscoping: pulling out big Unions and Intersections. *} + +lemma UN_extend_simps: + "!!a B C. insert a (UN x:C. B x) = (if C={} then {a} else (UN x:C. insert a (B x)))" + "!!A B C. (UN x:C. A x) Un B = (if C={} then B else (UN x:C. A x Un B))" + "!!A B C. A Un (UN x:C. B x) = (if C={} then A else (UN x:C. A Un B x))" + "!!A B C. ((UN x:C. A x) Int B) = (UN x:C. A x Int B)" + "!!A B C. (A Int (UN x:C. B x)) = (UN x:C. A Int B x)" + "!!A B C. ((UN x:C. A x) - B) = (UN x:C. A x - B)" + "!!A B C. (A - (INT x:C. B x)) = (UN x:C. A - B x)" + "!!A B. (UN y:A. UN x:y. B x) = (UN x: Union A. B x)" + "!!A B C. (UN x:A. UN z: B(x). C z) = (UN z: UNION A B. C z)" + "!!A B f. (UN a:A. B (f a)) = (UN x:f`A. B x)" + by auto + +lemma INT_extend_simps: + "!!A B C. (INT x:C. A x) Int B = (if C={} then B else (INT x:C. A x Int B))" + "!!A B C. A Int (INT x:C. B x) = (if C={} then A else (INT x:C. A Int B x))" + "!!A B C. (INT x:C. A x) - B = (if C={} then UNIV-B else (INT x:C. A x - B))" + "!!A B C. A - (UN x:C. B x) = (if C={} then A else (INT x:C. A - B x))" + "!!a B C. insert a (INT x:C. B x) = (INT x:C. insert a (B x))" + "!!A B C. ((INT x:C. A x) Un B) = (INT x:C. A x Un B)" + "!!A B C. A Un (INT x:C. B x) = (INT x:C. A Un B x)" + "!!A B. (INT y:A. INT x:y. B x) = (INT x: Union A. B x)" + "!!A B C. (INT x:A. INT z: B(x). C z) = (INT z: UNION A B. C z)" + "!!A B f. (INT a:A. B (f a)) = (INT x:f`A. B x)" + by auto + + +no_notation + less_eq (infix "\" 50) and + less (infix "\" 50) and + inf (infixl "\" 70) and + sup (infixl "\" 65) and + Inf ("\_" [900] 900) and + Sup ("\_" [900] 900) + +lemmas mem_simps = + insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff + mem_Collect_eq UN_iff Union_iff INT_iff Inter_iff + -- {* Each of these has ALREADY been added @{text "[simp]"} above. *} + +end diff -r 63686057cbe8 -r abda97d2deea src/HOL/Decision_Procs/Dense_Linear_Order.thy --- a/src/HOL/Decision_Procs/Dense_Linear_Order.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Decision_Procs/Dense_Linear_Order.thy Thu Jul 23 21:13:21 2009 +0200 @@ -692,7 +692,7 @@ val clt = Thm.dest_fun2 ct val cz = Thm.dest_arg ct val neg = cr nat" -primrec - fact_0: "fact 0 = 1" - fact_Suc: "fact (Suc n) = (Suc n) * fact n" +class fact = + +fixes + fact :: "'a \ 'a" + +instantiation nat :: fact + +begin + +fun + fact_nat :: "nat \ nat" +where + fact_0_nat: "fact_nat 0 = Suc 0" +| fact_Suc: "fact_nat (Suc x) = Suc x * fact x" + +instance proof qed + +end + +(* definitions for the integers *) + +instantiation int :: fact + +begin + +definition + fact_int :: "int \ int" +where + "fact_int x = (if x >= 0 then int (fact (nat x)) else 0)" + +instance proof qed + +end + + +subsection {* Set up Transfer *} + +lemma transfer_nat_int_factorial: + "(x::int) >= 0 \ fact (nat x) = nat (fact x)" + unfolding fact_int_def + by auto + + +lemma transfer_nat_int_factorial_closure: + "x >= (0::int) \ fact x >= 0" + by (auto simp add: fact_int_def) + +declare TransferMorphism_nat_int[transfer add return: + transfer_nat_int_factorial transfer_nat_int_factorial_closure] + +lemma transfer_int_nat_factorial: + "fact (int x) = int (fact x)" + unfolding fact_int_def by auto + +lemma transfer_int_nat_factorial_closure: + "is_nat x \ fact x >= 0" + by (auto simp add: fact_int_def) + +declare TransferMorphism_int_nat[transfer add return: + transfer_int_nat_factorial transfer_int_nat_factorial_closure] -lemma fact_gt_zero [simp]: "0 < fact n" -by (induct n) auto +subsection {* Factorial *} + +lemma fact_0_int [simp]: "fact (0::int) = 1" + by (simp add: fact_int_def) + +lemma fact_1_nat [simp]: "fact (1::nat) = 1" + by simp + +lemma fact_Suc_0_nat [simp]: "fact (Suc 0) = Suc 0" + by simp + +lemma fact_1_int [simp]: "fact (1::int) = 1" + by (simp add: fact_int_def) + +lemma fact_plus_one_nat: "fact ((n::nat) + 1) = (n + 1) * fact n" + by simp + +lemma fact_plus_one_int: + assumes "n >= 0" + shows "fact ((n::int) + 1) = (n + 1) * fact n" + + using prems unfolding fact_int_def + by (simp add: nat_add_distrib algebra_simps int_mult) + +lemma fact_reduce_nat: "(n::nat) > 0 \ fact n = n * fact (n - 1)" + apply (subgoal_tac "n = Suc (n - 1)") + apply (erule ssubst) + apply (subst fact_Suc) + apply simp_all +done + +lemma fact_reduce_int: "(n::int) > 0 \ fact n = n * fact (n - 1)" + apply (subgoal_tac "n = (n - 1) + 1") + apply (erule ssubst) + apply (subst fact_plus_one_int) + apply simp_all +done + +lemma fact_nonzero_nat [simp]: "fact (n::nat) \ 0" + apply (induct n) + apply (auto simp add: fact_plus_one_nat) +done + +lemma fact_nonzero_int [simp]: "n >= 0 \ fact (n::int) ~= 0" + by (simp add: fact_int_def) + +lemma fact_gt_zero_nat [simp]: "fact (n :: nat) > 0" + by (insert fact_nonzero_nat [of n], arith) + +lemma fact_gt_zero_int [simp]: "n >= 0 \ fact (n :: int) > 0" + by (auto simp add: fact_int_def) + +lemma fact_ge_one_nat [simp]: "fact (n :: nat) >= 1" + by (insert fact_nonzero_nat [of n], arith) + +lemma fact_ge_Suc_0_nat [simp]: "fact (n :: nat) >= Suc 0" + by (insert fact_nonzero_nat [of n], arith) + +lemma fact_ge_one_int [simp]: "n >= 0 \ fact (n :: int) >= 1" + apply (auto simp add: fact_int_def) + apply (subgoal_tac "1 = int 1") + apply (erule ssubst) + apply (subst zle_int) + apply auto +done + +lemma dvd_fact_nat [rule_format]: "1 <= m \ m <= n \ m dvd fact (n::nat)" + apply (induct n) + apply force + apply (auto simp only: fact_Suc) + apply (subgoal_tac "m = Suc n") + apply (erule ssubst) + apply (rule dvd_triv_left) + apply auto +done + +lemma dvd_fact_int [rule_format]: "1 <= m \ m <= n \ m dvd fact (n::int)" + apply (case_tac "1 <= n") + apply (induct n rule: int_ge_induct) + apply (auto simp add: fact_plus_one_int) + apply (subgoal_tac "m = i + 1") + apply auto +done + +lemma interval_plus_one_nat: "(i::nat) <= j + 1 \ + {i..j+1} = {i..j} Un {j+1}" + by auto + +lemma interval_Suc: "i <= Suc j \ {i..Suc j} = {i..j} Un {Suc j}" + by auto + +lemma interval_plus_one_int: "(i::int) <= j + 1 \ {i..j+1} = {i..j} Un {j+1}" + by auto -lemma fact_not_eq_zero [simp]: "fact n \ 0" -by simp +lemma fact_altdef_nat: "fact (n::nat) = (PROD i:{1..n}. i)" + apply (induct n) + apply force + apply (subst fact_Suc) + apply (subst interval_Suc) + apply auto +done + +lemma fact_altdef_int: "n >= 0 \ fact (n::int) = (PROD i:{1..n}. i)" + apply (induct n rule: int_ge_induct) + apply force + apply (subst fact_plus_one_int, assumption) + apply (subst interval_plus_one_int) + apply auto +done + +lemma fact_mono_nat: "(m::nat) \ n \ fact m \ fact n" +apply (drule le_imp_less_or_eq) +apply (auto dest!: less_imp_Suc_add) +apply (induct_tac k, auto) +done + +lemma fact_neg_int [simp]: "m < (0::int) \ fact m = 0" + unfolding fact_int_def by auto + +lemma fact_ge_zero_int [simp]: "fact m >= (0::int)" + apply (case_tac "m >= 0") + apply auto + apply (frule fact_gt_zero_int) + apply arith +done + +lemma fact_mono_int_aux [rule_format]: "k >= (0::int) \ + fact (m + k) >= fact m" + apply (case_tac "m < 0") + apply auto + apply (induct k rule: int_ge_induct) + apply auto + apply (subst add_assoc [symmetric]) + apply (subst fact_plus_one_int) + apply auto + apply (erule order_trans) + apply (subst mult_le_cancel_right1) + apply (subgoal_tac "fact (m + i) >= 0") + apply arith + apply auto +done + +lemma fact_mono_int: "(m::int) <= n \ fact m <= fact n" + apply (insert fact_mono_int_aux [of "n - m" "m"]) + apply auto +done + +text{*Note that @{term "fact 0 = fact 1"}*} +lemma fact_less_mono_nat: "[| (0::nat) < m; m < n |] ==> fact m < fact n" +apply (drule_tac m = m in less_imp_Suc_add, auto) +apply (induct_tac k, auto) +done + +lemma fact_less_mono_int_aux: "k >= 0 \ (0::int) < m \ + fact m < fact ((m + 1) + k)" + apply (induct k rule: int_ge_induct) + apply (simp add: fact_plus_one_int) + apply (subst mult_less_cancel_right1) + apply (insert fact_gt_zero_int [of m], arith) + apply (subst (2) fact_reduce_int) + apply (auto simp add: add_ac) + apply (erule order_less_le_trans) + apply (subst mult_le_cancel_right1) + apply auto + apply (subgoal_tac "fact (i + (1 + m)) >= 0") + apply force + apply (rule fact_ge_zero_int) +done + +lemma fact_less_mono_int: "(0::int) < m \ m < n \ fact m < fact n" + apply (insert fact_less_mono_int_aux [of "n - (m + 1)" "m"]) + apply auto +done + +lemma fact_num_eq_if_nat: "fact (m::nat) = + (if m=0 then 1 else m * fact (m - 1))" +by (cases m) auto + +lemma fact_add_num_eq_if_nat: + "fact ((m::nat) + n) = (if m + n = 0 then 1 else (m + n) * fact (m + n - 1))" +by (cases "m + n") auto + +lemma fact_add_num_eq_if2_nat: + "fact ((m::nat) + n) = + (if m = 0 then fact n else (m + n) * fact ((m - 1) + n))" +by (cases m) auto + + +subsection {* @{term fact} and @{term of_nat} *} lemma of_nat_fact_not_zero [simp]: "of_nat (fact n) \ (0::'a::semiring_char_0)" by auto @@ -33,46 +274,10 @@ lemma of_nat_fact_ge_zero [simp]: "(0::'a::ordered_semidom) \ of_nat(fact n)" by simp -lemma fact_ge_one [simp]: "1 \ fact n" -by (induct n) auto - -lemma fact_mono: "m \ n ==> fact m \ fact n" -apply (drule le_imp_less_or_eq) -apply (auto dest!: less_imp_Suc_add) -apply (induct_tac k, auto) -done - -text{*Note that @{term "fact 0 = fact 1"}*} -lemma fact_less_mono: "[| 0 < m; m < n |] ==> fact m < fact n" -apply (drule_tac m = m in less_imp_Suc_add, auto) -apply (induct_tac k, auto) -done - lemma inv_of_nat_fact_gt_zero [simp]: "(0::'a::ordered_field) < inverse (of_nat (fact n))" by (auto simp add: positive_imp_inverse_positive) lemma inv_of_nat_fact_ge_zero [simp]: "(0::'a::ordered_field) \ inverse (of_nat (fact n))" by (auto intro: order_less_imp_le) -lemma fact_diff_Suc [rule_format]: - "n < Suc m ==> fact (Suc m - n) = (Suc m - n) * fact (m - n)" -apply (induct n arbitrary: m) -apply auto -apply (drule_tac x = "m - Suc 0" in meta_spec, auto) -done - -lemma fact_num0: "fact 0 = 1" -by auto - -lemma fact_num_eq_if: "fact m = (if m=0 then 1 else m * fact (m - 1))" -by (cases m) auto - -lemma fact_add_num_eq_if: - "fact (m + n) = (if m + n = 0 then 1 else (m + n) * fact (m + n - 1))" -by (cases "m + n") auto - -lemma fact_add_num_eq_if2: - "fact (m + n) = (if m = 0 then fact n else (m + n) * fact ((m - 1) + n))" -by (cases m) auto - end diff -r 63686057cbe8 -r abda97d2deea src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Finite_Set.thy Thu Jul 23 21:13:21 2009 +0200 @@ -93,6 +93,7 @@ qed qed + text{* A finite choice principle. Does not need the SOME choice operator. *} lemma finite_set_choice: "finite A \ ALL x:A. (EX y. P x y) \ EX f. ALL x:A. P x (f x)" @@ -811,7 +812,7 @@ by(rule fun_left_comm_idem.fold_insert_idem[OF ab_semigroup_idem_mult.fun_left_comm_idem[OF ab_semigroup_idem_mult_inf]]) lemma inf_le_fold_inf: "finite A \ ALL a:A. b \ a \ inf b c \ fold inf c A" -by (induct pred:finite) auto +by (induct pred: finite) (auto intro: le_infI1) lemma fold_inf_le_inf: "finite A \ a \ A \ fold inf b A \ inf a b" proof(induct arbitrary: a pred:finite) @@ -822,7 +823,7 @@ proof cases assume "A = {}" thus ?thesis using insert by simp next - assume "A \ {}" thus ?thesis using insert by auto + assume "A \ {}" thus ?thesis using insert by (auto intro: le_infI2) qed qed @@ -2122,6 +2123,18 @@ \ x \ \ F = {}" by auto +lemma finite_psubset_induct[consumes 1, case_names psubset]: + assumes "finite A" and "!!A. finite A \ (!!B. finite B \ B \ A \ P(B)) \ P(A)" shows "P A" +using assms(1) +proof (induct A rule: measure_induct_rule[where f=card]) + case (less A) + show ?case + proof(rule assms(2)[OF less(2)]) + fix B assume "finite B" "B \ A" + show "P B" by(rule less(1)[OF psubset_card_mono[OF less(2) `B \ A`] `finite B`]) + qed +qed + text{* main cardinality theorem *} lemma card_partition [rule_format]: "finite C ==> @@ -3252,13 +3265,13 @@ by (simp add: Max fold1_antimono [folded dual_max]) qed -lemma finite_linorder_induct[consumes 1, case_names empty insert]: +lemma finite_linorder_max_induct[consumes 1, case_names empty insert]: "finite A \ P {} \ (!!A b. finite A \ ALL a:A. a < b \ P A \ P(insert b A)) \ P A" -proof (induct A rule: measure_induct_rule[where f=card]) +proof (induct rule: finite_psubset_induct) fix A :: "'a set" - assume IH: "!! B. card B < card A \ finite B \ P {} \ + assume IH: "!! B. finite B \ B < A \ P {} \ (!!A b. finite A \ (\a\A. a P A \ P (insert b A)) \ P B" and "finite A" and "P {}" @@ -3271,16 +3284,17 @@ assume "A \ {}" with `finite A` have "Max A : A" by auto hence A: "?A = A" using insert_Diff_single insert_absorb by auto - note card_Diff1_less[OF `finite A` `Max A : A`] moreover have "finite ?B" using `finite A` by simp ultimately have "P ?B" using `P {}` step IH by blast - moreover have "\a\?B. a < Max A" - using Max_ge [OF `finite A`] by fastsimp - ultimately show "P A" - using A insert_Diff_single step[OF `finite ?B`] by fastsimp + moreover have "\a\?B. a < Max A" using Max_ge [OF `finite A`] by fastsimp + ultimately show "P A" using A insert_Diff_single step[OF `finite ?B`] by fastsimp qed qed +lemma finite_linorder_min_induct[consumes 1, case_names empty insert]: + "\finite A; P {}; \A b. \finite A; \a\A. b < a; P A\ \ P (insert b A)\ \ P A" +by(rule linorder.finite_linorder_max_induct[OF dual_linorder]) + end context ordered_ab_semigroup_add diff -r 63686057cbe8 -r abda97d2deea src/HOL/Fun.thy --- a/src/HOL/Fun.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Fun.thy Thu Jul 23 21:13:21 2009 +0200 @@ -6,7 +6,7 @@ header {* Notions about functions *} theory Fun -imports Set +imports Complete_Lattice begin text{*As a simplification rule, it replaces all function equalities by diff -r 63686057cbe8 -r abda97d2deea src/HOL/GCD.thy --- a/src/HOL/GCD.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/GCD.thy Thu Jul 23 21:13:21 2009 +0200 @@ -20,6 +20,9 @@ the congruence relations on the integers. The notion was extended to the natural numbers by Chiaeb. +Jeremy Avigad combined all of these, made everything uniform for the +natural numbers and the integers, and added a number of new theorems. + Tobias Nipkow cleaned up a lot. *) @@ -27,7 +30,7 @@ header {* GCD *} theory GCD -imports NatTransfer +imports Fact begin declare One_nat_def [simp del] @@ -159,7 +162,6 @@ transfer_int_nat_gcd transfer_int_nat_gcd_closures] - subsection {* GCD *} (* was gcd_induct *) @@ -1505,6 +1507,252 @@ lemma lcm_1_iff_int[simp]: "lcm (m::int) n = 1 \ (m=1 \ m = -1) \ (n=1 \ n = -1)" by (auto simp add: abs_mult_self trans [OF lcm_unique_int eq_commute, symmetric] zmult_eq_1_iff) +subsubsection {* The complete divisibility lattice *} + + +interpretation gcd_semilattice_nat: lower_semilattice "op dvd" "(%m n::nat. m dvd n & ~ n dvd m)" gcd +proof + case goal3 thus ?case by(metis gcd_unique_nat) +qed auto + +interpretation lcm_semilattice_nat: upper_semilattice "op dvd" "(%m n::nat. m dvd n & ~ n dvd m)" lcm +proof + case goal3 thus ?case by(metis lcm_unique_nat) +qed auto + +interpretation gcd_lcm_lattice_nat: lattice "op dvd" "(%m n::nat. m dvd n & ~ n dvd m)" gcd lcm .. + +text{* Lifting gcd and lcm to finite (Gcd/Lcm) and infinite sets (GCD/LCM). +GCD is defined via LCM to facilitate the proof that we have a complete lattice. +Later on we show that GCD and Gcd coincide on finite sets. +*} +context gcd +begin + +definition Gcd :: "'a set \ 'a" +where "Gcd = fold gcd 0" + +definition Lcm :: "'a set \ 'a" +where "Lcm = fold lcm 1" + +definition LCM :: "'a set \ 'a" where +"LCM M = (if finite M then Lcm M else 0)" + +definition GCD :: "'a set \ 'a" where +"GCD M = LCM(INT m:M. {d. d dvd m})" + +end + +lemma Gcd_empty[simp]: "Gcd {} = 0" +by(simp add:Gcd_def) + +lemma Lcm_empty[simp]: "Lcm {} = 1" +by(simp add:Lcm_def) + +lemma GCD_empty_nat[simp]: "GCD {} = (0::nat)" +by(simp add:GCD_def LCM_def) + +lemma LCM_eq_Lcm[simp]: "finite M \ LCM M = Lcm M" +by(simp add:LCM_def) + +lemma Lcm_insert_nat [simp]: + assumes "finite N" + shows "Lcm (insert (n::nat) N) = lcm n (Lcm N)" +proof - + interpret fun_left_comm_idem "lcm::nat\nat\nat" + by (rule fun_left_comm_idem_lcm_nat) + from assms show ?thesis by(simp add: Lcm_def) +qed + +lemma Lcm_insert_int [simp]: + assumes "finite N" + shows "Lcm (insert (n::int) N) = lcm n (Lcm N)" +proof - + interpret fun_left_comm_idem "lcm::int\int\int" + by (rule fun_left_comm_idem_lcm_int) + from assms show ?thesis by(simp add: Lcm_def) +qed + +lemma Gcd_insert_nat [simp]: + assumes "finite N" + shows "Gcd (insert (n::nat) N) = gcd n (Gcd N)" +proof - + interpret fun_left_comm_idem "gcd::nat\nat\nat" + by (rule fun_left_comm_idem_gcd_nat) + from assms show ?thesis by(simp add: Gcd_def) +qed + +lemma Gcd_insert_int [simp]: + assumes "finite N" + shows "Gcd (insert (n::int) N) = gcd n (Gcd N)" +proof - + interpret fun_left_comm_idem "gcd::int\int\int" + by (rule fun_left_comm_idem_gcd_int) + from assms show ?thesis by(simp add: Gcd_def) +qed + +lemma Lcm0_iff[simp]: "finite (M::nat set) \ M \ {} \ Lcm M = 0 \ 0 : M" +by(induct rule:finite_ne_induct) auto + +lemma Lcm_eq_0[simp]: "finite (M::nat set) \ 0 : M \ Lcm M = 0" +by (metis Lcm0_iff empty_iff) + +lemma Gcd_dvd_nat [simp]: + assumes "finite M" and "(m::nat) \ M" + shows "Gcd M dvd m" +proof - + show ?thesis using gcd_semilattice_nat.fold_inf_le_inf[OF assms, of 0] by (simp add: Gcd_def) +qed + +lemma dvd_Gcd_nat[simp]: + assumes "finite M" and "ALL (m::nat) : M. n dvd m" + shows "n dvd Gcd M" +proof - + show ?thesis using gcd_semilattice_nat.inf_le_fold_inf[OF assms, of 0] by (simp add: Gcd_def) +qed + +lemma dvd_Lcm_nat [simp]: + assumes "finite M" and "(m::nat) \ M" + shows "m dvd Lcm M" +proof - + show ?thesis using lcm_semilattice_nat.sup_le_fold_sup[OF assms, of 1] by (simp add: Lcm_def) +qed + +lemma Lcm_dvd_nat[simp]: + assumes "finite M" and "ALL (m::nat) : M. m dvd n" + shows "Lcm M dvd n" +proof - + show ?thesis using lcm_semilattice_nat.fold_sup_le_sup[OF assms, of 1] by (simp add: Lcm_def) +qed + +interpretation gcd_lcm_complete_lattice_nat: + complete_lattice "op dvd" "%m n::nat. m dvd n & ~ n dvd m" gcd lcm 1 0 GCD LCM +proof + case goal1 show ?case by simp +next + case goal2 show ?case by simp +next + case goal5 thus ?case by (auto simp: LCM_def) +next + case goal6 thus ?case + by(auto simp: LCM_def)(metis finite_nat_set_iff_bounded_le gcd_proj2_if_dvd_nat gcd_le1_nat) +next + case goal3 thus ?case by (auto simp: GCD_def LCM_def)(metis finite_INT finite_divisors_nat) +next + case goal4 thus ?case by(auto simp: LCM_def GCD_def) +qed + +text{* Alternative characterizations of Gcd and GCD: *} + +lemma Gcd_eq_Max: "finite(M::nat set) \ M \ {} \ 0 \ M \ Gcd M = Max(\m\M. {d. d dvd m})" +apply(rule antisym) + apply(rule Max_ge) + apply (metis all_not_in_conv finite_divisors_nat finite_INT) + apply simp +apply (rule Max_le_iff[THEN iffD2]) + apply (metis all_not_in_conv finite_divisors_nat finite_INT) + apply fastsimp +apply clarsimp +apply (metis Gcd_dvd_nat Max_in dvd_0_left dvd_Gcd_nat dvd_imp_le linorder_antisym_conv3 not_less0) +done + +lemma Gcd_remove0_nat: "finite M \ Gcd M = Gcd (M - {0::nat})" +apply(induct pred:finite) + apply simp +apply(case_tac "x=0") + apply simp +apply(subgoal_tac "insert x F - {0} = insert x (F - {0})") + apply simp +apply blast +done + +lemma Lcm_in_lcm_closed_set_nat: + "finite M \ M \ {} \ ALL m n :: nat. m:M \ n:M \ lcm m n : M \ Lcm M : M" +apply(induct rule:finite_linorder_min_induct) + apply simp +apply simp +apply(subgoal_tac "ALL m n :: nat. m:A \ n:A \ lcm m n : A") + apply simp + apply(case_tac "A={}") + apply simp + apply simp +apply (metis lcm_pos_nat lcm_unique_nat linorder_neq_iff nat_dvd_not_less not_less0) +done + +lemma Lcm_eq_Max_nat: + "finite M \ M \ {} \ 0 \ M \ ALL m n :: nat. m:M \ n:M \ lcm m n : M \ Lcm M = Max M" +apply(rule antisym) + apply(rule Max_ge, assumption) + apply(erule (2) Lcm_in_lcm_closed_set_nat) +apply clarsimp +apply (metis Lcm0_iff dvd_Lcm_nat dvd_imp_le neq0_conv) +done + +text{* Finally GCD is Gcd: *} + +lemma GCD_eq_Gcd[simp]: assumes "finite(M::nat set)" shows "GCD M = Gcd M" +proof- + have divisors_remove0_nat: "(\m\M. {d::nat. d dvd m}) = (\m\M-{0}. {d::nat. d dvd m})" by auto + show ?thesis + proof cases + assume "M={}" thus ?thesis by simp + next + assume "M\{}" + show ?thesis + proof cases + assume "M={0}" thus ?thesis by(simp add:GCD_def LCM_def) + next + assume "M\{0}" + with `M\{}` assms show ?thesis + apply(subst Gcd_remove0_nat[OF assms]) + apply(simp add:GCD_def) + apply(subst divisors_remove0_nat) + apply(simp add:LCM_def) + apply rule + apply rule + apply(subst Gcd_eq_Max) + apply simp + apply blast + apply blast + apply(rule Lcm_eq_Max_nat) + apply simp + apply blast + apply fastsimp + apply clarsimp + apply(fastsimp intro: finite_divisors_nat intro!: finite_INT) + done + qed + qed +qed + +lemma Lcm_set_nat [code_unfold]: + "Lcm (set ns) = foldl lcm (1::nat) ns" +proof - + interpret fun_left_comm_idem "lcm::nat\nat\nat" by (rule fun_left_comm_idem_lcm_nat) + show ?thesis by(simp add: Lcm_def fold_set lcm_commute_nat) +qed + +lemma Lcm_set_int [code_unfold]: + "Lcm (set is) = foldl lcm (1::int) is" +proof - + interpret fun_left_comm_idem "lcm::int\int\int" by (rule fun_left_comm_idem_lcm_int) + show ?thesis by(simp add: Lcm_def fold_set lcm_commute_int) +qed + +lemma Gcd_set_nat [code_unfold]: + "Gcd (set ns) = foldl gcd (0::nat) ns" +proof - + interpret fun_left_comm_idem "gcd::nat\nat\nat" by (rule fun_left_comm_idem_gcd_nat) + show ?thesis by(simp add: Gcd_def fold_set gcd_commute_nat) +qed + +lemma Gcd_set_int [code_unfold]: + "Gcd (set ns) = foldl gcd (0::int) ns" +proof - + interpret fun_left_comm_idem "gcd::int\int\int" by (rule fun_left_comm_idem_gcd_int) + show ?thesis by(simp add: Gcd_def fold_set gcd_commute_int) +qed + subsection {* Primes *} @@ -1553,32 +1801,6 @@ apply (case_tac "p >= 0") by (blast, auto simp add: prime_ge_0_int) -(* To do: determine primality of any numeral *) - -lemma zero_not_prime_nat [simp]: "~prime (0::nat)" - by (simp add: prime_nat_def) - -lemma zero_not_prime_int [simp]: "~prime (0::int)" - by (simp add: prime_int_def) - -lemma one_not_prime_nat [simp]: "~prime (1::nat)" - by (simp add: prime_nat_def) - -lemma Suc_0_not_prime_nat [simp]: "~prime (Suc 0)" - by (simp add: prime_nat_def One_nat_def) - -lemma one_not_prime_int [simp]: "~prime (1::int)" - by (simp add: prime_int_def) - -lemma two_is_prime_nat [simp]: "prime (2::nat)" - apply (auto simp add: prime_nat_def) - apply (case_tac m) - apply (auto dest!: dvd_imp_le) - done - -lemma two_is_prime_int [simp]: "prime (2::int)" - by (rule two_is_prime_nat [transferred direction: nat "op <= (0::int)"]) - lemma prime_imp_coprime_nat: "prime (p::nat) \ \ p dvd n \ coprime p n" apply (unfold prime_nat_def) apply (metis gcd_dvd1_nat gcd_dvd2_nat) @@ -1625,15 +1847,84 @@ apply auto done -lemma prime_imp_power_coprime_nat: "prime (p::nat) \ ~ p dvd a \ - coprime a (p^m)" +subsubsection{* Make prime naively executable *} + +lemma zero_not_prime_nat [simp]: "~prime (0::nat)" + by (simp add: prime_nat_def) + +lemma zero_not_prime_int [simp]: "~prime (0::int)" + by (simp add: prime_int_def) + +lemma one_not_prime_nat [simp]: "~prime (1::nat)" + by (simp add: prime_nat_def) + +lemma Suc_0_not_prime_nat [simp]: "~prime (Suc 0)" + by (simp add: prime_nat_def One_nat_def) + +lemma one_not_prime_int [simp]: "~prime (1::int)" + by (simp add: prime_int_def) + +lemma prime_nat_code[code]: + "prime(p::nat) = (p > 1 & (ALL n : {1<.. 1 & (list_all (%n. ~ n dvd p) [2.. 1 & (ALL n : {1<.. 1 & (list_all (%n. ~ n dvd p) [2..p - 1]))" +apply(simp only:prime_int_code list_ball_code greaterThanLessThan_upto) +apply simp +done + +lemmas prime_int_simp_number_of[simp] = prime_int_simp[of "number_of m", standard] + +declare successor_int_def[simp] + +lemma two_is_prime_nat [simp]: "prime (2::nat)" +by simp + +lemma two_is_prime_int [simp]: "prime (2::int)" +by simp + +text{* A bit of regression testing: *} + +lemma "prime(97::nat)" +by simp + +lemma "prime(97::int)" +by simp + +lemma "prime(997::nat)" +by eval + +lemma "prime(997::int)" +by eval + + +lemma prime_imp_power_coprime_nat: "prime (p::nat) \ ~ p dvd a \ coprime a (p^m)" apply (rule coprime_exp_nat) apply (subst gcd_commute_nat) apply (erule (1) prime_imp_coprime_nat) done -lemma prime_imp_power_coprime_int: "prime (p::int) \ ~ p dvd a \ - coprime a (p^m)" +lemma prime_imp_power_coprime_int: "prime (p::int) \ ~ p dvd a \ coprime a (p^m)" apply (rule coprime_exp_int) apply (subst gcd_commute_int) apply (erule (1) prime_imp_coprime_int) @@ -1652,12 +1943,10 @@ apply auto done -lemma primes_imp_powers_coprime_nat: "prime (p::nat) \ prime q \ p ~= q \ - coprime (p^m) (q^n)" +lemma primes_imp_powers_coprime_nat: "prime (p::nat) \ prime q \ p ~= q \ coprime (p^m) (q^n)" by (rule coprime_exp2_nat, rule primes_coprime_nat) -lemma primes_imp_powers_coprime_int: "prime (p::int) \ prime q \ p ~= q \ - coprime (p^m) (q^n)" +lemma primes_imp_powers_coprime_int: "prime (p::int) \ prime q \ p ~= q \ coprime (p^m) (q^n)" by (rule coprime_exp2_int, rule primes_coprime_int) lemma prime_factor_nat: "n \ (1::nat) \ \ p. prime p \ p dvd n" @@ -1751,4 +2040,42 @@ ultimately show ?thesis by blast qed +subsection {* Infinitely many primes *} + +lemma next_prime_bound: "\(p::nat). prime p \ n < p \ p <= fact n + 1" +proof- + have f1: "fact n + 1 \ 1" using fact_ge_one_nat [of n] by arith + from prime_factor_nat [OF f1] + obtain p where "prime p" and "p dvd fact n + 1" by auto + hence "p \ fact n + 1" + by (intro dvd_imp_le, auto) + {assume "p \ n" + from `prime p` have "p \ 1" + by (cases p, simp_all) + with `p <= n` have "p dvd fact n" + by (intro dvd_fact_nat) + with `p dvd fact n + 1` have "p dvd fact n + 1 - fact n" + by (rule dvd_diff_nat) + hence "p dvd 1" by simp + hence "p <= 1" by auto + moreover from `prime p` have "p > 1" by auto + ultimately have False by auto} + hence "n < p" by arith + with `prime p` and `p <= fact n + 1` show ?thesis by auto +qed + +lemma bigger_prime: "\p. prime p \ p > (n::nat)" +using next_prime_bound by auto + +lemma primes_infinite: "\ (finite {(p::nat). prime p})" +proof + assume "finite {(p::nat). prime p}" + with Max_ge have "(EX b. (ALL x : {(p::nat). prime p}. x <= b))" + by auto + then obtain b where "ALL (x::nat). prime x \ x <= b" + by auto + with bigger_prime [of b] show False by auto +qed + + end diff -r 63686057cbe8 -r abda97d2deea src/HOL/HOL.thy --- a/src/HOL/HOL.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/HOL.thy Thu Jul 23 21:13:21 2009 +0200 @@ -107,7 +107,6 @@ notation (xsymbols) iff (infixr "\" 25) - nonterminals letbinds letbind case_syn cases_syn @@ -187,8 +186,8 @@ True_or_False: "(P=True) | (P=False)" defs - Let_def: "Let s f == f(s)" - if_def: "If P x y == THE z::'a. (P=True --> z=x) & (P=False --> z=y)" + Let_def [code]: "Let s f == f(s)" + if_def: "If P x y == THE z::'a. (P=True --> z=x) & (P=False --> z=y)" finalconsts "op =" @@ -198,96 +197,9 @@ axiomatization undefined :: 'a - -subsubsection {* Generic classes and algebraic operations *} - class default = fixes default :: 'a -class zero = - fixes zero :: 'a ("0") - -class one = - fixes one :: 'a ("1") - -hide (open) const zero one - -class plus = - fixes plus :: "'a \ 'a \ 'a" (infixl "+" 65) - -class minus = - fixes minus :: "'a \ 'a \ 'a" (infixl "-" 65) - -class uminus = - fixes uminus :: "'a \ 'a" ("- _" [81] 80) - -class times = - fixes times :: "'a \ 'a \ 'a" (infixl "*" 70) - -class inverse = - fixes inverse :: "'a \ 'a" - and divide :: "'a \ 'a \ 'a" (infixl "'/" 70) - -class abs = - fixes abs :: "'a \ 'a" -begin - -notation (xsymbols) - abs ("\_\") - -notation (HTML output) - abs ("\_\") - -end - -class sgn = - fixes sgn :: "'a \ 'a" - -class ord = - fixes less_eq :: "'a \ 'a \ bool" - and less :: "'a \ 'a \ bool" -begin - -notation - less_eq ("op <=") and - less_eq ("(_/ <= _)" [51, 51] 50) and - less ("op <") and - less ("(_/ < _)" [51, 51] 50) - -notation (xsymbols) - less_eq ("op \") and - less_eq ("(_/ \ _)" [51, 51] 50) - -notation (HTML output) - less_eq ("op \") and - less_eq ("(_/ \ _)" [51, 51] 50) - -abbreviation (input) - greater_eq (infix ">=" 50) where - "x >= y \ y <= x" - -notation (input) - greater_eq (infix "\" 50) - -abbreviation (input) - greater (infix ">" 50) where - "x > y \ y < x" - -end - -syntax - "_index1" :: index ("\<^sub>1") -translations - (index) "\<^sub>1" => (index) "\<^bsub>\\<^esub>" - -typed_print_translation {* -let - fun tr' c = (c, fn show_sorts => fn T => fn ts => - if (not o null) ts orelse T = dummyT orelse not (! show_types) andalso can Term.dest_Type T then raise Match - else Syntax.const Syntax.constrainC $ Syntax.const c $ Syntax.term_of_typ show_sorts T); -in map tr' [@{const_syntax HOL.one}, @{const_syntax HOL.zero}] end; -*} -- {* show types that are presumably too general *} - subsection {* Fundamental rules *} @@ -921,7 +833,7 @@ open BasicClassical; ML_Antiquote.value "claset" - (Scan.succeed "Classical.local_claset_of (ML_Context.the_local_context ())"); + (Scan.succeed "Classical.claset_of (ML_Context.the_local_context ())"); structure ResAtpset = Named_Thms (val name = "atp" val description = "ATP rules"); @@ -1029,8 +941,8 @@ "(P ~= Q) = (P = (~Q))" "(P | ~P) = True" "(~P | P) = True" "(x = x) = True" - and not_True_eq_False: "(\ True) = False" - and not_False_eq_True: "(\ False) = True" + and not_True_eq_False [code]: "(\ True) = False" + and not_False_eq_True [code]: "(\ False) = True" and "(~P) ~= P" "P ~= (~P)" "(True=P) = P" @@ -1157,10 +1069,10 @@ text {* \medskip if-then-else rules *} -lemma if_True: "(if True then x else y) = x" +lemma if_True [code]: "(if True then x else y) = x" by (unfold if_def) blast -lemma if_False: "(if False then x else y) = y" +lemma if_False [code]: "(if False then x else y) = y" by (unfold if_def) blast lemma if_P: "P ==> (if P then x else y) = x" @@ -1605,25 +1517,9 @@ setup ReorientProc.init -setup {* - ReorientProc.add - (fn Const(@{const_name HOL.zero}, _) => true - | Const(@{const_name HOL.one}, _) => true - | _ => false) -*} - -simproc_setup reorient_zero ("0 = x") = ReorientProc.proc -simproc_setup reorient_one ("1 = x") = ReorientProc.proc - subsection {* Other simple lemmas and lemma duplicates *} -lemma Let_0 [simp]: "Let 0 f = f 0" - unfolding Let_def .. - -lemma Let_1 [simp]: "Let 1 f = f 1" - unfolding Let_def .. - lemma ex1_eq [iff]: "EX! x. x = t" "EX! x. t = x" by blast+ @@ -1643,13 +1539,6 @@ apply (drule_tac [3] x = x in fun_cong, simp_all) done -lemma mk_left_commute: - fixes f (infix "\" 60) - assumes a: "\x y z. (x \ y) \ z = x \ (y \ z)" and - c: "\x y. x \ y = y \ x" - shows "x \ (y \ z) = y \ (x \ z)" - by (rule trans [OF trans [OF c a] arg_cong [OF c, of "f y"]]) - lemmas eq_sym_conv = eq_commute lemma nnf_simps: @@ -1659,6 +1548,118 @@ by blast+ +subsection {* Generic classes and algebraic operations *} + +class zero = + fixes zero :: 'a ("0") + +class one = + fixes one :: 'a ("1") + +lemma Let_0 [simp]: "Let 0 f = f 0" + unfolding Let_def .. + +lemma Let_1 [simp]: "Let 1 f = f 1" + unfolding Let_def .. + +setup {* + ReorientProc.add + (fn Const(@{const_name HOL.zero}, _) => true + | Const(@{const_name HOL.one}, _) => true + | _ => false) +*} + +simproc_setup reorient_zero ("0 = x") = ReorientProc.proc +simproc_setup reorient_one ("1 = x") = ReorientProc.proc + +typed_print_translation {* +let + fun tr' c = (c, fn show_sorts => fn T => fn ts => + if (not o null) ts orelse T = dummyT + orelse not (! show_types) andalso can Term.dest_Type T + then raise Match + else Syntax.const Syntax.constrainC $ Syntax.const c $ Syntax.term_of_typ show_sorts T); +in map tr' [@{const_syntax HOL.one}, @{const_syntax HOL.zero}] end; +*} -- {* show types that are presumably too general *} + +hide (open) const zero one + +class plus = + fixes plus :: "'a \ 'a \ 'a" (infixl "+" 65) + +class minus = + fixes minus :: "'a \ 'a \ 'a" (infixl "-" 65) + +class uminus = + fixes uminus :: "'a \ 'a" ("- _" [81] 80) + +class times = + fixes times :: "'a \ 'a \ 'a" (infixl "*" 70) + +class inverse = + fixes inverse :: "'a \ 'a" + and divide :: "'a \ 'a \ 'a" (infixl "'/" 70) + +class abs = + fixes abs :: "'a \ 'a" +begin + +notation (xsymbols) + abs ("\_\") + +notation (HTML output) + abs ("\_\") + +end + +class sgn = + fixes sgn :: "'a \ 'a" + +class ord = + fixes less_eq :: "'a \ 'a \ bool" + and less :: "'a \ 'a \ bool" +begin + +notation + less_eq ("op <=") and + less_eq ("(_/ <= _)" [51, 51] 50) and + less ("op <") and + less ("(_/ < _)" [51, 51] 50) + +notation (xsymbols) + less_eq ("op \") and + less_eq ("(_/ \ _)" [51, 51] 50) + +notation (HTML output) + less_eq ("op \") and + less_eq ("(_/ \ _)" [51, 51] 50) + +abbreviation (input) + greater_eq (infix ">=" 50) where + "x >= y \ y <= x" + +notation (input) + greater_eq (infix "\" 50) + +abbreviation (input) + greater (infix ">" 50) where + "x > y \ y < x" + +end + +syntax + "_index1" :: index ("\<^sub>1") +translations + (index) "\<^sub>1" => (index) "\<^bsub>\\<^esub>" + +lemma mk_left_commute: + fixes f (infix "\" 60) + assumes a: "\x y z. (x \ y) \ z = x \ (y \ z)" and + c: "\x y. x \ y = y \ x" + shows "x \ (y \ z) = y \ (x \ z)" + by (rule trans [OF trans [OF c a] arg_cong [OF c, of "f y"]]) + + subsection {* Basic ML bindings *} ML {* @@ -1722,6 +1723,7 @@ setup {* Codegen.setup #> RecfunCodegen.setup + #> Codegen.map_unfold (K HOL_basic_ss) *} types_code @@ -1841,13 +1843,7 @@ 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 +declare imp_conv_disj [code, code_unfold_post] instantiation itself :: (type) eq begin diff -r 63686057cbe8 -r abda97d2deea src/HOL/Hoare/Hoare.thy --- a/src/HOL/Hoare/Hoare.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Hoare/Hoare.thy Thu Jul 23 21:13:21 2009 +0200 @@ -161,7 +161,7 @@ (* assn_tr' & bexp_tr'*) ML{* fun assn_tr' (Const ("Collect",_) $ T) = dest_abstuple T - | assn_tr' (Const (@{const_name Set.Int}, _) $ (Const ("Collect",_) $ T1) $ + | assn_tr' (Const (@{const_name inter}, _) $ (Const ("Collect",_) $ T1) $ (Const ("Collect",_) $ T2)) = Syntax.const "Set.Int" $ dest_abstuple T1 $ dest_abstuple T2 | assn_tr' t = t; @@ -239,7 +239,7 @@ method_setup vcg_simp = {* Scan.succeed (fn ctxt => - SIMPLE_METHOD' (hoare_tac ctxt (asm_full_simp_tac (local_simpset_of ctxt)))) *} + SIMPLE_METHOD' (hoare_tac ctxt (asm_full_simp_tac (simpset_of ctxt)))) *} "verification condition generator plus simplification" end diff -r 63686057cbe8 -r abda97d2deea src/HOL/Hoare/HoareAbort.thy --- a/src/HOL/Hoare/HoareAbort.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Hoare/HoareAbort.thy Thu Jul 23 21:13:21 2009 +0200 @@ -163,7 +163,7 @@ (* assn_tr' & bexp_tr'*) ML{* fun assn_tr' (Const ("Collect",_) $ T) = dest_abstuple T - | assn_tr' (Const (@{const_name Int},_) $ (Const ("Collect",_) $ T1) $ + | assn_tr' (Const (@{const_name inter},_) $ (Const ("Collect",_) $ T1) $ (Const ("Collect",_) $ T2)) = Syntax.const "Set.Int" $ dest_abstuple T1 $ dest_abstuple T2 | assn_tr' t = t; @@ -251,7 +251,7 @@ method_setup vcg_simp = {* Scan.succeed (fn ctxt => - SIMPLE_METHOD' (hoare_tac ctxt (asm_full_simp_tac (local_simpset_of ctxt)))) *} + SIMPLE_METHOD' (hoare_tac ctxt (asm_full_simp_tac (simpset_of ctxt)))) *} "verification condition generator plus simplification" (* Special syntax for guarded statements and guarded array updates: *) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Hoare/hoare_tac.ML --- a/src/HOL/Hoare/hoare_tac.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Hoare/hoare_tac.ML Thu Jul 23 21:13:21 2009 +0200 @@ -73,7 +73,7 @@ val MsetT = fastype_of big_Collect; fun Mset_incl t = mk_Trueprop (inclt MsetT $ Free (Mset, MsetT) $ t); val impl = Logic.mk_implies (Mset_incl big_Collect, Mset_incl small_Collect); - val th = Goal.prove ctxt [Mset, P] [] impl (fn _ => blast_tac (local_claset_of ctxt) 1); + val th = Goal.prove ctxt [Mset, P] [] impl (fn _ => blast_tac (claset_of ctxt) 1); in (vars, th) end; end; diff -r 63686057cbe8 -r abda97d2deea src/HOL/HoareParallel/Mul_Gar_Coll.thy --- a/src/HOL/HoareParallel/Mul_Gar_Coll.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/HoareParallel/Mul_Gar_Coll.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1263,7 +1263,7 @@ apply(simp add: Mul_Mutator_def mul_collector_defs mul_mutator_defs nth_append) apply force apply clarify -apply(case_tac xa) +apply(case_tac i) apply(simp add:Mul_Collector_def mul_mutator_defs mul_collector_defs nth_append) apply(simp add: Mul_Mutator_def mul_mutator_defs mul_collector_defs nth_append nth_map_upt) --{* Collector *} diff -r 63686057cbe8 -r abda97d2deea src/HOL/HoareParallel/OG_Hoare.thy --- a/src/HOL/HoareParallel/OG_Hoare.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/HoareParallel/OG_Hoare.thy Thu Jul 23 21:13:21 2009 +0200 @@ -441,7 +441,7 @@ apply clarify apply(frule Parallel_length_post_PStar) apply clarify - apply(drule_tac j=xa in Parallel_Strong_Soundness) + apply(drule_tac j=xb in Parallel_Strong_Soundness) apply clarify apply simp apply force diff -r 63686057cbe8 -r abda97d2deea src/HOL/Imperative_HOL/Heap_Monad.thy --- a/src/HOL/Imperative_HOL/Heap_Monad.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Imperative_HOL/Heap_Monad.thy Thu Jul 23 21:13:21 2009 +0200 @@ -283,7 +283,7 @@ where [code del]: "raise_exc e = raise []" -lemma raise_raise_exc [code, code_inline]: +lemma raise_raise_exc [code, code_unfold]: "raise s = raise_exc (Fail (STR s))" unfolding Fail_def raise_exc_def raise_def .. diff -r 63686057cbe8 -r abda97d2deea src/HOL/Import/import.ML --- a/src/HOL/Import/import.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Import/import.ML Thu Jul 23 21:13:21 2009 +0200 @@ -44,9 +44,9 @@ val thm = equal_elim rew thm val prew = ProofKernel.rewrite_hol4_term prem thy val prem' = #2 (Logic.dest_equals (prop_of prew)) - val _ = message ("Import proved " ^ Display.string_of_thm thm) + val _ = message ("Import proved " ^ Display.string_of_thm ctxt thm) val thm = ProofKernel.disambiguate_frees thm - val _ = message ("Disambiguate: " ^ Display.string_of_thm thm) + val _ = message ("Disambiguate: " ^ Display.string_of_thm ctxt thm) in case Shuffler.set_prop thy prem' [("",thm)] of SOME (_,thm) => diff -r 63686057cbe8 -r abda97d2deea src/HOL/Import/proof_kernel.ML --- a/src/HOL/Import/proof_kernel.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Import/proof_kernel.ML Thu Jul 23 21:13:21 2009 +0200 @@ -243,7 +243,7 @@ val smart_string_of_thm = smart_string_of_cterm o cprop_of -fun prth th = writeln (PrintMode.setmp [] Display.string_of_thm th) +fun prth th = writeln (PrintMode.setmp [] Display.string_of_thm_without_context th) fun prc ct = writeln (PrintMode.setmp [] Display.string_of_cterm ct) fun prin t = writeln (PrintMode.setmp [] (fn () => Syntax.string_of_term_global (the_context ()) t) ()); fun pth (HOLThm(ren,thm)) = diff -r 63686057cbe8 -r abda97d2deea src/HOL/Import/shuffler.ML --- a/src/HOL/Import/shuffler.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Import/shuffler.ML Thu Jul 23 21:13:21 2009 +0200 @@ -40,7 +40,7 @@ case e of THM (msg,i,thms) => (writeln ("Exception THM " ^ string_of_int i ^ " raised:\n" ^ msg); - List.app Display.print_thm thms) + List.app (writeln o Display.string_of_thm_global sign) thms) | THEORY (msg,thys) => (writeln ("Exception THEORY raised:\n" ^ msg); List.app (writeln o Context.str_of_thy) thys) @@ -56,7 +56,7 @@ (*Prints an exception, then fails*) fun print_sign_exn sign e = (print_sign_exn_unit sign e; raise e) -val string_of_thm = PrintMode.setmp [] Display.string_of_thm; +val string_of_thm = PrintMode.setmp [] Display.string_of_thm_without_context; val string_of_cterm = PrintMode.setmp [] Display.string_of_cterm; fun mk_meta_eq th = @@ -84,7 +84,7 @@ fun print_shuffles thy = Pretty.writeln (Pretty.big_list "Shuffle theorems:" - (map Display.pretty_thm (ShuffleData.get thy))) + (map (Display.pretty_thm_global thy) (ShuffleData.get thy))) val weaken = let diff -r 63686057cbe8 -r abda97d2deea src/HOL/Int.thy --- a/src/HOL/Int.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Int.thy Thu Jul 23 21:13:21 2009 +0200 @@ -2126,11 +2126,11 @@ hide (open) const nat_aux -lemma zero_is_num_zero [code, code_inline, symmetric, code_post]: +lemma zero_is_num_zero [code, code_unfold_post]: "(0\int) = Numeral0" by simp -lemma one_is_num_one [code, code_inline, symmetric, code_post]: +lemma one_is_num_one [code, code_unfold_post]: "(1\int) = Numeral1" by simp diff -r 63686057cbe8 -r abda97d2deea src/HOL/IntDiv.thy --- a/src/HOL/IntDiv.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/IntDiv.thy Thu Jul 23 21:13:21 2009 +0200 @@ -36,7 +36,7 @@ text{*algorithm for the general case @{term "b\0"}*} definition negateSnd :: "int \ int \ int \ int" where - [code_inline]: "negateSnd = apsnd uminus" + [code_unfold]: "negateSnd = apsnd uminus" definition divmod :: "int \ int \ int \ int" where --{*The full division algorithm considers all possible signs for a, b @@ -266,7 +266,7 @@ in -val cancel_div_mod_int_proc = Simplifier.simproc (the_context ()) +val cancel_div_mod_int_proc = Simplifier.simproc @{theory} "cancel_zdiv_zmod" ["(k::int) + l"] (K CancelDivMod.proc); val _ = Addsimprocs [cancel_div_mod_int_proc]; diff -r 63686057cbe8 -r abda97d2deea src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/IsaMakefile Thu Jul 23 21:13:21 2009 +0200 @@ -117,6 +117,7 @@ @$(ISABELLE_TOOL) usedir -b -f base.ML -d false -g false $(OUT)/Pure HOL-Base PLAIN_DEPENDENCIES = $(BASE_DEPENDENCIES)\ + Complete_Lattice.thy \ Datatype.thy \ Divides.thy \ Extraction.thy \ diff -r 63686057cbe8 -r abda97d2deea src/HOL/Lambda/WeakNorm.thy --- a/src/HOL/Lambda/WeakNorm.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Lambda/WeakNorm.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1,5 +1,4 @@ (* Title: HOL/Lambda/WeakNorm.thy - ID: $Id$ Author: Stefan Berghofer Copyright 2003 TU Muenchen *) @@ -430,11 +429,11 @@ val ct1 = @{cterm "%f. ((%f x. f (f (f x))) ((%f x. f (f (f (f x)))) f))"}; val (dB1, _) = @{code type_NF} (typing_of_term [] dummyf (term_of ct1)); -val ct1' = cterm_of (the_context ()) (term_of_dB [] (#T (rep_cterm ct1)) dB1); +val ct1' = cterm_of @{theory} (term_of_dB [] (#T (rep_cterm ct1)) dB1); val ct2 = @{cterm "%f x. (%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) x)))))"}; val (dB2, _) = @{code type_NF} (typing_of_term [] dummyf (term_of ct2)); -val ct2' = cterm_of (the_context ()) (term_of_dB [] (#T (rep_cterm ct2)) dB2); +val ct2' = cterm_of @{theory} (term_of_dB [] (#T (rep_cterm ct2)) dB2); *} text {* @@ -505,11 +504,11 @@ ML {* val ct1 = @{cterm "%f. ((%f x. f (f (f x))) ((%f x. f (f (f (f x)))) f))"}; val (dB1, _) = Norm.type_NF (typing_of_term [] dummyf (term_of ct1)); -val ct1' = cterm_of (the_context ()) (term_of_dB [] (#T (rep_cterm ct1)) dB1); +val ct1' = cterm_of @{theory} (term_of_dB [] (#T (rep_cterm ct1)) dB1); val ct2 = @{cterm "%f x. (%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) x)))))"}; val (dB2, _) = Norm.type_NF (typing_of_term [] dummyf (term_of ct2)); -val ct2' = cterm_of (the_context ()) (term_of_dB [] (#T (rep_cterm ct2)) dB2); +val ct2' = cterm_of @{theory} (term_of_dB [] (#T (rep_cterm ct2)) dB2); *} end diff -r 63686057cbe8 -r abda97d2deea src/HOL/Lattices.thy --- a/src/HOL/Lattices.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Lattices.thy Thu Jul 23 21:13:21 2009 +0200 @@ -44,38 +44,27 @@ context lower_semilattice begin -lemma le_infI1[intro]: - assumes "a \ x" - shows "a \ b \ x" -proof (rule order_trans) - from assms show "a \ x" . - show "a \ b \ a" by simp -qed -lemmas (in -) [rule del] = le_infI1 +lemma le_infI1: + "a \ x \ a \ b \ x" + by (rule order_trans) auto -lemma le_infI2[intro]: - assumes "b \ x" - shows "a \ b \ x" -proof (rule order_trans) - from assms show "b \ x" . - show "a \ b \ b" by simp -qed -lemmas (in -) [rule del] = le_infI2 +lemma le_infI2: + "b \ x \ a \ b \ x" + by (rule order_trans) auto -lemma le_infI[intro!]: "x \ a \ x \ b \ x \ a \ b" -by(blast intro: inf_greatest) -lemmas (in -) [rule del] = le_infI +lemma le_infI: "x \ a \ x \ b \ x \ a \ b" + by (blast intro: inf_greatest) -lemma le_infE [elim!]: "x \ a \ b \ (x \ a \ x \ b \ P) \ P" - by (blast intro: order_trans) -lemmas (in -) [rule del] = le_infE +lemma le_infE: "x \ a \ b \ (x \ a \ x \ b \ P) \ P" + by (blast intro: order_trans le_infI1 le_infI2) lemma le_inf_iff [simp]: - "x \ y \ z = (x \ y \ x \ z)" -by blast + "x \ y \ z \ x \ y \ x \ z" + by (blast intro: le_infI elim: le_infE) -lemma le_iff_inf: "(x \ y) = (x \ y = x)" - by (blast intro: antisym dest: eq_iff [THEN iffD1]) +lemma le_iff_inf: + "x \ y \ x \ y = x" + by (auto intro: le_infI1 antisym dest: eq_iff [THEN iffD1]) lemma mono_inf: fixes f :: "'a \ 'b\lower_semilattice" @@ -87,28 +76,29 @@ context upper_semilattice begin -lemma le_supI1[intro]: "x \ a \ x \ a \ b" +lemma le_supI1: + "x \ a \ x \ a \ b" by (rule order_trans) auto -lemmas (in -) [rule del] = le_supI1 -lemma le_supI2[intro]: "x \ b \ x \ a \ b" +lemma le_supI2: + "x \ b \ x \ a \ b" by (rule order_trans) auto -lemmas (in -) [rule del] = le_supI2 -lemma le_supI[intro!]: "a \ x \ b \ x \ a \ b \ x" +lemma le_supI: + "a \ x \ b \ x \ a \ b \ x" by (blast intro: sup_least) -lemmas (in -) [rule del] = le_supI -lemma le_supE[elim!]: "a \ b \ x \ (a \ x \ b \ x \ P) \ P" - by (blast intro: order_trans) -lemmas (in -) [rule del] = le_supE +lemma le_supE: + "a \ b \ x \ (a \ x \ b \ x \ P) \ P" + by (blast intro: le_supI1 le_supI2 order_trans) -lemma ge_sup_conv[simp]: - "x \ y \ z = (x \ z \ y \ z)" -by blast +lemma le_sup_iff [simp]: + "x \ y \ z \ x \ z \ y \ z" + by (blast intro: le_supI elim: le_supE) -lemma le_iff_sup: "(x \ y) = (x \ y = y)" - by (blast intro: antisym dest: eq_iff [THEN iffD1]) +lemma le_iff_sup: + "x \ y \ x \ y = y" + by (auto intro: le_supI2 antisym dest: eq_iff [THEN iffD1]) lemma mono_sup: fixes f :: "'a \ 'b\upper_semilattice" @@ -118,62 +108,61 @@ end -subsubsection{* Equational laws *} +subsubsection {* Equational laws *} context lower_semilattice begin lemma inf_commute: "(x \ y) = (y \ x)" - by (blast intro: antisym) + by (rule antisym) auto lemma inf_assoc: "(x \ y) \ z = x \ (y \ z)" - by (blast intro: antisym) + by (rule antisym) (auto intro: le_infI1 le_infI2) lemma inf_idem[simp]: "x \ x = x" - by (blast intro: antisym) + by (rule antisym) auto lemma inf_left_idem[simp]: "x \ (x \ y) = x \ y" - by (blast intro: antisym) + by (rule antisym) (auto intro: le_infI2) lemma inf_absorb1: "x \ y \ x \ y = x" - by (blast intro: antisym) + by (rule antisym) auto lemma inf_absorb2: "y \ x \ x \ y = y" - by (blast intro: antisym) + by (rule antisym) auto lemma inf_left_commute: "x \ (y \ z) = y \ (x \ z)" - by (blast intro: antisym) - -lemmas inf_ACI = inf_commute inf_assoc inf_left_commute inf_left_idem + by (rule mk_left_commute [of inf]) (fact inf_assoc inf_commute)+ + +lemmas inf_aci = inf_commute inf_assoc inf_left_commute inf_left_idem end - context upper_semilattice begin lemma sup_commute: "(x \ y) = (y \ x)" - by (blast intro: antisym) + by (rule antisym) auto lemma sup_assoc: "(x \ y) \ z = x \ (y \ z)" - by (blast intro: antisym) + by (rule antisym) (auto intro: le_supI1 le_supI2) lemma sup_idem[simp]: "x \ x = x" - by (blast intro: antisym) + by (rule antisym) auto lemma sup_left_idem[simp]: "x \ (x \ y) = x \ y" - by (blast intro: antisym) + by (rule antisym) (auto intro: le_supI2) lemma sup_absorb1: "y \ x \ x \ y = x" - by (blast intro: antisym) + by (rule antisym) auto lemma sup_absorb2: "x \ y \ x \ y = y" - by (blast intro: antisym) + by (rule antisym) auto lemma sup_left_commute: "x \ (y \ z) = y \ (x \ z)" - by (blast intro: antisym) + by (rule mk_left_commute [of sup]) (fact sup_assoc sup_commute)+ -lemmas sup_ACI = sup_commute sup_assoc sup_left_commute sup_left_idem +lemmas sup_aci = sup_commute sup_assoc sup_left_commute sup_left_idem end @@ -191,18 +180,17 @@ lemma sup_inf_absorb: "x \ (x \ y) = x" by (blast intro: antisym sup_ge1 sup_least inf_le1) -lemmas ACI = inf_ACI sup_ACI +lemmas inf_sup_aci = inf_aci sup_aci lemmas inf_sup_ord = inf_le1 inf_le2 sup_ge1 sup_ge2 text{* Towards distributivity *} lemma distrib_sup_le: "x \ (y \ z) \ (x \ y) \ (x \ z)" - by blast + by (auto intro: le_infI1 le_infI2 le_supI1 le_supI2) lemma distrib_inf_le: "(x \ y) \ (x \ z) \ x \ (y \ z)" - by blast - + by (auto intro: le_infI1 le_infI2 le_supI1 le_supI2) text{* If you have one of them, you have them all. *} @@ -230,10 +218,6 @@ finally show ?thesis . qed -(* seems unused *) -lemma modular_le: "x \ z \ x \ (y \ z) \ (x \ y) \ z" -by blast - end @@ -247,7 +231,7 @@ lemma sup_inf_distrib2: "(y \ z) \ x = (y \ x) \ (z \ x)" -by(simp add:ACI sup_inf_distrib1) +by(simp add: inf_sup_aci sup_inf_distrib1) lemma inf_sup_distrib1: "x \ (y \ z) = (x \ y) \ (x \ z)" @@ -255,7 +239,7 @@ lemma inf_sup_distrib2: "(y \ z) \ x = (y \ x) \ (z \ x)" -by(simp add:ACI inf_sup_distrib1) +by(simp add: inf_sup_aci inf_sup_distrib1) lemma dual_distrib_lattice: "distrib_lattice (op \) (op >) sup inf" @@ -458,16 +442,6 @@ lemmas min_ac = min_max.inf_assoc min_max.inf_commute mk_left_commute [of min, OF min_max.inf_assoc min_max.inf_commute] -text {* - Now we have inherited antisymmetry as an intro-rule on all - linear orders. This is a problem because it applies to bool, which is - undesirable. -*} - -lemmas [rule del] = min_max.le_infI min_max.le_supI - min_max.le_supE min_max.le_infE min_max.le_supI1 min_max.le_supI2 - min_max.le_infI1 min_max.le_infI2 - subsection {* Bool as lattice *} @@ -548,8 +522,6 @@ text {* redundant bindings *} -lemmas inf_aci = inf_ACI -lemmas sup_aci = sup_ACI no_notation less_eq (infix "\" 50) and diff -r 63686057cbe8 -r abda97d2deea src/HOL/Library/Binomial.thy --- a/src/HOL/Library/Binomial.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Library/Binomial.thy Thu Jul 23 21:13:21 2009 +0200 @@ -243,14 +243,8 @@ ultimately show ?thesis by blast qed -lemma fact_setprod: "fact n = setprod id {1 .. n}" - apply (induct n, simp) - apply (simp only: fact_Suc atLeastAtMostSuc_conv) - apply (subst setprod_insert) - by simp_all - lemma pochhammer_fact: "of_nat (fact n) = pochhammer 1 n" - unfolding fact_setprod + unfolding fact_altdef_nat apply (cases n, simp_all add: of_nat_setprod pochhammer_Suc_setprod) apply (rule setprod_reindex_cong[where f=Suc]) @@ -397,7 +391,7 @@ assumes kn: "k \ n" shows "(of_nat (n choose k) :: 'a::field_char_0) = of_nat (fact n) / (of_nat (fact k) * of_nat (fact (n - k)))" using binomial_fact_lemma[OF kn] - by (simp add: field_simps fact_not_eq_zero of_nat_mult[symmetric]) + by (simp add: field_simps of_nat_mult[symmetric]) lemma binomial_gbinomial: "of_nat (n choose k) = of_nat n gchoose k" proof- @@ -427,7 +421,7 @@ have ?thesis using kn apply (simp add: binomial_fact[OF kn, where ?'a = 'a] gbinomial_pochhammer field_simps pochhammer_Suc_setprod) - apply (simp add: pochhammer_Suc_setprod fact_setprod h of_nat_setprod setprod_timesf[symmetric] eq' del: One_nat_def power_Suc) + apply (simp add: pochhammer_Suc_setprod fact_altdef_nat h of_nat_setprod setprod_timesf[symmetric] eq' del: One_nat_def power_Suc) unfolding setprod_Un_disjoint[OF th0, unfolded eq3, of "of_nat:: nat \ 'a"] eq[unfolded h] unfolding mult_assoc[symmetric] unfolding setprod_timesf[symmetric] diff -r 63686057cbe8 -r abda97d2deea src/HOL/Library/Code_Char_chr.thy --- a/src/HOL/Library/Code_Char_chr.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Library/Code_Char_chr.thy Thu Jul 23 21:13:21 2009 +0200 @@ -24,11 +24,11 @@ lemmas [code del] = char.recs char.cases char.size -lemma [code, code_inline]: +lemma [code, code_unfold]: "char_rec f c = split f (nibble_pair_of_nat (nat_of_char c))" by (cases c) (auto simp add: nibble_pair_of_nat_char) -lemma [code, code_inline]: +lemma [code, code_unfold]: "char_case f c = split f (nibble_pair_of_nat (nat_of_char c))" by (cases c) (auto simp add: nibble_pair_of_nat_char) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Library/Efficient_Nat.thy --- a/src/HOL/Library/Efficient_Nat.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Library/Efficient_Nat.thy Thu Jul 23 21:13:21 2009 +0200 @@ -26,15 +26,13 @@ code_datatype number_nat_inst.number_of_nat -lemma zero_nat_code [code, code_inline]: +lemma zero_nat_code [code, code_unfold_post]: "0 = (Numeral0 :: nat)" by simp -lemmas [code_post] = zero_nat_code [symmetric] -lemma one_nat_code [code, code_inline]: +lemma one_nat_code [code, code_unfold_post]: "1 = (Numeral1 :: nat)" by simp -lemmas [code_post] = one_nat_code [symmetric] lemma Suc_code [code]: "Suc n = n + 1" @@ -302,7 +300,7 @@ Natural numerals. *} -lemma [code_inline, symmetric, code_post]: +lemma [code_unfold_post]: "nat (number_of i) = number_nat_inst.number_of_nat i" -- {* this interacts as desired with @{thm nat_number_of_def} *} by (simp add: number_nat_inst.number_of_nat) @@ -322,9 +320,7 @@ returns @{text "0"}. *} -definition - int :: "nat \ int" -where +definition int :: "nat \ int" where [code del]: "int = of_nat" lemma int_code' [code]: @@ -335,9 +331,12 @@ "nat (number_of l) = (if neg (number_of l \ int) then 0 else number_of l)" unfolding nat_number_of_def number_of_is_id neg_def by simp -lemma of_nat_int [code_unfold]: +lemma of_nat_int [code_unfold_post]: "of_nat = int" by (simp add: int_def) -declare of_nat_int [symmetric, code_post] + +lemma of_nat_aux_int [code_unfold]: + "of_nat_aux (\i. i + 1) k 0 = int k" + by (simp add: int_def Nat.of_nat_code) code_const int (SML "_") diff -r 63686057cbe8 -r abda97d2deea src/HOL/Library/Executable_Set.thy --- a/src/HOL/Library/Executable_Set.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Library/Executable_Set.thy Thu Jul 23 21:13:21 2009 +0200 @@ -75,8 +75,8 @@ "op \" ("{*Fset.union*}") "op \" ("{*Fset.inter*}") "op - \ 'a set \ 'a set \ 'a set" ("{*flip Fset.subtract*}") - "Set.Union" ("{*Fset.Union*}") - "Set.Inter" ("{*Fset.Inter*}") + "Complete_Lattice.Union" ("{*Fset.Union*}") + "Complete_Lattice.Inter" ("{*Fset.Inter*}") card ("{*Fset.card*}") fold ("{*foldl o flip*}") diff -r 63686057cbe8 -r abda97d2deea src/HOL/Library/Float.thy --- a/src/HOL/Library/Float.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Library/Float.thy Thu Jul 23 21:13:21 2009 +0200 @@ -42,7 +42,7 @@ instance .. end -lemma number_of_float_Float [code_inline, symmetric, code_post]: +lemma number_of_float_Float [code_unfold_post]: "number_of k = Float (number_of k) 0" by (simp add: number_of_float_def number_of_is_id) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Library/Formal_Power_Series.thy --- a/src/HOL/Library/Formal_Power_Series.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Library/Formal_Power_Series.thy Thu Jul 23 21:13:21 2009 +0200 @@ -2528,8 +2528,8 @@ apply (induct n) apply simp unfolding th - using fact_gt_zero - apply (simp add: field_simps del: of_nat_Suc fact.simps) + using fact_gt_zero_nat + apply (simp add: field_simps del: of_nat_Suc fact_Suc) apply (drule sym) by (simp add: ring_simps of_nat_mult power_Suc)} note th' = this @@ -3041,9 +3041,6 @@ finally show ?thesis . qed -lemma fact_1 [simp]: "fact 1 = 1" -unfolding One_nat_def fact_Suc by simp - lemma divide_eq_iff: "a \ (0::'a::field) \ x / a = y \ x = y * a" by auto diff -r 63686057cbe8 -r abda97d2deea src/HOL/Library/Fset.thy --- a/src/HOL/Library/Fset.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Library/Fset.thy Thu Jul 23 21:13:21 2009 +0200 @@ -160,7 +160,7 @@ qed definition Inter :: "'a fset fset \ 'a fset" where - [simp]: "Inter A = Fset (Set.Inter (member ` member A))" + [simp]: "Inter A = Fset (Complete_Lattice.Inter (member ` member A))" lemma Inter_inter [code]: "Inter (Set (A # As)) = foldl inter A As" @@ -174,7 +174,7 @@ qed definition Union :: "'a fset fset \ 'a fset" where - [simp]: "Union A = Fset (Set.Union (member ` member A))" + [simp]: "Union A = Fset (Complete_Lattice.Union (member ` member A))" lemma Union_union [code]: "Union (Set As) = foldl union empty As" diff -r 63686057cbe8 -r abda97d2deea src/HOL/Library/Lattice_Syntax.thy --- a/src/HOL/Library/Lattice_Syntax.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Library/Lattice_Syntax.thy Thu Jul 23 21:13:21 2009 +0200 @@ -4,16 +4,16 @@ (*<*) theory Lattice_Syntax -imports Set +imports Complete_Lattice begin notation + top ("\") and + bot ("\") and inf (infixl "\" 70) and sup (infixl "\" 65) and Inf ("\_" [900] 900) and - Sup ("\_" [900] 900) and - top ("\") and - bot ("\") + Sup ("\_" [900] 900) end (*>*) \ No newline at end of file diff -r 63686057cbe8 -r abda97d2deea src/HOL/Library/Nat_Infinity.thy --- a/src/HOL/Library/Nat_Infinity.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Library/Nat_Infinity.thy Thu Jul 23 21:13:21 2009 +0200 @@ -40,10 +40,10 @@ "0 = Fin 0" definition - [code_inline]: "1 = Fin 1" + [code_unfold]: "1 = Fin 1" definition - [code_inline, code del]: "number_of k = Fin (number_of k)" + [code_unfold, code del]: "number_of k = Fin (number_of k)" instance .. diff -r 63686057cbe8 -r abda97d2deea src/HOL/Library/comm_ring.ML --- a/src/HOL/Library/comm_ring.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Library/comm_ring.ML Thu Jul 23 21:13:21 2009 +0200 @@ -89,7 +89,7 @@ fun comm_ring_tac ctxt = SUBGOAL (fn (g, i) => let val thy = ProofContext.theory_of ctxt; - val cring_ss = Simplifier.local_simpset_of ctxt (*FIXME really the full simpset!?*) + val cring_ss = Simplifier.simpset_of ctxt (*FIXME really the full simpset!?*) addsimps cring_simps; val (ca, cvs, clhs, crhs) = reif_eq thy (HOLogic.dest_Trueprop g) val norm_eq_th = diff -r 63686057cbe8 -r abda97d2deea src/HOL/Library/reflection.ML --- a/src/HOL/Library/reflection.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Library/reflection.ML Thu Jul 23 21:13:21 2009 +0200 @@ -153,8 +153,8 @@ val certy = ctyp_of thy val (tyenv, tmenv) = Pattern.match thy - ((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (concl_of cong), t) - (Envir.type_env (Envir.empty 0), Vartab.empty) + ((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (concl_of cong), t) + (Vartab.empty, Vartab.empty) val (fnvs,invs) = List.partition (fn ((vn,_),_) => vn mem vns) (Vartab.dest tmenv) val (fts,its) = (map (snd o snd) fnvs, @@ -204,11 +204,9 @@ val xns_map = (fst (split_list nths)) ~~ xns val subst = map (fn (nt, xn) => (nt, Var ((xn,0), fastype_of nt))) xns_map val rhs_P = subst_free subst rhs - val (tyenv, tmenv) = Pattern.match - thy (rhs_P, t) - (Envir.type_env (Envir.empty 0), Vartab.empty) - val sbst = Envir.subst_vars (tyenv, tmenv) - val sbsT = Envir.typ_subst_TVars tyenv + val (tyenv, tmenv) = Pattern.match thy (rhs_P, t) (Vartab.empty, Vartab.empty) + val sbst = Envir.subst_term (tyenv, tmenv) + val sbsT = Envir.subst_type tyenv val subst_ty = map (fn (n,(s,t)) => (certT (TVar (n, s)), certT t)) (Vartab.dest tyenv) val tml = Vartab.dest tmenv @@ -287,7 +285,7 @@ val th' = instantiate ([], cvs) th val t' = (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) th' val th'' = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t'))) - (fn _ => simp_tac (local_simpset_of ctxt) 1) + (fn _ => simp_tac (simpset_of ctxt) 1) in FWD trans [th'',th'] end diff -r 63686057cbe8 -r abda97d2deea src/HOL/Library/sum_of_squares.ML --- a/src/HOL/Library/sum_of_squares.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Library/sum_of_squares.ML Thu Jul 23 21:13:21 2009 +0200 @@ -1686,7 +1686,7 @@ NONE => no_tac | SOME (d,ord) => let - val ss = simpset_of (ProofContext.theory_of ctxt) addsimps @{thms field_simps} + val ss = simpset_of ctxt addsimps @{thms field_simps} addsimps [@{thm nonzero_power_divide}, @{thm power_divide}] val th = instantiate' [] [SOME d, SOME (Thm.dest_arg P)] (if ord then @{lemma "(d=0 --> P) & (d>0 --> P) & (d<(0::real) --> P) ==> P" by auto} diff -r 63686057cbe8 -r abda97d2deea src/HOL/List.thy --- a/src/HOL/List.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/List.thy Thu Jul 23 21:13:21 2009 +0200 @@ -679,7 +679,7 @@ in val list_eq_simproc = - Simplifier.simproc (the_context ()) "list_eq" ["(xs::'a list) = ys"] (K list_eq); + Simplifier.simproc @{theory} "list_eq" ["(xs::'a list) = ys"] (K list_eq); end; @@ -2268,6 +2268,8 @@ -- {* simp does not terminate! *} by (induct j) auto +lemmas upt_rec_number_of[simp] = upt_rec[of "number_of m" "number_of n", standard] + lemma upt_conv_Nil [simp]: "j <= i ==> [i..int. i+1)" +successor_int_def[simp]: "successor = (%i\int. i+1)" instance by intro_classes (simp_all add: successor_int_def) @@ -3202,6 +3204,8 @@ lemma upto_rec2_int: "(i::int) \ j \ [i..j+1] = [i..j] @ [j+1]" by(rule upto_rec2[where 'a = int, simplified successor_int_def]) +lemmas upto_rec_number_of_int[simp] = upto_rec[of "number_of m :: int" "number_of n", standard] + subsubsection {* @{text lists}: the list-forming operator over sets *} @@ -3756,7 +3760,7 @@ "xs = [] \ null xs" by (cases xs) simp_all -lemma [code_inline]: +lemma [code_unfold]: "eq_class.eq xs [] \ null xs" by (simp add: eq empty_null) @@ -3804,7 +3808,7 @@ "map_filter f P xs = map f (filter P xs)" by (induct xs) auto -lemma length_remdups_length_unique [code_inline]: +lemma length_remdups_length_unique [code_unfold]: "length (remdups xs) = length_unique xs" by (induct xs) simp_all diff -r 63686057cbe8 -r abda97d2deea src/HOL/Ln.thy --- a/src/HOL/Ln.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Ln.thy Thu Jul 23 21:13:21 2009 +0200 @@ -31,13 +31,13 @@ qed lemma aux1: assumes a: "0 <= x" and b: "x <= 1" - shows "inverse (real (fact (n + 2))) * x ^ (n + 2) <= (x^2/2) * ((1/2)^n)" + shows "inverse (real (fact ((n::nat) + 2))) * x ^ (n + 2) <= (x^2/2) * ((1/2)^n)" proof (induct n) - show "inverse (real (fact (0 + 2))) * x ^ (0 + 2) <= + show "inverse (real (fact ((0::nat) + 2))) * x ^ (0 + 2) <= x ^ 2 / 2 * (1 / 2) ^ 0" by (simp add: real_of_nat_Suc power2_eq_square) next - fix n + fix n :: nat assume c: "inverse (real (fact (n + 2))) * x ^ (n + 2) <= x ^ 2 / 2 * (1 / 2) ^ n" show "inverse (real (fact (Suc n + 2))) * x ^ (Suc n + 2) diff -r 63686057cbe8 -r abda97d2deea src/HOL/MacLaurin.thy --- a/src/HOL/MacLaurin.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/MacLaurin.thy Thu Jul 23 21:13:21 2009 +0200 @@ -27,6 +27,10 @@ lemma eq_diff_eq': "(x = y - z) = (y = x + (z::real))" by arith +lemma fact_diff_Suc [rule_format]: + "n < Suc m ==> fact (Suc m - n) = (Suc m - n) * fact (m - n)" + by (subst fact_reduce_nat, auto) + lemma Maclaurin_lemma2: assumes diff: "\m t. m < n \ 0\t \ t\h \ DERIV (diff m) t :> diff (Suc m) t" assumes n: "n = Suc k" @@ -47,7 +51,9 @@ apply (rule_tac [2] DERIV_quotient) apply (rule_tac [3] DERIV_const) apply (rule_tac [2] DERIV_pow) - prefer 3 apply (simp add: fact_diff_Suc) + prefer 3 + +apply (simp add: fact_diff_Suc) prefer 2 apply simp apply (frule less_iff_Suc_add [THEN iffD1], clarify) apply (simp del: setsum_op_ivl_Suc) diff -r 63686057cbe8 -r abda97d2deea src/HOL/MicroJava/J/JTypeSafe.thy --- a/src/HOL/MicroJava/J/JTypeSafe.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/MicroJava/J/JTypeSafe.thy Thu Jul 23 21:13:21 2009 +0200 @@ -204,7 +204,7 @@ apply( simp_all) apply( tactic "ALLGOALS strip_tac") apply( tactic {* ALLGOALS (eresolve_tac [thm "ty_expr.cases", thm "ty_exprs.cases", thm "wt_stmt.cases"] - THEN_ALL_NEW (full_simp_tac (simpset_of @{theory Conform}))) *}) + THEN_ALL_NEW (full_simp_tac (global_simpset_of @{theory Conform}))) *}) apply(tactic "ALLGOALS (EVERY' [REPEAT o (etac conjE), REPEAT o hyp_subst_tac])") -- "Level 7" diff -r 63686057cbe8 -r abda97d2deea src/HOL/Modelcheck/EindhovenSyn.thy --- a/src/HOL/Modelcheck/EindhovenSyn.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Modelcheck/EindhovenSyn.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1,5 +1,4 @@ (* Title: HOL/Modelcheck/EindhovenSyn.thy - ID: $Id$ Author: Olaf Mueller, Jan Philipps, Robert Sandner Copyright 1997 TU Muenchen *) @@ -70,7 +69,7 @@ val pair_eta_expand = Thm.symmetric (mk_meta_eq (thm "split_eta")); val pair_eta_expand_proc = - Simplifier.simproc (the_context ()) "pair_eta_expand" ["f::'a*'b=>'c"] + Simplifier.simproc @{theory} "pair_eta_expand" ["f::'a*'b=>'c"] (fn _ => fn _ => fn t => case t of Abs _ => SOME pair_eta_expand | _ => NONE); val Eindhoven_ss = diff -r 63686057cbe8 -r abda97d2deea src/HOL/Modelcheck/MuckeSyn.thy --- a/src/HOL/Modelcheck/MuckeSyn.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Modelcheck/MuckeSyn.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1,5 +1,4 @@ (* Title: HOL/Modelcheck/MuckeSyn.thy - ID: $Id$ Author: Tobias Hamberger Copyright 1999 TU Muenchen *) @@ -119,7 +118,7 @@ local - val move_thm = prove_goal (the_context ()) "[| a = b ==> P a; a = b |] ==> P b" + val move_thm = prove_goal @{theory} "[| a = b ==> P a; a = b |] ==> P b" (fn prems => [cut_facts_tac prems 1, dtac sym 1, hyp_subst_tac 1, REPEAT (resolve_tac prems 1)]); @@ -214,7 +213,7 @@ val pair_eta_expand = Thm.symmetric (mk_meta_eq (thm "split_eta")); val pair_eta_expand_proc = - Simplifier.simproc (the_context ()) "pair_eta_expand" ["f::'a*'b=>'c"] + Simplifier.simproc @{theory} "pair_eta_expand" ["f::'a*'b=>'c"] (fn _ => fn _ => fn t => case t of Abs _ => SOME pair_eta_expand | _ => NONE); val Mucke_ss = @{simpset} addsimprocs [pair_eta_expand_proc] addsimps [Let_def]; diff -r 63686057cbe8 -r abda97d2deea src/HOL/Modelcheck/mucke_oracle.ML --- a/src/HOL/Modelcheck/mucke_oracle.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Modelcheck/mucke_oracle.ML Thu Jul 23 21:13:21 2009 +0200 @@ -109,7 +109,7 @@ ( OldGoals.push_proof(); OldGoals.goalw_cterm [] (cterm_of sign trm); -by (simp_tac (simpset_of sign) 1); +by (simp_tac (global_simpset_of sign) 1); let val if_tmp_result = result() in diff -r 63686057cbe8 -r abda97d2deea src/HOL/NatTransfer.thy --- a/src/HOL/NatTransfer.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/NatTransfer.thy Thu Jul 23 21:13:21 2009 +0200 @@ -380,12 +380,16 @@ "(even (int x)) = (even x)" by (auto simp add: zdvd_int even_nat_def) +lemma UNIV_apply: + "UNIV x = True" + by (simp add: top_set_eq [symmetric] top_fun_eq top_bool_eq) + declare TransferMorphism_int_nat[transfer add return: transfer_int_nat_numerals transfer_int_nat_functions transfer_int_nat_function_closures transfer_int_nat_relations - UNIV_code + UNIV_apply ] diff -r 63686057cbe8 -r abda97d2deea src/HOL/Nat_Numeral.thy --- a/src/HOL/Nat_Numeral.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Nat_Numeral.thy Thu Jul 23 21:13:21 2009 +0200 @@ -20,7 +20,7 @@ begin definition - nat_number_of_def [code_inline, code del]: "number_of v = nat (number_of v)" + nat_number_of_def [code_unfold, code del]: "number_of v = nat (number_of v)" instance .. diff -r 63686057cbe8 -r abda97d2deea src/HOL/NewNumberTheory/Binomial.thy --- a/src/HOL/NewNumberTheory/Binomial.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/NewNumberTheory/Binomial.thy Thu Jul 23 21:13:21 2009 +0200 @@ -2,7 +2,7 @@ Authors: Lawrence C. Paulson, Jeremy Avigad, Tobias Nipkow -Defines factorial and the "choose" function, and establishes basic properties. +Defines the "choose" function, and establishes basic properties. The original theory "Binomial" was by Lawrence C. Paulson, based on the work of Andy Gordon and Florian Kammueller. The approach here, @@ -16,7 +16,7 @@ header {* Binomial *} theory Binomial -imports Cong +imports Cong Fact begin @@ -25,7 +25,6 @@ class binomial = fixes - fact :: "'a \ 'a" and binomial :: "'a \ 'a \ 'a" (infixl "choose" 65) (* definitions for the natural numbers *) @@ -35,13 +34,6 @@ begin fun - fact_nat :: "nat \ nat" -where - "fact_nat x = - (if x = 0 then 1 else - x * fact(x - 1))" - -fun binomial_nat :: "nat \ nat \ nat" where "binomial_nat n k = @@ -60,11 +52,6 @@ begin definition - fact_int :: "int \ int" -where - "fact_int x = (if x >= 0 then int (fact (nat x)) else 0)" - -definition binomial_int :: "int => int \ int" where "binomial_int n k = (if n \ 0 \ k \ 0 then int (binomial (nat n) (nat k)) @@ -76,182 +63,29 @@ subsection {* Set up Transfer *} - lemma transfer_nat_int_binomial: - "(x::int) >= 0 \ fact (nat x) = nat (fact x)" "(n::int) >= 0 \ k >= 0 \ binomial (nat n) (nat k) = nat (binomial n k)" - unfolding fact_int_def binomial_int_def + unfolding binomial_int_def by auto - -lemma transfer_nat_int_binomial_closures: - "x >= (0::int) \ fact x >= 0" +lemma transfer_nat_int_binomial_closure: "n >= (0::int) \ k >= 0 \ binomial n k >= 0" - by (auto simp add: fact_int_def binomial_int_def) + by (auto simp add: binomial_int_def) declare TransferMorphism_nat_int[transfer add return: - transfer_nat_int_binomial transfer_nat_int_binomial_closures] + transfer_nat_int_binomial transfer_nat_int_binomial_closure] lemma transfer_int_nat_binomial: - "fact (int x) = int (fact x)" "binomial (int n) (int k) = int (binomial n k)" unfolding fact_int_def binomial_int_def by auto -lemma transfer_int_nat_binomial_closures: - "is_nat x \ fact x >= 0" +lemma transfer_int_nat_binomial_closure: "is_nat n \ is_nat k \ binomial n k >= 0" - by (auto simp add: fact_int_def binomial_int_def) + by (auto simp add: binomial_int_def) declare TransferMorphism_int_nat[transfer add return: - transfer_int_nat_binomial transfer_int_nat_binomial_closures] - - -subsection {* Factorial *} - -lemma fact_zero_nat [simp]: "fact (0::nat) = 1" - by simp - -lemma fact_zero_int [simp]: "fact (0::int) = 1" - by (simp add: fact_int_def) - -lemma fact_one_nat [simp]: "fact (1::nat) = 1" - by simp - -lemma fact_Suc_0_nat [simp]: "fact (Suc 0) = Suc 0" - by (simp add: One_nat_def) - -lemma fact_one_int [simp]: "fact (1::int) = 1" - by (simp add: fact_int_def) - -lemma fact_plus_one_nat: "fact ((n::nat) + 1) = (n + 1) * fact n" - by simp - -lemma fact_Suc_nat: "fact (Suc n) = (Suc n) * fact n" - by (simp add: One_nat_def) - -lemma fact_plus_one_int: - assumes "n >= 0" - shows "fact ((n::int) + 1) = (n + 1) * fact n" - - using prems by (rule fact_plus_one_nat [transferred]) - -lemma fact_reduce_nat: "(n::nat) > 0 \ fact n = n * fact (n - 1)" - by simp - -lemma fact_reduce_int: - assumes "(n::int) > 0" - shows "fact n = n * fact (n - 1)" - - using prems apply (subst tsub_eq [symmetric], auto) - apply (rule fact_reduce_nat [transferred]) - using prems apply auto -done - -declare fact_nat.simps [simp del] - -lemma fact_nonzero_nat [simp]: "fact (n::nat) \ 0" - apply (induct n rule: induct'_nat) - apply (auto simp add: fact_plus_one_nat) -done - -lemma fact_nonzero_int [simp]: "n >= 0 \ fact (n::int) ~= 0" - by (simp add: fact_int_def) - -lemma fact_gt_zero_nat [simp]: "fact (n :: nat) > 0" - by (insert fact_nonzero_nat [of n], arith) - -lemma fact_gt_zero_int [simp]: "n >= 0 \ fact (n :: int) > 0" - by (auto simp add: fact_int_def) - -lemma fact_ge_one_nat [simp]: "fact (n :: nat) >= 1" - by (insert fact_nonzero_nat [of n], arith) - -lemma fact_ge_Suc_0_nat [simp]: "fact (n :: nat) >= Suc 0" - by (insert fact_nonzero_nat [of n], arith) - -lemma fact_ge_one_int [simp]: "n >= 0 \ fact (n :: int) >= 1" - apply (auto simp add: fact_int_def) - apply (subgoal_tac "1 = int 1") - apply (erule ssubst) - apply (subst zle_int) - apply auto -done - -lemma dvd_fact_nat [rule_format]: "1 <= m \ m <= n \ m dvd fact (n::nat)" - apply (induct n rule: induct'_nat) - apply (auto simp add: fact_plus_one_nat) - apply (subgoal_tac "m = n + 1") - apply auto -done - -lemma dvd_fact_int [rule_format]: "1 <= m \ m <= n \ m dvd fact (n::int)" - apply (case_tac "1 <= n") - apply (induct n rule: int_ge_induct) - apply (auto simp add: fact_plus_one_int) - apply (subgoal_tac "m = i + 1") - apply auto -done - -lemma interval_plus_one_nat: "(i::nat) <= j + 1 \ - {i..j+1} = {i..j} Un {j+1}" - by auto - -lemma interval_plus_one_int: "(i::int) <= j + 1 \ {i..j+1} = {i..j} Un {j+1}" - by auto - -lemma fact_altdef_nat: "fact (n::nat) = (PROD i:{1..n}. i)" - apply (induct n rule: induct'_nat) - apply force - apply (subst fact_plus_one_nat) - apply (subst interval_plus_one_nat) - apply auto -done - -lemma fact_altdef_int: "n >= 0 \ fact (n::int) = (PROD i:{1..n}. i)" - apply (induct n rule: int_ge_induct) - apply force - apply (subst fact_plus_one_int, assumption) - apply (subst interval_plus_one_int) - apply auto -done - -subsection {* Infinitely many primes *} - -lemma next_prime_bound: "\(p::nat). prime p \ n < p \ p <= fact n + 1" -proof- - have f1: "fact n + 1 \ 1" using fact_ge_one_nat [of n] by arith - from prime_factor_nat [OF f1] - obtain p where "prime p" and "p dvd fact n + 1" by auto - hence "p \ fact n + 1" - by (intro dvd_imp_le, auto) - {assume "p \ n" - from `prime p` have "p \ 1" - by (cases p, simp_all) - with `p <= n` have "p dvd fact n" - by (intro dvd_fact_nat) - with `p dvd fact n + 1` have "p dvd fact n + 1 - fact n" - by (rule dvd_diff_nat) - hence "p dvd 1" by simp - hence "p <= 1" by auto - moreover from `prime p` have "p > 1" by auto - ultimately have False by auto} - hence "n < p" by arith - with `prime p` and `p <= fact n + 1` show ?thesis by auto -qed - -lemma bigger_prime: "\p. prime p \ p > (n::nat)" -using next_prime_bound by auto - -lemma primes_infinite: "\ (finite {(p::nat). prime p})" -proof - assume "finite {(p::nat). prime p}" - with Max_ge have "(EX b. (ALL x : {(p::nat). prime p}. x <= b))" - by auto - then obtain b where "ALL (x::nat). prime x \ x <= b" - by auto - with bigger_prime [of b] show False by auto -qed + transfer_int_nat_binomial transfer_int_nat_binomial_closure] subsection {* Binomial coefficients *} diff -r 63686057cbe8 -r abda97d2deea src/HOL/Nominal/Examples/Fsub.thy --- a/src/HOL/Nominal/Examples/Fsub.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Nominal/Examples/Fsub.thy Thu Jul 23 21:13:21 2009 +0200 @@ -91,7 +91,7 @@ pairs of type-variables and types (this is sufficient for Part 1A). *} text {* In order to state validity-conditions for typing-contexts, the notion of - a @{text "domain"} of a typing-context is handy. *} + a @{text "dom"} of a typing-context is handy. *} nominal_primrec "tyvrs_of" :: "binding \ tyvrs set" @@ -108,16 +108,16 @@ by auto consts - "ty_domain" :: "env \ tyvrs set" + "ty_dom" :: "env \ tyvrs set" primrec - "ty_domain [] = {}" - "ty_domain (X#\) = (tyvrs_of X)\(ty_domain \)" + "ty_dom [] = {}" + "ty_dom (X#\) = (tyvrs_of X)\(ty_dom \)" consts - "trm_domain" :: "env \ vrs set" + "trm_dom" :: "env \ vrs set" primrec - "trm_domain [] = {}" - "trm_domain (X#\) = (vrs_of X)\(trm_domain \)" + "trm_dom [] = {}" + "trm_dom (X#\) = (vrs_of X)\(trm_dom \)" lemma vrs_of_eqvt[eqvt]: fixes pi ::"tyvrs prm" @@ -128,13 +128,13 @@ and "pi'\(vrs_of x) = vrs_of (pi'\x)" by (nominal_induct x rule: binding.strong_induct) (simp_all add: tyvrs_of.simps eqvts) -lemma domains_eqvt[eqvt]: +lemma doms_eqvt[eqvt]: fixes pi::"tyvrs prm" and pi'::"vrs prm" - shows "pi \(ty_domain \) = ty_domain (pi\\)" - and "pi'\(ty_domain \) = ty_domain (pi'\\)" - and "pi \(trm_domain \) = trm_domain (pi\\)" - and "pi'\(trm_domain \) = trm_domain (pi'\\)" + shows "pi \(ty_dom \) = ty_dom (pi\\)" + and "pi'\(ty_dom \) = ty_dom (pi'\\)" + and "pi \(trm_dom \) = trm_dom (pi\\)" + and "pi'\(trm_dom \) = trm_dom (pi'\\)" by (induct \) (simp_all add: eqvts) lemma finite_vrs: @@ -142,19 +142,19 @@ and "finite (vrs_of x)" by (nominal_induct rule:binding.strong_induct, auto) -lemma finite_domains: - shows "finite (ty_domain \)" - and "finite (trm_domain \)" +lemma finite_doms: + shows "finite (ty_dom \)" + and "finite (trm_dom \)" by (induct \, auto simp add: finite_vrs) -lemma ty_domain_supp: - shows "(supp (ty_domain \)) = (ty_domain \)" - and "(supp (trm_domain \)) = (trm_domain \)" -by (simp only: at_fin_set_supp at_tyvrs_inst at_vrs_inst finite_domains)+ +lemma ty_dom_supp: + shows "(supp (ty_dom \)) = (ty_dom \)" + and "(supp (trm_dom \)) = (trm_dom \)" +by (simp only: at_fin_set_supp at_tyvrs_inst at_vrs_inst finite_doms)+ -lemma ty_domain_inclusion: +lemma ty_dom_inclusion: assumes a: "(TVarB X T)\set \" - shows "X\(ty_domain \)" + shows "X\(ty_dom \)" using a by (induct \, auto) lemma ty_binding_existence: @@ -163,8 +163,8 @@ using assms by (nominal_induct a rule: binding.strong_induct, auto) -lemma ty_domain_existence: - assumes a: "X\(ty_domain \)" +lemma ty_dom_existence: + assumes a: "X\(ty_dom \)" shows "\T.(TVarB X T)\set \" using a apply (induct \, auto) @@ -173,9 +173,9 @@ apply (auto simp add: ty_binding_existence) done -lemma domains_append: - shows "ty_domain (\@\) = ((ty_domain \) \ (ty_domain \))" - and "trm_domain (\@\) = ((trm_domain \) \ (trm_domain \))" +lemma doms_append: + shows "ty_dom (\@\) = ((ty_dom \) \ (ty_dom \))" + and "trm_dom (\@\) = ((trm_dom \) \ (trm_dom \))" by (induct \, auto) lemma ty_vrs_prm_simp: @@ -184,15 +184,15 @@ shows "pi\S = S" by (induct S rule: ty.induct) (auto simp add: calc_atm) -lemma fresh_ty_domain_cons: +lemma fresh_ty_dom_cons: fixes X::"tyvrs" - shows "X\(ty_domain (Y#\)) = (X\(tyvrs_of Y) \ X\(ty_domain \))" + shows "X\(ty_dom (Y#\)) = (X\(tyvrs_of Y) \ X\(ty_dom \))" apply (nominal_induct rule:binding.strong_induct) apply (auto) apply (simp add: fresh_def supp_def eqvts) - apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_domains) + apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms) apply (simp add: fresh_def supp_def eqvts) - apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_domains)+ + apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms)+ done lemma tyvrs_fresh: @@ -206,26 +206,26 @@ apply (fresh_guess)+ done -lemma fresh_domain: +lemma fresh_dom: fixes X::"tyvrs" assumes a: "X\\" - shows "X\(ty_domain \)" + shows "X\(ty_dom \)" using a apply(induct \) apply(simp add: fresh_set_empty) -apply(simp only: fresh_ty_domain_cons) +apply(simp only: fresh_ty_dom_cons) apply(auto simp add: fresh_prod fresh_list_cons tyvrs_fresh) done text {* Not all lists of type @{typ "env"} are well-formed. One condition requires that in @{term "TVarB X S#\"} all free variables of @{term "S"} must be - in the @{term "ty_domain"} of @{term "\"}, that is @{term "S"} must be @{text "closed"} + in the @{term "ty_dom"} of @{term "\"}, that is @{term "S"} must be @{text "closed"} in @{term "\"}. The set of free variables of @{term "S"} is the @{text "support"} of @{term "S"}. *} constdefs "closed_in" :: "ty \ env \ bool" ("_ closed'_in _" [100,100] 100) - "S closed_in \ \ (supp S)\(ty_domain \)" + "S closed_in \ \ (supp S)\(ty_dom \)" lemma closed_in_eqvt[eqvt]: fixes pi::"tyvrs prm" @@ -251,10 +251,10 @@ shows "x \ T" by (simp add: fresh_def supp_def ty_vrs_prm_simp) -lemma ty_domain_vrs_prm_simp: +lemma ty_dom_vrs_prm_simp: fixes pi::"vrs prm" and \::"env" - shows "(ty_domain (pi\\)) = (ty_domain \)" + shows "(ty_dom (pi\\)) = (ty_dom \)" apply(induct \) apply (simp add: eqvts) apply(simp add: tyvrs_vrs_prm_simp) @@ -265,7 +265,7 @@ assumes a: "S closed_in \" shows "(pi\S) closed_in (pi\\)" using a -by (simp add: closed_in_def ty_domain_vrs_prm_simp ty_vrs_prm_simp) +by (simp add: closed_in_def ty_dom_vrs_prm_simp ty_vrs_prm_simp) lemma fresh_vrs_of: fixes x::"vrs" @@ -273,16 +273,16 @@ by (nominal_induct b rule: binding.strong_induct) (simp_all add: fresh_singleton [OF pt_vrs_inst at_vrs_inst] fresh_set_empty ty_vrs_fresh fresh_atm) -lemma fresh_trm_domain: +lemma fresh_trm_dom: fixes x::"vrs" - shows "x\ trm_domain \ = x\\" + shows "x\ trm_dom \ = x\\" by (induct \) (simp_all add: fresh_set_empty fresh_list_cons fresh_fin_union [OF pt_vrs_inst at_vrs_inst fs_vrs_inst] - finite_domains finite_vrs fresh_vrs_of fresh_list_nil) + finite_doms finite_vrs fresh_vrs_of fresh_list_nil) -lemma closed_in_fresh: "(X::tyvrs) \ ty_domain \ \ T closed_in \ \ X \ T" - by (auto simp add: closed_in_def fresh_def ty_domain_supp) +lemma closed_in_fresh: "(X::tyvrs) \ ty_dom \ \ T closed_in \ \ X \ T" + by (auto simp add: closed_in_def fresh_def ty_dom_supp) text {* Now validity of a context is a straightforward inductive definition. *} @@ -290,15 +290,18 @@ valid_rel :: "env \ bool" ("\ _ ok" [100] 100) where valid_nil[simp]: "\ [] ok" -| valid_consT[simp]: "\\ \ ok; X\(ty_domain \); T closed_in \\ \ \ (TVarB X T#\) ok" -| valid_cons [simp]: "\\ \ ok; x\(trm_domain \); T closed_in \\ \ \ (VarB x T#\) ok" +| valid_consT[simp]: "\\ \ ok; X\(ty_dom \); T closed_in \\ \ \ (TVarB X T#\) ok" +| valid_cons [simp]: "\\ \ ok; x\(trm_dom \); T closed_in \\ \ \ (VarB x T#\) ok" equivariance valid_rel declare binding.inject [simp add] declare trm.inject [simp add] -inductive_cases validE[elim]: "\ (TVarB X T#\) ok" "\ (VarB x T#\) ok" "\ (b#\) ok" +inductive_cases validE[elim]: + "\ (TVarB X T#\) ok" + "\ (VarB x T#\) ok" + "\ (b#\) ok" declare binding.inject [simp del] declare trm.inject [simp del] @@ -321,12 +324,12 @@ using a b proof(induct \) case Nil - then show ?case by (auto elim: validE intro: valid_cons simp add: domains_append closed_in_def) + then show ?case by (auto elim: validE intro: valid_cons simp add: doms_append closed_in_def) next case (Cons a \') then show ?case by (nominal_induct a rule:binding.strong_induct) - (auto elim: validE intro!: valid_cons simp add: domains_append closed_in_def) + (auto elim: validE intro!: valid_cons simp add: doms_append closed_in_def) qed text {* Well-formed contexts have a unique type-binding for a type-variable. *} @@ -342,14 +345,14 @@ case (valid_consT \ X' T') moreover { fix \'::"env" - assume a: "X'\(ty_domain \')" + assume a: "X'\(ty_dom \')" have "\(\T.(TVarB X' T)\(set \'))" using a proof (induct \') case (Cons Y \') thus "\ (\T.(TVarB X' T)\set(Y#\'))" - by (simp add: fresh_ty_domain_cons + by (simp add: fresh_ty_dom_cons fresh_fin_union[OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] - finite_vrs finite_domains, + finite_vrs finite_doms, auto simp add: fresh_atm fresh_singleton [OF pt_tyvrs_inst at_tyvrs_inst]) qed (simp) } @@ -367,13 +370,13 @@ case (valid_cons \ x' T') moreover { fix \'::"env" - assume a: "x'\(trm_domain \')" + assume a: "x'\(trm_dom \')" have "\(\T.(VarB x' T)\(set \'))" using a proof (induct \') case (Cons y \') thus "\ (\T.(VarB x' T)\set(y#\'))" by (simp add: fresh_fin_union[OF pt_vrs_inst at_vrs_inst fs_vrs_inst] - finite_vrs finite_domains, + finite_vrs finite_doms, auto simp add: fresh_atm fresh_singleton [OF pt_vrs_inst at_vrs_inst]) qed (simp) } @@ -401,7 +404,7 @@ "(Tvar X)[Y \ T]\<^sub>\ = (if X=Y then T else Tvar X)" | "(Top)[Y \ T]\<^sub>\ = Top" | "(T\<^isub>1 \ T\<^isub>2)[Y \ T]\<^sub>\ = T\<^isub>1[Y \ T]\<^sub>\ \ T\<^isub>2[Y \ T]\<^sub>\" -| "\X\(Y,T); X\T\<^isub>1\ \ (\X<:T\<^isub>1. T\<^isub>2)[Y \ T]\<^sub>\ = (\X<:T\<^isub>1[Y \ T]\<^sub>\. T\<^isub>2[Y \ T]\<^sub>\)" +| "X\(Y,T,T\<^isub>1) \ (\X<:T\<^isub>1. T\<^isub>2)[Y \ T]\<^sub>\ = (\X<:T\<^isub>1[Y \ T]\<^sub>\. T\<^isub>2[Y \ T]\<^sub>\)" apply (finite_guess)+ apply (rule TrueI)+ apply (simp add: abs_fresh) @@ -424,8 +427,8 @@ lemma type_subst_fresh: fixes X::"tyvrs" - assumes "X \ T" and "X \ P" - shows "X \ T[Y \ P]\<^sub>\" + assumes "X\T" and "X\P" + shows "X\T[Y \ P]\<^sub>\" using assms by (nominal_induct T avoiding: X Y P rule:ty.strong_induct) (auto simp add: abs_fresh) @@ -437,17 +440,18 @@ by (nominal_induct T avoiding: X T' rule: ty.strong_induct) (auto simp add: fresh_atm abs_fresh fresh_nat) -lemma type_subst_identity: "X \ T \ T[X \ U]\<^sub>\ = T" +lemma type_subst_identity: + "X\T \ T[X \ U]\<^sub>\ = T" by (nominal_induct T avoiding: X U rule: ty.strong_induct) (simp_all add: fresh_atm abs_fresh) lemma type_substitution_lemma: - "X \ Y \ X \ L \ M[X \ N]\<^sub>\[Y \ L]\<^sub>\ = M[Y \ L]\<^sub>\[X \ N[Y \ L]\<^sub>\]\<^sub>\" + "X \ Y \ X\L \ M[X \ N]\<^sub>\[Y \ L]\<^sub>\ = M[Y \ L]\<^sub>\[X \ N[Y \ L]\<^sub>\]\<^sub>\" by (nominal_induct M avoiding: X Y N L rule: ty.strong_induct) (auto simp add: type_subst_fresh type_subst_identity) lemma type_subst_rename: - "Y \ T \ ([(Y, X)] \ T)[Y \ U]\<^sub>\ = T[X \ U]\<^sub>\" + "Y\T \ ([(Y,X)]\T)[Y \ U]\<^sub>\ = T[X \ U]\<^sub>\" by (nominal_induct T avoiding: X Y U rule: ty.strong_induct) (simp_all add: fresh_atm calc_atm abs_fresh fresh_aux) @@ -460,15 +464,15 @@ lemma binding_subst_fresh: fixes X::"tyvrs" - assumes "X \ a" - and "X \ P" - shows "X \ a[Y \ P]\<^sub>b" + 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: type_subst_fresh) lemma binding_subst_identity: - shows "X \ B \ B[X \ U]\<^sub>b = B" + shows "X\B \ B[X \ U]\<^sub>b = B" by (induct B rule: binding.induct) (simp_all add: fresh_atm type_subst_identity) @@ -481,9 +485,9 @@ lemma ctxt_subst_fresh': fixes X::"tyvrs" - assumes "X \ \" - and "X \ P" - shows "X \ \[Y \ P]\<^sub>e" + assumes "X\\" + and "X\P" + shows "X\\[Y \ P]\<^sub>e" using assms by (induct \) (auto simp add: fresh_list_cons binding_subst_fresh) @@ -494,7 +498,7 @@ lemma ctxt_subst_mem_VarB: "VarB x T \ set \ \ VarB x (T[Y \ U]\<^sub>\) \ set (\[Y \ U]\<^sub>e)" by (induct \) auto -lemma ctxt_subst_identity: "X \ \ \ \[X \ U]\<^sub>e = \" +lemma ctxt_subst_identity: "X\\ \ \[X \ U]\<^sub>e = \" by (induct \) (simp_all add: fresh_list_cons binding_subst_identity) lemma ctxt_subst_append: "(\ @ \)[X \ T]\<^sub>e = \[X \ T]\<^sub>e @ \[X \ T]\<^sub>e" @@ -515,11 +519,13 @@ done lemma subst_trm_fresh_tyvar: - "(X::tyvrs) \ t \ X \ u \ X \ t[x \ u]" + fixes X::"tyvrs" + shows "X\t \ X\u \ X\t[x \ u]" by (nominal_induct t avoiding: x u rule: trm.strong_induct) (auto simp add: trm.fresh abs_fresh) -lemma subst_trm_fresh_var: "x \ u \ x \ t[x \ u]" +lemma subst_trm_fresh_var: + "x\u \ x\t[x \ u]" by (nominal_induct t avoiding: x u rule: trm.strong_induct) (simp_all add: abs_fresh fresh_atm ty_vrs_fresh) @@ -538,7 +544,7 @@ (perm_simp add: fresh_left)+ lemma subst_trm_rename: - "y \ t \ ([(y, x)] \ t)[y \ u] = t[x \ u]" + "y\t \ ([(y, x)] \ t)[y \ u] = t[x \ u]" by (nominal_induct t avoiding: x y u rule: trm.strong_induct) (simp_all add: fresh_atm calc_atm abs_fresh fresh_aux ty_vrs_fresh perm_fresh_fresh) @@ -558,12 +564,13 @@ done lemma subst_trm_ty_fresh: - "(X::tyvrs) \ t \ X \ T \ X \ t[Y \\<^sub>\ T]" + fixes X::"tyvrs" + shows "X\t \ X\T \ X\t[Y \\<^sub>\ T]" by (nominal_induct t avoiding: Y T rule: trm.strong_induct) (auto simp add: abs_fresh type_subst_fresh) lemma subst_trm_ty_fresh': - "X \ T \ X \ t[X \\<^sub>\ T]" + "X\T \ X\t[X \\<^sub>\ T]" by (nominal_induct t avoiding: X T rule: trm.strong_induct) (simp_all add: abs_fresh fresh_type_subst_fresh fresh_atm) @@ -582,7 +589,7 @@ (perm_simp add: fresh_left subst_eqvt')+ lemma subst_trm_ty_rename: - "Y \ t \ ([(Y, X)] \ t)[Y \\<^sub>\ U] = t[X \\<^sub>\ U]" + "Y\t \ ([(Y, X)] \ t)[Y \\<^sub>\ U] = t[X \\<^sub>\ U]" by (nominal_induct t avoiding: X Y U rule: trm.strong_induct) (simp_all add: fresh_atm calc_atm abs_fresh fresh_aux type_subst_rename) @@ -599,7 +606,7 @@ subtype_of :: "env \ ty \ ty \ bool" ("_\_<:_" [100,100,100] 100) where SA_Top[intro]: "\\ \ ok; S closed_in \\ \ \ \ S <: Top" -| SA_refl_TVar[intro]: "\\ \ ok; X \ ty_domain \\\ \ \ Tvar X <: Tvar X" +| SA_refl_TVar[intro]: "\\ \ ok; X \ ty_dom \\\ \ \ Tvar X <: Tvar X" | SA_trans_TVar[intro]: "\(TVarB X S) \ set \; \ \ S <: T\ \ \ \ (Tvar X) <: T" | SA_arrow[intro]: "\\ \ T\<^isub>1 <: S\<^isub>1; \ \ S\<^isub>2 <: T\<^isub>2\ \ \ \ (S\<^isub>1 \ S\<^isub>2) <: (T\<^isub>1 \ T\<^isub>2)" | SA_all[intro]: "\\ \ T\<^isub>1 <: S\<^isub>1; ((TVarB X T\<^isub>1)#\) \ S\<^isub>2 <: T\<^isub>2\ \ \ \ (\X<:S\<^isub>1. S\<^isub>2) <: (\X<:T\<^isub>1. T\<^isub>2)" @@ -623,7 +630,7 @@ next case (SA_trans_TVar X S \ T) have "(TVarB X S)\set \" by fact - hence "X \ ty_domain \" by (rule ty_domain_inclusion) + hence "X \ ty_dom \" by (rule ty_dom_inclusion) hence "(Tvar X) closed_in \" by (simp add: closed_in_def ty.supp supp_atm) moreover have "S closed_in \ \ T closed_in \" by fact @@ -638,23 +645,23 @@ shows "X\S \ X\T" proof - from a1 have "\ \ ok" by (rule subtype_implies_ok) - with a2 have "X\ty_domain(\)" by (simp add: fresh_domain) + with a2 have "X\ty_dom(\)" by (simp add: fresh_dom) moreover from a1 have "S closed_in \ \ T closed_in \" by (rule subtype_implies_closed) - hence "supp S \ ((supp (ty_domain \))::tyvrs set)" - and "supp T \ ((supp (ty_domain \))::tyvrs set)" by (simp_all add: ty_domain_supp closed_in_def) + hence "supp S \ ((supp (ty_dom \))::tyvrs set)" + and "supp T \ ((supp (ty_dom \))::tyvrs set)" by (simp_all add: ty_dom_supp closed_in_def) ultimately show "X\S \ X\T" by (force simp add: supp_prod fresh_def) qed -lemma valid_ty_domain_fresh: +lemma valid_ty_dom_fresh: fixes X::"tyvrs" assumes valid: "\ \ ok" - shows "X\(ty_domain \) = X\\" + shows "X\(ty_dom \) = X\\" using valid apply induct apply (simp add: fresh_list_nil fresh_set_empty) apply (simp_all add: binding.fresh fresh_list_cons - fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_domains fresh_atm) + fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms fresh_atm) apply (auto simp add: closed_in_fresh) done @@ -662,7 +669,7 @@ nominal_inductive subtype_of apply (simp_all add: abs_fresh) - apply (fastsimp simp add: valid_ty_domain_fresh dest: subtype_implies_ok) + apply (fastsimp simp add: valid_ty_dom_fresh dest: subtype_implies_ok) apply (force simp add: closed_in_fresh dest: subtype_implies_closed subtype_implies_ok)+ done @@ -678,12 +685,12 @@ have ih_T\<^isub>1: "\\. \\ \ ok; T\<^isub>1 closed_in \\ \ \ \ T\<^isub>1 <: T\<^isub>1" by fact have ih_T\<^isub>2: "\\. \\ \ ok; T\<^isub>2 closed_in \\ \ \ \ T\<^isub>2 <: T\<^isub>2" by fact have fresh_cond: "X\\" by fact - hence fresh_ty_domain: "X\(ty_domain \)" by (simp add: fresh_domain) + hence fresh_ty_dom: "X\(ty_dom \)" by (simp add: fresh_dom) have "(\X<:T\<^isub>2. T\<^isub>1) closed_in \" by fact hence closed\<^isub>T\<^isub>2: "T\<^isub>2 closed_in \" and closed\<^isub>T\<^isub>1: "T\<^isub>1 closed_in ((TVarB X T\<^isub>2)#\)" by (auto simp add: closed_in_def ty.supp abs_supp) have ok: "\ \ ok" by fact - hence ok': "\ ((TVarB X T\<^isub>2)#\) ok" using closed\<^isub>T\<^isub>2 fresh_ty_domain by simp + hence ok': "\ ((TVarB X T\<^isub>2)#\) ok" using closed\<^isub>T\<^isub>2 fresh_ty_dom by simp have "\ \ T\<^isub>2 <: T\<^isub>2" using ih_T\<^isub>2 closed\<^isub>T\<^isub>2 ok by simp moreover have "((TVarB X T\<^isub>2)#\) \ T\<^isub>1 <: T\<^isub>1" using ih_T\<^isub>1 closed\<^isub>T\<^isub>1 ok' by simp @@ -702,7 +709,7 @@ \isakeyword{auto}; \isakeyword{blast} would find it if we had not used an explicit definition for @{text "closed_in_def"}. *} apply(drule_tac x="(TVarB tyvrs ty2)#\" in meta_spec) -apply(force dest: fresh_domain simp add: closed_in_def) +apply(force dest: fresh_dom simp add: closed_in_def) done section {* Weakening *} @@ -715,13 +722,13 @@ extends :: "env \ env \ bool" ("_ extends _" [100,100] 100) "\ extends \ \ \X Q. (TVarB X Q)\set \ \ (TVarB X Q)\set \" -lemma extends_ty_domain: +lemma extends_ty_dom: assumes a: "\ extends \" - shows "ty_domain \ \ ty_domain \" + shows "ty_dom \ \ ty_dom \" using a apply (auto simp add: extends_def) - apply (drule ty_domain_existence) - apply (force simp add: ty_domain_inclusion) + apply (drule ty_dom_existence) + apply (force simp add: ty_dom_inclusion) done lemma extends_closed: @@ -729,7 +736,7 @@ and a2: "\ extends \" shows "T closed_in \" using a1 a2 - by (auto dest: extends_ty_domain simp add: closed_in_def) + by (auto dest: extends_ty_dom simp add: closed_in_def) lemma extends_memb: assumes a: "\ extends \" @@ -763,18 +770,18 @@ ultimately show "\ \ Tvar X <: T" using ok by force next case (SA_refl_TVar \ X) - have lh_drv_prem: "X \ ty_domain \" by fact + have lh_drv_prem: "X \ ty_dom \" by fact have "\ \ ok" by fact moreover have "\ extends \" by fact - hence "X \ ty_domain \" using lh_drv_prem by (force dest: extends_ty_domain) + hence "X \ ty_dom \" using lh_drv_prem by (force dest: extends_ty_dom) ultimately show "\ \ Tvar X <: Tvar X" by force next case (SA_arrow \ T\<^isub>1 S\<^isub>1 S\<^isub>2 T\<^isub>2) thus "\ \ S\<^isub>1 \ S\<^isub>2 <: T\<^isub>1 \ T\<^isub>2" by blast next case (SA_all \ T\<^isub>1 S\<^isub>1 X S\<^isub>2 T\<^isub>2) have fresh_cond: "X\\" by fact - hence fresh_domain: "X\(ty_domain \)" by (simp add: fresh_domain) + hence fresh_dom: "X\(ty_dom \)" by (simp add: fresh_dom) have ih\<^isub>1: "\\. \ \ ok \ \ extends \ \ \ \ T\<^isub>1 <: S\<^isub>1" by fact have ih\<^isub>2: "\\. \ \ ok \ \ extends ((TVarB X T\<^isub>1)#\) \ \ \ S\<^isub>2 <: T\<^isub>2" by fact have lh_drv_prem: "\ \ T\<^isub>1 <: S\<^isub>1" by fact @@ -782,7 +789,7 @@ have ok: "\ \ ok" by fact have ext: "\ extends \" by fact have "T\<^isub>1 closed_in \" using ext closed\<^isub>T\<^isub>1 by (simp only: extends_closed) - hence "\ ((TVarB X T\<^isub>1)#\) ok" using fresh_domain ok by force + hence "\ ((TVarB X T\<^isub>1)#\) ok" using fresh_dom ok by force moreover have "((TVarB X T\<^isub>1)#\) extends ((TVarB X T\<^isub>1)#\)" using ext by (force simp add: extends_def) ultimately have "((TVarB X T\<^isub>1)#\) \ S\<^isub>2 <: T\<^isub>2" using ih\<^isub>2 by simp @@ -802,7 +809,7 @@ proof (nominal_induct \ S T avoiding: \ rule: subtype_of.strong_induct) case (SA_all \ T\<^isub>1 S\<^isub>1 X S\<^isub>2 T\<^isub>2) have fresh_cond: "X\\" by fact - hence fresh_domain: "X\(ty_domain \)" by (simp add: fresh_domain) + hence fresh_dom: "X\(ty_dom \)" by (simp add: fresh_dom) have ih\<^isub>1: "\\. \ \ ok \ \ extends \ \ \ \ T\<^isub>1 <: S\<^isub>1" by fact have ih\<^isub>2: "\\. \ \ ok \ \ extends ((TVarB X T\<^isub>1)#\) \ \ \ S\<^isub>2 <: T\<^isub>2" by fact have lh_drv_prem: "\ \ T\<^isub>1 <: S\<^isub>1" by fact @@ -810,14 +817,14 @@ have ok: "\ \ ok" by fact have ext: "\ extends \" by fact have "T\<^isub>1 closed_in \" using ext closed\<^isub>T\<^isub>1 by (simp only: extends_closed) - hence "\ ((TVarB X T\<^isub>1)#\) ok" using fresh_domain ok by force + hence "\ ((TVarB X T\<^isub>1)#\) ok" using fresh_dom ok by force moreover have "((TVarB X T\<^isub>1)#\) extends ((TVarB X T\<^isub>1)#\)" using ext by (force simp add: extends_def) ultimately have "((TVarB X T\<^isub>1)#\) \ S\<^isub>2 <: T\<^isub>2" using ih\<^isub>2 by simp moreover have "\ \ T\<^isub>1 <: S\<^isub>1" using ok ext ih\<^isub>1 by simp ultimately show "\ \ (\X<:S\<^isub>1. S\<^isub>2) <: (\X<:T\<^isub>1. T\<^isub>2)" using ok by (force intro: SA_all) -qed (blast intro: extends_closed extends_memb dest: extends_ty_domain)+ +qed (blast intro: extends_closed extends_memb dest: extends_ty_dom)+ section {* Transitivity and Narrowing *} @@ -831,38 +838,11 @@ declare ty.inject [simp del] lemma S_ForallE_left: - shows "\\ \ (\X<:S\<^isub>1. S\<^isub>2) <: T; X\\; X\S\<^isub>1\ + shows "\\ \ (\X<:S\<^isub>1. S\<^isub>2) <: T; X\\; X\S\<^isub>1; X\T\ \ T = Top \ (\T\<^isub>1 T\<^isub>2. T = (\X<:T\<^isub>1. T\<^isub>2) \ \ \ T\<^isub>1 <: S\<^isub>1 \ ((TVarB X T\<^isub>1)#\) \ S\<^isub>2 <: T\<^isub>2)" - apply(frule subtype_implies_ok) - apply(ind_cases "\ \ (\X<:S\<^isub>1. S\<^isub>2) <: T") - apply(auto simp add: ty.inject alpha) - apply(rule_tac x="[(X,Xa)]\T\<^isub>2" in exI) - apply(rule conjI) - apply(rule sym) - apply(rule pt_bij2[OF pt_tyvrs_inst, OF at_tyvrs_inst]) - apply(rule pt_tyvrs3) - apply(simp) - apply(rule at_ds5[OF at_tyvrs_inst]) - apply(rule conjI) - apply(simp add: pt_fresh_left[OF pt_tyvrs_inst, OF at_tyvrs_inst] calc_atm) - apply(drule_tac \="((TVarB Xa T\<^isub>1)#\)" in subtype_implies_closed)+ - apply(simp add: closed_in_def) - apply(drule fresh_domain)+ - apply(simp add: fresh_def) - apply(subgoal_tac "X \ (insert Xa (ty_domain \))")(*A*) - apply(force) - (*A*)apply(simp add: at_fin_set_supp[OF at_tyvrs_inst, OF finite_domains(1)]) - (* 2nd conjunct *)apply(frule_tac X="X" in subtype_implies_fresh) - apply(assumption) - apply (frule_tac \="TVarB Xa T\<^isub>1 # \" in subtype_implies_ok) - apply (erule validE) - apply (simp add: valid_ty_domain_fresh) - apply(drule_tac X="Xa" in subtype_implies_fresh) - apply(assumption) - apply(drule_tac pi="[(X,Xa)]" in subtype_of.eqvt(2)) - apply(simp add: calc_atm) - apply(simp add: pt_fresh_fresh[OF pt_tyvrs_inst, OF at_tyvrs_inst]) - done +apply(erule subtype_of.strong_cases[where X="X"]) +apply(auto simp add: abs_fresh ty.inject alpha) +done text {* Next we prove the transitivity and narrowing for the subtyping-relation. The POPLmark-paper says the following: @@ -898,75 +878,38 @@ and subtype_narrow: "(\@[(TVarB X Q)]@\)\M<:N \ \\P<:Q \ (\@[(TVarB X P)]@\)\M<:N" proof (induct Q arbitrary: \ S T \ X P M N taking: "size_ty" rule: measure_induct_rule) case (less Q) - --{* \begin{minipage}[t]{0.9\textwidth} - First we mention the induction hypotheses of the outer induction for later - reference:\end{minipage}*} have IH_trans: "\Q' \ S T. \size_ty Q' < size_ty Q; \\S<:Q'; \\Q'<:T\ \ \\S<:T" by fact have IH_narrow: "\Q' \ \ X M N P. \size_ty Q' < size_ty Q; (\@[(TVarB X Q')]@\)\M<:N; \\P<:Q'\ \ (\@[(TVarB X P)]@\)\M<:N" by fact - --{* \begin{minipage}[t]{0.9\textwidth} - We proceed with the transitivity proof as an auxiliary lemma, because it needs - to be referenced in the narrowing proof.\end{minipage}*} - have transitivity_aux: - "\\ S T. \\ \ S <: Q; \ \ Q <: T\ \ \ \ S <: T" - proof - - fix \' S' T - assume "\' \ S' <: Q" --{* left-hand derivation *} - and "\' \ Q <: T" --{* right-hand derivation *} - thus "\' \ S' <: T" - proof (nominal_induct \' S' Q\Q rule: subtype_of.strong_induct) + + { fix \ S T + have "\\ \ S <: Q; \ \ Q <: T\ \ \ \ S <: T" + proof (induct \ S Q\Q rule: subtype_of.induct) case (SA_Top \ S) - --{* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "\ \ S <: Top"}, giving - us @{term "\ \ ok"} and @{term "S closed_in \"}. This case is straightforward, - because the right-hand derivation must be of the form @{term "\ \ Top <: Top"} - giving us the equation @{term "T = Top"}.\end{minipage}*} - hence rh_drv: "\ \ Top <: T" by simp - hence T_inst: "T = Top" by (auto elim: S_TopE) + then have rh_drv: "\ \ Top <: T" by simp + then have T_inst: "T = Top" by (auto elim: S_TopE) from `\ \ ok` and `S closed_in \` - have "\ \ S <: Top" by (simp add: subtype_of.SA_Top) - thus "\ \ S <: T" using T_inst by simp + have "\ \ S <: Top" by auto + then show "\ \ S <: T" using T_inst by simp next case (SA_trans_TVar Y U \) - -- {* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "\ \ Tvar Y <: Q"} - with @{term "S = Tvar Y"}. We have therefore @{term "(Y,U)"} - is in @{term "\"} and by inner induction hypothesis that @{term "\ \ U <: T"}. - By @{text "S_Var"} follows @{term "\ \ Tvar Y <: T"}.\end{minipage}*} - hence IH_inner: "\ \ U <: T" by simp + then have IH_inner: "\ \ U <: T" by simp have "(TVarB Y U) \ set \" by fact - with IH_inner show "\ \ Tvar Y <: T" by (simp add: subtype_of.SA_trans_TVar) + with IH_inner show "\ \ Tvar Y <: T" by auto next case (SA_refl_TVar \ X) - --{* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "\\(Tvar X) <: (Tvar X)"} with - @{term "Q=Tvar X"}. The goal then follows immediately from the right-hand - derivation.\end{minipage}*} - thus "\ \ Tvar X <: T" by simp + then show "\ \ Tvar X <: T" by simp next case (SA_arrow \ Q\<^isub>1 S\<^isub>1 S\<^isub>2 Q\<^isub>2) - --{* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "\ \ S\<^isub>1 \ S\<^isub>2 <: Q\<^isub>1 \ Q\<^isub>2"} with - @{term "S\<^isub>1\S\<^isub>2=S"} and @{term "Q\<^isub>1\Q\<^isub>2=Q"}. We know that the @{text "size_ty"} of - @{term Q\<^isub>1} and @{term Q\<^isub>2} is smaller than that of @{term Q}; - so we can apply the outer induction hypotheses for @{term Q\<^isub>1} and @{term Q\<^isub>2}. - We also have the sub-derivations @{term "\\Q\<^isub>1<:S\<^isub>1"} and @{term "\\S\<^isub>2<:Q\<^isub>2"}. - The right-hand derivation is @{term "\ \ Q\<^isub>1 \ Q\<^isub>2 <: T"}. There exist types - @{text "T\<^isub>1,T\<^isub>2"} such that @{term "T=Top \ T=T\<^isub>1\T\<^isub>2"}. The @{term "Top"}-case is - straightforward once we know @{term "(S\<^isub>1 \ S\<^isub>2) closed_in \"} and @{term "\ \ ok"}. - In the other case we have the sub-derivations @{term "\\T\<^isub>1<:Q\<^isub>1"} and @{term "\\Q\<^isub>2<:T\<^isub>2"}. - Using the outer induction hypothesis for transitivity we can derive @{term "\\T\<^isub>1<:S\<^isub>1"} - and @{term "\\S\<^isub>2<:T\<^isub>2"}. By rule @{text "S_Arrow"} follows @{term "\ \ S\<^isub>1 \ S\<^isub>2 <: T\<^isub>1 \ T\<^isub>2"}, - which is @{term "\ \ S\<^isub>1 \ S\<^isub>2 <: T\<^isub>"}.\end{minipage}*} - hence rh_drv: "\ \ Q\<^isub>1 \ Q\<^isub>2 <: T" by simp + then have rh_drv: "\ \ Q\<^isub>1 \ Q\<^isub>2 <: T" by simp from `Q\<^isub>1 \ Q\<^isub>2 = Q` have Q\<^isub>1\<^isub>2_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q" by auto have lh_drv_prm\<^isub>1: "\ \ Q\<^isub>1 <: S\<^isub>1" by fact have lh_drv_prm\<^isub>2: "\ \ S\<^isub>2 <: Q\<^isub>2" by fact from rh_drv have "T=Top \ (\T\<^isub>1 T\<^isub>2. T=T\<^isub>1\T\<^isub>2 \ \\T\<^isub>1<:Q\<^isub>1 \ \\Q\<^isub>2<:T\<^isub>2)" - by (auto elim: S_ArrowE_left) + by (auto elim: S_ArrowE_left) moreover have "S\<^isub>1 closed_in \" and "S\<^isub>2 closed_in \" using lh_drv_prm\<^isub>1 lh_drv_prm\<^isub>2 by (simp_all add: subtype_implies_closed) @@ -974,7 +917,7 @@ moreover have "\ \ ok" using rh_drv by (rule subtype_implies_ok) moreover - { assume "\T\<^isub>1 T\<^isub>2. T=T\<^isub>1\T\<^isub>2 \ \\T\<^isub>1<:Q\<^isub>1 \ \\Q\<^isub>2<:T\<^isub>2" + { assume "\T\<^isub>1 T\<^isub>2. T = T\<^isub>1\T\<^isub>2 \ \ \ T\<^isub>1 <: Q\<^isub>1 \ \ \ Q\<^isub>2 <: T\<^isub>2" then obtain T\<^isub>1 T\<^isub>2 where T_inst: "T = T\<^isub>1 \ T\<^isub>2" and rh_drv_prm\<^isub>1: "\ \ T\<^isub>1 <: Q\<^isub>1" @@ -984,46 +927,26 @@ moreover from IH_trans[of "Q\<^isub>2"] have "\ \ S\<^isub>2 <: T\<^isub>2" using Q\<^isub>1\<^isub>2_less rh_drv_prm\<^isub>2 lh_drv_prm\<^isub>2 by simp - ultimately have "\ \ S\<^isub>1 \ S\<^isub>2 <: T\<^isub>1 \ T\<^isub>2" by (simp add: subtype_of.SA_arrow) - hence "\ \ S\<^isub>1 \ S\<^isub>2 <: T" using T_inst by simp + ultimately have "\ \ S\<^isub>1 \ S\<^isub>2 <: T\<^isub>1 \ T\<^isub>2" by auto + then have "\ \ S\<^isub>1 \ S\<^isub>2 <: T" using T_inst by simp } ultimately show "\ \ S\<^isub>1 \ S\<^isub>2 <: T" by blast next case (SA_all \ Q\<^isub>1 S\<^isub>1 X S\<^isub>2 Q\<^isub>2) - --{* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "\\(\X<:S\<^isub>1. S\<^isub>2) <: (\X<:Q\<^isub>1. Q\<^isub>2)"} with - @{term "(\X<:S\<^isub>1. S\<^isub>2)=S"} and @{term "(\X<:Q\<^isub>1. Q\<^isub>2)=Q"}. We therefore have the sub-derivations - @{term "\\Q\<^isub>1<:S\<^isub>1"} and @{term "((TVarB X Q\<^isub>1)#\)\S\<^isub>2<:Q\<^isub>2"}. Since @{term "X"} is a binder, we - assume that it is sufficiently fresh; in particular we have the freshness conditions - @{term "X\\"} and @{term "X\Q\<^isub>1"} (these assumptions are provided by the strong - induction-rule @{text "subtype_of_induct"}). We know that the @{text "size_ty"} of - @{term Q\<^isub>1} and @{term Q\<^isub>2} is smaller than that of @{term Q}; - so we can apply the outer induction hypotheses for @{term Q\<^isub>1} and @{term Q\<^isub>2}. - The right-hand derivation is @{term "\ \ (\X<:Q\<^isub>1. Q\<^isub>2) <: T"}. Since @{term "X\\"} - and @{term "X\Q\<^isub>1"} there exists types @{text "T\<^isub>1,T\<^isub>2"} such that - @{term "T=Top \ T=(\X<:T\<^isub>1. T\<^isub>2)"}. The @{term "Top"}-case is straightforward once we know - @{term "(\X<:S\<^isub>1. S\<^isub>2) closed_in \"} and @{term "\ \ ok"}. In the other case we have - the sub-derivations @{term "\\T\<^isub>1<:Q\<^isub>1"} and @{term "((TVarB X T\<^isub>1)#\)\Q\<^isub>2<:T\<^isub>2"}. Using the outer - induction hypothesis for transitivity we can derive @{term "\\T\<^isub>1<:S\<^isub>1"}. From the outer - induction for narrowing we get @{term "((TVarB X T\<^isub>1)#\) \ S\<^isub>2 <: Q\<^isub>2"} and then using again - induction for transitivity we obtain @{term "((TVarB X T\<^isub>1)#\) \ S\<^isub>2 <: T\<^isub>2"}. By rule - @{text "S_Forall"} and the freshness condition @{term "X\\"} follows - @{term "\ \ (\X<:S\<^isub>1. S\<^isub>2) <: (\X<:T\<^isub>1. T\<^isub>2)"}, which is @{term "\ \ (\X<:S\<^isub>1. S\<^isub>2) <: T\<^isub>"}. - \end{minipage}*} - hence rh_drv: "\ \ (\X<:Q\<^isub>1. Q\<^isub>2) <: T" by simp + then have rh_drv: "\ \ (\X<:Q\<^isub>1. Q\<^isub>2) <: T" by simp have lh_drv_prm\<^isub>1: "\ \ Q\<^isub>1 <: S\<^isub>1" by fact have lh_drv_prm\<^isub>2: "((TVarB X Q\<^isub>1)#\) \ S\<^isub>2 <: Q\<^isub>2" by fact - then have "X\\" by (force dest: subtype_implies_ok simp add: valid_ty_domain_fresh) - then have fresh_cond: "X\\" "X\Q\<^isub>1" using lh_drv_prm\<^isub>1 by (simp_all add: subtype_implies_fresh) - from `(\X<:Q\<^isub>1. Q\<^isub>2) = Q` - have Q\<^isub>1\<^isub>2_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q " using fresh_cond by auto + then have "X\\" by (force dest: subtype_implies_ok simp add: valid_ty_dom_fresh) + then have fresh_cond: "X\\" "X\Q\<^isub>1" "X\T" using rh_drv lh_drv_prm\<^isub>1 + by (simp_all add: subtype_implies_fresh) from rh_drv - have "T=Top \ (\T\<^isub>1 T\<^isub>2. T=(\X<:T\<^isub>1. T\<^isub>2) \ \\T\<^isub>1<:Q\<^isub>1 \ ((TVarB X T\<^isub>1)#\)\Q\<^isub>2<:T\<^isub>2)" + have "T = Top \ + (\T\<^isub>1 T\<^isub>2. T = (\X<:T\<^isub>1. T\<^isub>2) \ \ \ T\<^isub>1 <: Q\<^isub>1 \ ((TVarB X T\<^isub>1)#\) \ Q\<^isub>2 <: T\<^isub>2)" using fresh_cond by (simp add: S_ForallE_left) moreover have "S\<^isub>1 closed_in \" and "S\<^isub>2 closed_in ((TVarB X Q\<^isub>1)#\)" using lh_drv_prm\<^isub>1 lh_drv_prm\<^isub>2 by (simp_all add: subtype_implies_closed) - hence "(\X<:S\<^isub>1. S\<^isub>2) closed_in \" by (force simp add: closed_in_def ty.supp abs_supp) + then have "(\X<:S\<^isub>1. S\<^isub>2) closed_in \" by (force simp add: closed_in_def ty.supp abs_supp) moreover have "\ \ ok" using rh_drv by (rule subtype_implies_ok) moreover @@ -1032,6 +955,9 @@ where T_inst: "T = (\X<:T\<^isub>1. T\<^isub>2)" and rh_drv_prm\<^isub>1: "\ \ T\<^isub>1 <: Q\<^isub>1" and rh_drv_prm\<^isub>2:"((TVarB X T\<^isub>1)#\) \ Q\<^isub>2 <: T\<^isub>2" by force + have "(\X<:Q\<^isub>1. Q\<^isub>2) = Q" by fact + then have Q\<^isub>1\<^isub>2_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q" + using fresh_cond by auto from IH_trans[of "Q\<^isub>1"] have "\ \ T\<^isub>1 <: S\<^isub>1" using lh_drv_prm\<^isub>1 rh_drv_prm\<^isub>1 Q\<^isub>1\<^isub>2_less by blast moreover @@ -1045,52 +971,35 @@ } ultimately show "\ \ (\X<:S\<^isub>1. S\<^isub>2) <: T" by blast qed - qed + } note transitivity_lemma = this { --{* The transitivity proof is now by the auxiliary lemma. *} case 1 from `\ \ S <: Q` and `\ \ Q <: T` - show "\ \ S <: T" by (rule transitivity_aux) + show "\ \ S <: T" by (rule transitivity_lemma) next - --{* The narrowing proof proceeds by an induction over @{term "(\@[(TVarB X Q)]@\) \ M <: N"}. *} case 2 - from `(\@[(TVarB X Q)]@\) \ M <: N` --{* left-hand derivation *} - and `\ \ P<:Q` --{* right-hand derivation *} + from `(\@[(TVarB X Q)]@\) \ M <: N` + and `\ \ P<:Q` show "(\@[(TVarB X P)]@\) \ M <: N" - proof (nominal_induct \\"\@[(TVarB X Q)]@\" M N avoiding: \ \ X rule: subtype_of.strong_induct) - case (SA_Top _ S \ \ X) - --{* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "(\@[(TVarB X Q)]@\) \ S <: Top"}. We show - that the context @{term "\@[(TVarB X P)]@\"} is ok and that @{term S} is closed in - @{term "\@[(TVarB X P)]@\"}. Then we can apply the @{text "S_Top"}-rule.\end{minipage}*} - hence lh_drv_prm\<^isub>1: "\ (\@[(TVarB X Q)]@\) ok" + proof (induct \\"\@[(TVarB X Q)]@\" M N arbitrary: \ X \ rule: subtype_of.induct) + case (SA_Top _ S \ X \) + then have lh_drv_prm\<^isub>1: "\ (\@[(TVarB X Q)]@\) ok" and lh_drv_prm\<^isub>2: "S closed_in (\@[(TVarB X Q)]@\)" by simp_all have rh_drv: "\ \ P <: Q" by fact hence "P closed_in \" by (simp add: subtype_implies_closed) with lh_drv_prm\<^isub>1 have "\ (\@[(TVarB X P)]@\) ok" by (simp add: replace_type) moreover from lh_drv_prm\<^isub>2 have "S closed_in (\@[(TVarB X P)]@\)" - by (simp add: closed_in_def domains_append) + by (simp add: closed_in_def doms_append) ultimately show "(\@[(TVarB X P)]@\) \ S <: Top" by (simp add: subtype_of.SA_Top) next - case (SA_trans_TVar Y S _ N \ \ X) - --{* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "(\@[(TVarB X Q)]@\) \ Tvar Y <: N"} and - by inner induction hypothesis we have @{term "(\@[(TVarB X P)]@\) \ S <: N"}. We therefore - know that the contexts @{term "\@[(TVarB X Q)]@\"} and @{term "\@[(TVarB X P)]@\"} are ok, and that - @{term "(Y,S)"} is in @{term "\@[(TVarB X Q)]@\"}. We need to show that - @{term "(\@[(TVarB X P)]@\) \ Tvar Y <: N"} holds. In case @{term "X\Y"} we know that - @{term "(Y,S)"} is in @{term "\@[(TVarB X P)]@\"} and can use the inner induction hypothesis - and rule @{text "S_Var"} to conclude. In case @{term "X=Y"} we can infer that - @{term "S=Q"}; moreover we have that @{term "(\@[(TVarB X P)]@\) extends \"} and therefore - by @{text "weakening"} that @{term "(\@[(TVarB X P)]@\) \ P <: Q"} holds. By transitivity we - obtain then @{term "(\@[(TVarB X P)]@\) \ P <: N"} and can conclude by applying rule - @{text "S_Var"}.\end{minipage}*} - hence IH_inner: "(\@[(TVarB X P)]@\) \ S <: N" + case (SA_trans_TVar Y S _ N \ X \) + then have IH_inner: "(\@[(TVarB X P)]@\) \ S <: N" and lh_drv_prm: "(TVarB Y S) \ set (\@[(TVarB X Q)]@\)" and rh_drv: "\ \ P<:Q" and ok\<^isub>Q: "\ (\@[(TVarB X Q)]@\) ok" by (simp_all add: subtype_implies_ok) - hence ok\<^isub>P: "\ (\@[(TVarB X P)]@\) ok" by (simp add: subtype_implies_ok) + then have ok\<^isub>P: "\ (\@[(TVarB X P)]@\) ok" by (simp add: subtype_implies_ok) show "(\@[(TVarB X P)]@\) \ Tvar Y <: N" proof (cases "X=Y") case False @@ -1107,48 +1016,28 @@ moreover have "(\@[(TVarB X P)]@\) extends \" by (simp add: extends_def) hence "(\@[(TVarB X P)]@\) \ P <: Q" using rh_drv ok\<^isub>P by (simp only: weakening) - ultimately have "(\@[(TVarB X P)]@\) \ P <: N" by (simp add: transitivity_aux) - thus "(\@[(TVarB X P)]@\) \ Tvar Y <: N" using memb\<^isub>X\<^isub>P eq by (simp only: subtype_of.SA_trans_TVar) + ultimately have "(\@[(TVarB X P)]@\) \ P <: N" by (simp add: transitivity_lemma) + then show "(\@[(TVarB X P)]@\) \ Tvar Y <: N" using memb\<^isub>X\<^isub>P eq by auto qed next - case (SA_refl_TVar _ Y \ \ X) - --{* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "(\@[(TVarB X Q)]@\) \ Tvar Y <: Tvar Y"} and we - therefore know that @{term "\@[(TVarB X Q)]@\"} is ok and that @{term "Y"} is in - the domain of @{term "\@[(TVarB X Q)]@\"}. We therefore know that @{term "\@[(TVarB X P)]@\"} is ok - and that @{term Y} is in the domain of @{term "\@[(TVarB X P)]@\"}. We can conclude by applying - rule @{text "S_Refl"}.\end{minipage}*} - hence lh_drv_prm\<^isub>1: "\ (\@[(TVarB X Q)]@\) ok" - and lh_drv_prm\<^isub>2: "Y \ ty_domain (\@[(TVarB X Q)]@\)" by simp_all + case (SA_refl_TVar _ Y \ X \) + then have lh_drv_prm\<^isub>1: "\ (\@[(TVarB X Q)]@\) ok" + and lh_drv_prm\<^isub>2: "Y \ ty_dom (\@[(TVarB X Q)]@\)" by simp_all have "\ \ P <: Q" by fact hence "P closed_in \" by (simp add: subtype_implies_closed) with lh_drv_prm\<^isub>1 have "\ (\@[(TVarB X P)]@\) ok" by (simp add: replace_type) moreover - from lh_drv_prm\<^isub>2 have "Y \ ty_domain (\@[(TVarB X P)]@\)" by (simp add: domains_append) + from lh_drv_prm\<^isub>2 have "Y \ ty_dom (\@[(TVarB X P)]@\)" by (simp add: doms_append) ultimately show "(\@[(TVarB X P)]@\) \ Tvar Y <: Tvar Y" by (simp add: subtype_of.SA_refl_TVar) next - case (SA_arrow _ S\<^isub>1 Q\<^isub>1 Q\<^isub>2 S\<^isub>2 \ \ X) - --{* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "(\@[(TVarB X Q)]@\) \ Q\<^isub>1 \ Q\<^isub>2 <: S\<^isub>1 \ S\<^isub>2"} - and the proof is trivial.\end{minipage}*} - thus "(\@[(TVarB X P)]@\) \ Q\<^isub>1 \ Q\<^isub>2 <: S\<^isub>1 \ S\<^isub>2" by blast + case (SA_arrow _ S\<^isub>1 Q\<^isub>1 Q\<^isub>2 S\<^isub>2 \ X \) + then show "(\@[(TVarB X P)]@\) \ Q\<^isub>1 \ Q\<^isub>2 <: S\<^isub>1 \ S\<^isub>2" by blast next - case (SA_all \' T\<^isub>1 S\<^isub>1 Y S\<^isub>2 T\<^isub>2 \ \ X) - --{* \begin{minipage}[t]{0.9\textwidth} - In this case the left-hand derivation is @{term "(\@[(TVarB X Q)]@\) \ (\Y<:S\<^isub>1. S\<^isub>2) <: (\Y<:T\<^isub>1. T\<^isub>2)"} - and therfore we know that the binder @{term Y} is fresh for @{term "\@[(TVarB X Q)]@\"}. By - the inner induction hypothesis we have that @{term "(\@[(TVarB X P)]@\) \ T\<^isub>1 <: S\<^isub>1"} and - @{term "((TVarB Y T\<^isub>1)#\@[(TVarB X P)]@\) \ S\<^isub>2 <: T\<^isub>2"}. Since @{term P} is a subtype of @{term Q} - we can infer that @{term Y} is fresh for @{term P} and thus also fresh for - @{term "\@[(TVarB X P)]@\"}. We can then conclude by applying rule @{text "S_Forall"}. - \end{minipage}*} - hence rh_drv: "\ \ P <: Q" - and IH_inner\<^isub>1: "(\@[(TVarB X P)]@\) \ T\<^isub>1 <: S\<^isub>1" - and "TVarB Y T\<^isub>1 # \' = ((TVarB Y T\<^isub>1)#\) @ [TVarB X Q] @ \" by auto - moreover have " \\\P<:Q; TVarB Y T\<^isub>1 # \' = ((TVarB Y T\<^isub>1)#\) @ [TVarB X Q] @ \\ \ (((TVarB Y T\<^isub>1)#\) @ [TVarB X P] @ \)\S\<^isub>2<:T\<^isub>2" by fact - ultimately have IH_inner\<^isub>2: "(((TVarB Y T\<^isub>1)#\) @ [TVarB X P] @ \)\S\<^isub>2<:T\<^isub>2" by auto - with IH_inner\<^isub>1 IH_inner\<^isub>2 - show "(\@[(TVarB X P)]@\) \ (\Y<:S\<^isub>1. S\<^isub>2) <: (\Y<:T\<^isub>1. T\<^isub>2)" by (simp add: subtype_of.SA_all) + case (SA_all _ T\<^isub>1 S\<^isub>1 Y S\<^isub>2 T\<^isub>2 \ X \) + from SA_all(2,4,5,6) + have IH_inner\<^isub>1: "(\@[(TVarB X P)]@\) \ T\<^isub>1 <: S\<^isub>1" + and IH_inner\<^isub>2: "(((TVarB Y T\<^isub>1)#\)@[(TVarB X P)]@\) \ S\<^isub>2 <: T\<^isub>2" by force+ + then show "(\@[(TVarB X P)]@\) \ (\Y<:S\<^isub>1. S\<^isub>2) <: (\Y<:T\<^isub>1. T\<^isub>2)" by auto qed } qed @@ -1163,7 +1052,7 @@ | T_Abs[intro]: "\ VarB x T\<^isub>1 # \ \ t\<^isub>2 : T\<^isub>2 \ \ \ \ (\x:T\<^isub>1. t\<^isub>2) : T\<^isub>1 \ T\<^isub>2" | T_Sub[intro]: "\ \ \ t : S; \ \ S <: T \ \ \ \ t : T" | T_TAbs[intro]:"\ TVarB X T\<^isub>1 # \ \ t\<^isub>2 : T\<^isub>2 \ \ \ \ (\X<:T\<^isub>1. t\<^isub>2) : (\X<:T\<^isub>1. T\<^isub>2)" -| T_TApp[intro]:"\ X \ (\, t\<^isub>1, T\<^isub>2); \ \ t\<^isub>1 : (\X<:T\<^isub>1\<^isub>1. T\<^isub>1\<^isub>2); \ \ T\<^isub>2 <: T\<^isub>1\<^isub>1 \ \ \ \ t\<^isub>1 \\<^sub>\ T\<^isub>2 : (T\<^isub>1\<^isub>2[X \ T\<^isub>2]\<^sub>\)" +| T_TApp[intro]:"\X\(\,t\<^isub>1,T\<^isub>2); \ \ t\<^isub>1 : (\X<:T\<^isub>1\<^isub>1. T\<^isub>1\<^isub>2); \ \ T\<^isub>2 <: T\<^isub>1\<^isub>1\ \ \ \ t\<^isub>1 \\<^sub>\ T\<^isub>2 : (T\<^isub>1\<^isub>2[X \ T\<^isub>2]\<^sub>\)" equivariance typing @@ -1189,8 +1078,8 @@ using assms by (induct, auto) nominal_inductive typing -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) +by (auto dest!: typing_ok intro: closed_in_fresh fresh_dom type_subst_fresh + simp: abs_fresh fresh_type_subst_fresh ty_vrs_fresh valid_ty_dom_fresh fresh_trm_dom) lemma ok_imp_VarB_closed_in: assumes ok: "\ \ ok" @@ -1200,19 +1089,19 @@ lemma tyvrs_of_subst: "tyvrs_of (B[X \ T]\<^sub>b) = tyvrs_of B" by (nominal_induct B rule: binding.strong_induct) simp_all -lemma ty_domain_subst: "ty_domain (\[X \ T]\<^sub>e) = ty_domain \" +lemma ty_dom_subst: "ty_dom (\[X \ T]\<^sub>e) = ty_dom \" by (induct \) (simp_all add: tyvrs_of_subst) lemma vrs_of_subst: "vrs_of (B[X \ T]\<^sub>b) = vrs_of B" by (nominal_induct B rule: binding.strong_induct) simp_all -lemma trm_domain_subst: "trm_domain (\[X \ T]\<^sub>e) = trm_domain \" +lemma trm_dom_subst: "trm_dom (\[X \ T]\<^sub>e) = trm_dom \" by (induct \) (simp_all add: vrs_of_subst) lemma subst_closed_in: "T closed_in (\ @ TVarB X S # \) \ U closed_in \ \ T[X \ U]\<^sub>\ closed_in (\[X \ U]\<^sub>e @ \)" apply (nominal_induct T avoiding: X U \ rule: ty.strong_induct) - apply (simp add: closed_in_def ty.supp supp_atm domains_append ty_domain_subst) + apply (simp add: closed_in_def ty.supp supp_atm doms_append ty_dom_subst) apply blast apply (simp add: closed_in_def ty.supp) apply (simp add: closed_in_def ty.supp) @@ -1220,7 +1109,7 @@ apply (drule_tac x = X in meta_spec) apply (drule_tac x = U in meta_spec) apply (drule_tac x = "(TVarB tyvrs ty2) # \" in meta_spec) - apply (simp add: domains_append ty_domain_subst) + apply (simp add: doms_append ty_dom_subst) apply blast done @@ -1323,20 +1212,20 @@ "(\x:T. t) \ t'" "(\X<:T. t) \ t'" -lemma ty_domain_cons: - shows "ty_domain (\@[VarB X Q]@\) = ty_domain (\@\)" +lemma ty_dom_cons: + shows "ty_dom (\@[VarB X Q]@\) = ty_dom (\@\)" by (induct \, auto) lemma closed_in_cons: assumes "S closed_in (\ @ VarB X Q # \)" shows "S closed_in (\@\)" -using assms ty_domain_cons closed_in_def by auto +using assms ty_dom_cons closed_in_def by auto lemma closed_in_weaken: "T closed_in (\ @ \) \ T closed_in (\ @ B # \)" - by (auto simp add: closed_in_def domains_append) + by (auto simp add: closed_in_def doms_append) lemma closed_in_weaken': "T closed_in \ \ T closed_in (\ @ \)" - by (auto simp add: closed_in_def domains_append) + by (auto simp add: closed_in_def doms_append) lemma valid_subst: assumes ok: "\ (\ @ TVarB X Q # \) ok" @@ -1350,24 +1239,24 @@ apply simp apply (rule valid_consT) apply assumption - apply (simp add: domains_append ty_domain_subst) - apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_domains) + apply (simp add: doms_append ty_dom_subst) + apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms) apply (rule_tac S=Q in subst_closed_in') - apply (simp add: closed_in_def domains_append ty_domain_subst) - apply (simp add: closed_in_def domains_append) + apply (simp add: closed_in_def doms_append ty_dom_subst) + apply (simp add: closed_in_def doms_append) apply blast apply simp apply (rule valid_cons) apply assumption - apply (simp add: domains_append trm_domain_subst) + apply (simp add: doms_append trm_dom_subst) apply (rule_tac S=Q in subst_closed_in') - apply (simp add: closed_in_def domains_append ty_domain_subst) - apply (simp add: closed_in_def domains_append) + apply (simp add: closed_in_def doms_append ty_dom_subst) + apply (simp add: closed_in_def doms_append) apply blast done -lemma ty_domain_vrs: - shows "ty_domain (G @ [VarB x Q] @ D) = ty_domain (G @ D)" +lemma ty_dom_vrs: + shows "ty_dom (G @ [VarB x Q] @ D) = ty_dom (G @ D)" by (induct G, auto) lemma valid_cons': @@ -1389,10 +1278,10 @@ case (Cons b bs) with valid_consT have "\ (bs @ \) ok" by simp - moreover from Cons and valid_consT have "X \ ty_domain (bs @ \)" - by (simp add: domains_append) + moreover from Cons and valid_consT have "X \ ty_dom (bs @ \)" + by (simp add: doms_append) moreover from Cons and valid_consT have "T closed_in (bs @ \)" - by (simp add: closed_in_def domains_append) + by (simp add: closed_in_def doms_append) ultimately have "\ (TVarB X T # bs @ \) ok" by (rule valid_rel.valid_consT) with Cons and valid_consT show ?thesis by simp @@ -1407,11 +1296,11 @@ case (Cons b bs) with valid_cons have "\ (bs @ \) ok" by simp - moreover from Cons and valid_cons have "x \ trm_domain (bs @ \)" - by (simp add: domains_append finite_domains + moreover from Cons and valid_cons have "x \ trm_dom (bs @ \)" + by (simp add: doms_append finite_doms fresh_fin_insert [OF pt_vrs_inst at_vrs_inst fs_vrs_inst]) moreover from Cons and valid_cons have "T closed_in (bs @ \)" - by (simp add: closed_in_def domains_append) + by (simp add: closed_in_def doms_append) ultimately have "\ (VarB x T # bs @ \) ok" by (rule valid_rel.valid_cons) with Cons and valid_cons show ?thesis by simp @@ -1439,10 +1328,10 @@ have "\ (VarB y T\<^isub>1 # \ @ B # \) ok" apply (rule valid_cons) apply (rule T_Abs) - apply (simp add: domains_append + apply (simp add: doms_append fresh_fin_insert [OF pt_vrs_inst at_vrs_inst fs_vrs_inst] fresh_fin_union [OF pt_vrs_inst at_vrs_inst fs_vrs_inst] - finite_domains finite_vrs fresh_vrs_of T_Abs fresh_trm_domain) + finite_doms finite_vrs fresh_vrs_of T_Abs fresh_trm_dom) apply (rule closed_in_weaken) apply (rule closed) done @@ -1467,10 +1356,10 @@ have "\ (TVarB X T\<^isub>1 # \ @ B # \) ok" apply (rule valid_consT) apply (rule T_TAbs) - apply (simp add: domains_append + apply (simp add: doms_append fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] fresh_fin_union [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] - finite_domains finite_vrs tyvrs_fresh T_TAbs fresh_domain) + finite_doms finite_vrs tyvrs_fresh T_TAbs fresh_dom) apply (rule closed_in_weaken) apply (rule closed) done @@ -1513,8 +1402,8 @@ case (SA_refl_TVar G X' G') then have "\ (G' @ VarB x Q # \) ok" by simp then have h1:"\ (G' @ \) ok" by (auto dest: valid_cons') - have "X' \ ty_domain (G' @ VarB x Q # \)" using SA_refl_TVar by auto - then have h2:"X' \ ty_domain (G' @ \)" using ty_domain_vrs by auto + have "X' \ ty_dom (G' @ VarB x Q # \)" using SA_refl_TVar by auto + then have h2:"X' \ ty_dom (G' @ \)" using ty_dom_vrs by auto show ?case using h1 h2 by auto next case (SA_all G T1 S1 X S2 T2 G') @@ -1592,7 +1481,7 @@ next case (T_TAbs X T1 G t2 T2 x u D) from `TVarB X T1 # G \ t2 : T2` have "X \ T1" - by (auto simp add: valid_ty_domain_fresh dest: typing_ok intro!: closed_in_fresh) + by (auto simp add: valid_ty_dom_fresh dest: typing_ok intro!: closed_in_fresh) with `X \ u` and T_TAbs show ?case by fastsimp next case (T_TApp X G t1 T2 T11 T12 x u D) @@ -1627,7 +1516,7 @@ assume eq: "X = Y" from eq and SA_trans_TVar have "TVarB Y Q \ set G" by simp with G_ok have QS: "Q = S" using `TVarB Y S \ set G` by (rule uniqueness_of_ctxt) - from X\_ok have "X \ ty_domain \" and "Q closed_in \" by auto + from X\_ok have "X \ ty_dom \" and "Q closed_in \" by auto then have XQ: "X \ Q" by (rule closed_in_fresh) note `\\P<:Q` moreover from ST have "\ (D[X \ P]\<^sub>e @ \) ok" by (rule subtype_implies_ok) @@ -1652,8 +1541,8 @@ with neq and ST show ?thesis by auto next assume Y: "TVarB Y S \ set \" - from X\_ok have "X \ ty_domain \" and "\ \ ok" by auto - then have "X \ \" by (simp add: valid_ty_domain_fresh) + from X\_ok have "X \ ty_dom \" and "\ \ ok" by auto + then have "X \ \" by (simp add: valid_ty_dom_fresh) with Y have "X \ S" by (induct \) (auto simp add: fresh_list_nil fresh_list_cons) with ST have "(D[X \ P]\<^sub>e @ \)\S<:T[X \ P]\<^sub>\" @@ -1677,8 +1566,8 @@ by (auto intro: subtype_reflexivity) next assume neq: "X \ Y" - with SA_refl_TVar have "Y \ ty_domain (D[X \ P]\<^sub>e @ \)" - by (simp add: ty_domain_subst domains_append) + with SA_refl_TVar have "Y \ ty_dom (D[X \ P]\<^sub>e @ \)" + by (simp add: ty_dom_subst doms_append) with neq and ok show ?thesis by auto qed next @@ -1688,8 +1577,8 @@ show ?case using subtype_of.SA_arrow h1 h2 by auto next case (SA_all G T1 S1 Y S2 T2 X P D) - then have Y: "Y \ ty_domain (D @ TVarB X Q # \)" - by (auto dest: subtype_implies_ok intro: fresh_domain) + then have Y: "Y \ ty_dom (D @ TVarB X Q # \)" + by (auto dest: subtype_implies_ok intro: fresh_dom) moreover from SA_all have "S1 closed_in (D @ TVarB X Q # \)" by (auto dest: subtype_implies_closed) ultimately have S1: "Y \ S1" by (rule closed_in_fresh) @@ -1722,8 +1611,8 @@ next assume x: "VarB x T \ set G" from T_Var have ok: "\ G ok" by (auto dest: subtype_implies_ok) - then have "X \ ty_domain G" using T_Var by (auto dest: validE_append) - with ok have "X \ G" by (simp add: valid_ty_domain_fresh) + then have "X \ ty_dom G" using T_Var by (auto dest: validE_append) + with ok have "X \ G" by (simp add: valid_ty_dom_fresh) moreover from x have "VarB x T \ set (D' @ G)" by simp then have "VarB x (T[X \ P]\<^sub>\) \ set ((D' @ G)[X \ P]\<^sub>e)" by (rule ctxt_subst_mem_VarB) @@ -1744,15 +1633,15 @@ then show ?case using substT_subtype by force next case (T_TAbs X' G' T1 t2 T2 X P D') - then have "X' \ ty_domain (D' @ TVarB X Q # G)" + then have "X' \ ty_dom (D' @ TVarB X Q # G)" and "G' closed_in (D' @ TVarB X Q # G)" by (auto dest: typing_ok) then have "X' \ G'" by (rule closed_in_fresh) with T_TAbs show ?case by force next case (T_TApp X' G' t1 T2 T11 T12 X P D') - then have "X' \ ty_domain (D' @ TVarB X Q # G)" - by (simp add: fresh_domain) + then have "X' \ ty_dom (D' @ TVarB X Q # G)" + by (simp add: fresh_dom) moreover from T_TApp have "T11 closed_in (D' @ TVarB X Q # G)" by (auto dest: subtype_implies_closed) ultimately have X': "X' \ T11" by (rule closed_in_fresh) @@ -1783,7 +1672,7 @@ then have "[(y, x)] \ (VarB y S # \) \ [(y, x)] \ [(y, x)] \ s : [(y, x)] \ T\<^isub>2" by (rule typing.eqvt) moreover from T_Abs have "y \ \" - by (auto dest!: typing_ok simp add: fresh_trm_domain) + by (auto dest!: typing_ok simp add: fresh_trm_dom) ultimately have "VarB x S # \ \ s : T\<^isub>2" using T_Abs by (perm_simp add: ty_vrs_prm_simp) with ty1 show ?case using ty2 by (rule T_Abs) @@ -1819,7 +1708,7 @@ proof (nominal_induct \ t \ "\X<:S. s" T avoiding: X U U' S arbitrary: s rule: typing.strong_induct) case (T_TAbs Y T\<^isub>1 \ t\<^isub>2 T\<^isub>2) from `TVarB Y T\<^isub>1 # \ \ t\<^isub>2 : T\<^isub>2` have Y: "Y \ \" - by (auto dest!: typing_ok simp add: valid_ty_domain_fresh) + by (auto dest!: typing_ok simp add: valid_ty_dom_fresh) from `Y \ U'` and `Y \ X` have "(\X<:U. U') = (\Y<:U. [(Y, X)] \ U')" by (simp add: ty.inject alpha' fresh_atm) @@ -1907,7 +1796,7 @@ with T_TApp have "\ \ (\X<:T\<^isub>1\<^isub>1'. t\<^isub>1\<^isub>2) : (\X<:T\<^isub>1\<^isub>1. T\<^isub>1\<^isub>2)" and "X \ \" and "X \ T\<^isub>1\<^isub>1'" by (simp_all add: trm.inject) moreover from `\\T\<^isub>2<:T\<^isub>1\<^isub>1` and `X \ \` have "X \ T\<^isub>1\<^isub>1" - by (blast intro: closed_in_fresh fresh_domain dest: subtype_implies_closed) + by (blast intro: closed_in_fresh fresh_dom dest: subtype_implies_closed) ultimately obtain S' where "TVarB X T\<^isub>1\<^isub>1 # \ \ t\<^isub>1\<^isub>2 : S'" and "(TVarB X T\<^isub>1\<^isub>1 # \) \ S' <: T\<^isub>1\<^isub>2" diff -r 63686057cbe8 -r abda97d2deea src/HOL/Nominal/Examples/W.thy --- a/src/HOL/Nominal/Examples/W.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Nominal/Examples/W.thy Thu Jul 23 21:13:21 2009 +0200 @@ -25,6 +25,15 @@ using a by (induct Xs) (auto simp add: fresh_list_nil fresh_list_cons fresh_atm) +lemma difference_supset: + fixes xs::"'a list" + and ys::"'a list" + and zs::"'a list" + assumes asm: "set xs \ set ys" + shows "xs - ys = []" +using asm +by (induct xs) (auto) + nominal_datatype ty = TVar "tvar" | Fun "ty" "ty" ("_\_" [100,100] 100) @@ -49,65 +58,44 @@ text {* free type variables *} -class ftv = - fixes ftv :: "'a \ tvar list" +consts ftv :: "'a \ tvar list" -instantiation * :: (ftv, ftv) ftv +overloading + ftv_prod \ "ftv :: ('a \ 'b) \ tvar list" + ftv_tvar \ "ftv :: tvar \ tvar list" + ftv_var \ "ftv :: var \ tvar list" + ftv_list \ "ftv :: 'a list \ tvar list" + ftv_ty \ "ftv :: ty \ tvar list" begin -primrec ftv_prod +primrec + ftv_prod where - "ftv (x::'a::ftv, y::'b::ftv) = (ftv x)@(ftv y)" - -instance .. - -end - -instantiation tvar :: ftv -begin - -definition - ftv_of_tvar[simp]: "ftv X \ [(X::tvar)]" - -instance .. - -end - -instantiation var :: ftv -begin + "ftv_prod (x, y) = (ftv x) @ (ftv y)" definition - ftv_of_var[simp]: "ftv (x::var) \ []" - -instance .. - -end + ftv_tvar :: "tvar \ tvar list" +where +[simp]: "ftv_tvar X \ [(X::tvar)]" -instantiation list :: (ftv) ftv -begin - -primrec ftv_list +definition + ftv_var :: "var \ tvar list" where - "ftv [] = []" -| "ftv (x#xs) = (ftv x)@(ftv xs)" +[simp]: "ftv_var x \ []" -instance .. - -end - -(* free type-variables of types *) +primrec + ftv_list +where + "ftv_list [] = []" +| "ftv_list (x#xs) = (ftv x)@(ftv_list xs)" -instantiation ty :: ftv -begin - -nominal_primrec ftv_ty +nominal_primrec + ftv_ty :: "ty \ tvar list" where - "ftv (TVar X) = [X]" -| "ftv (T\<^isub>1\T\<^isub>2) = (ftv T\<^isub>1)@(ftv T\<^isub>2)" + "ftv_ty (TVar X) = [X]" +| "ftv_ty (T1 \T2) = (ftv_ty T1) @ (ftv_ty T2)" by (rule TrueI)+ -instance .. - end lemma ftv_ty_eqvt[eqvt]: @@ -117,13 +105,15 @@ by (nominal_induct T rule: ty.strong_induct) (perm_simp add: append_eqvt)+ -instantiation tyS :: ftv +overloading + ftv_tyS \ "ftv :: tyS \ tvar list" begin -nominal_primrec ftv_tyS +nominal_primrec + ftv_tyS :: "tyS \ tvar list" where - "ftv (Ty T) = ftv T" -| "ftv (\[X].S) = (ftv S) - [X]" + "ftv_tyS (Ty T) = ((ftv (T::ty))::tvar list)" +| "ftv_tyS (\[X].S) = (ftv_tyS S) - [X]" apply(finite_guess add: ftv_ty_eqvt fs_tvar1)+ apply(rule TrueI)+ apply(rule difference_fresh) @@ -131,8 +121,6 @@ apply(fresh_guess add: ftv_ty_eqvt fs_tvar1)+ done -instance .. - end lemma ftv_tyS_eqvt[eqvt]: @@ -174,6 +162,8 @@ shows "pi\(gen T Xs) = gen (pi\T) (pi\Xs)" by (induct Xs) (simp_all add: eqvts) + + abbreviation close :: "Ctxt \ ty \ tyS" where @@ -188,11 +178,11 @@ types Subst = "(tvar\ty) list" -class psubst = - fixes psubst :: "Subst \ 'a \ 'a" ("_<_>" [100,60] 120) +consts + psubst :: "Subst \ 'a \ 'a" ("_<_>" [100,60] 120) abbreviation - subst :: "'a::psubst \ tvar \ ty \ 'a" ("_[_::=_]" [100,100,100] 100) + subst :: "'a \ tvar \ ty \ 'a" ("_[_::=_]" [100,100,100] 100) where "smth[X::=T] \ ([(X,T)])" @@ -207,34 +197,163 @@ shows "pi\(lookup \ X) = lookup (pi\\) (pi\X)" by (induct \) (auto simp add: eqvts) -instantiation ty :: psubst +lemma lookup_fresh: + fixes X::"tvar" + assumes a: "X\\" + shows "lookup \ X = TVar X" +using a +by (induct \) + (auto simp add: fresh_list_cons fresh_prod fresh_atm) + +overloading + psubst_ty \ "psubst :: Subst \ ty \ ty" begin -nominal_primrec psubst_ty +nominal_primrec + psubst_ty where "\ = lookup \ X" | "\1 \ T\<^isub>2> = (\1>) \ (\2>)" by (rule TrueI)+ -instance .. - end lemma psubst_ty_eqvt[eqvt]: - fixes pi1::"tvar prm" + fixes pi::"tvar prm" and \::"Subst" and T::"ty" - shows "pi1\(\) = (pi1\\)<(pi1\T)>" + shows "pi\(\) = (pi\\)<(pi\T)>" by (induct T rule: ty.induct) (simp_all add: eqvts) -text {* instance *} -inductive - general :: "ty \ tyS \ bool"("_ \ _" [50,51] 50) +overloading + psubst_tyS \ "psubst :: Subst \ tyS \ tyS" +begin + +nominal_primrec + psubst_tyS :: "Subst \ tyS \ tyS" +where + "\<(Ty T)> = Ty (\)" +| "X\\ \ \<(\[X].S)> = \[X].(\)" +apply(finite_guess add: psubst_ty_eqvt fs_tvar1)+ +apply(rule TrueI)+ +apply(simp add: abs_fresh) +apply(fresh_guess add: psubst_ty_eqvt fs_tvar1)+ +done + +end + +overloading + psubst_Ctxt \ "psubst :: Subst \ Ctxt \ Ctxt" +begin + +fun + psubst_Ctxt :: "Subst \ Ctxt \ Ctxt" where - G_Ty[intro]: "T \ (Ty T)" -| G_All[intro]: "\X\T'; (T::ty) \ S\ \ T[X::=T'] \ \[X].S" + "psubst_Ctxt \ [] = []" +| "psubst_Ctxt \ ((x,S)#\) = (x,\)#(psubst_Ctxt \ \)" + +end + +lemma fresh_lookup: + fixes X::"tvar" + and \::"Subst" + and Y::"tvar" + assumes asms: "X\Y" "X\\" + shows "X\(lookup \ Y)" + using asms + by (induct \) + (auto simp add: fresh_list_cons fresh_prod fresh_atm) + +lemma fresh_psubst_ty: + fixes X::"tvar" + and \::"Subst" + and T::"ty" + assumes asms: "X\\" "X\T" + shows "X\\" + using asms + by (nominal_induct T rule: ty.strong_induct) + (auto simp add: fresh_list_append fresh_list_cons fresh_prod fresh_lookup) + +lemma fresh_psubst_tyS: + fixes X::"tvar" + and \::"Subst" + and S::"tyS" + assumes asms: "X\\" "X\S" + shows "X\\" + using asms + by (nominal_induct S avoiding: \ X rule: tyS.strong_induct) + (auto simp add: fresh_psubst_ty abs_fresh) + +lemma fresh_psubst_Ctxt: + fixes X::"tvar" + and \::"Subst" + and \::"Ctxt" + assumes asms: "X\\" "X\\" + shows "X\\<\>" +using asms +by (induct \) + (auto simp add: fresh_psubst_tyS fresh_list_cons) -equivariance general[tvar] +lemma subst_freshfact2_ty: + fixes X::"tvar" + and Y::"tvar" + and T::"ty" + assumes asms: "X\S" + shows "X\T[X::=S]" + using asms +by (nominal_induct T rule: ty.strong_induct) + (auto simp add: fresh_atm) + +text {* instance of a type scheme *} +inductive + inst :: "ty \ tyS \ bool"("_ \ _" [50,51] 50) +where + I_Ty[intro]: "T \ (Ty T)" +| I_All[intro]: "\X\T'; T \ S\ \ T[X::=T'] \ \[X].S" + +equivariance inst[tvar] + +nominal_inductive inst + by (simp_all add: abs_fresh subst_freshfact2_ty) + +lemma subst_forget_ty: + fixes T::"ty" + and X::"tvar" + assumes a: "X\T" + shows "T[X::=S] = T" + using a + by (nominal_induct T rule: ty.strong_induct) + (auto simp add: fresh_atm) + +lemma psubst_ty_lemma: + fixes \::"Subst" + and X::"tvar" + and T'::"ty" + and T::"ty" + assumes a: "X\\" + shows "\ = (\)[X::=\]" +using a +apply(nominal_induct T avoiding: \ X T' rule: ty.strong_induct) +apply(auto simp add: ty.inject lookup_fresh) +apply(rule sym) +apply(rule subst_forget_ty) +apply(rule fresh_lookup) +apply(simp_all add: fresh_atm) +done + +lemma general_preserved: + fixes \::"Subst" + assumes a: "T \ S" + shows "\ \ \" +using a +apply(nominal_induct T S avoiding: \ rule: inst.strong_induct) +apply(auto)[1] +apply(simp add: psubst_ty_lemma) +apply(rule_tac I_All) +apply(simp add: fresh_psubst_ty) +apply(simp) +done + text{* typing judgements *} inductive @@ -243,46 +362,64 @@ T_VAR[intro]: "\valid \; (x,S)\set \; T \ S\\ \ \ Var x : T" | T_APP[intro]: "\\ \ t\<^isub>1 : T\<^isub>1\T\<^isub>2; \ \ t\<^isub>2 : T\<^isub>1\\ \ \ App t\<^isub>1 t\<^isub>2 : T\<^isub>2" | T_LAM[intro]: "\x\\;((x,Ty T\<^isub>1)#\) \ t : T\<^isub>2\ \ \ \ Lam [x].t : T\<^isub>1\T\<^isub>2" -| T_LET[intro]: "\x\\; \ \ t\<^isub>1 : T\<^isub>1; ((x,close \ T\<^isub>1)#\) \ t\<^isub>2 : T\<^isub>2; set (ftv T\<^isub>1 - ftv \) \* T\<^isub>2\ \ \ \ Let x be t\<^isub>1 in t\<^isub>2 : T\<^isub>2" - -lemma fresh_star_tvar_eqvt[eqvt]: - "(pi::tvar prm) \ ((Xs::tvar set) \* (x::'a::{cp_tvar_tvar,pt_tvar})) = (pi \ Xs) \* (pi \ x)" - by (simp only: pt_fresh_star_bij_ineq(1) [OF pt_tvar_inst pt_tvar_inst at_tvar_inst cp_tvar_tvar_inst] perm_bool) +| T_LET[intro]: "\x\\; \ \ t\<^isub>1 : T\<^isub>1; ((x,close \ T\<^isub>1)#\) \ t\<^isub>2 : T\<^isub>2; set (ftv T\<^isub>1 - ftv \) \* T\<^isub>2\ + \ \ \ Let x be t\<^isub>1 in t\<^isub>2 : T\<^isub>2" equivariance typing[tvar] -lemma fresh_tvar_trm: "(X::tvar) \ (t::trm)" - by (nominal_induct t rule: trm.strong_induct) (simp_all add: fresh_atm abs_fresh) +lemma fresh_tvar_trm: + fixes X::"tvar" + and t::"trm" + shows "X\t" +by (nominal_induct t rule: trm.strong_induct) + (simp_all add: fresh_atm abs_fresh) -lemma ftv_ty: "supp (T::ty) = set (ftv T)" - by (nominal_induct T rule: ty.strong_induct) (simp_all add: ty.supp supp_atm) +lemma ftv_ty: + fixes T::"ty" + shows "supp T = set (ftv T)" +by (nominal_induct T rule: ty.strong_induct) + (simp_all add: ty.supp supp_atm) -lemma ftv_tyS: "supp (s::tyS) = set (ftv s)" - by (nominal_induct s rule: tyS.strong_induct) (auto simp add: tyS.supp abs_supp ftv_ty) +lemma ftv_tyS: + fixes S::"tyS" + shows "supp S = set (ftv S)" +by (nominal_induct S rule: tyS.strong_induct) + (auto simp add: tyS.supp abs_supp ftv_ty) -lemma ftv_Ctxt: "supp (\::Ctxt) = set (ftv \)" - apply (induct \) - apply (simp_all add: supp_list_nil supp_list_cons) - apply (case_tac a) - apply (simp add: supp_prod supp_atm ftv_tyS) - done +lemma ftv_Ctxt: + fixes \::"Ctxt" + shows "supp \ = set (ftv \)" +apply (induct \) +apply (simp_all add: supp_list_nil supp_list_cons) +apply (case_tac a) +apply (simp add: supp_prod supp_atm ftv_tyS) +done -lemma ftv_tvars: "supp (Tvs::tvar list) = set Tvs" - by (induct Tvs) (simp_all add: supp_list_nil supp_list_cons supp_atm) +lemma ftv_tvars: + fixes Tvs::"tvar list" + shows "supp Tvs = set Tvs" +by (induct Tvs) + (simp_all add: supp_list_nil supp_list_cons supp_atm) -lemma difference_supp: "((supp ((xs::tvar list) - ys))::tvar set) = supp xs - supp ys" - by (induct xs) (auto simp add: supp_list_nil supp_list_cons ftv_tvars) +lemma difference_supp: + fixes xs ys::"tvar list" + shows "((supp (xs - ys))::tvar set) = supp xs - supp ys" +by (induct xs) + (auto simp add: supp_list_nil supp_list_cons ftv_tvars) -lemma set_supp_eq: "set (xs::tvar list) = supp xs" - by (induct xs) (simp_all add: supp_list_nil supp_list_cons supp_atm) +lemma set_supp_eq: + fixes xs::"tvar list" + shows "set xs = supp xs" +by (induct xs) + (simp_all add: supp_list_nil supp_list_cons supp_atm) nominal_inductive2 typing avoids T_LET: "set (ftv T\<^isub>1 - ftv \)" - apply (simp add: fresh_star_def fresh_def ftv_Ctxt) - apply (simp add: fresh_star_def fresh_tvar_trm) - apply assumption - apply simp - done +apply (simp add: fresh_star_def fresh_def ftv_Ctxt) +apply (simp add: fresh_star_def fresh_tvar_trm) +apply assumption +apply simp +done lemma perm_fresh_fresh_aux: "\(x,y)\set (pi::tvar prm). x \ z \ y \ z \ pi \ (z::'a::pt_tvar) = z" @@ -294,8 +431,12 @@ apply (simp add: perm_fresh_fresh) done -lemma freshs_mem: "x \ (S::tvar set) \ S \* z \ x \ z" - by (simp add: fresh_star_def) +lemma freshs_mem: + fixes S::"tvar set" + assumes "x \ S" + and "S \* z" + shows "x \ z" +using prems by (simp add: fresh_star_def) lemma fresh_gen_set: fixes X::"tvar" @@ -315,13 +456,17 @@ shows "\(X::tvar)\set ((ftv T) - (ftv \)). X\(close \ T)" by (simp add: fresh_gen_set) -lemma gen_supp: "(supp (gen T Xs)::tvar set) = supp T - supp Xs" - by (induct Xs) (auto simp add: supp_list_nil supp_list_cons tyS.supp abs_supp supp_atm) +lemma gen_supp: + shows "(supp (gen T Xs)::tvar set) = supp T - supp Xs" +by (induct Xs) + (auto simp add: supp_list_nil supp_list_cons tyS.supp abs_supp supp_atm) -lemma minus_Int_eq: "T - (T - U) = T \ U" +lemma minus_Int_eq: + shows "T - (T - U) = T \ U" by blast -lemma close_supp: "supp (close \ T) = set (ftv T) \ set (ftv \)" +lemma close_supp: + shows "supp (close \ T) = set (ftv T) \ set (ftv \)" apply (simp add: gen_supp difference_supp ftv_ty ftv_Ctxt) apply (simp only: set_supp_eq minus_Int_eq) done @@ -370,4 +515,162 @@ ultimately show ?thesis by (rule T_LET) qed +lemma ftv_ty_subst: + fixes T::"ty" + and \::"Subst" + and X Y ::"tvar" + assumes a1: "X \ set (ftv T)" + and a2: "Y \ set (ftv (lookup \ X))" + shows "Y \ set (ftv (\))" +using a1 a2 +by (nominal_induct T rule: ty.strong_induct) (auto) + +lemma ftv_tyS_subst: + fixes S::"tyS" + and \::"Subst" + and X Y::"tvar" + assumes a1: "X \ set (ftv S)" + and a2: "Y \ set (ftv (lookup \ X))" + shows "Y \ set (ftv (\))" +using a1 a2 +by (nominal_induct S avoiding: \ Y rule: tyS.strong_induct) + (auto simp add: ftv_ty_subst fresh_atm) + +lemma ftv_Ctxt_subst: + fixes \::"Ctxt" + and \::"Subst" + assumes a1: "X \ set (ftv \)" + and a2: "Y \ set (ftv (lookup \ X))" + shows "Y \ set (ftv (\<\>))" +using a1 a2 +by (induct \) + (auto simp add: ftv_tyS_subst) + +lemma gen_preserved1: + assumes asm: "Xs \* \" + shows "\ = gen (\) Xs" +using asm +by (induct Xs) + (auto simp add: fresh_star_def) + +lemma gen_preserved2: + fixes T::"ty" + and \::"Ctxt" + assumes asm: "((ftv T) - (ftv \)) \* \" + shows "((ftv (\)) - (ftv (\<\>))) = ((ftv T) - (ftv \))" +using asm +apply(nominal_induct T rule: ty.strong_induct) +apply(auto simp add: fresh_star_def) +apply(simp add: lookup_fresh) +apply(simp add: ftv_Ctxt[symmetric]) +apply(fold fresh_def) +apply(rule fresh_psubst_Ctxt) +apply(assumption) +apply(assumption) +apply(rule difference_supset) +apply(auto) +apply(simp add: ftv_Ctxt_subst) +done + +lemma close_preserved: + fixes \::"Ctxt" + assumes asm: "((ftv T) - (ftv \)) \* \" + shows "\ T> = close (\<\>) (\)" +using asm +by (simp add: gen_preserved1 gen_preserved2) + +lemma var_fresh_for_ty: + fixes x::"var" + and T::"ty" + shows "x\T" +by (nominal_induct T rule: ty.strong_induct) + (simp_all add: fresh_atm) + +lemma var_fresh_for_tyS: + fixes x::"var" + and S::"tyS" + shows "x\S" +by (nominal_induct S rule: tyS.strong_induct) + (simp_all add: abs_fresh var_fresh_for_ty) + +lemma psubst_fresh_Ctxt: + fixes x::"var" + and \::"Ctxt" + and \::"Subst" + shows "x\\<\> = x\\" +by (induct \) + (auto simp add: fresh_list_cons fresh_list_nil fresh_prod var_fresh_for_tyS) + +lemma psubst_valid: + fixes \::Subst + and \::Ctxt + assumes a: "valid \" + shows "valid (\<\>)" +using a +by (induct) + (auto simp add: psubst_fresh_Ctxt) + +lemma psubst_in: + fixes \::"Ctxt" + and \::"Subst" + and pi::"tvar prm" + and S::"tyS" + assumes a: "(x,S)\set \" + shows "(x,\)\set (\<\>)" +using a +by (induct \) + (auto simp add: calc_atm) + + +lemma typing_preserved: + fixes \::"Subst" + and pi::"tvar prm" + assumes a: "\ \ t : T" + shows "(\<\>) \ t : (\)" +using a +proof (nominal_induct \ t T avoiding: \ rule: typing.strong_induct) + case (T_VAR \ x S T) + have a1: "valid \" by fact + have a2: "(x, S) \ set \" by fact + have a3: "T \ S" by fact + have "valid (\<\>)" using a1 by (simp add: psubst_valid) + moreover + have "(x,\)\set (\<\>)" using a2 by (simp add: psubst_in) + moreover + have "\ \ \" using a3 by (simp add: general_preserved) + ultimately show "(\<\>) \ Var x : (\)" by (simp add: typing.T_VAR) +next + case (T_APP \ t1 T1 T2 t2) + have "\<\> \ t1 : \ T2>" by fact + then have "\<\> \ t1 : (\) \ (\)" by simp + moreover + have "\<\> \ t2 : \" by fact + ultimately show "\<\> \ App t1 t2 : \" by (simp add: typing.T_APP) +next + case (T_LAM x \ T1 t T2) + fix pi::"tvar prm" and \::"Subst" + have "x\\" by fact + then have "x\\<\>" by (simp add: psubst_fresh_Ctxt) + moreover + have "\<((x, Ty T1)#\)> \ t : \" by fact + then have "((x, Ty (\))#(\<\>)) \ t : \" by (simp add: calc_atm) + ultimately show "\<\> \ Lam [x].t : \ T2>" by (simp add: typing.T_LAM) +next + case (T_LET x \ t1 T1 t2 T2) + have vc: "((ftv T1) - (ftv \)) \* \" by fact + have "x\\" by fact + then have a1: "x\\<\>" by (simp add: calc_atm psubst_fresh_Ctxt) + have a2: "\<\> \ t1 : \" by fact + have a3: "\<((x, close \ T1)#\)> \ t2 : \" by fact + from a2 a3 show "\<\> \ Let x be t1 in t2 : \" + apply - + apply(rule better_T_LET) + apply(rule a1) + apply(rule a2) + apply(simp add: close_preserved vc) + done +qed + + + end diff -r 63686057cbe8 -r abda97d2deea src/HOL/Nominal/nominal_datatype.ML --- a/src/HOL/Nominal/nominal_datatype.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Nominal/nominal_datatype.ML Thu Jul 23 21:13:21 2009 +0200 @@ -34,7 +34,7 @@ val In1_not_In0 = thm "In1_not_In0"; val Un_assoc = thm "Un_assoc"; val Collect_disj_eq = thm "Collect_disj_eq"; -val empty_def = thm "empty_def"; +val Collect_False_empty = @{thm empty_def [THEN sym, THEN eq_reflection]}; val empty_iff = thm "empty_iff"; open DatatypeAux; @@ -147,7 +147,7 @@ | perm_simproc' thy ss _ = NONE; val perm_simproc = - Simplifier.simproc (the_context ()) "perm_simp" ["pi1 \ (pi2 \ x)"] perm_simproc'; + Simplifier.simproc @{theory} "perm_simp" ["pi1 \ (pi2 \ x)"] perm_simproc'; val meta_spec = thm "meta_spec"; @@ -324,7 +324,7 @@ (perm_names_types ~~ perm_indnames)))) (fn _ => EVERY [indtac induction perm_indnames 1, ALLGOALS (asm_full_simp_tac - (simpset_of thy2 addsimps [perm_fun_def]))])), + (global_simpset_of thy2 addsimps [perm_fun_def]))])), length new_type_names)); (**** prove [] \ t = t ****) @@ -344,7 +344,7 @@ (perm_names ~~ map body_type perm_types ~~ perm_indnames))))) (fn _ => EVERY [indtac induction perm_indnames 1, - ALLGOALS (asm_full_simp_tac (simpset_of thy2))])), + ALLGOALS (asm_full_simp_tac (global_simpset_of thy2))])), length new_type_names)) end) atoms); @@ -379,7 +379,7 @@ (perm_names ~~ map body_type perm_types ~~ perm_indnames))))) (fn _ => EVERY [indtac induction perm_indnames 1, - ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt2', pt2_ax]))]))), + ALLGOALS (asm_full_simp_tac (global_simpset_of thy2 addsimps [pt2', pt2_ax]))]))), length new_type_names) end) atoms); @@ -415,7 +415,7 @@ (perm_names ~~ map body_type perm_types ~~ perm_indnames)))))) (fn _ => EVERY [indtac induction perm_indnames 1, - ALLGOALS (asm_full_simp_tac (simpset_of thy2 addsimps [pt3', pt3_rev', pt3_ax]))]))), + ALLGOALS (asm_full_simp_tac (global_simpset_of thy2 addsimps [pt3', pt3_rev', pt3_ax]))]))), length new_type_names) end) atoms); @@ -437,7 +437,7 @@ val permT2 = mk_permT (Type (name2, [])); val Ts = map body_type perm_types; val cp_inst = cp_inst_of thy name1 name2; - val simps = simpset_of thy addsimps (perm_fun_def :: + val simps = global_simpset_of thy addsimps (perm_fun_def :: (if name1 <> name2 then let val dj = dj_thm_of thy name2 name1 in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end @@ -601,7 +601,7 @@ end) (rep_set_names'' ~~ recTs' ~~ perm_indnames'))))) (fn _ => EVERY [indtac rep_induct [] 1, - ALLGOALS (simp_tac (simpset_of thy4 addsimps + ALLGOALS (simp_tac (global_simpset_of thy4 addsimps (symmetric perm_fun_def :: abs_perm))), ALLGOALS (resolve_tac rep_intrs THEN_ALL_NEW assume_tac)])), length new_type_names)); @@ -661,8 +661,8 @@ map (inter_sort thy sort o snd) tvs, [pt_class]) (EVERY [Class.intro_classes_tac [], rewrite_goals_tac [perm_def], - asm_full_simp_tac (simpset_of thy addsimps [Rep_inverse]) 1, - asm_full_simp_tac (simpset_of thy addsimps + asm_full_simp_tac (global_simpset_of thy addsimps [Rep_inverse]) 1, + asm_full_simp_tac (global_simpset_of thy addsimps [Rep RS perm_closed RS Abs_inverse]) 1, asm_full_simp_tac (HOL_basic_ss addsimps [PureThy.get_thm thy ("pt_" ^ Long_Name.base_name atom ^ "3")]) 1]) thy @@ -690,7 +690,7 @@ map (inter_sort thy sort o snd) tvs, [cp_class]) (EVERY [Class.intro_classes_tac [], rewrite_goals_tac [perm_def], - asm_full_simp_tac (simpset_of thy addsimps + asm_full_simp_tac (global_simpset_of thy addsimps ((Rep RS perm_closed1 RS Abs_inverse) :: (if atom1 = atom2 then [] else [Rep RS perm_closed2 RS Abs_inverse]))) 1, @@ -815,7 +815,7 @@ val (thy', defs', eqns') = Library.foldl (make_constr_def tname T T') ((Sign.add_path tname thy, defs, []), constrs ~~ constrs' ~~ constr_syntax) in - (parent_path (#flat_names config) thy', defs', eqns @ [eqns'], dist_lemmas @ [dist]) + (Sign.parent_path thy', defs', eqns @ [eqns'], dist_lemmas @ [dist]) end; val (thy8, constr_defs, constr_rep_eqns, dist_lemmas) = Library.foldl dt_constr_defs @@ -875,7 +875,7 @@ | prove_distinct_thms (p as (rep_thms, dist_lemma)) (k, t :: ts) = let val dist_thm = Goal.prove_global thy8 [] [] t (fn _ => - simp_tac (simpset_of thy8 addsimps (dist_lemma :: rep_thms)) 1) + simp_tac (global_simpset_of thy8 addsimps (dist_lemma :: rep_thms)) 1) in dist_thm :: standard (dist_thm RS not_sym) :: prove_distinct_thms p (k, ts) end; @@ -920,7 +920,7 @@ (HOLogic.mk_Trueprop (HOLogic.mk_eq (perm (list_comb (c, l_args)), list_comb (c, r_args))))) (fn _ => EVERY - [simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: perm_defs)) 1, + [simp_tac (global_simpset_of thy8 addsimps (constr_rep_thm :: perm_defs)) 1, simp_tac (HOL_basic_ss addsimps (Rep_thms @ Abs_inverse_thms @ constr_defs @ perm_closed_thms)) 1, TRY (simp_tac (HOL_basic_ss addsimps @@ -973,7 +973,7 @@ (HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)), foldr1 HOLogic.mk_conj eqs)))) (fn _ => EVERY - [asm_full_simp_tac (simpset_of thy8 addsimps (constr_rep_thm :: + [asm_full_simp_tac (global_simpset_of thy8 addsimps (constr_rep_thm :: rep_inject_thms')) 1, TRY (asm_full_simp_tac (HOL_basic_ss addsimps (fresh_def :: supp_def :: alpha @ abs_perm @ abs_fresh @ rep_inject_thms @ @@ -1013,11 +1013,11 @@ (HOLogic.mk_Trueprop (HOLogic.mk_eq (supp c, if null dts then HOLogic.mk_set atomT [] - else foldr1 (HOLogic.mk_binop @{const_name Un}) (map supp args2))))) + else foldr1 (HOLogic.mk_binop @{const_name union}) (map supp args2))))) (fn _ => simp_tac (HOL_basic_ss addsimps (supp_def :: Un_assoc :: de_Morgan_conj :: Collect_disj_eq :: finite_Un :: - symmetric empty_def :: finite_emptyI :: simp_thms @ + Collect_False_empty :: finite_emptyI :: simp_thms @ abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1) in (supp_thm, @@ -1100,7 +1100,7 @@ (Const ("Nominal.supp", T --> HOLogic.mk_setT atomT) $ Free (s, T))) (indnames ~~ recTs))))) (fn _ => indtac dt_induct indnames 1 THEN - ALLGOALS (asm_full_simp_tac (simpset_of thy8 addsimps + ALLGOALS (asm_full_simp_tac (global_simpset_of thy8 addsimps (abs_supp @ supp_atm @ PureThy.get_thms thy8 ("fs_" ^ Long_Name.base_name atom ^ "1") @ List.concat supp_thms))))), diff -r 63686057cbe8 -r abda97d2deea src/HOL/Nominal/nominal_fresh_fun.ML --- a/src/HOL/Nominal/nominal_fresh_fun.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Nominal/nominal_fresh_fun.ML Thu Jul 23 21:13:21 2009 +0200 @@ -143,7 +143,7 @@ let val thy = theory_of_thm thm; val ctx = Context.init_proof thy; - val ss = simpset_of thy; + val ss = global_simpset_of thy; val abs_fresh = PureThy.get_thms thy "abs_fresh"; val fresh_perm_app = PureThy.get_thms thy "fresh_perm_app"; val ss' = ss addsimps fresh_prod::abs_fresh; diff -r 63686057cbe8 -r abda97d2deea src/HOL/Nominal/nominal_inductive.ML --- a/src/HOL/Nominal/nominal_inductive.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Nominal/nominal_inductive.ML Thu Jul 23 21:13:21 2009 +0200 @@ -342,7 +342,7 @@ val env = Pattern.first_order_match thy (ihypt, prop_of ihyp) (Vartab.empty, Vartab.empty); val ihyp' = Thm.instantiate ([], map (pairself (cterm_of thy)) - (map (Envir.subst_vars env) vs ~~ + (map (Envir.subst_term env) vs ~~ map (fold_rev (NominalDatatype.mk_perm []) (rev pis' @ pis)) params' @ [z])) ihyp; fun mk_pi th = @@ -403,7 +403,7 @@ cut_facts_tac [th] 1 THEN REPEAT (etac conjE 1) THEN REPEAT (REPEAT (resolve_tac [conjI, impI] 1) THEN etac impE 1 THEN atac 1 THEN REPEAT (etac @{thm allE_Nil} 1) THEN - asm_full_simp_tac (simpset_of thy) 1) + asm_full_simp_tac (simpset_of ctxt) 1) end) |> singleton (ProofContext.export ctxt' ctxt); (** strong case analysis rule **) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Nominal/nominal_inductive2.ML --- a/src/HOL/Nominal/nominal_inductive2.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Nominal/nominal_inductive2.ML Thu Jul 23 21:13:21 2009 +0200 @@ -147,7 +147,7 @@ let val env = Pattern.first_order_match thy (p, prop_of th) (Vartab.empty, Vartab.empty) in Thm.instantiate ([], - map (Envir.subst_vars env #> cterm_of thy) vs ~~ cts) th + map (Envir.subst_term env #> cterm_of thy) vs ~~ cts) th end; fun prove_strong_ind s avoids ctxt = @@ -196,7 +196,7 @@ | add_set (t, T) ((u, U) :: ps) = if T = U then let val S = HOLogic.mk_setT T - in (Const (@{const_name "Un"}, S --> S --> S) $ u $ t, T) :: ps + in (Const (@{const_name union}, S --> S --> S) $ u $ t, T) :: ps end else (u, U) :: add_set (t, T) ps in @@ -438,7 +438,7 @@ cut_facts_tac [th] 1 THEN REPEAT (etac conjE 1) THEN REPEAT (REPEAT (resolve_tac [conjI, impI] 1) THEN etac impE 1 THEN atac 1 THEN REPEAT (etac @{thm allE_Nil} 1) THEN - asm_full_simp_tac (simpset_of thy) 1) + asm_full_simp_tac (simpset_of ctxt) 1) end) |> fresh_postprocess |> singleton (ProofContext.export ctxt' ctxt); diff -r 63686057cbe8 -r abda97d2deea src/HOL/Nominal/nominal_permeq.ML --- a/src/HOL/Nominal/nominal_permeq.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Nominal/nominal_permeq.ML Thu Jul 23 21:13:21 2009 +0200 @@ -117,7 +117,7 @@ | _ => NONE end -val perm_simproc_app = Simplifier.simproc (the_context ()) "perm_simproc_app" +val perm_simproc_app = Simplifier.simproc @{theory} "perm_simproc_app" ["Nominal.perm pi x"] perm_simproc_app'; (* a simproc that deals with permutation instances in front of functions *) @@ -137,7 +137,7 @@ | _ => NONE end -val perm_simproc_fun = Simplifier.simproc (the_context ()) "perm_simproc_fun" +val perm_simproc_fun = Simplifier.simproc @{theory} "perm_simproc_fun" ["Nominal.perm pi x"] perm_simproc_fun'; (* function for simplyfying permutations *) @@ -217,7 +217,7 @@ end | _ => NONE); -val perm_compose_simproc = Simplifier.simproc (the_context ()) "perm_compose" +val perm_compose_simproc = Simplifier.simproc @{theory} "perm_compose" ["Nominal.perm pi1 (Nominal.perm pi2 t)"] perm_compose_simproc'; fun perm_compose_tac ss i = @@ -404,19 +404,19 @@ Method.sections (Simplifier.simp_modifiers') >> (fn (prems, tac) => fn ctxt => METHOD (fn facts => HEADGOAL (Method.insert_tac (prems @ facts) THEN' - (CHANGED_PROP oo tac) (local_simpset_of ctxt)))); + (CHANGED_PROP oo tac) (simpset_of ctxt)))); (* setup so that the simpset is used which is active at the moment when the tactic is called *) fun local_simp_meth_setup tac = Method.sections (Simplifier.simp_modifiers' @ Splitter.split_modifiers) >> - (K (SIMPLE_METHOD' o tac o local_simpset_of)); + (K (SIMPLE_METHOD' o tac o simpset_of)); (* uses HOL_basic_ss only and fails if the tactic does not solve the subgoal *) fun basic_simp_meth_setup debug tac = Scan.depend (fn ctxt => Scan.succeed (Simplifier.map_ss (fn _ => HOL_basic_ss) ctxt, ())) -- Method.sections (Simplifier.simp_modifiers' @ Splitter.split_modifiers) >> - (K (SIMPLE_METHOD' o (fn ss => if debug then tac ss else SOLVEI (tac ss)) o local_simpset_of)); + (K (SIMPLE_METHOD' o (fn ss => if debug then tac ss else SOLVEI (tac ss)) o simpset_of)); val perm_simp_meth_debug = local_simp_meth_setup dperm_simp_tac; val perm_extend_simp_meth = local_simp_meth_setup perm_extend_simp_tac; diff -r 63686057cbe8 -r abda97d2deea src/HOL/Nominal/nominal_thmdecls.ML --- a/src/HOL/Nominal/nominal_thmdecls.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Nominal/nominal_thmdecls.ML Thu Jul 23 21:13:21 2009 +0200 @@ -156,7 +156,7 @@ fold (fn thm => Data.map (flag thm)) thms_to_be_added context end handle EQVT_FORM s => - error (Display.string_of_thm orig_thm ^ + error (Display.string_of_thm (Context.proof_of context) orig_thm ^ " does not comply with the form of an equivariance lemma (" ^ s ^").") diff -r 63686057cbe8 -r abda97d2deea src/HOL/Option.thy --- a/src/HOL/Option.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Option.thy Thu Jul 23 21:13:21 2009 +0200 @@ -105,7 +105,7 @@ "is_none x \ x = None" by (simp add: is_none_def) -lemma [code_inline]: +lemma [code_unfold]: "eq_class.eq x None \ is_none x" by (simp add: eq is_none_none) diff -r 63686057cbe8 -r abda97d2deea src/HOL/OrderedGroup.thy --- a/src/HOL/OrderedGroup.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/OrderedGroup.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1075,16 +1075,16 @@ lemma nprt_0[simp]: "nprt 0 = 0" by (simp add: nprt_def) lemma pprt_eq_id [simp, noatp]: "0 \ x \ pprt x = x" -by (simp add: pprt_def le_iff_sup sup_ACI) +by (simp add: pprt_def le_iff_sup sup_aci) lemma nprt_eq_id [simp, noatp]: "x \ 0 \ nprt x = x" -by (simp add: nprt_def le_iff_inf inf_ACI) +by (simp add: nprt_def le_iff_inf inf_aci) lemma pprt_eq_0 [simp, noatp]: "x \ 0 \ pprt x = 0" -by (simp add: pprt_def le_iff_sup sup_ACI) +by (simp add: pprt_def le_iff_sup sup_aci) lemma nprt_eq_0 [simp, noatp]: "0 \ x \ nprt x = 0" -by (simp add: nprt_def le_iff_inf inf_ACI) +by (simp add: nprt_def le_iff_inf inf_aci) lemma sup_0_imp_0: "sup a (- a) = 0 \ a = 0" proof - @@ -1120,7 +1120,7 @@ assume "0 <= a + a" hence a:"inf (a+a) 0 = 0" by (simp add: le_iff_inf inf_commute) have "(inf a 0)+(inf a 0) = inf (inf (a+a) 0) a" (is "?l=_") - by (simp add: add_sup_inf_distribs inf_ACI) + by (simp add: add_sup_inf_distribs inf_aci) hence "?l = 0 + inf a 0" by (simp add: a, simp add: inf_commute) hence "inf a 0 = 0" by (simp only: add_right_cancel) then show "0 <= a" by (simp add: le_iff_inf inf_commute) @@ -1206,10 +1206,10 @@ by (simp add: le_iff_inf nprt_def inf_commute) lemma pprt_mono [simp, noatp]: "a \ b \ pprt a \ pprt b" -by (simp add: le_iff_sup pprt_def sup_ACI sup_assoc [symmetric, of a]) +by (simp add: le_iff_sup pprt_def sup_aci sup_assoc [symmetric, of a]) lemma nprt_mono [simp, noatp]: "a \ b \ nprt a \ nprt b" -by (simp add: le_iff_inf nprt_def inf_ACI inf_assoc [symmetric, of a]) +by (simp add: le_iff_inf nprt_def inf_aci inf_assoc [symmetric, of a]) end @@ -1230,7 +1230,7 @@ then have "0 \ sup a (- a)" unfolding abs_lattice . then have "sup (sup a (- a)) 0 = sup a (- a)" by (rule sup_absorb1) then show ?thesis - by (simp add: add_sup_inf_distribs sup_ACI + by (simp add: add_sup_inf_distribs sup_aci pprt_def nprt_def diff_minus abs_lattice) qed @@ -1254,7 +1254,7 @@ show "\a + b\ \ \a\ + \b\" proof - have g:"abs a + abs b = sup (a+b) (sup (-a-b) (sup (-a+b) (a + (-b))))" (is "_=sup ?m ?n") - by (simp add: abs_lattice add_sup_inf_distribs sup_ACI diff_minus) + by (simp add: abs_lattice add_sup_inf_distribs sup_aci diff_minus) have a:"a+b <= sup ?m ?n" by (simp) have b:"-a-b <= ?n" by (simp) have c:"?n <= sup ?m ?n" by (simp) @@ -1380,7 +1380,7 @@ if termless_agrp (y, x) then SOME ac2 else NONE | solve_add_ac thy _ _ = NONE in - val add_ac_proc = Simplifier.simproc (the_context ()) + val add_ac_proc = Simplifier.simproc @{theory} "add_ac_proc" ["x + y::'a::ab_semigroup_add"] solve_add_ac; end; diff -r 63686057cbe8 -r abda97d2deea src/HOL/PReal.thy --- a/src/HOL/PReal.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/PReal.thy Thu Jul 23 21:13:21 2009 +0200 @@ -52,7 +52,7 @@ definition psup :: "preal set => preal" where - "psup P = Abs_preal (\X \ P. Rep_preal X)" + [code del]: "psup P = Abs_preal (\X \ P. Rep_preal X)" definition add_set :: "[rat set,rat set] => rat set" where diff -r 63686057cbe8 -r abda97d2deea src/HOL/Product_Type.thy --- a/src/HOL/Product_Type.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Product_Type.thy Thu Jul 23 21:13:21 2009 +0200 @@ -470,8 +470,8 @@ | NONE => NONE) | eta_proc _ _ = NONE; in - val split_beta_proc = Simplifier.simproc (the_context ()) "split_beta" ["split f z"] (K beta_proc); - val split_eta_proc = Simplifier.simproc (the_context ()) "split_eta" ["split f"] (K eta_proc); + val split_beta_proc = Simplifier.simproc @{theory} "split_beta" ["split f z"] (K beta_proc); + val split_eta_proc = Simplifier.simproc @{theory} "split_eta" ["split f"] (K eta_proc); end; Addsimprocs [split_beta_proc, split_eta_proc]; diff -r 63686057cbe8 -r abda97d2deea src/HOL/Prolog/prolog.ML --- a/src/HOL/Prolog/prolog.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Prolog/prolog.ML Thu Jul 23 21:13:21 2009 +0200 @@ -1,5 +1,4 @@ (* Title: HOL/Prolog/prolog.ML - ID: $Id$ Author: David von Oheimb (based on a lecture on Lambda Prolog by Nadathur) *) @@ -60,7 +59,7 @@ in map zero_var_indexes (at thm) end; val atomize_ss = - Simplifier.theory_context (the_context ()) empty_ss + Simplifier.theory_context @{theory} empty_ss setmksimps (mksimps mksimps_pairs) addsimps [ all_conj_distrib, (* "(! x. P x & Q x) = ((! x. P x) & (! x. Q x))" *) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Rational.thy --- a/src/HOL/Rational.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Rational.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1003,7 +1003,7 @@ definition (in term_syntax) valterm_fract :: "int \ (unit \ Code_Eval.term) \ int \ (unit \ Code_Eval.term) \ rat \ (unit \ Code_Eval.term)" where - [code_inline]: "valterm_fract k l = Code_Eval.valtermify Fract {\} k {\} l" + [code_unfold]: "valterm_fract k l = Code_Eval.valtermify Fract {\} k {\} l" notation fcomp (infixl "o>" 60) notation scomp (infixl "o\" 60) diff -r 63686057cbe8 -r abda97d2deea src/HOL/RealDef.thy --- a/src/HOL/RealDef.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/RealDef.thy Thu Jul 23 21:13:21 2009 +0200 @@ -946,7 +946,7 @@ end -lemma [code_unfold, symmetric, code_post]: +lemma [code_unfold_post]: "number_of k = real_of_int (number_of k)" unfolding number_of_is_id real_number_of_def .. @@ -1129,7 +1129,7 @@ definition (in term_syntax) valterm_ratreal :: "rat \ (unit \ Code_Eval.term) \ real \ (unit \ Code_Eval.term)" where - [code_inline]: "valterm_ratreal k = Code_Eval.valtermify Ratreal {\} k" + [code_unfold]: "valterm_ratreal k = Code_Eval.valtermify Ratreal {\} k" notation fcomp (infixl "o>" 60) notation scomp (infixl "o\" 60) diff -r 63686057cbe8 -r abda97d2deea src/HOL/SEQ.thy --- a/src/HOL/SEQ.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/SEQ.thy Thu Jul 23 21:13:21 2009 +0200 @@ -113,8 +113,8 @@ lemma Bseq_conv_Bfun: "Bseq X \ Bfun X sequentially" unfolding Bfun_def eventually_sequentially apply (rule iffI) -apply (simp add: Bseq_def, fast) -apply (fast intro: BseqI2') +apply (simp add: Bseq_def) +apply (auto intro: BseqI2') done @@ -762,13 +762,25 @@ lemma lemma_NBseq_def: "(\K > 0. \n. norm (X n) \ K) = (\N. \n. norm (X n) \ real(Suc N))" -apply auto - prefer 2 apply force -apply (cut_tac x = K in reals_Archimedean2, clarify) -apply (rule_tac x = n in exI, clarify) -apply (drule_tac x = na in spec) -apply (auto simp add: real_of_nat_Suc) -done +proof auto + fix K :: real + from reals_Archimedean2 obtain n :: nat where "K < real n" .. + then have "K \ real (Suc n)" by auto + assume "\m. norm (X m) \ K" + have "\m. norm (X m) \ real (Suc n)" + proof + fix m :: 'a + from `\m. norm (X m) \ K` have "norm (X m) \ K" .. + with `K \ real (Suc n)` show "norm (X m) \ real (Suc n)" by auto + qed + then show "\N. \n. norm (X n) \ real (Suc N)" .. +next + fix N :: nat + have "real (Suc N) > 0" by (simp add: real_of_nat_Suc) + moreover assume "\n. norm (X n) \ real (Suc N)" + ultimately show "\K>0. \n. norm (X n) \ K" by blast +qed + text{* alternative definition for Bseq *} lemma Bseq_iff: "Bseq X = (\N. \n. norm (X n) \ real(Suc N))" @@ -1105,7 +1117,7 @@ lemma (in real_Cauchy) isLub_ex: "\u. isLub UNIV S u" proof (rule reals_complete) obtain N where "\m\N. \n\N. norm (X m - X n) < 1" - using CauchyD [OF X zero_less_one] by fast + using CauchyD [OF X zero_less_one] by auto hence N: "\n\N. norm (X n - X N) < 1" by simp show "\x. x \ S" proof @@ -1129,7 +1141,7 @@ fix r::real assume "0 < r" hence r: "0 < r/2" by simp obtain N where "\n\N. \m\N. norm (X n - X m) < r/2" - using CauchyD [OF X r] by fast + using CauchyD [OF X r] by auto hence "\n\N. norm (X n - X N) < r/2" by simp hence N: "\n\N. X N - r/2 < X n \ X n < X N + r/2" by (simp only: real_norm_def real_abs_diff_less_iff) diff -r 63686057cbe8 -r abda97d2deea src/HOL/SET-Protocol/MessageSET.thy --- a/src/HOL/SET-Protocol/MessageSET.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/SET-Protocol/MessageSET.thy Thu Jul 23 21:13:21 2009 +0200 @@ -854,6 +854,8 @@ Abstraction over i is ESSENTIAL: it delays the dereferencing of claset DEPENDS UPON "X" REFERRING TO THE FRADULENT MESSAGE *) +fun impOfSubs th = th RSN (2, @{thm rev_subsetD}) + (*Apply rules to break down assumptions of the form Y \ parts(insert X H) and Y \ analz(insert X H) *) @@ -937,17 +939,17 @@ method_setup spy_analz = {* Scan.succeed (fn ctxt => - SIMPLE_METHOD' (MessageSET.spy_analz_tac (local_clasimpset_of ctxt))) *} + SIMPLE_METHOD' (MessageSET.spy_analz_tac (clasimpset_of ctxt))) *} "for proving the Fake case when analz is involved" method_setup atomic_spy_analz = {* Scan.succeed (fn ctxt => - SIMPLE_METHOD' (MessageSET.atomic_spy_analz_tac (local_clasimpset_of ctxt))) *} + SIMPLE_METHOD' (MessageSET.atomic_spy_analz_tac (clasimpset_of ctxt))) *} "for debugging spy_analz" method_setup Fake_insert_simp = {* Scan.succeed (fn ctxt => - SIMPLE_METHOD' (MessageSET.Fake_insert_simp_tac (local_simpset_of ctxt))) *} + SIMPLE_METHOD' (MessageSET.Fake_insert_simp_tac (simpset_of ctxt))) *} "for debugging spy_analz" end diff -r 63686057cbe8 -r abda97d2deea src/HOL/SET-Protocol/PublicSET.thy --- a/src/HOL/SET-Protocol/PublicSET.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/SET-Protocol/PublicSET.thy Thu Jul 23 21:13:21 2009 +0200 @@ -350,7 +350,7 @@ (*Tactic for possibility theorems*) fun possibility_tac ctxt = REPEAT (*omit used_Says so that Nonces start from different traces!*) - (ALLGOALS (simp_tac (local_simpset_of ctxt delsimps [@{thm used_Says}, @{thm used_Notes}])) + (ALLGOALS (simp_tac (simpset_of ctxt delsimps [@{thm used_Says}, @{thm used_Notes}])) THEN REPEAT_FIRST (eq_assume_tac ORELSE' resolve_tac [refl, conjI, @{thm Nonce_supply}])) @@ -359,7 +359,7 @@ nonces and keys initially*) fun basic_possibility_tac ctxt = REPEAT - (ALLGOALS (asm_simp_tac (local_simpset_of ctxt setSolver safe_solver)) + (ALLGOALS (asm_simp_tac (simpset_of ctxt setSolver safe_solver)) THEN REPEAT_FIRST (resolve_tac [refl, conjI])) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Set.thy --- a/src/HOL/Set.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Set.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1,6 +1,4 @@ -(* Title: HOL/Set.thy - Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel -*) +(* Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel *) header {* Set theory for higher-order logic *} @@ -8,9 +6,7 @@ imports Lattices begin -text {* A set in HOL is simply a predicate. *} - -subsection {* Basic syntax *} +subsection {* Sets as predicates *} global @@ -19,9 +15,6 @@ consts Collect :: "('a => bool) => 'a set" -- "comprehension" "op :" :: "'a => 'a set => bool" -- "membership" - Ball :: "'a set => ('a => bool) => bool" -- "bounded universal quantifiers" - Bex :: "'a set => ('a => bool) => bool" -- "bounded existential quantifiers" - Bex1 :: "'a set => ('a => bool) => bool" -- "bounded unique existential quantifiers" local @@ -29,6 +22,10 @@ "op :" ("op :") and "op :" ("(_/ : _)" [50, 51] 50) +defs + mem_def [code]: "x : S == S x" + Collect_def [code]: "Collect P == P" + abbreviation "not_mem x A == ~ (x : A)" -- "non-membership" @@ -48,34 +45,66 @@ not_mem ("op \") and not_mem ("(_/ \ _)" [50, 51] 50) +text {* Set comprehensions *} + syntax "@Coll" :: "pttrn => bool => 'a set" ("(1{_./ _})") translations "{x. P}" == "Collect (%x. P)" -definition Int :: "'a set \ 'a set \ 'a set" (infixl "Int" 70) where - "A Int B \ {x. x \ A \ x \ B}" +syntax + "@SetCompr" :: "'a => idts => bool => 'a set" ("(1{_ |/_./ _})") + "@Collect" :: "idt => 'a set => bool => 'a set" ("(1{_ :/ _./ _})") + +syntax (xsymbols) + "@Collect" :: "idt => 'a set => bool => 'a set" ("(1{_ \/ _./ _})") + +translations + "{x:A. P}" => "{x. x:A & P}" -definition Un :: "'a set \ 'a set \ 'a set" (infixl "Un" 65) where - "A Un B \ {x. x \ A \ x \ B}" +lemma mem_Collect_eq [iff]: "(a : {x. P(x)}) = P(a)" + by (simp add: Collect_def mem_def) + +lemma Collect_mem_eq [simp]: "{x. x:A} = A" + by (simp add: Collect_def mem_def) + +lemma CollectI: "P(a) ==> a : {x. P(x)}" + by simp + +lemma CollectD: "a : {x. P(x)} ==> P(a)" + by simp -notation (xsymbols) - "Int" (infixl "\" 70) and - "Un" (infixl "\" 65) +lemma Collect_cong: "(!!x. P x = Q x) ==> {x. P(x)} = {x. Q(x)}" + by simp + +text {* +Simproc for pulling @{text "x=t"} in @{text "{x. \ & x=t & \}"} +to the front (and similarly for @{text "t=x"}): +*} -notation (HTML output) - "Int" (infixl "\" 70) and - "Un" (infixl "\" 65) +setup {* +let + val Coll_perm_tac = rtac @{thm Collect_cong} 1 THEN rtac @{thm iffI} 1 THEN + ALLGOALS(EVERY'[REPEAT_DETERM o (etac @{thm conjE}), + DEPTH_SOLVE_1 o (ares_tac [@{thm conjI}])]) + val defColl_regroup = Simplifier.simproc @{theory} + "defined Collect" ["{x. P x & Q x}"] + (Quantifier1.rearrange_Coll Coll_perm_tac) +in + Simplifier.map_simpset (fn ss => ss addsimprocs [defColl_regroup]) +end +*} + +lemmas CollectE = CollectD [elim_format] + +text {* Set enumerations *} definition empty :: "'a set" ("{}") where - "empty \ {x. False}" + bot_set_eq [symmetric]: "{} = bot" definition insert :: "'a \ 'a set \ 'a set" where - "insert a B \ {x. x = a} \ B" - -definition UNIV :: "'a set" where - "UNIV \ {x. True}" + insert_compr: "insert a B = {x. x = a \ x \ B}" syntax "@Finset" :: "args => 'a set" ("{(_)}") @@ -84,93 +113,8 @@ "{x, xs}" == "CONST insert x {xs}" "{x}" == "CONST insert x {}" -syntax - "_Ball" :: "pttrn => 'a set => bool => bool" ("(3ALL _:_./ _)" [0, 0, 10] 10) - "_Bex" :: "pttrn => 'a set => bool => bool" ("(3EX _:_./ _)" [0, 0, 10] 10) - "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3EX! _:_./ _)" [0, 0, 10] 10) - "_Bleast" :: "id => 'a set => bool => 'a" ("(3LEAST _:_./ _)" [0, 0, 10] 10) -syntax (HOL) - "_Ball" :: "pttrn => 'a set => bool => bool" ("(3! _:_./ _)" [0, 0, 10] 10) - "_Bex" :: "pttrn => 'a set => bool => bool" ("(3? _:_./ _)" [0, 0, 10] 10) - "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3?! _:_./ _)" [0, 0, 10] 10) - -syntax (xsymbols) - "_Ball" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) - "_Bex" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) - "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3\!_\_./ _)" [0, 0, 10] 10) - "_Bleast" :: "id => 'a set => bool => 'a" ("(3LEAST_\_./ _)" [0, 0, 10] 10) - -syntax (HTML output) - "_Ball" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) - "_Bex" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) - "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3\!_\_./ _)" [0, 0, 10] 10) - -translations - "ALL x:A. P" == "Ball A (%x. P)" - "EX x:A. P" == "Bex A (%x. P)" - "EX! x:A. P" == "Bex1 A (%x. P)" - "LEAST x:A. P" => "LEAST x. x:A & P" - -definition INTER :: "'a set \ ('a \ 'b set) \ 'b set" where - "INTER A B \ {y. \x\A. y \ B x}" - -definition UNION :: "'a set \ ('a \ 'b set) \ 'b set" where - "UNION A B \ {y. \x\A. y \ B x}" - -definition Inter :: "'a set set \ 'a set" where - "Inter S \ INTER S (\x. x)" - -definition Union :: "'a set set \ 'a set" where - "Union S \ UNION S (\x. x)" - -notation (xsymbols) - Inter ("\_" [90] 90) and - Union ("\_" [90] 90) - - -subsection {* Additional concrete syntax *} - -syntax - "@SetCompr" :: "'a => idts => bool => 'a set" ("(1{_ |/_./ _})") - "@Collect" :: "idt => 'a set => bool => 'a set" ("(1{_ :/ _./ _})") - "@INTER1" :: "pttrns => 'b set => 'b set" ("(3INT _./ _)" [0, 10] 10) - "@UNION1" :: "pttrns => 'b set => 'b set" ("(3UN _./ _)" [0, 10] 10) - "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3INT _:_./ _)" [0, 10] 10) - "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3UN _:_./ _)" [0, 10] 10) - -syntax (xsymbols) - "@Collect" :: "idt => 'a set => bool => 'a set" ("(1{_ \/ _./ _})") - "@INTER1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) - "@UNION1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) - "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 10] 10) - "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 10] 10) - -syntax (latex output) - "@INTER1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) - "@UNION1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) - "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 10] 10) - "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 10] 10) - -translations - "{x:A. P}" => "{x. x:A & P}" - "INT x y. B" == "INT x. INT y. B" - "INT x. B" == "CONST INTER CONST UNIV (%x. B)" - "INT x. B" == "INT x:CONST UNIV. B" - "INT x:A. B" == "CONST INTER A (%x. B)" - "UN x y. B" == "UN x. UN y. B" - "UN x. B" == "CONST UNION CONST UNIV (%x. B)" - "UN x. B" == "UN x:CONST UNIV. B" - "UN x:A. B" == "CONST UNION A (%x. B)" - -text {* - Note the difference between ordinary xsymbol syntax of indexed - unions and intersections (e.g.\ @{text"\a\<^isub>1\A\<^isub>1. B"}) - and their \LaTeX\ rendition: @{term"\a\<^isub>1\A\<^isub>1. B"}. The - former does not make the index expression a subscript of the - union/intersection symbol because this leads to problems with nested - subscripts in Proof General. -*} +subsection {* Subsets and bounded quantifiers *} abbreviation subset :: "'a set \ 'a set \ bool" where @@ -212,9 +156,47 @@ supset_eq ("op \") and supset_eq ("(_/ \ _)" [50, 51] 50) +global +consts + Ball :: "'a set => ('a => bool) => bool" -- "bounded universal quantifiers" + Bex :: "'a set => ('a => bool) => bool" -- "bounded existential quantifiers" + Bex1 :: "'a set => ('a => bool) => bool" -- "bounded unique existential quantifiers" + +local + +defs + Ball_def: "Ball A P == ALL x. x:A --> P(x)" + Bex_def: "Bex A P == EX x. x:A & P(x)" + Bex1_def: "Bex1 A P == EX! x. x:A & P(x)" + +syntax + "_Ball" :: "pttrn => 'a set => bool => bool" ("(3ALL _:_./ _)" [0, 0, 10] 10) + "_Bex" :: "pttrn => 'a set => bool => bool" ("(3EX _:_./ _)" [0, 0, 10] 10) + "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3EX! _:_./ _)" [0, 0, 10] 10) + "_Bleast" :: "id => 'a set => bool => 'a" ("(3LEAST _:_./ _)" [0, 0, 10] 10) -subsubsection "Bounded quantifiers" +syntax (HOL) + "_Ball" :: "pttrn => 'a set => bool => bool" ("(3! _:_./ _)" [0, 0, 10] 10) + "_Bex" :: "pttrn => 'a set => bool => bool" ("(3? _:_./ _)" [0, 0, 10] 10) + "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3?! _:_./ _)" [0, 0, 10] 10) + +syntax (xsymbols) + "_Ball" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) + "_Bex" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) + "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3\!_\_./ _)" [0, 0, 10] 10) + "_Bleast" :: "id => 'a set => bool => 'a" ("(3LEAST_\_./ _)" [0, 0, 10] 10) + +syntax (HTML output) + "_Ball" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) + "_Bex" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) + "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3\!_\_./ _)" [0, 0, 10] 10) + +translations + "ALL x:A. P" == "Ball A (%x. P)" + "EX x:A. P" == "Bex A (%x. P)" + "EX! x:A. P" == "Bex1 A (%x. P)" + "LEAST x:A. P" => "LEAST x. x:A & P" syntax (output) "_setlessAll" :: "[idt, 'a, bool] => bool" ("(3ALL _<_./ _)" [0, 0, 10] 10) @@ -307,17 +289,10 @@ in [("@SetCompr", setcompr_tr)] end; *} -(* To avoid eta-contraction of body: *) -print_translation {* -let - fun btr' syn [A, Abs abs] = - let val (x, t) = atomic_abs_tr' abs - in Syntax.const syn $ x $ A $ t end -in -[(@{const_syntax Ball}, btr' "_Ball"), (@{const_syntax Bex}, btr' "_Bex"), - (@{const_syntax UNION}, btr' "@UNION"),(@{const_syntax INTER}, btr' "@INTER")] -end -*} +print_translation {* [ +Syntax.preserve_binder_abs2_tr' @{const_syntax Ball} "_Ball", +Syntax.preserve_binder_abs2_tr' @{const_syntax Bex} "_Bex" +] *} -- {* to avoid eta-contraction of body *} print_translation {* let @@ -348,54 +323,22 @@ in [("Collect", setcompr_tr')] end; *} - -subsection {* Rules and definitions *} - -text {* Isomorphisms between predicates and sets. *} - -defs - mem_def [code]: "x : S == S x" - Collect_def [code]: "Collect P == P" - -defs - Ball_def: "Ball A P == ALL x. x:A --> P(x)" - Bex_def: "Bex A P == EX x. x:A & P(x)" - Bex1_def: "Bex1 A P == EX! x. x:A & P(x)" - -definition Pow :: "'a set => 'a set set" where - Pow_def: "Pow A = {B. B \ A}" - -definition image :: "('a => 'b) => 'a set => 'b set" (infixr "`" 90) where - image_def: "f ` A = {y. EX x:A. y = f(x)}" - -abbreviation - range :: "('a => 'b) => 'b set" where -- "of function" - "range f == f ` UNIV" - - -subsection {* Lemmas and proof tool setup *} - -subsubsection {* Relating predicates and sets *} - -lemma mem_Collect_eq [iff]: "(a : {x. P(x)}) = P(a)" - by (simp add: Collect_def mem_def) - -lemma Collect_mem_eq [simp]: "{x. x:A} = A" - by (simp add: Collect_def mem_def) - -lemma CollectI: "P(a) ==> a : {x. P(x)}" - by simp - -lemma CollectD: "a : {x. P(x)} ==> P(a)" - by simp - -lemma Collect_cong: "(!!x. P x = Q x) ==> {x. P(x)} = {x. Q(x)}" - by simp - -lemmas CollectE = CollectD [elim_format] - - -subsubsection {* Bounded quantifiers *} +setup {* +let + val unfold_bex_tac = unfold_tac @{thms "Bex_def"}; + fun prove_bex_tac ss = unfold_bex_tac ss THEN Quantifier1.prove_one_point_ex_tac; + val rearrange_bex = Quantifier1.rearrange_bex prove_bex_tac; + val unfold_ball_tac = unfold_tac @{thms "Ball_def"}; + fun prove_ball_tac ss = unfold_ball_tac ss THEN Quantifier1.prove_one_point_all_tac; + val rearrange_ball = Quantifier1.rearrange_ball prove_ball_tac; + val defBEX_regroup = Simplifier.simproc @{theory} + "defined BEX" ["EX x:A. P x & Q x"] rearrange_bex; + val defBALL_regroup = Simplifier.simproc @{theory} + "defined BALL" ["ALL x:A. P x --> Q x"] rearrange_ball; +in + Simplifier.map_simpset (fn ss => ss addsimprocs [defBALL_regroup, defBEX_regroup]) +end +*} lemma ballI [intro!]: "(!!x. x:A ==> P x) ==> ALL x:A. P x" by (simp add: Ball_def) @@ -405,20 +348,6 @@ lemma bspec [dest?]: "ALL x:A. P x ==> x:A ==> P x" by (simp add: Ball_def) -lemma ballE [elim]: "ALL x:A. P x ==> (P x ==> Q) ==> (x ~: A ==> Q) ==> Q" - by (unfold Ball_def) blast - -ML {* bind_thm ("rev_ballE", Thm.permute_prems 1 1 @{thm ballE}) *} - -text {* - \medskip This tactic takes assumptions @{prop "ALL x:A. P x"} and - @{prop "a:A"}; creates assumption @{prop "P a"}. -*} - -ML {* - fun ball_tac i = etac @{thm ballE} i THEN contr_tac (i + 1) -*} - text {* Gives better instantiation for bound: *} @@ -427,6 +356,26 @@ Classical.map_cs (fn cs => cs addbefore ("bspec", datac @{thm bspec} 1)) *} +ML {* +structure Simpdata = +struct + +open Simpdata; + +val mksimps_pairs = [(@{const_name Ball}, @{thms bspec})] @ mksimps_pairs; + +end; + +open Simpdata; +*} + +declaration {* fn _ => + Simplifier.map_ss (fn ss => ss setmksimps (mksimps mksimps_pairs)) +*} + +lemma ballE [elim]: "ALL x:A. P x ==> (P x ==> Q) ==> (x ~: A ==> Q) ==> Q" + by (unfold Ball_def) blast + lemma bexI [intro]: "P x ==> x:A ==> EX x:A. P x" -- {* Normally the best argument order: @{prop "P x"} constrains the choice of @{prop "x:A"}. *} @@ -468,27 +417,8 @@ lemma ball_one_point2 [simp]: "(ALL x:A. a = x --> P x) = (a:A --> P a)" by blast -ML {* - local - val unfold_bex_tac = unfold_tac @{thms "Bex_def"}; - fun prove_bex_tac ss = unfold_bex_tac ss THEN Quantifier1.prove_one_point_ex_tac; - val rearrange_bex = Quantifier1.rearrange_bex prove_bex_tac; - val unfold_ball_tac = unfold_tac @{thms "Ball_def"}; - fun prove_ball_tac ss = unfold_ball_tac ss THEN Quantifier1.prove_one_point_all_tac; - val rearrange_ball = Quantifier1.rearrange_ball prove_ball_tac; - in - val defBEX_regroup = Simplifier.simproc (the_context ()) - "defined BEX" ["EX x:A. P x & Q x"] rearrange_bex; - val defBALL_regroup = Simplifier.simproc (the_context ()) - "defined BALL" ["ALL x:A. P x --> Q x"] rearrange_ball; - end; - - Addsimprocs [defBALL_regroup, defBEX_regroup]; -*} - - -subsubsection {* Congruence rules *} +text {* Congruence rules *} lemma ball_cong: "A = B ==> (!!x. x:B ==> P x = Q x) ==> @@ -511,6 +441,8 @@ by (simp add: simp_implies_def Bex_def cong: conj_cong) +subsection {* Basic operations *} + subsubsection {* Subsets *} lemma subsetI [atp,intro!]: "(!!x. x:A ==> x:B) ==> A \ B" @@ -535,33 +467,29 @@ \medskip Converts @{prop "A \ B"} to @{prop "x \ A ==> x \ B"}. *} -ML {* - fun impOfSubs th = th RSN (2, @{thm rev_subsetD}) -*} - lemma subsetCE [elim]: "A \ B ==> (c \ A ==> P) ==> (c \ B ==> P) ==> P" -- {* Classical elimination rule. *} by (unfold mem_def) blast lemma subset_eq: "A \ B = (\x\A. x \ B)" by blast -text {* - \medskip Takes assumptions @{prop "A \ B"}; @{prop "c \ A"} and - creates the assumption @{prop "c \ B"}. -*} - -ML {* - fun set_mp_tac i = etac @{thm subsetCE} i THEN mp_tac i -*} - lemma contra_subsetD: "A \ B ==> c \ B ==> c \ A" by blast lemma subset_refl [simp,atp]: "A \ A" - by fast + by (fact order_refl) lemma subset_trans: "A \ B ==> B \ C ==> A \ C" - by blast + by (fact order_trans) + +lemma set_rev_mp: "x:A ==> A \ B ==> x:B" + by (rule subsetD) + +lemma set_mp: "A \ B ==> x:A ==> x:B" + by (rule subsetD) + +lemmas basic_trans_rules [trans] = + order_trans_rules set_rev_mp set_mp subsubsection {* Equality *} @@ -613,6 +541,13 @@ subsubsection {* The universal set -- UNIV *} +definition UNIV :: "'a set" where + top_set_eq [symmetric]: "UNIV = top" + +lemma UNIV_def: + "UNIV = {x. True}" + by (simp add: top_set_eq [symmetric] top_fun_eq top_bool_eq Collect_def) + lemma UNIV_I [simp]: "x : UNIV" by (simp add: UNIV_def) @@ -642,6 +577,10 @@ subsubsection {* The empty set *} +lemma empty_def: + "{} = {x. False}" + by (simp add: bot_set_eq [symmetric] bot_fun_eq bot_bool_eq Collect_def) + lemma empty_iff [simp]: "(c : {}) = False" by (simp add: empty_def) @@ -656,7 +595,7 @@ by blast lemma equals0D: "A = {} ==> a \ A" - -- {* Use for reasoning about disjointness: @{prop "A Int B = {}"} *} + -- {* Use for reasoning about disjointness: @{text "A Int B = {}"} *} by blast lemma ball_empty [simp]: "Ball {} P = True" @@ -671,6 +610,9 @@ subsubsection {* The Powerset operator -- Pow *} +definition Pow :: "'a set => 'a set set" where + Pow_def: "Pow A = {B. B \ A}" + lemma Pow_iff [iff]: "(A \ Pow B) = (A \ B)" by (simp add: Pow_def) @@ -710,6 +652,19 @@ subsubsection {* Binary union -- Un *} +definition union :: "'a set \ 'a set \ 'a set" (infixl "Un" 65) where + sup_set_eq [symmetric]: "A Un B = sup A B" + +notation (xsymbols) + union (infixl "\" 65) + +notation (HTML output) + union (infixl "\" 65) + +lemma Un_def: + "A \ B = {x. x \ A \ x \ B}" + by (simp add: sup_fun_eq sup_bool_eq sup_set_eq [symmetric] Collect_def mem_def) + lemma Un_iff [simp]: "(c : A Un B) = (c:A | c:B)" by (unfold Un_def) blast @@ -730,9 +685,30 @@ lemma UnE [elim!]: "c : A Un B ==> (c:A ==> P) ==> (c:B ==> P) ==> P" by (unfold Un_def) blast +lemma insert_def: "insert a B = {x. x = a} \ B" + by (simp add: Collect_def mem_def insert_compr Un_def) + +lemma mono_Un: "mono f \ f A \ f B \ f (A \ B)" + apply (fold sup_set_eq) + apply (erule mono_sup) + done + subsubsection {* Binary intersection -- Int *} +definition inter :: "'a set \ 'a set \ 'a set" (infixl "Int" 70) where + inf_set_eq [symmetric]: "A Int B = inf A B" + +notation (xsymbols) + inter (infixl "\" 70) + +notation (HTML output) + inter (infixl "\" 70) + +lemma Int_def: + "A \ B = {x. x \ A \ x \ B}" + by (simp add: inf_fun_eq inf_bool_eq inf_set_eq [symmetric] Collect_def mem_def) + lemma Int_iff [simp]: "(c : A Int B) = (c:A & c:B)" by (unfold Int_def) blast @@ -748,6 +724,11 @@ lemma IntE [elim!]: "c : A Int B ==> (c:A ==> c:B ==> P) ==> P" by simp +lemma mono_Int: "mono f \ f (A \ B) \ f A \ f B" + apply (fold inf_set_eq) + apply (erule mono_inf) + done + subsubsection {* Set difference *} @@ -846,98 +827,18 @@ by (blast elim: equalityE) -subsubsection {* Unions of families *} +subsubsection {* Image of a set under a function *} text {* - @{term [source] "UN x:A. B x"} is @{term "Union (B`A)"}. + Frequently @{term b} does not have the syntactic form of @{term "f x"}. *} -declare UNION_def [noatp] - -lemma UN_iff [simp]: "(b: (UN x:A. B x)) = (EX x:A. b: B x)" - by (unfold UNION_def) blast - -lemma UN_I [intro]: "a:A ==> b: B a ==> b: (UN x:A. B x)" - -- {* The order of the premises presupposes that @{term A} is rigid; - @{term b} may be flexible. *} - by auto - -lemma UN_E [elim!]: "b : (UN x:A. B x) ==> (!!x. x:A ==> b: B x ==> R) ==> R" - by (unfold UNION_def) blast - -lemma UN_cong [cong]: - "A = B ==> (!!x. x:B ==> C x = D x) ==> (UN x:A. C x) = (UN x:B. D x)" - by (simp add: UNION_def) - -lemma strong_UN_cong: - "A = B ==> (!!x. x:B =simp=> C x = D x) ==> (UN x:A. C x) = (UN x:B. D x)" - by (simp add: UNION_def simp_implies_def) - - -subsubsection {* Intersections of families *} - -text {* @{term [source] "INT x:A. B x"} is @{term "Inter (B`A)"}. *} - -lemma INT_iff [simp]: "(b: (INT x:A. B x)) = (ALL x:A. b: B x)" - by (unfold INTER_def) blast - -lemma INT_I [intro!]: "(!!x. x:A ==> b: B x) ==> b : (INT x:A. B x)" - by (unfold INTER_def) blast - -lemma INT_D [elim]: "b : (INT x:A. B x) ==> a:A ==> b: B a" - by auto - -lemma INT_E [elim]: "b : (INT x:A. B x) ==> (b: B a ==> R) ==> (a~:A ==> R) ==> R" - -- {* "Classical" elimination -- by the Excluded Middle on @{prop "a:A"}. *} - by (unfold INTER_def) blast - -lemma INT_cong [cong]: - "A = B ==> (!!x. x:B ==> C x = D x) ==> (INT x:A. C x) = (INT x:B. D x)" - by (simp add: INTER_def) +definition image :: "('a => 'b) => 'a set => 'b set" (infixr "`" 90) where + image_def [noatp]: "f ` A = {y. EX x:A. y = f(x)}" - -subsubsection {* Union *} - -lemma Union_iff [simp,noatp]: "(A : Union C) = (EX X:C. A:X)" - by (unfold Union_def) blast - -lemma UnionI [intro]: "X:C ==> A:X ==> A : Union C" - -- {* The order of the premises presupposes that @{term C} is rigid; - @{term A} may be flexible. *} - by auto - -lemma UnionE [elim!]: "A : Union C ==> (!!X. A:X ==> X:C ==> R) ==> R" - by (unfold Union_def) blast - - -subsubsection {* Inter *} - -lemma Inter_iff [simp,noatp]: "(A : Inter C) = (ALL X:C. A:X)" - by (unfold Inter_def) blast - -lemma InterI [intro!]: "(!!X. X:C ==> A:X) ==> A : Inter C" - by (simp add: Inter_def) - -text {* - \medskip A ``destruct'' rule -- every @{term X} in @{term C} - contains @{term A} as an element, but @{prop "A:X"} can hold when - @{prop "X:C"} does not! This rule is analogous to @{text spec}. -*} - -lemma InterD [elim]: "A : Inter C ==> X:C ==> A:X" - by auto - -lemma InterE [elim]: "A : Inter C ==> (X~:C ==> R) ==> (A:X ==> R) ==> R" - -- {* ``Classical'' elimination rule -- does not require proving - @{prop "X:C"}. *} - by (unfold Inter_def) blast - -text {* - \medskip Image of a set under a function. Frequently @{term b} does - not have the syntactic form of @{term "f x"}. -*} - -declare image_def [noatp] +abbreviation + range :: "('a => 'b) => 'b set" where -- "of function" + "range f == f ` UNIV" lemma image_eqI [simp, intro]: "b = f x ==> x:A ==> b : f`A" by (unfold image_def) blast @@ -958,9 +859,6 @@ lemma image_Un: "f`(A Un B) = f`A Un f`B" by blast -lemma image_eq_UN: "f`A = (UN x:A. {f x})" - by blast - lemma image_iff: "(z : f`A) = (EX x:A. z = f x)" by blast @@ -993,34 +891,15 @@ by blast -subsubsection {* Set reasoning tools *} +subsubsection {* Some rules with @{text "if"} *} text{* Elimination of @{text"{x. \ & x=t & \}"}. *} lemma Collect_conv_if: "{x. x=a & P x} = (if P a then {a} else {})" -by auto + by auto lemma Collect_conv_if2: "{x. a=x & P x} = (if P a then {a} else {})" -by auto - -text {* -Simproc for pulling @{text "x=t"} in @{text "{x. \ & x=t & \}"} -to the front (and similarly for @{text "t=x"}): -*} - -ML{* - local - val Coll_perm_tac = rtac @{thm Collect_cong} 1 THEN rtac @{thm iffI} 1 THEN - ALLGOALS(EVERY'[REPEAT_DETERM o (etac @{thm conjE}), - DEPTH_SOLVE_1 o (ares_tac [@{thm conjI}])]) - in - val defColl_regroup = Simplifier.simproc (the_context ()) - "defined Collect" ["{x. P x & Q x}"] - (Quantifier1.rearrange_Coll Coll_perm_tac) - end; - - Addsimprocs [defColl_regroup]; -*} + by auto text {* Rewrite rules for boolean case-splitting: faster than @{text @@ -1055,13 +934,8 @@ ("Collect", [CollectD]), ("Inter", [InterD]), ("INTER", [INT_D])] *) -ML {* - val mksimps_pairs = [(@{const_name Ball}, @{thms bspec})] @ mksimps_pairs; -*} -declaration {* fn _ => - Simplifier.map_ss (fn ss => ss setmksimps (mksimps mksimps_pairs)) -*} +subsection {* Further operations and lemmas *} subsubsection {* The ``proper subset'' relation *} @@ -1108,9 +982,6 @@ lemmas [symmetric, rulify] = atomize_ball and [symmetric, defn] = atomize_ball - -subsection {* Further set-theory lemmas *} - subsubsection {* Derived rules involving subsets. *} text {* @{text insert}. *} @@ -1125,43 +996,6 @@ by blast -text {* \medskip Big Union -- least upper bound of a set. *} - -lemma Union_upper: "B \ A ==> B \ Union A" - by (iprover intro: subsetI UnionI) - -lemma Union_least: "(!!X. X \ A ==> X \ C) ==> Union A \ C" - by (iprover intro: subsetI elim: UnionE dest: subsetD) - - -text {* \medskip General union. *} - -lemma UN_upper: "a \ A ==> B a \ (\x\A. B x)" - by blast - -lemma UN_least: "(!!x. x \ A ==> B x \ C) ==> (\x\A. B x) \ C" - by (iprover intro: subsetI elim: UN_E dest: subsetD) - - -text {* \medskip Big Intersection -- greatest lower bound of a set. *} - -lemma Inter_lower: "B \ A ==> Inter A \ B" - by blast - -lemma Inter_subset: - "[| !!X. X \ A ==> X \ B; A ~= {} |] ==> \A \ B" - by blast - -lemma Inter_greatest: "(!!X. X \ A ==> C \ X) ==> C \ Inter A" - by (iprover intro: InterI subsetI dest: subsetD) - -lemma INT_lower: "a \ A ==> (\x\A. B x) \ B a" - by blast - -lemma INT_greatest: "(!!x. x \ A ==> C \ B x) ==> C \ (\x\A. B x)" - by (iprover intro: INT_I subsetI dest: subsetD) - - text {* \medskip Finite Union -- the least upper bound of two sets. *} lemma Un_upper1: "A \ A \ B" @@ -1227,18 +1061,6 @@ lemma Collect_conj_eq: "{x. P x & Q x} = {x. P x} \ {x. Q x}" by blast -lemma Collect_all_eq: "{x. \y. P x y} = (\y. {x. P x y})" - by blast - -lemma Collect_ball_eq: "{x. \y\A. P x y} = (\y\A. {x. P x y})" - by blast - -lemma Collect_ex_eq [noatp]: "{x. \y. P x y} = (\y. {x. P x y})" - by blast - -lemma Collect_bex_eq [noatp]: "{x. \y\A. P x y} = (\y\A. {x. P x y})" - by blast - text {* \medskip @{text insert}. *} @@ -1274,9 +1096,6 @@ lemma insert_Collect: "insert a (Collect P) = {u. u \ a --> P u}" by auto -lemma UN_insert_distrib: "u \ A ==> (\x\A. insert a (B x)) = insert a (\x\A. B x)" - by blast - lemma insert_inter_insert[simp]: "insert a A \ insert a B = insert a (A \ B)" by blast @@ -1385,9 +1204,6 @@ lemma Int_UNIV_right [simp]: "A \ UNIV = A" by blast -lemma Int_eq_Inter: "A \ B = \{A, B}" - by blast - lemma Int_Un_distrib: "A \ (B \ C) = (A \ B) \ (A \ C)" by blast @@ -1442,9 +1258,6 @@ lemma Un_UNIV_right [simp]: "A \ UNIV = UNIV" by blast -lemma Un_eq_Union: "A \ B = \{A, B}" - by blast - lemma Un_insert_left [simp]: "(insert a B) \ C = insert a (B \ C)" by blast @@ -1508,12 +1321,6 @@ lemma Compl_Int [simp]: "-(A \ B) = (-A) \ (-B)" by blast -lemma Compl_UN [simp]: "-(\x\A. B x) = (\x\A. -B x)" - by blast - -lemma Compl_INT [simp]: "-(\x\A. B x) = (\x\A. -B x)" - by blast - lemma subset_Compl_self_eq: "(A \ -A) = (A = {})" by blast @@ -1533,179 +1340,6 @@ lemma Compl_eq_Compl_iff [iff]: "(-A = -B) = (A = (B::'a set))" by blast - -text {* \medskip @{text Union}. *} - -lemma Union_empty [simp]: "Union({}) = {}" - by blast - -lemma Union_UNIV [simp]: "Union UNIV = UNIV" - by blast - -lemma Union_insert [simp]: "Union (insert a B) = a \ \B" - by blast - -lemma Union_Un_distrib [simp]: "\(A Un B) = \A \ \B" - by blast - -lemma Union_Int_subset: "\(A \ B) \ \A \ \B" - by blast - -lemma Union_empty_conv [simp,noatp]: "(\A = {}) = (\x\A. x = {})" - by blast - -lemma empty_Union_conv [simp,noatp]: "({} = \A) = (\x\A. x = {})" - by blast - -lemma Union_disjoint: "(\C \ A = {}) = (\B\C. B \ A = {})" - by blast - - -text {* \medskip @{text Inter}. *} - -lemma Inter_empty [simp]: "\{} = UNIV" - by blast - -lemma Inter_UNIV [simp]: "\UNIV = {}" - by blast - -lemma Inter_insert [simp]: "\(insert a B) = a \ \B" - by blast - -lemma Inter_Un_subset: "\A \ \B \ \(A \ B)" - by blast - -lemma Inter_Un_distrib: "\(A \ B) = \A \ \B" - by blast - -lemma Inter_UNIV_conv [simp,noatp]: - "(\A = UNIV) = (\x\A. x = UNIV)" - "(UNIV = \A) = (\x\A. x = UNIV)" - by blast+ - - -text {* - \medskip @{text UN} and @{text INT}. - - Basic identities: *} - -lemma UN_empty [simp,noatp]: "(\x\{}. B x) = {}" - by blast - -lemma UN_empty2 [simp]: "(\x\A. {}) = {}" - by blast - -lemma UN_singleton [simp]: "(\x\A. {x}) = A" - by blast - -lemma UN_absorb: "k \ I ==> A k \ (\i\I. A i) = (\i\I. A i)" - by auto - -lemma INT_empty [simp]: "(\x\{}. B x) = UNIV" - by blast - -lemma INT_absorb: "k \ I ==> A k \ (\i\I. A i) = (\i\I. A i)" - by blast - -lemma UN_insert [simp]: "(\x\insert a A. B x) = B a \ UNION A B" - by blast - -lemma UN_Un[simp]: "(\i \ A \ B. M i) = (\i\A. M i) \ (\i\B. M i)" - by blast - -lemma UN_UN_flatten: "(\x \ (\y\A. B y). C x) = (\y\A. \x\B y. C x)" - by blast - -lemma UN_subset_iff: "((\i\I. A i) \ B) = (\i\I. A i \ B)" - by blast - -lemma INT_subset_iff: "(B \ (\i\I. A i)) = (\i\I. B \ A i)" - by blast - -lemma INT_insert [simp]: "(\x \ insert a A. B x) = B a \ INTER A B" - by blast - -lemma INT_Un: "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" - by blast - -lemma INT_insert_distrib: - "u \ A ==> (\x\A. insert a (B x)) = insert a (\x\A. B x)" - by blast - -lemma Union_image_eq [simp]: "\(B`A) = (\x\A. B x)" - by blast - -lemma image_Union: "f ` \S = (\x\S. f ` x)" - by blast - -lemma Inter_image_eq [simp]: "\(B`A) = (\x\A. B x)" - by blast - -lemma UN_constant [simp]: "(\y\A. c) = (if A = {} then {} else c)" - by auto - -lemma INT_constant [simp]: "(\y\A. c) = (if A = {} then UNIV else c)" - by auto - -lemma UN_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" - by blast - -lemma INT_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" - -- {* Look: it has an \emph{existential} quantifier *} - by blast - -lemma UNION_empty_conv[simp]: - "({} = (UN x:A. B x)) = (\x\A. B x = {})" - "((UN x:A. B x) = {}) = (\x\A. B x = {})" -by blast+ - -lemma INTER_UNIV_conv[simp]: - "(UNIV = (INT x:A. B x)) = (\x\A. B x = UNIV)" - "((INT x:A. B x) = UNIV) = (\x\A. B x = UNIV)" -by blast+ - - -text {* \medskip Distributive laws: *} - -lemma Int_Union: "A \ \B = (\C\B. A \ C)" - by blast - -lemma Int_Union2: "\B \ A = (\C\B. C \ A)" - by blast - -lemma Un_Union_image: "(\x\C. A x \ B x) = \(A`C) \ \(B`C)" - -- {* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: *} - -- {* Union of a family of unions *} - by blast - -lemma UN_Un_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" - -- {* Equivalent version *} - by blast - -lemma Un_Inter: "A \ \B = (\C\B. A \ C)" - by blast - -lemma Int_Inter_image: "(\x\C. A x \ B x) = \(A`C) \ \(B`C)" - by blast - -lemma INT_Int_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" - -- {* Equivalent version *} - by blast - -lemma Int_UN_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" - -- {* Halmos, Naive Set Theory, page 35. *} - by blast - -lemma Un_INT_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" - by blast - -lemma Int_UN_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" - by blast - -lemma Un_INT_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" - by blast - - text {* \medskip Bounded quantifiers. The following are not added to the default simpset because @@ -1717,12 +1351,6 @@ lemma bex_Un: "(\x \ A \ B. P x) = ((\x\A. P x) | (\x\B. P x))" by blast -lemma ball_UN: "(\z \ UNION A B. P z) = (\x\A. \z \ B x. P z)" - by blast - -lemma bex_UN: "(\z \ UNION A B. P z) = (\x\A. \z\B x. P z)" - by blast - text {* \medskip Set difference. *} @@ -1830,15 +1458,6 @@ lemma ex_bool_eq: "(\b. P b) \ P True \ P False" by (auto intro: bool_contrapos) -lemma Un_eq_UN: "A \ B = (\b. if b then A else B)" - by (auto simp add: split_if_mem2) - -lemma UN_bool_eq: "(\b::bool. A b) = (A True \ A False)" - by (auto intro: bool_contrapos) - -lemma INT_bool_eq: "(\b::bool. A b) = (A True \ A False)" - by (auto intro: bool_induct) - text {* \medskip @{text Pow} *} lemma Pow_empty [simp]: "Pow {} = {{}}" @@ -1856,21 +1475,9 @@ lemma Un_Pow_subset: "Pow A \ Pow B \ Pow (A \ B)" by blast -lemma UN_Pow_subset: "(\x\A. Pow (B x)) \ Pow (\x\A. B x)" - by blast - -lemma subset_Pow_Union: "A \ Pow (\A)" - by blast - -lemma Union_Pow_eq [simp]: "\(Pow A) = A" - by blast - lemma Pow_Int_eq [simp]: "Pow (A \ B) = Pow A \ Pow B" by blast -lemma Pow_INT_eq: "Pow (\x\A. B x) = (\x\A. Pow (B x))" - by blast - text {* \medskip Miscellany. *} @@ -1893,101 +1500,6 @@ by iprover -text {* \medskip Miniscoping: pushing in quantifiers and big Unions - and Intersections. *} - -lemma UN_simps [simp]: - "!!a B C. (UN x:C. insert a (B x)) = (if C={} then {} else insert a (UN x:C. B x))" - "!!A B C. (UN x:C. A x Un B) = ((if C={} then {} else (UN x:C. A x) Un B))" - "!!A B C. (UN x:C. A Un B x) = ((if C={} then {} else A Un (UN x:C. B x)))" - "!!A B C. (UN x:C. A x Int B) = ((UN x:C. A x) Int B)" - "!!A B C. (UN x:C. A Int B x) = (A Int (UN x:C. B x))" - "!!A B C. (UN x:C. A x - B) = ((UN x:C. A x) - B)" - "!!A B C. (UN x:C. A - B x) = (A - (INT x:C. B x))" - "!!A B. (UN x: Union A. B x) = (UN y:A. UN x:y. B x)" - "!!A B C. (UN z: UNION A B. C z) = (UN x:A. UN z: B(x). C z)" - "!!A B f. (UN x:f`A. B x) = (UN a:A. B (f a))" - by auto - -lemma INT_simps [simp]: - "!!A B C. (INT x:C. A x Int B) = (if C={} then UNIV else (INT x:C. A x) Int B)" - "!!A B C. (INT x:C. A Int B x) = (if C={} then UNIV else A Int (INT x:C. B x))" - "!!A B C. (INT x:C. A x - B) = (if C={} then UNIV else (INT x:C. A x) - B)" - "!!A B C. (INT x:C. A - B x) = (if C={} then UNIV else A - (UN x:C. B x))" - "!!a B C. (INT x:C. insert a (B x)) = insert a (INT x:C. B x)" - "!!A B C. (INT x:C. A x Un B) = ((INT x:C. A x) Un B)" - "!!A B C. (INT x:C. A Un B x) = (A Un (INT x:C. B x))" - "!!A B. (INT x: Union A. B x) = (INT y:A. INT x:y. B x)" - "!!A B C. (INT z: UNION A B. C z) = (INT x:A. INT z: B(x). C z)" - "!!A B f. (INT x:f`A. B x) = (INT a:A. B (f a))" - by auto - -lemma ball_simps [simp,noatp]: - "!!A P Q. (ALL x:A. P x | Q) = ((ALL x:A. P x) | Q)" - "!!A P Q. (ALL x:A. P | Q x) = (P | (ALL x:A. Q x))" - "!!A P Q. (ALL x:A. P --> Q x) = (P --> (ALL x:A. Q x))" - "!!A P Q. (ALL x:A. P x --> Q) = ((EX x:A. P x) --> Q)" - "!!P. (ALL x:{}. P x) = True" - "!!P. (ALL x:UNIV. P x) = (ALL x. P x)" - "!!a B P. (ALL x:insert a B. P x) = (P a & (ALL x:B. P x))" - "!!A P. (ALL x:Union A. P x) = (ALL y:A. ALL x:y. P x)" - "!!A B P. (ALL x: UNION A B. P x) = (ALL a:A. ALL x: B a. P x)" - "!!P Q. (ALL x:Collect Q. P x) = (ALL x. Q x --> P x)" - "!!A P f. (ALL x:f`A. P x) = (ALL x:A. P (f x))" - "!!A P. (~(ALL x:A. P x)) = (EX x:A. ~P x)" - by auto - -lemma bex_simps [simp,noatp]: - "!!A P Q. (EX x:A. P x & Q) = ((EX x:A. P x) & Q)" - "!!A P Q. (EX x:A. P & Q x) = (P & (EX x:A. Q x))" - "!!P. (EX x:{}. P x) = False" - "!!P. (EX x:UNIV. P x) = (EX x. P x)" - "!!a B P. (EX x:insert a B. P x) = (P(a) | (EX x:B. P x))" - "!!A P. (EX x:Union A. P x) = (EX y:A. EX x:y. P x)" - "!!A B P. (EX x: UNION A B. P x) = (EX a:A. EX x:B a. P x)" - "!!P Q. (EX x:Collect Q. P x) = (EX x. Q x & P x)" - "!!A P f. (EX x:f`A. P x) = (EX x:A. P (f x))" - "!!A P. (~(EX x:A. P x)) = (ALL x:A. ~P x)" - by auto - -lemma ball_conj_distrib: - "(ALL x:A. P x & Q x) = ((ALL x:A. P x) & (ALL x:A. Q x))" - by blast - -lemma bex_disj_distrib: - "(EX x:A. P x | Q x) = ((EX x:A. P x) | (EX x:A. Q x))" - by blast - - -text {* \medskip Maxiscoping: pulling out big Unions and Intersections. *} - -lemma UN_extend_simps: - "!!a B C. insert a (UN x:C. B x) = (if C={} then {a} else (UN x:C. insert a (B x)))" - "!!A B C. (UN x:C. A x) Un B = (if C={} then B else (UN x:C. A x Un B))" - "!!A B C. A Un (UN x:C. B x) = (if C={} then A else (UN x:C. A Un B x))" - "!!A B C. ((UN x:C. A x) Int B) = (UN x:C. A x Int B)" - "!!A B C. (A Int (UN x:C. B x)) = (UN x:C. A Int B x)" - "!!A B C. ((UN x:C. A x) - B) = (UN x:C. A x - B)" - "!!A B C. (A - (INT x:C. B x)) = (UN x:C. A - B x)" - "!!A B. (UN y:A. UN x:y. B x) = (UN x: Union A. B x)" - "!!A B C. (UN x:A. UN z: B(x). C z) = (UN z: UNION A B. C z)" - "!!A B f. (UN a:A. B (f a)) = (UN x:f`A. B x)" - by auto - -lemma INT_extend_simps: - "!!A B C. (INT x:C. A x) Int B = (if C={} then B else (INT x:C. A x Int B))" - "!!A B C. A Int (INT x:C. B x) = (if C={} then A else (INT x:C. A Int B x))" - "!!A B C. (INT x:C. A x) - B = (if C={} then UNIV-B else (INT x:C. A x - B))" - "!!A B C. A - (UN x:C. B x) = (if C={} then A else (INT x:C. A - B x))" - "!!a B C. insert a (INT x:C. B x) = (INT x:C. insert a (B x))" - "!!A B C. ((INT x:C. A x) Un B) = (INT x:C. A x Un B)" - "!!A B C. A Un (INT x:C. B x) = (INT x:C. A Un B x)" - "!!A B. (INT y:A. INT x:y. B x) = (INT x: Union A. B x)" - "!!A B C. (INT x:A. INT z: B(x). C z) = (INT z: UNION A B. C z)" - "!!A B f. (INT a:A. B (f a)) = (INT x:f`A. B x)" - by auto - - subsubsection {* Monotonicity of various operations *} lemma image_mono: "A \ B ==> f`A \ f`B" @@ -1996,23 +1508,6 @@ lemma Pow_mono: "A \ B ==> Pow A \ Pow B" by blast -lemma Union_mono: "A \ B ==> \A \ \B" - by blast - -lemma Inter_anti_mono: "B \ A ==> \A \ \B" - by blast - -lemma UN_mono: - "A \ B ==> (!!x. x \ A ==> f x \ g x) ==> - (\x\A. f x) \ (\x\B. g x)" - by (blast dest: subsetD) - -lemma INT_anti_mono: - "B \ A ==> (!!x. x \ A ==> f x \ g x) ==> - (\x\A. f x) \ (\x\A. g x)" - -- {* The last inclusion is POSITIVE! *} - by (blast dest: subsetD) - lemma insert_mono: "C \ D ==> insert a C \ insert a D" by blast @@ -2070,15 +1565,12 @@ by iprover -subsection {* Inverse image of a function *} +subsubsection {* Inverse image of a function *} constdefs vimage :: "('a => 'b) => 'b set => 'a set" (infixr "-`" 90) [code del]: "f -` B == {x. f x : B}" - -subsubsection {* Basic rules *} - lemma vimage_eq [simp]: "(a : f -` B) = (f a : B)" by (unfold vimage_def) blast @@ -2097,9 +1589,6 @@ lemma vimageD: "a : f -` A ==> f a : A" by (unfold vimage_def) fast - -subsubsection {* Equations *} - lemma vimage_empty [simp]: "f -` {} = {}" by blast @@ -2112,15 +1601,6 @@ lemma vimage_Int [simp]: "f -` (A Int B) = (f -` A) Int (f -` B)" by fast -lemma vimage_Union: "f -` (Union A) = (UN X:A. f -` X)" - by blast - -lemma vimage_UN: "f-`(UN x:A. B x) = (UN x:A. f -` B x)" - by blast - -lemma vimage_INT: "f-`(INT x:A. B x) = (INT x:A. f -` B x)" - by blast - lemma vimage_Collect_eq [simp]: "f -` Collect P = {y. P (f y)}" by blast @@ -2137,10 +1617,6 @@ lemma vimage_UNIV [simp]: "f -` UNIV = UNIV" by blast -lemma vimage_eq_UN: "f-`B = (UN y: B. f-`{y})" - -- {* NOT suitable for rewriting *} - by blast - lemma vimage_mono: "A \ B ==> f -` A \ f -` B" -- {* monotonicity *} by blast @@ -2160,11 +1636,8 @@ lemma image_diff_subset: "f`A - f`B <= f`(A - B)" by blast -lemma image_UN: "(f ` (UNION A B)) = (UN x:A.(f ` (B x)))" -by blast - -subsection {* Getting the Contents of a Singleton Set *} +subsubsection {* Getting the Contents of a Singleton Set *} definition contents :: "'a set \ 'a" where [code del]: "contents X = (THE x. X = {x})" @@ -2173,19 +1646,7 @@ by (simp add: contents_def) -subsection {* Transitivity rules for calculational reasoning *} - -lemma set_rev_mp: "x:A ==> A \ B ==> x:B" - by (rule subsetD) - -lemma set_mp: "A \ B ==> x:A ==> x:B" - by (rule subsetD) - -lemmas basic_trans_rules [trans] = - order_trans_rules set_rev_mp set_mp - - -subsection {* Least value operator *} +subsubsection {* Least value operator *} lemma Least_mono: "mono (f::'a::order => 'b::order) ==> EX x:S. ALL y:S. x <= y @@ -2197,285 +1658,20 @@ apply (auto elim: monoD intro!: order_antisym) done - -subsection {* Rudimentary code generation *} +subsection {* Misc *} -lemma empty_code [code]: "{} x \ False" - unfolding empty_def Collect_def .. - -lemma UNIV_code [code]: "UNIV x \ True" - unfolding UNIV_def Collect_def .. +text {* Rudimentary code generation *} lemma insert_code [code]: "insert y A x \ y = x \ A x" - unfolding insert_def Collect_def mem_def Un_def by auto - -lemma inter_code [code]: "(A \ B) x \ A x \ B x" - unfolding Int_def Collect_def mem_def .. - -lemma union_code [code]: "(A \ B) x \ A x \ B x" - unfolding Un_def Collect_def mem_def .. + by (auto simp add: insert_compr Collect_def mem_def) lemma vimage_code [code]: "(f -` A) x = A (f x)" - unfolding vimage_def Collect_def mem_def .. + by (simp add: vimage_def Collect_def mem_def) -subsection {* Complete lattices *} - -notation - less_eq (infix "\" 50) and - less (infix "\" 50) and - inf (infixl "\" 70) and - sup (infixl "\" 65) - -class complete_lattice = lattice + bot + top + - fixes Inf :: "'a set \ 'a" ("\_" [900] 900) - and Sup :: "'a set \ 'a" ("\_" [900] 900) - assumes Inf_lower: "x \ A \ \A \ x" - and Inf_greatest: "(\x. x \ A \ z \ x) \ z \ \A" - assumes Sup_upper: "x \ A \ x \ \A" - and Sup_least: "(\x. x \ A \ x \ z) \ \A \ z" -begin - -lemma Inf_Sup: "\A = \{b. \a \ A. b \ a}" - by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) - -lemma Sup_Inf: "\A = \{b. \a \ A. a \ b}" - by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) - -lemma Inf_Univ: "\UNIV = \{}" - unfolding Sup_Inf by auto - -lemma Sup_Univ: "\UNIV = \{}" - unfolding Inf_Sup by auto - -lemma Inf_insert: "\insert a A = a \ \A" - by (auto intro: antisym Inf_greatest Inf_lower) - -lemma Sup_insert: "\insert a A = a \ \A" - by (auto intro: antisym Sup_least Sup_upper) - -lemma Inf_singleton [simp]: - "\{a} = a" - by (auto intro: antisym Inf_lower Inf_greatest) - -lemma Sup_singleton [simp]: - "\{a} = a" - by (auto intro: antisym Sup_upper Sup_least) - -lemma Inf_insert_simp: - "\insert a A = (if A = {} then a else a \ \A)" - by (cases "A = {}") (simp_all, simp add: Inf_insert) - -lemma Sup_insert_simp: - "\insert a A = (if A = {} then a else a \ \A)" - by (cases "A = {}") (simp_all, simp add: Sup_insert) - -lemma Inf_binary: - "\{a, b} = a \ b" - by (simp add: Inf_insert_simp) - -lemma Sup_binary: - "\{a, b} = a \ b" - by (simp add: Sup_insert_simp) - -lemma bot_def: - "bot = \{}" - by (auto intro: antisym Sup_least) - -lemma top_def: - "top = \{}" - by (auto intro: antisym Inf_greatest) - -lemma sup_bot [simp]: - "x \ bot = x" - using bot_least [of x] by (simp add: le_iff_sup sup_commute) - -lemma inf_top [simp]: - "x \ top = x" - using top_greatest [of x] by (simp add: le_iff_inf inf_commute) - -definition SUPR :: "'b set \ ('b \ 'a) \ 'a" where - "SUPR A f == \ (f ` A)" - -definition INFI :: "'b set \ ('b \ 'a) \ 'a" where - "INFI A f == \ (f ` A)" - -end - -syntax - "_SUP1" :: "pttrns => 'b => 'b" ("(3SUP _./ _)" [0, 10] 10) - "_SUP" :: "pttrn => 'a set => 'b => 'b" ("(3SUP _:_./ _)" [0, 10] 10) - "_INF1" :: "pttrns => 'b => 'b" ("(3INF _./ _)" [0, 10] 10) - "_INF" :: "pttrn => 'a set => 'b => 'b" ("(3INF _:_./ _)" [0, 10] 10) - -translations - "SUP x y. B" == "SUP x. SUP y. B" - "SUP x. B" == "CONST SUPR CONST UNIV (%x. B)" - "SUP x. B" == "SUP x:CONST UNIV. B" - "SUP x:A. B" == "CONST SUPR A (%x. B)" - "INF x y. B" == "INF x. INF y. B" - "INF x. B" == "CONST INFI CONST UNIV (%x. B)" - "INF x. B" == "INF x:CONST UNIV. B" - "INF x:A. B" == "CONST INFI A (%x. B)" - -(* To avoid eta-contraction of body: *) -print_translation {* -let - fun btr' syn (A :: Abs abs :: ts) = - let val (x,t) = atomic_abs_tr' abs - in list_comb (Syntax.const syn $ x $ A $ t, ts) end - val const_syntax_name = Sign.const_syntax_name @{theory} o fst o dest_Const -in -[(const_syntax_name @{term SUPR}, btr' "_SUP"),(const_syntax_name @{term "INFI"}, btr' "_INF")] -end -*} - -context complete_lattice -begin - -lemma le_SUPI: "i : A \ M i \ (SUP i:A. M i)" - by (auto simp add: SUPR_def intro: Sup_upper) - -lemma SUP_leI: "(\i. i : A \ M i \ u) \ (SUP i:A. M i) \ u" - by (auto simp add: SUPR_def intro: Sup_least) - -lemma INF_leI: "i : A \ (INF i:A. M i) \ M i" - by (auto simp add: INFI_def intro: Inf_lower) - -lemma le_INFI: "(\i. i : A \ u \ M i) \ u \ (INF i:A. M i)" - by (auto simp add: INFI_def intro: Inf_greatest) - -lemma SUP_const[simp]: "A \ {} \ (SUP i:A. M) = M" - by (auto intro: antisym SUP_leI le_SUPI) - -lemma INF_const[simp]: "A \ {} \ (INF i:A. M) = M" - by (auto intro: antisym INF_leI le_INFI) - -end - - -subsection {* Bool as complete lattice *} - -instantiation bool :: complete_lattice -begin - -definition - Inf_bool_def: "\A \ (\x\A. x)" - -definition - Sup_bool_def: "\A \ (\x\A. x)" - -instance - by intro_classes (auto simp add: Inf_bool_def Sup_bool_def le_bool_def) - -end - -lemma Inf_empty_bool [simp]: - "\{}" - unfolding Inf_bool_def by auto - -lemma not_Sup_empty_bool [simp]: - "\ \{}" - unfolding Sup_bool_def by auto - - -subsection {* Fun as complete lattice *} - -instantiation "fun" :: (type, complete_lattice) complete_lattice -begin - -definition - Inf_fun_def [code del]: "\A = (\x. \{y. \f\A. y = f x})" - -definition - Sup_fun_def [code del]: "\A = (\x. \{y. \f\A. y = f x})" - -instance - by intro_classes - (auto simp add: Inf_fun_def Sup_fun_def le_fun_def - intro: Inf_lower Sup_upper Inf_greatest Sup_least) - -end - -lemma Inf_empty_fun: - "\{} = (\_. \{})" - by rule (auto simp add: Inf_fun_def) - -lemma Sup_empty_fun: - "\{} = (\_. \{})" - by rule (auto simp add: Sup_fun_def) - - -subsection {* Set as lattice *} - -lemma inf_set_eq: "A \ B = A \ B" - apply (rule subset_antisym) - apply (rule Int_greatest) - apply (rule inf_le1) - apply (rule inf_le2) - apply (rule inf_greatest) - apply (rule Int_lower1) - apply (rule Int_lower2) - done - -lemma sup_set_eq: "A \ B = A \ B" - apply (rule subset_antisym) - apply (rule sup_least) - apply (rule Un_upper1) - apply (rule Un_upper2) - apply (rule Un_least) - apply (rule sup_ge1) - apply (rule sup_ge2) - done - -lemma mono_Int: "mono f \ f (A \ B) \ f A \ f B" - apply (fold inf_set_eq sup_set_eq) - apply (erule mono_inf) - done - -lemma mono_Un: "mono f \ f A \ f B \ f (A \ B)" - apply (fold inf_set_eq sup_set_eq) - apply (erule mono_sup) - done - -lemma Inf_set_eq: "\S = \S" - apply (rule subset_antisym) - apply (rule Inter_greatest) - apply (erule Inf_lower) - apply (rule Inf_greatest) - apply (erule Inter_lower) - done - -lemma Sup_set_eq: "\S = \S" - apply (rule subset_antisym) - apply (rule Sup_least) - apply (erule Union_upper) - apply (rule Union_least) - apply (erule Sup_upper) - done - -lemma top_set_eq: "top = UNIV" - by (iprover intro!: subset_antisym subset_UNIV top_greatest) - -lemma bot_set_eq: "bot = {}" - by (iprover intro!: subset_antisym empty_subsetI bot_least) - -no_notation - less_eq (infix "\" 50) and - less (infix "\" 50) and - inf (infixl "\" 70) and - sup (infixl "\" 65) and - Inf ("\_" [900] 900) and - Sup ("\_" [900] 900) - - -subsection {* Misc theorem and ML bindings *} +text {* Misc theorem and ML bindings *} lemmas equalityI = subset_antisym -lemmas mem_simps = - insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff - mem_Collect_eq UN_iff Union_iff INT_iff Inter_iff - -- {* Each of these has ALREADY been added @{text "[simp]"} above. *} ML {* val Ball_def = @{thm Ball_def} diff -r 63686057cbe8 -r abda97d2deea src/HOL/SetInterval.thy --- a/src/HOL/SetInterval.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/SetInterval.thy Thu Jul 23 21:13:21 2009 +0200 @@ -425,7 +425,7 @@ proof cases assume "finite A" thus "PROP ?P" - proof(induct A rule:finite_linorder_induct) + proof(induct A rule:finite_linorder_max_induct) case empty thus ?case by auto next case (insert A b) diff -r 63686057cbe8 -r abda97d2deea src/HOL/SizeChange/sct.ML --- a/src/HOL/SizeChange/sct.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/SizeChange/sct.ML Thu Jul 23 21:13:21 2009 +0200 @@ -184,7 +184,7 @@ fun setup_probe_goal ctxt domT Dtab Mtab (i, j) = let - val css = local_clasimpset_of ctxt + val css = clasimpset_of ctxt val thy = ProofContext.theory_of ctxt val RD1 = get_elem (Dtab i) val RD2 = get_elem (Dtab j) @@ -269,7 +269,7 @@ fun in_graph_tac ctxt = simp_tac (HOL_basic_ss addsimps has_edge_simps) 1 - THEN (simp_tac (local_simpset_of ctxt) 1) (* FIXME reduce simpset *) + THEN (simp_tac (simpset_of ctxt) 1) (* FIXME reduce simpset *) fun approx_tac _ (NoStep thm) = rtac disjI1 1 THEN rtac thm 1 | approx_tac ctxt (Graph (G, thm)) = @@ -334,7 +334,7 @@ val sound_int_goal = HOLogic.mk_Trueprop (mk_sound_int ACG RDlist mfuns) val tac = - unfold_tac [sound_int_def, len_simp] (local_simpset_of ctxt) + unfold_tac [sound_int_def, len_simp] (simpset_of ctxt) THEN all_less_tac (map (all_less_tac o map (approx_tac ctxt)) parts) in tac (instantiate' [] [SOME (cterm_of thy ACG), SOME (cterm_of thy mfuns)] st) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Statespace/DistinctTreeProver.thy --- a/src/HOL/Statespace/DistinctTreeProver.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Statespace/DistinctTreeProver.thy Thu Jul 23 21:13:21 2009 +0200 @@ -676,7 +676,7 @@ ML {* - val ct' = cterm_of (the_context ()) t'; + val ct' = cterm_of @{theory} t'; *} ML {* @@ -706,7 +706,7 @@ ML {* -val cdist' = cterm_of (the_context ()) dist'; +val cdist' = cterm_of @{theory} dist'; DistinctTreeProver.distinct_implProver (!da) cdist'; *} diff -r 63686057cbe8 -r abda97d2deea src/HOL/Statespace/state_fun.ML --- a/src/HOL/Statespace/state_fun.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Statespace/state_fun.ML Thu Jul 23 21:13:21 2009 +0200 @@ -115,7 +115,7 @@ Context.theory_map (StateFunData.put (lookup_ss,ex_lookup_ss,false)); val lookup_simproc = - Simplifier.simproc (the_context ()) "lookup_simp" ["lookup d n (update d' c m v s)"] + Simplifier.simproc @{theory} "lookup_simp" ["lookup d n (update d' c m v s)"] (fn thy => fn ss => fn t => (case t of (Const ("StateFun.lookup",lT)$destr$n$ (s as Const ("StateFun.update",uT)$_$_$_$_$_)) => @@ -171,7 +171,7 @@ addcongs [thm "block_conj_cong"]); in val update_simproc = - Simplifier.simproc (the_context ()) "update_simp" ["update d c n v s"] + Simplifier.simproc @{theory} "update_simp" ["update d c n v s"] (fn thy => fn ss => fn t => (case t of ((upd as Const ("StateFun.update", uT)) $ d $ c $ n $ v $ s) => let diff -r 63686057cbe8 -r abda97d2deea src/HOL/String.thy --- a/src/HOL/String.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/String.thy Thu Jul 23 21:13:21 2009 +0200 @@ -58,11 +58,11 @@ end *} -lemma char_case_nibble_pair [code, code_inline]: +lemma char_case_nibble_pair [code, code_unfold]: "char_case f = split f o nibble_pair_of_char" by (simp add: expand_fun_eq split: char.split) -lemma char_rec_nibble_pair [code, code_inline]: +lemma char_rec_nibble_pair [code, code_unfold]: "char_rec f = split f o nibble_pair_of_char" unfolding char_case_nibble_pair [symmetric] by (simp add: expand_fun_eq split: char.split) diff -r 63686057cbe8 -r abda97d2deea src/HOL/TLA/Memory/MemoryImplementation.thy --- a/src/HOL/TLA/Memory/MemoryImplementation.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/TLA/Memory/MemoryImplementation.thy Thu Jul 23 21:13:21 2009 +0200 @@ -763,7 +763,7 @@ *) ML {* fun split_idle_tac ctxt simps i = - let val ss = Simplifier.local_simpset_of ctxt in + let val ss = simpset_of ctxt in TRY (rtac @{thm actionI} i) THEN InductTacs.case_tac ctxt "(s,t) |= unchanged (e p, c p, r p, m p, rmhist!p)" i THEN rewrite_goals_tac @{thms action_rews} THEN diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/Datatype/datatype.ML --- a/src/HOL/Tools/Datatype/datatype.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/Datatype/datatype.ML Thu Jul 23 21:13:21 2009 +0200 @@ -552,8 +552,7 @@ val _ = (case fold (curry OldTerm.add_typ_tfree_names) cargs' [] \\ tvs of [] => () | vs => error ("Extra type variables on rhs: " ^ commas vs)) - in (constrs @ [((if #flat_names config then Sign.full_name tmp_thy else - Sign.full_name_path tmp_thy tname') + in (constrs @ [(Sign.full_name_path tmp_thy tname' (Binding.map_name (Syntax.const_name mx') cname), map (dtyp_of_typ new_dts) cargs')], constr_syntax' @ [(cname, mx')], sorts'') diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/Datatype/datatype_abs_proofs.ML --- a/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Thu Jul 23 21:13:21 2009 +0200 @@ -93,7 +93,7 @@ val _ = message config "Constructing primrec combinators ..."; val big_name = space_implode "_" new_type_names; - val thy0 = add_path (#flat_names config) big_name thy; + val thy0 = Sign.add_path big_name thy; val descr' = List.concat descr; val recTs = get_rec_types descr' sorts; @@ -243,7 +243,7 @@ Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T', set $ Free ("x", T) $ Free ("y", T')))))) (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts)) - ||> parent_path (#flat_names config) + ||> Sign.parent_path ||> Theory.checkpoint; @@ -277,7 +277,7 @@ let val _ = message config "Proving characteristic theorems for case combinators ..."; - val thy1 = add_path (#flat_names config) (space_implode "_" new_type_names) thy; + val thy1 = Sign.add_path (space_implode "_" new_type_names) thy; val descr' = List.concat descr; val recTs = get_rec_types descr' sorts; @@ -339,7 +339,7 @@ thy2 |> Context.the_theory o fold (fold Nitpick_Const_Simps.add_thm) case_thms o Context.Theory - |> parent_path (#flat_names config) + |> Sign.parent_path |> store_thmss "cases" new_type_names case_thms |-> (fn thmss => pair (thmss, case_names)) end; diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/Datatype/datatype_aux.ML --- a/src/HOL/Tools/Datatype/datatype_aux.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/Datatype/datatype_aux.ML Thu Jul 23 21:13:21 2009 +0200 @@ -22,9 +22,6 @@ val message : config -> string -> unit - val add_path : bool -> string -> theory -> theory - val parent_path : bool -> theory -> theory - val store_thmss_atts : string -> string list -> attribute list list -> thm list list -> theory -> thm list list * theory val store_thmss : string -> string list -> thm list list -> theory -> thm list list * theory @@ -76,15 +73,12 @@ (* datatype option flags *) -type config = { strict: bool, flat_names: bool, quiet: bool }; +type config = { strict: bool, quiet: bool }; val default_config : config = - { strict = true, flat_names = false, quiet = false }; + { strict = true, quiet = false }; fun message ({ quiet, ...} : config) s = if quiet then () else writeln s; -fun add_path flat_names s = if flat_names then I else Sign.add_path s; -fun parent_path flat_names = if flat_names then I else Sign.parent_path; - (* store theorems in theory *) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/Datatype/datatype_case.ML --- a/src/HOL/Tools/Datatype/datatype_case.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/Datatype/datatype_case.ML Thu Jul 23 21:13:21 2009 +0200 @@ -293,7 +293,7 @@ end; fun make_case tab ctxt = gen_make_case - (match_type (ProofContext.theory_of ctxt)) Envir.subst_TVars fastype_of tab ctxt; + (match_type (ProofContext.theory_of ctxt)) Envir.subst_term_types fastype_of tab ctxt; val make_case_untyped = gen_make_case (K (K Vartab.empty)) (K (Term.map_types (K dummyT))) (K dummyT); diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/Datatype/datatype_realizer.ML --- a/src/HOL/Tools/Datatype/datatype_realizer.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/Datatype/datatype_realizer.ML Thu Jul 23 21:13:21 2009 +0200 @@ -180,7 +180,7 @@ val (rs, prems) = split_list (map (make_casedist_prem T) constrs); val r = Const (case_name, map fastype_of rs ---> T --> rT); - val y = Var (("y", 0), Logic.legacy_varifyT T); + val y = Var (("y", 0), Logic.varifyT T); val y' = Free ("y", T); val thm = OldGoals.prove_goalw_cterm [] (cert (Logic.list_implies (prems, @@ -201,10 +201,10 @@ val P = Var (("P", 0), rT' --> HOLogic.boolT); val prf = forall_intr_prf (y, forall_intr_prf (P, List.foldr (fn ((p, r), prf) => - forall_intr_prf (Logic.legacy_varify r, AbsP ("H", SOME (Logic.varify p), + forall_intr_prf (Logic.varify r, AbsP ("H", SOME (Logic.varify p), prf))) (Proofterm.proof_combP (prf_of thm', map PBound (length prems - 1 downto 0))) (prems ~~ rs))); - val r' = Logic.legacy_varify (Abs ("y", Logic.legacy_varifyT T, + val r' = Logic.varify (Abs ("y", T, list_abs (map dest_Free rs, list_comb (r, map Bound ((length rs - 1 downto 0) @ [length rs]))))); diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/Datatype/datatype_rep_proofs.ML --- a/src/HOL/Tools/Datatype/datatype_rep_proofs.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/Datatype/datatype_rep_proofs.ML Thu Jul 23 21:13:21 2009 +0200 @@ -66,7 +66,7 @@ val descr' = flat descr; val big_name = space_implode "_" new_type_names; - val thy1 = add_path (#flat_names config) big_name thy; + val thy1 = Sign.add_path big_name thy; val big_rec_name = big_name ^ "_rep_set"; val rep_set_names' = (if length descr' = 1 then [big_rec_name] else @@ -187,7 +187,7 @@ (********************************* typedef ********************************) val (typedefs, thy3) = thy2 |> - parent_path (#flat_names config) |> + Sign.parent_path |> fold_map (fn ((((name, mx), tvs), c), name') => Typedef.add_typedef false (SOME (Binding.name name')) (name, tvs, mx) (Collect $ Const (c, UnivT')) NONE @@ -196,7 +196,7 @@ (resolve_tac rep_intrs 1))) (types_syntax ~~ tyvars ~~ (Library.take (length newTs, rep_set_names)) ~~ new_type_names) ||> - add_path (#flat_names config) big_name; + Sign.add_path big_name; (*********************** definition of constructors ***********************) @@ -254,14 +254,14 @@ val cong' = standard (cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong); val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma); val (thy', defs', eqns', _) = Library.foldl ((make_constr_def tname T) (length constrs)) - ((add_path (#flat_names config) tname thy, defs, [], 1), constrs ~~ constr_syntax) + ((Sign.add_path tname thy, defs, [], 1), constrs ~~ constr_syntax) in - (parent_path (#flat_names config) thy', defs', eqns @ [eqns'], + (Sign.parent_path thy', defs', eqns @ [eqns'], rep_congs @ [cong'], dist_lemmas @ [dist]) end; val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) = Library.foldl dt_constr_defs - ((thy3 |> Sign.add_consts_i iso_decls |> parent_path (#flat_names config), [], [], [], []), + ((thy3 |> Sign.add_consts_i iso_decls |> Sign.parent_path, [], [], [], []), hd descr ~~ new_type_names ~~ newTs ~~ constr_syntax); (*********** isomorphisms for new types (introduced by typedef) ***********) @@ -355,7 +355,7 @@ in (thy', char_thms' @ char_thms) end; val (thy5, iso_char_thms) = apfst Theory.checkpoint (List.foldr make_iso_defs - (add_path (#flat_names config) big_name thy4, []) (tl descr)); + (Sign.add_path big_name thy4, []) (tl descr)); (* prove isomorphism properties *) @@ -565,7 +565,7 @@ val ((constr_inject', distinct_thms'), thy6) = thy5 - |> parent_path (#flat_names config) + |> Sign.parent_path |> store_thmss "inject" new_type_names constr_inject ||>> store_thmss "distinct" new_type_names distinct_thms; diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/Function/context_tree.ML --- a/src/HOL/Tools/Function/context_tree.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/Function/context_tree.ML Thu Jul 23 21:13:21 2009 +0200 @@ -116,8 +116,9 @@ val (c, subs) = (concl_of r, prems_of r) val subst = Pattern.match (ProofContext.theory_of ctx) (c, tt') (Vartab.empty, Vartab.empty) - val branches = map (mk_branch ctx o Envir.beta_norm o Envir.subst_vars subst) subs - val inst = map (fn v => (cterm_of thy (Var v), cterm_of thy (Envir.subst_vars subst (Var v)))) (Term.add_vars c []) + val branches = map (mk_branch ctx o Envir.beta_norm o Envir.subst_term subst) subs + val inst = map (fn v => + (cterm_of thy (Var v), cterm_of thy (Envir.subst_term subst (Var v)))) (Term.add_vars c []) in (cterm_instantiate inst r, dep, branches) end diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/Function/decompose.ML --- a/src/HOL/Tools/Function/decompose.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/Function/decompose.ML Thu Jul 23 21:13:21 2009 +0200 @@ -98,7 +98,7 @@ fun auto_decompose_tac ctxt = Termination.TERMINATION ctxt - (decompose_tac ctxt (auto_tac (local_clasimpset_of ctxt)) + (decompose_tac ctxt (auto_tac (clasimpset_of ctxt)) (K (K all_tac)) (K (K no_tac))) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/Function/fundef_core.ML --- a/src/HOL/Tools/Function/fundef_core.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/Function/fundef_core.ML Thu Jul 23 21:13:21 2009 +0200 @@ -176,19 +176,18 @@ end -(* lowlevel term function *) +(* lowlevel term function. FIXME: remove *) fun abstract_over_list vs body = let - exception SAME; fun abs lev v tm = if v aconv tm then Bound lev else (case tm of Abs (a, T, t) => Abs (a, T, abs (lev + 1) v t) - | t $ u => (abs lev v t $ (abs lev v u handle SAME => u) handle SAME => t $ abs lev v u) - | _ => raise SAME); + | t $ u => abs lev v t $ abs lev v u + | t => t); in - fold_index (fn (i,v) => fn t => abs i v t handle SAME => t) vs body + fold_index (fn (i, v) => fn t => abs i v t) vs body end @@ -703,7 +702,7 @@ Goal.init goal |> (SINGLE (resolve_tac [accI] 1)) |> the |> (SINGLE (eresolve_tac [Thm.forall_elim_vars 0 R_cases] 1)) |> the - |> (SINGLE (auto_tac (local_clasimpset_of ctxt))) |> the + |> (SINGLE (auto_tac (clasimpset_of ctxt))) |> the |> Goal.conclude |> fold_rev forall_intr_rename (map fst oqs ~~ cqs) end @@ -832,7 +831,7 @@ ((rtac (G_induct OF [a])) THEN_ALL_NEW (rtac accI) THEN_ALL_NEW (etac R_cases) - THEN_ALL_NEW (asm_full_simp_tac (local_simpset_of octxt))) 1) + THEN_ALL_NEW (asm_full_simp_tac (simpset_of octxt))) 1) val default_thm = (forall_intr_vars graph_implies_dom) COMP (f_def COMP fundef_default_value) @@ -850,9 +849,9 @@ (fn _ => rtac (instantiate' [] [SOME (cterm_of thy lhs_acc)] case_split) 1 THEN (rtac (Thm.forall_elim_vars 0 psimp) THEN_ALL_NEW assume_tac) 1 - THEN (simp_default_tac (local_simpset_of ctxt) 1) + THEN (simp_default_tac (simpset_of ctxt) 1) THEN (etac not_acc_down 1) - THEN ((etac R_cases) THEN_ALL_NEW (simp_default_tac (local_simpset_of ctxt))) 1) + THEN ((etac R_cases) THEN_ALL_NEW (simp_default_tac (simpset_of ctxt))) 1) |> fold_rev forall_intr_rename (map fst oqs ~~ cqs) end in diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/Function/fundef_datatype.ML --- a/src/HOL/Tools/Function/fundef_datatype.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/Function/fundef_datatype.ML Thu Jul 23 21:13:21 2009 +0200 @@ -102,7 +102,8 @@ fun inst_constrs_of thy (T as Type (name, _)) = - map (fn (Cn,CT) => Envir.subst_TVars (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT))) + map (fn (Cn,CT) => + Envir.subst_term_types (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT))) (the (Datatype.get_constrs thy name)) | inst_constrs_of thy _ = raise Match diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/Function/fundef_lib.ML --- a/src/HOL/Tools/Function/fundef_lib.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/Function/fundef_lib.ML Thu Jul 23 21:13:21 2009 +0200 @@ -167,10 +167,10 @@ end (* instance for unions *) -fun regroup_union_conv t = regroup_conv @{const_name Set.empty} @{const_name Un} - (map (fn t => t RS eq_reflection) (@{thms "Un_ac"} @ - @{thms "Un_empty_right"} @ - @{thms "Un_empty_left"})) t +fun regroup_union_conv t = regroup_conv @{const_name Set.empty} @{const_name Set.union} + (map (fn t => t RS eq_reflection) (@{thms Un_ac} @ + @{thms Un_empty_right} @ + @{thms Un_empty_left})) t end diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/Function/lexicographic_order.ML --- a/src/HOL/Tools/Function/lexicographic_order.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/Function/lexicographic_order.ML Thu Jul 23 21:13:21 2009 +0200 @@ -140,7 +140,7 @@ fun pr_table table = writeln (cat_lines (map (fn r => concat (map pr_cell r)) table)) fun pr_goals ctxt st = - Display.pretty_goals_aux (Syntax.pp ctxt) Markup.none (true, false) (Thm.nprems_of st) st + Display_Goal.pretty_goals ctxt Markup.none (true, false) (Thm.nprems_of st) st |> Pretty.chunks |> Pretty.string_of @@ -217,7 +217,7 @@ fun lexicographic_order_tac ctxt = TRY (FundefCommon.apply_termination_rule ctxt 1) THEN lex_order_tac ctxt - (auto_tac (local_clasimpset_of ctxt addsimps2 FundefCommon.Termination_Simps.get ctxt)) + (auto_tac (clasimpset_of ctxt addsimps2 FundefCommon.Termination_Simps.get ctxt)) val lexicographic_order = SIMPLE_METHOD o lexicographic_order_tac diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/Function/mutual.ML --- a/src/HOL/Tools/Function/mutual.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/Function/mutual.ML Thu Jul 23 21:13:21 2009 +0200 @@ -100,7 +100,7 @@ val (caTss, resultTs) = split_list (map curried_types fs) val argTs = map (foldr1 HOLogic.mk_prodT) caTss - val dresultTs = distinct (Type.eq_type Vartab.empty) resultTs + val dresultTs = distinct (op =) resultTs val n' = length dresultTs val RST = BalancedTree.make (uncurry SumTree.mk_sumT) dresultTs @@ -114,7 +114,7 @@ fun define (fvar as (n, T)) caTs resultT i = let val vars = map_index (fn (j,T) => Free ("x" ^ string_of_int j, T)) caTs (* FIXME: Bind xs properly *) - val i' = find_index (fn Ta => Type.eq_type Vartab.empty (Ta, resultT)) dresultTs + 1 + val i' = find_index (fn Ta => Ta = resultT) dresultTs + 1 val f_exp = SumTree.mk_proj RST n' i' (Free fsum_var $ SumTree.mk_inj ST num i (foldr1 HOLogic.mk_prod vars)) val def = Term.abstract_over (Free fsum_var, fold_rev lambda vars f_exp) @@ -206,7 +206,7 @@ (HOLogic.Trueprop $ HOLogic.mk_eq (list_comb (f, args), rhs)) (fn _ => (LocalDefs.unfold_tac ctxt all_orig_fdefs) THEN EqSubst.eqsubst_tac ctxt [0] [simp] 1 - THEN (simp_tac (local_simpset_of ctxt addsimps SumTree.proj_in_rules)) 1) + THEN (simp_tac (simpset_of ctxt addsimps SumTree.proj_in_rules)) 1) |> restore_cond |> export end diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/Function/pattern_split.ML --- a/src/HOL/Tools/Function/pattern_split.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/Function/pattern_split.ML Thu Jul 23 21:13:21 2009 +0200 @@ -39,7 +39,8 @@ (* This is copied from "fundef_datatype.ML" *) fun inst_constrs_of thy (T as Type (name, _)) = - map (fn (Cn,CT) => Envir.subst_TVars (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT))) + map (fn (Cn,CT) => + Envir.subst_term_types (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT))) (the (Datatype.get_constrs thy name)) | inst_constrs_of thy T = raise TYPE ("inst_constrs_of", [T], []) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/Function/scnp_reconstruct.ML --- a/src/HOL/Tools/Function/scnp_reconstruct.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/Function/scnp_reconstruct.ML Thu Jul 23 21:13:21 2009 +0200 @@ -291,7 +291,7 @@ THEN (rtac @{thm rp_inv_image_rp} 1) THEN (rtac (order_rpair ms_rp label) 1) THEN PRIMITIVE (instantiate' [] [SOME level_mapping]) - THEN unfold_tac @{thms rp_inv_image_def} (local_simpset_of ctxt) + THEN unfold_tac @{thms rp_inv_image_def} (simpset_of ctxt) THEN LocalDefs.unfold_tac ctxt (@{thms split_conv} @ @{thms fst_conv} @ @{thms snd_conv}) THEN REPEAT (SOMEGOAL (resolve_tac [@{thm Un_least}, @{thm empty_subsetI}])) @@ -406,7 +406,7 @@ fun decomp_scnp orders ctxt = let val extra_simps = FundefCommon.Termination_Simps.get ctxt - val autom_tac = auto_tac (local_clasimpset_of ctxt addsimps2 extra_simps) + val autom_tac = auto_tac (clasimpset_of ctxt addsimps2 extra_simps) in SIMPLE_METHOD (gen_sizechange_tac orders autom_tac ctxt (print_error ctxt)) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/Function/termination.ML --- a/src/HOL/Tools/Function/termination.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/Function/termination.ML Thu Jul 23 21:13:21 2009 +0200 @@ -145,7 +145,7 @@ fun mk_sum_skel rel = let - val cs = FundefLib.dest_binop_list @{const_name Un} rel + val cs = FundefLib.dest_binop_list @{const_name union} rel fun collect_pats (Const ("Collect", _) $ Abs (_, _, c)) = let val (Const ("op &", _) $ (Const ("op =", _) $ _ $ (Const ("Pair", _) $ r $ l)) $ Gam) @@ -233,7 +233,7 @@ fun CALLS tac i st = if Thm.no_prems st then all_tac st else case Thm.term_of (Thm.cprem_of st i) of - (_ $ (_ $ rel)) => tac (FundefLib.dest_binop_list @{const_name Un} rel, i) st + (_ $ (_ $ rel)) => tac (FundefLib.dest_binop_list @{const_name union} rel, i) st |_ => no_tac st type ttac = (data -> int -> tactic) -> (data -> int -> tactic) -> data -> int -> tactic @@ -293,7 +293,7 @@ if null ineqs then Const (@{const_name Set.empty}, fastype_of rel) else - foldr1 (HOLogic.mk_binop @{const_name Un}) (map mk_compr ineqs) + foldr1 (HOLogic.mk_binop @{const_name union}) (map mk_compr ineqs) fun solve_membership_tac i = (EVERY' (replicate (i - 2) (rtac @{thm UnI2})) (* pick the right component of the union *) @@ -303,7 +303,7 @@ THEN' ((rtac @{thm refl}) (* unification instantiates all Vars *) ORELSE' ((rtac @{thm conjI}) THEN' (rtac @{thm refl}) - THEN' (blast_tac (local_claset_of ctxt)))) (* Solve rest of context... not very elegant *) + THEN' (blast_tac (claset_of ctxt)))) (* Solve rest of context... not very elegant *) ) i in ((PRIMITIVE (Drule.cterm_instantiate [(cert rel, cert relation)]) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/Qelim/ferrante_rackoff.ML --- a/src/HOL/Tools/Qelim/ferrante_rackoff.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/Qelim/ferrante_rackoff.ML Thu Jul 23 21:13:21 2009 +0200 @@ -228,6 +228,6 @@ ObjectLogic.full_atomize_tac i THEN simp_tac (#simpset (snd instance)) i THEN (* FIXME already part of raw_ferrack_qe_conv? *) CONVERSION (ObjectLogic.judgment_conv (raw_ferrack_qe_conv ctxt instance)) i THEN - simp_tac (Simplifier.local_simpset_of ctxt) i)); (* FIXME really? *) + simp_tac (simpset_of ctxt) i)); (* FIXME really? *) end; diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/TFL/casesplit.ML --- a/src/HOL/Tools/TFL/casesplit.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/TFL/casesplit.ML Thu Jul 23 21:13:21 2009 +0200 @@ -129,8 +129,8 @@ | _ => error "not a valid case thm"; val type_cinsts = map (fn (ixn, (S, T)) => (ctypify (TVar (ixn, S)), ctypify T)) (Vartab.dest type_insts); - val cPv = ctermify (Envir.subst_TVars type_insts Pv); - val cDv = ctermify (Envir.subst_TVars type_insts Dv); + val cPv = ctermify (Envir.subst_term_types type_insts Pv); + val cDv = ctermify (Envir.subst_term_types type_insts Dv); in (beta_eta_contract (case_thm @@ -269,9 +269,9 @@ fun split th = (case find_thms_split splitths 1 th of NONE => - (writeln "th:"; - Display.print_thm th; writeln "split ths:"; - Display.print_thms splitths; writeln "\n--"; + (writeln (cat_lines + (["th:", Display.string_of_thm_without_context th, "split ths:"] @ + map Display.string_of_thm_without_context splitths @ ["\n--"])); error "splitto: cannot find variable to split on") | SOME v => let diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/TFL/rules.ML --- a/src/HOL/Tools/TFL/rules.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/TFL/rules.ML Thu Jul 23 21:13:21 2009 +0200 @@ -552,7 +552,7 @@ fun say s = if !tracing then writeln s else (); fun print_thms s L = - say (cat_lines (s :: map Display.string_of_thm L)); + say (cat_lines (s :: map Display.string_of_thm_without_context L)); fun print_cterms s L = say (cat_lines (s :: map Display.string_of_cterm L)); @@ -677,7 +677,7 @@ val cntxt = Simplifier.prems_of_ss ss val dummy = print_thms "cntxt:" cntxt val dummy = say "cong rule:" - val dummy = say (Display.string_of_thm thm) + val dummy = say (Display.string_of_thm_without_context thm) val dummy = thm_ref := (thm :: !thm_ref) val dummy = ss_ref := (ss :: !ss_ref) (* Unquantified eliminate *) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/TFL/tfl.ML --- a/src/HOL/Tools/TFL/tfl.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/TFL/tfl.ML Thu Jul 23 21:13:21 2009 +0200 @@ -529,9 +529,8 @@ val f = #lhs(S.dest_eq proto_def) val (extractants,TCl) = ListPair.unzip extracta val dummy = if !trace - then (writeln "Extractants = "; - Display.prths extractants; - ()) + then writeln (cat_lines ("Extractants =" :: + map (Display.string_of_thm_global thy) extractants)) else () val TCs = List.foldr (gen_union (op aconv)) [] TCl val full_rqt = WFR::TCs @@ -551,8 +550,9 @@ |> PureThy.add_defs false [Thm.no_attributes (Binding.name (fid ^ "_def"), defn)] val def = Thm.freezeT def0; - val dummy = if !trace then writeln ("DEF = " ^ Display.string_of_thm def) - else () + val dummy = + if !trace then writeln ("DEF = " ^ Display.string_of_thm_global theory def) + else () (* val fconst = #lhs(S.dest_eq(concl def)) *) val tych = Thry.typecheck theory val full_rqt_prop = map (Dcterm.mk_prop o tych) full_rqt @@ -560,7 +560,7 @@ val baz = R.DISCH_ALL (fold_rev R.DISCH full_rqt_prop (R.LIST_CONJ extractants)) - val dum = if !trace then writeln ("baz = " ^ Display.string_of_thm baz) + val dum = if !trace then writeln ("baz = " ^ Display.string_of_thm_global theory baz) else () val f_free = Free (fid, fastype_of f) (*'cos f is a Const*) val SV' = map tych SV; @@ -909,7 +909,7 @@ fun trace_thms s L = - if !trace then writeln (cat_lines (s :: map Display.string_of_thm L)) + if !trace then writeln (cat_lines (s :: map Display.string_of_thm_without_context L)) else (); fun trace_cterms s L = diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/TFL/thry.ML --- a/src/HOL/Tools/TFL/thry.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/TFL/thry.ML Thu Jul 23 21:13:21 2009 +0200 @@ -35,7 +35,7 @@ fun match_term thry pat ob = let val (ty_theta, tm_theta) = Pattern.match thry (pat,ob) (Vartab.empty, Vartab.empty); - fun tmbind (ixn, (T, t)) = (Var (ixn, Envir.typ_subst_TVars ty_theta T), t) + fun tmbind (ixn, (T, t)) = (Var (ixn, Envir.subst_type ty_theta T), t) in (map tmbind (Vartab.dest tm_theta), map tybind (Vartab.dest ty_theta)) end; diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/atp_minimal.ML --- a/src/HOL/Tools/atp_minimal.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/atp_minimal.ML Thu Jul 23 21:13:21 2009 +0200 @@ -125,7 +125,8 @@ println ("Minimize called with " ^ length_string name_thms_pairs ^ " theorems, prover: " ^ prover_name ^ ", time limit: " ^ Int.toString time_limit ^ " seconds") val _ = debug_fn (fn () => app (fn (n, tl) => - (debug n; app (fn t => debug (" " ^ Display.string_of_thm t)) tl)) name_thms_pairs) + (debug n; app (fn t => + debug (" " ^ Display.string_of_thm (Proof.context_of state) t)) tl)) name_thms_pairs) val test_thms_fun = sh_test_thms prover prover_name time_limit 1 state fun test_thms filtered thms = case test_thms_fun filtered thms of (Success _, _) => true | _ => false diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/atp_wrapper.ML --- a/src/HOL/Tools/atp_wrapper.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/atp_wrapper.ML Thu Jul 23 21:13:21 2009 +0200 @@ -62,7 +62,7 @@ val goal_cls = #1 (ResAxioms.neg_conjecture_clauses th subgoalno) handle THM ("assume: variables", _, _) => error "Sledgehammer: Goal contains type variables (TVars)" - val _ = app (fn th => Output.debug (fn _ => Display.string_of_thm th)) goal_cls + val _ = app (fn th => Output.debug (fn _ => Display.string_of_thm ctxt th)) goal_cls val the_filtered_clauses = case filtered_clauses of NONE => relevance_filter goal goal_cls diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/choice_specification.ML --- a/src/HOL/Tools/choice_specification.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/choice_specification.ML Thu Jul 23 21:13:21 2009 +0200 @@ -183,7 +183,7 @@ fun add_final (arg as (thy, thm)) = if name = "" then arg |> Library.swap - else (writeln (" " ^ name ^ ": " ^ (Display.string_of_thm thm)); + else (writeln (" " ^ name ^ ": " ^ Display.string_of_thm_global thy thm); PureThy.store_thm (Binding.name name, thm) thy) in args |> apsnd (remove_alls frees) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/inductive.ML --- a/src/HOL/Tools/inductive.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/inductive.ML Thu Jul 23 21:13:21 2009 +0200 @@ -140,7 +140,7 @@ val space = Consts.space_of (ProofContext.consts_of ctxt); in [Pretty.strs ("(co)inductives:" :: map #1 (NameSpace.extern_table (space, tab))), - Pretty.big_list "monotonicity rules:" (map (ProofContext.pretty_thm ctxt) monos)] + Pretty.big_list "monotonicity rules:" (map (Display.pretty_thm ctxt) monos)] |> Pretty.chunks |> Pretty.writeln end; @@ -179,7 +179,8 @@ [dest_less_concl (Seq.hd (REPEAT (FIRSTGOAL (resolve_tac [le_funI, le_boolI'])) thm))] | _ => [thm] - end handle THM _ => error ("Bad monotonicity theorem:\n" ^ Display.string_of_thm thm); + end handle THM _ => + error ("Bad monotonicity theorem:\n" ^ Display.string_of_thm_without_context thm); val mono_add = Thm.declaration_attribute (map_monos o fold Thm.add_thm o mk_mono); val mono_del = Thm.declaration_attribute (map_monos o fold Thm.del_thm o mk_mono); @@ -424,7 +425,7 @@ fun mk_cases ctxt prop = let val thy = ProofContext.theory_of ctxt; - val ss = Simplifier.local_simpset_of ctxt; + val ss = simpset_of ctxt; fun err msg = error (Pretty.string_of (Pretty.block @@ -922,7 +923,7 @@ val tab = fold (Pattern.first_order_match thy) (ts ~~ us) (Vartab.empty, Vartab.empty); in - map (Envir.subst_vars tab) vars + map (Envir.subst_term tab) vars end in map (mtch o apsnd prop_of) (cases ~~ intros) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/inductive_codegen.ML --- a/src/HOL/Tools/inductive_codegen.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/inductive_codegen.ML Thu Jul 23 21:13:21 2009 +0200 @@ -39,7 +39,7 @@ fun warn thm = warning ("InductiveCodegen: Not a proper clause:\n" ^ - Display.string_of_thm thm); + Display.string_of_thm_without_context thm); fun add_node (g, x) = Graph.new_node (x, ()) g handle Graph.DUP _ => g; diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/inductive_realizer.ML --- a/src/HOL/Tools/inductive_realizer.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/inductive_realizer.ML Thu Jul 23 21:13:21 2009 +0200 @@ -19,7 +19,7 @@ (case Proofterm.fold_proof_atoms false (fn PThm (_, ((name, _, _), _)) => cons name | _ => I) [Thm.proof_of thm] [] of [name] => name - | _ => error ("name_of_thm: bad proof of theorem\n" ^ Display.string_of_thm thm)); + | _ => error ("name_of_thm: bad proof of theorem\n" ^ Display.string_of_thm_without_context thm)); val all_simps = map (symmetric o mk_meta_eq) (thms "HOL.all_simps"); @@ -308,7 +308,7 @@ val ((dummies, some_dt_names), thy2) = thy1 |> add_dummies (Datatype.add_datatype - { strict = false, flat_names = false, quiet = false } (map (Binding.name_of o #2) dts)) + { strict = false, quiet = false } (map (Binding.name_of o #2) dts)) (map (pair false) dts) [] ||> Extraction.add_typeof_eqns_i ty_eqs ||> Extraction.add_realizes_eqns_i rlz_eqs; diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/inductive_set.ML --- a/src/HOL/Tools/inductive_set.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/inductive_set.ML Thu Jul 23 21:13:21 2009 +0200 @@ -73,8 +73,8 @@ in Drule.instantiate' [] (rev (map (SOME o cterm_of thy o Var) vs)) (p (fold (Logic.all o Var) vs t) f) end; - fun mkop "op &" T x = SOME (Const (@{const_name "Int"}, T --> T --> T), x) - | mkop "op |" T x = SOME (Const (@{const_name "Un"}, T --> T --> T), x) + fun mkop "op &" T x = SOME (Const (@{const_name inter}, T --> T --> T), x) + | mkop "op |" T x = SOME (Const (@{const_name union}, T --> T --> T), x) | mkop _ _ _ = NONE; fun mk_collect p T t = let val U = HOLogic.dest_setT T @@ -324,7 +324,7 @@ fun to_pred_proc thy rules t = case lookup_rule thy I rules t of NONE => NONE | SOME (lhs, rhs) => - SOME (Envir.subst_vars + SOME (Envir.subst_term (Pattern.match thy (lhs, t) (Vartab.empty, Vartab.empty)) rhs); fun to_pred thms ctxt thm = diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/lin_arith.ML --- a/src/HOL/Tools/lin_arith.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/lin_arith.ML Thu Jul 23 21:13:21 2009 +0200 @@ -286,7 +286,7 @@ (* checks if splitting with 'thm' is implemented *) -fun is_split_thm (thm : thm) : bool = +fun is_split_thm thm = case concl_of thm of _ $ (_ $ (_ $ lhs) $ _) => ( (* Trueprop $ ((op =) $ (?P $ lhs) $ rhs) *) case head_of lhs of @@ -298,10 +298,10 @@ "Divides.div_class.mod", "Divides.div_class.div"] a | _ => (warning ("Lin. Arith.: wrong format for split rule " ^ - Display.string_of_thm thm); + Display.string_of_thm_without_context thm); false)) | _ => (warning ("Lin. Arith.: wrong format for split rule " ^ - Display.string_of_thm thm); + Display.string_of_thm_without_context thm); false); (* substitute new for occurrences of old in a term, incrementing bound *) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/meson.ML --- a/src/HOL/Tools/meson.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/meson.ML Thu Jul 23 21:13:21 2009 +0200 @@ -87,14 +87,12 @@ fun typ_pair_of (ix, (sort,ty)) = (TVar (ix,sort), ty); fun term_pair_of (ix, (ty,t)) = (Var (ix,ty), t); -val Envir.Envir {asol = tenv0, iTs = tyenv0, ...} = Envir.empty 0 - (*FIXME: currently does not "rename variables apart"*) fun first_order_resolve thA thB = let val thy = theory_of_thm thA val tmA = concl_of thA val Const("==>",_) $ tmB $ _ = prop_of thB - val (tyenv,tenv) = Pattern.first_order_match thy (tmB,tmA) (tyenv0,tenv0) + val (tyenv, tenv) = Pattern.first_order_match thy (tmB, tmA) (Vartab.empty, Vartab.empty) val ct_pairs = map (pairself (cterm_of thy) o term_pair_of) (Vartab.dest tenv) in thA RS (cterm_instantiate ct_pairs thB) end handle _ => raise THM ("first_order_resolve", 0, [thA,thB]); (* FIXME avoid handle _ *) @@ -104,8 +102,8 @@ [] => th | pairs => let val thy = theory_of_thm th - val (tyenv,tenv) = - List.foldl (uncurry (Pattern.first_order_match thy)) (tyenv0,tenv0) pairs + val (tyenv, tenv) = + fold (Pattern.first_order_match thy) pairs (Vartab.empty, Vartab.empty) val t_pairs = map term_pair_of (Vartab.dest tenv) val th' = Thm.instantiate ([], map (pairself (cterm_of thy)) t_pairs) th in th' end @@ -129,10 +127,10 @@ fun forward_res nf st = let fun forward_tacf [prem] = rtac (nf prem) 1 | forward_tacf prems = - error ("Bad proof state in forward_res, please inform lcp@cl.cam.ac.uk:\n" ^ - Display.string_of_thm st ^ - "\nPremises:\n" ^ - ML_Syntax.print_list Display.string_of_thm prems) + error (cat_lines + ("Bad proof state in forward_res, please inform lcp@cl.cam.ac.uk:" :: + Display.string_of_thm_without_context st :: + "Premises:" :: map Display.string_of_thm_without_context prems)) in case Seq.pull (ALLGOALS (METAHYPS forward_tacf) st) of SOME(th,_) => th @@ -344,7 +342,7 @@ and cnf_nil th = cnf_aux (th,[]) val cls = if too_many_clauses (SOME ctxt) (concl_of th) - then (warning ("cnf is ignoring: " ^ Display.string_of_thm th); ths) + then (warning ("cnf is ignoring: " ^ Display.string_of_thm ctxt th); ths) else cnf_aux (th,ths) in (cls, !ctxtr) end; @@ -547,7 +545,7 @@ | skolemize_nnf_list (th::ths) = skolemize th :: skolemize_nnf_list ths handle THM _ => (*RS can fail if Unify.search_bound is too small*) - (warning ("Failed to Skolemize " ^ Display.string_of_thm th); + (warning ("Failed to Skolemize " ^ Display.string_of_thm_without_context th); skolemize_nnf_list ths); fun add_clauses (th,cls) = @@ -630,19 +628,17 @@ ITER_DEEPEN (has_fewer_prems 1) (prolog_step_tac' horns); fun iter_deepen_meson_tac ths = MESON make_clauses - (fn cls => - case (gocls (cls@ths)) of - [] => no_tac (*no goal clauses*) - | goes => - let val horns = make_horns (cls@ths) - val _ = - Output.debug (fn () => ("meson method called:\n" ^ - ML_Syntax.print_list Display.string_of_thm (cls@ths) ^ - "\nclauses:\n" ^ - ML_Syntax.print_list Display.string_of_thm horns)) - in THEN_ITER_DEEPEN (resolve_tac goes 1) (has_fewer_prems 1) (prolog_step_tac' horns) - end - ); + (fn cls => + (case (gocls (cls @ ths)) of + [] => no_tac (*no goal clauses*) + | goes => + let + val horns = make_horns (cls @ ths) + val _ = Output.debug (fn () => + cat_lines ("meson method called:" :: + map Display.string_of_thm_without_context (cls @ ths) @ + ["clauses:"] @ map Display.string_of_thm_without_context horns)) + in THEN_ITER_DEEPEN (resolve_tac goes 1) (has_fewer_prems 1) (prolog_step_tac' horns) end)); fun meson_claset_tac ths cs = SELECT_GOAL (TRY (safe_tac cs) THEN TRYALL (iter_deepen_meson_tac ths)); diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/metis_tools.ML --- a/src/HOL/Tools/metis_tools.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/metis_tools.ML Thu Jul 23 21:13:21 2009 +0200 @@ -270,7 +270,7 @@ fun print_thpair (fth,th) = (Output.debug (fn () => "============================================="); Output.debug (fn () => "Metis: " ^ Metis.Thm.toString fth); - Output.debug (fn () => "Isabelle: " ^ Display.string_of_thm th)); + Output.debug (fn () => "Isabelle: " ^ Display.string_of_thm_without_context th)); fun lookth thpairs (fth : Metis.Thm.thm) = valOf (AList.lookup (uncurry Metis.Thm.equal) thpairs fth) @@ -310,7 +310,7 @@ in SOME (cterm_of thy (Var v), t) end handle Option => (Output.debug (fn() => "List.find failed for the variable " ^ x ^ - " in " ^ Display.string_of_thm i_th); + " in " ^ Display.string_of_thm ctxt i_th); NONE) fun remove_typeinst (a, t) = case Recon.strip_prefix ResClause.schematic_var_prefix a of @@ -318,7 +318,7 @@ | NONE => case Recon.strip_prefix ResClause.tvar_prefix a of SOME _ => NONE (*type instantiations are forbidden!*) | NONE => SOME (a,t) (*internal Metis var?*) - val _ = Output.debug (fn () => " isa th: " ^ Display.string_of_thm i_th) + val _ = Output.debug (fn () => " isa th: " ^ Display.string_of_thm ctxt i_th) val substs = List.mapPartial remove_typeinst (Metis.Subst.toList fsubst) val (vars,rawtms) = ListPair.unzip (List.mapPartial subst_translation substs) val tms = infer_types ctxt rawtms; @@ -350,8 +350,8 @@ let val thy = ProofContext.theory_of ctxt val i_th1 = lookth thpairs th1 and i_th2 = lookth thpairs th2 - val _ = Output.debug (fn () => " isa th1 (pos): " ^ Display.string_of_thm i_th1) - val _ = Output.debug (fn () => " isa th2 (neg): " ^ Display.string_of_thm i_th2) + val _ = Output.debug (fn () => " isa th1 (pos): " ^ Display.string_of_thm ctxt i_th1) + val _ = Output.debug (fn () => " isa th2 (neg): " ^ Display.string_of_thm ctxt i_th2) in if is_TrueI i_th1 then i_th2 (*Trivial cases where one operand is type info*) else if is_TrueI i_th2 then i_th1 @@ -371,7 +371,7 @@ end; (* INFERENCE RULE: REFL *) - val refl_x = cterm_of (the_context ()) (Var (hd (Term.add_vars (prop_of REFL_THM) []))); + val refl_x = cterm_of @{theory} (Var (hd (Term.add_vars (prop_of REFL_THM) []))); val refl_idx = 1 + Thm.maxidx_of REFL_THM; fun refl_inf ctxt t = @@ -428,7 +428,7 @@ val _ = Output.debug (fn () => "located term: " ^ Syntax.string_of_term ctxt tm_subst) val imax = maxidx_of_term (i_tm $ tm_abs $ tm_subst) (*ill typed but gives right max*) val subst' = incr_indexes (imax+1) (if pos then subst_em else ssubst_em) - val _ = Output.debug (fn () => "subst' " ^ Display.string_of_thm subst') + val _ = Output.debug (fn () => "subst' " ^ Display.string_of_thm ctxt subst') val eq_terms = map (pairself (cterm_of thy)) (ListPair.zip (OldTerm.term_vars (prop_of subst'), [tm_abs, tm_subst, i_tm])) in cterm_instantiate eq_terms subst' end; @@ -456,7 +456,7 @@ val _ = Output.debug (fn () => "METIS THM: " ^ Metis.Thm.toString fol_th) val _ = Output.debug (fn () => "INFERENCE: " ^ Metis.Proof.inferenceToString inf) val th = Meson.flexflex_first_order (step ctxt isFO thpairs (fol_th, inf)) - val _ = Output.debug (fn () => "ISABELLE THM: " ^ Display.string_of_thm th) + val _ = Output.debug (fn () => "ISABELLE THM: " ^ Display.string_of_thm ctxt th) val _ = Output.debug (fn () => "=============================================") val n_metis_lits = length (filter real_literal (Metis.LiteralSet.toList (Metis.Thm.clause fol_th))) in @@ -475,14 +475,14 @@ fun cnf_th thy th = hd (ResAxioms.cnf_axiom thy th); - val equal_imp_fequal' = cnf_th (the_context ()) (thm"equal_imp_fequal"); - val fequal_imp_equal' = cnf_th (the_context ()) (thm"fequal_imp_equal"); + val equal_imp_fequal' = cnf_th @{theory} @{thm equal_imp_fequal}; + val fequal_imp_equal' = cnf_th @{theory} @{thm fequal_imp_equal}; - val comb_I = cnf_th (the_context ()) ResHolClause.comb_I; - val comb_K = cnf_th (the_context ()) ResHolClause.comb_K; - val comb_B = cnf_th (the_context ()) ResHolClause.comb_B; - val comb_C = cnf_th (the_context ()) ResHolClause.comb_C; - val comb_S = cnf_th (the_context ()) ResHolClause.comb_S; + val comb_I = cnf_th @{theory} ResHolClause.comb_I; + val comb_K = cnf_th @{theory} ResHolClause.comb_K; + val comb_B = cnf_th @{theory} ResHolClause.comb_B; + val comb_C = cnf_th @{theory} ResHolClause.comb_C; + val comb_S = cnf_th @{theory} ResHolClause.comb_S; fun type_ext thy tms = let val subs = ResAtp.tfree_classes_of_terms tms @@ -576,9 +576,9 @@ val th_cls_pairs = map (fn th => (Thm.get_name_hint th, ResAxioms.cnf_axiom thy th)) ths0 val ths = List.concat (map #2 th_cls_pairs) val _ = Output.debug (fn () => "FOL_SOLVE: CONJECTURE CLAUSES") - val _ = app (fn th => Output.debug (fn () => Display.string_of_thm th)) cls + val _ = app (fn th => Output.debug (fn () => Display.string_of_thm ctxt th)) cls val _ = Output.debug (fn () => "THEOREM CLAUSES") - val _ = app (fn th => Output.debug (fn () => Display.string_of_thm th)) ths + val _ = app (fn th => Output.debug (fn () => Display.string_of_thm ctxt th)) ths val {isFO,axioms,tfrees} = build_map mode ctxt cls ths val _ = if null tfrees then () else (Output.debug (fn () => "TFREE CLAUSES"); @@ -604,14 +604,14 @@ val result = translate isFO ctxt' axioms proof and used = List.mapPartial (used_axioms axioms) proof val _ = Output.debug (fn () => "METIS COMPLETED...clauses actually used:") - val _ = app (fn th => Output.debug (fn () => Display.string_of_thm th)) used + val _ = app (fn th => Output.debug (fn () => Display.string_of_thm ctxt th)) used val unused = filter (fn (a,cls) => not (common_thm used cls)) th_cls_pairs in if null unused then () else warning ("Metis: unused theorems " ^ commas_quote (map #1 unused)); case result of (_,ith)::_ => - (Output.debug (fn () => "success: " ^ Display.string_of_thm ith); + (Output.debug (fn () => "success: " ^ Display.string_of_thm ctxt ith); [ith]) | _ => (Output.debug (fn () => "Metis: no result"); []) @@ -623,7 +623,7 @@ fun metis_general_tac mode ctxt ths i st0 = let val _ = Output.debug (fn () => - "Metis called with theorems " ^ cat_lines (map Display.string_of_thm ths)) + "Metis called with theorems " ^ cat_lines (map (Display.string_of_thm ctxt) ths)) in if exists_type ResAxioms.type_has_empty_sort (prop_of st0) then (warning "Proof state contains the empty sort"; Seq.empty) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/nat_arith.ML --- a/src/HOL/Tools/nat_arith.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/nat_arith.ML Thu Jul 23 21:13:21 2009 +0200 @@ -91,18 +91,18 @@ end); val nat_cancel_sums_add = - [Simplifier.simproc (the_context ()) "nateq_cancel_sums" + [Simplifier.simproc @{theory} "nateq_cancel_sums" ["(l::nat) + m = n", "(l::nat) = m + n", "Suc m = n", "m = Suc n"] (K EqCancelSums.proc), - Simplifier.simproc (the_context ()) "natless_cancel_sums" + Simplifier.simproc @{theory} "natless_cancel_sums" ["(l::nat) + m < n", "(l::nat) < m + n", "Suc m < n", "m < Suc n"] (K LessCancelSums.proc), - Simplifier.simproc (the_context ()) "natle_cancel_sums" + Simplifier.simproc @{theory} "natle_cancel_sums" ["(l::nat) + m <= n", "(l::nat) <= m + n", "Suc m <= n", "m <= Suc n"] (K LeCancelSums.proc)]; val nat_cancel_sums = nat_cancel_sums_add @ - [Simplifier.simproc (the_context ()) "natdiff_cancel_sums" + [Simplifier.simproc @{theory} "natdiff_cancel_sums" ["((l::nat) + m) - n", "(l::nat) - (m + n)", "Suc m - n", "m - Suc n"] (K DiffCancelSums.proc)]; diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/nat_numeral_simprocs.ML --- a/src/HOL/Tools/nat_numeral_simprocs.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/nat_numeral_simprocs.ML Thu Jul 23 21:13:21 2009 +0200 @@ -19,7 +19,7 @@ val numeral_sym_ss = HOL_ss addsimps numeral_syms; fun rename_numerals th = - simplify numeral_sym_ss (Thm.transfer (the_context ()) th); + simplify numeral_sym_ss (Thm.transfer @{theory} th); (*Utilities*) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/numeral.ML --- a/src/HOL/Tools/numeral.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/numeral.ML Thu Jul 23 21:13:21 2009 +0200 @@ -39,7 +39,7 @@ val oneT = Thm.ctyp_of_term one; val number_of = @{cpat "number_of"}; -val numberT = Thm.ctyp_of (the_context ()) (Term.range_type (Thm.typ_of (Thm.ctyp_of_term number_of))); +val numberT = Thm.ctyp_of @{theory} (Term.range_type (Thm.typ_of (Thm.ctyp_of_term number_of))); fun instT T V = Thm.instantiate_cterm ([(V, T)], []); diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/recdef.ML --- a/src/HOL/Tools/recdef.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/recdef.ML Thu Jul 23 21:13:21 2009 +0200 @@ -48,9 +48,9 @@ fun map_wfs f = map_hints (fn (simps, congs, wfs) => (simps, congs, f wfs)); fun pretty_hints ({simps, congs, wfs}: hints) = - [Pretty.big_list "recdef simp hints:" (map Display.pretty_thm simps), - Pretty.big_list "recdef cong hints:" (map Display.pretty_thm (map snd congs)), - Pretty.big_list "recdef wf hints:" (map Display.pretty_thm wfs)]; + [Pretty.big_list "recdef simp hints:" (map Display.pretty_thm_without_context simps), + Pretty.big_list "recdef cong hints:" (map Display.pretty_thm_without_context (map snd congs)), + Pretty.big_list "recdef wf hints:" (map Display.pretty_thm_without_context wfs)]; (* congruence rules *) @@ -172,15 +172,15 @@ NONE => ctxt0 | SOME src => #2 (Method.syntax (Method.sections recdef_modifiers) src ctxt0)); val {simps, congs, wfs} = get_hints ctxt; - val cs = local_claset_of ctxt; - val ss = local_simpset_of ctxt addsimps simps; + val cs = claset_of ctxt; + val ss = simpset_of ctxt addsimps simps; in (cs, ss, rev (map snd congs), wfs) end; fun prepare_hints_i thy () = let val ctxt0 = ProofContext.init thy; val {simps, congs, wfs} = get_global_hints thy; - in (local_claset_of ctxt0, local_simpset_of ctxt0 addsimps simps, rev (map snd congs), wfs) end; + in (claset_of ctxt0, simpset_of ctxt0 addsimps simps, rev (map snd congs), wfs) end; diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/record.ML --- a/src/HOL/Tools/record.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/record.ML Thu Jul 23 21:13:21 2009 +0200 @@ -105,7 +105,7 @@ (* messages *) fun trace_thm str thm = - tracing (str ^ (Pretty.string_of (Display.pretty_thm thm))); + tracing (str ^ (Pretty.string_of (Display.pretty_thm_without_context thm))); fun trace_thms str thms = (tracing str; map (trace_thm "") thms); diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/res_atp.ML --- a/src/HOL/Tools/res_atp.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/res_atp.ML Thu Jul 23 21:13:21 2009 +0200 @@ -401,8 +401,9 @@ (List.foldl add_single_names (List.foldl add_multi_names [] mults) singles) end; -fun check_named ("",th) = (warning ("No name for theorem " ^ Display.string_of_thm th); false) - | check_named (_,th) = true; +fun check_named ("", th) = + (warning ("No name for theorem " ^ Display.string_of_thm_without_context th); false) + | check_named (_, th) = true; (* get lemmas from claset, simpset, atpset and extra supplied rules *) fun get_clasimp_atp_lemmas ctxt = @@ -538,7 +539,7 @@ val extra_cls = white_cls @ extra_cls val axcls = white_cls @ axcls val ccls = subtract_cls goal_cls extra_cls - val _ = app (fn th => Output.debug (fn _ => Display.string_of_thm th)) ccls + val _ = app (fn th => Output.debug (fn _ => Display.string_of_thm_global thy th)) ccls val ccltms = map prop_of ccls and axtms = map (prop_of o #1) extra_cls val subs = tfree_classes_of_terms ccltms diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/res_axioms.ML --- a/src/HOL/Tools/res_axioms.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/res_axioms.ML Thu Jul 23 21:13:21 2009 +0200 @@ -154,9 +154,9 @@ val monomorphic = not o Term.exists_type (Term.exists_subtype Term.is_TVar); -val [f_B,g_B] = map (cterm_of (the_context ())) (OldTerm.term_vars (prop_of @{thm abs_B})); -val [g_C,f_C] = map (cterm_of (the_context ())) (OldTerm.term_vars (prop_of @{thm abs_C})); -val [f_S,g_S] = map (cterm_of (the_context ())) (OldTerm.term_vars (prop_of @{thm abs_S})); +val [f_B,g_B] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_B})); +val [g_C,f_C] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_C})); +val [f_S,g_S] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_S})); (*FIXME: requires more use of cterm constructors*) fun abstract ct = @@ -227,8 +227,9 @@ val eqth = combinators_aux (cprop_of th) in equal_elim eqth th end handle THM (msg,_,_) => - (warning ("Error in the combinator translation of " ^ Display.string_of_thm th); - warning (" Exception message: " ^ msg); + (warning (cat_lines + ["Error in the combinator translation of " ^ Display.string_of_thm_without_context th, + " Exception message: " ^ msg]); TrueI); (*A type variable of sort {} will cause make abstraction fail.*) (*cterms are used throughout for efficiency*) @@ -280,7 +281,7 @@ fun assert_lambda_free ths msg = case filter (not o lambda_free o prop_of) ths of [] => () - | ths' => error (msg ^ "\n" ^ cat_lines (map Display.string_of_thm ths')); + | ths' => error (cat_lines (msg :: map Display.string_of_thm_without_context ths')); (*** Blacklisting (duplicated in ResAtp?) ***) @@ -336,7 +337,7 @@ fun name_or_string th = if Thm.has_name_hint th then Thm.get_name_hint th - else Display.string_of_thm th; + else Display.string_of_thm_without_context th; (*Skolemize a named theorem, with Skolem functions as additional premises.*) fun skolem_thm (s, th) = diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/res_clause.ML --- a/src/HOL/Tools/res_clause.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/res_clause.ML Thu Jul 23 21:13:21 2009 +0200 @@ -110,8 +110,8 @@ (@{const_name "op -->"}, "implies"), (@{const_name Set.empty}, "emptyset"), (@{const_name "op :"}, "in"), - (@{const_name Un}, "union"), - (@{const_name Int}, "inter"), + (@{const_name union}, "union"), + (@{const_name inter}, "inter"), ("List.append", "append"), ("ATP_Linkup.fequal", "fequal"), ("ATP_Linkup.COMBI", "COMBI"), diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/res_reconstruct.ML --- a/src/HOL/Tools/res_reconstruct.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/res_reconstruct.ML Thu Jul 23 21:13:21 2009 +0200 @@ -31,7 +31,7 @@ fun trace s = if !Output.debugging then File.append (File.tmp_path trace_path) s else (); -val string_of_thm = PrintMode.setmp [] Display.string_of_thm; +fun string_of_thm ctxt = PrintMode.setmp [] (Display.string_of_thm ctxt); (*For generating structured proofs: keep every nth proof line*) val (modulus, modulus_setup) = Attrib.config_int "sledgehammer_modulus" 1; @@ -445,8 +445,9 @@ val (ccls,fixes) = ResAxioms.neg_conjecture_clauses th sgno val _ = trace (Int.toString (length ccls) ^ " conjecture clauses\n") val ccls = map forall_intr_vars ccls - val _ = if !Output.debugging then app (fn th => trace ("\nccl: " ^ string_of_thm th)) ccls - else () + val _ = + if !Output.debugging then app (fn th => trace ("\nccl: " ^ string_of_thm ctxt th)) ccls + else () val ilines = isar_lines ctxt (map prop_of ccls) (stringify_deps thm_names [] lines) val _ = trace "\ndecode_tstp_file: finishing\n" in diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/rewrite_hol_proof.ML --- a/src/HOL/Tools/rewrite_hol_proof.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/rewrite_hol_proof.ML Thu Jul 23 21:13:21 2009 +0200 @@ -15,8 +15,8 @@ open Proofterm; -val rews = map (pairself (ProofSyntax.proof_of_term (the_context ()) true) o - Logic.dest_equals o Logic.varify o ProofSyntax.read_term (the_context ()) propT) +val rews = map (pairself (ProofSyntax.proof_of_term @{theory} true) o + Logic.dest_equals o Logic.varify o ProofSyntax.read_term @{theory} propT) (** eliminate meta-equality rules **) diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/sat_funcs.ML --- a/src/HOL/Tools/sat_funcs.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/sat_funcs.ML Thu Jul 23 21:13:21 2009 +0200 @@ -66,16 +66,10 @@ val counter = ref 0; -val resolution_thm = (* "[| ?P ==> False; ~ ?P ==> False |] ==> False" *) - let - val cterm = cterm_of (the_context ()) - val Q = Var (("Q", 0), HOLogic.boolT) - val False = HOLogic.false_const - in - Thm.instantiate ([], [(cterm Q, cterm False)]) @{thm case_split} - end; +val resolution_thm = + @{lemma "(P ==> False) ==> (~ P ==> False) ==> False" by (rule case_split)} -val cP = cterm_of (theory_of_thm resolution_thm) (Var (("P", 0), HOLogic.boolT)); +val cP = cterm_of @{theory} (Var (("P", 0), HOLogic.boolT)); (* ------------------------------------------------------------------------- *) (* lit_ord: an order on integers that considers their absolute values only, *) @@ -125,7 +119,7 @@ (* (Thm.thm * (int * Thm.cterm) list) list -> Thm.thm * (int * Thm.cterm) list *) fun resolve_raw_clauses [] = - raise THM ("Proof reconstruction failed (empty list of resolvents)!", 0, []) + raise THM ("Proof reconstruction failed (empty list of resolvents)!", 0, []) | resolve_raw_clauses (c::cs) = let (* merges two sorted lists wrt. 'lit_ord', suppressing duplicates *) @@ -140,9 +134,9 @@ (* find out which two hyps are used in the resolution *) (* (int * Thm.cterm) list * (int * Thm.cterm) list -> (int * Thm.cterm) list -> bool * Thm.cterm * Thm.cterm * (int * Thm.cterm) list *) fun find_res_hyps ([], _) _ = - raise THM ("Proof reconstruction failed (no literal for resolution)!", 0, []) + raise THM ("Proof reconstruction failed (no literal for resolution)!", 0, []) | find_res_hyps (_, []) _ = - raise THM ("Proof reconstruction failed (no literal for resolution)!", 0, []) + raise THM ("Proof reconstruction failed (no literal for resolution)!", 0, []) | find_res_hyps (h1 :: hyps1, h2 :: hyps2) acc = (case (lit_ord o pairself fst) (h1, h2) of LESS => find_res_hyps (hyps1, h2 :: hyps2) (h1 :: acc) @@ -160,9 +154,12 @@ (* Thm.thm * (int * Thm.cterm) list -> Thm.thm * (int * Thm.cterm) list -> Thm.thm * (int * Thm.cterm) list *) fun resolution (c1, hyps1) (c2, hyps2) = let - val _ = if !trace_sat then - tracing ("Resolving clause: " ^ Display.string_of_thm c1 ^ " (hyps: " ^ ML_Syntax.print_list (Syntax.string_of_term_global (theory_of_thm c1)) (#hyps (rep_thm c1)) - ^ ")\nwith clause: " ^ Display.string_of_thm c2 ^ " (hyps: " ^ ML_Syntax.print_list (Syntax.string_of_term_global (theory_of_thm c2)) (#hyps (rep_thm c2)) ^ ")") + val _ = + if ! trace_sat then + tracing ("Resolving clause: " ^ Display.string_of_thm_without_context c1 ^ + " (hyps: " ^ ML_Syntax.print_list (Syntax.string_of_term_global (theory_of_thm c1)) (#hyps (rep_thm c1)) + ^ ")\nwith clause: " ^ Display.string_of_thm_without_context c2 ^ + " (hyps: " ^ ML_Syntax.print_list (Syntax.string_of_term_global (theory_of_thm c2)) (#hyps (rep_thm c2)) ^ ")") else () (* the two literals used for resolution *) @@ -178,8 +175,9 @@ Thm.instantiate ([], [(cP, cLit)]) resolution_thm end - val _ = if !trace_sat then - tracing ("Resolution theorem: " ^ Display.string_of_thm res_thm) + val _ = + if !trace_sat then + tracing ("Resolution theorem: " ^ Display.string_of_thm_without_context res_thm) else () (* Gamma1, Gamma2 |- False *) @@ -187,8 +185,11 @@ (Thm.implies_elim res_thm (if hyp1_is_neg then c2' else c1')) (if hyp1_is_neg then c1' else c2') - val _ = if !trace_sat then - tracing ("Resulting clause: " ^ Display.string_of_thm c_new ^ " (hyps: " ^ ML_Syntax.print_list (Syntax.string_of_term_global (theory_of_thm c_new)) (#hyps (rep_thm c_new)) ^ ")") + val _ = + if !trace_sat then + tracing ("Resulting clause: " ^ Display.string_of_thm_without_context c_new ^ + " (hyps: " ^ ML_Syntax.print_list + (Syntax.string_of_term_global (theory_of_thm c_new)) (#hyps (rep_thm c_new)) ^ ")") else () val _ = inc counter in diff -r 63686057cbe8 -r abda97d2deea src/HOL/Tools/simpdata.ML --- a/src/HOL/Tools/simpdata.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Tools/simpdata.ML Thu Jul 23 21:13:21 2009 +0200 @@ -158,7 +158,7 @@ open Clasimp; val _ = ML_Antiquote.value "clasimpset" - (Scan.succeed "Clasimp.local_clasimpset_of (ML_Context.the_local_context ())"); + (Scan.succeed "Clasimp.clasimpset_of (ML_Context.the_local_context ())"); val mksimps_pairs = [("op -->", [@{thm mp}]), ("op &", [@{thm conjunct1}, @{thm conjunct2}]), @@ -181,11 +181,11 @@ in fn ss => ALLGOALS (full_simp_tac (Simplifier.inherit_context ss ss0)) end; val defALL_regroup = - Simplifier.simproc (the_context ()) + Simplifier.simproc @{theory} "defined ALL" ["ALL x. P x"] Quantifier1.rearrange_all; val defEX_regroup = - Simplifier.simproc (the_context ()) + Simplifier.simproc @{theory} "defined EX" ["EX x. P x"] Quantifier1.rearrange_ex; diff -r 63686057cbe8 -r abda97d2deea src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Transcendental.thy Thu Jul 23 21:13:21 2009 +0200 @@ -621,19 +621,19 @@ subsection{* Some properties of factorials *} -lemma real_of_nat_fact_not_zero [simp]: "real (fact n) \ 0" +lemma real_of_nat_fact_not_zero [simp]: "real (fact (n::nat)) \ 0" by auto -lemma real_of_nat_fact_gt_zero [simp]: "0 < real(fact n)" +lemma real_of_nat_fact_gt_zero [simp]: "0 < real(fact (n::nat))" by auto -lemma real_of_nat_fact_ge_zero [simp]: "0 \ real(fact n)" +lemma real_of_nat_fact_ge_zero [simp]: "0 \ real(fact (n::nat))" by simp -lemma inv_real_of_nat_fact_gt_zero [simp]: "0 < inverse (real (fact n))" +lemma inv_real_of_nat_fact_gt_zero [simp]: "0 < inverse (real (fact (n::nat)))" by (auto simp add: positive_imp_inverse_positive) -lemma inv_real_of_nat_fact_ge_zero [simp]: "0 \ inverse (real (fact n))" +lemma inv_real_of_nat_fact_ge_zero [simp]: "0 \ inverse (real (fact (n::nat)))" by (auto intro: order_less_imp_le) subsection {* Derivability of power series *} @@ -1670,7 +1670,7 @@ apply (rule_tac y = "\n=0..< Suc(Suc(Suc 0)). - (-1 ^ n / (real(fact (2*n))) * 2 ^ (2*n))" in order_less_trans) -apply (simp (no_asm) add: fact_num_eq_if realpow_num_eq_if del: fact_Suc) +apply (simp (no_asm) add: fact_num_eq_if_nat realpow_num_eq_if del: fact_Suc) apply (simp (no_asm) add: mult_assoc del: setsum_op_ivl_Suc) apply (rule sumr_pos_lt_pair) apply (erule sums_summable, safe) @@ -1687,7 +1687,7 @@ apply (rule_tac [3] real_of_nat_ge_zero) prefer 2 apply force apply (rule real_of_nat_less_iff [THEN iffD2]) -apply (rule fact_less_mono, auto) +apply (rule fact_less_mono_nat, auto) done lemmas cos_two_neq_zero [simp] = cos_two_less_zero [THEN less_imp_neq] diff -r 63686057cbe8 -r abda97d2deea src/HOL/UNITY/Comp/Alloc.thy --- a/src/HOL/UNITY/Comp/Alloc.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/UNITY/Comp/Alloc.thy Thu Jul 23 21:13:21 2009 +0200 @@ -321,7 +321,7 @@ *} method_setup record_auto = {* - Scan.succeed (fn ctxt => SIMPLE_METHOD (record_auto_tac (local_clasimpset_of ctxt))) + Scan.succeed (fn ctxt => SIMPLE_METHOD (record_auto_tac (clasimpset_of ctxt))) *} "" lemma inj_sysOfAlloc [iff]: "inj sysOfAlloc" @@ -714,7 +714,7 @@ method_setup rename_client_map = {* Scan.succeed (fn ctxt => - SIMPLE_METHOD (rename_client_map_tac (local_simpset_of ctxt))) + SIMPLE_METHOD (rename_client_map_tac (simpset_of ctxt))) *} "" text{*Lifting @{text Client_Increasing} to @{term systemState}*} @@ -1021,7 +1021,7 @@ LeadsTo {s. h pfixLe (sub i o allocGiv) s})" apply (simp only: o_apply sub_def) apply (insert Alloc_Progress [THEN rename_guarantees_sysOfAlloc_I]) - apply (simp add: o_def del: Set.INT_iff); + apply (simp add: o_def del: INT_iff) apply (erule component_guaranteesD) apply (auto simp add: System_Increasing_allocRel [simplified sub_apply o_def] diff -r 63686057cbe8 -r abda97d2deea src/HOL/UNITY/ProgressSets.thy --- a/src/HOL/UNITY/ProgressSets.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/UNITY/ProgressSets.thy Thu Jul 23 21:13:21 2009 +0200 @@ -44,7 +44,7 @@ lemma UN_in_lattice: "[|lattice L; !!i. i\I ==> r i \ L|] ==> (\i\I. r i) \ L" -apply (simp add: Set.UN_eq) +apply (simp add: UN_eq) apply (blast intro: Union_in_lattice) done diff -r 63686057cbe8 -r abda97d2deea src/HOL/UNITY/Simple/NSP_Bad.thy --- a/src/HOL/UNITY/Simple/NSP_Bad.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/UNITY/Simple/NSP_Bad.thy Thu Jul 23 21:13:21 2009 +0200 @@ -125,7 +125,7 @@ method_setup ns_induct = {* Scan.succeed (fn ctxt => - SIMPLE_METHOD' (ns_induct_tac (local_clasimpset_of ctxt))) *} + SIMPLE_METHOD' (ns_induct_tac (clasimpset_of ctxt))) *} "for inductive reasoning about the Needham-Schroeder protocol" text{*Converts invariants into statements about reachable states*} diff -r 63686057cbe8 -r abda97d2deea src/HOL/UNITY/UNITY_Main.thy --- a/src/HOL/UNITY/UNITY_Main.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/UNITY/UNITY_Main.thy Thu Jul 23 21:13:21 2009 +0200 @@ -11,12 +11,12 @@ method_setup safety = {* Scan.succeed (fn ctxt => - SIMPLE_METHOD' (constrains_tac (local_clasimpset_of ctxt))) *} + SIMPLE_METHOD' (constrains_tac (clasimpset_of ctxt))) *} "for proving safety properties" method_setup ensures_tac = {* Args.goal_spec -- Scan.lift Args.name_source >> - (fn (quant, s) => fn ctxt => SIMPLE_METHOD'' quant (ensures_tac (local_clasimpset_of ctxt) s)) + (fn (quant, s) => fn ctxt => SIMPLE_METHOD'' quant (ensures_tac (clasimpset_of ctxt) s)) *} "for proving progress properties" end diff -r 63686057cbe8 -r abda97d2deea src/HOL/Word/WordArith.thy --- a/src/HOL/Word/WordArith.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/Word/WordArith.thy Thu Jul 23 21:13:21 2009 +0200 @@ -513,8 +513,8 @@ fun uint_arith_tacs ctxt = let fun arith_tac' n t = Arith_Data.verbose_arith_tac ctxt n t handle COOPER => Seq.empty; - val cs = local_claset_of ctxt; - val ss = local_simpset_of ctxt; + val cs = claset_of ctxt; + val ss = simpset_of ctxt; in [ clarify_tac cs 1, full_simp_tac (uint_arith_ss_of ss) 1, @@ -1075,8 +1075,8 @@ fun unat_arith_tacs ctxt = let fun arith_tac' n t = Arith_Data.verbose_arith_tac ctxt n t handle COOPER => Seq.empty; - val cs = local_claset_of ctxt; - val ss = local_simpset_of ctxt; + val cs = claset_of ctxt; + val ss = simpset_of ctxt; in [ clarify_tac cs 1, full_simp_tac (unat_arith_ss_of ss) 1, diff -r 63686057cbe8 -r abda97d2deea src/HOL/ex/Eval_Examples.thy --- a/src/HOL/ex/Eval_Examples.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/ex/Eval_Examples.thy Thu Jul 23 21:13:21 2009 +0200 @@ -3,7 +3,7 @@ header {* Small examples for evaluation mechanisms *} theory Eval_Examples -imports Code_Eval Rational +imports Complex_Main begin text {* evaluation oracle *} diff -r 63686057cbe8 -r abda97d2deea src/HOL/ex/Meson_Test.thy --- a/src/HOL/ex/Meson_Test.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/ex/Meson_Test.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1,4 +1,3 @@ -(*ID: $Id$*) header {* Meson test cases *} @@ -11,7 +10,7 @@ below and constants declared in HOL! *} -hide const subset member quotient +hide (open) const subset member quotient union inter text {* Test data for the MESON proof procedure diff -r 63686057cbe8 -r abda97d2deea src/HOL/ex/Numeral.thy --- a/src/HOL/ex/Numeral.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/ex/Numeral.thy Thu Jul 23 21:13:21 2009 +0200 @@ -758,7 +758,7 @@ code_datatype "0::int" Pls Mns -lemmas [code_inline] = Pls_def [symmetric] Mns_def [symmetric] +lemmas [code_unfold] = Pls_def [symmetric] Mns_def [symmetric] definition sub :: "num \ num \ int" where [simp, code del]: "sub m n = (of_num m - of_num n)" @@ -874,10 +874,10 @@ using of_num_pos [of l, where ?'a = int] of_num_pos [of k, where ?'a = int] by (simp_all add: of_num_less_iff) -lemma [code_inline del]: "(0::int) \ Numeral0" by simp -lemma [code_inline del]: "(1::int) \ Numeral1" by simp -declare zero_is_num_zero [code_inline del] -declare one_is_num_one [code_inline del] +lemma [code_unfold del]: "(0::int) \ Numeral0" by simp +lemma [code_unfold del]: "(1::int) \ Numeral1" by simp +declare zero_is_num_zero [code_unfold del] +declare one_is_num_one [code_unfold del] hide (open) const sub dup diff -r 63686057cbe8 -r abda97d2deea src/HOL/ex/ROOT.ML --- a/src/HOL/ex/ROOT.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/ex/ROOT.ML Thu Jul 23 21:13:21 2009 +0200 @@ -83,6 +83,7 @@ (* requires a proof-generating SAT solver (zChaff or MiniSAT) to be *) (* installed: *) try use_thy "SAT_Examples"; +Future.shutdown (); (* requires zChaff (or some other reasonably fast SAT solver) to be *) (* installed: *) diff -r 63686057cbe8 -r abda97d2deea src/HOL/ex/predicate_compile.ML --- a/src/HOL/ex/predicate_compile.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOL/ex/predicate_compile.ML Thu Jul 23 21:13:21 2009 +0200 @@ -236,11 +236,11 @@ val _ = writeln ("predicate: " ^ pred) val _ = writeln ("number of parameters: " ^ string_of_int (nparams_of thy pred)) val _ = writeln ("introrules: ") - val _ = fold (fn thm => fn u => writeln (Display.string_of_thm thm)) + val _ = fold (fn thm => fn u => writeln (Display.string_of_thm_global thy thm)) (rev (intros_of thy pred)) () in if (has_elim thy pred) then - writeln ("elimrule: " ^ Display.string_of_thm (the_elim_of thy pred)) + writeln ("elimrule: " ^ Display.string_of_thm_global thy (the_elim_of thy pred)) else writeln ("no elimrule defined") end diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/Domain.thy --- a/src/HOLCF/Domain.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOLCF/Domain.thy Thu Jul 23 21:13:21 2009 +0200 @@ -9,11 +9,11 @@ uses ("Tools/cont_consts.ML") ("Tools/cont_proc.ML") - ("Tools/domain/domain_library.ML") - ("Tools/domain/domain_syntax.ML") - ("Tools/domain/domain_axioms.ML") - ("Tools/domain/domain_theorems.ML") - ("Tools/domain/domain_extender.ML") + ("Tools/Domain/domain_library.ML") + ("Tools/Domain/domain_syntax.ML") + ("Tools/Domain/domain_axioms.ML") + ("Tools/Domain/domain_theorems.ML") + ("Tools/Domain/domain_extender.ML") begin defaultsort pcpo @@ -274,10 +274,10 @@ use "Tools/cont_consts.ML" use "Tools/cont_proc.ML" -use "Tools/domain/domain_library.ML" -use "Tools/domain/domain_syntax.ML" -use "Tools/domain/domain_axioms.ML" -use "Tools/domain/domain_theorems.ML" -use "Tools/domain/domain_extender.ML" +use "Tools/Domain/domain_library.ML" +use "Tools/Domain/domain_syntax.ML" +use "Tools/Domain/domain_axioms.ML" +use "Tools/Domain/domain_theorems.ML" +use "Tools/Domain/domain_extender.ML" end diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/FOCUS/Buffer.thy --- a/src/HOLCF/FOCUS/Buffer.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOLCF/FOCUS/Buffer.thy Thu Jul 23 21:13:21 2009 +0200 @@ -101,7 +101,7 @@ fun B_prover s tac eqs = let val thy = the_context () in prove_goal thy s (fn prems => [cut_facts_tac prems 1, - tac 1, auto_tac (claset_of thy, simpset_of thy addsimps eqs)]) + tac 1, auto_tac (global_claset_of thy, global_simpset_of thy addsimps eqs)]) end; fun prove_forw s thm = B_prover s (dtac (thm RS iffD1)) []; diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/IOA/Modelcheck/MuIOA.thy --- a/src/HOLCF/IOA/Modelcheck/MuIOA.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOLCF/IOA/Modelcheck/MuIOA.thy Thu Jul 23 21:13:21 2009 +0200 @@ -278,14 +278,14 @@ by (REPEAT (rtac impI 1)); by (REPEAT (dtac eq_reflection 1)); (* Bis hierher wird im Kapitel 4 erl"autert, ab hier im Kapitel 5 *) -by (full_simp_tac ((simpset_of sign delcongs (if_weak_cong :: weak_case_congs) +by (full_simp_tac ((global_simpset_of sign delcongs (if_weak_cong :: weak_case_congs) delsimps [not_iff,split_part]) addsimps (thl @ comp_simps @ restrict_simps @ hide_simps @ rename_simps @ ioa_simps @ asig_simps)) 1); by (full_simp_tac (Mucke_ss delcongs (if_weak_cong :: weak_case_congs) delsimps [not_iff,split_part]) 1); by (REPEAT (if_full_simp_tac - (simpset_of sign delcongs (if_weak_cong :: weak_case_congs) delsimps [not_iff,split_part]) 1)); + (global_simpset_of sign delcongs (if_weak_cong :: weak_case_congs) delsimps [not_iff,split_part]) 1)); by (call_mucke_tac 1); (* Bis hierher wird im Kapitel 5 erl"autert, ab hier wieder im Kapitel 4 *) by (atac 1); diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/IOA/meta_theory/Abstraction.thy --- a/src/HOLCF/IOA/meta_theory/Abstraction.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOLCF/IOA/meta_theory/Abstraction.thy Thu Jul 23 21:13:21 2009 +0200 @@ -605,7 +605,7 @@ ML {* fun abstraction_tac ctxt = - let val (cs, ss) = local_clasimpset_of ctxt in + let val (cs, ss) = clasimpset_of ctxt in SELECT_GOAL (auto_tac (cs addSIs @{thms weak_strength_lemmas}, ss addsimps [@{thm state_strengthening_def}, @{thm state_weakening_def}])) end diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/IOA/meta_theory/Automata.thy --- a/src/HOLCF/IOA/meta_theory/Automata.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOLCF/IOA/meta_theory/Automata.thy Thu Jul 23 21:13:21 2009 +0200 @@ -399,7 +399,7 @@ ==> input_enabled (A||B)" apply (unfold input_enabled_def) apply (simp add: Let_def inputs_of_par trans_of_par) -apply (tactic "safe_tac (claset_of @{theory Fun})") +apply (tactic "safe_tac (global_claset_of @{theory Fun})") apply (simp add: inp_is_act) prefer 2 apply (simp add: inp_is_act) diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/IOA/meta_theory/CompoScheds.thy --- a/src/HOLCF/IOA/meta_theory/CompoScheds.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOLCF/IOA/meta_theory/CompoScheds.thy Thu Jul 23 21:13:21 2009 +0200 @@ -299,10 +299,10 @@ in fun mkex_induct_tac ctxt sch exA exB = - let val ss = Simplifier.local_simpset_of ctxt in + let val ss = simpset_of ctxt in EVERY1[Seq_induct_tac ctxt sch defs, asm_full_simp_tac ss, - SELECT_GOAL (safe_tac (claset_of @{theory Fun})), + SELECT_GOAL (safe_tac (global_claset_of @{theory Fun})), Seq_case_simp_tac ctxt exA, Seq_case_simp_tac ctxt exB, asm_full_simp_tac ss, diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/IOA/meta_theory/Sequence.thy --- a/src/HOLCF/IOA/meta_theory/Sequence.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOLCF/IOA/meta_theory/Sequence.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1091,7 +1091,7 @@ (* on a>>s only simp_tac, as full_simp_tac is uncomplete and often causes errors *) fun Seq_case_simp_tac ctxt s i = - let val ss = Simplifier.local_simpset_of ctxt in + let val ss = simpset_of ctxt in Seq_case_tac ctxt s i THEN asm_simp_tac ss (i+2) THEN asm_full_simp_tac ss (i+1) @@ -1100,7 +1100,7 @@ (* rws are definitions to be unfolded for admissibility check *) fun Seq_induct_tac ctxt s rws i = - let val ss = Simplifier.local_simpset_of ctxt in + let val ss = simpset_of ctxt in res_inst_tac ctxt [(("x", 0), s)] @{thm Seq_induct} i THEN (REPEAT_DETERM (CHANGED (asm_simp_tac ss (i+1)))) THEN simp_tac (ss addsimps rws) i @@ -1108,15 +1108,15 @@ fun Seq_Finite_induct_tac ctxt i = etac @{thm Seq_Finite_ind} i - THEN (REPEAT_DETERM (CHANGED (asm_simp_tac (Simplifier.local_simpset_of ctxt) i))); + THEN (REPEAT_DETERM (CHANGED (asm_simp_tac (simpset_of ctxt) i))); fun pair_tac ctxt s = res_inst_tac ctxt [(("p", 0), s)] PairE - THEN' hyp_subst_tac THEN' asm_full_simp_tac (Simplifier.local_simpset_of ctxt); + THEN' hyp_subst_tac THEN' asm_full_simp_tac (simpset_of ctxt); (* induction on a sequence of pairs with pairsplitting and simplification *) fun pair_induct_tac ctxt s rws i = - let val ss = Simplifier.local_simpset_of ctxt in + let val ss = simpset_of ctxt in res_inst_tac ctxt [(("x", 0), s)] @{thm Seq_induct} i THEN pair_tac ctxt "a" (i+3) THEN (REPEAT_DETERM (CHANGED (simp_tac ss (i+1)))) diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/IsaMakefile --- a/src/HOLCF/IsaMakefile Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOLCF/IsaMakefile Thu Jul 23 21:13:21 2009 +0200 @@ -62,11 +62,11 @@ Tools/adm_tac.ML \ Tools/cont_consts.ML \ Tools/cont_proc.ML \ - Tools/domain/domain_extender.ML \ - Tools/domain/domain_axioms.ML \ - Tools/domain/domain_library.ML \ - Tools/domain/domain_syntax.ML \ - Tools/domain/domain_theorems.ML \ + Tools/Domain/domain_extender.ML \ + Tools/Domain/domain_axioms.ML \ + Tools/Domain/domain_library.ML \ + Tools/Domain/domain_syntax.ML \ + Tools/Domain/domain_theorems.ML \ Tools/fixrec.ML \ Tools/pcpodef.ML \ holcf_logic.ML \ diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/Lift.thy --- a/src/HOLCF/Lift.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOLCF/Lift.thy Thu Jul 23 21:13:21 2009 +0200 @@ -61,7 +61,7 @@ method_setup defined = {* Scan.succeed (fn ctxt => SIMPLE_METHOD' - (etac @{thm lift_definedE} THEN' asm_simp_tac (local_simpset_of ctxt))) + (etac @{thm lift_definedE} THEN' asm_simp_tac (simpset_of ctxt))) *} "" lemma DefE: "Def x = \ \ R" diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/Tools/Domain/domain_axioms.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/Tools/Domain/domain_axioms.ML Thu Jul 23 21:13:21 2009 +0200 @@ -0,0 +1,235 @@ +(* Title: HOLCF/Tools/Domain/domain_axioms.ML + Author: David von Oheimb + +Syntax generator for domain command. +*) + +signature DOMAIN_AXIOMS = +sig + val copy_of_dtyp : (int -> term) -> Datatype.dtyp -> term + + val calc_axioms : + string -> Domain_Library.eq list -> int -> Domain_Library.eq -> + string * (string * term) list * (string * term) list + + val add_axioms : + bstring -> Domain_Library.eq list -> theory -> theory +end; + + +structure Domain_Axioms :> DOMAIN_AXIOMS = +struct + +open Domain_Library; + +infixr 0 ===>;infixr 0 ==>;infix 0 == ; +infix 1 ===; infix 1 ~= ; infix 1 <<; infix 1 ~<<; +infix 9 ` ; infix 9 `% ; infix 9 `%%; infixr 9 oo; + +(* FIXME: use theory data for this *) +val copy_tab : string Symtab.table = + Symtab.make [(@{type_name "->"}, @{const_name "cfun_fun"}), + (@{type_name "++"}, @{const_name "ssum_fun"}), + (@{type_name "**"}, @{const_name "sprod_fun"}), + (@{type_name "*"}, @{const_name "cprod_fun"}), + (@{type_name "u"}, @{const_name "u_fun"})]; + +fun copy_of_dtyp r dt = if DatatypeAux.is_rec_type dt then copy r dt else ID +and copy r (DatatypeAux.DtRec i) = r i + | copy r (DatatypeAux.DtTFree a) = ID + | copy r (DatatypeAux.DtType (c, ds)) = + case Symtab.lookup copy_tab c of + SOME f => list_ccomb (%%:f, map (copy_of_dtyp r) ds) + | NONE => (warning ("copy_of_dtyp: unknown type constructor " ^ c); ID); + +fun calc_axioms + (comp_dname : string) + (eqs : eq list) + (n : int) + (eqn as ((dname,_),cons) : eq) + : string * (string * term) list * (string * term) list = + let + + (* ----- axioms and definitions concerning the isomorphism ------------------ *) + + val dc_abs = %%:(dname^"_abs"); + val dc_rep = %%:(dname^"_rep"); + val x_name'= "x"; + val x_name = idx_name eqs x_name' (n+1); + val dnam = Long_Name.base_name dname; + + val abs_iso_ax = ("abs_iso", mk_trp(dc_rep`(dc_abs`%x_name') === %:x_name')); + val rep_iso_ax = ("rep_iso", mk_trp(dc_abs`(dc_rep`%x_name') === %:x_name')); + + val when_def = ("when_def",%%:(dname^"_when") == + List.foldr (uncurry /\ ) (/\x_name'((when_body cons (fn (x,y) => + Bound(1+length cons+x-y)))`(dc_rep`Bound 0))) (when_funs cons)); + + val copy_def = + let fun r i = cproj (Bound 0) eqs i; + in ("copy_def", %%:(dname^"_copy") == + /\ "f" (dc_abs oo (copy_of_dtyp r (dtyp_of_eq eqn)) oo dc_rep)) end; + + (* -- definitions concerning the constructors, discriminators and selectors - *) + + fun con_def m n (_,args) = let + fun idxs z x arg = (if is_lazy arg then mk_up else I) (Bound(z-x)); + fun parms vs = mk_stuple (mapn (idxs(length vs)) 1 vs); + fun inj y 1 _ = y + | inj y _ 0 = mk_sinl y + | inj y i j = mk_sinr (inj y (i-1) (j-1)); + in List.foldr /\# (dc_abs`(inj (parms args) m n)) args end; + + val con_defs = mapn (fn n => fn (con,args) => + (extern_name con ^"_def", %%:con == con_def (length cons) n (con,args))) 0 cons; + + val dis_defs = let + fun ddef (con,_) = (dis_name con ^"_def",%%:(dis_name con) == + list_ccomb(%%:(dname^"_when"),map + (fn (con',args) => (List.foldr /\# + (if con'=con then TT else FF) args)) cons)) + in map ddef cons end; + + val mat_defs = + let + fun mdef (con,_) = + let + val k = Bound 0 + val x = Bound 1 + fun one_con (con', args') = + if con'=con then k else List.foldr /\# mk_fail args' + val w = list_ccomb(%%:(dname^"_when"), map one_con cons) + val rhs = /\ "x" (/\ "k" (w ` x)) + in (mat_name con ^"_def", %%:(mat_name con) == rhs) end + in map mdef cons end; + + val pat_defs = + let + fun pdef (con,args) = + let + val ps = mapn (fn n => fn _ => %:("pat" ^ string_of_int n)) 1 args; + val xs = map (bound_arg args) args; + val r = Bound (length args); + val rhs = case args of [] => mk_return HOLogic.unit + | _ => mk_ctuple_pat ps ` mk_ctuple xs; + fun one_con (con',args') = List.foldr /\# (if con'=con then rhs else mk_fail) args'; + in (pat_name con ^"_def", list_comb (%%:(pat_name con), ps) == + list_ccomb(%%:(dname^"_when"), map one_con cons)) + end + in map pdef cons end; + + val sel_defs = let + fun sdef con n arg = Option.map (fn sel => (sel^"_def",%%:sel == + list_ccomb(%%:(dname^"_when"),map + (fn (con',args) => if con'<>con then UU else + List.foldr /\# (Bound (length args - n)) args) cons))) (sel_of arg); + in List.mapPartial I (List.concat(map (fn (con,args) => mapn (sdef con) 1 args) cons)) end; + + + (* ----- axiom and definitions concerning induction ------------------------- *) + + val reach_ax = ("reach", mk_trp(cproj (mk_fix (%%:(comp_dname^"_copy"))) eqs n + `%x_name === %:x_name)); + val take_def = + ("take_def", + %%:(dname^"_take") == + mk_lam("n",cproj + (mk_iterate (Bound 0, %%:(comp_dname^"_copy"), UU)) eqs n)); + val finite_def = + ("finite_def", + %%:(dname^"_finite") == + mk_lam(x_name, + mk_ex("n",(%%:(dname^"_take") $ Bound 0)`Bound 1 === Bound 1))); + + in (dnam, + [abs_iso_ax, rep_iso_ax, reach_ax], + [when_def, copy_def] @ + con_defs @ dis_defs @ mat_defs @ pat_defs @ sel_defs @ + [take_def, finite_def]) + end; (* let (calc_axioms) *) + + +(* legacy type inference *) + +fun legacy_infer_term thy t = + singleton (Syntax.check_terms (ProofContext.init thy)) (Sign.intern_term thy t); + +fun legacy_infer_prop thy t = legacy_infer_term thy (TypeInfer.constrain propT t); + +fun infer_props thy = map (apsnd (legacy_infer_prop thy)); + + +fun add_axioms_i x = snd o PureThy.add_axioms (map (Thm.no_attributes o apfst Binding.name) x); +fun add_axioms_infer axms thy = add_axioms_i (infer_props thy axms) thy; + +fun add_defs_i x = snd o (PureThy.add_defs false) (map (Thm.no_attributes o apfst Binding.name) x); +fun add_defs_infer defs thy = add_defs_i (infer_props thy defs) thy; + +fun add_matchers (((dname,_),cons) : eq) thy = + let + val con_names = map fst cons; + val mat_names = map mat_name con_names; + fun qualify n = Sign.full_name thy (Binding.name n); + val ms = map qualify con_names ~~ map qualify mat_names; + in Fixrec.add_matchers ms thy end; + +fun add_axioms comp_dnam (eqs : eq list) thy' = + let + val comp_dname = Sign.full_bname thy' comp_dnam; + val dnames = map (fst o fst) eqs; + val x_name = idx_name dnames "x"; + fun copy_app dname = %%:(dname^"_copy")`Bound 0; + val copy_def = ("copy_def" , %%:(comp_dname^"_copy") == + /\ "f"(mk_ctuple (map copy_app dnames))); + + fun one_con (con,args) = let + val nonrec_args = filter_out is_rec args; + val rec_args = List.filter is_rec args; + val recs_cnt = length rec_args; + val allargs = nonrec_args @ rec_args + @ map (upd_vname (fn s=> s^"'")) rec_args; + val allvns = map vname allargs; + fun vname_arg s arg = if is_rec arg then vname arg^s else vname arg; + val vns1 = map (vname_arg "" ) args; + val vns2 = map (vname_arg "'") args; + val allargs_cnt = length nonrec_args + 2*recs_cnt; + val rec_idxs = (recs_cnt-1) downto 0; + val nonlazy_idxs = map snd (filter_out (fn (arg,_) => is_lazy arg) + (allargs~~((allargs_cnt-1) downto 0))); + fun rel_app i ra = proj (Bound(allargs_cnt+2)) eqs (rec_of ra) $ + Bound (2*recs_cnt-i) $ Bound (recs_cnt-i); + val capps = + List.foldr mk_conj + (mk_conj( + Bound(allargs_cnt+1)===list_ccomb(%%:con,map (bound_arg allvns) vns1), + Bound(allargs_cnt+0)===list_ccomb(%%:con,map (bound_arg allvns) vns2))) + (mapn rel_app 1 rec_args); + in List.foldr mk_ex + (Library.foldr mk_conj + (map (defined o Bound) nonlazy_idxs,capps)) allvns + end; + fun one_comp n (_,cons) = + mk_all(x_name(n+1), + mk_all(x_name(n+1)^"'", + mk_imp(proj (Bound 2) eqs n $ Bound 1 $ Bound 0, + foldr1 mk_disj (mk_conj(Bound 1 === UU,Bound 0 === UU) + ::map one_con cons)))); + val bisim_def = + ("bisim_def", + %%:(comp_dname^"_bisim")==mk_lam("R", foldr1 mk_conj (mapn one_comp 0 eqs))); + + fun add_one (thy,(dnam,axs,dfs)) = + thy |> Sign.add_path dnam + |> add_defs_infer dfs + |> add_axioms_infer axs + |> Sign.parent_path; + + val thy = Library.foldl add_one (thy', mapn (calc_axioms comp_dname eqs) 0 eqs); + + in thy |> Sign.add_path comp_dnam + |> add_defs_infer (bisim_def::(if length eqs>1 then [copy_def] else [])) + |> Sign.parent_path + |> fold add_matchers eqs + end; (* let (add_axioms) *) + +end; (* struct *) diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/Tools/Domain/domain_extender.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/Tools/Domain/domain_extender.ML Thu Jul 23 21:13:21 2009 +0200 @@ -0,0 +1,204 @@ +(* Title: HOLCF/Tools/Domain/domain_extender.ML + Author: David von Oheimb + +Theory extender for domain command, including theory syntax. +*) + +signature DOMAIN_EXTENDER = +sig + val add_domain_cmd: string -> + ((string * string option) list * binding * mixfix * + (binding * (bool * binding option * string) list * mixfix) list) list + -> theory -> theory + val add_domain: string -> + ((string * string option) list * binding * mixfix * + (binding * (bool * binding option * typ) list * mixfix) list) list + -> theory -> theory +end; + +structure Domain_Extender :> DOMAIN_EXTENDER = +struct + +open Domain_Library; + +(* ----- general testing and preprocessing of constructor list -------------- *) +fun check_and_sort_domain + (dtnvs : (string * typ list) list) + (cons'' : (binding * (bool * binding option * typ) list * mixfix) list list) + (sg : theory) + : ((string * typ list) * + (binding * (bool * binding option * typ) list * mixfix) list) list = + let + val defaultS = Sign.defaultS sg; + val test_dupl_typs = (case duplicates (op =) (map fst dtnvs) of + [] => false | dups => error ("Duplicate types: " ^ commas_quote dups)); + val test_dupl_cons = + (case duplicates (op =) (map (Binding.name_of o first) (List.concat cons'')) of + [] => false | dups => error ("Duplicate constructors: " + ^ commas_quote dups)); + val test_dupl_sels = + (case duplicates (op =) (map Binding.name_of (List.mapPartial second + (List.concat (map second (List.concat cons''))))) of + [] => false | dups => error("Duplicate selectors: "^commas_quote dups)); + val test_dupl_tvars = + exists(fn s=>case duplicates (op =) (map(fst o dest_TFree)s)of + [] => false | dups => error("Duplicate type arguments: " + ^commas_quote dups)) (map snd dtnvs); + (* test for free type variables, illegal sort constraints on rhs, + non-pcpo-types and invalid use of recursive type; + replace sorts in type variables on rhs *) + fun analyse_equation ((dname,typevars),cons') = + let + val tvars = map dest_TFree typevars; + val distinct_typevars = map TFree tvars; + fun rm_sorts (TFree(s,_)) = TFree(s,[]) + | rm_sorts (Type(s,ts)) = Type(s,remove_sorts ts) + | rm_sorts (TVar(s,_)) = TVar(s,[]) + and remove_sorts l = map rm_sorts l; + val indirect_ok = ["*","Cfun.->","Ssum.++","Sprod.**","Up.u"] + fun analyse indirect (TFree(v,s)) = + (case AList.lookup (op =) tvars v of + NONE => error ("Free type variable " ^ quote v ^ " on rhs.") + | SOME sort => if eq_set_string (s,defaultS) orelse + eq_set_string (s,sort ) + then TFree(v,sort) + else error ("Inconsistent sort constraint" ^ + " for type variable " ^ quote v)) + | analyse indirect (t as Type(s,typl)) = + (case AList.lookup (op =) dtnvs s of + NONE => if s mem indirect_ok + then Type(s,map (analyse false) typl) + else Type(s,map (analyse true) typl) + | SOME typevars => if indirect + then error ("Indirect recursion of type " ^ + quote (string_of_typ sg t)) + else if dname <> s orelse + (** BUG OR FEATURE?: + mutual recursion may use different arguments **) + remove_sorts typevars = remove_sorts typl + then Type(s,map (analyse true) typl) + else error ("Direct recursion of type " ^ + quote (string_of_typ sg t) ^ + " with different arguments")) + | analyse indirect (TVar _) = Imposs "extender:analyse"; + fun check_pcpo lazy T = + let val ok = if lazy then cpo_type else pcpo_type + in if ok sg T then T else error + ("Constructor argument type is not of sort pcpo: " ^ + string_of_typ sg T) + end; + fun analyse_arg (lazy, sel, T) = + (lazy, sel, check_pcpo lazy (analyse false T)); + fun analyse_con (b, args, mx) = (b, map analyse_arg args, mx); + in ((dname,distinct_typevars), map analyse_con cons') end; + in ListPair.map analyse_equation (dtnvs,cons'') + end; (* let *) + +(* ----- calls for building new thy and thms -------------------------------- *) + +fun gen_add_domain + (prep_typ : theory -> 'a -> typ) + (comp_dnam : string) + (eqs''' : ((string * string option) list * binding * mixfix * + (binding * (bool * binding option * 'a) list * mixfix) list) list) + (thy''' : theory) = + let + fun readS (SOME s) = Syntax.read_sort_global thy''' s + | readS NONE = Sign.defaultS thy'''; + fun readTFree (a, s) = TFree (a, readS s); + + val dtnvs = map (fn (vs,dname:binding,mx,_) => + (dname, map readTFree vs, mx)) eqs'''; + val cons''' = map (fn (_,_,_,cons) => cons) eqs'''; + fun thy_type (dname,tvars,mx) = (dname, length tvars, mx); + fun thy_arity (dname,tvars,mx) = (Sign.full_name thy''' dname, map (snd o dest_TFree) tvars, pcpoS); + val thy'' = thy''' |> Sign.add_types (map thy_type dtnvs) + |> fold (AxClass.axiomatize_arity o thy_arity) dtnvs; + val cons'' = map (map (upd_second (map (upd_third (prep_typ thy''))))) cons'''; + val dtnvs' = map (fn (dname,vs,mx) => (Sign.full_name thy''' dname,vs)) dtnvs; + val eqs' : ((string * typ list) * (binding * (bool * binding option * typ) list * mixfix) list) list = + check_and_sort_domain dtnvs' cons'' thy''; + val thy' = thy'' |> Domain_Syntax.add_syntax comp_dnam eqs'; + val dts = map (Type o fst) eqs'; + val new_dts = map (fn ((s,Ts),_) => (s, map (fst o dest_TFree) Ts)) eqs'; + fun strip ss = Library.drop (find_index (fn s => s = "'") ss + 1, ss); + fun typid (Type (id,_)) = + let val c = hd (Symbol.explode (Long_Name.base_name id)) + in if Symbol.is_letter c then c else "t" end + | typid (TFree (id,_) ) = hd (strip (tl (Symbol.explode id))) + | typid (TVar ((id,_),_)) = hd (tl (Symbol.explode id)); + fun one_con (con,args,mx) = + ((Syntax.const_name mx (Binding.name_of con)), + ListPair.map (fn ((lazy,sel,tp),vn) => mk_arg ((lazy, + DatatypeAux.dtyp_of_typ new_dts tp), + Option.map Binding.name_of sel,vn)) + (args,(mk_var_names(map (typid o third) args))) + ) : cons; + val eqs = map (fn (dtnvs,cons') => (dtnvs, map one_con cons')) eqs' : eq list; + val thy = thy' |> Domain_Axioms.add_axioms comp_dnam eqs; + val ((rewss, take_rews), theorems_thy) = + thy |> fold_map (fn eq => Domain_Theorems.theorems (eq, eqs)) eqs + ||>> Domain_Theorems.comp_theorems (comp_dnam, eqs); + in + theorems_thy + |> Sign.add_path (Long_Name.base_name comp_dnam) + |> (snd o (PureThy.add_thmss [((Binding.name "rews", List.concat rewss @ take_rews), [])])) + |> Sign.parent_path + end; + +val add_domain = gen_add_domain Sign.certify_typ; +val add_domain_cmd = gen_add_domain Syntax.read_typ_global; + + +(** outer syntax **) + +local structure P = OuterParse and K = OuterKeyword in + +val _ = OuterKeyword.keyword "lazy"; + +val dest_decl : (bool * binding option * string) parser = + P.$$$ "(" |-- Scan.optional (P.$$$ "lazy" >> K true) false -- + (P.binding >> SOME) -- (P.$$$ "::" |-- P.typ) --| P.$$$ ")" >> P.triple1 + || P.$$$ "(" |-- P.$$$ "lazy" |-- P.typ --| P.$$$ ")" + >> (fn t => (true,NONE,t)) + || P.typ >> (fn t => (false,NONE,t)); + +val cons_decl = + P.binding -- Scan.repeat dest_decl -- P.opt_mixfix; + +val type_var' : (string * string option) parser = + (P.type_ident -- Scan.option (P.$$$ "::" |-- P.!!! P.sort)); + +val type_args' : (string * string option) list parser = + type_var' >> single || + P.$$$ "(" |-- P.!!! (P.list1 type_var' --| P.$$$ ")") || + Scan.succeed []; + +val domain_decl = + (type_args' -- P.binding -- P.opt_infix) -- + (P.$$$ "=" |-- P.enum1 "|" cons_decl); + +val domains_decl = + Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- + P.and_list1 domain_decl; + +fun mk_domain (opt_name : string option, + doms : ((((string * string option) list * binding) * mixfix) * + ((binding * (bool * binding option * string) list) * mixfix) list) list ) = + let + val names = map (fn (((_, t), _), _) => Binding.name_of t) doms; + val specs : ((string * string option) list * binding * mixfix * + (binding * (bool * binding option * string) list * mixfix) list) list = + map (fn (((vs, t), mx), cons) => + (vs, t, mx, map (fn ((c, ds), mx) => (c, ds, mx)) cons)) doms; + val comp_dnam = + case opt_name of NONE => space_implode "_" names | SOME s => s; + in add_domain_cmd comp_dnam specs end; + +val _ = + OuterSyntax.command "domain" "define recursive domains (HOLCF)" K.thy_decl + (domains_decl >> (Toplevel.theory o mk_domain)); + +end; + +end; diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/Tools/Domain/domain_library.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/Tools/Domain/domain_library.ML Thu Jul 23 21:13:21 2009 +0200 @@ -0,0 +1,399 @@ +(* Title: HOLCF/Tools/Domain/domain_library.ML + Author: David von Oheimb + +Library for domain command. +*) + + +(* ----- general support ---------------------------------------------------- *) + +fun mapn f n [] = [] + | mapn f n (x::xs) = (f n x) :: mapn f (n+1) xs; + +fun foldr'' f (l,f2) = + let fun itr [] = raise Fail "foldr''" + | itr [a] = f2 a + | itr (a::l) = f(a, itr l) + in itr l end; + +fun map_cumulr f start xs = + List.foldr (fn (x,(ys,res))=>case f(x,res) of (y,res2) => + (y::ys,res2)) ([],start) xs; + +fun first (x,_,_) = x; fun second (_,x,_) = x; fun third (_,_,x) = x; +fun upd_first f (x,y,z) = (f x, y, z); +fun upd_second f (x,y,z) = ( x, f y, z); +fun upd_third f (x,y,z) = ( x, y, f z); + +fun atomize ctxt thm = + let + val r_inst = read_instantiate ctxt; + fun at thm = + case concl_of thm of + _$(Const("op &",_)$_$_) => at(thm RS conjunct1)@at(thm RS conjunct2) + | _$(Const("All" ,_)$Abs(s,_,_))=> at(thm RS (r_inst [(("x", 0), "?" ^ s)] spec)) + | _ => [thm]; + in map zero_var_indexes (at thm) end; + +(* infix syntax *) + +infixr 5 -->; +infixr 6 ->>; +infixr 0 ===>; +infixr 0 ==>; +infix 0 ==; +infix 1 ===; +infix 1 ~=; +infix 1 <<; +infix 1 ~<<; + +infix 9 ` ; +infix 9 `% ; +infix 9 `%%; + + +(* ----- specific support for domain ---------------------------------------- *) + +signature DOMAIN_LIBRARY = +sig + val Imposs : string -> 'a; + val cpo_type : theory -> typ -> bool; + val pcpo_type : theory -> typ -> bool; + val string_of_typ : theory -> typ -> string; + + (* Creating HOLCF types *) + val mk_cfunT : typ * typ -> typ; + val ->> : typ * typ -> typ; + val mk_ssumT : typ * typ -> typ; + val mk_sprodT : typ * typ -> typ; + val mk_uT : typ -> typ; + val oneT : typ; + val trT : typ; + val mk_maybeT : typ -> typ; + val mk_ctupleT : typ list -> typ; + val mk_TFree : string -> typ; + val pcpoS : sort; + + (* Creating HOLCF terms *) + val %: : string -> term; + val %%: : string -> term; + val ` : term * term -> term; + val `% : term * string -> term; + val /\ : string -> term -> term; + val UU : term; + val TT : term; + val FF : term; + val ID : term; + val oo : term * term -> term; + val mk_up : term -> term; + val mk_sinl : term -> term; + val mk_sinr : term -> term; + val mk_stuple : term list -> term; + val mk_ctuple : term list -> term; + val mk_fix : term -> term; + val mk_iterate : term * term * term -> term; + val mk_fail : term; + val mk_return : term -> term; + val cproj : term -> 'a list -> int -> term; + val list_ccomb : term * term list -> term; + (* + val con_app : string -> ('a * 'b * string) list -> term; + *) + val con_app2 : string -> ('a -> term) -> 'a list -> term; + val proj : term -> 'a list -> int -> term; + val prj : ('a -> 'b -> 'a) -> ('a -> 'b -> 'a) -> 'a -> 'b list -> int -> 'a; + val mk_ctuple_pat : term list -> term; + val mk_branch : term -> term; + + (* Creating propositions *) + val mk_conj : term * term -> term; + val mk_disj : term * term -> term; + val mk_imp : term * term -> term; + val mk_lam : string * term -> term; + val mk_all : string * term -> term; + val mk_ex : string * term -> term; + val mk_constrain : typ * term -> term; + val mk_constrainall : string * typ * term -> term; + val === : term * term -> term; + val << : term * term -> term; + val ~<< : term * term -> term; + val strict : term -> term; + val defined : term -> term; + val mk_adm : term -> term; + val mk_compact : term -> term; + val lift : ('a -> term) -> 'a list * term -> term; + val lift_defined : ('a -> term) -> 'a list * term -> term; + + (* Creating meta-propositions *) + val mk_trp : term -> term; (* HOLogic.mk_Trueprop *) + val == : term * term -> term; + val ===> : term * term -> term; + val ==> : term * term -> term; + val mk_All : string * term -> term; + + (* Domain specifications *) + eqtype arg; + type cons = string * arg list; + type eq = (string * typ list) * cons list; + val mk_arg : (bool * Datatype.dtyp) * string option * string -> arg; + val is_lazy : arg -> bool; + val rec_of : arg -> int; + val dtyp_of : arg -> Datatype.dtyp; + val sel_of : arg -> string option; + val vname : arg -> string; + val upd_vname : (string -> string) -> arg -> arg; + val is_rec : arg -> bool; + val is_nonlazy_rec : arg -> bool; + val nonlazy : arg list -> string list; + val nonlazy_rec : arg list -> string list; + val %# : arg -> term; + val /\# : arg * term -> term; + val when_body : cons list -> (int * int -> term) -> term; + val when_funs : 'a list -> string list; + val bound_arg : ''a list -> ''a -> term; (* ''a = arg or string *) + val idx_name : 'a list -> string -> int -> string; + val app_rec_arg : (int -> term) -> arg -> term; + val con_app : string -> arg list -> term; + val dtyp_of_eq : eq -> Datatype.dtyp; + + + (* Name mangling *) + val strip_esc : string -> string; + val extern_name : string -> string; + val dis_name : string -> string; + val mat_name : string -> string; + val pat_name : string -> string; + val mk_var_names : string list -> string list; +end; + +structure Domain_Library :> DOMAIN_LIBRARY = +struct + +exception Impossible of string; +fun Imposs msg = raise Impossible ("Domain:"^msg); + +(* ----- name handling ----- *) + +val strip_esc = + let fun strip ("'" :: c :: cs) = c :: strip cs + | strip ["'"] = [] + | strip (c :: cs) = c :: strip cs + | strip [] = []; + in implode o strip o Symbol.explode end; + +fun extern_name con = + case Symbol.explode con of + ("o"::"p"::" "::rest) => implode rest + | _ => con; +fun dis_name con = "is_"^ (extern_name con); +fun dis_name_ con = "is_"^ (strip_esc con); +fun mat_name con = "match_"^ (extern_name con); +fun mat_name_ con = "match_"^ (strip_esc con); +fun pat_name con = (extern_name con) ^ "_pat"; +fun pat_name_ con = (strip_esc con) ^ "_pat"; + +(* make distinct names out of the type list, + forbidding "o","n..","x..","f..","P.." as names *) +(* a number string is added if necessary *) +fun mk_var_names ids : string list = + let + fun nonreserved s = if s mem ["n","x","f","P"] then s^"'" else s; + fun index_vnames(vn::vns,occupied) = + (case AList.lookup (op =) occupied vn of + NONE => if vn mem vns + then (vn^"1") :: index_vnames(vns,(vn,1) ::occupied) + else vn :: index_vnames(vns, occupied) + | SOME(i) => (vn^(string_of_int (i+1))) + :: index_vnames(vns,(vn,i+1)::occupied)) + | index_vnames([],occupied) = []; + in index_vnames(map nonreserved ids, [("O",0),("o",0)]) end; + +fun cpo_type sg t = Sign.of_sort sg (Sign.certify_typ sg t, @{sort cpo}); +fun pcpo_type sg t = Sign.of_sort sg (Sign.certify_typ sg t, @{sort pcpo}); +fun string_of_typ sg = Syntax.string_of_typ_global sg o Sign.certify_typ sg; + +(* ----- constructor list handling ----- *) + +type arg = + (bool * Datatype.dtyp) * (* (lazy, recursive element) *) + string option * (* selector name *) + string; (* argument name *) + +type cons = + string * (* operator name of constr *) + arg list; (* argument list *) + +type eq = + (string * (* name of abstracted type *) + typ list) * (* arguments of abstracted type *) + cons list; (* represented type, as a constructor list *) + +val mk_arg = I; + +fun rec_of ((_,dtyp),_,_) = + case dtyp of DatatypeAux.DtRec i => i | _ => ~1; +(* FIXME: what about indirect recursion? *) + +fun is_lazy arg = fst (first arg); +fun dtyp_of arg = snd (first arg); +val sel_of = second; +val vname = third; +val upd_vname = upd_third; +fun is_rec arg = rec_of arg >=0; +fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg); +fun nonlazy args = map vname (filter_out is_lazy args); +fun nonlazy_rec args = map vname (List.filter is_nonlazy_rec args); + + +(* ----- combinators for making dtyps ----- *) + +fun mk_uD T = DatatypeAux.DtType(@{type_name "u"}, [T]); +fun mk_sprodD (T, U) = DatatypeAux.DtType(@{type_name "**"}, [T, U]); +fun mk_ssumD (T, U) = DatatypeAux.DtType(@{type_name "++"}, [T, U]); +fun mk_liftD T = DatatypeAux.DtType(@{type_name "lift"}, [T]); +val unitD = DatatypeAux.DtType(@{type_name "unit"}, []); +val boolD = DatatypeAux.DtType(@{type_name "bool"}, []); +val oneD = mk_liftD unitD; +val trD = mk_liftD boolD; +fun big_sprodD ds = case ds of [] => oneD | _ => foldr1 mk_sprodD ds; +fun big_ssumD ds = case ds of [] => unitD | _ => foldr1 mk_ssumD ds; + +fun dtyp_of_arg ((lazy, D), _, _) = if lazy then mk_uD D else D; +fun dtyp_of_cons (_, args) = big_sprodD (map dtyp_of_arg args); +fun dtyp_of_eq (_, cons) = big_ssumD (map dtyp_of_cons cons); + + +(* ----- support for type and mixfix expressions ----- *) + +fun mk_uT T = Type(@{type_name "u"}, [T]); +fun mk_cfunT (T, U) = Type(@{type_name "->"}, [T, U]); +fun mk_sprodT (T, U) = Type(@{type_name "**"}, [T, U]); +fun mk_ssumT (T, U) = Type(@{type_name "++"}, [T, U]); +val oneT = @{typ one}; +val trT = @{typ tr}; + +val op ->> = mk_cfunT; + +fun mk_TFree s = TFree ("'" ^ s, @{sort pcpo}); + +(* ----- support for term expressions ----- *) + +fun %: s = Free(s,dummyT); +fun %# arg = %:(vname arg); +fun %%: s = Const(s,dummyT); + +local open HOLogic in +val mk_trp = mk_Trueprop; +fun mk_conj (S,T) = conj $ S $ T; +fun mk_disj (S,T) = disj $ S $ T; +fun mk_imp (S,T) = imp $ S $ T; +fun mk_lam (x,T) = Abs(x,dummyT,T); +fun mk_all (x,P) = HOLogic.mk_all (x,dummyT,P); +fun mk_ex (x,P) = mk_exists (x,dummyT,P); +val mk_constrain = uncurry TypeInfer.constrain; +fun mk_constrainall (x,typ,P) = %%:"All" $ (TypeInfer.constrain (typ --> boolT) (mk_lam(x,P))); +end + +fun mk_All (x,P) = %%:"all" $ mk_lam(x,P); (* meta universal quantification *) + +infixr 0 ===>; fun S ===> T = %%:"==>" $ S $ T; +infixr 0 ==>; fun S ==> T = mk_trp S ===> mk_trp T; +infix 0 ==; fun S == T = %%:"==" $ S $ T; +infix 1 ===; fun S === T = %%:"op =" $ S $ T; +infix 1 ~=; fun S ~= T = HOLogic.mk_not (S === T); +infix 1 <<; fun S << T = %%: @{const_name Porder.below} $ S $ T; +infix 1 ~<<; fun S ~<< T = HOLogic.mk_not (S << T); + +infix 9 ` ; fun f ` x = %%: @{const_name Rep_CFun} $ f $ x; +infix 9 `% ; fun f`% s = f` %: s; +infix 9 `%%; fun f`%%s = f` %%:s; + +fun mk_adm t = %%: @{const_name adm} $ t; +fun mk_compact t = %%: @{const_name compact} $ t; +val ID = %%: @{const_name ID}; +fun mk_strictify t = %%: @{const_name strictify}`t; +fun mk_cfst t = %%: @{const_name cfst}`t; +fun mk_csnd t = %%: @{const_name csnd}`t; +(*val csplitN = "Cprod.csplit";*) +(*val sfstN = "Sprod.sfst";*) +(*val ssndN = "Sprod.ssnd";*) +fun mk_ssplit t = %%: @{const_name ssplit}`t; +fun mk_sinl t = %%: @{const_name sinl}`t; +fun mk_sinr t = %%: @{const_name sinr}`t; +fun mk_sscase (x, y) = %%: @{const_name sscase}`x`y; +fun mk_up t = %%: @{const_name up}`t; +fun mk_fup (t,u) = %%: @{const_name fup} ` t ` u; +val ONE = @{term ONE}; +val TT = @{term TT}; +val FF = @{term FF}; +fun mk_iterate (n,f,z) = %%: @{const_name iterate} $ n ` f ` z; +fun mk_fix t = %%: @{const_name fix}`t; +fun mk_return t = %%: @{const_name Fixrec.return}`t; +val mk_fail = %%: @{const_name Fixrec.fail}; + +fun mk_branch t = %%: @{const_name Fixrec.branch} $ t; + +val pcpoS = @{sort pcpo}; + +val list_ccomb = Library.foldl (op `); (* continuous version of list_comb *) +fun con_app2 con f args = list_ccomb(%%:con,map f args); +fun con_app con = con_app2 con %#; +fun if_rec arg f y = if is_rec arg then f (rec_of arg) else y; +fun app_rec_arg p arg = if_rec arg (fn n => fn x => (p n)`x) I (%# arg); +fun prj _ _ x ( _::[]) _ = x + | prj f1 _ x (_::y::ys) 0 = f1 x y + | prj f1 f2 x (y:: ys) j = prj f1 f2 (f2 x y) ys (j-1); +fun proj x = prj (fn S => K(%%:"fst" $S)) (fn S => K(%%:"snd" $S)) x; +fun cproj x = prj (fn S => K(mk_cfst S)) (fn S => K(mk_csnd S)) x; +fun lift tfn = Library.foldr (fn (x,t)=> (mk_trp(tfn x) ===> t)); + +fun /\ v T = %%: @{const_name Abs_CFun} $ mk_lam(v,T); +fun /\# (arg,T) = /\ (vname arg) T; +infixr 9 oo; fun S oo T = %%: @{const_name cfcomp}`S`T; +val UU = %%: @{const_name UU}; +fun strict f = f`UU === UU; +fun defined t = t ~= UU; +fun cpair (t,u) = %%: @{const_name cpair}`t`u; +fun spair (t,u) = %%: @{const_name spair}`t`u; +fun mk_ctuple [] = HOLogic.unit (* used in match_defs *) + | mk_ctuple ts = foldr1 cpair ts; +fun mk_stuple [] = ONE + | mk_stuple ts = foldr1 spair ts; +fun mk_ctupleT [] = HOLogic.unitT (* used in match_defs *) + | mk_ctupleT Ts = foldr1 HOLogic.mk_prodT Ts; +fun mk_maybeT T = Type ("Fixrec.maybe",[T]); +fun cpair_pat (p1,p2) = %%: @{const_name cpair_pat} $ p1 $ p2; +val mk_ctuple_pat = foldr1 cpair_pat; +fun lift_defined f = lift (fn x => defined (f x)); +fun bound_arg vns v = Bound (length vns - find_index (fn v' => v' = v) vns - 1); + +fun cont_eta_contract (Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body)) = + (case cont_eta_contract body of + body' as (Const("Cfun.Rep_CFun",Ta) $ f $ Bound 0) => + if not (0 mem loose_bnos f) then incr_boundvars ~1 f + else Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body') + | body' => Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body')) + | cont_eta_contract(f$t) = cont_eta_contract f $ cont_eta_contract t + | cont_eta_contract t = t; + +fun idx_name dnames s n = s^(if length dnames = 1 then "" else string_of_int n); +fun when_funs cons = if length cons = 1 then ["f"] + else mapn (fn n => K("f"^(string_of_int n))) 1 cons; +fun when_body cons funarg = + let + fun one_fun n (_,[] ) = /\ "dummy" (funarg(1,n)) + | one_fun n (_,args) = let + val l2 = length args; + fun idxs m arg = (if is_lazy arg then (fn t => mk_fup (ID, t)) + else I) (Bound(l2-m)); + in cont_eta_contract + (foldr'' + (fn (a,t) => mk_ssplit (/\# (a,t))) + (args, + fn a=> /\#(a,(list_ccomb(funarg(l2,n),mapn idxs 1 args)))) + ) end; + in (if length cons = 1 andalso length(snd(hd cons)) <= 1 + then mk_strictify else I) + (foldr1 mk_sscase (mapn one_fun 1 cons)) end; + +end; (* struct *) diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/Tools/Domain/domain_syntax.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/Tools/Domain/domain_syntax.ML Thu Jul 23 21:13:21 2009 +0200 @@ -0,0 +1,181 @@ +(* Title: HOLCF/Tools/Domain/domain_syntax.ML + Author: David von Oheimb + +Syntax generator for domain command. +*) + +signature DOMAIN_SYNTAX = +sig + val calc_syntax: + typ -> + (string * typ list) * + (binding * (bool * binding option * typ) list * mixfix) list -> + (binding * typ * mixfix) list * ast Syntax.trrule list + + val add_syntax: + string -> + ((string * typ list) * + (binding * (bool * binding option * typ) list * mixfix) list) list -> + theory -> theory +end; + + +structure Domain_Syntax :> DOMAIN_SYNTAX = +struct + +open Domain_Library; +infixr 5 -->; infixr 6 ->>; + +fun calc_syntax + (dtypeprod : typ) + ((dname : string, typevars : typ list), + (cons': (binding * (bool * binding option * typ) list * mixfix) list)) + : (binding * typ * mixfix) list * ast Syntax.trrule list = + let + (* ----- constants concerning the isomorphism ------------------------------- *) + + local + fun opt_lazy (lazy,_,t) = if lazy then mk_uT t else t + fun prod (_,args,_) = case args of [] => oneT + | _ => foldr1 mk_sprodT (map opt_lazy args); + fun freetvar s = let val tvar = mk_TFree s in + if tvar mem typevars then freetvar ("t"^s) else tvar end; + fun when_type (_,args,_) = List.foldr (op ->>) (freetvar "t") (map third args); + in + val dtype = Type(dname,typevars); + val dtype2 = foldr1 mk_ssumT (map prod cons'); + val dnam = Long_Name.base_name dname; + fun dbind s = Binding.name (dnam ^ s); + val const_rep = (dbind "_rep" , dtype ->> dtype2, NoSyn); + val const_abs = (dbind "_abs" , dtype2 ->> dtype , NoSyn); + val const_when = (dbind "_when", List.foldr (op ->>) (dtype ->> freetvar "t") (map when_type cons'), NoSyn); + val const_copy = (dbind "_copy", dtypeprod ->> dtype ->> dtype , NoSyn); + end; + + (* ----- constants concerning constructors, discriminators, and selectors --- *) + + local + val escape = let + fun esc (c::cs) = if c mem ["'","_","(",")","/"] then "'"::c::esc cs + else c::esc cs + | esc [] = [] + in implode o esc o Symbol.explode end; + fun dis_name_ con = Binding.name ("is_" ^ strip_esc (Binding.name_of con)); + fun mat_name_ con = Binding.name ("match_" ^ strip_esc (Binding.name_of con)); + fun pat_name_ con = Binding.name (strip_esc (Binding.name_of con) ^ "_pat"); + fun con (name,args,mx) = (name, List.foldr (op ->>) dtype (map third args), mx); + fun dis (con,args,mx) = (dis_name_ con, dtype->>trT, + Mixfix(escape ("is_" ^ Binding.name_of con), [], Syntax.max_pri)); + (* strictly speaking, these constants have one argument, + but the mixfix (without arguments) is introduced only + to generate parse rules for non-alphanumeric names*) + fun freetvar s n = let val tvar = mk_TFree (s ^ string_of_int n) in + if tvar mem typevars then freetvar ("t"^s) n else tvar end; + fun mk_matT (a,bs,c) = a ->> foldr (op ->>) (mk_maybeT c) bs ->> mk_maybeT c; + fun mat (con,args,mx) = (mat_name_ con, + mk_matT(dtype, map third args, freetvar "t" 1), + Mixfix(escape ("match_" ^ Binding.name_of con), [], Syntax.max_pri)); + fun sel1 (_,sel,typ) = Option.map (fn s => (s,dtype ->> typ,NoSyn)) sel; + fun sel (con,args,mx) = List.mapPartial sel1 args; + fun mk_patT (a,b) = a ->> mk_maybeT b; + fun pat_arg_typ n arg = mk_patT (third arg, freetvar "t" n); + fun pat (con,args,mx) = (pat_name_ con, + (mapn pat_arg_typ 1 args) + ---> + mk_patT (dtype, mk_ctupleT (map (freetvar "t") (1 upto length args))), + Mixfix(escape (Binding.name_of con ^ "_pat"), [], Syntax.max_pri)); + + in + val consts_con = map con cons'; + val consts_dis = map dis cons'; + val consts_mat = map mat cons'; + val consts_pat = map pat cons'; + val consts_sel = List.concat(map sel cons'); + end; + + (* ----- constants concerning induction ------------------------------------- *) + + val const_take = (dbind "_take" , HOLogic.natT-->dtype->>dtype, NoSyn); + val const_finite = (dbind "_finite", dtype-->HOLogic.boolT , NoSyn); + + (* ----- case translation --------------------------------------------------- *) + + local open Syntax in + local + fun c_ast con mx = Constant (Syntax.const_name mx (Binding.name_of con)); + fun expvar n = Variable ("e"^(string_of_int n)); + fun argvar n m _ = Variable ("a"^(string_of_int n)^"_"^ + (string_of_int m)); + fun argvars n args = mapn (argvar n) 1 args; + fun app s (l,r) = mk_appl (Constant s) [l,r]; + val cabs = app "_cabs"; + val capp = app "Rep_CFun"; + fun con1 n (con,args,mx) = Library.foldl capp (c_ast con mx, argvars n args); + fun case1 n (con,args,mx) = app "_case1" (con1 n (con,args,mx), expvar n); + fun arg1 n (con,args,_) = List.foldr cabs (expvar n) (argvars n args); + fun when1 n m = if n = m then arg1 n else K (Constant "UU"); + + fun app_var x = mk_appl (Constant "_variable") [x, Variable "rhs"]; + fun app_pat x = mk_appl (Constant "_pat") [x]; + fun args_list [] = Constant "_noargs" + | args_list xs = foldr1 (app "_args") xs; + in + val case_trans = + ParsePrintRule + (app "_case_syntax" (Variable "x", foldr1 (app "_case2") (mapn case1 1 cons')), + capp (Library.foldl capp (Constant (dnam^"_when"), mapn arg1 1 cons'), Variable "x")); + + fun one_abscon_trans n (con,mx,args) = + ParsePrintRule + (cabs (con1 n (con,mx,args), expvar n), + Library.foldl capp (Constant (dnam^"_when"), mapn (when1 n) 1 cons')); + val abscon_trans = mapn one_abscon_trans 1 cons'; + + fun one_case_trans (con,args,mx) = + let + val cname = c_ast con mx; + val pname = Constant (strip_esc (Binding.name_of con) ^ "_pat"); + val ns = 1 upto length args; + val xs = map (fn n => Variable ("x"^(string_of_int n))) ns; + val ps = map (fn n => Variable ("p"^(string_of_int n))) ns; + val vs = map (fn n => Variable ("v"^(string_of_int n))) ns; + in + [ParseRule (app_pat (Library.foldl capp (cname, xs)), + mk_appl pname (map app_pat xs)), + ParseRule (app_var (Library.foldl capp (cname, xs)), + app_var (args_list xs)), + PrintRule (Library.foldl capp (cname, ListPair.map (app "_match") (ps,vs)), + app "_match" (mk_appl pname ps, args_list vs))] + end; + val Case_trans = List.concat (map one_case_trans cons'); + end; + end; + + in ([const_rep, const_abs, const_when, const_copy] @ + consts_con @ consts_dis @ consts_mat @ consts_pat @ consts_sel @ + [const_take, const_finite], + (case_trans::(abscon_trans @ Case_trans))) + end; (* let *) + +(* ----- putting all the syntax stuff together ------------------------------ *) + +fun add_syntax + (comp_dnam : string) + (eqs' : ((string * typ list) * + (binding * (bool * binding option * typ) list * mixfix) list) list) + (thy'' : theory) = + let + val dtypes = map (Type o fst) eqs'; + val boolT = HOLogic.boolT; + val funprod = foldr1 HOLogic.mk_prodT (map (fn tp => tp ->> tp ) dtypes); + val relprod = foldr1 HOLogic.mk_prodT (map (fn tp => tp --> tp --> boolT) dtypes); + val const_copy = (Binding.name (comp_dnam^"_copy"), funprod ->> funprod, NoSyn); + val const_bisim = (Binding.name (comp_dnam^"_bisim"), relprod --> boolT, NoSyn); + val ctt : ((binding * typ * mixfix) list * ast Syntax.trrule list) list = map (calc_syntax funprod) eqs'; + in thy'' |> ContConsts.add_consts_i (List.concat (map fst ctt) @ + (if length eqs'>1 then [const_copy] else[])@ + [const_bisim]) + |> Sign.add_trrules_i (List.concat(map snd ctt)) + end; (* let *) + +end; (* struct *) diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/Tools/Domain/domain_theorems.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/Tools/Domain/domain_theorems.ML Thu Jul 23 21:13:21 2009 +0200 @@ -0,0 +1,1053 @@ +(* Title: HOLCF/Tools/Domain/domain_theorems.ML + Author: David von Oheimb + New proofs/tactics by Brian Huffman + +Proof generator for domain command. +*) + +val HOLCF_ss = @{simpset}; + +signature DOMAIN_THEOREMS = +sig + val theorems: Domain_Library.eq * Domain_Library.eq list -> theory -> thm list * theory; + val comp_theorems: bstring * Domain_Library.eq list -> theory -> thm list * theory; + val quiet_mode: bool ref; + val trace_domain: bool ref; +end; + +structure Domain_Theorems :> DOMAIN_THEOREMS = +struct + +val quiet_mode = ref false; +val trace_domain = ref false; + +fun message s = if !quiet_mode then () else writeln s; +fun trace s = if !trace_domain then tracing s else (); + +local + +val adm_impl_admw = @{thm adm_impl_admw}; +val adm_all = @{thm adm_all}; +val adm_conj = @{thm adm_conj}; +val adm_subst = @{thm adm_subst}; +val antisym_less_inverse = @{thm below_antisym_inverse}; +val beta_cfun = @{thm beta_cfun}; +val cfun_arg_cong = @{thm cfun_arg_cong}; +val ch2ch_Rep_CFunL = @{thm ch2ch_Rep_CFunL}; +val ch2ch_Rep_CFunR = @{thm ch2ch_Rep_CFunR}; +val chain_iterate = @{thm chain_iterate}; +val compact_ONE = @{thm compact_ONE}; +val compact_sinl = @{thm compact_sinl}; +val compact_sinr = @{thm compact_sinr}; +val compact_spair = @{thm compact_spair}; +val compact_up = @{thm compact_up}; +val contlub_cfun_arg = @{thm contlub_cfun_arg}; +val contlub_cfun_fun = @{thm contlub_cfun_fun}; +val fix_def2 = @{thm fix_def2}; +val injection_eq = @{thm injection_eq}; +val injection_less = @{thm injection_below}; +val lub_equal = @{thm lub_equal}; +val monofun_cfun_arg = @{thm monofun_cfun_arg}; +val retraction_strict = @{thm retraction_strict}; +val spair_eq = @{thm spair_eq}; +val spair_less = @{thm spair_below}; +val sscase1 = @{thm sscase1}; +val ssplit1 = @{thm ssplit1}; +val strictify1 = @{thm strictify1}; +val wfix_ind = @{thm wfix_ind}; + +val iso_intro = @{thm iso.intro}; +val iso_abs_iso = @{thm iso.abs_iso}; +val iso_rep_iso = @{thm iso.rep_iso}; +val iso_abs_strict = @{thm iso.abs_strict}; +val iso_rep_strict = @{thm iso.rep_strict}; +val iso_abs_defin' = @{thm iso.abs_defin'}; +val iso_rep_defin' = @{thm iso.rep_defin'}; +val iso_abs_defined = @{thm iso.abs_defined}; +val iso_rep_defined = @{thm iso.rep_defined}; +val iso_compact_abs = @{thm iso.compact_abs}; +val iso_compact_rep = @{thm iso.compact_rep}; +val iso_iso_swap = @{thm iso.iso_swap}; + +val exh_start = @{thm exh_start}; +val ex_defined_iffs = @{thms ex_defined_iffs}; +val exh_casedist0 = @{thm exh_casedist0}; +val exh_casedists = @{thms exh_casedists}; + +open Domain_Library; +infixr 0 ===>; +infixr 0 ==>; +infix 0 == ; +infix 1 ===; +infix 1 ~= ; +infix 1 <<; +infix 1 ~<<; +infix 9 ` ; +infix 9 `% ; +infix 9 `%%; +infixr 9 oo; + +(* ----- general proof facilities ------------------------------------------- *) + +fun legacy_infer_term thy t = + let val ctxt = ProofContext.set_mode ProofContext.mode_schematic (ProofContext.init thy) + in singleton (Syntax.check_terms ctxt) (Sign.intern_term thy t) end; + +fun pg'' thy defs t tacs = + let + val t' = legacy_infer_term thy t; + val asms = Logic.strip_imp_prems t'; + val prop = Logic.strip_imp_concl t'; + fun tac {prems, context} = + rewrite_goals_tac defs THEN + EVERY (tacs {prems = map (rewrite_rule defs) prems, context = context}) + in Goal.prove_global thy [] asms prop tac end; + +fun pg' thy defs t tacsf = + let + fun tacs {prems, context} = + if null prems then tacsf context + else cut_facts_tac prems 1 :: tacsf context; + in pg'' thy defs t tacs end; + +fun case_UU_tac ctxt rews i v = + InductTacs.case_tac ctxt (v^"=UU") i THEN + asm_simp_tac (HOLCF_ss addsimps rews) i; + +val chain_tac = + REPEAT_DETERM o resolve_tac + [chain_iterate, ch2ch_Rep_CFunR, ch2ch_Rep_CFunL]; + +(* ----- general proofs ----------------------------------------------------- *) + +val all2E = @{lemma "!x y . P x y ==> (P x y ==> R) ==> R" by simp} + +val dist_eqI = @{lemma "!!x::'a::po. ~ x << y ==> x ~= y" by (blast dest!: below_antisym_inverse)} + +in + +fun theorems (((dname, _), cons) : eq, eqs : eq list) thy = +let + +val _ = message ("Proving isomorphism properties of domain "^dname^" ..."); +val pg = pg' thy; + +(* ----- getting the axioms and definitions --------------------------------- *) + +local + fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s); +in + val ax_abs_iso = ga "abs_iso" dname; + val ax_rep_iso = ga "rep_iso" dname; + val ax_when_def = ga "when_def" dname; + fun get_def mk_name (con,_) = ga (mk_name con^"_def") dname; + val axs_con_def = map (get_def extern_name) cons; + val axs_dis_def = map (get_def dis_name) cons; + val axs_mat_def = map (get_def mat_name) cons; + val axs_pat_def = map (get_def pat_name) cons; + val axs_sel_def = + let + fun def_of_sel sel = ga (sel^"_def") dname; + fun def_of_arg arg = Option.map def_of_sel (sel_of arg); + fun defs_of_con (_, args) = List.mapPartial def_of_arg args; + in + maps defs_of_con cons + end; + val ax_copy_def = ga "copy_def" dname; +end; (* local *) + +(* ----- theorems concerning the isomorphism -------------------------------- *) + +val dc_abs = %%:(dname^"_abs"); +val dc_rep = %%:(dname^"_rep"); +val dc_copy = %%:(dname^"_copy"); +val x_name = "x"; + +val iso_locale = iso_intro OF [ax_abs_iso, ax_rep_iso]; +val abs_strict = ax_rep_iso RS (allI RS retraction_strict); +val rep_strict = ax_abs_iso RS (allI RS retraction_strict); +val abs_defin' = iso_locale RS iso_abs_defin'; +val rep_defin' = iso_locale RS iso_rep_defin'; +val iso_rews = map standard [ax_abs_iso,ax_rep_iso,abs_strict,rep_strict]; + +(* ----- generating beta reduction rules from definitions-------------------- *) + +val _ = trace " Proving beta reduction rules..."; + +local + fun arglist (Const _ $ Abs (s, _, t)) = + let + val (vars,body) = arglist t; + in (s :: vars, body) end + | arglist t = ([], t); + fun bind_fun vars t = Library.foldr mk_All (vars, t); + fun bound_vars 0 = [] + | bound_vars i = Bound (i-1) :: bound_vars (i - 1); +in + fun appl_of_def def = + let + val (_ $ con $ lam) = concl_of def; + val (vars, rhs) = arglist lam; + val lhs = list_ccomb (con, bound_vars (length vars)); + val appl = bind_fun vars (lhs == rhs); + val cs = ContProc.cont_thms lam; + val betas = map (fn c => mk_meta_eq (c RS beta_cfun)) cs; + in pg (def::betas) appl (K [rtac reflexive_thm 1]) end; +end; + +val _ = trace "Proving when_appl..."; +val when_appl = appl_of_def ax_when_def; +val _ = trace "Proving con_appls..."; +val con_appls = map appl_of_def axs_con_def; + +local + fun arg2typ n arg = + let val t = TVar (("'a", n), pcpoS) + in (n + 1, if is_lazy arg then mk_uT t else t) end; + + fun args2typ n [] = (n, oneT) + | args2typ n [arg] = arg2typ n arg + | args2typ n (arg::args) = + let + val (n1, t1) = arg2typ n arg; + val (n2, t2) = args2typ n1 args + in (n2, mk_sprodT (t1, t2)) end; + + fun cons2typ n [] = (n,oneT) + | cons2typ n [con] = args2typ n (snd con) + | cons2typ n (con::cons) = + let + val (n1, t1) = args2typ n (snd con); + val (n2, t2) = cons2typ n1 cons + in (n2, mk_ssumT (t1, t2)) end; +in + fun cons2ctyp cons = ctyp_of thy (snd (cons2typ 1 cons)); +end; + +local + val iso_swap = iso_locale RS iso_iso_swap; + fun one_con (con, args) = + let + val vns = map vname args; + val eqn = %:x_name === con_app2 con %: vns; + val conj = foldr1 mk_conj (eqn :: map (defined o %:) (nonlazy args)); + in Library.foldr mk_ex (vns, conj) end; + + val conj_assoc = @{thm conj_assoc}; + val exh = foldr1 mk_disj ((%:x_name === UU) :: map one_con cons); + val thm1 = instantiate' [SOME (cons2ctyp cons)] [] exh_start; + val thm2 = rewrite_rule (map mk_meta_eq ex_defined_iffs) thm1; + val thm3 = rewrite_rule [mk_meta_eq @{thm conj_assoc}] thm2; + + (* first 3 rules replace "x = UU \/ P" with "rep$x = UU \/ P" *) + val tacs = [ + rtac disjE 1, + etac (rep_defin' RS disjI1) 2, + etac disjI2 2, + rewrite_goals_tac [mk_meta_eq iso_swap], + rtac thm3 1]; +in + val _ = trace " Proving exhaust..."; + val exhaust = pg con_appls (mk_trp exh) (K tacs); + val _ = trace " Proving casedist..."; + val casedist = + standard (rewrite_rule exh_casedists (exhaust RS exh_casedist0)); +end; + +local + fun bind_fun t = Library.foldr mk_All (when_funs cons, t); + fun bound_fun i _ = Bound (length cons - i); + val when_app = list_ccomb (%%:(dname^"_when"), mapn bound_fun 1 cons); +in + val _ = trace " Proving when_strict..."; + val when_strict = + let + val axs = [when_appl, mk_meta_eq rep_strict]; + val goal = bind_fun (mk_trp (strict when_app)); + val tacs = [resolve_tac [sscase1, ssplit1, strictify1] 1]; + in pg axs goal (K tacs) end; + + val _ = trace " Proving when_apps..."; + val when_apps = + let + fun one_when n (con,args) = + let + val axs = when_appl :: con_appls; + val goal = bind_fun (lift_defined %: (nonlazy args, + mk_trp (when_app`(con_app con args) === + list_ccomb (bound_fun n 0, map %# args)))); + val tacs = [asm_simp_tac (HOLCF_ss addsimps [ax_abs_iso]) 1]; + in pg axs goal (K tacs) end; + in mapn one_when 1 cons end; +end; +val when_rews = when_strict :: when_apps; + +(* ----- theorems concerning the constructors, discriminators and selectors - *) + +local + fun dis_strict (con, _) = + let + val goal = mk_trp (strict (%%:(dis_name con))); + in pg axs_dis_def goal (K [rtac when_strict 1]) end; + + fun dis_app c (con, args) = + let + val lhs = %%:(dis_name c) ` con_app con args; + val rhs = if con = c then TT else FF; + val goal = lift_defined %: (nonlazy args, mk_trp (lhs === rhs)); + val tacs = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; + in pg axs_dis_def goal (K tacs) end; + + val _ = trace " Proving dis_apps..."; + val dis_apps = maps (fn (c,_) => map (dis_app c) cons) cons; + + fun dis_defin (con, args) = + let + val goal = defined (%:x_name) ==> defined (%%:(dis_name con) `% x_name); + val tacs = + [rtac casedist 1, + contr_tac 1, + DETERM_UNTIL_SOLVED (CHANGED + (asm_simp_tac (HOLCF_ss addsimps dis_apps) 1))]; + in pg [] goal (K tacs) end; + + val _ = trace " Proving dis_stricts..."; + val dis_stricts = map dis_strict cons; + val _ = trace " Proving dis_defins..."; + val dis_defins = map dis_defin cons; +in + val dis_rews = dis_stricts @ dis_defins @ dis_apps; +end; + +local + fun mat_strict (con, _) = + let + val goal = mk_trp (%%:(mat_name con) ` UU ` %:"rhs" === UU); + val tacs = [asm_simp_tac (HOLCF_ss addsimps [when_strict]) 1]; + in pg axs_mat_def goal (K tacs) end; + + val _ = trace " Proving mat_stricts..."; + val mat_stricts = map mat_strict cons; + + fun one_mat c (con, args) = + let + val lhs = %%:(mat_name c) ` con_app con args ` %:"rhs"; + val rhs = + if con = c + then list_ccomb (%:"rhs", map %# args) + else mk_fail; + val goal = lift_defined %: (nonlazy args, mk_trp (lhs === rhs)); + val tacs = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; + in pg axs_mat_def goal (K tacs) end; + + val _ = trace " Proving mat_apps..."; + val mat_apps = + maps (fn (c,_) => map (one_mat c) cons) cons; +in + val mat_rews = mat_stricts @ mat_apps; +end; + +local + fun ps args = mapn (fn n => fn _ => %:("pat" ^ string_of_int n)) 1 args; + + fun pat_lhs (con,args) = mk_branch (list_comb (%%:(pat_name con), ps args)); + + fun pat_rhs (con,[]) = mk_return ((%:"rhs") ` HOLogic.unit) + | pat_rhs (con,args) = + (mk_branch (mk_ctuple_pat (ps args))) + `(%:"rhs")`(mk_ctuple (map %# args)); + + fun pat_strict c = + let + val axs = @{thm branch_def} :: axs_pat_def; + val goal = mk_trp (strict (pat_lhs c ` (%:"rhs"))); + val tacs = [simp_tac (HOLCF_ss addsimps [when_strict]) 1]; + in pg axs goal (K tacs) end; + + fun pat_app c (con, args) = + let + val axs = @{thm branch_def} :: axs_pat_def; + val lhs = (pat_lhs c)`(%:"rhs")`(con_app con args); + val rhs = if con = fst c then pat_rhs c else mk_fail; + val goal = lift_defined %: (nonlazy args, mk_trp (lhs === rhs)); + val tacs = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; + in pg axs goal (K tacs) end; + + val _ = trace " Proving pat_stricts..."; + val pat_stricts = map pat_strict cons; + val _ = trace " Proving pat_apps..."; + val pat_apps = maps (fn c => map (pat_app c) cons) cons; +in + val pat_rews = pat_stricts @ pat_apps; +end; + +local + fun con_strict (con, args) = + let + val rules = abs_strict :: @{thms con_strict_rules}; + fun one_strict vn = + let + fun f arg = if vname arg = vn then UU else %# arg; + val goal = mk_trp (con_app2 con f args === UU); + val tacs = [simp_tac (HOL_basic_ss addsimps rules) 1]; + in pg con_appls goal (K tacs) end; + in map one_strict (nonlazy args) end; + + fun con_defin (con, args) = + let + fun iff_disj (t, []) = HOLogic.mk_not t + | iff_disj (t, ts) = t === foldr1 HOLogic.mk_disj ts; + val lhs = con_app con args === UU; + val rhss = map (fn x => %:x === UU) (nonlazy args); + val goal = mk_trp (iff_disj (lhs, rhss)); + val rule1 = iso_locale RS @{thm iso.abs_defined_iff}; + val rules = rule1 :: @{thms con_defined_iff_rules}; + val tacs = [simp_tac (HOL_ss addsimps rules) 1]; + in pg con_appls goal (K tacs) end; +in + val _ = trace " Proving con_stricts..."; + val con_stricts = maps con_strict cons; + val _ = trace " Proving con_defins..."; + val con_defins = map con_defin cons; + val con_rews = con_stricts @ con_defins; +end; + +local + val rules = + [compact_sinl, compact_sinr, compact_spair, compact_up, compact_ONE]; + fun con_compact (con, args) = + let + val concl = mk_trp (mk_compact (con_app con args)); + val goal = lift (fn x => mk_compact (%#x)) (args, concl); + val tacs = [ + rtac (iso_locale RS iso_compact_abs) 1, + REPEAT (resolve_tac rules 1 ORELSE atac 1)]; + in pg con_appls goal (K tacs) end; +in + val _ = trace " Proving con_compacts..."; + val con_compacts = map con_compact cons; +end; + +local + fun one_sel sel = + pg axs_sel_def (mk_trp (strict (%%:sel))) + (K [simp_tac (HOLCF_ss addsimps when_rews) 1]); + + fun sel_strict (_, args) = + List.mapPartial (Option.map one_sel o sel_of) args; +in + val _ = trace " Proving sel_stricts..."; + val sel_stricts = maps sel_strict cons; +end; + +local + fun sel_app_same c n sel (con, args) = + let + val nlas = nonlazy args; + val vns = map vname args; + val vnn = List.nth (vns, n); + val nlas' = List.filter (fn v => v <> vnn) nlas; + val lhs = (%%:sel)`(con_app con args); + val goal = lift_defined %: (nlas', mk_trp (lhs === %:vnn)); + fun tacs1 ctxt = + if vnn mem nlas + then [case_UU_tac ctxt (when_rews @ con_stricts) 1 vnn] + else []; + val tacs2 = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; + in pg axs_sel_def goal (fn ctxt => (tacs1 ctxt @ tacs2)) end; + + fun sel_app_diff c n sel (con, args) = + let + val nlas = nonlazy args; + val goal = mk_trp (%%:sel ` con_app con args === UU); + fun tacs1 ctxt = map (case_UU_tac ctxt (when_rews @ con_stricts) 1) nlas; + val tacs2 = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; + in pg axs_sel_def goal (fn ctxt => (tacs1 ctxt @ tacs2)) end; + + fun sel_app c n sel (con, args) = + if con = c + then sel_app_same c n sel (con, args) + else sel_app_diff c n sel (con, args); + + fun one_sel c n sel = map (sel_app c n sel) cons; + fun one_sel' c n arg = Option.map (one_sel c n) (sel_of arg); + fun one_con (c, args) = + flat (map_filter I (mapn (one_sel' c) 0 args)); +in + val _ = trace " Proving sel_apps..."; + val sel_apps = maps one_con cons; +end; + +local + fun sel_defin sel = + let + val goal = defined (%:x_name) ==> defined (%%:sel`%x_name); + val tacs = [ + rtac casedist 1, + contr_tac 1, + DETERM_UNTIL_SOLVED (CHANGED + (asm_simp_tac (HOLCF_ss addsimps sel_apps) 1))]; + in pg [] goal (K tacs) end; +in + val _ = trace " Proving sel_defins..."; + val sel_defins = + if length cons = 1 + then List.mapPartial (fn arg => Option.map sel_defin (sel_of arg)) + (filter_out is_lazy (snd (hd cons))) + else []; +end; + +val sel_rews = sel_stricts @ sel_defins @ sel_apps; + +val _ = trace " Proving dist_les..."; +val distincts_le = + let + fun dist (con1, args1) (con2, args2) = + let + val goal = lift_defined %: (nonlazy args1, + mk_trp (con_app con1 args1 ~<< con_app con2 args2)); + fun tacs ctxt = [ + rtac @{thm rev_contrapos} 1, + eres_inst_tac ctxt [(("f", 0), dis_name con1)] monofun_cfun_arg 1] + @ map (case_UU_tac ctxt (con_stricts @ dis_rews) 1) (nonlazy args2) + @ [asm_simp_tac (HOLCF_ss addsimps dis_rews) 1]; + in pg [] goal tacs end; + + fun distinct (con1, args1) (con2, args2) = + let + val arg1 = (con1, args1); + val arg2 = + (con2, ListPair.map (fn (arg,vn) => upd_vname (K vn) arg) + (args2, Name.variant_list (map vname args1) (map vname args2))); + in [dist arg1 arg2, dist arg2 arg1] end; + fun distincts [] = [] + | distincts (c::cs) = (map (distinct c) cs) :: distincts cs; + in distincts cons end; +val dist_les = flat (flat distincts_le); + +val _ = trace " Proving dist_eqs..."; +val dist_eqs = + let + fun distinct (_,args1) ((_,args2), leqs) = + let + val (le1,le2) = (hd leqs, hd(tl leqs)); + val (eq1,eq2) = (le1 RS dist_eqI, le2 RS dist_eqI) + in + if nonlazy args1 = [] then [eq1, eq1 RS not_sym] else + if nonlazy args2 = [] then [eq2, eq2 RS not_sym] else + [eq1, eq2] + end; + fun distincts [] = [] + | distincts ((c,leqs)::cs) = + flat + (ListPair.map (distinct c) ((map #1 cs),leqs)) @ + distincts cs; + in map standard (distincts (cons ~~ distincts_le)) end; + +local + fun pgterm rel con args = + let + fun append s = upd_vname (fn v => v^s); + val (largs, rargs) = (args, map (append "'") args); + val concl = + foldr1 mk_conj (ListPair.map rel (map %# largs, map %# rargs)); + val prem = rel (con_app con largs, con_app con rargs); + val sargs = case largs of [_] => [] | _ => nonlazy args; + val prop = lift_defined %: (sargs, mk_trp (prem === concl)); + in pg con_appls prop end; + val cons' = List.filter (fn (_,args) => args<>[]) cons; +in + val _ = trace " Proving inverts..."; + val inverts = + let + val abs_less = ax_abs_iso RS (allI RS injection_less); + val tacs = + [asm_full_simp_tac (HOLCF_ss addsimps [abs_less, spair_less]) 1]; + in map (fn (con, args) => pgterm (op <<) con args (K tacs)) cons' end; + + val _ = trace " Proving injects..."; + val injects = + let + val abs_eq = ax_abs_iso RS (allI RS injection_eq); + val tacs = [asm_full_simp_tac (HOLCF_ss addsimps [abs_eq, spair_eq]) 1]; + in map (fn (con, args) => pgterm (op ===) con args (K tacs)) cons' end; +end; + +(* ----- theorems concerning one induction step ----------------------------- *) + +val copy_strict = + let + val _ = trace " Proving copy_strict..."; + val goal = mk_trp (strict (dc_copy `% "f")); + val rules = [abs_strict, rep_strict] @ @{thms domain_fun_stricts}; + val tacs = [asm_simp_tac (HOLCF_ss addsimps rules) 1]; + in pg [ax_copy_def] goal (K tacs) end; + +local + fun copy_app (con, args) = + let + val lhs = dc_copy`%"f"`(con_app con args); + fun one_rhs arg = + if DatatypeAux.is_rec_type (dtyp_of arg) + then Domain_Axioms.copy_of_dtyp (cproj (%:"f") eqs) (dtyp_of arg) ` (%# arg) + else (%# arg); + val rhs = con_app2 con one_rhs args; + val goal = lift_defined %: (nonlazy_rec args, mk_trp (lhs === rhs)); + val args' = List.filter (fn a => not (is_rec a orelse is_lazy a)) args; + val stricts = abs_strict :: rep_strict :: @{thms domain_fun_stricts}; + fun tacs1 ctxt = map (case_UU_tac ctxt stricts 1 o vname) args'; + val rules = [ax_abs_iso] @ @{thms domain_fun_simps}; + val tacs2 = [asm_simp_tac (HOLCF_ss addsimps rules) 1]; + in pg (ax_copy_def::con_appls) goal (fn ctxt => (tacs1 ctxt @ tacs2)) end; +in + val _ = trace " Proving copy_apps..."; + val copy_apps = map copy_app cons; +end; + +local + fun one_strict (con, args) = + let + val goal = mk_trp (dc_copy`UU`(con_app con args) === UU); + val rews = copy_strict :: copy_apps @ con_rews; + fun tacs ctxt = map (case_UU_tac ctxt rews 1) (nonlazy args) @ + [asm_simp_tac (HOLCF_ss addsimps rews) 1]; + in pg [] goal tacs end; + + fun has_nonlazy_rec (_, args) = exists is_nonlazy_rec args; +in + val _ = trace " Proving copy_stricts..."; + val copy_stricts = map one_strict (List.filter has_nonlazy_rec cons); +end; + +val copy_rews = copy_strict :: copy_apps @ copy_stricts; + +in + thy + |> Sign.add_path (Long_Name.base_name dname) + |> snd o PureThy.add_thmss [ + ((Binding.name "iso_rews" , iso_rews ), [Simplifier.simp_add]), + ((Binding.name "exhaust" , [exhaust] ), []), + ((Binding.name "casedist" , [casedist] ), [Induct.cases_type dname]), + ((Binding.name "when_rews" , when_rews ), [Simplifier.simp_add]), + ((Binding.name "compacts" , con_compacts), [Simplifier.simp_add]), + ((Binding.name "con_rews" , con_rews ), [Simplifier.simp_add]), + ((Binding.name "sel_rews" , sel_rews ), [Simplifier.simp_add]), + ((Binding.name "dis_rews" , dis_rews ), [Simplifier.simp_add]), + ((Binding.name "pat_rews" , pat_rews ), [Simplifier.simp_add]), + ((Binding.name "dist_les" , dist_les ), [Simplifier.simp_add]), + ((Binding.name "dist_eqs" , dist_eqs ), [Simplifier.simp_add]), + ((Binding.name "inverts" , inverts ), [Simplifier.simp_add]), + ((Binding.name "injects" , injects ), [Simplifier.simp_add]), + ((Binding.name "copy_rews" , copy_rews ), [Simplifier.simp_add]), + ((Binding.name "match_rews", mat_rews ), [Simplifier.simp_add])] + |> Sign.parent_path + |> pair (iso_rews @ when_rews @ con_rews @ sel_rews @ dis_rews @ + pat_rews @ dist_les @ dist_eqs @ copy_rews) +end; (* let *) + +fun comp_theorems (comp_dnam, eqs: eq list) thy = +let +val global_ctxt = ProofContext.init thy; + +val dnames = map (fst o fst) eqs; +val conss = map snd eqs; +val comp_dname = Sign.full_bname thy comp_dnam; + +val _ = message ("Proving induction properties of domain "^comp_dname^" ..."); +val pg = pg' thy; + +(* ----- getting the composite axiom and definitions ------------------------ *) + +local + fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s); +in + val axs_reach = map (ga "reach" ) dnames; + val axs_take_def = map (ga "take_def" ) dnames; + val axs_finite_def = map (ga "finite_def") dnames; + val ax_copy2_def = ga "copy_def" comp_dnam; + val ax_bisim_def = ga "bisim_def" comp_dnam; +end; + +local + fun gt s dn = PureThy.get_thm thy (dn ^ "." ^ s); + fun gts s dn = PureThy.get_thms thy (dn ^ "." ^ s); +in + val cases = map (gt "casedist" ) dnames; + val con_rews = maps (gts "con_rews" ) dnames; + val copy_rews = maps (gts "copy_rews") dnames; +end; + +fun dc_take dn = %%:(dn^"_take"); +val x_name = idx_name dnames "x"; +val P_name = idx_name dnames "P"; +val n_eqs = length eqs; + +(* ----- theorems concerning finite approximation and finite induction ------ *) + +local + val iterate_Cprod_ss = global_simpset_of @{theory Fix}; + val copy_con_rews = copy_rews @ con_rews; + val copy_take_defs = + (if n_eqs = 1 then [] else [ax_copy2_def]) @ axs_take_def; + val _ = trace " Proving take_stricts..."; + val take_stricts = + let + fun one_eq ((dn, args), _) = strict (dc_take dn $ %:"n"); + val goal = mk_trp (foldr1 mk_conj (map one_eq eqs)); + fun tacs ctxt = [ + InductTacs.induct_tac ctxt [[SOME "n"]] 1, + simp_tac iterate_Cprod_ss 1, + asm_simp_tac (iterate_Cprod_ss addsimps copy_rews) 1]; + in pg copy_take_defs goal tacs end; + + val take_stricts' = rewrite_rule copy_take_defs take_stricts; + fun take_0 n dn = + let + val goal = mk_trp ((dc_take dn $ %%:"HOL.zero") `% x_name n === UU); + in pg axs_take_def goal (K [simp_tac iterate_Cprod_ss 1]) end; + val take_0s = mapn take_0 1 dnames; + fun c_UU_tac ctxt = case_UU_tac ctxt (take_stricts'::copy_con_rews) 1; + val _ = trace " Proving take_apps..."; + val take_apps = + let + fun mk_eqn dn (con, args) = + let + fun mk_take n = dc_take (List.nth (dnames, n)) $ %:"n"; + fun one_rhs arg = + if DatatypeAux.is_rec_type (dtyp_of arg) + then Domain_Axioms.copy_of_dtyp mk_take (dtyp_of arg) ` (%# arg) + else (%# arg); + val lhs = (dc_take dn $ (%%:"Suc" $ %:"n"))`(con_app con args); + val rhs = con_app2 con one_rhs args; + in Library.foldr mk_all (map vname args, lhs === rhs) end; + fun mk_eqns ((dn, _), cons) = map (mk_eqn dn) cons; + val goal = mk_trp (foldr1 mk_conj (maps mk_eqns eqs)); + val simps = List.filter (has_fewer_prems 1) copy_rews; + fun con_tac ctxt (con, args) = + if nonlazy_rec args = [] + then all_tac + else EVERY (map (c_UU_tac ctxt) (nonlazy_rec args)) THEN + asm_full_simp_tac (HOLCF_ss addsimps copy_rews) 1; + fun eq_tacs ctxt ((dn, _), cons) = map (con_tac ctxt) cons; + fun tacs ctxt = + simp_tac iterate_Cprod_ss 1 :: + InductTacs.induct_tac ctxt [[SOME "n"]] 1 :: + simp_tac (iterate_Cprod_ss addsimps copy_con_rews) 1 :: + asm_full_simp_tac (HOLCF_ss addsimps simps) 1 :: + TRY (safe_tac HOL_cs) :: + maps (eq_tacs ctxt) eqs; + in pg copy_take_defs goal tacs end; +in + val take_rews = map standard + (atomize global_ctxt take_stricts @ take_0s @ atomize global_ctxt take_apps); +end; (* local *) + +local + fun one_con p (con,args) = + let + fun ind_hyp arg = %:(P_name (1 + rec_of arg)) $ bound_arg args arg; + val t1 = mk_trp (%:p $ con_app2 con (bound_arg args) args); + val t2 = lift ind_hyp (List.filter is_rec args, t1); + val t3 = lift_defined (bound_arg (map vname args)) (nonlazy args, t2); + in Library.foldr mk_All (map vname args, t3) end; + + fun one_eq ((p, cons), concl) = + mk_trp (%:p $ UU) ===> Logic.list_implies (map (one_con p) cons, concl); + + fun ind_term concf = Library.foldr one_eq + (mapn (fn n => fn x => (P_name n, x)) 1 conss, + mk_trp (foldr1 mk_conj (mapn concf 1 dnames))); + val take_ss = HOL_ss addsimps take_rews; + fun quant_tac ctxt i = EVERY + (mapn (fn n => fn _ => res_inst_tac ctxt [(("x", 0), x_name n)] spec i) 1 dnames); + + fun ind_prems_tac prems = EVERY + (maps (fn cons => + (resolve_tac prems 1 :: + maps (fn (_,args) => + resolve_tac prems 1 :: + map (K(atac 1)) (nonlazy args) @ + map (K(atac 1)) (List.filter is_rec args)) + cons)) + conss); + local + (* check whether every/exists constructor of the n-th part of the equation: + it has a possibly indirectly recursive argument that isn't/is possibly + indirectly lazy *) + fun rec_to quant nfn rfn ns lazy_rec (n,cons) = quant (exists (fn arg => + is_rec arg andalso not(rec_of arg mem ns) andalso + ((rec_of arg = n andalso nfn(lazy_rec orelse is_lazy arg)) orelse + rec_of arg <> n andalso rec_to quant nfn rfn (rec_of arg::ns) + (lazy_rec orelse is_lazy arg) (n, (List.nth(conss,rec_of arg)))) + ) o snd) cons; + fun all_rec_to ns = rec_to forall not all_rec_to ns; + fun warn (n,cons) = + if all_rec_to [] false (n,cons) + then (warning ("domain "^List.nth(dnames,n)^" is empty!"); true) + else false; + fun lazy_rec_to ns = rec_to exists I lazy_rec_to ns; + + in + val n__eqs = mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs; + val is_emptys = map warn n__eqs; + val is_finite = forall (not o lazy_rec_to [] false) n__eqs; + end; +in (* local *) + val _ = trace " Proving finite_ind..."; + val finite_ind = + let + fun concf n dn = %:(P_name n) $ (dc_take dn $ %:"n" `%(x_name n)); + val goal = ind_term concf; + + fun tacf {prems, context} = + let + val tacs1 = [ + quant_tac context 1, + simp_tac HOL_ss 1, + InductTacs.induct_tac context [[SOME "n"]] 1, + simp_tac (take_ss addsimps prems) 1, + TRY (safe_tac HOL_cs)]; + fun arg_tac arg = + case_UU_tac context (prems @ con_rews) 1 + (List.nth (dnames, rec_of arg) ^ "_take n$" ^ vname arg); + fun con_tacs (con, args) = + asm_simp_tac take_ss 1 :: + map arg_tac (List.filter is_nonlazy_rec args) @ + [resolve_tac prems 1] @ + map (K (atac 1)) (nonlazy args) @ + map (K (etac spec 1)) (List.filter is_rec args); + fun cases_tacs (cons, cases) = + res_inst_tac context [(("x", 0), "x")] cases 1 :: + asm_simp_tac (take_ss addsimps prems) 1 :: + maps con_tacs cons; + in + tacs1 @ maps cases_tacs (conss ~~ cases) + end; + in pg'' thy [] goal tacf + handle ERROR _ => (warning "Proof of finite_ind failed."; TrueI) + end; + + val _ = trace " Proving take_lemmas..."; + val take_lemmas = + let + fun take_lemma n (dn, ax_reach) = + let + val lhs = dc_take dn $ Bound 0 `%(x_name n); + val rhs = dc_take dn $ Bound 0 `%(x_name n^"'"); + val concl = mk_trp (%:(x_name n) === %:(x_name n^"'")); + val goal = mk_All ("n", mk_trp (lhs === rhs)) ===> concl; + fun tacf {prems, context} = [ + res_inst_tac context [(("t", 0), x_name n )] (ax_reach RS subst) 1, + res_inst_tac context [(("t", 0), x_name n^"'")] (ax_reach RS subst) 1, + stac fix_def2 1, + REPEAT (CHANGED + (rtac (contlub_cfun_arg RS ssubst) 1 THEN chain_tac 1)), + stac contlub_cfun_fun 1, + stac contlub_cfun_fun 2, + rtac lub_equal 3, + chain_tac 1, + rtac allI 1, + resolve_tac prems 1]; + in pg'' thy axs_take_def goal tacf end; + in mapn take_lemma 1 (dnames ~~ axs_reach) end; + +(* ----- theorems concerning finiteness and induction ----------------------- *) + + val _ = trace " Proving finites, ind..."; + val (finites, ind) = + ( + if is_finite + then (* finite case *) + let + fun take_enough dn = mk_ex ("n",dc_take dn $ Bound 0 ` %:"x" === %:"x"); + fun dname_lemma dn = + let + val prem1 = mk_trp (defined (%:"x")); + val disj1 = mk_all ("n", dc_take dn $ Bound 0 ` %:"x" === UU); + val prem2 = mk_trp (mk_disj (disj1, take_enough dn)); + val concl = mk_trp (take_enough dn); + val goal = prem1 ===> prem2 ===> concl; + val tacs = [ + etac disjE 1, + etac notE 1, + resolve_tac take_lemmas 1, + asm_simp_tac take_ss 1, + atac 1]; + in pg [] goal (K tacs) end; + val _ = trace " Proving finite_lemmas1a"; + val finite_lemmas1a = map dname_lemma dnames; + + val _ = trace " Proving finite_lemma1b"; + val finite_lemma1b = + let + fun mk_eqn n ((dn, args), _) = + let + val disj1 = dc_take dn $ Bound 1 ` Bound 0 === UU; + val disj2 = dc_take dn $ Bound 1 ` Bound 0 === Bound 0; + in + mk_constrainall + (x_name n, Type (dn,args), mk_disj (disj1, disj2)) + end; + val goal = + mk_trp (mk_all ("n", foldr1 mk_conj (mapn mk_eqn 1 eqs))); + fun arg_tacs ctxt vn = [ + eres_inst_tac ctxt [(("x", 0), vn)] all_dupE 1, + etac disjE 1, + asm_simp_tac (HOL_ss addsimps con_rews) 1, + asm_simp_tac take_ss 1]; + fun con_tacs ctxt (con, args) = + asm_simp_tac take_ss 1 :: + maps (arg_tacs ctxt) (nonlazy_rec args); + fun foo_tacs ctxt n (cons, cases) = + simp_tac take_ss 1 :: + rtac allI 1 :: + res_inst_tac ctxt [(("x", 0), x_name n)] cases 1 :: + asm_simp_tac take_ss 1 :: + maps (con_tacs ctxt) cons; + fun tacs ctxt = + rtac allI 1 :: + InductTacs.induct_tac ctxt [[SOME "n"]] 1 :: + simp_tac take_ss 1 :: + TRY (safe_tac (empty_cs addSEs [conjE] addSIs [conjI])) :: + flat (mapn (foo_tacs ctxt) 1 (conss ~~ cases)); + in pg [] goal tacs end; + + fun one_finite (dn, l1b) = + let + val goal = mk_trp (%%:(dn^"_finite") $ %:"x"); + fun tacs ctxt = [ + case_UU_tac ctxt take_rews 1 "x", + eresolve_tac finite_lemmas1a 1, + step_tac HOL_cs 1, + step_tac HOL_cs 1, + cut_facts_tac [l1b] 1, + fast_tac HOL_cs 1]; + in pg axs_finite_def goal tacs end; + + val _ = trace " Proving finites"; + val finites = map one_finite (dnames ~~ atomize global_ctxt finite_lemma1b); + val _ = trace " Proving ind"; + val ind = + let + fun concf n dn = %:(P_name n) $ %:(x_name n); + fun tacf {prems, context} = + let + fun finite_tacs (finite, fin_ind) = [ + rtac(rewrite_rule axs_finite_def finite RS exE)1, + etac subst 1, + rtac fin_ind 1, + ind_prems_tac prems]; + in + TRY (safe_tac HOL_cs) :: + maps finite_tacs (finites ~~ atomize global_ctxt finite_ind) + end; + in pg'' thy [] (ind_term concf) tacf end; + in (finites, ind) end (* let *) + + else (* infinite case *) + let + fun one_finite n dn = + read_instantiate global_ctxt [(("P", 0), dn ^ "_finite " ^ x_name n)] excluded_middle; + val finites = mapn one_finite 1 dnames; + + val goal = + let + fun one_adm n _ = mk_trp (mk_adm (%:(P_name n))); + fun concf n dn = %:(P_name n) $ %:(x_name n); + in Logic.list_implies (mapn one_adm 1 dnames, ind_term concf) end; + fun tacf {prems, context} = + map (fn ax_reach => rtac (ax_reach RS subst) 1) axs_reach @ [ + quant_tac context 1, + rtac (adm_impl_admw RS wfix_ind) 1, + REPEAT_DETERM (rtac adm_all 1), + REPEAT_DETERM ( + TRY (rtac adm_conj 1) THEN + rtac adm_subst 1 THEN + cont_tacR 1 THEN resolve_tac prems 1), + strip_tac 1, + rtac (rewrite_rule axs_take_def finite_ind) 1, + ind_prems_tac prems]; + val ind = (pg'' thy [] goal tacf + handle ERROR _ => + (warning "Cannot prove infinite induction rule"; refl)); + in (finites, ind) end + ) + handle THM _ => + (warning "Induction proofs failed (THM raised)."; ([], TrueI)) + | ERROR _ => + (warning "Induction proofs failed (ERROR raised)."; ([], TrueI)); + + +end; (* local *) + +(* ----- theorem concerning coinduction ------------------------------------- *) + +local + val xs = mapn (fn n => K (x_name n)) 1 dnames; + fun bnd_arg n i = Bound(2*(n_eqs - n)-i-1); + val take_ss = HOL_ss addsimps take_rews; + val sproj = prj (fn s => K("fst("^s^")")) (fn s => K("snd("^s^")")); + val _ = trace " Proving coind_lemma..."; + val coind_lemma = + let + fun mk_prj n _ = proj (%:"R") eqs n $ bnd_arg n 0 $ bnd_arg n 1; + fun mk_eqn n dn = + (dc_take dn $ %:"n" ` bnd_arg n 0) === + (dc_take dn $ %:"n" ` bnd_arg n 1); + fun mk_all2 (x,t) = mk_all (x, mk_all (x^"'", t)); + val goal = + mk_trp (mk_imp (%%:(comp_dname^"_bisim") $ %:"R", + Library.foldr mk_all2 (xs, + Library.foldr mk_imp (mapn mk_prj 0 dnames, + foldr1 mk_conj (mapn mk_eqn 0 dnames))))); + fun x_tacs ctxt n x = [ + rotate_tac (n+1) 1, + etac all2E 1, + eres_inst_tac ctxt [(("P", 1), sproj "R" eqs n^" "^x^" "^x^"'")] (mp RS disjE) 1, + TRY (safe_tac HOL_cs), + REPEAT (CHANGED (asm_simp_tac take_ss 1))]; + fun tacs ctxt = [ + rtac impI 1, + InductTacs.induct_tac ctxt [[SOME "n"]] 1, + simp_tac take_ss 1, + safe_tac HOL_cs] @ + flat (mapn (x_tacs ctxt) 0 xs); + in pg [ax_bisim_def] goal tacs end; +in + val _ = trace " Proving coind..."; + val coind = + let + fun mk_prj n x = mk_trp (proj (%:"R") eqs n $ %:x $ %:(x^"'")); + fun mk_eqn x = %:x === %:(x^"'"); + val goal = + mk_trp (%%:(comp_dname^"_bisim") $ %:"R") ===> + Logic.list_implies (mapn mk_prj 0 xs, + mk_trp (foldr1 mk_conj (map mk_eqn xs))); + val tacs = + TRY (safe_tac HOL_cs) :: + maps (fn take_lemma => [ + rtac take_lemma 1, + cut_facts_tac [coind_lemma] 1, + fast_tac HOL_cs 1]) + take_lemmas; + in pg [] goal (K tacs) end; +end; (* local *) + +val inducts = ProjectRule.projections (ProofContext.init thy) ind; +fun ind_rule (dname, rule) = ((Binding.empty, [rule]), [Induct.induct_type dname]); +val induct_failed = (Thm.prop_of ind = Thm.prop_of TrueI); + +in thy |> Sign.add_path comp_dnam + |> snd o PureThy.add_thmss [ + ((Binding.name "take_rews" , take_rews ), [Simplifier.simp_add]), + ((Binding.name "take_lemmas", take_lemmas ), []), + ((Binding.name "finites" , finites ), []), + ((Binding.name "finite_ind" , [finite_ind]), []), + ((Binding.name "ind" , [ind] ), []), + ((Binding.name "coind" , [coind] ), [])] + |> (if induct_failed then I + else snd o PureThy.add_thmss (map ind_rule (dnames ~~ inducts))) + |> Sign.parent_path |> pair take_rews +end; (* let *) +end; (* local *) +end; (* struct *) diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/Tools/adm_tac.ML --- a/src/HOLCF/Tools/adm_tac.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOLCF/Tools/adm_tac.ML Thu Jul 23 21:13:21 2009 +0200 @@ -117,8 +117,8 @@ val tye' = Sign.typ_match thy (PT, #T (rep_cterm Pt)) tye; val ctye = map (fn (ixn, (S, T)) => (ctyp_of thy (TVar (ixn, S)), ctyp_of thy T)) (Vartab.dest tye'); - val tv = cterm_of thy (Var (("t", j), Envir.typ_subst_TVars tye' tT)); - val Pv = cterm_of thy (Var (("P", j), Envir.typ_subst_TVars tye' PT)); + val tv = cterm_of thy (Var (("t", j), Envir.subst_type tye' tT)); + val Pv = cterm_of thy (Var (("P", j), Envir.subst_type tye' PT)); val rule' = instantiate (ctye, [(tv, tt), (Pv, Pt)]) rule in rule' end; diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/Tools/domain/domain_axioms.ML --- a/src/HOLCF/Tools/domain/domain_axioms.ML Thu Jul 23 21:12:57 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,235 +0,0 @@ -(* Title: HOLCF/Tools/domain/domain_axioms.ML - Author: David von Oheimb - -Syntax generator for domain command. -*) - -signature DOMAIN_AXIOMS = -sig - val copy_of_dtyp : (int -> term) -> Datatype.dtyp -> term - - val calc_axioms : - string -> Domain_Library.eq list -> int -> Domain_Library.eq -> - string * (string * term) list * (string * term) list - - val add_axioms : - bstring -> Domain_Library.eq list -> theory -> theory -end; - - -structure Domain_Axioms :> DOMAIN_AXIOMS = -struct - -open Domain_Library; - -infixr 0 ===>;infixr 0 ==>;infix 0 == ; -infix 1 ===; infix 1 ~= ; infix 1 <<; infix 1 ~<<; -infix 9 ` ; infix 9 `% ; infix 9 `%%; infixr 9 oo; - -(* FIXME: use theory data for this *) -val copy_tab : string Symtab.table = - Symtab.make [(@{type_name "->"}, @{const_name "cfun_fun"}), - (@{type_name "++"}, @{const_name "ssum_fun"}), - (@{type_name "**"}, @{const_name "sprod_fun"}), - (@{type_name "*"}, @{const_name "cprod_fun"}), - (@{type_name "u"}, @{const_name "u_fun"})]; - -fun copy_of_dtyp r dt = if DatatypeAux.is_rec_type dt then copy r dt else ID -and copy r (DatatypeAux.DtRec i) = r i - | copy r (DatatypeAux.DtTFree a) = ID - | copy r (DatatypeAux.DtType (c, ds)) = - case Symtab.lookup copy_tab c of - SOME f => list_ccomb (%%:f, map (copy_of_dtyp r) ds) - | NONE => (warning ("copy_of_dtyp: unknown type constructor " ^ c); ID); - -fun calc_axioms - (comp_dname : string) - (eqs : eq list) - (n : int) - (eqn as ((dname,_),cons) : eq) - : string * (string * term) list * (string * term) list = - let - - (* ----- axioms and definitions concerning the isomorphism ------------------ *) - - val dc_abs = %%:(dname^"_abs"); - val dc_rep = %%:(dname^"_rep"); - val x_name'= "x"; - val x_name = idx_name eqs x_name' (n+1); - val dnam = Long_Name.base_name dname; - - val abs_iso_ax = ("abs_iso", mk_trp(dc_rep`(dc_abs`%x_name') === %:x_name')); - val rep_iso_ax = ("rep_iso", mk_trp(dc_abs`(dc_rep`%x_name') === %:x_name')); - - val when_def = ("when_def",%%:(dname^"_when") == - List.foldr (uncurry /\ ) (/\x_name'((when_body cons (fn (x,y) => - Bound(1+length cons+x-y)))`(dc_rep`Bound 0))) (when_funs cons)); - - val copy_def = - let fun r i = cproj (Bound 0) eqs i; - in ("copy_def", %%:(dname^"_copy") == - /\ "f" (dc_abs oo (copy_of_dtyp r (dtyp_of_eq eqn)) oo dc_rep)) end; - - (* -- definitions concerning the constructors, discriminators and selectors - *) - - fun con_def m n (_,args) = let - fun idxs z x arg = (if is_lazy arg then mk_up else I) (Bound(z-x)); - fun parms vs = mk_stuple (mapn (idxs(length vs)) 1 vs); - fun inj y 1 _ = y - | inj y _ 0 = mk_sinl y - | inj y i j = mk_sinr (inj y (i-1) (j-1)); - in List.foldr /\# (dc_abs`(inj (parms args) m n)) args end; - - val con_defs = mapn (fn n => fn (con,args) => - (extern_name con ^"_def", %%:con == con_def (length cons) n (con,args))) 0 cons; - - val dis_defs = let - fun ddef (con,_) = (dis_name con ^"_def",%%:(dis_name con) == - list_ccomb(%%:(dname^"_when"),map - (fn (con',args) => (List.foldr /\# - (if con'=con then TT else FF) args)) cons)) - in map ddef cons end; - - val mat_defs = - let - fun mdef (con,_) = - let - val k = Bound 0 - val x = Bound 1 - fun one_con (con', args') = - if con'=con then k else List.foldr /\# mk_fail args' - val w = list_ccomb(%%:(dname^"_when"), map one_con cons) - val rhs = /\ "x" (/\ "k" (w ` x)) - in (mat_name con ^"_def", %%:(mat_name con) == rhs) end - in map mdef cons end; - - val pat_defs = - let - fun pdef (con,args) = - let - val ps = mapn (fn n => fn _ => %:("pat" ^ string_of_int n)) 1 args; - val xs = map (bound_arg args) args; - val r = Bound (length args); - val rhs = case args of [] => mk_return HOLogic.unit - | _ => mk_ctuple_pat ps ` mk_ctuple xs; - fun one_con (con',args') = List.foldr /\# (if con'=con then rhs else mk_fail) args'; - in (pat_name con ^"_def", list_comb (%%:(pat_name con), ps) == - list_ccomb(%%:(dname^"_when"), map one_con cons)) - end - in map pdef cons end; - - val sel_defs = let - fun sdef con n arg = Option.map (fn sel => (sel^"_def",%%:sel == - list_ccomb(%%:(dname^"_when"),map - (fn (con',args) => if con'<>con then UU else - List.foldr /\# (Bound (length args - n)) args) cons))) (sel_of arg); - in List.mapPartial I (List.concat(map (fn (con,args) => mapn (sdef con) 1 args) cons)) end; - - - (* ----- axiom and definitions concerning induction ------------------------- *) - - val reach_ax = ("reach", mk_trp(cproj (mk_fix (%%:(comp_dname^"_copy"))) eqs n - `%x_name === %:x_name)); - val take_def = - ("take_def", - %%:(dname^"_take") == - mk_lam("n",cproj - (mk_iterate (Bound 0, %%:(comp_dname^"_copy"), UU)) eqs n)); - val finite_def = - ("finite_def", - %%:(dname^"_finite") == - mk_lam(x_name, - mk_ex("n",(%%:(dname^"_take") $ Bound 0)`Bound 1 === Bound 1))); - - in (dnam, - [abs_iso_ax, rep_iso_ax, reach_ax], - [when_def, copy_def] @ - con_defs @ dis_defs @ mat_defs @ pat_defs @ sel_defs @ - [take_def, finite_def]) - end; (* let (calc_axioms) *) - - -(* legacy type inference *) - -fun legacy_infer_term thy t = - singleton (Syntax.check_terms (ProofContext.init thy)) (Sign.intern_term thy t); - -fun legacy_infer_prop thy t = legacy_infer_term thy (TypeInfer.constrain propT t); - -fun infer_props thy = map (apsnd (legacy_infer_prop thy)); - - -fun add_axioms_i x = snd o PureThy.add_axioms (map (Thm.no_attributes o apfst Binding.name) x); -fun add_axioms_infer axms thy = add_axioms_i (infer_props thy axms) thy; - -fun add_defs_i x = snd o (PureThy.add_defs false) (map (Thm.no_attributes o apfst Binding.name) x); -fun add_defs_infer defs thy = add_defs_i (infer_props thy defs) thy; - -fun add_matchers (((dname,_),cons) : eq) thy = - let - val con_names = map fst cons; - val mat_names = map mat_name con_names; - fun qualify n = Sign.full_name thy (Binding.name n); - val ms = map qualify con_names ~~ map qualify mat_names; - in Fixrec.add_matchers ms thy end; - -fun add_axioms comp_dnam (eqs : eq list) thy' = - let - val comp_dname = Sign.full_bname thy' comp_dnam; - val dnames = map (fst o fst) eqs; - val x_name = idx_name dnames "x"; - fun copy_app dname = %%:(dname^"_copy")`Bound 0; - val copy_def = ("copy_def" , %%:(comp_dname^"_copy") == - /\ "f"(mk_ctuple (map copy_app dnames))); - - fun one_con (con,args) = let - val nonrec_args = filter_out is_rec args; - val rec_args = List.filter is_rec args; - val recs_cnt = length rec_args; - val allargs = nonrec_args @ rec_args - @ map (upd_vname (fn s=> s^"'")) rec_args; - val allvns = map vname allargs; - fun vname_arg s arg = if is_rec arg then vname arg^s else vname arg; - val vns1 = map (vname_arg "" ) args; - val vns2 = map (vname_arg "'") args; - val allargs_cnt = length nonrec_args + 2*recs_cnt; - val rec_idxs = (recs_cnt-1) downto 0; - val nonlazy_idxs = map snd (filter_out (fn (arg,_) => is_lazy arg) - (allargs~~((allargs_cnt-1) downto 0))); - fun rel_app i ra = proj (Bound(allargs_cnt+2)) eqs (rec_of ra) $ - Bound (2*recs_cnt-i) $ Bound (recs_cnt-i); - val capps = - List.foldr mk_conj - (mk_conj( - Bound(allargs_cnt+1)===list_ccomb(%%:con,map (bound_arg allvns) vns1), - Bound(allargs_cnt+0)===list_ccomb(%%:con,map (bound_arg allvns) vns2))) - (mapn rel_app 1 rec_args); - in List.foldr mk_ex - (Library.foldr mk_conj - (map (defined o Bound) nonlazy_idxs,capps)) allvns - end; - fun one_comp n (_,cons) = - mk_all(x_name(n+1), - mk_all(x_name(n+1)^"'", - mk_imp(proj (Bound 2) eqs n $ Bound 1 $ Bound 0, - foldr1 mk_disj (mk_conj(Bound 1 === UU,Bound 0 === UU) - ::map one_con cons)))); - val bisim_def = - ("bisim_def", - %%:(comp_dname^"_bisim")==mk_lam("R", foldr1 mk_conj (mapn one_comp 0 eqs))); - - fun add_one (thy,(dnam,axs,dfs)) = - thy |> Sign.add_path dnam - |> add_defs_infer dfs - |> add_axioms_infer axs - |> Sign.parent_path; - - val thy = Library.foldl add_one (thy', mapn (calc_axioms comp_dname eqs) 0 eqs); - - in thy |> Sign.add_path comp_dnam - |> add_defs_infer (bisim_def::(if length eqs>1 then [copy_def] else [])) - |> Sign.parent_path - |> fold add_matchers eqs - end; (* let (add_axioms) *) - -end; (* struct *) diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/Tools/domain/domain_extender.ML --- a/src/HOLCF/Tools/domain/domain_extender.ML Thu Jul 23 21:12:57 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,204 +0,0 @@ -(* Title: HOLCF/Tools/domain/domain_extender.ML - Author: David von Oheimb - -Theory extender for domain command, including theory syntax. -*) - -signature DOMAIN_EXTENDER = -sig - val add_domain_cmd: string -> - ((string * string option) list * binding * mixfix * - (binding * (bool * binding option * string) list * mixfix) list) list - -> theory -> theory - val add_domain: string -> - ((string * string option) list * binding * mixfix * - (binding * (bool * binding option * typ) list * mixfix) list) list - -> theory -> theory -end; - -structure Domain_Extender :> DOMAIN_EXTENDER = -struct - -open Domain_Library; - -(* ----- general testing and preprocessing of constructor list -------------- *) -fun check_and_sort_domain - (dtnvs : (string * typ list) list) - (cons'' : (binding * (bool * binding option * typ) list * mixfix) list list) - (sg : theory) - : ((string * typ list) * - (binding * (bool * binding option * typ) list * mixfix) list) list = - let - val defaultS = Sign.defaultS sg; - val test_dupl_typs = (case duplicates (op =) (map fst dtnvs) of - [] => false | dups => error ("Duplicate types: " ^ commas_quote dups)); - val test_dupl_cons = - (case duplicates (op =) (map (Binding.name_of o first) (List.concat cons'')) of - [] => false | dups => error ("Duplicate constructors: " - ^ commas_quote dups)); - val test_dupl_sels = - (case duplicates (op =) (map Binding.name_of (List.mapPartial second - (List.concat (map second (List.concat cons''))))) of - [] => false | dups => error("Duplicate selectors: "^commas_quote dups)); - val test_dupl_tvars = - exists(fn s=>case duplicates (op =) (map(fst o dest_TFree)s)of - [] => false | dups => error("Duplicate type arguments: " - ^commas_quote dups)) (map snd dtnvs); - (* test for free type variables, illegal sort constraints on rhs, - non-pcpo-types and invalid use of recursive type; - replace sorts in type variables on rhs *) - fun analyse_equation ((dname,typevars),cons') = - let - val tvars = map dest_TFree typevars; - val distinct_typevars = map TFree tvars; - fun rm_sorts (TFree(s,_)) = TFree(s,[]) - | rm_sorts (Type(s,ts)) = Type(s,remove_sorts ts) - | rm_sorts (TVar(s,_)) = TVar(s,[]) - and remove_sorts l = map rm_sorts l; - val indirect_ok = ["*","Cfun.->","Ssum.++","Sprod.**","Up.u"] - fun analyse indirect (TFree(v,s)) = - (case AList.lookup (op =) tvars v of - NONE => error ("Free type variable " ^ quote v ^ " on rhs.") - | SOME sort => if eq_set_string (s,defaultS) orelse - eq_set_string (s,sort ) - then TFree(v,sort) - else error ("Inconsistent sort constraint" ^ - " for type variable " ^ quote v)) - | analyse indirect (t as Type(s,typl)) = - (case AList.lookup (op =) dtnvs s of - NONE => if s mem indirect_ok - then Type(s,map (analyse false) typl) - else Type(s,map (analyse true) typl) - | SOME typevars => if indirect - then error ("Indirect recursion of type " ^ - quote (string_of_typ sg t)) - else if dname <> s orelse - (** BUG OR FEATURE?: - mutual recursion may use different arguments **) - remove_sorts typevars = remove_sorts typl - then Type(s,map (analyse true) typl) - else error ("Direct recursion of type " ^ - quote (string_of_typ sg t) ^ - " with different arguments")) - | analyse indirect (TVar _) = Imposs "extender:analyse"; - fun check_pcpo lazy T = - let val ok = if lazy then cpo_type else pcpo_type - in if ok sg T then T else error - ("Constructor argument type is not of sort pcpo: " ^ - string_of_typ sg T) - end; - fun analyse_arg (lazy, sel, T) = - (lazy, sel, check_pcpo lazy (analyse false T)); - fun analyse_con (b, args, mx) = (b, map analyse_arg args, mx); - in ((dname,distinct_typevars), map analyse_con cons') end; - in ListPair.map analyse_equation (dtnvs,cons'') - end; (* let *) - -(* ----- calls for building new thy and thms -------------------------------- *) - -fun gen_add_domain - (prep_typ : theory -> 'a -> typ) - (comp_dnam : string) - (eqs''' : ((string * string option) list * binding * mixfix * - (binding * (bool * binding option * 'a) list * mixfix) list) list) - (thy''' : theory) = - let - fun readS (SOME s) = Syntax.read_sort_global thy''' s - | readS NONE = Sign.defaultS thy'''; - fun readTFree (a, s) = TFree (a, readS s); - - val dtnvs = map (fn (vs,dname:binding,mx,_) => - (dname, map readTFree vs, mx)) eqs'''; - val cons''' = map (fn (_,_,_,cons) => cons) eqs'''; - fun thy_type (dname,tvars,mx) = (dname, length tvars, mx); - fun thy_arity (dname,tvars,mx) = (Sign.full_name thy''' dname, map (snd o dest_TFree) tvars, pcpoS); - val thy'' = thy''' |> Sign.add_types (map thy_type dtnvs) - |> fold (AxClass.axiomatize_arity o thy_arity) dtnvs; - val cons'' = map (map (upd_second (map (upd_third (prep_typ thy''))))) cons'''; - val dtnvs' = map (fn (dname,vs,mx) => (Sign.full_name thy''' dname,vs)) dtnvs; - val eqs' : ((string * typ list) * (binding * (bool * binding option * typ) list * mixfix) list) list = - check_and_sort_domain dtnvs' cons'' thy''; - val thy' = thy'' |> Domain_Syntax.add_syntax comp_dnam eqs'; - val dts = map (Type o fst) eqs'; - val new_dts = map (fn ((s,Ts),_) => (s, map (fst o dest_TFree) Ts)) eqs'; - fun strip ss = Library.drop (find_index (fn s => s = "'") ss + 1, ss); - fun typid (Type (id,_)) = - let val c = hd (Symbol.explode (Long_Name.base_name id)) - in if Symbol.is_letter c then c else "t" end - | typid (TFree (id,_) ) = hd (strip (tl (Symbol.explode id))) - | typid (TVar ((id,_),_)) = hd (tl (Symbol.explode id)); - fun one_con (con,args,mx) = - ((Syntax.const_name mx (Binding.name_of con)), - ListPair.map (fn ((lazy,sel,tp),vn) => mk_arg ((lazy, - DatatypeAux.dtyp_of_typ new_dts tp), - Option.map Binding.name_of sel,vn)) - (args,(mk_var_names(map (typid o third) args))) - ) : cons; - val eqs = map (fn (dtnvs,cons') => (dtnvs, map one_con cons')) eqs' : eq list; - val thy = thy' |> Domain_Axioms.add_axioms comp_dnam eqs; - val ((rewss, take_rews), theorems_thy) = - thy |> fold_map (fn eq => Domain_Theorems.theorems (eq, eqs)) eqs - ||>> Domain_Theorems.comp_theorems (comp_dnam, eqs); - in - theorems_thy - |> Sign.add_path (Long_Name.base_name comp_dnam) - |> (snd o (PureThy.add_thmss [((Binding.name "rews", List.concat rewss @ take_rews), [])])) - |> Sign.parent_path - end; - -val add_domain = gen_add_domain Sign.certify_typ; -val add_domain_cmd = gen_add_domain Syntax.read_typ_global; - - -(** outer syntax **) - -local structure P = OuterParse and K = OuterKeyword in - -val _ = OuterKeyword.keyword "lazy"; - -val dest_decl : (bool * binding option * string) parser = - P.$$$ "(" |-- Scan.optional (P.$$$ "lazy" >> K true) false -- - (P.binding >> SOME) -- (P.$$$ "::" |-- P.typ) --| P.$$$ ")" >> P.triple1 - || P.$$$ "(" |-- P.$$$ "lazy" |-- P.typ --| P.$$$ ")" - >> (fn t => (true,NONE,t)) - || P.typ >> (fn t => (false,NONE,t)); - -val cons_decl = - P.binding -- Scan.repeat dest_decl -- P.opt_mixfix; - -val type_var' : (string * string option) parser = - (P.type_ident -- Scan.option (P.$$$ "::" |-- P.!!! P.sort)); - -val type_args' : (string * string option) list parser = - type_var' >> single || - P.$$$ "(" |-- P.!!! (P.list1 type_var' --| P.$$$ ")") || - Scan.succeed []; - -val domain_decl = - (type_args' -- P.binding -- P.opt_infix) -- - (P.$$$ "=" |-- P.enum1 "|" cons_decl); - -val domains_decl = - Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- - P.and_list1 domain_decl; - -fun mk_domain (opt_name : string option, - doms : ((((string * string option) list * binding) * mixfix) * - ((binding * (bool * binding option * string) list) * mixfix) list) list ) = - let - val names = map (fn (((_, t), _), _) => Binding.name_of t) doms; - val specs : ((string * string option) list * binding * mixfix * - (binding * (bool * binding option * string) list * mixfix) list) list = - map (fn (((vs, t), mx), cons) => - (vs, t, mx, map (fn ((c, ds), mx) => (c, ds, mx)) cons)) doms; - val comp_dnam = - case opt_name of NONE => space_implode "_" names | SOME s => s; - in add_domain_cmd comp_dnam specs end; - -val _ = - OuterSyntax.command "domain" "define recursive domains (HOLCF)" K.thy_decl - (domains_decl >> (Toplevel.theory o mk_domain)); - -end; - -end; diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/Tools/domain/domain_library.ML --- a/src/HOLCF/Tools/domain/domain_library.ML Thu Jul 23 21:12:57 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,399 +0,0 @@ -(* Title: HOLCF/Tools/domain/domain_library.ML - Author: David von Oheimb - -Library for domain command. -*) - - -(* ----- general support ---------------------------------------------------- *) - -fun mapn f n [] = [] - | mapn f n (x::xs) = (f n x) :: mapn f (n+1) xs; - -fun foldr'' f (l,f2) = - let fun itr [] = raise Fail "foldr''" - | itr [a] = f2 a - | itr (a::l) = f(a, itr l) - in itr l end; - -fun map_cumulr f start xs = - List.foldr (fn (x,(ys,res))=>case f(x,res) of (y,res2) => - (y::ys,res2)) ([],start) xs; - -fun first (x,_,_) = x; fun second (_,x,_) = x; fun third (_,_,x) = x; -fun upd_first f (x,y,z) = (f x, y, z); -fun upd_second f (x,y,z) = ( x, f y, z); -fun upd_third f (x,y,z) = ( x, y, f z); - -fun atomize ctxt thm = - let - val r_inst = read_instantiate ctxt; - fun at thm = - case concl_of thm of - _$(Const("op &",_)$_$_) => at(thm RS conjunct1)@at(thm RS conjunct2) - | _$(Const("All" ,_)$Abs(s,_,_))=> at(thm RS (r_inst [(("x", 0), "?" ^ s)] spec)) - | _ => [thm]; - in map zero_var_indexes (at thm) end; - -(* infix syntax *) - -infixr 5 -->; -infixr 6 ->>; -infixr 0 ===>; -infixr 0 ==>; -infix 0 ==; -infix 1 ===; -infix 1 ~=; -infix 1 <<; -infix 1 ~<<; - -infix 9 ` ; -infix 9 `% ; -infix 9 `%%; - - -(* ----- specific support for domain ---------------------------------------- *) - -signature DOMAIN_LIBRARY = -sig - val Imposs : string -> 'a; - val cpo_type : theory -> typ -> bool; - val pcpo_type : theory -> typ -> bool; - val string_of_typ : theory -> typ -> string; - - (* Creating HOLCF types *) - val mk_cfunT : typ * typ -> typ; - val ->> : typ * typ -> typ; - val mk_ssumT : typ * typ -> typ; - val mk_sprodT : typ * typ -> typ; - val mk_uT : typ -> typ; - val oneT : typ; - val trT : typ; - val mk_maybeT : typ -> typ; - val mk_ctupleT : typ list -> typ; - val mk_TFree : string -> typ; - val pcpoS : sort; - - (* Creating HOLCF terms *) - val %: : string -> term; - val %%: : string -> term; - val ` : term * term -> term; - val `% : term * string -> term; - val /\ : string -> term -> term; - val UU : term; - val TT : term; - val FF : term; - val ID : term; - val oo : term * term -> term; - val mk_up : term -> term; - val mk_sinl : term -> term; - val mk_sinr : term -> term; - val mk_stuple : term list -> term; - val mk_ctuple : term list -> term; - val mk_fix : term -> term; - val mk_iterate : term * term * term -> term; - val mk_fail : term; - val mk_return : term -> term; - val cproj : term -> 'a list -> int -> term; - val list_ccomb : term * term list -> term; - (* - val con_app : string -> ('a * 'b * string) list -> term; - *) - val con_app2 : string -> ('a -> term) -> 'a list -> term; - val proj : term -> 'a list -> int -> term; - val prj : ('a -> 'b -> 'a) -> ('a -> 'b -> 'a) -> 'a -> 'b list -> int -> 'a; - val mk_ctuple_pat : term list -> term; - val mk_branch : term -> term; - - (* Creating propositions *) - val mk_conj : term * term -> term; - val mk_disj : term * term -> term; - val mk_imp : term * term -> term; - val mk_lam : string * term -> term; - val mk_all : string * term -> term; - val mk_ex : string * term -> term; - val mk_constrain : typ * term -> term; - val mk_constrainall : string * typ * term -> term; - val === : term * term -> term; - val << : term * term -> term; - val ~<< : term * term -> term; - val strict : term -> term; - val defined : term -> term; - val mk_adm : term -> term; - val mk_compact : term -> term; - val lift : ('a -> term) -> 'a list * term -> term; - val lift_defined : ('a -> term) -> 'a list * term -> term; - - (* Creating meta-propositions *) - val mk_trp : term -> term; (* HOLogic.mk_Trueprop *) - val == : term * term -> term; - val ===> : term * term -> term; - val ==> : term * term -> term; - val mk_All : string * term -> term; - - (* Domain specifications *) - eqtype arg; - type cons = string * arg list; - type eq = (string * typ list) * cons list; - val mk_arg : (bool * Datatype.dtyp) * string option * string -> arg; - val is_lazy : arg -> bool; - val rec_of : arg -> int; - val dtyp_of : arg -> Datatype.dtyp; - val sel_of : arg -> string option; - val vname : arg -> string; - val upd_vname : (string -> string) -> arg -> arg; - val is_rec : arg -> bool; - val is_nonlazy_rec : arg -> bool; - val nonlazy : arg list -> string list; - val nonlazy_rec : arg list -> string list; - val %# : arg -> term; - val /\# : arg * term -> term; - val when_body : cons list -> (int * int -> term) -> term; - val when_funs : 'a list -> string list; - val bound_arg : ''a list -> ''a -> term; (* ''a = arg or string *) - val idx_name : 'a list -> string -> int -> string; - val app_rec_arg : (int -> term) -> arg -> term; - val con_app : string -> arg list -> term; - val dtyp_of_eq : eq -> Datatype.dtyp; - - - (* Name mangling *) - val strip_esc : string -> string; - val extern_name : string -> string; - val dis_name : string -> string; - val mat_name : string -> string; - val pat_name : string -> string; - val mk_var_names : string list -> string list; -end; - -structure Domain_Library :> DOMAIN_LIBRARY = -struct - -exception Impossible of string; -fun Imposs msg = raise Impossible ("Domain:"^msg); - -(* ----- name handling ----- *) - -val strip_esc = - let fun strip ("'" :: c :: cs) = c :: strip cs - | strip ["'"] = [] - | strip (c :: cs) = c :: strip cs - | strip [] = []; - in implode o strip o Symbol.explode end; - -fun extern_name con = - case Symbol.explode con of - ("o"::"p"::" "::rest) => implode rest - | _ => con; -fun dis_name con = "is_"^ (extern_name con); -fun dis_name_ con = "is_"^ (strip_esc con); -fun mat_name con = "match_"^ (extern_name con); -fun mat_name_ con = "match_"^ (strip_esc con); -fun pat_name con = (extern_name con) ^ "_pat"; -fun pat_name_ con = (strip_esc con) ^ "_pat"; - -(* make distinct names out of the type list, - forbidding "o","n..","x..","f..","P.." as names *) -(* a number string is added if necessary *) -fun mk_var_names ids : string list = - let - fun nonreserved s = if s mem ["n","x","f","P"] then s^"'" else s; - fun index_vnames(vn::vns,occupied) = - (case AList.lookup (op =) occupied vn of - NONE => if vn mem vns - then (vn^"1") :: index_vnames(vns,(vn,1) ::occupied) - else vn :: index_vnames(vns, occupied) - | SOME(i) => (vn^(string_of_int (i+1))) - :: index_vnames(vns,(vn,i+1)::occupied)) - | index_vnames([],occupied) = []; - in index_vnames(map nonreserved ids, [("O",0),("o",0)]) end; - -fun cpo_type sg t = Sign.of_sort sg (Sign.certify_typ sg t, @{sort cpo}); -fun pcpo_type sg t = Sign.of_sort sg (Sign.certify_typ sg t, @{sort pcpo}); -fun string_of_typ sg = Syntax.string_of_typ_global sg o Sign.certify_typ sg; - -(* ----- constructor list handling ----- *) - -type arg = - (bool * Datatype.dtyp) * (* (lazy, recursive element) *) - string option * (* selector name *) - string; (* argument name *) - -type cons = - string * (* operator name of constr *) - arg list; (* argument list *) - -type eq = - (string * (* name of abstracted type *) - typ list) * (* arguments of abstracted type *) - cons list; (* represented type, as a constructor list *) - -val mk_arg = I; - -fun rec_of ((_,dtyp),_,_) = - case dtyp of DatatypeAux.DtRec i => i | _ => ~1; -(* FIXME: what about indirect recursion? *) - -fun is_lazy arg = fst (first arg); -fun dtyp_of arg = snd (first arg); -val sel_of = second; -val vname = third; -val upd_vname = upd_third; -fun is_rec arg = rec_of arg >=0; -fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg); -fun nonlazy args = map vname (filter_out is_lazy args); -fun nonlazy_rec args = map vname (List.filter is_nonlazy_rec args); - - -(* ----- combinators for making dtyps ----- *) - -fun mk_uD T = DatatypeAux.DtType(@{type_name "u"}, [T]); -fun mk_sprodD (T, U) = DatatypeAux.DtType(@{type_name "**"}, [T, U]); -fun mk_ssumD (T, U) = DatatypeAux.DtType(@{type_name "++"}, [T, U]); -fun mk_liftD T = DatatypeAux.DtType(@{type_name "lift"}, [T]); -val unitD = DatatypeAux.DtType(@{type_name "unit"}, []); -val boolD = DatatypeAux.DtType(@{type_name "bool"}, []); -val oneD = mk_liftD unitD; -val trD = mk_liftD boolD; -fun big_sprodD ds = case ds of [] => oneD | _ => foldr1 mk_sprodD ds; -fun big_ssumD ds = case ds of [] => unitD | _ => foldr1 mk_ssumD ds; - -fun dtyp_of_arg ((lazy, D), _, _) = if lazy then mk_uD D else D; -fun dtyp_of_cons (_, args) = big_sprodD (map dtyp_of_arg args); -fun dtyp_of_eq (_, cons) = big_ssumD (map dtyp_of_cons cons); - - -(* ----- support for type and mixfix expressions ----- *) - -fun mk_uT T = Type(@{type_name "u"}, [T]); -fun mk_cfunT (T, U) = Type(@{type_name "->"}, [T, U]); -fun mk_sprodT (T, U) = Type(@{type_name "**"}, [T, U]); -fun mk_ssumT (T, U) = Type(@{type_name "++"}, [T, U]); -val oneT = @{typ one}; -val trT = @{typ tr}; - -val op ->> = mk_cfunT; - -fun mk_TFree s = TFree ("'" ^ s, @{sort pcpo}); - -(* ----- support for term expressions ----- *) - -fun %: s = Free(s,dummyT); -fun %# arg = %:(vname arg); -fun %%: s = Const(s,dummyT); - -local open HOLogic in -val mk_trp = mk_Trueprop; -fun mk_conj (S,T) = conj $ S $ T; -fun mk_disj (S,T) = disj $ S $ T; -fun mk_imp (S,T) = imp $ S $ T; -fun mk_lam (x,T) = Abs(x,dummyT,T); -fun mk_all (x,P) = HOLogic.mk_all (x,dummyT,P); -fun mk_ex (x,P) = mk_exists (x,dummyT,P); -val mk_constrain = uncurry TypeInfer.constrain; -fun mk_constrainall (x,typ,P) = %%:"All" $ (TypeInfer.constrain (typ --> boolT) (mk_lam(x,P))); -end - -fun mk_All (x,P) = %%:"all" $ mk_lam(x,P); (* meta universal quantification *) - -infixr 0 ===>; fun S ===> T = %%:"==>" $ S $ T; -infixr 0 ==>; fun S ==> T = mk_trp S ===> mk_trp T; -infix 0 ==; fun S == T = %%:"==" $ S $ T; -infix 1 ===; fun S === T = %%:"op =" $ S $ T; -infix 1 ~=; fun S ~= T = HOLogic.mk_not (S === T); -infix 1 <<; fun S << T = %%: @{const_name Porder.below} $ S $ T; -infix 1 ~<<; fun S ~<< T = HOLogic.mk_not (S << T); - -infix 9 ` ; fun f ` x = %%: @{const_name Rep_CFun} $ f $ x; -infix 9 `% ; fun f`% s = f` %: s; -infix 9 `%%; fun f`%%s = f` %%:s; - -fun mk_adm t = %%: @{const_name adm} $ t; -fun mk_compact t = %%: @{const_name compact} $ t; -val ID = %%: @{const_name ID}; -fun mk_strictify t = %%: @{const_name strictify}`t; -fun mk_cfst t = %%: @{const_name cfst}`t; -fun mk_csnd t = %%: @{const_name csnd}`t; -(*val csplitN = "Cprod.csplit";*) -(*val sfstN = "Sprod.sfst";*) -(*val ssndN = "Sprod.ssnd";*) -fun mk_ssplit t = %%: @{const_name ssplit}`t; -fun mk_sinl t = %%: @{const_name sinl}`t; -fun mk_sinr t = %%: @{const_name sinr}`t; -fun mk_sscase (x, y) = %%: @{const_name sscase}`x`y; -fun mk_up t = %%: @{const_name up}`t; -fun mk_fup (t,u) = %%: @{const_name fup} ` t ` u; -val ONE = @{term ONE}; -val TT = @{term TT}; -val FF = @{term FF}; -fun mk_iterate (n,f,z) = %%: @{const_name iterate} $ n ` f ` z; -fun mk_fix t = %%: @{const_name fix}`t; -fun mk_return t = %%: @{const_name Fixrec.return}`t; -val mk_fail = %%: @{const_name Fixrec.fail}; - -fun mk_branch t = %%: @{const_name Fixrec.branch} $ t; - -val pcpoS = @{sort pcpo}; - -val list_ccomb = Library.foldl (op `); (* continuous version of list_comb *) -fun con_app2 con f args = list_ccomb(%%:con,map f args); -fun con_app con = con_app2 con %#; -fun if_rec arg f y = if is_rec arg then f (rec_of arg) else y; -fun app_rec_arg p arg = if_rec arg (fn n => fn x => (p n)`x) I (%# arg); -fun prj _ _ x ( _::[]) _ = x - | prj f1 _ x (_::y::ys) 0 = f1 x y - | prj f1 f2 x (y:: ys) j = prj f1 f2 (f2 x y) ys (j-1); -fun proj x = prj (fn S => K(%%:"fst" $S)) (fn S => K(%%:"snd" $S)) x; -fun cproj x = prj (fn S => K(mk_cfst S)) (fn S => K(mk_csnd S)) x; -fun lift tfn = Library.foldr (fn (x,t)=> (mk_trp(tfn x) ===> t)); - -fun /\ v T = %%: @{const_name Abs_CFun} $ mk_lam(v,T); -fun /\# (arg,T) = /\ (vname arg) T; -infixr 9 oo; fun S oo T = %%: @{const_name cfcomp}`S`T; -val UU = %%: @{const_name UU}; -fun strict f = f`UU === UU; -fun defined t = t ~= UU; -fun cpair (t,u) = %%: @{const_name cpair}`t`u; -fun spair (t,u) = %%: @{const_name spair}`t`u; -fun mk_ctuple [] = HOLogic.unit (* used in match_defs *) - | mk_ctuple ts = foldr1 cpair ts; -fun mk_stuple [] = ONE - | mk_stuple ts = foldr1 spair ts; -fun mk_ctupleT [] = HOLogic.unitT (* used in match_defs *) - | mk_ctupleT Ts = foldr1 HOLogic.mk_prodT Ts; -fun mk_maybeT T = Type ("Fixrec.maybe",[T]); -fun cpair_pat (p1,p2) = %%: @{const_name cpair_pat} $ p1 $ p2; -val mk_ctuple_pat = foldr1 cpair_pat; -fun lift_defined f = lift (fn x => defined (f x)); -fun bound_arg vns v = Bound (length vns - find_index (fn v' => v' = v) vns - 1); - -fun cont_eta_contract (Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body)) = - (case cont_eta_contract body of - body' as (Const("Cfun.Rep_CFun",Ta) $ f $ Bound 0) => - if not (0 mem loose_bnos f) then incr_boundvars ~1 f - else Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body') - | body' => Const("Cfun.Abs_CFun",TT) $ Abs(a,T,body')) - | cont_eta_contract(f$t) = cont_eta_contract f $ cont_eta_contract t - | cont_eta_contract t = t; - -fun idx_name dnames s n = s^(if length dnames = 1 then "" else string_of_int n); -fun when_funs cons = if length cons = 1 then ["f"] - else mapn (fn n => K("f"^(string_of_int n))) 1 cons; -fun when_body cons funarg = - let - fun one_fun n (_,[] ) = /\ "dummy" (funarg(1,n)) - | one_fun n (_,args) = let - val l2 = length args; - fun idxs m arg = (if is_lazy arg then (fn t => mk_fup (ID, t)) - else I) (Bound(l2-m)); - in cont_eta_contract - (foldr'' - (fn (a,t) => mk_ssplit (/\# (a,t))) - (args, - fn a=> /\#(a,(list_ccomb(funarg(l2,n),mapn idxs 1 args)))) - ) end; - in (if length cons = 1 andalso length(snd(hd cons)) <= 1 - then mk_strictify else I) - (foldr1 mk_sscase (mapn one_fun 1 cons)) end; - -end; (* struct *) diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/Tools/domain/domain_syntax.ML --- a/src/HOLCF/Tools/domain/domain_syntax.ML Thu Jul 23 21:12:57 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,181 +0,0 @@ -(* Title: HOLCF/Tools/domain/domain_syntax.ML - Author: David von Oheimb - -Syntax generator for domain command. -*) - -signature DOMAIN_SYNTAX = -sig - val calc_syntax: - typ -> - (string * typ list) * - (binding * (bool * binding option * typ) list * mixfix) list -> - (binding * typ * mixfix) list * ast Syntax.trrule list - - val add_syntax: - string -> - ((string * typ list) * - (binding * (bool * binding option * typ) list * mixfix) list) list -> - theory -> theory -end; - - -structure Domain_Syntax :> DOMAIN_SYNTAX = -struct - -open Domain_Library; -infixr 5 -->; infixr 6 ->>; - -fun calc_syntax - (dtypeprod : typ) - ((dname : string, typevars : typ list), - (cons': (binding * (bool * binding option * typ) list * mixfix) list)) - : (binding * typ * mixfix) list * ast Syntax.trrule list = - let - (* ----- constants concerning the isomorphism ------------------------------- *) - - local - fun opt_lazy (lazy,_,t) = if lazy then mk_uT t else t - fun prod (_,args,_) = case args of [] => oneT - | _ => foldr1 mk_sprodT (map opt_lazy args); - fun freetvar s = let val tvar = mk_TFree s in - if tvar mem typevars then freetvar ("t"^s) else tvar end; - fun when_type (_,args,_) = List.foldr (op ->>) (freetvar "t") (map third args); - in - val dtype = Type(dname,typevars); - val dtype2 = foldr1 mk_ssumT (map prod cons'); - val dnam = Long_Name.base_name dname; - fun dbind s = Binding.name (dnam ^ s); - val const_rep = (dbind "_rep" , dtype ->> dtype2, NoSyn); - val const_abs = (dbind "_abs" , dtype2 ->> dtype , NoSyn); - val const_when = (dbind "_when", List.foldr (op ->>) (dtype ->> freetvar "t") (map when_type cons'), NoSyn); - val const_copy = (dbind "_copy", dtypeprod ->> dtype ->> dtype , NoSyn); - end; - - (* ----- constants concerning constructors, discriminators, and selectors --- *) - - local - val escape = let - fun esc (c::cs) = if c mem ["'","_","(",")","/"] then "'"::c::esc cs - else c::esc cs - | esc [] = [] - in implode o esc o Symbol.explode end; - fun dis_name_ con = Binding.name ("is_" ^ strip_esc (Binding.name_of con)); - fun mat_name_ con = Binding.name ("match_" ^ strip_esc (Binding.name_of con)); - fun pat_name_ con = Binding.name (strip_esc (Binding.name_of con) ^ "_pat"); - fun con (name,args,mx) = (name, List.foldr (op ->>) dtype (map third args), mx); - fun dis (con,args,mx) = (dis_name_ con, dtype->>trT, - Mixfix(escape ("is_" ^ Binding.name_of con), [], Syntax.max_pri)); - (* strictly speaking, these constants have one argument, - but the mixfix (without arguments) is introduced only - to generate parse rules for non-alphanumeric names*) - fun freetvar s n = let val tvar = mk_TFree (s ^ string_of_int n) in - if tvar mem typevars then freetvar ("t"^s) n else tvar end; - fun mk_matT (a,bs,c) = a ->> foldr (op ->>) (mk_maybeT c) bs ->> mk_maybeT c; - fun mat (con,args,mx) = (mat_name_ con, - mk_matT(dtype, map third args, freetvar "t" 1), - Mixfix(escape ("match_" ^ Binding.name_of con), [], Syntax.max_pri)); - fun sel1 (_,sel,typ) = Option.map (fn s => (s,dtype ->> typ,NoSyn)) sel; - fun sel (con,args,mx) = List.mapPartial sel1 args; - fun mk_patT (a,b) = a ->> mk_maybeT b; - fun pat_arg_typ n arg = mk_patT (third arg, freetvar "t" n); - fun pat (con,args,mx) = (pat_name_ con, - (mapn pat_arg_typ 1 args) - ---> - mk_patT (dtype, mk_ctupleT (map (freetvar "t") (1 upto length args))), - Mixfix(escape (Binding.name_of con ^ "_pat"), [], Syntax.max_pri)); - - in - val consts_con = map con cons'; - val consts_dis = map dis cons'; - val consts_mat = map mat cons'; - val consts_pat = map pat cons'; - val consts_sel = List.concat(map sel cons'); - end; - - (* ----- constants concerning induction ------------------------------------- *) - - val const_take = (dbind "_take" , HOLogic.natT-->dtype->>dtype, NoSyn); - val const_finite = (dbind "_finite", dtype-->HOLogic.boolT , NoSyn); - - (* ----- case translation --------------------------------------------------- *) - - local open Syntax in - local - fun c_ast con mx = Constant (Syntax.const_name mx (Binding.name_of con)); - fun expvar n = Variable ("e"^(string_of_int n)); - fun argvar n m _ = Variable ("a"^(string_of_int n)^"_"^ - (string_of_int m)); - fun argvars n args = mapn (argvar n) 1 args; - fun app s (l,r) = mk_appl (Constant s) [l,r]; - val cabs = app "_cabs"; - val capp = app "Rep_CFun"; - fun con1 n (con,args,mx) = Library.foldl capp (c_ast con mx, argvars n args); - fun case1 n (con,args,mx) = app "_case1" (con1 n (con,args,mx), expvar n); - fun arg1 n (con,args,_) = List.foldr cabs (expvar n) (argvars n args); - fun when1 n m = if n = m then arg1 n else K (Constant "UU"); - - fun app_var x = mk_appl (Constant "_variable") [x, Variable "rhs"]; - fun app_pat x = mk_appl (Constant "_pat") [x]; - fun args_list [] = Constant "_noargs" - | args_list xs = foldr1 (app "_args") xs; - in - val case_trans = - ParsePrintRule - (app "_case_syntax" (Variable "x", foldr1 (app "_case2") (mapn case1 1 cons')), - capp (Library.foldl capp (Constant (dnam^"_when"), mapn arg1 1 cons'), Variable "x")); - - fun one_abscon_trans n (con,mx,args) = - ParsePrintRule - (cabs (con1 n (con,mx,args), expvar n), - Library.foldl capp (Constant (dnam^"_when"), mapn (when1 n) 1 cons')); - val abscon_trans = mapn one_abscon_trans 1 cons'; - - fun one_case_trans (con,args,mx) = - let - val cname = c_ast con mx; - val pname = Constant (strip_esc (Binding.name_of con) ^ "_pat"); - val ns = 1 upto length args; - val xs = map (fn n => Variable ("x"^(string_of_int n))) ns; - val ps = map (fn n => Variable ("p"^(string_of_int n))) ns; - val vs = map (fn n => Variable ("v"^(string_of_int n))) ns; - in - [ParseRule (app_pat (Library.foldl capp (cname, xs)), - mk_appl pname (map app_pat xs)), - ParseRule (app_var (Library.foldl capp (cname, xs)), - app_var (args_list xs)), - PrintRule (Library.foldl capp (cname, ListPair.map (app "_match") (ps,vs)), - app "_match" (mk_appl pname ps, args_list vs))] - end; - val Case_trans = List.concat (map one_case_trans cons'); - end; - end; - - in ([const_rep, const_abs, const_when, const_copy] @ - consts_con @ consts_dis @ consts_mat @ consts_pat @ consts_sel @ - [const_take, const_finite], - (case_trans::(abscon_trans @ Case_trans))) - end; (* let *) - -(* ----- putting all the syntax stuff together ------------------------------ *) - -fun add_syntax - (comp_dnam : string) - (eqs' : ((string * typ list) * - (binding * (bool * binding option * typ) list * mixfix) list) list) - (thy'' : theory) = - let - val dtypes = map (Type o fst) eqs'; - val boolT = HOLogic.boolT; - val funprod = foldr1 HOLogic.mk_prodT (map (fn tp => tp ->> tp ) dtypes); - val relprod = foldr1 HOLogic.mk_prodT (map (fn tp => tp --> tp --> boolT) dtypes); - val const_copy = (Binding.name (comp_dnam^"_copy"), funprod ->> funprod, NoSyn); - val const_bisim = (Binding.name (comp_dnam^"_bisim"), relprod --> boolT, NoSyn); - val ctt : ((binding * typ * mixfix) list * ast Syntax.trrule list) list = map (calc_syntax funprod) eqs'; - in thy'' |> ContConsts.add_consts_i (List.concat (map fst ctt) @ - (if length eqs'>1 then [const_copy] else[])@ - [const_bisim]) - |> Sign.add_trrules_i (List.concat(map snd ctt)) - end; (* let *) - -end; (* struct *) diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/Tools/domain/domain_theorems.ML --- a/src/HOLCF/Tools/domain/domain_theorems.ML Thu Jul 23 21:12:57 2009 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1054 +0,0 @@ -(* Title: HOLCF/Tools/domain/domain_theorems.ML - ID: $Id$ - Author: David von Oheimb - New proofs/tactics by Brian Huffman - -Proof generator for domain command. -*) - -val HOLCF_ss = @{simpset}; - -signature DOMAIN_THEOREMS = -sig - val theorems: Domain_Library.eq * Domain_Library.eq list -> theory -> thm list * theory; - val comp_theorems: bstring * Domain_Library.eq list -> theory -> thm list * theory; - val quiet_mode: bool ref; - val trace_domain: bool ref; -end; - -structure Domain_Theorems :> DOMAIN_THEOREMS = -struct - -val quiet_mode = ref false; -val trace_domain = ref false; - -fun message s = if !quiet_mode then () else writeln s; -fun trace s = if !trace_domain then tracing s else (); - -local - -val adm_impl_admw = @{thm adm_impl_admw}; -val adm_all = @{thm adm_all}; -val adm_conj = @{thm adm_conj}; -val adm_subst = @{thm adm_subst}; -val antisym_less_inverse = @{thm below_antisym_inverse}; -val beta_cfun = @{thm beta_cfun}; -val cfun_arg_cong = @{thm cfun_arg_cong}; -val ch2ch_Rep_CFunL = @{thm ch2ch_Rep_CFunL}; -val ch2ch_Rep_CFunR = @{thm ch2ch_Rep_CFunR}; -val chain_iterate = @{thm chain_iterate}; -val compact_ONE = @{thm compact_ONE}; -val compact_sinl = @{thm compact_sinl}; -val compact_sinr = @{thm compact_sinr}; -val compact_spair = @{thm compact_spair}; -val compact_up = @{thm compact_up}; -val contlub_cfun_arg = @{thm contlub_cfun_arg}; -val contlub_cfun_fun = @{thm contlub_cfun_fun}; -val fix_def2 = @{thm fix_def2}; -val injection_eq = @{thm injection_eq}; -val injection_less = @{thm injection_below}; -val lub_equal = @{thm lub_equal}; -val monofun_cfun_arg = @{thm monofun_cfun_arg}; -val retraction_strict = @{thm retraction_strict}; -val spair_eq = @{thm spair_eq}; -val spair_less = @{thm spair_below}; -val sscase1 = @{thm sscase1}; -val ssplit1 = @{thm ssplit1}; -val strictify1 = @{thm strictify1}; -val wfix_ind = @{thm wfix_ind}; - -val iso_intro = @{thm iso.intro}; -val iso_abs_iso = @{thm iso.abs_iso}; -val iso_rep_iso = @{thm iso.rep_iso}; -val iso_abs_strict = @{thm iso.abs_strict}; -val iso_rep_strict = @{thm iso.rep_strict}; -val iso_abs_defin' = @{thm iso.abs_defin'}; -val iso_rep_defin' = @{thm iso.rep_defin'}; -val iso_abs_defined = @{thm iso.abs_defined}; -val iso_rep_defined = @{thm iso.rep_defined}; -val iso_compact_abs = @{thm iso.compact_abs}; -val iso_compact_rep = @{thm iso.compact_rep}; -val iso_iso_swap = @{thm iso.iso_swap}; - -val exh_start = @{thm exh_start}; -val ex_defined_iffs = @{thms ex_defined_iffs}; -val exh_casedist0 = @{thm exh_casedist0}; -val exh_casedists = @{thms exh_casedists}; - -open Domain_Library; -infixr 0 ===>; -infixr 0 ==>; -infix 0 == ; -infix 1 ===; -infix 1 ~= ; -infix 1 <<; -infix 1 ~<<; -infix 9 ` ; -infix 9 `% ; -infix 9 `%%; -infixr 9 oo; - -(* ----- general proof facilities ------------------------------------------- *) - -fun legacy_infer_term thy t = - let val ctxt = ProofContext.set_mode ProofContext.mode_schematic (ProofContext.init thy) - in singleton (Syntax.check_terms ctxt) (Sign.intern_term thy t) end; - -fun pg'' thy defs t tacs = - let - val t' = legacy_infer_term thy t; - val asms = Logic.strip_imp_prems t'; - val prop = Logic.strip_imp_concl t'; - fun tac {prems, context} = - rewrite_goals_tac defs THEN - EVERY (tacs {prems = map (rewrite_rule defs) prems, context = context}) - in Goal.prove_global thy [] asms prop tac end; - -fun pg' thy defs t tacsf = - let - fun tacs {prems, context} = - if null prems then tacsf context - else cut_facts_tac prems 1 :: tacsf context; - in pg'' thy defs t tacs end; - -fun case_UU_tac ctxt rews i v = - InductTacs.case_tac ctxt (v^"=UU") i THEN - asm_simp_tac (HOLCF_ss addsimps rews) i; - -val chain_tac = - REPEAT_DETERM o resolve_tac - [chain_iterate, ch2ch_Rep_CFunR, ch2ch_Rep_CFunL]; - -(* ----- general proofs ----------------------------------------------------- *) - -val all2E = @{lemma "!x y . P x y ==> (P x y ==> R) ==> R" by simp} - -val dist_eqI = @{lemma "!!x::'a::po. ~ x << y ==> x ~= y" by (blast dest!: below_antisym_inverse)} - -in - -fun theorems (((dname, _), cons) : eq, eqs : eq list) thy = -let - -val _ = message ("Proving isomorphism properties of domain "^dname^" ..."); -val pg = pg' thy; - -(* ----- getting the axioms and definitions --------------------------------- *) - -local - fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s); -in - val ax_abs_iso = ga "abs_iso" dname; - val ax_rep_iso = ga "rep_iso" dname; - val ax_when_def = ga "when_def" dname; - fun get_def mk_name (con,_) = ga (mk_name con^"_def") dname; - val axs_con_def = map (get_def extern_name) cons; - val axs_dis_def = map (get_def dis_name) cons; - val axs_mat_def = map (get_def mat_name) cons; - val axs_pat_def = map (get_def pat_name) cons; - val axs_sel_def = - let - fun def_of_sel sel = ga (sel^"_def") dname; - fun def_of_arg arg = Option.map def_of_sel (sel_of arg); - fun defs_of_con (_, args) = List.mapPartial def_of_arg args; - in - maps defs_of_con cons - end; - val ax_copy_def = ga "copy_def" dname; -end; (* local *) - -(* ----- theorems concerning the isomorphism -------------------------------- *) - -val dc_abs = %%:(dname^"_abs"); -val dc_rep = %%:(dname^"_rep"); -val dc_copy = %%:(dname^"_copy"); -val x_name = "x"; - -val iso_locale = iso_intro OF [ax_abs_iso, ax_rep_iso]; -val abs_strict = ax_rep_iso RS (allI RS retraction_strict); -val rep_strict = ax_abs_iso RS (allI RS retraction_strict); -val abs_defin' = iso_locale RS iso_abs_defin'; -val rep_defin' = iso_locale RS iso_rep_defin'; -val iso_rews = map standard [ax_abs_iso,ax_rep_iso,abs_strict,rep_strict]; - -(* ----- generating beta reduction rules from definitions-------------------- *) - -val _ = trace " Proving beta reduction rules..."; - -local - fun arglist (Const _ $ Abs (s, _, t)) = - let - val (vars,body) = arglist t; - in (s :: vars, body) end - | arglist t = ([], t); - fun bind_fun vars t = Library.foldr mk_All (vars, t); - fun bound_vars 0 = [] - | bound_vars i = Bound (i-1) :: bound_vars (i - 1); -in - fun appl_of_def def = - let - val (_ $ con $ lam) = concl_of def; - val (vars, rhs) = arglist lam; - val lhs = list_ccomb (con, bound_vars (length vars)); - val appl = bind_fun vars (lhs == rhs); - val cs = ContProc.cont_thms lam; - val betas = map (fn c => mk_meta_eq (c RS beta_cfun)) cs; - in pg (def::betas) appl (K [rtac reflexive_thm 1]) end; -end; - -val _ = trace "Proving when_appl..."; -val when_appl = appl_of_def ax_when_def; -val _ = trace "Proving con_appls..."; -val con_appls = map appl_of_def axs_con_def; - -local - fun arg2typ n arg = - let val t = TVar (("'a", n), pcpoS) - in (n + 1, if is_lazy arg then mk_uT t else t) end; - - fun args2typ n [] = (n, oneT) - | args2typ n [arg] = arg2typ n arg - | args2typ n (arg::args) = - let - val (n1, t1) = arg2typ n arg; - val (n2, t2) = args2typ n1 args - in (n2, mk_sprodT (t1, t2)) end; - - fun cons2typ n [] = (n,oneT) - | cons2typ n [con] = args2typ n (snd con) - | cons2typ n (con::cons) = - let - val (n1, t1) = args2typ n (snd con); - val (n2, t2) = cons2typ n1 cons - in (n2, mk_ssumT (t1, t2)) end; -in - fun cons2ctyp cons = ctyp_of thy (snd (cons2typ 1 cons)); -end; - -local - val iso_swap = iso_locale RS iso_iso_swap; - fun one_con (con, args) = - let - val vns = map vname args; - val eqn = %:x_name === con_app2 con %: vns; - val conj = foldr1 mk_conj (eqn :: map (defined o %:) (nonlazy args)); - in Library.foldr mk_ex (vns, conj) end; - - val conj_assoc = @{thm conj_assoc}; - val exh = foldr1 mk_disj ((%:x_name === UU) :: map one_con cons); - val thm1 = instantiate' [SOME (cons2ctyp cons)] [] exh_start; - val thm2 = rewrite_rule (map mk_meta_eq ex_defined_iffs) thm1; - val thm3 = rewrite_rule [mk_meta_eq @{thm conj_assoc}] thm2; - - (* first 3 rules replace "x = UU \/ P" with "rep$x = UU \/ P" *) - val tacs = [ - rtac disjE 1, - etac (rep_defin' RS disjI1) 2, - etac disjI2 2, - rewrite_goals_tac [mk_meta_eq iso_swap], - rtac thm3 1]; -in - val _ = trace " Proving exhaust..."; - val exhaust = pg con_appls (mk_trp exh) (K tacs); - val _ = trace " Proving casedist..."; - val casedist = - standard (rewrite_rule exh_casedists (exhaust RS exh_casedist0)); -end; - -local - fun bind_fun t = Library.foldr mk_All (when_funs cons, t); - fun bound_fun i _ = Bound (length cons - i); - val when_app = list_ccomb (%%:(dname^"_when"), mapn bound_fun 1 cons); -in - val _ = trace " Proving when_strict..."; - val when_strict = - let - val axs = [when_appl, mk_meta_eq rep_strict]; - val goal = bind_fun (mk_trp (strict when_app)); - val tacs = [resolve_tac [sscase1, ssplit1, strictify1] 1]; - in pg axs goal (K tacs) end; - - val _ = trace " Proving when_apps..."; - val when_apps = - let - fun one_when n (con,args) = - let - val axs = when_appl :: con_appls; - val goal = bind_fun (lift_defined %: (nonlazy args, - mk_trp (when_app`(con_app con args) === - list_ccomb (bound_fun n 0, map %# args)))); - val tacs = [asm_simp_tac (HOLCF_ss addsimps [ax_abs_iso]) 1]; - in pg axs goal (K tacs) end; - in mapn one_when 1 cons end; -end; -val when_rews = when_strict :: when_apps; - -(* ----- theorems concerning the constructors, discriminators and selectors - *) - -local - fun dis_strict (con, _) = - let - val goal = mk_trp (strict (%%:(dis_name con))); - in pg axs_dis_def goal (K [rtac when_strict 1]) end; - - fun dis_app c (con, args) = - let - val lhs = %%:(dis_name c) ` con_app con args; - val rhs = if con = c then TT else FF; - val goal = lift_defined %: (nonlazy args, mk_trp (lhs === rhs)); - val tacs = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; - in pg axs_dis_def goal (K tacs) end; - - val _ = trace " Proving dis_apps..."; - val dis_apps = maps (fn (c,_) => map (dis_app c) cons) cons; - - fun dis_defin (con, args) = - let - val goal = defined (%:x_name) ==> defined (%%:(dis_name con) `% x_name); - val tacs = - [rtac casedist 1, - contr_tac 1, - DETERM_UNTIL_SOLVED (CHANGED - (asm_simp_tac (HOLCF_ss addsimps dis_apps) 1))]; - in pg [] goal (K tacs) end; - - val _ = trace " Proving dis_stricts..."; - val dis_stricts = map dis_strict cons; - val _ = trace " Proving dis_defins..."; - val dis_defins = map dis_defin cons; -in - val dis_rews = dis_stricts @ dis_defins @ dis_apps; -end; - -local - fun mat_strict (con, _) = - let - val goal = mk_trp (%%:(mat_name con) ` UU ` %:"rhs" === UU); - val tacs = [asm_simp_tac (HOLCF_ss addsimps [when_strict]) 1]; - in pg axs_mat_def goal (K tacs) end; - - val _ = trace " Proving mat_stricts..."; - val mat_stricts = map mat_strict cons; - - fun one_mat c (con, args) = - let - val lhs = %%:(mat_name c) ` con_app con args ` %:"rhs"; - val rhs = - if con = c - then list_ccomb (%:"rhs", map %# args) - else mk_fail; - val goal = lift_defined %: (nonlazy args, mk_trp (lhs === rhs)); - val tacs = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; - in pg axs_mat_def goal (K tacs) end; - - val _ = trace " Proving mat_apps..."; - val mat_apps = - maps (fn (c,_) => map (one_mat c) cons) cons; -in - val mat_rews = mat_stricts @ mat_apps; -end; - -local - fun ps args = mapn (fn n => fn _ => %:("pat" ^ string_of_int n)) 1 args; - - fun pat_lhs (con,args) = mk_branch (list_comb (%%:(pat_name con), ps args)); - - fun pat_rhs (con,[]) = mk_return ((%:"rhs") ` HOLogic.unit) - | pat_rhs (con,args) = - (mk_branch (mk_ctuple_pat (ps args))) - `(%:"rhs")`(mk_ctuple (map %# args)); - - fun pat_strict c = - let - val axs = @{thm branch_def} :: axs_pat_def; - val goal = mk_trp (strict (pat_lhs c ` (%:"rhs"))); - val tacs = [simp_tac (HOLCF_ss addsimps [when_strict]) 1]; - in pg axs goal (K tacs) end; - - fun pat_app c (con, args) = - let - val axs = @{thm branch_def} :: axs_pat_def; - val lhs = (pat_lhs c)`(%:"rhs")`(con_app con args); - val rhs = if con = fst c then pat_rhs c else mk_fail; - val goal = lift_defined %: (nonlazy args, mk_trp (lhs === rhs)); - val tacs = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; - in pg axs goal (K tacs) end; - - val _ = trace " Proving pat_stricts..."; - val pat_stricts = map pat_strict cons; - val _ = trace " Proving pat_apps..."; - val pat_apps = maps (fn c => map (pat_app c) cons) cons; -in - val pat_rews = pat_stricts @ pat_apps; -end; - -local - fun con_strict (con, args) = - let - val rules = abs_strict :: @{thms con_strict_rules}; - fun one_strict vn = - let - fun f arg = if vname arg = vn then UU else %# arg; - val goal = mk_trp (con_app2 con f args === UU); - val tacs = [simp_tac (HOL_basic_ss addsimps rules) 1]; - in pg con_appls goal (K tacs) end; - in map one_strict (nonlazy args) end; - - fun con_defin (con, args) = - let - fun iff_disj (t, []) = HOLogic.mk_not t - | iff_disj (t, ts) = t === foldr1 HOLogic.mk_disj ts; - val lhs = con_app con args === UU; - val rhss = map (fn x => %:x === UU) (nonlazy args); - val goal = mk_trp (iff_disj (lhs, rhss)); - val rule1 = iso_locale RS @{thm iso.abs_defined_iff}; - val rules = rule1 :: @{thms con_defined_iff_rules}; - val tacs = [simp_tac (HOL_ss addsimps rules) 1]; - in pg con_appls goal (K tacs) end; -in - val _ = trace " Proving con_stricts..."; - val con_stricts = maps con_strict cons; - val _ = trace " Proving con_defins..."; - val con_defins = map con_defin cons; - val con_rews = con_stricts @ con_defins; -end; - -local - val rules = - [compact_sinl, compact_sinr, compact_spair, compact_up, compact_ONE]; - fun con_compact (con, args) = - let - val concl = mk_trp (mk_compact (con_app con args)); - val goal = lift (fn x => mk_compact (%#x)) (args, concl); - val tacs = [ - rtac (iso_locale RS iso_compact_abs) 1, - REPEAT (resolve_tac rules 1 ORELSE atac 1)]; - in pg con_appls goal (K tacs) end; -in - val _ = trace " Proving con_compacts..."; - val con_compacts = map con_compact cons; -end; - -local - fun one_sel sel = - pg axs_sel_def (mk_trp (strict (%%:sel))) - (K [simp_tac (HOLCF_ss addsimps when_rews) 1]); - - fun sel_strict (_, args) = - List.mapPartial (Option.map one_sel o sel_of) args; -in - val _ = trace " Proving sel_stricts..."; - val sel_stricts = maps sel_strict cons; -end; - -local - fun sel_app_same c n sel (con, args) = - let - val nlas = nonlazy args; - val vns = map vname args; - val vnn = List.nth (vns, n); - val nlas' = List.filter (fn v => v <> vnn) nlas; - val lhs = (%%:sel)`(con_app con args); - val goal = lift_defined %: (nlas', mk_trp (lhs === %:vnn)); - fun tacs1 ctxt = - if vnn mem nlas - then [case_UU_tac ctxt (when_rews @ con_stricts) 1 vnn] - else []; - val tacs2 = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; - in pg axs_sel_def goal (fn ctxt => (tacs1 ctxt @ tacs2)) end; - - fun sel_app_diff c n sel (con, args) = - let - val nlas = nonlazy args; - val goal = mk_trp (%%:sel ` con_app con args === UU); - fun tacs1 ctxt = map (case_UU_tac ctxt (when_rews @ con_stricts) 1) nlas; - val tacs2 = [asm_simp_tac (HOLCF_ss addsimps when_rews) 1]; - in pg axs_sel_def goal (fn ctxt => (tacs1 ctxt @ tacs2)) end; - - fun sel_app c n sel (con, args) = - if con = c - then sel_app_same c n sel (con, args) - else sel_app_diff c n sel (con, args); - - fun one_sel c n sel = map (sel_app c n sel) cons; - fun one_sel' c n arg = Option.map (one_sel c n) (sel_of arg); - fun one_con (c, args) = - flat (map_filter I (mapn (one_sel' c) 0 args)); -in - val _ = trace " Proving sel_apps..."; - val sel_apps = maps one_con cons; -end; - -local - fun sel_defin sel = - let - val goal = defined (%:x_name) ==> defined (%%:sel`%x_name); - val tacs = [ - rtac casedist 1, - contr_tac 1, - DETERM_UNTIL_SOLVED (CHANGED - (asm_simp_tac (HOLCF_ss addsimps sel_apps) 1))]; - in pg [] goal (K tacs) end; -in - val _ = trace " Proving sel_defins..."; - val sel_defins = - if length cons = 1 - then List.mapPartial (fn arg => Option.map sel_defin (sel_of arg)) - (filter_out is_lazy (snd (hd cons))) - else []; -end; - -val sel_rews = sel_stricts @ sel_defins @ sel_apps; - -val _ = trace " Proving dist_les..."; -val distincts_le = - let - fun dist (con1, args1) (con2, args2) = - let - val goal = lift_defined %: (nonlazy args1, - mk_trp (con_app con1 args1 ~<< con_app con2 args2)); - fun tacs ctxt = [ - rtac @{thm rev_contrapos} 1, - eres_inst_tac ctxt [(("f", 0), dis_name con1)] monofun_cfun_arg 1] - @ map (case_UU_tac ctxt (con_stricts @ dis_rews) 1) (nonlazy args2) - @ [asm_simp_tac (HOLCF_ss addsimps dis_rews) 1]; - in pg [] goal tacs end; - - fun distinct (con1, args1) (con2, args2) = - let - val arg1 = (con1, args1); - val arg2 = - (con2, ListPair.map (fn (arg,vn) => upd_vname (K vn) arg) - (args2, Name.variant_list (map vname args1) (map vname args2))); - in [dist arg1 arg2, dist arg2 arg1] end; - fun distincts [] = [] - | distincts (c::cs) = (map (distinct c) cs) :: distincts cs; - in distincts cons end; -val dist_les = flat (flat distincts_le); - -val _ = trace " Proving dist_eqs..."; -val dist_eqs = - let - fun distinct (_,args1) ((_,args2), leqs) = - let - val (le1,le2) = (hd leqs, hd(tl leqs)); - val (eq1,eq2) = (le1 RS dist_eqI, le2 RS dist_eqI) - in - if nonlazy args1 = [] then [eq1, eq1 RS not_sym] else - if nonlazy args2 = [] then [eq2, eq2 RS not_sym] else - [eq1, eq2] - end; - fun distincts [] = [] - | distincts ((c,leqs)::cs) = - flat - (ListPair.map (distinct c) ((map #1 cs),leqs)) @ - distincts cs; - in map standard (distincts (cons ~~ distincts_le)) end; - -local - fun pgterm rel con args = - let - fun append s = upd_vname (fn v => v^s); - val (largs, rargs) = (args, map (append "'") args); - val concl = - foldr1 mk_conj (ListPair.map rel (map %# largs, map %# rargs)); - val prem = rel (con_app con largs, con_app con rargs); - val sargs = case largs of [_] => [] | _ => nonlazy args; - val prop = lift_defined %: (sargs, mk_trp (prem === concl)); - in pg con_appls prop end; - val cons' = List.filter (fn (_,args) => args<>[]) cons; -in - val _ = trace " Proving inverts..."; - val inverts = - let - val abs_less = ax_abs_iso RS (allI RS injection_less); - val tacs = - [asm_full_simp_tac (HOLCF_ss addsimps [abs_less, spair_less]) 1]; - in map (fn (con, args) => pgterm (op <<) con args (K tacs)) cons' end; - - val _ = trace " Proving injects..."; - val injects = - let - val abs_eq = ax_abs_iso RS (allI RS injection_eq); - val tacs = [asm_full_simp_tac (HOLCF_ss addsimps [abs_eq, spair_eq]) 1]; - in map (fn (con, args) => pgterm (op ===) con args (K tacs)) cons' end; -end; - -(* ----- theorems concerning one induction step ----------------------------- *) - -val copy_strict = - let - val _ = trace " Proving copy_strict..."; - val goal = mk_trp (strict (dc_copy `% "f")); - val rules = [abs_strict, rep_strict] @ @{thms domain_fun_stricts}; - val tacs = [asm_simp_tac (HOLCF_ss addsimps rules) 1]; - in pg [ax_copy_def] goal (K tacs) end; - -local - fun copy_app (con, args) = - let - val lhs = dc_copy`%"f"`(con_app con args); - fun one_rhs arg = - if DatatypeAux.is_rec_type (dtyp_of arg) - then Domain_Axioms.copy_of_dtyp (cproj (%:"f") eqs) (dtyp_of arg) ` (%# arg) - else (%# arg); - val rhs = con_app2 con one_rhs args; - val goal = lift_defined %: (nonlazy_rec args, mk_trp (lhs === rhs)); - val args' = List.filter (fn a => not (is_rec a orelse is_lazy a)) args; - val stricts = abs_strict :: rep_strict :: @{thms domain_fun_stricts}; - fun tacs1 ctxt = map (case_UU_tac ctxt stricts 1 o vname) args'; - val rules = [ax_abs_iso] @ @{thms domain_fun_simps}; - val tacs2 = [asm_simp_tac (HOLCF_ss addsimps rules) 1]; - in pg (ax_copy_def::con_appls) goal (fn ctxt => (tacs1 ctxt @ tacs2)) end; -in - val _ = trace " Proving copy_apps..."; - val copy_apps = map copy_app cons; -end; - -local - fun one_strict (con, args) = - let - val goal = mk_trp (dc_copy`UU`(con_app con args) === UU); - val rews = copy_strict :: copy_apps @ con_rews; - fun tacs ctxt = map (case_UU_tac ctxt rews 1) (nonlazy args) @ - [asm_simp_tac (HOLCF_ss addsimps rews) 1]; - in pg [] goal tacs end; - - fun has_nonlazy_rec (_, args) = exists is_nonlazy_rec args; -in - val _ = trace " Proving copy_stricts..."; - val copy_stricts = map one_strict (List.filter has_nonlazy_rec cons); -end; - -val copy_rews = copy_strict :: copy_apps @ copy_stricts; - -in - thy - |> Sign.add_path (Long_Name.base_name dname) - |> snd o PureThy.add_thmss [ - ((Binding.name "iso_rews" , iso_rews ), [Simplifier.simp_add]), - ((Binding.name "exhaust" , [exhaust] ), []), - ((Binding.name "casedist" , [casedist] ), [Induct.cases_type dname]), - ((Binding.name "when_rews" , when_rews ), [Simplifier.simp_add]), - ((Binding.name "compacts" , con_compacts), [Simplifier.simp_add]), - ((Binding.name "con_rews" , con_rews ), [Simplifier.simp_add]), - ((Binding.name "sel_rews" , sel_rews ), [Simplifier.simp_add]), - ((Binding.name "dis_rews" , dis_rews ), [Simplifier.simp_add]), - ((Binding.name "pat_rews" , pat_rews ), [Simplifier.simp_add]), - ((Binding.name "dist_les" , dist_les ), [Simplifier.simp_add]), - ((Binding.name "dist_eqs" , dist_eqs ), [Simplifier.simp_add]), - ((Binding.name "inverts" , inverts ), [Simplifier.simp_add]), - ((Binding.name "injects" , injects ), [Simplifier.simp_add]), - ((Binding.name "copy_rews" , copy_rews ), [Simplifier.simp_add]), - ((Binding.name "match_rews", mat_rews ), [Simplifier.simp_add])] - |> Sign.parent_path - |> pair (iso_rews @ when_rews @ con_rews @ sel_rews @ dis_rews @ - pat_rews @ dist_les @ dist_eqs @ copy_rews) -end; (* let *) - -fun comp_theorems (comp_dnam, eqs: eq list) thy = -let -val global_ctxt = ProofContext.init thy; - -val dnames = map (fst o fst) eqs; -val conss = map snd eqs; -val comp_dname = Sign.full_bname thy comp_dnam; - -val _ = message ("Proving induction properties of domain "^comp_dname^" ..."); -val pg = pg' thy; - -(* ----- getting the composite axiom and definitions ------------------------ *) - -local - fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s); -in - val axs_reach = map (ga "reach" ) dnames; - val axs_take_def = map (ga "take_def" ) dnames; - val axs_finite_def = map (ga "finite_def") dnames; - val ax_copy2_def = ga "copy_def" comp_dnam; - val ax_bisim_def = ga "bisim_def" comp_dnam; -end; - -local - fun gt s dn = PureThy.get_thm thy (dn ^ "." ^ s); - fun gts s dn = PureThy.get_thms thy (dn ^ "." ^ s); -in - val cases = map (gt "casedist" ) dnames; - val con_rews = maps (gts "con_rews" ) dnames; - val copy_rews = maps (gts "copy_rews") dnames; -end; - -fun dc_take dn = %%:(dn^"_take"); -val x_name = idx_name dnames "x"; -val P_name = idx_name dnames "P"; -val n_eqs = length eqs; - -(* ----- theorems concerning finite approximation and finite induction ------ *) - -local - val iterate_Cprod_ss = simpset_of @{theory Fix}; - val copy_con_rews = copy_rews @ con_rews; - val copy_take_defs = - (if n_eqs = 1 then [] else [ax_copy2_def]) @ axs_take_def; - val _ = trace " Proving take_stricts..."; - val take_stricts = - let - fun one_eq ((dn, args), _) = strict (dc_take dn $ %:"n"); - val goal = mk_trp (foldr1 mk_conj (map one_eq eqs)); - fun tacs ctxt = [ - InductTacs.induct_tac ctxt [[SOME "n"]] 1, - simp_tac iterate_Cprod_ss 1, - asm_simp_tac (iterate_Cprod_ss addsimps copy_rews) 1]; - in pg copy_take_defs goal tacs end; - - val take_stricts' = rewrite_rule copy_take_defs take_stricts; - fun take_0 n dn = - let - val goal = mk_trp ((dc_take dn $ %%:"HOL.zero") `% x_name n === UU); - in pg axs_take_def goal (K [simp_tac iterate_Cprod_ss 1]) end; - val take_0s = mapn take_0 1 dnames; - fun c_UU_tac ctxt = case_UU_tac ctxt (take_stricts'::copy_con_rews) 1; - val _ = trace " Proving take_apps..."; - val take_apps = - let - fun mk_eqn dn (con, args) = - let - fun mk_take n = dc_take (List.nth (dnames, n)) $ %:"n"; - fun one_rhs arg = - if DatatypeAux.is_rec_type (dtyp_of arg) - then Domain_Axioms.copy_of_dtyp mk_take (dtyp_of arg) ` (%# arg) - else (%# arg); - val lhs = (dc_take dn $ (%%:"Suc" $ %:"n"))`(con_app con args); - val rhs = con_app2 con one_rhs args; - in Library.foldr mk_all (map vname args, lhs === rhs) end; - fun mk_eqns ((dn, _), cons) = map (mk_eqn dn) cons; - val goal = mk_trp (foldr1 mk_conj (maps mk_eqns eqs)); - val simps = List.filter (has_fewer_prems 1) copy_rews; - fun con_tac ctxt (con, args) = - if nonlazy_rec args = [] - then all_tac - else EVERY (map (c_UU_tac ctxt) (nonlazy_rec args)) THEN - asm_full_simp_tac (HOLCF_ss addsimps copy_rews) 1; - fun eq_tacs ctxt ((dn, _), cons) = map (con_tac ctxt) cons; - fun tacs ctxt = - simp_tac iterate_Cprod_ss 1 :: - InductTacs.induct_tac ctxt [[SOME "n"]] 1 :: - simp_tac (iterate_Cprod_ss addsimps copy_con_rews) 1 :: - asm_full_simp_tac (HOLCF_ss addsimps simps) 1 :: - TRY (safe_tac HOL_cs) :: - maps (eq_tacs ctxt) eqs; - in pg copy_take_defs goal tacs end; -in - val take_rews = map standard - (atomize global_ctxt take_stricts @ take_0s @ atomize global_ctxt take_apps); -end; (* local *) - -local - fun one_con p (con,args) = - let - fun ind_hyp arg = %:(P_name (1 + rec_of arg)) $ bound_arg args arg; - val t1 = mk_trp (%:p $ con_app2 con (bound_arg args) args); - val t2 = lift ind_hyp (List.filter is_rec args, t1); - val t3 = lift_defined (bound_arg (map vname args)) (nonlazy args, t2); - in Library.foldr mk_All (map vname args, t3) end; - - fun one_eq ((p, cons), concl) = - mk_trp (%:p $ UU) ===> Logic.list_implies (map (one_con p) cons, concl); - - fun ind_term concf = Library.foldr one_eq - (mapn (fn n => fn x => (P_name n, x)) 1 conss, - mk_trp (foldr1 mk_conj (mapn concf 1 dnames))); - val take_ss = HOL_ss addsimps take_rews; - fun quant_tac ctxt i = EVERY - (mapn (fn n => fn _ => res_inst_tac ctxt [(("x", 0), x_name n)] spec i) 1 dnames); - - fun ind_prems_tac prems = EVERY - (maps (fn cons => - (resolve_tac prems 1 :: - maps (fn (_,args) => - resolve_tac prems 1 :: - map (K(atac 1)) (nonlazy args) @ - map (K(atac 1)) (List.filter is_rec args)) - cons)) - conss); - local - (* check whether every/exists constructor of the n-th part of the equation: - it has a possibly indirectly recursive argument that isn't/is possibly - indirectly lazy *) - fun rec_to quant nfn rfn ns lazy_rec (n,cons) = quant (exists (fn arg => - is_rec arg andalso not(rec_of arg mem ns) andalso - ((rec_of arg = n andalso nfn(lazy_rec orelse is_lazy arg)) orelse - rec_of arg <> n andalso rec_to quant nfn rfn (rec_of arg::ns) - (lazy_rec orelse is_lazy arg) (n, (List.nth(conss,rec_of arg)))) - ) o snd) cons; - fun all_rec_to ns = rec_to forall not all_rec_to ns; - fun warn (n,cons) = - if all_rec_to [] false (n,cons) - then (warning ("domain "^List.nth(dnames,n)^" is empty!"); true) - else false; - fun lazy_rec_to ns = rec_to exists I lazy_rec_to ns; - - in - val n__eqs = mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs; - val is_emptys = map warn n__eqs; - val is_finite = forall (not o lazy_rec_to [] false) n__eqs; - end; -in (* local *) - val _ = trace " Proving finite_ind..."; - val finite_ind = - let - fun concf n dn = %:(P_name n) $ (dc_take dn $ %:"n" `%(x_name n)); - val goal = ind_term concf; - - fun tacf {prems, context} = - let - val tacs1 = [ - quant_tac context 1, - simp_tac HOL_ss 1, - InductTacs.induct_tac context [[SOME "n"]] 1, - simp_tac (take_ss addsimps prems) 1, - TRY (safe_tac HOL_cs)]; - fun arg_tac arg = - case_UU_tac context (prems @ con_rews) 1 - (List.nth (dnames, rec_of arg) ^ "_take n$" ^ vname arg); - fun con_tacs (con, args) = - asm_simp_tac take_ss 1 :: - map arg_tac (List.filter is_nonlazy_rec args) @ - [resolve_tac prems 1] @ - map (K (atac 1)) (nonlazy args) @ - map (K (etac spec 1)) (List.filter is_rec args); - fun cases_tacs (cons, cases) = - res_inst_tac context [(("x", 0), "x")] cases 1 :: - asm_simp_tac (take_ss addsimps prems) 1 :: - maps con_tacs cons; - in - tacs1 @ maps cases_tacs (conss ~~ cases) - end; - in pg'' thy [] goal tacf - handle ERROR _ => (warning "Proof of finite_ind failed."; TrueI) - end; - - val _ = trace " Proving take_lemmas..."; - val take_lemmas = - let - fun take_lemma n (dn, ax_reach) = - let - val lhs = dc_take dn $ Bound 0 `%(x_name n); - val rhs = dc_take dn $ Bound 0 `%(x_name n^"'"); - val concl = mk_trp (%:(x_name n) === %:(x_name n^"'")); - val goal = mk_All ("n", mk_trp (lhs === rhs)) ===> concl; - fun tacf {prems, context} = [ - res_inst_tac context [(("t", 0), x_name n )] (ax_reach RS subst) 1, - res_inst_tac context [(("t", 0), x_name n^"'")] (ax_reach RS subst) 1, - stac fix_def2 1, - REPEAT (CHANGED - (rtac (contlub_cfun_arg RS ssubst) 1 THEN chain_tac 1)), - stac contlub_cfun_fun 1, - stac contlub_cfun_fun 2, - rtac lub_equal 3, - chain_tac 1, - rtac allI 1, - resolve_tac prems 1]; - in pg'' thy axs_take_def goal tacf end; - in mapn take_lemma 1 (dnames ~~ axs_reach) end; - -(* ----- theorems concerning finiteness and induction ----------------------- *) - - val _ = trace " Proving finites, ind..."; - val (finites, ind) = - ( - if is_finite - then (* finite case *) - let - fun take_enough dn = mk_ex ("n",dc_take dn $ Bound 0 ` %:"x" === %:"x"); - fun dname_lemma dn = - let - val prem1 = mk_trp (defined (%:"x")); - val disj1 = mk_all ("n", dc_take dn $ Bound 0 ` %:"x" === UU); - val prem2 = mk_trp (mk_disj (disj1, take_enough dn)); - val concl = mk_trp (take_enough dn); - val goal = prem1 ===> prem2 ===> concl; - val tacs = [ - etac disjE 1, - etac notE 1, - resolve_tac take_lemmas 1, - asm_simp_tac take_ss 1, - atac 1]; - in pg [] goal (K tacs) end; - val _ = trace " Proving finite_lemmas1a"; - val finite_lemmas1a = map dname_lemma dnames; - - val _ = trace " Proving finite_lemma1b"; - val finite_lemma1b = - let - fun mk_eqn n ((dn, args), _) = - let - val disj1 = dc_take dn $ Bound 1 ` Bound 0 === UU; - val disj2 = dc_take dn $ Bound 1 ` Bound 0 === Bound 0; - in - mk_constrainall - (x_name n, Type (dn,args), mk_disj (disj1, disj2)) - end; - val goal = - mk_trp (mk_all ("n", foldr1 mk_conj (mapn mk_eqn 1 eqs))); - fun arg_tacs ctxt vn = [ - eres_inst_tac ctxt [(("x", 0), vn)] all_dupE 1, - etac disjE 1, - asm_simp_tac (HOL_ss addsimps con_rews) 1, - asm_simp_tac take_ss 1]; - fun con_tacs ctxt (con, args) = - asm_simp_tac take_ss 1 :: - maps (arg_tacs ctxt) (nonlazy_rec args); - fun foo_tacs ctxt n (cons, cases) = - simp_tac take_ss 1 :: - rtac allI 1 :: - res_inst_tac ctxt [(("x", 0), x_name n)] cases 1 :: - asm_simp_tac take_ss 1 :: - maps (con_tacs ctxt) cons; - fun tacs ctxt = - rtac allI 1 :: - InductTacs.induct_tac ctxt [[SOME "n"]] 1 :: - simp_tac take_ss 1 :: - TRY (safe_tac (empty_cs addSEs [conjE] addSIs [conjI])) :: - flat (mapn (foo_tacs ctxt) 1 (conss ~~ cases)); - in pg [] goal tacs end; - - fun one_finite (dn, l1b) = - let - val goal = mk_trp (%%:(dn^"_finite") $ %:"x"); - fun tacs ctxt = [ - case_UU_tac ctxt take_rews 1 "x", - eresolve_tac finite_lemmas1a 1, - step_tac HOL_cs 1, - step_tac HOL_cs 1, - cut_facts_tac [l1b] 1, - fast_tac HOL_cs 1]; - in pg axs_finite_def goal tacs end; - - val _ = trace " Proving finites"; - val finites = map one_finite (dnames ~~ atomize global_ctxt finite_lemma1b); - val _ = trace " Proving ind"; - val ind = - let - fun concf n dn = %:(P_name n) $ %:(x_name n); - fun tacf {prems, context} = - let - fun finite_tacs (finite, fin_ind) = [ - rtac(rewrite_rule axs_finite_def finite RS exE)1, - etac subst 1, - rtac fin_ind 1, - ind_prems_tac prems]; - in - TRY (safe_tac HOL_cs) :: - maps finite_tacs (finites ~~ atomize global_ctxt finite_ind) - end; - in pg'' thy [] (ind_term concf) tacf end; - in (finites, ind) end (* let *) - - else (* infinite case *) - let - fun one_finite n dn = - read_instantiate global_ctxt [(("P", 0), dn ^ "_finite " ^ x_name n)] excluded_middle; - val finites = mapn one_finite 1 dnames; - - val goal = - let - fun one_adm n _ = mk_trp (mk_adm (%:(P_name n))); - fun concf n dn = %:(P_name n) $ %:(x_name n); - in Logic.list_implies (mapn one_adm 1 dnames, ind_term concf) end; - fun tacf {prems, context} = - map (fn ax_reach => rtac (ax_reach RS subst) 1) axs_reach @ [ - quant_tac context 1, - rtac (adm_impl_admw RS wfix_ind) 1, - REPEAT_DETERM (rtac adm_all 1), - REPEAT_DETERM ( - TRY (rtac adm_conj 1) THEN - rtac adm_subst 1 THEN - cont_tacR 1 THEN resolve_tac prems 1), - strip_tac 1, - rtac (rewrite_rule axs_take_def finite_ind) 1, - ind_prems_tac prems]; - val ind = (pg'' thy [] goal tacf - handle ERROR _ => - (warning "Cannot prove infinite induction rule"; refl)); - in (finites, ind) end - ) - handle THM _ => - (warning "Induction proofs failed (THM raised)."; ([], TrueI)) - | ERROR _ => - (warning "Induction proofs failed (ERROR raised)."; ([], TrueI)); - - -end; (* local *) - -(* ----- theorem concerning coinduction ------------------------------------- *) - -local - val xs = mapn (fn n => K (x_name n)) 1 dnames; - fun bnd_arg n i = Bound(2*(n_eqs - n)-i-1); - val take_ss = HOL_ss addsimps take_rews; - val sproj = prj (fn s => K("fst("^s^")")) (fn s => K("snd("^s^")")); - val _ = trace " Proving coind_lemma..."; - val coind_lemma = - let - fun mk_prj n _ = proj (%:"R") eqs n $ bnd_arg n 0 $ bnd_arg n 1; - fun mk_eqn n dn = - (dc_take dn $ %:"n" ` bnd_arg n 0) === - (dc_take dn $ %:"n" ` bnd_arg n 1); - fun mk_all2 (x,t) = mk_all (x, mk_all (x^"'", t)); - val goal = - mk_trp (mk_imp (%%:(comp_dname^"_bisim") $ %:"R", - Library.foldr mk_all2 (xs, - Library.foldr mk_imp (mapn mk_prj 0 dnames, - foldr1 mk_conj (mapn mk_eqn 0 dnames))))); - fun x_tacs ctxt n x = [ - rotate_tac (n+1) 1, - etac all2E 1, - eres_inst_tac ctxt [(("P", 1), sproj "R" eqs n^" "^x^" "^x^"'")] (mp RS disjE) 1, - TRY (safe_tac HOL_cs), - REPEAT (CHANGED (asm_simp_tac take_ss 1))]; - fun tacs ctxt = [ - rtac impI 1, - InductTacs.induct_tac ctxt [[SOME "n"]] 1, - simp_tac take_ss 1, - safe_tac HOL_cs] @ - flat (mapn (x_tacs ctxt) 0 xs); - in pg [ax_bisim_def] goal tacs end; -in - val _ = trace " Proving coind..."; - val coind = - let - fun mk_prj n x = mk_trp (proj (%:"R") eqs n $ %:x $ %:(x^"'")); - fun mk_eqn x = %:x === %:(x^"'"); - val goal = - mk_trp (%%:(comp_dname^"_bisim") $ %:"R") ===> - Logic.list_implies (mapn mk_prj 0 xs, - mk_trp (foldr1 mk_conj (map mk_eqn xs))); - val tacs = - TRY (safe_tac HOL_cs) :: - maps (fn take_lemma => [ - rtac take_lemma 1, - cut_facts_tac [coind_lemma] 1, - fast_tac HOL_cs 1]) - take_lemmas; - in pg [] goal (K tacs) end; -end; (* local *) - -val inducts = ProjectRule.projections (ProofContext.init thy) ind; -fun ind_rule (dname, rule) = ((Binding.empty, [rule]), [Induct.induct_type dname]); -val induct_failed = (Thm.prop_of ind = Thm.prop_of TrueI); - -in thy |> Sign.add_path comp_dnam - |> snd o PureThy.add_thmss [ - ((Binding.name "take_rews" , take_rews ), [Simplifier.simp_add]), - ((Binding.name "take_lemmas", take_lemmas ), []), - ((Binding.name "finites" , finites ), []), - ((Binding.name "finite_ind" , [finite_ind]), []), - ((Binding.name "ind" , [ind] ), []), - ((Binding.name "coind" , [coind] ), [])] - |> (if induct_failed then I - else snd o PureThy.add_thmss (map ind_rule (dnames ~~ inducts))) - |> Sign.parent_path |> pair take_rews -end; (* let *) -end; (* local *) -end; (* struct *) diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/Tools/fixrec.ML --- a/src/HOLCF/Tools/fixrec.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOLCF/Tools/fixrec.ML Thu Jul 23 21:13:21 2009 +0200 @@ -158,13 +158,13 @@ val thy = ProofContext.theory_of lthy; val names = map (Binding.name_of o fst o fst) fixes; val all_names = space_implode "_" names; - val (lhss,rhss) = ListPair.unzip (map (dest_eqs o snd) spec); + val (lhss, rhss) = ListPair.unzip (map (dest_eqs o snd) spec); val functional = lambda_tuple lhss (mk_tuple rhss); val fixpoint = mk_fix (mk_cabs functional); val cont_thm = Goal.prove lthy [] [] (mk_trp (mk_cont functional)) - (K (simp_tac (local_simpset_of lthy) 1)); + (K (simp_tac (simpset_of lthy) 1)); fun one_def (l as Free(n,_)) r = let val b = Long_Name.base_name n @@ -311,12 +311,12 @@ (*************************************************************************) (* proves a block of pattern matching equations as theorems, using unfold *) -fun make_simps lthy (unfold_thm, eqns : (Attrib.binding * term) list) = +fun make_simps ctxt (unfold_thm, eqns : (Attrib.binding * term) list) = let val tacs = [rtac (unfold_thm RS @{thm ssubst_lhs}) 1, - asm_simp_tac (local_simpset_of lthy) 1]; - fun prove_term t = Goal.prove lthy [] [] t (K (EVERY tacs)); + asm_simp_tac (simpset_of ctxt) 1]; + fun prove_term t = Goal.prove ctxt [] [] t (K (EVERY tacs)); fun prove_eqn (bind, eqn_t) = (bind, prove_term eqn_t); in map prove_eqn eqns @@ -399,7 +399,7 @@ fixrec_err "function is not declared as constant in theory"; val unfold_thm = PureThy.get_thm thy (cname^"_unfold"); val simp = Goal.prove_global thy [] [] eq - (fn _ => EVERY [stac unfold_thm 1, simp_tac (simpset_of thy) 1]); + (fn _ => EVERY [stac unfold_thm 1, simp_tac (global_simpset_of thy) 1]); in simp end; fun gen_add_fixpat prep_term prep_attrib ((name, srcs), strings) thy = diff -r 63686057cbe8 -r abda97d2deea src/HOLCF/holcf_logic.ML --- a/src/HOLCF/holcf_logic.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/HOLCF/holcf_logic.ML Thu Jul 23 21:13:21 2009 +0200 @@ -1,5 +1,4 @@ (* Title: HOLCF/holcf_logic.ML - ID: $Id$ Author: David von Oheimb Abstract syntax operations for HOLCF. @@ -10,7 +9,6 @@ signature HOLCF_LOGIC = sig val mk_btyp: string -> typ * typ -> typ - val pcpoC: class val pcpoS: sort val mk_TFree: string -> typ val cfun_arrow: string @@ -27,8 +25,7 @@ (* sort pcpo *) -val pcpoC = Sign.intern_class (the_context ()) "pcpo"; -val pcpoS = [pcpoC]; +val pcpoS = @{sort pcpo}; fun mk_TFree s = TFree ("'" ^ s, pcpoS); diff -r 63686057cbe8 -r abda97d2deea src/Provers/Arith/fast_lin_arith.ML --- a/src/Provers/Arith/fast_lin_arith.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Provers/Arith/fast_lin_arith.ML Thu Jul 23 21:13:21 2009 +0200 @@ -417,8 +417,8 @@ (* Translate back a proof. *) (* ------------------------------------------------------------------------- *) -fun trace_thm msg th = - (if !trace then (tracing msg; tracing (Display.string_of_thm th)) else (); th); +fun trace_thm ctxt msg th = + (if !trace then (tracing msg; tracing (Display.string_of_thm ctxt th)) else (); th); fun trace_term ctxt msg t = (if !trace then tracing (cat_lines [msg, Syntax.string_of_term ctxt t]) else (); t) @@ -472,7 +472,7 @@ NONE => (the (try_add ([thm2] RL inj_thms) thm1) handle Option => - (trace_thm "" thm1; trace_thm "" thm2; + (trace_thm ctxt "" thm1; trace_thm ctxt "" thm2; sys_error "Linear arithmetic: failed to add thms")) | SOME thm => thm) | SOME thm => thm); @@ -511,24 +511,25 @@ else mult n thm; fun simp thm = - let val thm' = trace_thm "Simplified:" (full_simplify simpset' thm) + let val thm' = trace_thm ctxt "Simplified:" (full_simplify simpset' thm) in if LA_Logic.is_False thm' then raise FalseE thm' else thm' end; - fun mk (Asm i) = trace_thm ("Asm " ^ string_of_int i) (nth asms i) - | mk (Nat i) = trace_thm ("Nat " ^ string_of_int i) (LA_Logic.mk_nat_thm thy (nth atoms i)) - | mk (LessD j) = trace_thm "L" (hd ([mk j] RL lessD)) - | mk (NotLeD j) = trace_thm "NLe" (mk j RS LA_Logic.not_leD) - | mk (NotLeDD j) = trace_thm "NLeD" (hd ([mk j RS LA_Logic.not_leD] RL lessD)) - | mk (NotLessD j) = trace_thm "NL" (mk j RS LA_Logic.not_lessD) - | mk (Added (j1, j2)) = simp (trace_thm "+" (add_thms (mk j1) (mk j2))) - | mk (Multiplied (n, j)) = (trace_msg ("*" ^ string_of_int n); trace_thm "*" (mult_thm (n, mk j))) + fun mk (Asm i) = trace_thm ctxt ("Asm " ^ string_of_int i) (nth asms i) + | mk (Nat i) = trace_thm ctxt ("Nat " ^ string_of_int i) (LA_Logic.mk_nat_thm thy (nth atoms i)) + | mk (LessD j) = trace_thm ctxt "L" (hd ([mk j] RL lessD)) + | mk (NotLeD j) = trace_thm ctxt "NLe" (mk j RS LA_Logic.not_leD) + | mk (NotLeDD j) = trace_thm ctxt "NLeD" (hd ([mk j RS LA_Logic.not_leD] RL lessD)) + | mk (NotLessD j) = trace_thm ctxt "NL" (mk j RS LA_Logic.not_lessD) + | mk (Added (j1, j2)) = simp (trace_thm ctxt "+" (add_thms (mk j1) (mk j2))) + | mk (Multiplied (n, j)) = + (trace_msg ("*" ^ string_of_int n); trace_thm ctxt "*" (mult_thm (n, mk j))) in let val _ = trace_msg "mkthm"; - val thm = trace_thm "Final thm:" (mk just); + val thm = trace_thm ctxt "Final thm:" (mk just); val fls = simplify simpset' thm; - val _ = trace_thm "After simplification:" fls; + val _ = trace_thm ctxt "After simplification:" fls; val _ = if LA_Logic.is_False fls then () else @@ -536,15 +537,15 @@ if count > warning_count_max then () else (tracing (cat_lines - (["Assumptions:"] @ map Display.string_of_thm asms @ [""] @ - ["Proved:", Display.string_of_thm fls, ""] @ + (["Assumptions:"] @ map (Display.string_of_thm ctxt) asms @ [""] @ + ["Proved:", Display.string_of_thm ctxt fls, ""] @ (if count <> warning_count_max then [] else ["\n(Reached maximal message count -- disabling future warnings)"]))); warning "Linear arithmetic should have refuted the assumptions.\n\ \Please inform Tobias Nipkow (nipkow@in.tum.de).") end; in fls end - handle FalseE thm => trace_thm "False reached early:" thm + handle FalseE thm => trace_thm ctxt "False reached early:" thm end; end; @@ -775,8 +776,9 @@ fn state => let val ctxt = Simplifier.the_context ss; - val _ = trace_thm ("refute_tac (on subgoal " ^ string_of_int i ^ ", with " ^ - string_of_int (length justs) ^ " justification(s)):") state + val _ = trace_thm ctxt + ("refute_tac (on subgoal " ^ string_of_int i ^ ", with " ^ + string_of_int (length justs) ^ " justification(s)):") state val {neqE, ...} = get_data ctxt; fun just1 j = (* eliminate inequalities *) @@ -784,7 +786,7 @@ REPEAT_DETERM (eresolve_tac neqE i) else all_tac) THEN - PRIMITIVE (trace_thm "State after neqE:") THEN + PRIMITIVE (trace_thm ctxt "State after neqE:") THEN (* use theorems generated from the actual justifications *) METAHYPS (fn asms => rtac (mkthm ss asms j) 1) i in @@ -792,7 +794,7 @@ DETERM (resolve_tac [LA_Logic.notI, LA_Logic.ccontr] i) THEN (* user-defined preprocessing of the subgoal *) DETERM (LA_Data.pre_tac ctxt i) THEN - PRIMITIVE (trace_thm "State after pre_tac:") THEN + PRIMITIVE (trace_thm ctxt "State after pre_tac:") THEN (* prove every resulting subgoal, using its justification *) EVERY (map just1 justs) end state; @@ -881,7 +883,7 @@ val (Falsethm, _) = fwdproof ss tree js val contr = if pos then LA_Logic.ccontr else LA_Logic.notI val concl = implies_intr cnTconcl Falsethm COMP contr - in SOME (trace_thm "Proved by lin. arith. prover:" (LA_Logic.mk_Eq concl)) end + in SOME (trace_thm ctxt "Proved by lin. arith. prover:" (LA_Logic.mk_Eq concl)) end (*in case concl contains ?-var, which makes assume fail:*) (* FIXME Variable.import_terms *) handle THM _ => NONE; diff -r 63686057cbe8 -r abda97d2deea src/Provers/blast.ML --- a/src/Provers/blast.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Provers/blast.ML Thu Jul 23 21:13:21 2009 +0200 @@ -492,7 +492,9 @@ end; -fun TRACE rl tac st i = if !trace then (Display.prth rl; tac st i) else tac st i; +fun TRACE rl tac st i = + if !trace then (writeln (Display.string_of_thm_without_context rl); tac st i) + else tac st i; (*Resolution/matching tactics: if upd then the proof state may be updated. Matching makes the tactics more deterministic in the presence of Vars.*) @@ -509,14 +511,16 @@ THEN rot_subgoals_tac (rot, #2 trl) i in Option.SOME (trl, tac) end - handle ElimBadPrem => (*reject: prems don't preserve conclusion*) - (warning("Ignoring weak elimination rule\n" ^ Display.string_of_thm rl); - Option.NONE) - | ElimBadConcl => (*ignore: conclusion is not just a variable*) - (if !trace then (warning("Ignoring ill-formed elimination rule:\n" ^ - "conclusion should be a variable\n" ^ Display.string_of_thm rl)) - else (); - Option.NONE); + handle + ElimBadPrem => (*reject: prems don't preserve conclusion*) + (warning ("Ignoring weak elimination rule\n" ^ Display.string_of_thm_global thy rl); + Option.NONE) + | ElimBadConcl => (*ignore: conclusion is not just a variable*) + (if !trace then + (warning ("Ignoring ill-formed elimination rule:\n" ^ + "conclusion should be a variable\n" ^ Display.string_of_thm_global thy rl)) + else (); + Option.NONE); (*** Conversion of Introduction Rules ***) diff -r 63686057cbe8 -r abda97d2deea src/Provers/clasimp.ML --- a/src/Provers/clasimp.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Provers/clasimp.ML Thu Jul 23 21:13:21 2009 +0200 @@ -26,7 +26,7 @@ type clasimpset val get_css: Context.generic -> clasimpset val map_css: (clasimpset -> clasimpset) -> Context.generic -> Context.generic - val local_clasimpset_of: Proof.context -> clasimpset + val clasimpset_of: Proof.context -> clasimpset val addSIs2: clasimpset * thm list -> clasimpset val addSEs2: clasimpset * thm list -> clasimpset val addSDs2: clasimpset * thm list -> clasimpset @@ -73,8 +73,7 @@ let val (cs', ss') = f (get_css context) in context |> Classical.map_cs (K cs') |> Simplifier.map_ss (K ss') end; -fun local_clasimpset_of ctxt = - (Classical.local_claset_of ctxt, Simplifier.local_simpset_of ctxt); +fun clasimpset_of ctxt = (Classical.claset_of ctxt, Simplifier.simpset_of ctxt); (* clasimpset operations *) @@ -261,10 +260,10 @@ (* methods *) fun clasimp_meth tac prems ctxt = METHOD (fn facts => - ALLGOALS (Method.insert_tac (prems @ facts)) THEN tac (local_clasimpset_of ctxt)); + ALLGOALS (Method.insert_tac (prems @ facts)) THEN tac (clasimpset_of ctxt)); fun clasimp_meth' tac prems ctxt = METHOD (fn facts => - HEADGOAL (Method.insert_tac (prems @ facts) THEN' tac (local_clasimpset_of ctxt))); + HEADGOAL (Method.insert_tac (prems @ facts) THEN' tac (clasimpset_of ctxt))); fun clasimp_method tac = diff -r 63686057cbe8 -r abda97d2deea src/Provers/classical.ML --- a/src/Provers/classical.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Provers/classical.ML Thu Jul 23 21:13:21 2009 +0200 @@ -69,8 +69,8 @@ val appSWrappers : claset -> wrapper val appWrappers : claset -> wrapper - val claset_of: theory -> claset - val local_claset_of : Proof.context -> claset + val global_claset_of : theory -> claset + val claset_of : Proof.context -> claset val fast_tac : claset -> int -> tactic val slow_tac : claset -> int -> tactic @@ -256,7 +256,7 @@ xtra_netpair = empty_netpair}; fun print_cs (CS {safeIs, safeEs, hazIs, hazEs, swrappers, uwrappers, ...}) = - let val pretty_thms = map Display.pretty_thm in + let val pretty_thms = map Display.pretty_thm_without_context in [Pretty.big_list "safe introduction rules (intro!):" (pretty_thms safeIs), Pretty.big_list "introduction rules (intro):" (pretty_thms hazIs), Pretty.big_list "safe elimination rules (elim!):" (pretty_thms safeEs), @@ -313,14 +313,18 @@ (*Warn if the rule is already present ELSEWHERE in the claset. The addition is still allowed.*) fun warn_dup th (CS{safeIs, safeEs, hazIs, hazEs, ...}) = - if mem_thm safeIs th then - warning ("Rule already declared as safe introduction (intro!)\n" ^ Display.string_of_thm th) + if mem_thm safeIs th then + warning ("Rule already declared as safe introduction (intro!)\n" ^ + Display.string_of_thm_without_context th) else if mem_thm safeEs th then - warning ("Rule already declared as safe elimination (elim!)\n" ^ Display.string_of_thm th) + warning ("Rule already declared as safe elimination (elim!)\n" ^ + Display.string_of_thm_without_context th) else if mem_thm hazIs th then - warning ("Rule already declared as introduction (intro)\n" ^ Display.string_of_thm th) + warning ("Rule already declared as introduction (intro)\n" ^ + Display.string_of_thm_without_context th) else if mem_thm hazEs th then - warning ("Rule already declared as elimination (elim)\n" ^ Display.string_of_thm th) + warning ("Rule already declared as elimination (elim)\n" ^ + Display.string_of_thm_without_context th) else (); @@ -330,8 +334,8 @@ (cs as CS {safeIs, safeEs, hazIs, hazEs, swrappers, uwrappers, safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair}) = if mem_thm safeIs th then - (warning ("Ignoring duplicate safe introduction (intro!)\n" ^ Display.string_of_thm th); - cs) + (warning ("Ignoring duplicate safe introduction (intro!)\n" ^ + Display.string_of_thm_without_context th); cs) else let val th' = flat_rule th val (safe0_rls, safep_rls) = (*0 subgoals vs 1 or more*) @@ -356,10 +360,10 @@ (cs as CS {safeIs, safeEs, hazIs, hazEs, swrappers, uwrappers, safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair}) = if mem_thm safeEs th then - (warning ("Ignoring duplicate safe elimination (elim!)\n" ^ Display.string_of_thm th); - cs) + (warning ("Ignoring duplicate safe elimination (elim!)\n" ^ + Display.string_of_thm_without_context th); cs) else if has_fewer_prems 1 th then - error("Ill-formed elimination rule\n" ^ Display.string_of_thm th) + error ("Ill-formed elimination rule\n" ^ Display.string_of_thm_without_context th) else let val th' = classical_rule (flat_rule th) @@ -386,7 +390,7 @@ fun make_elim th = if has_fewer_prems 1 th then - error("Ill-formed destruction rule\n" ^ Display.string_of_thm th) + error ("Ill-formed destruction rule\n" ^ Display.string_of_thm_without_context th) else Tactic.make_elim th; fun cs addSDs ths = cs addSEs (map make_elim ths); @@ -398,8 +402,8 @@ (cs as CS {safeIs, safeEs, hazIs, hazEs, swrappers, uwrappers, safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair}) = if mem_thm hazIs th then - (warning ("Ignoring duplicate introduction (intro)\n" ^ Display.string_of_thm th); - cs) + (warning ("Ignoring duplicate introduction (intro)\n" ^ + Display.string_of_thm_without_context th); cs) else let val th' = flat_rule th val nI = length hazIs + 1 @@ -418,16 +422,16 @@ xtra_netpair = insert' (the_default 1 w) (nI,nE) ([th], []) xtra_netpair} end handle THM("RSN: no unifiers",_,_) => (*from dup_intr*) - error ("Ill-formed introduction rule\n" ^ Display.string_of_thm th); + error ("Ill-formed introduction rule\n" ^ Display.string_of_thm_without_context th); fun addE w th (cs as CS {safeIs, safeEs, hazIs, hazEs, swrappers, uwrappers, safe0_netpair, safep_netpair, haz_netpair, dup_netpair, xtra_netpair}) = if mem_thm hazEs th then - (warning ("Ignoring duplicate elimination (elim)\n" ^ Display.string_of_thm th); - cs) + (warning ("Ignoring duplicate elimination (elim)\n" ^ + Display.string_of_thm_without_context th); cs) else if has_fewer_prems 1 th then - error("Ill-formed elimination rule\n" ^ Display.string_of_thm th) + error("Ill-formed elimination rule\n" ^ Display.string_of_thm_without_context th) else let val th' = classical_rule (flat_rule th) @@ -519,7 +523,7 @@ end else cs handle THM("RSN: no unifiers",_,_) => (*from dup_intr*) - error ("Ill-formed introduction rule\n" ^ Display.string_of_thm th); + error ("Ill-formed introduction rule\n" ^ Display.string_of_thm_without_context th); fun delE th @@ -548,7 +552,7 @@ mem_thm hazIs th orelse mem_thm hazEs th orelse mem_thm safeEs th' orelse mem_thm hazEs th' then delSI th (delSE th (delI th (delE th (delSE th' (delE th' cs))))) - else (warning ("Undeclared classical rule\n" ^ Display.string_of_thm th); cs) + else (warning ("Undeclared classical rule\n" ^ Display.string_of_thm_without_context th); cs) end; fun cs delrules ths = fold delrule ths cs; @@ -848,7 +852,7 @@ fun map_context_cs f = GlobalClaset.map (apsnd (fn ContextCS {swrappers, uwrappers} => make_context_cs (f (swrappers, uwrappers)))); -fun claset_of thy = +fun global_claset_of thy = let val (cs, ctxt_cs) = GlobalClaset.get thy in context_cs (ProofContext.init thy) cs (ctxt_cs) end; @@ -870,13 +874,13 @@ val init = get_claset; ); -fun local_claset_of ctxt = +fun claset_of ctxt = context_cs ctxt (LocalClaset.get ctxt) (get_context_cs ctxt); (* generic clasets *) -val get_cs = Context.cases claset_of local_claset_of; +val get_cs = Context.cases global_claset_of claset_of; fun map_cs f = Context.mapping (map_claset f) (LocalClaset.map f); @@ -926,7 +930,7 @@ fun some_rule_tac ctxt facts = SUBGOAL (fn (goal, i) => let val [rules1, rules2, rules4] = ContextRules.find_rules false facts goal ctxt; - val CS {xtra_netpair, ...} = local_claset_of ctxt; + val CS {xtra_netpair, ...} = claset_of ctxt; val rules3 = ContextRules.find_rules_netpair true facts goal xtra_netpair; val rules = rules1 @ rules2 @ rules3 @ rules4; val ruleq = Drule.multi_resolves facts rules; @@ -965,10 +969,10 @@ Args.del -- Args.colon >> K (I, rule_del)]; fun cla_meth tac prems ctxt = METHOD (fn facts => - ALLGOALS (Method.insert_tac (prems @ facts)) THEN tac (local_claset_of ctxt)); + ALLGOALS (Method.insert_tac (prems @ facts)) THEN tac (claset_of ctxt)); fun cla_meth' tac prems ctxt = METHOD (fn facts => - HEADGOAL (Method.insert_tac (prems @ facts) THEN' tac (local_claset_of ctxt))); + HEADGOAL (Method.insert_tac (prems @ facts) THEN' tac (claset_of ctxt))); fun cla_method tac = Args.bang_facts --| Method.sections cla_modifiers >> (cla_meth tac); fun cla_method' tac = Args.bang_facts --| Method.sections cla_modifiers >> (cla_meth' tac); @@ -1010,6 +1014,6 @@ OuterSyntax.improper_command "print_claset" "print context of Classical Reasoner" OuterKeyword.diag (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_context o - Toplevel.keep (print_cs o local_claset_of o Toplevel.context_of))); + Toplevel.keep (print_cs o claset_of o Toplevel.context_of))); end; diff -r 63686057cbe8 -r abda97d2deea src/Pure/Concurrent/future.ML --- a/src/Pure/Concurrent/future.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Concurrent/future.ML Thu Jul 23 21:13:21 2009 +0200 @@ -29,7 +29,8 @@ val enabled: unit -> bool type task = Task_Queue.task type group = Task_Queue.group - val thread_data: unit -> (string * task) option + val is_worker: unit -> bool + val worker_group: unit -> Task_Queue.group option type 'a future val task_of: 'a future -> task val group_of: 'a future -> group @@ -66,11 +67,17 @@ type task = Task_Queue.task; type group = Task_Queue.group; -local val tag = Universal.tag () : (string * task) option Universal.tag in +local + val tag = Universal.tag () : (string * task * group) option Universal.tag; +in fun thread_data () = the_default NONE (Thread.getLocal tag); - fun setmp_thread_data data f x = Library.setmp_thread_data tag (thread_data ()) (SOME data) f x; + fun setmp_thread_data data f x = + Library.setmp_thread_data tag (thread_data ()) (SOME data) f x; end; +val is_worker = is_some o thread_data; +val worker_group = Option.map #3 o thread_data; + (* datatype future *) @@ -87,7 +94,7 @@ fun value x = Future {task = Task_Queue.new_task 0, - group = Task_Queue.new_group (), + group = Task_Queue.new_group NONE, result = ref (SOME (Exn.Result x))}; @@ -128,18 +135,43 @@ (* worker activity *) -fun trace_active () = +fun count_active ws = + fold (fn (_, active) => fn i => if active then i + 1 else i) ws 0; + +fun trace_active () = Multithreading.tracing 1 (fn () => let val ws = ! workers; val m = string_of_int (length ws); - val n = string_of_int (length (filter #2 ws)); - in Multithreading.tracing 1 (fn () => "SCHEDULE: " ^ m ^ " workers, " ^ n ^ " active") end; + val n = string_of_int (count_active ws); + in "SCHEDULE: " ^ m ^ " workers, " ^ n ^ " active" end); fun change_active active = (*requires SYNCHRONIZED*) change workers (AList.update Thread.equal (Thread.self (), active)); +fun overloaded () = + count_active (! workers) > Multithreading.max_threads_value (); -(* execute jobs *) + +(* execute future jobs *) + +fun future_job group (e: unit -> 'a) = + let + val result = ref (NONE: 'a Exn.result option); + fun job ok = + let + val res = + if ok then + Exn.capture + (Multithreading.with_attributes Multithreading.restricted_interrupts + (fn _ => fn () => e ())) () + else Exn.Exn Exn.Interrupt; + val _ = result := SOME res; + in + (case res of + Exn.Exn exn => (Task_Queue.cancel_group group exn; false) + | Exn.Result _ => true) + end; + in (result, job) end; fun do_cancel group = (*requires SYNCHRONIZED*) change canceled (insert Task_Queue.eq_group group); @@ -147,8 +179,8 @@ fun execute name (task, group, jobs) = let val _ = trace_active (); - val valid = Task_Queue.is_valid group; - val ok = setmp_thread_data (name, task) (fn () => + val valid = not (Task_Queue.is_canceled group); + val ok = setmp_thread_data (name, task, group) (fn () => fold (fn job => fn ok => job valid andalso ok) jobs true) (); val _ = SYNCHRONIZED "execute" (fn () => (change queue (Task_Queue.finish task); @@ -170,13 +202,14 @@ change workers (filter_out (fn (thread, _) => Thread.equal (thread, Thread.self ()))); notify_all (); NONE) + else if overloaded () then (worker_wait (); worker_next ()) else (case change_result queue Task_Queue.dequeue of NONE => (worker_wait (); worker_next ()) | some => some); fun worker_loop name = - (case SYNCHRONIZED name worker_next of + (case SYNCHRONIZED name (fn () => worker_next ()) of NONE => () | SOME work => (execute name work; worker_loop name)); @@ -188,27 +221,41 @@ fun scheduler_next () = (*requires SYNCHRONIZED*) let + (*queue status*) + val _ = Multithreading.tracing 1 (fn () => + let val {ready, pending, running} = Task_Queue.status (! queue) in + "SCHEDULE: " ^ + string_of_int ready ^ " ready, " ^ + string_of_int pending ^ " pending, " ^ + string_of_int running ^ " running" + end); + (*worker threads*) + val ws = ! workers; val _ = - (case List.partition (Thread.isActive o #1) (! workers) of - (_, []) => () - | (active, inactive) => - (workers := active; Multithreading.tracing 0 (fn () => - "SCHEDULE: disposed " ^ string_of_int (length inactive) ^ " dead worker threads"))); + if forall (Thread.isActive o #1) ws then () + else + (case List.partition (Thread.isActive o #1) ws of + (_, []) => () + | (active, inactive) => + (workers := active; Multithreading.tracing 0 (fn () => + "SCHEDULE: disposed " ^ string_of_int (length inactive) ^ " dead worker threads"))); val _ = trace_active (); val m = if ! do_shutdown then 0 else Multithreading.max_threads_value (); - val l = length (! workers); - val _ = excessive := l - m; + val mm = (m * 3) div 2; + val l = length ws; + val _ = excessive := l - mm; val _ = - if m > l then funpow (m - l) (fn () => worker_start ("worker " ^ string_of_int (inc next))) () + if mm > l then + funpow (mm - l) (fn () => worker_start ("worker " ^ string_of_int (inc next))) () else (); (*canceled groups*) - val _ = change canceled (filter_out (Task_Queue.cancel (! queue))); + val _ = change canceled (filter_out (Task_Queue.cancel (! queue))); (*shutdown*) - val continue = not (! do_shutdown andalso null (! workers)); + val continue = not (! do_shutdown andalso null ws); val _ = if continue then () else scheduler := NONE; val _ = notify_all (); @@ -218,7 +265,7 @@ in continue end; fun scheduler_loop () = - while SYNCHRONIZED "scheduler" scheduler_next do (); + while SYNCHRONIZED "scheduler" (fn () => scheduler_next ()) do (); fun scheduler_active () = (*requires SYNCHRONIZED*) (case ! scheduler of NONE => false | SOME thread => Thread.isActive thread); @@ -233,32 +280,16 @@ (** futures **) -(* future job: fill result *) - -fun future_job group (e: unit -> 'a) = - let - val result = ref (NONE: 'a Exn.result option); - val job = Multithreading.with_attributes Multithreading.restricted_interrupts - (fn _ => fn ok => - let - val res = if ok then Exn.capture e () else Exn.Exn Exn.Interrupt; - val _ = result := SOME res; - val res_ok = - (case res of - Exn.Result _ => true - | Exn.Exn Exn.Interrupt => (Task_Queue.invalidate_group group; true) - | _ => false); - in res_ok end); - in (result, job) end; - - (* fork *) fun fork_future opt_group deps pri e = let val _ = scheduler_check "future check"; - val group = (case opt_group of SOME group => group | NONE => Task_Queue.new_group ()); + val group = + (case opt_group of + SOME group => group + | NONE => Task_Queue.new_group (worker_group ())); val (result, job) = future_job group e; val task = SYNCHRONIZED "future" (fn () => change_result queue (Task_Queue.enqueue group deps pri job) before notify_all ()); @@ -274,23 +305,21 @@ local -fun get_result x = the_default (Exn.Exn (SYS_ERROR "unfinished future")) (peek x); - -fun join_wait x = - while not (is_finished x) - do SYNCHRONIZED "join_wait" (fn () => wait ()); +fun get_result x = + (case peek x of + NONE => Exn.Exn (SYS_ERROR "unfinished future") + | SOME (Exn.Exn Exn.Interrupt) => + Exn.Exn (Exn.EXCEPTIONS (Exn.flatten_list (Task_Queue.group_status (group_of x)))) + | SOME res => res); -fun join_next x = (*requires SYNCHRONIZED*) - if is_finished x then NONE - else - (case change_result queue Task_Queue.dequeue of - NONE => (worker_wait (); join_next x) - | some => some); +fun join_next deps = (*requires SYNCHRONIZED*) + if overloaded () then (worker_wait (); join_next deps) + else change_result queue (Task_Queue.dequeue_towards deps); -fun join_loop x = - (case SYNCHRONIZED "join" (fn () => join_next x) of +fun join_deps deps = + (case SYNCHRONIZED "join" (fn () => join_next deps) of NONE => () - | SOME work => (execute "join" work; join_loop x)); + | SOME (work, deps') => (execute "join" work; join_deps deps')); in @@ -301,10 +330,16 @@ val _ = scheduler_check "join check"; val _ = Multithreading.self_critical () andalso error "Cannot join future values within critical section"; - val _ = - if is_some (thread_data ()) - then List.app join_loop xs (*proper task -- continue work*) - else List.app join_wait xs; (*alien thread -- refrain from contending for resources*) + + val worker = is_worker (); + val _ = if worker then join_deps (map task_of xs) else (); + + fun join_wait x = + if SYNCHRONIZED "join_wait" (fn () => + is_finished x orelse (if worker then worker_wait () else wait (); false)) + then () else join_wait x; + val _ = List.app join_wait xs; + in map get_result xs end) (); end; @@ -320,7 +355,7 @@ val _ = scheduler_check "map_future check"; val task = task_of x; - val group = Task_Queue.new_group (); + val group = Task_Queue.new_group (SOME (group_of x)); val (result, job) = future_job group (fn () => f (join x)); val extended = SYNCHRONIZED "map_future" (fn () => @@ -329,7 +364,7 @@ | NONE => false)); in if extended then Future {task = task, group = group, result = result} - else fork_future NONE [task] (Task_Queue.pri_of_task task) (fn () => f (join x)) + else fork_future (SOME group) [task] (Task_Queue.pri_of_task task) (fn () => f (join x)) end; @@ -338,7 +373,7 @@ fun interruptible_task f x = if Multithreading.available then Multithreading.with_attributes - (if is_some (thread_data ()) + (if is_worker () then Multithreading.restricted_interrupts else Multithreading.regular_interrupts) (fn _ => f) x diff -r 63686057cbe8 -r abda97d2deea src/Pure/Concurrent/par_list.ML --- a/src/Pure/Concurrent/par_list.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Concurrent/par_list.ML Thu Jul 23 21:13:21 2009 +0200 @@ -28,11 +28,8 @@ fun raw_map f xs = if Future.enabled () then - let - val group = Task_Queue.new_group (); - val futures = map (fn x => Future.fork_group group (fn () => f x)) xs; - val _ = List.app (ignore o Future.join_result) futures; - in Future.join_results futures end + let val group = Task_Queue.new_group (Future.worker_group ()) + in Future.join_results (map (fn x => Future.fork_group group (fn () => f x)) xs) end else map (Exn.capture f) xs; fun map f xs = Exn.release_first (raw_map f xs); diff -r 63686057cbe8 -r abda97d2deea src/Pure/Concurrent/simple_thread.ML --- a/src/Pure/Concurrent/simple_thread.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Concurrent/simple_thread.ML Thu Jul 23 21:13:21 2009 +0200 @@ -25,13 +25,16 @@ fun synchronized name lock e = Exn.release (uninterruptible (fn restore_attributes => fn () => let - val _ = - if Mutex.trylock lock then () + val immediate = + if Mutex.trylock lock then true else (Multithreading.tracing 3 (fn () => name ^ ": locking ..."); Mutex.lock lock; - Multithreading.tracing 3 (fn () => name ^ ": ... locked")); + Multithreading.tracing 3 (fn () => name ^ ": locked"); + false); val result = Exn.capture (restore_attributes e) (); + val _ = + if immediate then () else Multithreading.tracing 3 (fn () => name ^ ": unlocking ..."); val _ = Mutex.unlock lock; in result end) ()); diff -r 63686057cbe8 -r abda97d2deea src/Pure/Concurrent/task_queue.ML --- a/src/Pure/Concurrent/task_queue.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Concurrent/task_queue.ML Thu Jul 23 21:13:21 2009 +0200 @@ -11,19 +11,24 @@ val pri_of_task: task -> int val str_of_task: task -> string type group + val group_id: group -> int val eq_group: group * group -> bool - val new_group: unit -> group - val is_valid: group -> bool - val invalidate_group: group -> unit + val new_group: group option -> group + val group_status: group -> exn list val str_of_group: group -> string type queue val empty: queue val is_empty: queue -> bool + val status: queue -> {ready: int, pending: int, running: int} val enqueue: group -> task list -> int -> (bool -> bool) -> queue -> task * queue val extend: task -> (bool -> bool) -> queue -> queue option val dequeue: queue -> (task * group * (bool -> bool) list) option * queue + val dequeue_towards: task list -> queue -> + (((task * group * (bool -> bool) list) * task list) option * queue) val interrupt: queue -> task -> unit val interrupt_external: queue -> string -> unit + val is_canceled: group -> bool + val cancel_group: group -> exn -> unit val cancel: queue -> group -> bool val cancel_all: queue -> group list val finish: task -> queue -> queue @@ -44,18 +49,37 @@ structure Task_Graph = Graph(type key = task val ord = task_ord); -(* groups *) +(* nested groups *) + +datatype group = Group of + {parent: group option, + id: serial, + status: exn list ref}; -datatype group = Group of serial * bool ref; -fun eq_group (Group (gid1, _), Group (gid2, _)) = gid1 = gid2; +fun make_group (parent, id, status) = Group {parent = parent, id = id, status = status}; + +fun new_group parent = make_group (parent, serial (), ref []); + +fun group_id (Group {id, ...}) = id; +fun eq_group (group1, group2) = group_id group1 = group_id group2; -fun new_group () = Group (serial (), ref true); +fun group_ancestry (Group {parent, id, ...}) = + id :: (case parent of NONE => [] | SOME group => group_ancestry group); + + +fun cancel_group (Group {status, ...}) exn = CRITICAL (fn () => + (case exn of + Exn.Interrupt => if null (! status) then status := [exn] else () + | _ => change status (cons exn))); -fun is_valid (Group (_, ref ok)) = ok; -fun invalidate_group (Group (_, ok)) = ok := false; +fun group_status (Group {parent, status, ...}) = (*non-critical*) + ! status @ (case parent of NONE => [] | SOME group => group_status group); -fun str_of_group (Group (i, ref ok)) = - if ok then string_of_int i else enclose "(" ")" (string_of_int i); +fun is_canceled (Group {parent, status, ...}) = (*non-critical*) + not (null (! status)) orelse (case parent of NONE => false | SOME group => is_canceled group); + +fun str_of_group group = + (is_canceled group ? enclose "(" ")") (string_of_int (group_id group)); (* jobs *) @@ -89,14 +113,51 @@ fun is_empty (Queue {jobs, ...}) = Task_Graph.is_empty jobs; +(* queue status *) + +fun status (Queue {jobs, ...}) = + let + val (x, y, z) = + Task_Graph.fold (fn (task, ((_, job), (deps, _))) => fn (x, y, z) => + (case job of + Job _ => if null deps then (x + 1, y, z) else (x, y + 1, z) + | Running _ => (x, y, z + 1))) + jobs (0, 0, 0); + in {ready = x, pending = y, running = z} end; + + +(* cancel -- peers and sub-groups *) + +fun cancel (Queue {groups, jobs, ...}) group = + let + val _ = cancel_group group Exn.Interrupt; + val tasks = Inttab.lookup_list groups (group_id group); + val running = + fold (get_job jobs #> (fn Running t => insert Thread.equal t | _ => I)) tasks []; + val _ = List.app SimpleThread.interrupt running; + in null running end; + +fun cancel_all (Queue {jobs, ...}) = + let + fun cancel_job (group, job) (groups, running) = + (cancel_group group Exn.Interrupt; + (case job of Running t => (insert eq_group group groups, insert Thread.equal t running) + | _ => (groups, running))); + val (groups, running) = Task_Graph.fold (cancel_job o #1 o #2) jobs ([], []); + val _ = List.app SimpleThread.interrupt running; + in groups end; + + (* enqueue *) -fun enqueue (group as Group (gid, _)) deps pri job (Queue {groups, jobs, cache}) = +fun enqueue group deps pri job (Queue {groups, jobs, cache}) = let val task = new_task pri; - val groups' = Inttab.cons_list (gid, task) groups; + val groups' = groups + |> fold (fn gid => Inttab.cons_list (gid, task)) (group_ancestry group); val jobs' = jobs - |> Task_Graph.new_node (task, (group, Job [job])) |> fold (add_job task) deps; + |> Task_Graph.new_node (task, (group, Job [job])) + |> fold (add_job task) deps; val cache' = (case cache of Result last => @@ -133,10 +194,40 @@ end; +(* dequeue_towards -- adhoc dependencies *) + +fun dequeue_towards deps (queue as Queue {groups, jobs, ...}) = + let + fun ready task = + (case Task_Graph.get_node jobs task of + (group, Job list) => + if null (Task_Graph.imm_preds jobs task) + then SOME (task, group, rev list) + else NONE + | _ => NONE); + + val tasks = filter (can (Task_Graph.get_node jobs)) deps; + fun result (res as (task, _, _)) = + let + val jobs' = set_job task (Running (Thread.self ())) jobs; + val cache' = Unknown; + in (SOME (res, tasks), make_queue groups jobs' cache') end; + in + (case get_first ready tasks of + SOME res => result res + | NONE => + (case get_first ready (Task_Graph.all_preds jobs tasks) of + SOME res => result res + | NONE => (NONE, queue))) + end; + + (* sporadic interrupts *) fun interrupt (Queue {jobs, ...}) task = - (case try (get_job jobs) task of SOME (Running thread) => SimpleThread.interrupt thread | _ => ()); + (case try (get_job jobs) task of + SOME (Running thread) => SimpleThread.interrupt thread + | _ => ()); fun interrupt_external (queue as Queue {jobs, ...}) str = (case Int.fromString str of @@ -149,28 +240,11 @@ (* termination *) -fun cancel (Queue {groups, jobs, ...}) (group as Group (gid, _)) = - let - val _ = invalidate_group group; - val tasks = Inttab.lookup_list groups gid; - val running = fold (get_job jobs #> (fn Running t => insert Thread.equal t | _ => I)) tasks []; - val _ = List.app SimpleThread.interrupt running; - in null running end; - -fun cancel_all (Queue {jobs, ...}) = - let - fun cancel_job (group, job) (groups, running) = - (invalidate_group group; - (case job of Running t => (insert eq_group group groups, insert Thread.equal t running) - | _ => (groups, running))); - val (groups, running) = Task_Graph.fold (cancel_job o #1 o #2) jobs ([], []); - val _ = List.app SimpleThread.interrupt running; - in groups end; - fun finish task (Queue {groups, jobs, cache}) = let - val Group (gid, _) = get_group jobs task; - val groups' = Inttab.remove_list (op =) (gid, task) groups; + val group = get_group jobs task; + val groups' = groups + |> fold (fn gid => Inttab.remove_list (op =) (gid, task)) (group_ancestry group); val jobs' = Task_Graph.del_node task jobs; val cache' = if null (Task_Graph.imm_succs jobs task) then cache diff -r 63686057cbe8 -r abda97d2deea src/Pure/General/same.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/General/same.ML Thu Jul 23 21:13:21 2009 +0200 @@ -0,0 +1,45 @@ +(* Title: Pure/General/same.ML + Author: Makarius + +Support for copy-avoiding functions on pure values, at the cost of +readability. +*) + +signature SAME = +sig + exception SAME + type ('a, 'b) function = 'a -> 'b (*exception SAME*) + type 'a operation = ('a, 'a) function (*exception SAME*) + val same: ('a, 'b) function + val commit: 'a operation -> 'a -> 'a + val function: ('a -> 'b option) -> ('a, 'b) function + val capture: ('a, 'b) function -> 'a -> 'b option + val map: 'a operation -> 'a list operation + val map_option: ('a, 'b) function -> ('a option, 'b option) function +end; + +structure Same: SAME = +struct + +exception SAME; + +type ('a, 'b) function = 'a -> 'b; +type 'a operation = ('a, 'a) function; + +fun same _ = raise SAME; +fun commit f x = f x handle SAME => x; + +fun capture f x = SOME (f x) handle SAME => NONE; + +fun function f x = + (case f x of + NONE => raise SAME + | SOME y => y); + +fun map f [] = raise SAME + | map f (x :: xs) = (f x :: commit (map f) xs handle SAME => x :: map f xs); + +fun map_option f NONE = raise SAME + | map_option f (SOME x) = SOME (f x); + +end; diff -r 63686057cbe8 -r abda97d2deea src/Pure/IsaMakefile --- a/src/Pure/IsaMakefile Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/IsaMakefile Thu Jul 23 21:13:21 2009 +0200 @@ -52,12 +52,12 @@ General/long_name.ML General/markup.ML General/name_space.ML \ General/ord_list.ML General/output.ML General/path.ML \ General/position.ML General/pretty.ML General/print_mode.ML \ - General/properties.ML General/queue.ML General/scan.ML \ - General/secure.ML General/seq.ML General/source.ML General/stack.ML \ - General/symbol.ML General/symbol_pos.ML General/table.ML \ - General/url.ML General/xml.ML General/yxml.ML Isar/args.ML \ - Isar/attrib.ML Isar/auto_bind.ML Isar/calculation.ML Isar/class.ML \ - Isar/class_target.ML Isar/code.ML Isar/constdefs.ML \ + General/properties.ML General/queue.ML General/same.ML \ + General/scan.ML General/secure.ML General/seq.ML General/source.ML \ + General/stack.ML General/symbol.ML General/symbol_pos.ML \ + General/table.ML General/url.ML General/xml.ML General/yxml.ML \ + Isar/args.ML Isar/attrib.ML Isar/auto_bind.ML Isar/calculation.ML \ + Isar/class.ML Isar/class_target.ML Isar/code.ML Isar/constdefs.ML \ Isar/context_rules.ML Isar/element.ML Isar/expression.ML \ Isar/isar_cmd.ML Isar/isar_document.ML Isar/isar_syn.ML \ Isar/local_defs.ML Isar/local_syntax.ML Isar/local_theory.ML \ @@ -90,13 +90,13 @@ Tools/find_theorems.ML Tools/named_thms.ML Tools/xml_syntax.ML \ assumption.ML axclass.ML codegen.ML config.ML conjunction.ML \ consts.ML context.ML context_position.ML conv.ML defs.ML display.ML \ - drule.ML envir.ML facts.ML goal.ML interpretation.ML item_net.ML \ - library.ML logic.ML meta_simplifier.ML more_thm.ML morphism.ML \ - name.ML net.ML old_goals.ML old_term.ML pattern.ML primitive_defs.ML \ - proofterm.ML pure_setup.ML pure_thy.ML search.ML sign.ML \ - simplifier.ML sorts.ML subgoal.ML tactic.ML tctical.ML term.ML \ - term_ord.ML term_subst.ML theory.ML thm.ML type.ML type_infer.ML \ - unify.ML variable.ML + display_goal.ML drule.ML envir.ML facts.ML goal.ML interpretation.ML \ + item_net.ML library.ML logic.ML meta_simplifier.ML more_thm.ML \ + morphism.ML name.ML net.ML old_goals.ML old_term.ML pattern.ML \ + primitive_defs.ML proofterm.ML pure_setup.ML pure_thy.ML search.ML \ + sign.ML simplifier.ML sorts.ML subgoal.ML tactic.ML tctical.ML \ + term.ML term_ord.ML term_subst.ML theory.ML thm.ML type.ML \ + type_infer.ML unify.ML variable.ML @./mk diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/args.ML --- a/src/Pure/Isar/args.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/args.ML Thu Jul 23 21:13:21 2009 +0200 @@ -88,7 +88,7 @@ fun pretty_src ctxt src = let - val prt_thm = Pretty.backquote o ProofContext.pretty_thm ctxt; + val prt_thm = Pretty.backquote o Display.pretty_thm ctxt; fun prt arg = (case T.get_value arg of SOME (T.Text s) => Pretty.str (quote s) diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/calculation.ML --- a/src/Pure/Isar/calculation.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/calculation.ML Thu Jul 23 21:13:21 2009 +0200 @@ -40,8 +40,8 @@ fun print_rules ctxt = let val ((trans, sym), _) = CalculationData.get (Context.Proof ctxt) in [Pretty.big_list "transitivity rules:" - (map (ProofContext.pretty_thm ctxt) (Item_Net.content trans)), - Pretty.big_list "symmetry rules:" (map (ProofContext.pretty_thm ctxt) sym)] + (map (Display.pretty_thm ctxt) (Item_Net.content trans)), + Pretty.big_list "symmetry rules:" (map (Display.pretty_thm ctxt) sym)] |> Pretty.chunks |> Pretty.writeln end; diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/class.ML --- a/src/Pure/Isar/class.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/class.ML Thu Jul 23 21:13:21 2009 +0200 @@ -68,9 +68,8 @@ val base_morph = inst_morph $> Morphism.binding_morphism (Binding.prefix false (class_prefix class)) $> Element.satisfy_morphism (the_list wit); - val defs = these_defs thy sups; - val eq_morph = Element.eq_morphism thy defs; - val morph = base_morph $> eq_morph; + val eqs = these_defs thy sups; + val eq_morph = Element.eq_morphism thy eqs; (* assm_intro *) fun prove_assm_intro thm = @@ -97,7 +96,7 @@ ORELSE' Tactic.assume_tac)); val of_class = SkipProof.prove_global thy [] [] of_class_prop (K tac); - in (base_morph, morph, export_morph, axiom, assm_intro, of_class) end; + in (base_morph, eqs, export_morph, axiom, assm_intro, of_class) end; (* reading and processing class specifications *) @@ -284,9 +283,8 @@ ||> Theory.checkpoint |-> (fn (param_map, params, assm_axiom) => `(fn thy => calculate thy class sups base_sort param_map assm_axiom) - #-> (fn (base_morph, morph, export_morph, axiom, assm_intro, of_class) => - Locale.add_registration (class, (morph, export_morph)) - #> Context.theory_map (Locale.activate_facts (class, morph $> export_morph)) + #-> (fn (base_morph, eqs, export_morph, axiom, assm_intro, of_class) => + Locale.add_registration_eqs (class, base_morph) eqs export_morph #> register class sups params base_sort base_morph axiom assm_intro of_class)) |> TheoryTarget.init (SOME class) |> pair class diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/class_target.ML --- a/src/Pure/Isar/class_target.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/class_target.ML Thu Jul 23 21:13:21 2009 +0200 @@ -242,15 +242,16 @@ val diff_sort = Sign.complete_sort thy [sup] |> subtract (op =) (Sign.complete_sort thy [sub]) |> filter (is_class thy); - val deps_witss = case some_dep_morph - of SOME dep_morph => [((sup, dep_morph), the_list some_wit)] - | NONE => [] + val add_dependency = case some_dep_morph + of SOME dep_morph => Locale.add_dependency sub + (sup, dep_morph $> Element.satisfy_morphism (the_list some_wit)) export + | NONE => I in thy |> AxClass.add_classrel classrel |> ClassData.map (Graph.add_edge (sub, sup)) |> activate_defs sub (these_defs thy diff_sort) - |> Locale.add_dependencies sub deps_witss export + |> add_dependency end; diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/code.ML --- a/src/Pure/Isar/code.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/code.ML Thu Jul 23 21:13:21 2009 +0200 @@ -243,7 +243,7 @@ (*default flag, theorems with proper flag (perhaps lazy)*) fun pretty_lthms ctxt r = case Lazy.peek r - of SOME thms => map (ProofContext.pretty_thm ctxt o fst) (Exn.release thms) + of SOME thms => map (Display.pretty_thm ctxt o fst) (Exn.release thms) | NONE => [Pretty.str "[...]"]; fun certificate thy f r = @@ -263,7 +263,8 @@ Pattern.matchess thy (args, (map incr_idx o curry Library.take (length args)) args'); fun drop (thm', proper') = if (proper orelse not proper') andalso matches_args (args_of thm') then - (warning ("Code generator: dropping redundant code equation\n" ^ Display.string_of_thm thm'); true) + (warning ("Code generator: dropping redundant code equation\n" ^ + Display.string_of_thm_global thy thm'); true) else false; in (thm, proper) :: filter_out drop thms end; @@ -567,16 +568,16 @@ fun gen_assert_eqn thy is_constr_pat (thm, proper) = let val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm - handle TERM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm) - | THM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm thm); + handle TERM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm_global thy thm) + | THM _ => bad_thm ("Not an equation: " ^ Display.string_of_thm_global thy thm); fun vars_of t = fold_aterms (fn Var (v, _) => insert (op =) v | Free _ => bad_thm ("Illegal free variable in equation\n" - ^ Display.string_of_thm thm) + ^ Display.string_of_thm_global thy thm) | _ => I) t []; fun tvars_of t = fold_term_types (fn _ => fold_atyps (fn TVar (v, _) => insert (op =) v | TFree _ => bad_thm - ("Illegal free type variable in equation\n" ^ Display.string_of_thm thm))) t []; + ("Illegal free type variable in equation\n" ^ Display.string_of_thm_global thy thm))) t []; val lhs_vs = vars_of lhs; val rhs_vs = vars_of rhs; val lhs_tvs = tvars_of lhs; @@ -584,47 +585,48 @@ val _ = if null (subtract (op =) lhs_vs rhs_vs) then () else bad_thm ("Free variables on right hand side of equation\n" - ^ Display.string_of_thm thm); + ^ Display.string_of_thm_global thy thm); val _ = if null (subtract (op =) lhs_tvs rhs_tvs) then () else bad_thm ("Free type variables on right hand side of equation\n" - ^ Display.string_of_thm thm) val (head, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm; + ^ Display.string_of_thm_global thy thm) + val (head, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm; val (c, ty) = case head of Const (c_ty as (_, ty)) => (AxClass.unoverload_const thy c_ty, ty) - | _ => bad_thm ("Equation not headed by constant\n" ^ Display.string_of_thm thm); + | _ => bad_thm ("Equation not headed by constant\n" ^ Display.string_of_thm_global thy thm); fun check _ (Abs _) = bad_thm ("Abstraction on left hand side of equation\n" - ^ Display.string_of_thm thm) + ^ Display.string_of_thm_global thy thm) | check 0 (Var _) = () | check _ (Var _) = bad_thm ("Variable with application on left hand side of equation\n" - ^ Display.string_of_thm thm) + ^ Display.string_of_thm_global thy thm) | check n (t1 $ t2) = (check (n+1) t1; check 0 t2) | check n (Const (c_ty as (c, ty))) = if n = (length o fst o strip_type) ty then if not proper orelse is_constr_pat (AxClass.unoverload_const thy c_ty) then () else bad_thm (quote c ^ " is not a constructor, on left hand side of equation\n" - ^ Display.string_of_thm thm) + ^ Display.string_of_thm_global thy thm) else bad_thm ("Partially applied constant " ^ quote c ^ " on left hand side of equation\n" - ^ Display.string_of_thm thm); + ^ Display.string_of_thm_global thy thm); val _ = map (check 0) args; val _ = if not proper orelse is_linear thm then () else bad_thm ("Duplicate variables on left hand side of equation\n" - ^ Display.string_of_thm thm); + ^ Display.string_of_thm_global thy thm); val _ = if (is_none o AxClass.class_of_param thy) c then () else bad_thm ("Polymorphic constant as head in equation\n" - ^ Display.string_of_thm thm) + ^ Display.string_of_thm_global thy thm) val _ = if not (is_constr thy c) then () else bad_thm ("Constructor as head in equation\n" - ^ Display.string_of_thm thm) + ^ Display.string_of_thm_global thy thm) val ty_decl = Sign.the_const_type thy c; val _ = if Sign.typ_equiv thy (Type.strip_sorts ty_decl, Type.strip_sorts ty) then () else bad_thm ("Type\n" ^ string_of_typ thy ty ^ "\nof equation\n" - ^ Display.string_of_thm thm + ^ Display.string_of_thm_global thy thm ^ "\nis incompatible with declared function type\n" ^ string_of_typ thy ty_decl) in (thm, proper) end; @@ -657,7 +659,7 @@ let fun cert (eqn as (thm, _)) = if c = const_eqn thy thm then eqn else error ("Wrong head of code equation,\nexpected constant " - ^ string_of_const thy c ^ "\n" ^ Display.string_of_thm thm) + ^ string_of_const thy c ^ "\n" ^ Display.string_of_thm_global thy thm) in map (cert o assert_eqn thy) eqns end; fun common_typ_eqns thy [] = [] @@ -674,7 +676,7 @@ fun unify ty env = Sign.typ_unify thy (ty1, ty) env handle Type.TUNIFY => error ("Type unificaton failed, while unifying code equations\n" - ^ (cat_lines o map Display.string_of_thm) thms + ^ (cat_lines o map (Display.string_of_thm_global thy)) thms ^ "\nwith types\n" ^ (cat_lines o map (string_of_typ thy)) (ty1 :: tys)); val (env, _) = fold unify tys (Vartab.empty, maxidx) @@ -796,15 +798,17 @@ let val (ofclass, eqn) = case try Logic.dest_equals (Thm.prop_of thm) of SOME ofclass_eq => ofclass_eq - | _ => error ("Bad certificate: " ^ Display.string_of_thm thm); + | _ => error ("Bad certificate: " ^ Display.string_of_thm_global thy thm); val (T, class) = case try Logic.dest_of_class ofclass of SOME T_class => T_class - | _ => error ("Bad certificate: " ^ Display.string_of_thm thm); + | _ => error ("Bad certificate: " ^ Display.string_of_thm_global thy thm); val tvar = case try Term.dest_TVar T - of SOME tvar => tvar - | _ => error ("Bad type: " ^ Display.string_of_thm thm); + of SOME (tvar as (_, sort)) => if null (filter (can (AxClass.get_info thy)) sort) + then tvar + else error ("Bad sort: " ^ Display.string_of_thm_global thy thm) + | _ => error ("Bad type: " ^ Display.string_of_thm_global thy thm); val _ = if Term.add_tvars eqn [] = [tvar] then () - else error ("Inconsistent type: " ^ Display.string_of_thm thm); + else error ("Inconsistent type: " ^ Display.string_of_thm_global thy thm); val lhs_rhs = case try Logic.dest_equals eqn of SOME lhs_rhs => lhs_rhs | _ => error ("Not an equation: " ^ Syntax.string_of_term_global thy eqn); @@ -813,7 +817,7 @@ | _ => error ("Not an equation with two constants: " ^ Syntax.string_of_term_global thy eqn); val _ = if the_list (AxClass.class_of_param thy (snd c_c')) = [class] then () - else error ("Inconsistent class: " ^ Display.string_of_thm thm); + else error ("Inconsistent class: " ^ Display.string_of_thm_global thy thm); in thy |> (map_exec_purge NONE o map_aliasses) (fn (alias, classes) => ((c_c', thm) :: alias, insert (op =) class classes)) @@ -886,7 +890,9 @@ of SOME (thm, _) => change_eqns true (const_eqn thy thm) (del_thm thm) thy | NONE => thy; -structure Code_Attr = TheoryDataFun ( +(* c.f. src/HOL/Tools/recfun_codegen.ML *) + +structure Code_Target_Attr = TheoryDataFun ( type T = (string -> thm -> theory -> theory) option; val empty = NONE; val copy = I; @@ -895,7 +901,14 @@ | merge _ (f1, _) = f1; ); -fun set_code_target_attr f = Code_Attr.map (K (SOME f)); +fun set_code_target_attr f = Code_Target_Attr.map (K (SOME f)); + +fun code_target_attr prefix thm thy = + let + val attr = the_default ((K o K) I) (Code_Target_Attr.get thy); + in thy |> add_warning_eqn thm |> attr prefix thm end; + +(* setup *) val _ = Context.>> (Context.map_theory (let @@ -904,9 +917,7 @@ Args.del |-- Scan.succeed (mk_attribute del_eqn) || Args.$$$ "nbe" |-- Scan.succeed (mk_attribute add_nbe_eqn) || (Args.$$$ "target" |-- Args.colon |-- Args.name >> - (fn prefix => mk_attribute (fn thm => fn thy => thy - |> add_warning_eqn thm - |> the_default ((K o K) I) (Code_Attr.get thy) prefix thm))) + (mk_attribute o code_target_attr)) || Scan.succeed (mk_attribute add_warning_eqn); in Type_Interpretation.init diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/context_rules.ML --- a/src/Pure/Isar/context_rules.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/context_rules.ML Thu Jul 23 21:13:21 2009 +0200 @@ -116,7 +116,7 @@ fun prt_kind (i, b) = Pretty.big_list ((the o AList.lookup (op =) kind_names) (i, b) ^ ":") (map_filter (fn (_, (k, th)) => - if k = (i, b) then SOME (ProofContext.pretty_thm ctxt th) else NONE) + if k = (i, b) then SOME (Display.pretty_thm ctxt th) else NONE) (sort (int_ord o pairself fst) rules)); in Pretty.writeln (Pretty.chunks (map prt_kind rule_kinds)) end; diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/element.ML --- a/src/Pure/Isar/element.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/element.ML Thu Jul 23 21:13:21 2009 +0200 @@ -163,7 +163,7 @@ let val prt_typ = Pretty.quote o Syntax.pretty_typ ctxt; val prt_term = Pretty.quote o Syntax.pretty_term ctxt; - val prt_thm = Pretty.backquote o ProofContext.pretty_thm ctxt; + val prt_thm = Pretty.backquote o Display.pretty_thm ctxt; val prt_name_atts = pretty_name_atts ctxt; fun prt_mixfix NoSyn = [] diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/expression.ML --- a/src/Pure/Isar/expression.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/expression.ML Thu Jul 23 21:13:21 2009 +0200 @@ -788,27 +788,6 @@ (*** Interpretation ***) -(** Interpretation between locales: declaring sublocale relationships **) - -local - -fun gen_sublocale prep_expr intern raw_target expression thy = - let - val target = intern thy raw_target; - val target_ctxt = Locale.init target thy; - val ((propss, deps, export), goal_ctxt) = prep_expr expression target_ctxt; - fun after_qed witss = ProofContext.theory - (Locale.add_dependencies target (deps ~~ witss) export); - in Element.witness_proof after_qed propss goal_ctxt end; - -in - -fun sublocale x = gen_sublocale cert_goal_expression (K I) x; -fun sublocale_cmd x = gen_sublocale read_goal_expression Locale.intern x; - -end; - - (** Interpretation in theories **) local @@ -816,46 +795,31 @@ fun gen_interpretation prep_expr parse_prop prep_attr expression equations theory = let - val ((propss, regs, export), expr_ctxt) = ProofContext.init theory + val ((propss, deps, export), expr_ctxt) = ProofContext.init theory |> prep_expr expression; val eqns = map (parse_prop expr_ctxt o snd) equations |> Syntax.check_terms expr_ctxt; - val eq_attns = map ((apsnd o map) (prep_attr theory) o fst) equations; + val attrss = map ((apsnd o map) (prep_attr theory) o fst) equations; val goal_ctxt = fold Variable.auto_fixes eqns expr_ctxt; val export' = Variable.export_morphism goal_ctxt expr_ctxt; - fun store_reg ((name, morph), wits) thy = + fun note_eqns raw_eqns thy = let - val wits' = map (Element.morph_witness export') wits; - val morph' = morph $> Element.satisfy_morphism wits'; + val eqns = map (Morphism.thm (export' $> export)) raw_eqns; + val eqn_attrss = map2 (fn attrs => fn eqn => + ((apsnd o map) (Attrib.attribute_i thy) attrs, [([eqn], [])])) attrss eqns; + fun meta_rewrite thy = map (LocalDefs.meta_rewrite_rule (ProofContext.init thy) #> + Drule.abs_def) o maps snd; in thy - |> Locale.add_registration (name, (morph', export)) - |> pair (name, morph') + |> PureThy.note_thmss Thm.lemmaK eqn_attrss + |-> (fn facts => `(fn thy => meta_rewrite thy facts)) end; - fun store_eqns_activate regs [] thy = - thy - |> fold (fn (name, morph) => - Context.theory_map (Locale.activate_facts (name, morph $> export))) regs - | store_eqns_activate regs eqs thy = - let - val eqs' = eqs |> map (Morphism.thm (export' $> export)); - val morph_eqs = eqs' |> map (LocalDefs.meta_rewrite_rule (ProofContext.init thy) #> - Drule.abs_def); - val eq_morph = Element.eq_morphism thy morph_eqs; - val eq_attns' = map ((apsnd o map) (Attrib.attribute_i thy)) eq_attns; - in - thy - |> fold (fn (name, morph) => - Locale.amend_registration eq_morph (name, morph) #> - Context.theory_map (Locale.activate_facts (name, morph $> eq_morph $> export))) regs - |> PureThy.note_thmss Thm.lemmaK (eq_attns' ~~ map (fn eq => [([eq], [])]) eqs') - |> snd - end; - - fun after_qed wits eqs = ProofContext.theory (fold_map store_reg (regs ~~ wits) - #-> (fn regs => store_eqns_activate regs eqs)); + fun after_qed witss eqns = ProofContext.theory (note_eqns eqns + #-> (fn eqns => fold (fn ((dep, morph), wits) => + Locale.add_registration_eqs (dep, morph $> Element.satisfy_morphism + (map (Element.morph_witness export') wits)) eqns export) (deps ~~ witss))); in Element.witness_proof_eqs after_qed propss eqns goal_ctxt end; @@ -868,12 +832,33 @@ end; +(** Interpretation between locales: declaring sublocale relationships **) + +local + +fun gen_sublocale prep_expr intern raw_target expression thy = + let + val target = intern thy raw_target; + val target_ctxt = Locale.init target thy; + val ((propss, deps, export), goal_ctxt) = prep_expr expression target_ctxt; + fun after_qed witss = ProofContext.theory + (fold (fn ((dep, morph), wits) => Locale.add_dependency + target (dep, morph $> Element.satisfy_morphism wits) export) (deps ~~ witss)); + in Element.witness_proof after_qed propss goal_ctxt end; + +in + +fun sublocale x = gen_sublocale cert_goal_expression (K I) x; +fun sublocale_cmd x = gen_sublocale read_goal_expression Locale.intern x; + +end; + + (** Interpretation in proof contexts **) local -fun gen_interpret prep_expr - expression int state = +fun gen_interpret prep_expr expression int state = let val _ = Proof.assert_forward_or_chain state; val ctxt = Proof.context_of state; diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/isar_cmd.ML --- a/src/Pure/Isar/isar_cmd.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/isar_cmd.ML Thu Jul 23 21:13:21 2009 +0200 @@ -230,7 +230,7 @@ (* local endings *) fun local_qed m = Toplevel.proof (Proof.local_qed (m, true)); -val local_terminal_proof = Toplevel.proof o Proof.local_terminal_proof; +val local_terminal_proof = Toplevel.proof' o Proof.local_future_terminal_proof; val local_default_proof = Toplevel.proof Proof.local_default_proof; val local_immediate_proof = Toplevel.proof Proof.local_immediate_proof; val local_done_proof = Toplevel.proof Proof.local_done_proof; @@ -363,7 +363,7 @@ val print_simpset = Toplevel.unknown_context o Toplevel.keep (fn state => let val ctxt = Toplevel.context_of state - in Pretty.writeln (Simplifier.pretty_ss ctxt (Simplifier.local_simpset_of ctxt)) end); + in Pretty.writeln (Simplifier.pretty_ss ctxt (simpset_of ctxt)) end); val print_rules = Toplevel.unknown_context o Toplevel.keep (ContextRules.print_rules o Toplevel.context_of); @@ -442,8 +442,7 @@ |> Pretty.chunks2 |> Pretty.string_of; fun string_of_thms state args = - Pretty.string_of (ProofContext.pretty_thms (Proof.context_of state) - (Proof.get_thmss state args)); + Pretty.string_of (Display.pretty_thms (Proof.context_of state) (Proof.get_thmss state args)); fun string_of_prfs full state arg = Pretty.string_of (case arg of diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/local_defs.ML --- a/src/Pure/Isar/local_defs.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/local_defs.ML Thu Jul 23 21:13:21 2009 +0200 @@ -196,7 +196,7 @@ fun print_rules ctxt = Pretty.writeln (Pretty.big_list "definitional transformations:" - (map (ProofContext.pretty_thm ctxt) (Rules.get (Context.Proof ctxt)))); + (map (Display.pretty_thm ctxt) (Rules.get (Context.Proof ctxt)))); val defn_add = Thm.declaration_attribute (Rules.map o Thm.add_thm); val defn_del = Thm.declaration_attribute (Rules.map o Thm.del_thm); diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/locale.ML --- a/src/Pure/Isar/locale.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/locale.ML Thu Jul 23 21:13:21 2009 +0200 @@ -68,10 +68,9 @@ val intro_locales_tac: bool -> Proof.context -> thm list -> tactic (* Registrations and dependencies *) - val add_registration: string * (morphism * morphism) -> theory -> theory + val add_registration_eqs: string * morphism -> thm list -> morphism -> theory -> theory val amend_registration: morphism -> string * morphism -> theory -> theory - val add_dependencies: string -> ((string * morphism) * Element.witness list) list -> - morphism -> theory -> theory + val add_dependency: string -> string * morphism -> morphism -> theory -> theory (* Diagnostic *) val print_locales: theory -> unit @@ -295,8 +294,7 @@ fun activate_declarations dep = Context.proof_map (fn context => let val thy = Context.theory_of context; - val context' = roundup thy activate_decls dep (get_idents context, context) |-> put_idents; - in context' end); + in roundup thy activate_decls dep (get_idents context, context) |-> put_idents end); fun activate_facts dep context = let @@ -346,11 +344,6 @@ fun all_registrations thy = Registrations.get thy |> map reg_morph; -fun add_registration (name, (base_morph, export)) thy = - roundup thy (fn (name', morph') => Registrations.map (cons ((name', (morph', export)), stamp ()))) - (name, base_morph) (get_idents (Context.Theory thy), thy) - (* FIXME |-> (Context.theory_map o put_idents_diag)*) |> snd; - fun amend_registration morph (name, base_morph) thy = let val regs = map #1 (Registrations.get thy); @@ -368,14 +361,24 @@ (fn ((name, (base, export)), stamp) => ((name, (base $> morph, export)), stamp))) thy end; +fun add_registration_eqs (dep, proto_morph) eqns export thy = + let + val morph = if null eqns then proto_morph + else proto_morph $> Element.eq_morphism thy eqns; + in + (get_idents (Context.Theory thy), thy) + |> roundup thy (fn (dep', morph') => + Registrations.map (cons ((dep', (morph', export)), stamp ()))) (dep, morph) + |> snd + |> Context.theory_map (activate_facts (dep, morph $> export)) + end; + (*** Dependencies ***) -fun add_dependencies loc deps_witss export thy = +fun add_dependency loc (dep, morph) export thy = thy - |> fold (fn ((dep, morph), wits) => (change_locale loc o apsnd) - (cons ((dep, morph $> Element.satisfy_morphism wits $> export), stamp ()))) - deps_witss + |> (change_locale loc o apsnd) (cons ((dep, morph $> export), stamp ())) |> (fn thy => fold_rev (Context.theory_map o activate_facts) (all_registrations thy) thy); diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/method.ML --- a/src/Pure/Isar/method.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/method.ML Thu Jul 23 21:13:21 2009 +0200 @@ -220,7 +220,7 @@ fun trace ctxt rules = if ! trace_rules andalso not (null rules) then - Pretty.big_list "rules:" (map (ProofContext.pretty_thm ctxt) rules) + Pretty.big_list "rules:" (map (Display.pretty_thm ctxt) rules) |> Pretty.string_of |> tracing else (); diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/obtain.ML --- a/src/Pure/Isar/obtain.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/obtain.ML Thu Jul 23 21:13:21 2009 +0200 @@ -180,9 +180,9 @@ [prem] => if Thm.concl_of th aconv thesis andalso Logic.strip_assums_concl prem aconv thesis then th - else error ("Guessed a different clause:\n" ^ ProofContext.string_of_thm ctxt th) + else error ("Guessed a different clause:\n" ^ Display.string_of_thm ctxt th) | [] => error "Goal solved -- nothing guessed." - | _ => error ("Guess split into several cases:\n" ^ ProofContext.string_of_thm ctxt th)); + | _ => error ("Guess split into several cases:\n" ^ Display.string_of_thm ctxt th)); fun result tac facts ctxt = let @@ -218,7 +218,7 @@ val string_of_typ = Syntax.string_of_typ ctxt; val string_of_term = setmp show_types true (Syntax.string_of_term ctxt); - fun err msg th = error (msg ^ ":\n" ^ ProofContext.string_of_thm ctxt th); + fun err msg th = error (msg ^ ":\n" ^ Display.string_of_thm ctxt th); val maxidx = fold (Term.maxidx_typ o snd o fst) vars ~1; val rule = Thm.incr_indexes (maxidx + 1) raw_rule; diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/overloading.ML --- a/src/Pure/Isar/overloading.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/overloading.ML Thu Jul 23 21:13:21 2009 +0200 @@ -76,7 +76,7 @@ | _ => I) | accumulate_improvements _ = I; val improvements = (fold o fold_aterms) accumulate_improvements ts Vartab.empty; - val ts' = (map o map_types) (Envir.typ_subst_TVars improvements) ts; + val ts' = (map o map_types) (Envir.subst_type improvements) ts; fun apply_subst t = Envir.expand_term (fn Const (c, ty) => (case subst (c, ty) of SOME (ty', t') => if Type.typ_instance tsig (ty, ty') diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/proof.ML --- a/src/Pure/Isar/proof.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/proof.ML Thu Jul 23 21:13:21 2009 +0200 @@ -318,11 +318,9 @@ val show_main_goal = ref false; val verbose = ProofContext.verbose; -val pretty_goals_local = Display.pretty_goals_aux o Syntax.pp; - fun pretty_facts _ _ NONE = [] | pretty_facts s ctxt (SOME ths) = - [Pretty.big_list (s ^ "this:") (map (ProofContext.pretty_thm ctxt) ths), Pretty.str ""]; + [Pretty.big_list (s ^ "this:") (map (Display.pretty_thm ctxt) ths), Pretty.str ""]; fun pretty_state nr state = let @@ -346,7 +344,7 @@ (if mode <> Backward orelse null using then NONE else SOME using) @ [Pretty.str ("goal" ^ description [levels_up (i div 2), subgoals (Thm.nprems_of goal)] ^ ":")] @ - pretty_goals_local ctxt Markup.subgoal + Display_Goal.pretty_goals ctxt Markup.subgoal (true, ! show_main_goal) (! Display.goals_limit) goal @ (map (fn msg => Position.setmp_thread_data pos msg ()) (rev messages)) | prt_goal NONE = []; @@ -368,7 +366,7 @@ fun pretty_goals main state = let val (ctxt, (_, goal)) = get_goal state - in pretty_goals_local ctxt Markup.none (false, main) (! Display.goals_limit) goal end; + in Display_Goal.pretty_goals ctxt Markup.none (false, main) (! Display.goals_limit) goal end; @@ -470,11 +468,12 @@ let val thy = ProofContext.theory_of ctxt; val string_of_term = Syntax.string_of_term ctxt; - val string_of_thm = ProofContext.string_of_thm ctxt; + val string_of_thm = Display.string_of_thm ctxt; val ngoals = Thm.nprems_of goal; val _ = ngoals = 0 orelse error (Pretty.string_of (Pretty.chunks - (pretty_goals_local ctxt Markup.none (true, ! show_main_goal) (! Display.goals_limit) goal @ + (Display_Goal.pretty_goals ctxt Markup.none + (true, ! show_main_goal) (! Display.goals_limit) goal @ [Pretty.str (string_of_int ngoals ^ " unsolved goal(s)!")]))); val extra_hyps = Assumption.extra_hyps ctxt goal; @@ -1002,6 +1001,7 @@ val _ = assert_backward state; val (goal_ctxt, (_, goal)) = find_goal state; val {statement as (kind, propss, prop), messages, using, goal, before_qed, after_qed} = goal; + val goal_txt = prop :: map Thm.term_of (Assumption.all_assms_of goal_ctxt); val _ = is_relevant state andalso error "Cannot fork relevant proof"; @@ -1017,13 +1017,13 @@ val result_ctxt = state - |> map_contexts (Variable.declare_term prop) + |> map_contexts (fold Variable.declare_term goal_txt) |> map_goal I (K (statement', messages, using, goal', before_qed, after_qed')) |> fork_proof; val future_thm = result_ctxt |> Future.map (fn (_, x) => ProofContext.get_fact_single (get_context x) (Facts.named this_name)); - val finished_goal = exception_trace (fn () => Goal.future_result goal_ctxt future_thm prop'); + val finished_goal = Goal.future_result goal_ctxt future_thm prop'; val state' = state |> map_goal I (K (statement, messages, using, finished_goal, NONE, after_qed)) @@ -1043,9 +1043,9 @@ local fun future_terminal_proof proof1 proof2 meths int state = - if int orelse is_relevant state orelse not (Goal.future_enabled ()) + if int orelse is_relevant state orelse not (Goal.local_future_enabled ()) then proof1 meths state - else snd (state |> proof2 (fn state' => Future.fork_pri ~1 (fn () => ((), proof1 meths state')))); + else snd (proof2 (fn state' => Future.fork_pri ~1 (fn () => ((), proof1 meths state'))) state); in diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/proof_context.ML --- a/src/Pure/Isar/proof_context.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/proof_context.ML Thu Jul 23 21:13:21 2009 +0200 @@ -36,13 +36,8 @@ val theory_result: (theory -> 'a * theory) -> Proof.context -> 'a * Proof.context val extern_fact: Proof.context -> string -> xstring val pretty_term_abbrev: Proof.context -> term -> Pretty.T - val pretty_thm_aux: Proof.context -> bool -> thm -> Pretty.T - val pretty_thms_aux: Proof.context -> bool -> thm list -> Pretty.T val pretty_fact_aux: Proof.context -> bool -> string * thm list -> Pretty.T - val pretty_thm: Proof.context -> thm -> Pretty.T - val pretty_thms: Proof.context -> thm list -> Pretty.T val pretty_fact: Proof.context -> string * thm list -> Pretty.T - val string_of_thm: Proof.context -> thm -> string val read_typ: Proof.context -> string -> typ val read_typ_syntax: Proof.context -> string -> typ val read_typ_abbrev: Proof.context -> string -> typ @@ -293,31 +288,18 @@ fun pretty_term_abbrev ctxt = Syntax.pretty_term (set_mode mode_abbrev ctxt); -fun pretty_thm_aux ctxt show_status th = - let - val flags = {quote = false, show_hyps = true, show_status = show_status}; - val asms = map Thm.term_of (Assumption.all_assms_of ctxt); - in Display.pretty_thm_aux (Syntax.pp ctxt) flags asms th end; - -fun pretty_thms_aux ctxt flag [th] = pretty_thm_aux ctxt flag th - | pretty_thms_aux ctxt flag ths = - Pretty.blk (0, Pretty.fbreaks (map (pretty_thm_aux ctxt flag) ths)); - fun pretty_fact_name ctxt a = Pretty.block [Pretty.markup (Markup.fact a) [Pretty.str (extern_fact ctxt a)], Pretty.str ":"]; -fun pretty_fact_aux ctxt flag ("", ths) = pretty_thms_aux ctxt flag ths +fun pretty_fact_aux ctxt flag ("", ths) = + Display.pretty_thms_aux ctxt flag ths | pretty_fact_aux ctxt flag (a, [th]) = Pretty.block - [pretty_fact_name ctxt a, Pretty.brk 1, pretty_thm_aux ctxt flag th] + [pretty_fact_name ctxt a, Pretty.brk 1, Display.pretty_thm_aux ctxt flag th] | pretty_fact_aux ctxt flag (a, ths) = Pretty.block - (Pretty.fbreaks (pretty_fact_name ctxt a :: map (pretty_thm_aux ctxt flag) ths)); + (Pretty.fbreaks (pretty_fact_name ctxt a :: map (Display.pretty_thm_aux ctxt flag) ths)); -fun pretty_thm ctxt = pretty_thm_aux ctxt true; -fun pretty_thms ctxt = pretty_thms_aux ctxt true; fun pretty_fact ctxt = pretty_fact_aux ctxt true; -val string_of_thm = Pretty.string_of oo pretty_thm; - (** prepare types **) @@ -779,7 +761,7 @@ fun simult_matches ctxt (t, pats) = (case Seq.pull (Unify.matchers (theory_of ctxt) (map (rpair t) pats)) of NONE => error "Pattern match failed!" - | SOME (env, _) => map (apsnd snd) (Envir.alist_of env)); + | SOME (env, _) => Vartab.fold (fn (v, (_, t)) => cons (v, t)) (Envir.term_env env) []); (* bind_terms *) @@ -1369,7 +1351,7 @@ val suppressed = len - ! prems_limit; val prt_prems = if null prems then [] else [Pretty.big_list "prems:" ((if suppressed <= 0 then [] else [Pretty.str "..."]) @ - map (pretty_thm ctxt) (Library.drop (suppressed, prems)))]; + map (Display.pretty_thm ctxt) (Library.drop (suppressed, prems)))]; in prt_structs @ prt_fixes @ prt_prems end; diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/proof_display.ML --- a/src/Pure/Isar/proof_display.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/proof_display.ML Thu Jul 23 21:13:21 2009 +0200 @@ -35,7 +35,7 @@ let val thy = Thm.theory_of_thm th; val flags = {quote = true, show_hyps = false, show_status = true}; - in Display.pretty_thm_aux (Syntax.pp_global thy) flags [] th end; + in Display.pretty_thm_raw (Syntax.init_pretty_global thy) flags th end; val pp_typ = Pretty.quote oo Syntax.pretty_typ_global; val pp_term = Pretty.quote oo Syntax.pretty_term_global; @@ -68,7 +68,7 @@ fun pretty_rule ctxt s thm = Pretty.block [Pretty.str (s ^ " attempt to solve goal by exported rule:"), - Pretty.fbrk, ProofContext.pretty_thm_aux ctxt false thm]; + Pretty.fbrk, Display.pretty_thm_aux ctxt false thm]; val string_of_rule = Pretty.string_of ooo pretty_rule; diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/runtime.ML --- a/src/Pure/Isar/runtime.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/runtime.ML Thu Jul 23 21:13:21 2009 +0200 @@ -75,7 +75,7 @@ | exn_msg ctxt (exn as TERM (msg, ts)) = raised exn "TERM" (msg :: (if detailed then if_context ctxt Syntax.string_of_term ts else [])) | exn_msg ctxt (exn as THM (msg, i, thms)) = raised exn ("THM " ^ string_of_int i) (msg :: - (if detailed then if_context ctxt ProofContext.string_of_thm thms else [])) + (if detailed then if_context ctxt Display.string_of_thm thms else [])) | exn_msg _ exn = raised exn (General.exnMessage exn) []; in exn_msg NONE e end; diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/theory_target.ML --- a/src/Pure/Isar/theory_target.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Isar/theory_target.ML Thu Jul 23 21:13:21 2009 +0200 @@ -177,7 +177,6 @@ let val b' = Morphism.binding phi b; val rhs' = Morphism.term phi rhs; - val legacy_arg = (b', Term.close_schematic_term (Logic.legacy_varify rhs')); val arg = (b', Term.close_schematic_term rhs'); val similar_body = Type.similar_types (rhs, rhs'); (* FIXME workaround based on educated guess *) @@ -188,7 +187,7 @@ in not (is_class andalso (similar_body orelse class_global)) ? (Context.mapping_result - (Sign.add_abbrev PrintMode.internal tags legacy_arg) + (Sign.add_abbrev PrintMode.internal tags arg) (ProofContext.add_abbrev PrintMode.internal tags arg) #-> (fn (lhs' as Const (d, _), _) => similar_body ? diff -r 63686057cbe8 -r abda97d2deea src/Pure/Isar/toplevel.ML diff -r 63686057cbe8 -r abda97d2deea src/Pure/ML-Systems/exn.ML --- a/src/Pure/ML-Systems/exn.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/ML-Systems/exn.ML Thu Jul 23 21:13:21 2009 +0200 @@ -13,6 +13,8 @@ val release: 'a result -> 'a exception Interrupt exception EXCEPTIONS of exn list + val flatten: exn -> exn list + val flatten_list: exn list -> exn list val release_all: 'a result list -> 'a list val release_first: 'a result list -> 'a list end; @@ -43,19 +45,15 @@ exception Interrupt = Interrupt; exception EXCEPTIONS of exn list; -fun plain_exns (Result _) = [] - | plain_exns (Exn Interrupt) = [] - | plain_exns (Exn (EXCEPTIONS exns)) = List.concat (map (plain_exns o Exn) exns) - | plain_exns (Exn exn) = [exn]; - +fun flatten Interrupt = [] + | flatten (EXCEPTIONS exns) = flatten_list exns + | flatten exn = [exn] +and flatten_list exns = List.concat (map flatten exns); fun release_all results = if List.all (fn Result _ => true | _ => false) results then map (fn Result x => x) results - else - (case List.concat (map plain_exns results) of - [] => raise Interrupt - | exns => raise EXCEPTIONS exns); + else raise EXCEPTIONS (flatten_list (List.mapPartial get_exn results)); fun release_first results = release_all results handle EXCEPTIONS (exn :: _) => reraise exn; diff -r 63686057cbe8 -r abda97d2deea src/Pure/ML-Systems/multithreading_polyml.ML --- a/src/Pure/ML-Systems/multithreading_polyml.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/ML-Systems/multithreading_polyml.ML Thu Jul 23 21:13:21 2009 +0200 @@ -92,10 +92,12 @@ val _ = Thread.setAttributes orig_atts; in Exn.release result end; -fun interruptible f = with_attributes regular_interrupts (fn _ => f); +fun interruptible f = + with_attributes regular_interrupts (fn _ => fn x => f x); fun uninterruptible f = - with_attributes no_interrupts (fn atts => f (fn g => with_attributes atts (fn _ => g))); + with_attributes no_interrupts (fn atts => fn x => + f (fn g => with_attributes atts (fn _ => fn y => g y)) x); (* execution with time limit *) diff -r 63686057cbe8 -r abda97d2deea src/Pure/Proof/extraction.ML --- a/src/Pure/Proof/extraction.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Proof/extraction.ML Thu Jul 23 21:13:21 2009 +0200 @@ -103,11 +103,10 @@ fun ren t = the_default t (Term.rename_abs tm1 tm t); val inc = Logic.incr_indexes ([], maxidx_of_term tm + 1); val env as (Tenv, tenv) = Pattern.match thy (inc tm1, tm) (Vartab.empty, Vartab.empty); - val prems' = map (pairself (Envir.subst_vars env o inc o ren)) prems; + val prems' = map (pairself (Envir.subst_term env o inc o ren)) prems; val env' = Envir.Envir - {maxidx = Library.foldl Int.max - (~1, map (Int.max o pairself maxidx_of_term) prems'), - iTs = Tenv, asol = tenv}; + {maxidx = fold (fn (t, u) => Term.maxidx_term t #> Term.maxidx_term u) prems' ~1, + tenv = tenv, tyenv = Tenv}; val env'' = fold (Pattern.unify thy o pairself (lookup rew)) prems' env'; in SOME (Envir.norm_term env'' (inc (ren tm2))) end handle Pattern.MATCH => NONE | Pattern.Unif => NONE) diff -r 63686057cbe8 -r abda97d2deea src/Pure/Proof/reconstruct.ML --- a/src/Pure/Proof/reconstruct.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Proof/reconstruct.ML Thu Jul 23 21:13:21 2009 +0200 @@ -35,12 +35,6 @@ fun forall_intr_vfs_prf prop prf = fold_rev forall_intr_prf (vars_of prop @ frees_of prop) prf; -fun merge_envs (Envir.Envir {asol=asol1, iTs=iTs1, maxidx=maxidx1}) - (Envir.Envir {asol=asol2, iTs=iTs2, maxidx=maxidx2}) = - Envir.Envir {asol=Vartab.merge (op =) (asol1, asol2), - iTs=Vartab.merge (op =) (iTs1, iTs2), - maxidx=Int.max (maxidx1, maxidx2)}; - (**** generate constraints for proof term ****) @@ -48,31 +42,32 @@ let val (env', v) = Envir.genvar "a" (env, rev Ts ---> T) in (env', list_comb (v, map Bound (length Ts - 1 downto 0))) end; -fun mk_tvar (Envir.Envir {iTs, asol, maxidx}, s) = - (Envir.Envir {iTs = iTs, asol = asol, maxidx = maxidx+1}, - TVar (("'t", maxidx+1), s)); +fun mk_tvar (Envir.Envir {maxidx, tenv, tyenv}, s) = + (Envir.Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv}, + TVar (("'t", maxidx + 1), s)); val mk_abs = fold (fn T => fn u => Abs ("", T, u)); fun unifyT thy env T U = let - val Envir.Envir {asol, iTs, maxidx} = env; - val (iTs', maxidx') = Sign.typ_unify thy (T, U) (iTs, maxidx) - in Envir.Envir {asol=asol, iTs=iTs', maxidx=maxidx'} end - handle Type.TUNIFY => error ("Non-unifiable types:\n" ^ - Syntax.string_of_typ_global thy T ^ "\n\n" ^ Syntax.string_of_typ_global thy U); + val Envir.Envir {maxidx, tenv, tyenv} = env; + val (tyenv', maxidx') = Sign.typ_unify thy (T, U) (tyenv, maxidx); + in Envir.Envir {maxidx = maxidx', tenv = tenv, tyenv = tyenv'} end; -fun chaseT (env as Envir.Envir {iTs, ...}) (T as TVar ixnS) = - (case Type.lookup iTs ixnS of NONE => T | SOME T' => chaseT env T') +fun chaseT env (T as TVar v) = + (case Type.lookup (Envir.type_env env) v of + NONE => T + | SOME T' => chaseT env T') | chaseT _ T = T; -fun infer_type thy (env as Envir.Envir {maxidx, asol, iTs}) Ts vTs - (t as Const (s, T)) = if T = dummyT then (case Sign.const_type thy s of +fun infer_type thy (env as Envir.Envir {maxidx, tenv, tyenv}) Ts vTs + (t as Const (s, T)) = if T = dummyT then + (case Sign.const_type thy s of NONE => error ("reconstruct_proof: No such constant: " ^ quote s) | SOME T => let val T' = Type.strip_sorts (Logic.incr_tvar (maxidx + 1) T) in (Const (s, T'), T', vTs, - Envir.Envir {maxidx = maxidx + 1, asol = asol, iTs = iTs}) + Envir.Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv}) end) else (t, T, vTs, env) | infer_type thy env Ts vTs (t as Free (s, T)) = @@ -236,21 +231,23 @@ fun upd_constrs env cs = let - val Envir.Envir {asol, iTs, ...} = env; + val tenv = Envir.term_env env; + val tyenv = Envir.type_env env; val dom = [] - |> Vartab.fold (cons o #1) asol - |> Vartab.fold (cons o #1) iTs; + |> Vartab.fold (cons o #1) tenv + |> Vartab.fold (cons o #1) tyenv; val vran = [] - |> Vartab.fold (Term.add_var_names o #2 o #2) asol - |> Vartab.fold (Term.add_tvar_namesT o #2 o #2) iTs; + |> Vartab.fold (Term.add_var_names o #2 o #2) tenv + |> Vartab.fold (Term.add_tvar_namesT o #2 o #2) tyenv; fun check_cs [] = [] - | check_cs ((u, p, vs)::ps) = - let val vs' = subtract (op =) dom vs; - in if vs = vs' then (u, p, vs)::check_cs ps - else (true, p, fold (insert (op =)) vs' vran)::check_cs ps - end + | check_cs ((u, p, vs) :: ps) = + let val vs' = subtract (op =) dom vs in + if vs = vs' then (u, p, vs) :: check_cs ps + else (true, p, fold (insert op =) vs' vran) :: check_cs ps + end; in check_cs cs end; + (**** solution of constraints ****) fun solve _ [] bigenv = bigenv @@ -258,7 +255,7 @@ let fun search env [] = error ("Unsolvable constraints:\n" ^ Pretty.string_of (Pretty.chunks (map (fn (_, p, _) => - Display.pretty_flexpair (Syntax.pp_global thy) (pairself + Display_Goal.pretty_flexpair (Syntax.init_pretty_global thy) (pairself (Envir.norm_term bigenv) p)) cs))) | search env ((u, p as (t1, t2), vs)::ps) = if u then @@ -280,7 +277,7 @@ val Envir.Envir {maxidx, ...} = bigenv; val (env, cs') = search (Envir.empty maxidx) cs; in - solve thy (upd_constrs env cs') (merge_envs bigenv env) + solve thy (upd_constrs env cs') (Envir.merge (bigenv, env)) end; @@ -353,8 +350,6 @@ | (b, SOME prop') => a = b andalso prop = prop') thms) then (maxidx, prfs, prf) else let - fun inc i = - map_proof_terms (Logic.incr_indexes ([], i)) (Logic.incr_tvar i); val (maxidx', prf, prfs') = (case AList.lookup (op =) prfs (a, prop) of NONE => @@ -365,11 +360,11 @@ (reconstruct_proof thy prop (join_proof body)); val (maxidx', prfs', prf) = expand (maxidx_proof prf' ~1) prfs prf' - in (maxidx' + maxidx + 1, inc (maxidx + 1) prf, + in (maxidx' + maxidx + 1, incr_indexes (maxidx + 1) prf, ((a, prop), (maxidx', prf)) :: prfs') end | SOME (maxidx', prf) => (maxidx' + maxidx + 1, - inc (maxidx + 1) prf, prfs)); + incr_indexes (maxidx + 1) prf, prfs)); val tfrees = OldTerm.term_tfrees prop; val tye = map (fn ((s, j), _) => (s, maxidx + 1 + j)) (OldTerm.term_tvars prop) @ map (rpair ~1 o fst) tfrees ~~ Ts; diff -r 63686057cbe8 -r abda97d2deea src/Pure/ProofGeneral/preferences.ML --- a/src/Pure/ProofGeneral/preferences.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/ProofGeneral/preferences.ML Thu Jul 23 21:13:21 2009 +0200 @@ -84,6 +84,12 @@ 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) (); +val parallel_proof_pref = + let + fun get () = PgipTypes.bool_to_pgstring (! Goal.parallel_proofs >= 1); + fun set s = Goal.parallel_proofs := (if PgipTypes.read_pgipbool s then 1 else 0); + in mkpref get set PgipTypes.Pgipbool "parallel-proofs" "Check proofs in parallel" end; + val thm_depsN = "thm_deps"; val thm_deps_pref = let @@ -171,9 +177,7 @@ nat_pref Multithreading.max_threads "max-threads" "Maximum number of threads", - bool_pref Goal.parallel_proofs - "parallel-proofs" - "Check proofs in parallel"]; + parallel_proof_pref]; val pure_preferences = [(category_display, display_preferences), diff -r 63686057cbe8 -r abda97d2deea src/Pure/ProofGeneral/proof_general_pgip.ML --- a/src/Pure/ProofGeneral/proof_general_pgip.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/ProofGeneral/proof_general_pgip.ML Thu Jul 23 21:13:21 2009 +0200 @@ -655,11 +655,8 @@ text=[XML.Elem("pgml",[], maps YXML.parse_body strings)] }) - fun string_of_thm th = Pretty.string_of (Display.pretty_thm_aux - (Syntax.pp_global (Thm.theory_of_thm th)) - {quote = false, show_hyps = false, show_status = true} [] th) - - fun strings_of_thm (thy, name) = map string_of_thm (PureThy.get_thms thy name) + fun strings_of_thm (thy, name) = + map (Display.string_of_thm_global thy) (PureThy.get_thms thy name) val string_of_thy = Pretty.string_of o ProofDisplay.pretty_full_theory false in diff -r 63686057cbe8 -r abda97d2deea src/Pure/ROOT.ML --- a/src/Pure/ROOT.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/ROOT.ML Thu Jul 23 21:13:21 2009 +0200 @@ -33,11 +33,12 @@ use "ML/ml_lex.ML"; use "ML/ml_parse.ML"; use "General/secure.ML"; -(*----- basic ML bootstrap finished -----*) +(*^^^^^ end of basic ML bootstrap ^^^^^*) use "General/integer.ML"; use "General/stack.ML"; use "General/queue.ML"; use "General/heap.ML"; +use "General/same.ML"; use "General/ord_list.ML"; use "General/balanced_tree.ML"; use "General/graph.ML"; @@ -114,17 +115,18 @@ use "more_thm.ML"; use "facts.ML"; use "pure_thy.ML"; -use "display.ML"; use "drule.ML"; use "morphism.ML"; use "variable.ML"; use "conv.ML"; +use "display_goal.ML"; use "tctical.ML"; use "search.ML"; use "tactic.ML"; use "meta_simplifier.ML"; use "conjunction.ML"; use "assumption.ML"; +use "display.ML"; use "goal.ML"; use "axclass.ML"; diff -r 63686057cbe8 -r abda97d2deea src/Pure/Syntax/syn_trans.ML --- a/src/Pure/Syntax/syn_trans.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Syntax/syn_trans.ML Thu Jul 23 21:13:21 2009 +0200 @@ -8,6 +8,8 @@ sig val eta_contract: bool ref val atomic_abs_tr': string * typ * term -> term * term + val preserve_binder_abs_tr': string -> string -> string * (term list -> term) + val preserve_binder_abs2_tr': string -> string -> string * (term list -> term) val mk_binder_tr: string * string -> string * (term list -> term) val mk_binder_tr': string * string -> string * (term list -> term) val dependent_tr': string * string -> term list -> term @@ -309,6 +311,14 @@ ([], _) => raise Ast.AST ("abs_ast_tr'", asts) | (xs, body) => Ast.Appl [Ast.Constant "_lambda", Ast.fold_ast "_pttrns" xs, body]); +fun preserve_binder_abs_tr' name syn = (name, fn (Abs abs :: ts) => + let val (x, t) = atomic_abs_tr' abs + in list_comb (Lexicon.const syn $ x $ t, ts) end); + +fun preserve_binder_abs2_tr' name syn = (name, fn (A :: Abs abs :: ts) => + let val (x, t) = atomic_abs_tr' abs + in list_comb (Lexicon.const syn $ x $ A $ t, ts) end); + (* binder *) diff -r 63686057cbe8 -r abda97d2deea src/Pure/System/session.ML --- a/src/Pure/System/session.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/System/session.ML Thu Jul 23 21:13:21 2009 +0200 @@ -11,7 +11,7 @@ val welcome: unit -> string val use_dir: string -> string -> bool -> string list -> bool -> bool -> string -> bool -> string list -> string -> string -> bool * string -> - string -> int -> bool -> bool -> int -> int -> bool -> unit + string -> int -> bool -> bool -> int -> int -> int -> unit val finish: unit -> unit end; diff -r 63686057cbe8 -r abda97d2deea src/Pure/Thy/thy_info.ML --- a/src/Pure/Thy/thy_info.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Thy/thy_info.ML Thu Jul 23 21:13:21 2009 +0200 @@ -23,6 +23,7 @@ val get_parents: string -> string list val touch_thy: string -> unit val touch_child_thys: string -> unit + val thy_ord: theory * theory -> order val remove_thy: string -> unit val kill_thy: string -> unit val provide_file: Path.T -> string -> unit @@ -33,7 +34,6 @@ val use_thys: string list -> unit val use_thy: string -> unit val time_use_thy: string -> unit - val thy_ord: theory * theory -> order val begin_theory: string -> string list -> (Path.T * bool) list -> bool -> theory val end_theory: theory -> unit val register_thy: string -> unit @@ -231,17 +231,36 @@ end; -(* manage pending proofs *) +(* management data *) + +structure Management_Data = TheoryDataFun +( + type T = + Task_Queue.group option * (*worker thread group*) + int; (*abstract update time*) + val empty = (NONE, 0); + val copy = I; + fun extend _ = empty; + fun merge _ _ = empty; +); + +val thy_ord = int_ord o pairself (#2 o Management_Data.get); + + +(* pending proofs *) fun join_thy name = (case lookup_theory name of - NONE => [] + NONE => () | SOME thy => PureThy.join_proofs thy); fun cancel_thy name = (case lookup_theory name of NONE => () - | SOME thy => PureThy.cancel_proofs thy); + | SOME thy => + (case #1 (Management_Data.get thy) of + NONE => () + | SOME group => Future.cancel_group group)); (* remove theory *) @@ -350,7 +369,7 @@ val tasks = Graph.topological_order task_graph |> map_filter (fn name => (case Graph.get_node task_graph name of Task body => SOME (name, body) | _ => NONE)); - val par_proofs = ! parallel_proofs; + val par_proofs = ! parallel_proofs >= 1; fun fork (name, body) tab = let @@ -374,8 +393,7 @@ val exns = tasks |> maps (fn (name, _) => let val after_load = Future.join (the (Symtab.lookup futures name)); - val proof_exns = join_thy name; - val _ = null proof_exns orelse raise Exn.EXCEPTIONS proof_exns; + val _ = join_thy name; val _ = after_load (); in [] end handle exn => (kill_thy name; [exn])); @@ -509,20 +527,6 @@ end; -(* update_time *) - -structure UpdateTime = TheoryDataFun -( - type T = int; - val empty = 0; - val copy = I; - fun extend _ = empty; - fun merge _ _ = empty; -); - -val thy_ord = int_ord o pairself UpdateTime.get; - - (* begin / end theory *) fun begin_theory name parents uses int = @@ -542,7 +546,7 @@ val update_time = (case deps of NONE => 0 | SOME {update_time, ...} => update_time); val update_time = if update_time > 0 then update_time else serial (); val theory' = theory - |> UpdateTime.put update_time + |> Management_Data.put (Future.worker_group (), update_time) |> Present.begin_theory update_time dir uses; val uses_now = map_filter (fn (x, true) => SOME x | _ => NONE) uses; @@ -569,7 +573,7 @@ val _ = check_unfinished error name; val _ = touch_thy name; val master = #master (ThyLoad.deps_thy Path.current name); - val upd_time = UpdateTime.get thy; + val upd_time = #2 (Management_Data.get thy); in CRITICAL (fn () => (change_deps name (Option.map diff -r 63686057cbe8 -r abda97d2deea src/Pure/Tools/find_theorems.ML --- a/src/Pure/Tools/find_theorems.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/Tools/find_theorems.ML Thu Jul 23 21:13:21 2009 +0200 @@ -207,7 +207,7 @@ fun filter_simp ctxt t (_, thm) = let - val mksimps = Simplifier.mksimps (Simplifier.local_simpset_of ctxt); + val mksimps = Simplifier.mksimps (simpset_of ctxt); val extract_simp = (map Thm.full_prop_of o mksimps, #1 o Logic.dest_equals o Logic.strip_imp_concl); val ss = is_matching_thm false extract_simp ctxt false t thm; @@ -408,7 +408,7 @@ fun pretty_thm ctxt (thmref, thm) = Pretty.block [Pretty.str (Facts.string_of_ref thmref), Pretty.str ":", Pretty.brk 1, - ProofContext.pretty_thm ctxt thm]; + Display.pretty_thm ctxt thm]; fun print_theorems ctxt opt_goal opt_limit rem_dups raw_criteria = let diff -r 63686057cbe8 -r abda97d2deea src/Pure/codegen.ML --- a/src/Pure/codegen.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/codegen.ML Thu Jul 23 21:13:21 2009 +0200 @@ -64,7 +64,7 @@ val new_name: term -> string -> string val if_library: 'a -> 'a -> 'a val get_defn: theory -> deftab -> string -> typ -> - ((typ * (string * (term list * term))) * int option) option + ((typ * (string * thm)) * int option) option val is_instance: typ -> typ -> bool val parens: Pretty.T -> Pretty.T val mk_app: bool -> Pretty.T -> Pretty.T list -> Pretty.T @@ -85,7 +85,9 @@ val num_args_of: 'a mixfix list -> int val replace_quotes: 'b list -> 'a mixfix list -> 'b mixfix list val mk_deftab: theory -> deftab + val map_unfold: (simpset -> simpset) -> theory -> theory val add_unfold: thm -> theory -> theory + val del_unfold: thm -> theory -> theory val get_node: codegr -> string -> node val add_edge: string * string -> codegr -> codegr @@ -145,8 +147,7 @@ type deftab = (typ * (* type of constant *) (string * (* name of theory containing definition of constant *) - (term list * (* parameters *) - term))) (* right-hand side *) + thm)) (* definition theorem *) list Symtab.table; (* code dependency graph *) @@ -296,13 +297,9 @@ fun merge _ = merge_ss; ); -fun add_unfold eqn = - let - val ctxt = ProofContext.init (Thm.theory_of_thm eqn); - val eqn' = LocalDefs.meta_rewrite_rule ctxt eqn; - in - UnfoldData.map (fn ss => ss addsimps [eqn']) - end; +val map_unfold = UnfoldData.map; +val add_unfold = map_unfold o MetaSimplifier.add_simp; +val del_unfold = map_unfold o MetaSimplifier.del_simp; fun unfold_preprocessor thy = let val ss = Simplifier.theory_context thy (UnfoldData.get thy) @@ -492,35 +489,43 @@ fun get_aux_code xs = map_filter (fn (m, code) => if m = "" orelse member (op =) (!mode) m then SOME code else NONE) xs; +fun dest_prim_def t = + let + val (lhs, rhs) = Logic.dest_equals t; + val (c, args) = strip_comb lhs; + val (s, T) = dest_Const c + in if forall is_Var args then SOME (s, (T, (args, rhs))) else NONE + end handle TERM _ => NONE; + fun mk_deftab thy = let val axmss = map (fn thy' => (Context.theory_name thy', Theory.axiom_table thy')) (thy :: Theory.ancestors_of thy); fun prep_def def = (case preprocess thy [def] of [def'] => prop_of def' | _ => error "mk_deftab: bad preprocessor"); - fun dest t = - let - val (lhs, rhs) = Logic.dest_equals t; - val (c, args) = strip_comb lhs; - val (s, T) = dest_Const c - in if forall is_Var args then SOME (s, (T, (args, rhs))) else NONE - end handle TERM _ => NONE; - fun add_def thyname (name, t) = (case dest t of + fun add_def thyname (name, t) = (case dest_prim_def t of NONE => I - | SOME _ => (case dest (prep_def (Thm.axiom thy name)) of - NONE => I - | SOME (s, (T, (args, rhs))) => Symtab.map_default (s, []) - (cons (T, (thyname, split_last (rename_terms (args @ [rhs]))))))) + | SOME (s, (T, _)) => Symtab.map_default (s, []) + (cons (T, (thyname, Thm.axiom thy name)))); in fold (fn (thyname, axms) => Symtab.fold (add_def thyname) axms) axmss Symtab.empty end; +fun prep_prim_def thy thm = + let + val prop = case preprocess thy [thm] + of [thm'] => Thm.prop_of thm' + | _ => error "mk_deftab: bad preprocessor" + in ((Option.map o apsnd o apsnd) + (fn (args, rhs) => split_last (rename_terms (args @ [rhs]))) o dest_prim_def) prop + end; + fun get_defn thy defs s T = (case Symtab.lookup defs s of NONE => NONE | SOME ds => let val i = find_index (is_instance T o fst) ds in if i >= 0 then - SOME (List.nth (ds, i), if length ds = 1 then NONE else SOME i) + SOME (nth ds i, if length ds = 1 then NONE else SOME i) else NONE end); @@ -657,8 +662,8 @@ end | NONE => (case get_defn thy defs s T of NONE => NONE - | SOME ((U, (thyname, (args, rhs))), k) => - let + | SOME ((U, (thyname, thm)), k) => (case prep_prim_def thy thm + of SOME (_, (_, (args, rhs))) => let val module' = if_library thyname module; val suffix = (case k of NONE => "" | SOME i => " def" ^ string_of_int i); val node_id = s ^ suffix; @@ -688,7 +693,8 @@ [str " =", Pretty.brk 1, p', str ";"])) ^ "\n\n")) gr4) end | SOME _ => (p, add_edge (node_id, dep) gr')) - end)) + end + | NONE => NONE))) | Abs _ => let diff -r 63686057cbe8 -r abda97d2deea src/Pure/defs.ML --- a/src/Pure/defs.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/defs.ML Thu Jul 23 21:13:21 2009 +0200 @@ -10,10 +10,10 @@ val pretty_const: Pretty.pp -> string * typ list -> Pretty.T val plain_args: typ list -> bool type T + val all_specifications_of: T -> (string * {is_def: bool, name: string, + lhs: typ list, rhs: (string * typ list) list} list) list val specifications_of: T -> string -> {is_def: bool, name: string, lhs: typ list, rhs: (string * typ list) list} list - val all_specifications_of: T -> (string * {is_def: bool, name: string, - lhs: typ list, rhs: (string * typ list) list} list) list val dest: T -> {restricts: ((string * typ list) * string) list, reducts: ((string * typ list) * (string * typ list) list) list} @@ -46,7 +46,7 @@ handle Type.TUNIFY => true); fun match_args (Ts, Us) = - Option.map Envir.typ_subst_TVars + Option.map Envir.subst_type (SOME (Type.raw_matches (Ts, Us) Vartab.empty) handle Type.TYPE_MATCH => NONE); @@ -75,9 +75,11 @@ SOME (def: def) => which def | NONE => []); +fun all_specifications_of (Defs defs) = + (map o apsnd) (map snd o Inttab.dest o #specs) (Symtab.dest defs); + fun specifications_of (Defs defs) = lookup_list (map snd o Inttab.dest o #specs) defs; -fun all_specifications_of (Defs defs) = - ((map o apsnd) (map snd o Inttab.dest o #specs) o Symtab.dest) defs; + val restricts_of = lookup_list #restricts; val reducts_of = lookup_list #reducts; diff -r 63686057cbe8 -r abda97d2deea src/Pure/display.ML --- a/src/Pure/display.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/display.ML Thu Jul 23 21:13:21 2009 +0200 @@ -1,34 +1,32 @@ (* Title: Pure/display.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1993 University of Cambridge + Author: Makarius -Printing of theorems, goals, results etc. +Printing of theorems, results etc. *) signature BASIC_DISPLAY = sig val goals_limit: int ref + val show_consts: bool ref val show_hyps: bool ref val show_tags: bool ref - val show_consts: bool ref end; signature DISPLAY = sig include BASIC_DISPLAY - val pretty_flexpair: Pretty.pp -> term * term -> Pretty.T - val pretty_thm_aux: Pretty.pp -> {quote: bool, show_hyps: bool, show_status: bool} -> - term list -> thm -> Pretty.T - val pretty_thm: thm -> Pretty.T - val string_of_thm: thm -> string - val pretty_thms: thm list -> Pretty.T - val pretty_thm_sg: theory -> thm -> Pretty.T - val pretty_thms_sg: theory -> thm list -> Pretty.T - val print_thm: thm -> unit - val print_thms: thm list -> unit - val prth: thm -> thm - val prthq: thm Seq.seq -> thm Seq.seq - val prths: thm list -> thm list + val pretty_thm_raw: Proof.context -> {quote: bool, show_hyps: bool, show_status: bool} -> + thm -> Pretty.T + val pretty_thm_aux: Proof.context -> bool -> thm -> Pretty.T + val pretty_thm: Proof.context -> thm -> Pretty.T + val pretty_thm_global: theory -> thm -> Pretty.T + val pretty_thm_without_context: thm -> Pretty.T + val string_of_thm: Proof.context -> thm -> string + val string_of_thm_global: theory -> thm -> string + val string_of_thm_without_context: thm -> string + val pretty_thms_aux: Proof.context -> bool -> thm list -> Pretty.T + val pretty_thms: Proof.context -> thm list -> Pretty.T val pretty_ctyp: ctyp -> Pretty.T val string_of_ctyp: ctyp -> string val print_ctyp: ctyp -> unit @@ -37,27 +35,26 @@ val print_cterm: cterm -> unit val print_syntax: theory -> unit val pretty_full_theory: bool -> theory -> Pretty.T list - val pretty_goals_aux: Pretty.pp -> Markup.T -> bool * bool -> int -> thm -> Pretty.T list - val pretty_goals: int -> thm -> Pretty.T list - val print_goals: int -> thm -> unit end; structure Display: DISPLAY = struct +(** options **) + +val goals_limit = Display_Goal.goals_limit; +val show_consts = Display_Goal.show_consts; + +val show_hyps = ref false; (*false: print meta-hypotheses as dots*) +val show_tags = ref false; (*false: suppress tags*) + + (** print thm **) -val goals_limit = ref 10; (*max number of goals to print*) -val show_hyps = ref false; (*false: print meta-hypotheses as dots*) -val show_tags = ref false; (*false: suppress tags*) - fun pretty_tag (name, arg) = Pretty.strs [name, quote arg]; val pretty_tags = Pretty.list "[" "]" o map pretty_tag; -fun pretty_flexpair pp (t, u) = Pretty.block - [Pretty.term pp t, Pretty.str " =?=", Pretty.brk 1, Pretty.term pp u]; - fun display_status false _ = "" | display_status true th = let @@ -71,7 +68,7 @@ else "" end; -fun pretty_thm_aux pp {quote, show_hyps = show_hyps', show_status} asms raw_th = +fun pretty_thm_raw ctxt {quote, show_hyps = show_hyps', show_status} raw_th = let val th = Thm.strip_shyps raw_th; val {hyps, tpairs, prop, ...} = Thm.rep_thm th; @@ -79,8 +76,9 @@ val tags = Thm.get_tags th; val q = if quote then Pretty.quote else I; - val prt_term = q o Pretty.term pp; + val prt_term = q o Syntax.pretty_term ctxt; + val asms = map Thm.term_of (Assumption.all_assms_of ctxt); val hyps' = if ! show_hyps then hyps else subtract (op aconv) asms hyps; val status = display_status show_status th; @@ -89,8 +87,8 @@ if hlen = 0 andalso status = "" then [] else if ! show_hyps orelse show_hyps' then [Pretty.brk 2, Pretty.list "[" "]" - (map (q o pretty_flexpair pp) tpairs @ map prt_term hyps' @ - map (Pretty.sort pp) xshyps @ + (map (q o Display_Goal.pretty_flexpair ctxt) tpairs @ map prt_term hyps' @ + map (Syntax.pretty_sort ctxt) xshyps @ (if status = "" then [] else [Pretty.str status]))] else [Pretty.brk 2, Pretty.str ("[" ^ implode (replicate hlen ".") ^ status ^ "]")]; val tsymbs = @@ -98,27 +96,29 @@ else [Pretty.brk 1, pretty_tags tags]; in Pretty.block (prt_term prop :: (hsymbs @ tsymbs)) end; -fun pretty_thm th = - pretty_thm_aux (Syntax.pp_global (Thm.theory_of_thm th)) - {quote = true, show_hyps = false, show_status = true} [] th; +fun pretty_thm_aux ctxt show_status = + pretty_thm_raw ctxt {quote = false, show_hyps = true, show_status = show_status}; -val string_of_thm = Pretty.string_of o pretty_thm; +fun pretty_thm ctxt = pretty_thm_aux ctxt true; -fun pretty_thms [th] = pretty_thm th - | pretty_thms ths = Pretty.block (Pretty.fbreaks (map pretty_thm ths)); +fun pretty_thm_global thy = + pretty_thm_raw (Syntax.init_pretty_global thy) + {quote = false, show_hyps = false, show_status = true}; -val pretty_thm_sg = pretty_thm oo Thm.transfer; -val pretty_thms_sg = pretty_thms oo (map o Thm.transfer); +fun pretty_thm_without_context th = pretty_thm_global (Thm.theory_of_thm th) th; + +val string_of_thm = Pretty.string_of oo pretty_thm; +val string_of_thm_global = Pretty.string_of oo pretty_thm_global; +val string_of_thm_without_context = Pretty.string_of o pretty_thm_without_context; -(* top-level commands for printing theorems *) +(* multiple theorems *) -val print_thm = Pretty.writeln o pretty_thm; -val print_thms = Pretty.writeln o pretty_thms; +fun pretty_thms_aux ctxt flag [th] = pretty_thm_aux ctxt flag th + | pretty_thms_aux ctxt flag ths = + Pretty.blk (0, Pretty.fbreaks (map (pretty_thm_aux ctxt flag) ths)); -fun prth th = (print_thm th; th); -fun prthq thq = (Seq.print (K print_thm) 100000 thq; thq); -fun prths ths = (prthq (Seq.of_list ths); ths); +fun pretty_thms ctxt = pretty_thms_aux ctxt true; (* other printing commands *) @@ -242,109 +242,7 @@ Pretty.big_list "pattern restrictions:" (map pretty_restrict rests)]] end; - - -(** print_goals **) - -(* print_goals etc. *) - -val show_consts = ref false; (*true: show consts with types in proof state output*) - - -(*print thm A1,...,An/B in "goal style" -- premises as numbered subgoals*) - -local - -fun ins_entry (x, y) = - AList.default (op =) (x, []) #> - AList.map_entry (op =) x (insert (op =) y); - -val add_consts = Term.fold_aterms - (fn Const (c, T) => ins_entry (T, (c, T)) - | _ => I); - -val add_vars = Term.fold_aterms - (fn Free (x, T) => ins_entry (T, (x, ~1)) - | Var (xi, T) => ins_entry (T, xi) - | _ => I); - -val add_varsT = Term.fold_atyps - (fn TFree (x, S) => ins_entry (S, (x, ~1)) - | TVar (xi, S) => ins_entry (S, xi) - | _ => I); - -fun sort_idxs vs = map (apsnd (sort (prod_ord string_ord int_ord))) vs; -fun sort_cnsts cs = map (apsnd (sort_wrt fst)) cs; - -fun consts_of t = sort_cnsts (add_consts t []); -fun vars_of t = sort_idxs (add_vars t []); -fun varsT_of t = rev (sort_idxs (Term.fold_types add_varsT t [])); - -in - -fun pretty_goals_aux pp markup (msg, main) maxgoals state = - let - fun prt_atoms prt prtT (X, xs) = Pretty.block - [Pretty.block (Pretty.commas (map prt xs)), Pretty.str " ::", - Pretty.brk 1, prtT X]; - - fun prt_var (x, ~1) = Pretty.term pp (Syntax.free x) - | prt_var xi = Pretty.term pp (Syntax.var xi); - - fun prt_varT (x, ~1) = Pretty.typ pp (TFree (x, [])) - | prt_varT xi = Pretty.typ pp (TVar (xi, [])); - - val prt_consts = prt_atoms (Pretty.term pp o Const) (Pretty.typ pp); - val prt_vars = prt_atoms prt_var (Pretty.typ pp); - val prt_varsT = prt_atoms prt_varT (Pretty.sort pp); - - - fun pretty_list _ _ [] = [] - | pretty_list name prt lst = [Pretty.big_list name (map prt lst)]; - - fun pretty_subgoal (n, A) = Pretty.markup markup - [Pretty.str (" " ^ string_of_int n ^ ". "), Pretty.term pp A]; - fun pretty_subgoals As = map pretty_subgoal (1 upto length As ~~ As); - - val pretty_ffpairs = pretty_list "flex-flex pairs:" (pretty_flexpair pp); - - val pretty_consts = pretty_list "constants:" prt_consts o consts_of; - val pretty_vars = pretty_list "variables:" prt_vars o vars_of; - val pretty_varsT = pretty_list "type variables:" prt_varsT o varsT_of; - - - val {prop, tpairs, ...} = Thm.rep_thm state; - val (As, B) = Logic.strip_horn prop; - val ngoals = length As; - - fun pretty_gs (types, sorts) = - (if main then [Pretty.term pp B] else []) @ - (if ngoals = 0 then [Pretty.str "No subgoals!"] - else if ngoals > maxgoals then - pretty_subgoals (Library.take (maxgoals, As)) @ - (if msg then [Pretty.str ("A total of " ^ string_of_int ngoals ^ " subgoals...")] - else []) - else pretty_subgoals As) @ - pretty_ffpairs tpairs @ - (if ! show_consts then pretty_consts prop else []) @ - (if types then pretty_vars prop else []) @ - (if sorts then pretty_varsT prop else []); - in - setmp show_no_free_types true - (setmp show_types (! show_types orelse ! show_sorts orelse ! show_all_types) - (setmp show_sorts false pretty_gs)) - (! show_types orelse ! show_sorts orelse ! show_all_types, ! show_sorts) - end; - -fun pretty_goals n th = - pretty_goals_aux (Syntax.pp_global (Thm.theory_of_thm th)) Markup.none (true, true) n th; - -val print_goals = (Pretty.writeln o Pretty.chunks) oo pretty_goals; - end; - -end; - -structure BasicDisplay: BASIC_DISPLAY = Display; -open BasicDisplay; +structure Basic_Display: BASIC_DISPLAY = Display; +open Basic_Display; diff -r 63686057cbe8 -r abda97d2deea src/Pure/display_goal.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/display_goal.ML Thu Jul 23 21:13:21 2009 +0200 @@ -0,0 +1,126 @@ +(* Title: Pure/display_goal.ML + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Author: Makarius + +Display tactical goal state. +*) + +signature DISPLAY_GOAL = +sig + val goals_limit: int ref + val show_consts: bool ref + val pretty_flexpair: Proof.context -> term * term -> Pretty.T + val pretty_goals: Proof.context -> Markup.T -> bool * bool -> int -> thm -> Pretty.T list + val pretty_goals_without_context: int -> thm -> Pretty.T list + val print_goals_without_context: int -> thm -> unit +end; + +structure Display_Goal: DISPLAY_GOAL = +struct + +val goals_limit = ref 10; (*max number of goals to print*) +val show_consts = ref false; (*true: show consts with types in proof state output*) + +fun pretty_flexpair ctxt (t, u) = Pretty.block + [Syntax.pretty_term ctxt t, Pretty.str " =?=", Pretty.brk 1, Syntax.pretty_term ctxt u]; + + +(*print thm A1,...,An/B in "goal style" -- premises as numbered subgoals*) + +local + +fun ins_entry (x, y) = + AList.default (op =) (x, []) #> + AList.map_entry (op =) x (insert (op =) y); + +val add_consts = Term.fold_aterms + (fn Const (c, T) => ins_entry (T, (c, T)) + | _ => I); + +val add_vars = Term.fold_aterms + (fn Free (x, T) => ins_entry (T, (x, ~1)) + | Var (xi, T) => ins_entry (T, xi) + | _ => I); + +val add_varsT = Term.fold_atyps + (fn TFree (x, S) => ins_entry (S, (x, ~1)) + | TVar (xi, S) => ins_entry (S, xi) + | _ => I); + +fun sort_idxs vs = map (apsnd (sort (prod_ord string_ord int_ord))) vs; +fun sort_cnsts cs = map (apsnd (sort_wrt fst)) cs; + +fun consts_of t = sort_cnsts (add_consts t []); +fun vars_of t = sort_idxs (add_vars t []); +fun varsT_of t = rev (sort_idxs (Term.fold_types add_varsT t [])); + +in + +fun pretty_goals ctxt markup (msg, main) maxgoals state = + let + val prt_sort = Syntax.pretty_sort ctxt; + val prt_typ = Syntax.pretty_typ ctxt; + val prt_term = Syntax.pretty_term ctxt; + + fun prt_atoms prt prtT (X, xs) = Pretty.block + [Pretty.block (Pretty.commas (map prt xs)), Pretty.str " ::", + Pretty.brk 1, prtT X]; + + fun prt_var (x, ~1) = prt_term (Syntax.free x) + | prt_var xi = prt_term (Syntax.var xi); + + fun prt_varT (x, ~1) = prt_typ (TFree (x, [])) + | prt_varT xi = prt_typ (TVar (xi, [])); + + val prt_consts = prt_atoms (prt_term o Const) prt_typ; + val prt_vars = prt_atoms prt_var prt_typ; + val prt_varsT = prt_atoms prt_varT prt_sort; + + + fun pretty_list _ _ [] = [] + | pretty_list name prt lst = [Pretty.big_list name (map prt lst)]; + + fun pretty_subgoal (n, A) = Pretty.markup markup + [Pretty.str (" " ^ string_of_int n ^ ". "), prt_term A]; + fun pretty_subgoals As = map pretty_subgoal (1 upto length As ~~ As); + + val pretty_ffpairs = pretty_list "flex-flex pairs:" (pretty_flexpair ctxt); + + val pretty_consts = pretty_list "constants:" prt_consts o consts_of; + val pretty_vars = pretty_list "variables:" prt_vars o vars_of; + val pretty_varsT = pretty_list "type variables:" prt_varsT o varsT_of; + + + val {prop, tpairs, ...} = Thm.rep_thm state; + val (As, B) = Logic.strip_horn prop; + val ngoals = length As; + + fun pretty_gs (types, sorts) = + (if main then [prt_term B] else []) @ + (if ngoals = 0 then [Pretty.str "No subgoals!"] + else if ngoals > maxgoals then + pretty_subgoals (Library.take (maxgoals, As)) @ + (if msg then [Pretty.str ("A total of " ^ string_of_int ngoals ^ " subgoals...")] + else []) + else pretty_subgoals As) @ + pretty_ffpairs tpairs @ + (if ! show_consts then pretty_consts prop else []) @ + (if types then pretty_vars prop else []) @ + (if sorts then pretty_varsT prop else []); + in + setmp show_no_free_types true + (setmp show_types (! show_types orelse ! show_sorts orelse ! show_all_types) + (setmp show_sorts false pretty_gs)) + (! show_types orelse ! show_sorts orelse ! show_all_types, ! show_sorts) + end; + +fun pretty_goals_without_context n th = + pretty_goals (Syntax.init_pretty_global (Thm.theory_of_thm th)) Markup.none (true, true) n th; + +val print_goals_without_context = + (Pretty.writeln o Pretty.chunks) oo pretty_goals_without_context; + +end; + +end; + diff -r 63686057cbe8 -r abda97d2deea src/Pure/envir.ML --- a/src/Pure/envir.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/envir.ML Thu Jul 23 21:13:21 2009 +0200 @@ -1,249 +1,275 @@ (* Title: Pure/envir.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory -Environments. The type of a term variable / sort of a type variable is -part of its name. The lookup function must apply type substitutions, +Free-form environments. The type of a term variable / sort of a type variable is +part of its name. The lookup function must apply type substitutions, since they may change the identity of a variable. *) signature ENVIR = sig - type tenv - datatype env = Envir of {asol: tenv, iTs: Type.tyenv, maxidx: int} + type tenv = (typ * term) Vartab.table + datatype env = Envir of {maxidx: int, tenv: tenv, tyenv: Type.tyenv} + val maxidx_of: env -> int + val term_env: env -> tenv val type_env: env -> Type.tyenv + val is_empty: env -> bool + val empty: int -> env + val merge: env * env -> env val insert_sorts: env -> sort list -> sort list - exception SAME val genvars: string -> env * typ list -> env * term list val genvar: string -> env * typ -> env * term val lookup: env * (indexname * typ) -> term option val lookup': tenv * (indexname * typ) -> term option val update: ((indexname * typ) * term) * env -> env - val empty: int -> env - val is_empty: env -> bool val above: env -> int -> bool val vupdate: ((indexname * typ) * term) * env -> env - val alist_of: env -> (indexname * (typ * term)) list + val norm_type_same: Type.tyenv -> typ Same.operation + val norm_types_same: Type.tyenv -> typ list Same.operation + val norm_type: Type.tyenv -> typ -> typ + val norm_term_same: env -> term Same.operation val norm_term: env -> term -> term - val norm_term_same: env -> term -> term - val norm_type: Type.tyenv -> typ -> typ - val norm_type_same: Type.tyenv -> typ -> typ - val norm_types_same: Type.tyenv -> typ list -> typ list val beta_norm: term -> term val head_norm: env -> term -> term val eta_contract: term -> term val beta_eta_contract: term -> term val fastype: env -> typ list -> term -> typ - val typ_subst_TVars: Type.tyenv -> typ -> typ - val subst_TVars: Type.tyenv -> term -> term - val subst_Vars: tenv -> term -> term - val subst_vars: Type.tyenv * tenv -> term -> term + val subst_type_same: Type.tyenv -> typ Same.operation + val subst_term_types_same: Type.tyenv -> term Same.operation + val subst_term_same: Type.tyenv * tenv -> term Same.operation + val subst_type: Type.tyenv -> typ -> typ + val subst_term_types: Type.tyenv -> term -> term + val subst_term: Type.tyenv * tenv -> term -> term val expand_atom: typ -> typ * term -> term val expand_term: (term -> (typ * term) option) -> term -> term val expand_term_frees: ((string * typ) * term) list -> term -> term end; -structure Envir : ENVIR = +structure Envir: ENVIR = struct -(*updating can destroy environment in 2 ways!! - (1) variables out of range (2) circular assignments +(** datatype env **) + +(*Updating can destroy environment in 2 ways! + (1) variables out of range + (2) circular assignments *) -type tenv = (typ * term) Vartab.table + +type tenv = (typ * term) Vartab.table; datatype env = Envir of - {maxidx: int, (*maximum index of vars*) - asol: tenv, (*table of assignments to Vars*) - iTs: Type.tyenv} (*table of assignments to TVars*) + {maxidx: int, (*upper bound of maximum index of vars*) + tenv: tenv, (*assignments to Vars*) + tyenv: Type.tyenv}; (*assignments to TVars*) + +fun make_env (maxidx, tenv, tyenv) = Envir {maxidx = maxidx, tenv = tenv, tyenv = tyenv}; +fun map_env f (Envir {maxidx, tenv, tyenv}) = make_env (f (maxidx, tenv, tyenv)); + +fun maxidx_of (Envir {maxidx, ...}) = maxidx; +fun term_env (Envir {tenv, ...}) = tenv; +fun type_env (Envir {tyenv, ...}) = tyenv; + +fun is_empty env = + Vartab.is_empty (term_env env) andalso + Vartab.is_empty (type_env env); -fun type_env (Envir {iTs, ...}) = iTs; + +(* build env *) + +fun empty maxidx = make_env (maxidx, Vartab.empty, Vartab.empty); -(*NB: type unification may invent new sorts*) +fun merge + (Envir {maxidx = maxidx1, tenv = tenv1, tyenv = tyenv1}, + Envir {maxidx = maxidx2, tenv = tenv2, tyenv = tyenv2}) = + make_env (Int.max (maxidx1, maxidx2), + Vartab.merge (op =) (tenv1, tenv2), + Vartab.merge (op =) (tyenv1, tyenv2)); + + +(*NB: type unification may invent new sorts*) (* FIXME tenv!? *) val insert_sorts = Vartab.fold (fn (_, (_, T)) => Sorts.insert_typ T) o type_env; (*Generate a list of distinct variables. Increments index to make them distinct from ALL present variables. *) -fun genvars name (Envir{maxidx, asol, iTs}, Ts) : env * term list = - let fun genvs (_, [] : typ list) : term list = [] - | genvs (n, [T]) = [ Var((name, maxidx+1), T) ] - | genvs (n, T::Ts) = - Var((name ^ radixstring(26,"a",n), maxidx+1), T) - :: genvs(n+1,Ts) - in (Envir{maxidx=maxidx+1, asol=asol, iTs=iTs}, genvs (0,Ts)) end; +fun genvars name (Envir {maxidx, tenv, tyenv}, Ts) : env * term list = + let + fun genvs (_, [] : typ list) : term list = [] + | genvs (n, [T]) = [Var ((name, maxidx + 1), T)] + | genvs (n, T :: Ts) = + Var ((name ^ radixstring (26, "a" , n), maxidx + 1), T) + :: genvs (n + 1, Ts); + in (Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv}, genvs (0, Ts)) end; (*Generate a variable.*) -fun genvar name (env,T) : env * term = - let val (env',[v]) = genvars name (env,[T]) - in (env',v) end; +fun genvar name (env, T) : env * term = + let val (env', [v]) = genvars name (env, [T]) + in (env', v) end; -fun var_clash ixn T T' = raise TYPE ("Variable " ^ - quote (Term.string_of_vname ixn) ^ " has two distinct types", - [T', T], []); +fun var_clash xi T T' = + raise TYPE ("Variable " ^ quote (Term.string_of_vname xi) ^ " has two distinct types", + [T', T], []); -fun gen_lookup f asol (xname, T) = - (case Vartab.lookup asol xname of - NONE => NONE - | SOME (U, t) => if f (T, U) then SOME t - else var_clash xname T U); +fun lookup_check check tenv (xi, T) = + (case Vartab.lookup tenv xi of + NONE => NONE + | SOME (U, t) => if check (T, U) then SOME t else var_clash xi T U); (* When dealing with environments produced by matching instead *) (* of unification, there is no need to chase assigned TVars. *) (* In this case, we can simply ignore the type substitution *) (* and use = instead of eq_type. *) -fun lookup' (asol, p) = gen_lookup op = asol p; +fun lookup' (tenv, p) = lookup_check (op =) tenv p; -fun lookup2 (iTs, asol) p = - if Vartab.is_empty iTs then lookup' (asol, p) - else gen_lookup (Type.eq_type iTs) asol p; - -fun lookup (Envir {asol, iTs, ...}, p) = lookup2 (iTs, asol) p; +fun lookup2 (tyenv, tenv) = + lookup_check (Type.eq_type tyenv) tenv; -fun update (((xname, T), t), Envir {maxidx, asol, iTs}) = - Envir{maxidx=maxidx, asol=Vartab.update_new (xname, (T, t)) asol, iTs=iTs}; +fun lookup (Envir {tenv, tyenv, ...}, p) = lookup2 (tyenv, tenv) p; -(*The empty environment. New variables will start with the given index+1.*) -fun empty m = Envir{maxidx=m, asol=Vartab.empty, iTs=Vartab.empty}; - -(*Test for empty environment*) -fun is_empty (Envir {asol, iTs, ...}) = Vartab.is_empty asol andalso Vartab.is_empty iTs; +fun update (((xi, T), t), Envir {maxidx, tenv, tyenv}) = + Envir {maxidx = maxidx, tenv = Vartab.update_new (xi, (T, t)) tenv, tyenv = tyenv}; (*Determine if the least index updated exceeds lim*) -fun above (Envir {asol, iTs, ...}) lim = - (case Vartab.min_key asol of SOME (_, i) => i > lim | NONE => true) andalso - (case Vartab.min_key iTs of SOME (_, i) => i > lim | NONE => true); +fun above (Envir {tenv, tyenv, ...}) lim = + (case Vartab.min_key tenv of SOME (_, i) => i > lim | NONE => true) andalso + (case Vartab.min_key tyenv of SOME (_, i) => i > lim | NONE => true); (*Update, checking Var-Var assignments: try to suppress higher indexes*) -fun vupdate ((aU as (a, U), t), env as Envir {iTs, ...}) = case t of - Var (nT as (name', T)) => - if a = name' then env (*cycle!*) - else if TermOrd.indexname_ord (a, name') = LESS then - (case lookup (env, nT) of (*if already assigned, chase*) - NONE => update ((nT, Var (a, T)), env) - | SOME u => vupdate ((aU, u), env)) - else update ((aU, t), env) - | _ => update ((aU, t), env); +fun vupdate ((aU as (a, U), t), env as Envir {tyenv, ...}) = + (case t of + Var (nT as (name', T)) => + if a = name' then env (*cycle!*) + else if TermOrd.indexname_ord (a, name') = LESS then + (case lookup (env, nT) of (*if already assigned, chase*) + NONE => update ((nT, Var (a, T)), env) + | SOME u => vupdate ((aU, u), env)) + else update ((aU, t), env) + | _ => update ((aU, t), env)); -(*Convert environment to alist*) -fun alist_of (Envir{asol,...}) = Vartab.dest asol; - -(*** Beta normal form for terms (not eta normal form). - Chases variables in env; Does not exploit sharing of variable bindings - Does not check types, so could loop. ***) +(** beta normalization wrt. environment **) -(*raised when norm has no effect on a term, to do sharing instead of copying*) -exception SAME; +(*Chases variables in env. Does not exploit sharing of variable bindings + Does not check types, so could loop.*) + +local -fun norm_term1 same (asol,t) : term = - let fun norm (Var wT) = - (case lookup' (asol, wT) of - SOME u => (norm u handle SAME => u) - | NONE => raise SAME) - | norm (Abs(a,T,body)) = Abs(a, T, norm body) - | norm (Abs(_,_,body) $ t) = normh(subst_bound (t, body)) - | norm (f $ t) = - ((case norm f of - Abs(_,_,body) => normh(subst_bound (t, body)) - | nf => nf $ (norm t handle SAME => t)) - handle SAME => f $ norm t) - | norm _ = raise SAME - and normh t = norm t handle SAME => t - in (if same then norm else normh) t end +fun norm_type0 tyenv : typ Same.operation = + let + fun norm (Type (a, Ts)) = Type (a, Same.map norm Ts) + | norm (TFree _) = raise Same.SAME + | norm (TVar v) = + (case Type.lookup tyenv v of + SOME U => Same.commit norm U + | NONE => raise Same.SAME); + in norm end; + +fun norm_term1 tenv : term Same.operation = + let + fun norm (Var v) = + (case lookup' (tenv, v) of + SOME u => Same.commit norm u + | NONE => raise Same.SAME) + | norm (Abs (a, T, body)) = Abs (a, T, norm body) + | norm (Abs (_, _, body) $ t) = Same.commit norm (subst_bound (t, body)) + | norm (f $ t) = + ((case norm f of + Abs (_, _, body) => Same.commit norm (subst_bound (t, body)) + | nf => nf $ Same.commit norm t) + handle Same.SAME => f $ norm t) + | norm _ = raise Same.SAME; + in norm end; -fun normT iTs (Type (a, Ts)) = Type (a, normTs iTs Ts) - | normT iTs (TFree _) = raise SAME - | normT iTs (TVar vS) = (case Type.lookup iTs vS of - SOME U => normTh iTs U - | NONE => raise SAME) -and normTh iTs T = ((normT iTs T) handle SAME => T) -and normTs iTs [] = raise SAME - | normTs iTs (T :: Ts) = - ((normT iTs T :: (normTs iTs Ts handle SAME => Ts)) - handle SAME => T :: normTs iTs Ts); +fun norm_term2 tenv tyenv : term Same.operation = + let + val normT = norm_type0 tyenv; + fun norm (Const (a, T)) = Const (a, normT T) + | norm (Free (a, T)) = Free (a, normT T) + | norm (Var (xi, T)) = + (case lookup2 (tyenv, tenv) (xi, T) of + SOME u => Same.commit norm u + | NONE => Var (xi, normT T)) + | norm (Abs (a, T, body)) = + (Abs (a, normT T, Same.commit norm body) + handle Same.SAME => Abs (a, T, norm body)) + | norm (Abs (_, _, body) $ t) = Same.commit norm (subst_bound (t, body)) + | norm (f $ t) = + ((case norm f of + Abs (_, _, body) => Same.commit norm (subst_bound (t, body)) + | nf => nf $ Same.commit norm t) + handle Same.SAME => f $ norm t) + | norm _ = raise Same.SAME; + in norm end; -fun norm_term2 same (asol, iTs, t) : term = - let fun norm (Const (a, T)) = Const(a, normT iTs T) - | norm (Free (a, T)) = Free(a, normT iTs T) - | norm (Var (w, T)) = - (case lookup2 (iTs, asol) (w, T) of - SOME u => normh u - | NONE => Var(w, normT iTs T)) - | norm (Abs (a, T, body)) = - (Abs (a, normT iTs T, normh body) handle SAME => Abs (a, T, norm body)) - | norm (Abs(_, _, body) $ t) = normh (subst_bound (t, body)) - | norm (f $ t) = - ((case norm f of - Abs(_, _, body) => normh (subst_bound (t, body)) - | nf => nf $ normh t) - handle SAME => f $ norm t) - | norm _ = raise SAME - and normh t = (norm t) handle SAME => t - in (if same then norm else normh) t end; +in + +fun norm_type_same tyenv T = + if Vartab.is_empty tyenv then raise Same.SAME + else norm_type0 tyenv T; -fun gen_norm_term same (env as Envir{asol,iTs,...}) t : term = - if Vartab.is_empty iTs then norm_term1 same (asol, t) - else norm_term2 same (asol, iTs, t); +fun norm_types_same tyenv Ts = + if Vartab.is_empty tyenv then raise Same.SAME + else Same.map (norm_type0 tyenv) Ts; + +fun norm_type tyenv T = norm_type_same tyenv T handle Same.SAME => T; -val norm_term = gen_norm_term false; -val norm_term_same = gen_norm_term true; +fun norm_term_same (Envir {tenv, tyenv, ...}) = + if Vartab.is_empty tyenv then norm_term1 tenv + else norm_term2 tenv tyenv; +fun norm_term envir t = norm_term_same envir t handle Same.SAME => t; fun beta_norm t = if Term.has_abs t then norm_term (empty 0) t else t; -fun norm_type iTs = normTh iTs; -fun norm_type_same iTs = - if Vartab.is_empty iTs then raise SAME else normT iTs; - -fun norm_types_same iTs = - if Vartab.is_empty iTs then raise SAME else normTs iTs; +end; (*Put a term into head normal form for unification.*) -fun head_norm env t = +fun head_norm env = let - fun hnorm (Var vT) = (case lookup (env, vT) of + fun norm (Var v) = + (case lookup (env, v) of SOME u => head_norm env u - | NONE => raise SAME) - | hnorm (Abs (a, T, body)) = Abs (a, T, hnorm body) - | hnorm (Abs (_, _, body) $ t) = - head_norm env (subst_bound (t, body)) - | hnorm (f $ t) = (case hnorm f of - Abs (_, _, body) => head_norm env (subst_bound (t, body)) - | nf => nf $ t) - | hnorm _ = raise SAME - in hnorm t handle SAME => t end; + | NONE => raise Same.SAME) + | norm (Abs (a, T, body)) = Abs (a, T, norm body) + | norm (Abs (_, _, body) $ t) = Same.commit norm (subst_bound (t, body)) + | norm (f $ t) = + (case norm f of + Abs (_, _, body) => Same.commit norm (subst_bound (t, body)) + | nf => nf $ t) + | norm _ = raise Same.SAME; + in Same.commit norm end; (*Eta-contract a term (fully)*) local -fun decr lev (Bound i) = if i >= lev then Bound (i - 1) else raise SAME +fun decr lev (Bound i) = if i >= lev then Bound (i - 1) else raise Same.SAME | decr lev (Abs (a, T, body)) = Abs (a, T, decr (lev + 1) body) - | decr lev (t $ u) = (decr lev t $ decrh lev u handle SAME => t $ decr lev u) - | decr _ _ = raise SAME -and decrh lev t = (decr lev t handle SAME => t); + | decr lev (t $ u) = (decr lev t $ decrh lev u handle Same.SAME => t $ decr lev u) + | decr _ _ = raise Same.SAME +and decrh lev t = (decr lev t handle Same.SAME => t); fun eta (Abs (a, T, body)) = ((case eta body of body' as (f $ Bound 0) => if loose_bvar1 (f, 0) then Abs (a, T, body') else decrh 0 f - | body' => Abs (a, T, body')) handle SAME => + | body' => Abs (a, T, body')) handle Same.SAME => (case body of f $ Bound 0 => - if loose_bvar1 (f, 0) then raise SAME + if loose_bvar1 (f, 0) then raise Same.SAME else decrh 0 f - | _ => raise SAME)) - | eta (t $ u) = (eta t $ etah u handle SAME => t $ eta u) - | eta _ = raise SAME -and etah t = (eta t handle SAME => t); + | _ => raise Same.SAME)) + | eta (t $ u) = (eta t $ Same.commit eta u handle Same.SAME => t $ eta u) + | eta _ = raise Same.SAME; in fun eta_contract t = - if Term.has_abs t then etah t else t; + if Term.has_abs t then Same.commit eta t else t; val beta_eta_contract = eta_contract o beta_norm; @@ -252,64 +278,89 @@ (*finds type of term without checking that combinations are consistent Ts holds types of bound variables*) -fun fastype (Envir {iTs, ...}) = -let val funerr = "fastype: expected function type"; +fun fastype (Envir {tyenv, ...}) = + let + val funerr = "fastype: expected function type"; fun fast Ts (f $ u) = - (case fast Ts f of - Type ("fun", [_, T]) => T - | TVar ixnS => - (case Type.lookup iTs ixnS of - SOME (Type ("fun", [_, T])) => T - | _ => raise TERM (funerr, [f $ u])) - | _ => raise TERM (funerr, [f $ u])) + (case fast Ts f of + Type ("fun", [_, T]) => T + | TVar v => + (case Type.lookup tyenv v of + SOME (Type ("fun", [_, T])) => T + | _ => raise TERM (funerr, [f $ u])) + | _ => raise TERM (funerr, [f $ u])) | fast Ts (Const (_, T)) = T | fast Ts (Free (_, T)) = T | fast Ts (Bound i) = - (nth Ts i - handle Subscript => raise TERM ("fastype: Bound", [Bound i])) + (nth Ts i handle Subscript => raise TERM ("fastype: Bound", [Bound i])) | fast Ts (Var (_, T)) = T - | fast Ts (Abs (_, T, u)) = T --> fast (T :: Ts) u -in fast end; + | fast Ts (Abs (_, T, u)) = T --> fast (T :: Ts) u; + in fast end; + -(*Substitute for type Vars in a type*) -fun typ_subst_TVars iTs T = if Vartab.is_empty iTs then T else - let fun subst(Type(a, Ts)) = Type(a, map subst Ts) - | subst(T as TFree _) = T - | subst(T as TVar ixnS) = - (case Type.lookup iTs ixnS of NONE => T | SOME(U) => U) - in subst T end; +(** plain substitution -- without variable chasing **) + +local -(*Substitute for type Vars in a term*) -val subst_TVars = map_types o typ_subst_TVars; +fun subst_type0 tyenv = Term_Subst.map_atypsT_same + (fn TVar v => + (case Type.lookup tyenv v of + SOME U => U + | NONE => raise Same.SAME) + | _ => raise Same.SAME); -(*Substitute for Vars in a term *) -fun subst_Vars itms t = if Vartab.is_empty itms then t else - let fun subst (v as Var ixnT) = the_default v (lookup' (itms, ixnT)) - | subst (Abs (a, T, t)) = Abs (a, T, subst t) - | subst (f $ t) = subst f $ subst t - | subst t = t - in subst t end; +fun subst_term1 tenv = Term_Subst.map_aterms_same + (fn Var v => + (case lookup' (tenv, v) of + SOME u => u + | NONE => raise Same.SAME) + | _ => raise Same.SAME); -(*Substitute for type/term Vars in a term *) -fun subst_vars (iTs, itms) = - if Vartab.is_empty iTs then subst_Vars itms else - let fun subst (Const (a, T)) = Const(a, typ_subst_TVars iTs T) - | subst (Free (a, T)) = Free (a, typ_subst_TVars iTs T) - | subst (Var (ixn, T)) = (case lookup' (itms, (ixn, T)) of - NONE => Var (ixn, typ_subst_TVars iTs T) - | SOME t => t) - | subst (b as Bound _) = b - | subst (Abs (a, T, t)) = Abs(a, typ_subst_TVars iTs T, subst t) - | subst (f $ t) = subst f $ subst t +fun subst_term2 tenv tyenv : term Same.operation = + let + val substT = subst_type0 tyenv; + fun subst (Const (a, T)) = Const (a, substT T) + | subst (Free (a, T)) = Free (a, substT T) + | subst (Var (xi, T)) = + (case lookup' (tenv, (xi, T)) of + SOME u => u + | NONE => Var (xi, substT T)) + | subst (Bound _) = raise Same.SAME + | subst (Abs (a, T, t)) = + (Abs (a, substT T, Same.commit subst t) + handle Same.SAME => Abs (a, T, subst t)) + | subst (t $ u) = (subst t $ Same.commit subst u handle Same.SAME => t $ subst u); in subst end; +in -(* expand defined atoms -- with local beta reduction *) +fun subst_type_same tyenv T = + if Vartab.is_empty tyenv then raise Same.SAME + else subst_type0 tyenv T; + +fun subst_term_types_same tyenv t = + if Vartab.is_empty tyenv then raise Same.SAME + else Term_Subst.map_types_same (subst_type0 tyenv) t; + +fun subst_term_same (tyenv, tenv) = + if Vartab.is_empty tenv then subst_term_types_same tyenv + else if Vartab.is_empty tyenv then subst_term1 tenv + else subst_term2 tenv tyenv; + +fun subst_type tyenv T = subst_type_same tyenv T handle Same.SAME => T; +fun subst_term_types tyenv t = subst_term_types_same tyenv t handle Same.SAME => t; +fun subst_term envs t = subst_term_same envs t handle Same.SAME => t; + +end; + + + +(** expand defined atoms -- with local beta reduction **) fun expand_atom T (U, u) = - subst_TVars (Type.raw_match (U, T) Vartab.empty) u - handle Type.TYPE_MATCH => raise TYPE ("expand_atom: ill-typed replacement", [T, U], [u]); + subst_term_types (Type.raw_match (U, T) Vartab.empty) u + handle Type.TYPE_MATCH => raise TYPE ("expand_atom: ill-typed replacement", [T, U], [u]); fun expand_term get = let @@ -322,10 +373,9 @@ (case head of Abs (x, T, t) => comb (Abs (x, T, expand t)) | _ => - (case get head of - SOME def => Term.betapplys (expand_atom (Term.fastype_of head) def, args') - | NONE => comb head) - | _ => comb head) + (case get head of + SOME def => Term.betapplys (expand_atom (Term.fastype_of head) def, args') + | NONE => comb head)) end; in expand end; diff -r 63686057cbe8 -r abda97d2deea src/Pure/goal.ML --- a/src/Pure/goal.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/goal.ML Thu Jul 23 21:13:21 2009 +0200 @@ -6,7 +6,7 @@ signature BASIC_GOAL = sig - val parallel_proofs: bool ref + val parallel_proofs: int ref val SELECT_GOAL: tactic -> int -> tactic val CONJUNCTS: tactic -> int -> tactic val PRECISE_CONJUNCTS: int -> tactic -> int -> tactic @@ -21,6 +21,7 @@ val finish: thm -> thm val norm_result: thm -> thm val future_enabled: unit -> bool + val local_future_enabled: unit -> bool val future_result: Proof.context -> thm future -> term -> thm val prove_internal: cterm list -> cterm -> (thm list -> tactic) -> thm val prove_multi: Proof.context -> string list -> term list -> term list -> @@ -77,7 +78,7 @@ (case Thm.nprems_of th of 0 => conclude th | n => raise THM ("Proof failed.\n" ^ - Pretty.string_of (Pretty.chunks (Display.pretty_goals n th)) ^ + Pretty.string_of (Pretty.chunks (Display_Goal.pretty_goals_without_context n th)) ^ ("\n" ^ string_of_int n ^ " unsolved goal(s)!"), 0, [th])); @@ -95,10 +96,12 @@ (* future_enabled *) -val parallel_proofs = ref true; +val parallel_proofs = ref 1; fun future_enabled () = - Future.enabled () andalso ! parallel_proofs andalso is_some (Future.thread_data ()); + Future.enabled () andalso Future.is_worker () andalso ! parallel_proofs >= 1; + +fun local_future_enabled () = future_enabled () andalso ! parallel_proofs >= 2; (* future_result *) @@ -120,14 +123,15 @@ val instT = map (fn (a, S) => (certT (TVar ((a, 0), S)), certT (TFree (a, S)))) tfrees; val global_prop = - Term.map_types Logic.varifyT (fold_rev Logic.all xs (Logic.list_implies (As, prop))); + cert (Term.map_types Logic.varifyT (fold_rev Logic.all xs (Logic.list_implies (As, prop)))) + |> Thm.weaken_sorts (Variable.sorts_of ctxt); val global_result = result |> Future.map (Thm.adjust_maxidx_thm ~1 #> Drule.implies_intr_list assms #> Drule.forall_intr_list fixes #> Thm.generalize (map #1 tfrees, []) 0); val local_result = - Thm.future global_result (cert global_prop) + Thm.future global_result global_prop |> Thm.instantiate (instT, []) |> Drule.forall_elim_list fixes |> fold (Thm.elim_implies o Thm.assume) assms; diff -r 63686057cbe8 -r abda97d2deea src/Pure/logic.ML --- a/src/Pure/logic.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/logic.ML Thu Jul 23 21:13:21 2009 +0200 @@ -54,8 +54,10 @@ val close_form: term -> term val combound: term * int * int -> term val rlist_abs: (string * typ) list * term -> term + val incr_tvar_same: int -> typ Same.operation + val incr_tvar: int -> typ -> typ + val incr_indexes_same: typ list * int -> term Same.operation val incr_indexes: typ list * int -> term -> term - val incr_tvar: int -> typ -> typ val lift_abs: int -> term -> term -> term val lift_all: int -> term -> term -> term val strip_assums_hyp: term -> term list @@ -70,8 +72,6 @@ val unvarifyT: typ -> typ val varify: term -> term val unvarify: term -> term - val legacy_varifyT: typ -> typ - val legacy_varify: term -> term val get_goal: term -> int -> term val goal_params: term -> int -> term * term list val prems_of_goal: term -> int -> term list @@ -305,45 +305,35 @@ fun rlist_abs ([], body) = body | rlist_abs ((a,T)::pairs, body) = rlist_abs(pairs, Abs(a, T, body)); - -local exception SAME in +fun incr_tvar_same 0 = Same.same + | incr_tvar_same k = Term_Subst.map_atypsT_same + (fn TVar ((a, i), S) => TVar ((a, i + k), S) + | _ => raise Same.SAME); -fun incrT k = - let - fun incr (TVar ((a, i), S)) = TVar ((a, i + k), S) - | incr (Type (a, Ts)) = Type (a, incrs Ts) - | incr _ = raise SAME - and incrs (T :: Ts) = - (incr T :: (incrs Ts handle SAME => Ts) - handle SAME => T :: incrs Ts) - | incrs [] = raise SAME; - in incr end; +fun incr_tvar k T = incr_tvar_same k T handle Same.SAME => T; (*For all variables in the term, increment indexnames and lift over the Us result is ?Gidx(B.(lev+n-1),...,B.lev) where lev is abstraction level *) -fun incr_indexes ([], 0) t = t - | incr_indexes (Ts, k) t = - let - val n = length Ts; - val incrT = if k = 0 then I else incrT k; +fun incr_indexes_same ([], 0) = Same.same + | incr_indexes_same (Ts, k) = + let + val n = length Ts; + val incrT = incr_tvar_same k; - fun incr lev (Var ((x, i), T)) = - combound (Var ((x, i + k), Ts ---> (incrT T handle SAME => T)), lev, n) - | incr lev (Abs (x, T, body)) = - (Abs (x, incrT T, incr (lev + 1) body handle SAME => body) - handle SAME => Abs (x, T, incr (lev + 1) body)) - | incr lev (t $ u) = - (incr lev t $ (incr lev u handle SAME => u) - handle SAME => t $ incr lev u) - | incr _ (Const (c, T)) = Const (c, incrT T) - | incr _ (Free (x, T)) = Free (x, incrT T) - | incr _ (t as Bound _) = t; - in incr 0 t handle SAME => t end; + fun incr lev (Var ((x, i), T)) = + combound (Var ((x, i + k), Ts ---> Same.commit incrT T), lev, n) + | incr lev (Abs (x, T, body)) = + (Abs (x, incrT T, incr (lev + 1) body handle Same.SAME => body) + handle Same.SAME => Abs (x, T, incr (lev + 1) body)) + | incr lev (t $ u) = + (incr lev t $ (incr lev u handle Same.SAME => u) + handle Same.SAME => t $ incr lev u) + | incr _ (Const (c, T)) = Const (c, incrT T) + | incr _ (Free (x, T)) = Free (x, incrT T) + | incr _ (Bound _) = raise Same.SAME; + in incr 0 end; -fun incr_tvar 0 T = T - | incr_tvar k T = incrT k T handle SAME => T; - -end; +fun incr_indexes arg t = incr_indexes_same arg t handle Same.SAME => t; (* Lifting functions from subgoal and increment: @@ -424,6 +414,8 @@ a $ Abs(c, T, list_rename_params (cs, t)) | list_rename_params (cs, B) = B; + + (*** Treatment of "assume", "erule", etc. ***) (*Strips assumptions in goal yielding @@ -440,8 +432,7 @@ strip_assums_imp (nasms-1, H::Hs, B) | strip_assums_imp (_, Hs, B) = (Hs, B); (*recursion terminated by B*) - -(*Strips OUTER parameters only, unlike similar legacy versions.*) +(*Strips OUTER parameters only.*) fun strip_assums_all (params, Const("all",_)$Abs(a,T,t)) = strip_assums_all ((a,T)::params, t) | strip_assums_all (params, B) = (params, B); @@ -474,43 +465,37 @@ fun bad_schematic xi = "Illegal schematic variable: " ^ quote (Term.string_of_vname xi); fun bad_fixed x = "Illegal fixed variable: " ^ quote x; -fun varifyT_option ty = ty - |> Term_Subst.map_atypsT_option - (fn TFree (a, S) => SOME (TVar ((a, 0), S)) +fun varifyT_same ty = ty + |> Term_Subst.map_atypsT_same + (fn TFree (a, S) => TVar ((a, 0), S) | TVar (ai, _) => raise TYPE (bad_schematic ai, [ty], [])); -fun unvarifyT_option ty = ty - |> Term_Subst.map_atypsT_option - (fn TVar ((a, 0), S) => SOME (TFree (a, S)) +fun unvarifyT_same ty = ty + |> Term_Subst.map_atypsT_same + (fn TVar ((a, 0), S) => TFree (a, S) | TVar (ai, _) => raise TYPE (bad_schematic ai, [ty], []) | TFree (a, _) => raise TYPE (bad_fixed a, [ty], [])); -val varifyT = perhaps varifyT_option; -val unvarifyT = perhaps unvarifyT_option; +val varifyT = Same.commit varifyT_same; +val unvarifyT = Same.commit unvarifyT_same; fun varify tm = tm - |> perhaps (Term_Subst.map_aterms_option - (fn Free (x, T) => SOME (Var ((x, 0), T)) + |> Same.commit (Term_Subst.map_aterms_same + (fn Free (x, T) => Var ((x, 0), T) | Var (xi, _) => raise TERM (bad_schematic xi, [tm]) - | _ => NONE)) - |> perhaps (Term_Subst.map_types_option varifyT_option) + | _ => raise Same.SAME)) + |> Same.commit (Term_Subst.map_types_same varifyT_same) handle TYPE (msg, _, _) => raise TERM (msg, [tm]); fun unvarify tm = tm - |> perhaps (Term_Subst.map_aterms_option - (fn Var ((x, 0), T) => SOME (Free (x, T)) + |> Same.commit (Term_Subst.map_aterms_same + (fn Var ((x, 0), T) => Free (x, T) | Var (xi, _) => raise TERM (bad_schematic xi, [tm]) | Free (x, _) => raise TERM (bad_fixed x, [tm]) - | _ => NONE)) - |> perhaps (Term_Subst.map_types_option unvarifyT_option) + | _ => raise Same.SAME)) + |> Same.commit (Term_Subst.map_types_same unvarifyT_same) handle TYPE (msg, _, _) => raise TERM (msg, [tm]); -val legacy_varifyT = Term.map_atyps (fn TFree (a, S) => TVar ((a, 0), S) | T => T); - -val legacy_varify = - Term.map_aterms (fn Free (x, T) => Var ((x, 0), T) | t => t) #> - Term.map_types legacy_varifyT; - (* goal states *) diff -r 63686057cbe8 -r abda97d2deea src/Pure/old_goals.ML --- a/src/Pure/old_goals.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/old_goals.ML Thu Jul 23 21:13:21 2009 +0200 @@ -135,7 +135,8 @@ (*Default action is to print an error message; could be suppressed for special applications.*) fun result_error_default state msg : thm = - Pretty.str "Bad final proof state:" :: Display.pretty_goals (!goals_limit) state @ + Pretty.str "Bad final proof state:" :: + Display_Goal.pretty_goals_without_context (!goals_limit) state @ [Pretty.str msg, Pretty.str "Proof failed!"] |> Pretty.chunks |> Pretty.string_of |> error; val result_error_fn = ref result_error_default; @@ -199,7 +200,7 @@ case e of THM (msg,i,thms) => (writeln ("Exception THM " ^ string_of_int i ^ " raised:\n" ^ msg); - List.app Display.print_thm thms) + List.app (writeln o Display.string_of_thm_global thy) thms) | THEORY (msg,thys) => (writeln ("Exception THEORY raised:\n" ^ msg); List.app (writeln o Context.str_of_thy) thys) @@ -277,7 +278,7 @@ (if ngoals > 0 then " (" ^ string_of_int ngoals ^ " subgoal" ^ (if ngoals <> 1 then "s" else "") ^ ")" else ""))] @ - Display.pretty_goals m th + Display_Goal.pretty_goals_without_context m th end |> Pretty.chunks |> Pretty.writeln; (*Printing can raise exceptions, so the assignment occurs last. diff -r 63686057cbe8 -r abda97d2deea src/Pure/pattern.ML --- a/src/Pure/pattern.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/pattern.ML Thu Jul 23 21:13:21 2009 +0200 @@ -141,8 +141,10 @@ | split_type (Type ("fun",[T1,T2]),n,Ts) = split_type (T2,n-1,T1::Ts) | split_type _ = error("split_type"); -fun type_of_G (env as Envir.Envir {iTs, ...}) (T,n,is) = - let val (Ts, U) = split_type (Envir.norm_type iTs T, n, []) +fun type_of_G env (T, n, is) = + let + val tyenv = Envir.type_env env; + val (Ts, U) = split_type (Envir.norm_type tyenv T, n, []); in map (nth Ts) is ---> U end; fun mkhnf (binders,is,G,js) = mkabs (binders, is, app(G,js)); @@ -225,11 +227,12 @@ end; in if TermOrd.indexname_ord (G,F) = LESS then ff(F,Fty,is,G,Gty,js) else ff(G,Gty,js,F,Fty,is) end -fun unify_types thy (T,U) (env as Envir.Envir{asol,iTs,maxidx}) = - if T=U then env - else let val (iTs',maxidx') = Sign.typ_unify thy (U, T) (iTs, maxidx) - in Envir.Envir{asol=asol,maxidx=maxidx',iTs=iTs'} end - handle Type.TUNIFY => (typ_clash thy (iTs,T,U); raise Unif); +fun unify_types thy (T, U) (env as Envir.Envir {maxidx, tenv, tyenv}) = + if T = U then env + else + let val (tyenv', maxidx') = Sign.typ_unify thy (U, T) (tyenv, maxidx) + in Envir.Envir {maxidx = maxidx', tenv = tenv, tyenv = tyenv'} end + handle Type.TUNIFY => (typ_clash thy (tyenv, T, U); raise Unif); fun unif thy binders (s,t) env = case (Envir.head_norm env s, Envir.head_norm env t) of (Abs(ns,Ts,ts),Abs(nt,Tt,tt)) => @@ -351,7 +354,7 @@ Abs(ns,Ts,ts) => (case obj of Abs(nt,Tt,tt) => mtch ((nt,Tt)::binders) (ts,tt) env - | _ => let val Tt = Envir.typ_subst_TVars iTs Ts + | _ => let val Tt = Envir.subst_type iTs Ts in mtch((ns,Tt)::binders) (ts,(incr obj)$Bound(0)) env end) | _ => (case obj of Abs(nt,Tt,tt) => @@ -425,7 +428,7 @@ fun match_rew thy tm (tm1, tm2) = let val rtm = the_default tm2 (Term.rename_abs tm1 tm tm2) in - SOME (Envir.subst_vars (match thy (tm1, tm) (Vartab.empty, Vartab.empty)) rtm, rtm) + SOME (Envir.subst_term (match thy (tm1, tm) (Vartab.empty, Vartab.empty)) rtm, rtm) handle MATCH => NONE end; diff -r 63686057cbe8 -r abda97d2deea src/Pure/proofterm.ML --- a/src/Pure/proofterm.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/proofterm.ML Thu Jul 23 21:13:21 2009 +0200 @@ -37,12 +37,11 @@ type oracle = string * term type pthm = serial * (string * term * proof_body future) - val join_body: proof_body future -> - {oracles: oracle OrdList.T, thms: pthm OrdList.T, proof: proof} + val proof_of: proof_body -> proof val join_proof: proof_body future -> proof - val proof_of: proof_body -> proof val fold_proof_atoms: bool -> (proof -> 'a -> 'a) -> proof list -> 'a -> 'a val fold_body_thms: (string * term * proof_body -> 'a -> 'a) -> proof_body list -> 'a -> 'a + val join_bodies: proof_body list -> unit val status_of: proof_body list -> {failed: bool, oracle: bool, unfinished: bool} val oracle_ord: oracle * oracle -> order @@ -88,6 +87,7 @@ val instantiate: ((indexname * sort) * typ) list * ((indexname * typ) * term) list -> proof -> proof val lift_proof: term -> int -> term -> proof -> proof + val incr_indexes: int -> proof -> proof val assumption_proof: term list -> term -> int -> proof -> proof val bicompose_proof: bool -> term list -> term list -> term list -> term option -> int -> int -> proof -> proof -> proof @@ -127,9 +127,6 @@ structure Proofterm : PROOFTERM = struct -open Envir; - - (***** datatype proof *****) datatype proof = @@ -153,10 +150,8 @@ type oracle = string * term; type pthm = serial * (string * term * proof_body future); -val join_body = Future.join #> (fn PBody args => args); -val join_proof = #proof o join_body; - fun proof_of (PBody {proof, ...}) = proof; +val join_proof = Future.join #> proof_of; (***** proof atoms *****) @@ -178,15 +173,19 @@ fun fold_body_thms f = let - fun app (PBody {thms, ...}) = thms |> fold (fn (i, (name, prop, body)) => fn (x, seen) => - if Inttab.defined seen i then (x, seen) - else - let - val body' = Future.join body; - val (x', seen') = app body' (x, Inttab.update (i, ()) seen); - in (f (name, prop, body') x', seen') end); + fun app (PBody {thms, ...}) = + (Future.join_results (map (#3 o #2) thms); + thms |> fold (fn (i, (name, prop, body)) => fn (x, seen) => + if Inttab.defined seen i then (x, seen) + else + let + val body' = Future.join body; + val (x', seen') = app body' (x, Inttab.update (i, ()) seen); + in (f (name, prop, body') x', seen') end)); in fn bodies => fn x => #1 (fold app bodies (x, Inttab.empty)) end; +fun join_bodies bodies = fold_body_thms (fn _ => fn () => ()) bodies (); + fun status_of bodies = let fun status (PBody {oracles, thms, ...}) x = @@ -206,7 +205,8 @@ in ((oracle, true, failed), seen') end) end); in ((oracle orelse not (null oracles), unfinished, failed), seen) end; - val (oracle, unfinished, failed) = #1 (fold status bodies ((false, false, false), Inttab.empty)); + val (oracle, unfinished, failed) = + #1 (fold status bodies ((false, false, false), Inttab.empty)); in {oracle = oracle, unfinished = unfinished, failed = failed} end; @@ -220,13 +220,15 @@ val all_oracles_of = let - fun collect (PBody {oracles, thms, ...}) = thms |> fold (fn (i, (_, _, body)) => fn (x, seen) => - if Inttab.defined seen i then (x, seen) - else - let - val body' = Future.join body; - val (x', seen') = collect body' (x, Inttab.update (i, ()) seen); - in (merge_oracles oracles x', seen') end); + fun collect (PBody {oracles, thms, ...}) = + (Future.join_results (map (#3 o #2) thms); + thms |> fold (fn (i, (_, _, body)) => fn (x, seen) => + if Inttab.defined seen i then (x, seen) + else + let + val body' = Future.join body; + val (x', seen') = collect body' (x, Inttab.update (i, ()) seen); + in (merge_oracles oracles x', seen') end)); in fn body => #1 (collect body ([], Inttab.empty)) end; fun approximate_proof_body prf = @@ -269,47 +271,41 @@ val mk_Abst = fold_rev (fn (s, T:typ) => fn prf => Abst (s, NONE, prf)); fun mk_AbsP (i, prf) = funpow i (fn prf => AbsP ("H", NONE, prf)) prf; -fun apsome f NONE = raise SAME - | apsome f (SOME x) = (case f x of NONE => raise SAME | some => some); - -fun apsome' f NONE = raise SAME - | apsome' f (SOME x) = SOME (f x); - fun map_proof_terms_option f g = let - fun map_typs (T :: Ts) = - (case g T of - NONE => T :: map_typs Ts - | SOME T' => T' :: (map_typs Ts handle SAME => Ts)) - | map_typs [] = raise SAME; + val term = Same.function f; + val typ = Same.function g; + val typs = Same.map typ; - fun mapp (Abst (s, T, prf)) = (Abst (s, apsome g T, mapph prf) - handle SAME => Abst (s, T, mapp prf)) - | mapp (AbsP (s, t, prf)) = (AbsP (s, apsome f t, mapph prf) - handle SAME => AbsP (s, t, mapp prf)) - | mapp (prf % t) = (mapp prf % (apsome f t handle SAME => t) - handle SAME => prf % apsome f t) - | mapp (prf1 %% prf2) = (mapp prf1 %% mapph prf2 - handle SAME => prf1 %% mapp prf2) - | mapp (PAxm (a, prop, SOME Ts)) = PAxm (a, prop, SOME (map_typs Ts)) - | mapp (OfClass (T, c)) = OfClass (apsome g (SOME T) |> the, c) - | mapp (Oracle (a, prop, SOME Ts)) = Oracle (a, prop, SOME (map_typs Ts)) - | mapp (Promise (i, prop, Ts)) = Promise (i, prop, map_typs Ts) - | mapp (PThm (i, ((a, prop, SOME Ts), body))) = - PThm (i, ((a, prop, SOME (map_typs Ts)), body)) - | mapp _ = raise SAME - and mapph prf = (mapp prf handle SAME => prf) - - in mapph end; + fun proof (Abst (s, T, prf)) = + (Abst (s, Same.map_option typ T, Same.commit proof prf) + handle Same.SAME => Abst (s, T, proof prf)) + | proof (AbsP (s, t, prf)) = + (AbsP (s, Same.map_option term t, Same.commit proof prf) + handle Same.SAME => AbsP (s, t, proof prf)) + | proof (prf % t) = + (proof prf % Same.commit (Same.map_option term) t + handle Same.SAME => prf % Same.map_option term t) + | proof (prf1 %% prf2) = + (proof prf1 %% Same.commit proof prf2 + handle Same.SAME => prf1 %% proof prf2) + | proof (PAxm (a, prop, SOME Ts)) = PAxm (a, prop, SOME (typs Ts)) + | proof (OfClass (T, c)) = OfClass (typ T, c) + | proof (Oracle (a, prop, SOME Ts)) = Oracle (a, prop, SOME (typs Ts)) + | proof (Promise (i, prop, Ts)) = Promise (i, prop, typs Ts) + | proof (PThm (i, ((a, prop, SOME Ts), body))) = + PThm (i, ((a, prop, SOME (typs Ts)), body)) + | proof _ = raise Same.SAME; + in Same.commit proof end; fun same eq f x = let val x' = f x - in if eq (x, x') then raise SAME else x' end; + in if eq (x, x') then raise Same.SAME else x' end; fun map_proof_terms f g = map_proof_terms_option - (fn t => SOME (same (op =) f t) handle SAME => NONE) - (fn T => SOME (same (op =) g T) handle SAME => NONE); + (fn t => SOME (same (op =) f t) handle Same.SAME => NONE) + (fn T => SOME (same (op =) g T) handle Same.SAME => NONE); fun fold_proof_terms f g (Abst (_, SOME T, prf)) = g T #> fold_proof_terms f g prf | fold_proof_terms f g (Abst (_, NONE, prf)) = fold_proof_terms f g prf @@ -338,7 +334,8 @@ | change_type (SOME [T]) (OfClass (_, c)) = OfClass (T, c) | change_type opTs (Oracle (name, prop, _)) = Oracle (name, prop, opTs) | change_type opTs (Promise _) = error "change_type: unexpected promise" - | change_type opTs (PThm (i, ((name, prop, _), body))) = PThm (i, ((name, prop, opTs), body)) + | change_type opTs (PThm (i, ((name, prop, _), body))) = + PThm (i, ((name, prop, opTs), body)) | change_type _ prf = prf; @@ -359,20 +356,20 @@ fun abst' lev u = if v aconv u then Bound lev else (case u of Abs (a, T, t) => Abs (a, T, abst' (lev + 1) t) - | f $ t => (abst' lev f $ absth' lev t handle SAME => f $ abst' lev t) - | _ => raise SAME) - and absth' lev t = (abst' lev t handle SAME => t); + | f $ t => (abst' lev f $ absth' lev t handle Same.SAME => f $ abst' lev t) + | _ => raise Same.SAME) + and absth' lev t = (abst' lev t handle Same.SAME => t); fun abst lev (AbsP (a, t, prf)) = - (AbsP (a, apsome' (abst' lev) t, absth lev prf) - handle SAME => AbsP (a, t, abst lev prf)) + (AbsP (a, Same.map_option (abst' lev) t, absth lev prf) + handle Same.SAME => AbsP (a, t, abst lev prf)) | abst lev (Abst (a, T, prf)) = Abst (a, T, abst (lev + 1) prf) | abst lev (prf1 %% prf2) = (abst lev prf1 %% absth lev prf2 - handle SAME => prf1 %% abst lev prf2) + handle Same.SAME => prf1 %% abst lev prf2) | abst lev (prf % t) = (abst lev prf % Option.map (absth' lev) t - handle SAME => prf % apsome' (abst' lev) t) - | abst _ _ = raise SAME - and absth lev prf = (abst lev prf handle SAME => prf) + handle Same.SAME => prf % Same.map_option (abst' lev) t) + | abst _ _ = raise Same.SAME + and absth lev prf = (abst lev prf handle Same.SAME => prf); in absth 0 end; @@ -385,22 +382,22 @@ fun incr_bv' inct tlev t = incr_bv (inct, tlev, t); fun prf_incr_bv' incP inct Plev tlev (PBound i) = - if i >= Plev then PBound (i+incP) else raise SAME + if i >= Plev then PBound (i+incP) else raise Same.SAME | prf_incr_bv' incP inct Plev tlev (AbsP (a, t, body)) = - (AbsP (a, apsome' (same (op =) (incr_bv' inct tlev)) t, - prf_incr_bv incP inct (Plev+1) tlev body) handle SAME => + (AbsP (a, Same.map_option (same (op =) (incr_bv' inct tlev)) t, + prf_incr_bv incP inct (Plev+1) tlev body) handle Same.SAME => AbsP (a, t, prf_incr_bv' incP inct (Plev+1) tlev body)) | prf_incr_bv' incP inct Plev tlev (Abst (a, T, body)) = Abst (a, T, prf_incr_bv' incP inct Plev (tlev+1) body) | prf_incr_bv' incP inct Plev tlev (prf %% prf') = (prf_incr_bv' incP inct Plev tlev prf %% prf_incr_bv incP inct Plev tlev prf' - handle SAME => prf %% prf_incr_bv' incP inct Plev tlev prf') + handle Same.SAME => prf %% prf_incr_bv' incP inct Plev tlev prf') | prf_incr_bv' incP inct Plev tlev (prf % t) = (prf_incr_bv' incP inct Plev tlev prf % Option.map (incr_bv' inct tlev) t - handle SAME => prf % apsome' (same (op =) (incr_bv' inct tlev)) t) - | prf_incr_bv' _ _ _ _ _ = raise SAME + handle Same.SAME => prf % Same.map_option (same (op =) (incr_bv' inct tlev)) t) + | prf_incr_bv' _ _ _ _ _ = raise Same.SAME and prf_incr_bv incP inct Plev tlev prf = - (prf_incr_bv' incP inct Plev tlev prf handle SAME => prf); + (prf_incr_bv' incP inct Plev tlev prf handle Same.SAME => prf); fun incr_pboundvars 0 0 prf = prf | incr_pboundvars incP inct prf = prf_incr_bv incP inct 0 0 prf; @@ -448,7 +445,7 @@ fun del_conflicting_vars env t = Term_Subst.instantiate (map_filter (fn ixnS as (_, S) => - (Type.lookup (type_env env) ixnS; NONE) handle TYPE _ => + (Type.lookup (Envir.type_env env) ixnS; NONE) handle TYPE _ => SOME (ixnS, TFree ("'dummy", S))) (OldTerm.term_tvars t), map_filter (fn Var (ixnT as (_, T)) => (Envir.lookup (env, ixnT); NONE) handle TYPE _ => @@ -456,7 +453,7 @@ fun norm_proof env = let - val envT = type_env env; + val envT = Envir.type_env env; fun msg s = warning ("type conflict in norm_proof:\n" ^ s); fun htype f t = f env t handle TYPE (s, _, _) => (msg s; f env (del_conflicting_vars env t)); @@ -464,23 +461,31 @@ (msg s; f envT (del_conflicting_tvars envT T)); fun htypeTs f Ts = f envT Ts handle TYPE (s, _, _) => (msg s; f envT (map (del_conflicting_tvars envT) Ts)); - fun norm (Abst (s, T, prf)) = (Abst (s, apsome' (htypeT norm_type_same) T, normh prf) - handle SAME => Abst (s, T, norm prf)) - | norm (AbsP (s, t, prf)) = (AbsP (s, apsome' (htype norm_term_same) t, normh prf) - handle SAME => AbsP (s, t, norm prf)) - | norm (prf % t) = (norm prf % Option.map (htype norm_term) t - handle SAME => prf % apsome' (htype norm_term_same) t) - | norm (prf1 %% prf2) = (norm prf1 %% normh prf2 - handle SAME => prf1 %% norm prf2) - | norm (PAxm (s, prop, Ts)) = PAxm (s, prop, apsome' (htypeTs norm_types_same) Ts) - | norm (OfClass (T, c)) = OfClass (htypeT norm_type_same T, c) - | norm (Oracle (s, prop, Ts)) = Oracle (s, prop, apsome' (htypeTs norm_types_same) Ts) - | norm (Promise (i, prop, Ts)) = Promise (i, prop, htypeTs norm_types_same Ts) + + fun norm (Abst (s, T, prf)) = + (Abst (s, Same.map_option (htypeT Envir.norm_type_same) T, Same.commit norm prf) + handle Same.SAME => Abst (s, T, norm prf)) + | norm (AbsP (s, t, prf)) = + (AbsP (s, Same.map_option (htype Envir.norm_term_same) t, Same.commit norm prf) + handle Same.SAME => AbsP (s, t, norm prf)) + | norm (prf % t) = + (norm prf % Option.map (htype Envir.norm_term) t + handle Same.SAME => prf % Same.map_option (htype Envir.norm_term_same) t) + | norm (prf1 %% prf2) = + (norm prf1 %% Same.commit norm prf2 + handle Same.SAME => prf1 %% norm prf2) + | norm (PAxm (s, prop, Ts)) = + PAxm (s, prop, Same.map_option (htypeTs Envir.norm_types_same) Ts) + | norm (OfClass (T, c)) = + OfClass (htypeT Envir.norm_type_same T, c) + | norm (Oracle (s, prop, Ts)) = + Oracle (s, prop, Same.map_option (htypeTs Envir.norm_types_same) Ts) + | norm (Promise (i, prop, Ts)) = + Promise (i, prop, htypeTs Envir.norm_types_same Ts) | norm (PThm (i, ((s, t, Ts), body))) = - PThm (i, ((s, t, apsome' (htypeTs norm_types_same) Ts), body)) - | norm _ = raise SAME - and normh prf = (norm prf handle SAME => prf); - in normh end; + PThm (i, ((s, t, Same.map_option (htypeTs Envir.norm_types_same) Ts), body)) + | norm _ = raise Same.SAME; + in Same.commit norm end; (***** Remove some types in proof term (to save space) *****) @@ -490,9 +495,8 @@ | remove_types (Const (s, _)) = Const (s, dummyT) | remove_types t = t; -fun remove_types_env (Envir.Envir {iTs, asol, maxidx}) = - Envir.Envir {iTs = iTs, asol = Vartab.map (apsnd remove_types) asol, - maxidx = maxidx}; +fun remove_types_env (Envir.Envir {maxidx, tenv, tyenv}) = + Envir.Envir {maxidx = maxidx, tenv = Vartab.map (apsnd remove_types) tenv, tyenv = tyenv}; fun norm_proof' env prf = norm_proof (remove_types_env env) prf; @@ -503,40 +507,41 @@ let val n = length args; fun subst' lev (Bound i) = - (if i Bound (i-n)) (*loose: change it*) | subst' lev (Abs (a, T, body)) = Abs (a, T, subst' (lev+1) body) | subst' lev (f $ t) = (subst' lev f $ substh' lev t - handle SAME => f $ subst' lev t) - | subst' _ _ = raise SAME - and substh' lev t = (subst' lev t handle SAME => t); + handle Same.SAME => f $ subst' lev t) + | subst' _ _ = raise Same.SAME + and substh' lev t = (subst' lev t handle Same.SAME => t); - fun subst lev (AbsP (a, t, body)) = (AbsP (a, apsome' (subst' lev) t, substh lev body) - handle SAME => AbsP (a, t, subst lev body)) + fun subst lev (AbsP (a, t, body)) = + (AbsP (a, Same.map_option (subst' lev) t, substh lev body) + handle Same.SAME => AbsP (a, t, subst lev body)) | subst lev (Abst (a, T, body)) = Abst (a, T, subst (lev+1) body) | subst lev (prf %% prf') = (subst lev prf %% substh lev prf' - handle SAME => prf %% subst lev prf') + handle Same.SAME => prf %% subst lev prf') | subst lev (prf % t) = (subst lev prf % Option.map (substh' lev) t - handle SAME => prf % apsome' (subst' lev) t) - | subst _ _ = raise SAME - and substh lev prf = (subst lev prf handle SAME => prf) + handle Same.SAME => prf % Same.map_option (subst' lev) t) + | subst _ _ = raise Same.SAME + and substh lev prf = (subst lev prf handle Same.SAME => prf); in case args of [] => prf | _ => substh 0 prf end; fun prf_subst_pbounds args prf = let val n = length args; fun subst (PBound i) Plev tlev = - (if i < Plev then raise SAME (*var is locally bound*) + (if i < Plev then raise Same.SAME (*var is locally bound*) else incr_pboundvars Plev tlev (nth args (i-Plev)) handle Subscript => PBound (i-n) (*loose: change it*)) | subst (AbsP (a, t, body)) Plev tlev = AbsP (a, t, subst body (Plev+1) tlev) | subst (Abst (a, T, body)) Plev tlev = Abst (a, T, subst body Plev (tlev+1)) | subst (prf %% prf') Plev tlev = (subst prf Plev tlev %% substh prf' Plev tlev - handle SAME => prf %% subst prf' Plev tlev) + handle Same.SAME => prf %% subst prf' Plev tlev) | subst (prf % t) Plev tlev = subst prf Plev tlev % t - | subst prf _ _ = raise SAME - and substh prf Plev tlev = (subst prf Plev tlev handle SAME => prf) + | subst prf _ _ = raise Same.SAME + and substh prf Plev tlev = (subst prf Plev tlev handle Same.SAME => prf) in case args of [] => prf | _ => substh prf 0 0 end; @@ -598,14 +603,15 @@ fun implies_intr_proof h prf = let - fun abshyp i (Hyp t) = if h aconv t then PBound i else raise SAME + fun abshyp i (Hyp t) = if h aconv t then PBound i else raise Same.SAME | abshyp i (Abst (s, T, prf)) = Abst (s, T, abshyp i prf) - | abshyp i (AbsP (s, t, prf)) = AbsP (s, t, abshyp (i+1) prf) + | abshyp i (AbsP (s, t, prf)) = AbsP (s, t, abshyp (i + 1) prf) | abshyp i (prf % t) = abshyp i prf % t - | abshyp i (prf1 %% prf2) = (abshyp i prf1 %% abshyph i prf2 - handle SAME => prf1 %% abshyp i prf2) - | abshyp _ _ = raise SAME - and abshyph i prf = (abshyp i prf handle SAME => prf) + | abshyp i (prf1 %% prf2) = + (abshyp i prf1 %% abshyph i prf2 + handle Same.SAME => prf1 %% abshyp i prf2) + | abshyp _ _ = raise Same.SAME + and abshyph i prf = (abshyp i prf handle Same.SAME => prf); in AbsP ("H", NONE (*h*), abshyph 0 prf) end; @@ -624,7 +630,7 @@ (fn TFree v => if member (op =) fixed v then I else insert (op =) v | _ => I)) t []; val used = Name.context |> fold_types (fold_atyps (fn TVar ((a, _), _) => Name.declare a | _ => I)) t; - val fmap = fs ~~ (#1 (Name.variants (map fst fs) used)); + val fmap = fs ~~ #1 (Name.variants (map fst fs) used); fun thaw (f as (a, S)) = (case AList.lookup (op =) fmap f of NONE => TFree f @@ -705,30 +711,31 @@ fun lift_proof Bi inc prop prf = let - fun lift'' Us Ts t = strip_abs Ts (Logic.incr_indexes (Us, inc) (mk_abs Ts t)); + fun lift'' Us Ts t = + strip_abs Ts (Logic.incr_indexes (Us, inc) (mk_abs Ts t)); fun lift' Us Ts (Abst (s, T, prf)) = - (Abst (s, apsome' (same (op =) (Logic.incr_tvar inc)) T, lifth' Us (dummyT::Ts) prf) - handle SAME => Abst (s, T, lift' Us (dummyT::Ts) prf)) + (Abst (s, Same.map_option (Logic.incr_tvar_same inc) T, lifth' Us (dummyT::Ts) prf) + handle Same.SAME => Abst (s, T, lift' Us (dummyT::Ts) prf)) | lift' Us Ts (AbsP (s, t, prf)) = - (AbsP (s, apsome' (same (op =) (lift'' Us Ts)) t, lifth' Us Ts prf) - handle SAME => AbsP (s, t, lift' Us Ts prf)) + (AbsP (s, Same.map_option (same (op =) (lift'' Us Ts)) t, lifth' Us Ts prf) + handle Same.SAME => AbsP (s, t, lift' Us Ts prf)) | lift' Us Ts (prf % t) = (lift' Us Ts prf % Option.map (lift'' Us Ts) t - handle SAME => prf % apsome' (same (op =) (lift'' Us Ts)) t) + handle Same.SAME => prf % Same.map_option (same (op =) (lift'' Us Ts)) t) | lift' Us Ts (prf1 %% prf2) = (lift' Us Ts prf1 %% lifth' Us Ts prf2 - handle SAME => prf1 %% lift' Us Ts prf2) + handle Same.SAME => prf1 %% lift' Us Ts prf2) | lift' _ _ (PAxm (s, prop, Ts)) = - PAxm (s, prop, apsome' (same (op =) (map (Logic.incr_tvar inc))) Ts) + PAxm (s, prop, (Same.map_option o Same.map) (Logic.incr_tvar_same inc) Ts) | lift' _ _ (OfClass (T, c)) = - OfClass (same (op =) (Logic.incr_tvar inc) T, c) + OfClass (Logic.incr_tvar_same inc T, c) | lift' _ _ (Oracle (s, prop, Ts)) = - Oracle (s, prop, apsome' (same (op =) (map (Logic.incr_tvar inc))) Ts) + Oracle (s, prop, (Same.map_option o Same.map) (Logic.incr_tvar_same inc) Ts) | lift' _ _ (Promise (i, prop, Ts)) = - Promise (i, prop, same (op =) (map (Logic.incr_tvar inc)) Ts) + Promise (i, prop, Same.map (Logic.incr_tvar_same inc) Ts) | lift' _ _ (PThm (i, ((s, prop, Ts), body))) = - PThm (i, ((s, prop, apsome' (same (op =) (map (Logic.incr_tvar inc))) Ts), body)) - | lift' _ _ _ = raise SAME - and lifth' Us Ts prf = (lift' Us Ts prf handle SAME => prf); + PThm (i, ((s, prop, (Same.map_option o Same.map) (Logic.incr_tvar inc) Ts), body)) + | lift' _ _ _ = raise Same.SAME + and lifth' Us Ts prf = (lift' Us Ts prf handle Same.SAME => prf); val ps = map (Logic.lift_all inc Bi) (Logic.strip_imp_prems prop); val k = length ps; @@ -747,6 +754,11 @@ mk_AbsP (k, lift [] [] 0 0 Bi) end; +fun incr_indexes i = + map_proof_terms_option + (Same.capture (Logic.incr_indexes_same ([], i))) + (Same.capture (Logic.incr_tvar_same i)); + (***** proof by assumption *****) @@ -1082,34 +1094,47 @@ fun prf_subst (pinst, (tyinsts, insts)) = let - val substT = Envir.typ_subst_TVars tyinsts; + val substT = Envir.subst_type_same tyinsts; + val substTs = Same.map substT; - fun subst' lev (t as Var (ixn, _)) = (case AList.lookup (op =) insts ixn of - NONE => t + fun subst' lev (Var (xi, _)) = + (case AList.lookup (op =) insts xi of + NONE => raise Same.SAME | SOME u => incr_boundvars lev u) - | subst' lev (Const (s, T)) = Const (s, substT T) - | subst' lev (Free (s, T)) = Free (s, substT T) - | subst' lev (Abs (a, T, body)) = Abs (a, substT T, subst' (lev+1) body) - | subst' lev (f $ t) = subst' lev f $ subst' lev t - | subst' _ t = t; + | subst' _ (Const (s, T)) = Const (s, substT T) + | subst' _ (Free (s, T)) = Free (s, substT T) + | subst' lev (Abs (a, T, body)) = + (Abs (a, substT T, Same.commit (subst' (lev + 1)) body) + handle Same.SAME => Abs (a, T, subst' (lev + 1) body)) + | subst' lev (f $ t) = + (subst' lev f $ Same.commit (subst' lev) t + handle Same.SAME => f $ subst' lev t) + | subst' _ _ = raise Same.SAME; fun subst plev tlev (AbsP (a, t, body)) = - AbsP (a, Option.map (subst' tlev) t, subst (plev+1) tlev body) + (AbsP (a, Same.map_option (subst' tlev) t, Same.commit (subst (plev + 1) tlev) body) + handle Same.SAME => AbsP (a, t, subst (plev + 1) tlev body)) | subst plev tlev (Abst (a, T, body)) = - Abst (a, Option.map substT T, subst plev (tlev+1) body) - | subst plev tlev (prf %% prf') = subst plev tlev prf %% subst plev tlev prf' - | subst plev tlev (prf % t) = subst plev tlev prf % Option.map (subst' tlev) t - | subst plev tlev (prf as Hyp (Var (ixn, _))) = (case AList.lookup (op =) pinst ixn of - NONE => prf - | SOME prf' => incr_pboundvars plev tlev prf') - | subst _ _ (PAxm (id, prop, Ts)) = PAxm (id, prop, Option.map (map substT) Ts) + (Abst (a, Same.map_option substT T, Same.commit (subst plev (tlev + 1)) body) + handle Same.SAME => Abst (a, T, subst plev (tlev + 1) body)) + | subst plev tlev (prf %% prf') = + (subst plev tlev prf %% Same.commit (subst plev tlev) prf' + handle Same.SAME => prf %% subst plev tlev prf') + | subst plev tlev (prf % t) = + (subst plev tlev prf % Same.commit (Same.map_option (subst' tlev)) t + handle Same.SAME => prf % Same.map_option (subst' tlev) t) + | subst plev tlev (Hyp (Var (xi, _))) = + (case AList.lookup (op =) pinst xi of + NONE => raise Same.SAME + | SOME prf' => incr_pboundvars plev tlev prf') + | subst _ _ (PAxm (id, prop, Ts)) = PAxm (id, prop, Same.map_option substTs Ts) | subst _ _ (OfClass (T, c)) = OfClass (substT T, c) - | subst _ _ (Oracle (id, prop, Ts)) = Oracle (id, prop, Option.map (map substT) Ts) - | subst _ _ (Promise (i, prop, Ts)) = Promise (i, prop, (map substT) Ts) + | subst _ _ (Oracle (id, prop, Ts)) = Oracle (id, prop, Same.map_option substTs Ts) + | subst _ _ (Promise (i, prop, Ts)) = Promise (i, prop, substTs Ts) | subst _ _ (PThm (i, ((id, prop, Ts), body))) = - PThm (i, ((id, prop, Option.map (map substT) Ts), body)) - | subst _ _ t = t; - in subst 0 0 end; + PThm (i, ((id, prop, Same.map_option substTs Ts), body)) + | subst _ _ _ = raise Same.SAME; + in fn t => subst 0 0 t handle Same.SAME => t end; (*A fast unification filter: true unless the two terms cannot be unified. Terms must be NORMAL. Treats all Vars as distinct. *) @@ -1210,7 +1235,7 @@ (case skel of AbsP (_, _, skel') => skel' | _ => skel0) prf of SOME prf' => SOME (AbsP (s, t, prf')) | NONE => NONE) - | rew2 _ _ _ = NONE + | rew2 _ _ _ = NONE; in the_default prf (rew1 [] skel0 prf) end; @@ -1316,5 +1341,5 @@ end; -structure BasicProofterm : BASIC_PROOFTERM = Proofterm; -open BasicProofterm; +structure Basic_Proofterm : BASIC_PROOFTERM = Proofterm; +open Basic_Proofterm; diff -r 63686057cbe8 -r abda97d2deea src/Pure/pure_thy.ML --- a/src/Pure/pure_thy.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/pure_thy.ML Thu Jul 23 21:13:21 2009 +0200 @@ -10,8 +10,7 @@ val intern_fact: theory -> xstring -> string val defined_fact: theory -> string -> bool val hide_fact: bool -> string -> theory -> theory - val join_proofs: theory -> exn list - val cancel_proofs: theory -> unit + val join_proofs: theory -> unit val get_fact: Context.generic -> theory -> Facts.ref -> thm list val get_thms: theory -> xstring -> thm list val get_thm: theory -> xstring -> thm @@ -59,11 +58,11 @@ structure FactsData = TheoryDataFun ( - type T = Facts.T * (unit lazy list * Task_Queue.group list); - val empty = (Facts.empty, ([], [])); + type T = Facts.T * thm list; + val empty = (Facts.empty, []); val copy = I; - fun extend (facts, _) = (facts, ([], [])); - fun merge _ ((facts1, _), (facts2, _)) = (Facts.merge (facts1, facts2), ([], [])); + fun extend (facts, _) = (facts, []); + fun merge _ ((facts1, _), (facts2, _)) = (Facts.merge (facts1, facts2), []); ); @@ -79,12 +78,9 @@ (* proofs *) -fun lazy_proof thm = Lazy.lazy (fn () => Thm.join_proof thm); +fun register_proofs (thy, thms) = (FactsData.map (apsnd (append thms)) thy, thms); -fun join_proofs thy = - map_filter (Exn.get_exn o Lazy.force_result) (rev (#1 (#2 (FactsData.get thy)))); - -fun cancel_proofs thy = List.app Future.cancel_group (#2 (#2 (FactsData.get thy))); +fun join_proofs thy = Thm.join_proofs (rev (#2 (FactsData.get thy))); @@ -144,19 +140,15 @@ (* enter_thms *) -fun enter_proofs (thy, thms) = - (FactsData.map (apsnd (fn (proofs, groups) => - (fold (cons o lazy_proof) thms proofs, fold Thm.pending_groups thms groups))) thy, thms); - fun enter_thms pre_name post_name app_att (b, thms) thy = if Binding.is_empty b - then swap (enter_proofs (app_att (thy, thms))) + then swap (register_proofs (app_att (thy, thms))) else let val naming = Sign.naming_of thy; val name = NameSpace.full_name naming b; val (thy', thms') = - enter_proofs (apsnd (post_name name) (app_att (thy, pre_name name thms))); + register_proofs (apsnd (post_name name) (app_att (thy, pre_name name thms))); val thms'' = map (Thm.transfer thy') thms'; val thy'' = thy' |> (FactsData.map o apfst) (Facts.add_global naming (b, thms'') #> snd); in (thms'', thy'') end; diff -r 63686057cbe8 -r abda97d2deea src/Pure/simplifier.ML --- a/src/Pure/simplifier.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/simplifier.ML Thu Jul 23 21:13:21 2009 +0200 @@ -9,10 +9,10 @@ sig include BASIC_META_SIMPLIFIER val change_simpset: (simpset -> simpset) -> unit - val simpset_of: theory -> simpset + val global_simpset_of: theory -> simpset val Addsimprocs: simproc list -> unit val Delsimprocs: simproc list -> unit - val local_simpset_of: Proof.context -> simpset + val simpset_of: Proof.context -> simpset val generic_simp_tac: bool -> bool * bool * bool -> simpset -> int -> tactic val safe_asm_full_simp_tac: simpset -> int -> tactic val simp_tac: simpset -> int -> tactic @@ -79,7 +79,7 @@ fun pretty_ss ctxt ss = let val pretty_cterm = Syntax.pretty_term ctxt o Thm.term_of; - val pretty_thm = ProofContext.pretty_thm ctxt; + val pretty_thm = Display.pretty_thm ctxt; fun pretty_proc (name, lhss) = Pretty.big_list (name ^ ":") (map pretty_cterm lhss); fun pretty_cong (name, thm) = Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 1, pretty_thm thm]; @@ -125,7 +125,8 @@ fun map_simpset f = Context.theory_map (map_ss f); fun change_simpset f = Context.>> (Context.map_theory (map_simpset f)); -fun simpset_of thy = MetaSimplifier.context (ProofContext.init thy) (get_ss (Context.Theory thy)); +fun global_simpset_of thy = + MetaSimplifier.context (ProofContext.init thy) (get_ss (Context.Theory thy)); fun Addsimprocs args = change_simpset (fn ss => ss addsimprocs args); fun Delsimprocs args = change_simpset (fn ss => ss delsimprocs args); @@ -133,10 +134,10 @@ (* local simpset *) -fun local_simpset_of ctxt = MetaSimplifier.context ctxt (get_ss (Context.Proof ctxt)); +fun simpset_of ctxt = MetaSimplifier.context ctxt (get_ss (Context.Proof ctxt)); val _ = ML_Antiquote.value "simpset" - (Scan.succeed "Simplifier.local_simpset_of (ML_Context.the_local_context ())"); + (Scan.succeed "Simplifier.simpset_of (ML_Context.the_local_context ())"); @@ -331,7 +332,7 @@ val simplified = conv_mode -- Attrib.thms >> (fn (f, ths) => Thm.rule_attribute (fn context => f ((if null ths then I else MetaSimplifier.clear_ss) - (local_simpset_of (Context.proof_of context)) addsimps ths))); + (simpset_of (Context.proof_of context)) addsimps ths))); end; @@ -390,12 +391,12 @@ Method.setup (Binding.name simpN) (simp_method more_mods (fn ctxt => fn tac => fn facts => HEADGOAL (Method.insert_tac facts THEN' - (CHANGED_PROP oo tac) (local_simpset_of ctxt)))) + (CHANGED_PROP oo tac) (simpset_of ctxt)))) "simplification" #> Method.setup (Binding.name "simp_all") (simp_method more_mods (fn ctxt => fn tac => fn facts => ALLGOALS (Method.insert_tac facts) THEN - (CHANGED_PROP o ALLGOALS o tac) (local_simpset_of ctxt))) + (CHANGED_PROP o ALLGOALS o tac) (simpset_of ctxt))) "simplification (all goals)"; fun easy_setup reflect trivs = method_setup [] #> Context.theory_map (map_ss (fn _ => diff -r 63686057cbe8 -r abda97d2deea src/Pure/tctical.ML --- a/src/Pure/tctical.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/tctical.ML Thu Jul 23 21:13:21 2009 +0200 @@ -191,12 +191,11 @@ (*** Tracing tactics ***) (*Print the current proof state and pass it on.*) -fun print_tac msg = - (fn st => - (tracing msg; - tracing ((Pretty.string_of o Pretty.chunks o - Display.pretty_goals (! Display.goals_limit)) st); - Seq.single st)); +fun print_tac msg st = + (tracing (msg ^ "\n" ^ + Pretty.string_of (Pretty.chunks + (Display_Goal.pretty_goals_without_context (! Display_Goal.goals_limit) st))); + Seq.single st); (*Pause until a line is typed -- if non-empty then fail. *) fun pause_tac st = @@ -232,10 +231,10 @@ (*Extract from a tactic, a thm->thm seq function that handles tracing*) fun tracify flag tac st = - if !flag andalso not (!suppress_tracing) - then (Display.print_goals (! Display.goals_limit) st; - tracing "** Press RETURN to continue:"; - exec_trace_command flag (tac,st)) + if !flag andalso not (!suppress_tracing) then + (Display_Goal.print_goals_without_context (! Display_Goal.goals_limit) st; + tracing "** Press RETURN to continue:"; + exec_trace_command flag (tac, st)) else tac st; (*Create a tactic whose outcome is given by seqf, handling TRACE_EXIT*) diff -r 63686057cbe8 -r abda97d2deea src/Pure/term.ML --- a/src/Pure/term.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/term.ML Thu Jul 23 21:13:21 2009 +0200 @@ -634,31 +634,31 @@ *) fun subst_bounds (args: term list, t) : term = let - exception SAME; val n = length args; fun subst (t as Bound i, lev) = - (if i < lev then raise SAME (*var is locally bound*) + (if i < lev then raise Same.SAME (*var is locally bound*) else incr_boundvars lev (nth args (i - lev)) handle Subscript => Bound (i - n)) (*loose: change it*) | subst (Abs (a, T, body), lev) = Abs (a, T, subst (body, lev + 1)) | subst (f $ t, lev) = - (subst (f, lev) $ (subst (t, lev) handle SAME => t) handle SAME => f $ subst (t, lev)) - | subst _ = raise SAME; - in case args of [] => t | _ => (subst (t, 0) handle SAME => t) end; + (subst (f, lev) $ (subst (t, lev) handle Same.SAME => t) + handle Same.SAME => f $ subst (t, lev)) + | subst _ = raise Same.SAME; + in case args of [] => t | _ => (subst (t, 0) handle Same.SAME => t) end; (*Special case: one argument*) fun subst_bound (arg, t) : term = let - exception SAME; fun subst (Bound i, lev) = - if i < lev then raise SAME (*var is locally bound*) + if i < lev then raise Same.SAME (*var is locally bound*) else if i = lev then incr_boundvars lev arg else Bound (i - 1) (*loose: change it*) | subst (Abs (a, T, body), lev) = Abs (a, T, subst (body, lev + 1)) | subst (f $ t, lev) = - (subst (f, lev) $ (subst (t, lev) handle SAME => t) handle SAME => f $ subst (t, lev)) - | subst _ = raise SAME; - in subst (t, 0) handle SAME => t end; + (subst (f, lev) $ (subst (t, lev) handle Same.SAME => t) + handle Same.SAME => f $ subst (t, lev)) + | subst _ = raise Same.SAME; + in subst (t, 0) handle Same.SAME => t end; (*beta-reduce if possible, else form application*) fun betapply (Abs(_,_,t), u) = subst_bound (u,t) @@ -708,15 +708,16 @@ The resulting term is ready to become the body of an Abs.*) fun abstract_over (v, body) = let - exception SAME; fun abs lev tm = if v aconv tm then Bound lev else (case tm of Abs (a, T, t) => Abs (a, T, abs (lev + 1) t) - | t $ u => (abs lev t $ (abs lev u handle SAME => u) handle SAME => t $ abs lev u) - | _ => raise SAME); - in abs 0 body handle SAME => body end; + | t $ u => + (abs lev t $ (abs lev u handle Same.SAME => u) + handle Same.SAME => t $ abs lev u) + | _ => raise Same.SAME); + in abs 0 body handle Same.SAME => body end; fun lambda v t = let val x = diff -r 63686057cbe8 -r abda97d2deea src/Pure/term_subst.ML --- a/src/Pure/term_subst.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/term_subst.ML Thu Jul 23 21:13:21 2009 +0200 @@ -1,11 +1,14 @@ (* Title: Pure/term_subst.ML Author: Makarius -Efficient type/term substitution -- avoids copying. +Efficient type/term substitution. *) signature TERM_SUBST = sig + val map_atypsT_same: typ Same.operation -> typ Same.operation + val map_types_same: typ Same.operation -> term Same.operation + val map_aterms_same: term Same.operation -> term Same.operation val map_atypsT_option: (typ -> typ option) -> typ -> typ option val map_atyps_option: (typ -> typ option) -> term -> term option val map_types_option: (typ -> typ option) -> term -> term option @@ -32,29 +35,12 @@ structure Term_Subst: TERM_SUBST = struct -(* indication of same result *) - -exception SAME; - -fun same_fn f x = - (case f x of - NONE => raise SAME - | SOME y => y); - -fun option_fn f x = - SOME (f x) handle SAME => NONE; - - (* generic mapping *) -local - fun map_atypsT_same f = let - fun typ (Type (a, Ts)) = Type (a, typs Ts) - | typ T = f T - and typs (T :: Ts) = (typ T :: (typs Ts handle SAME => Ts) handle SAME => T :: typs Ts) - | typs [] = raise SAME; + fun typ (Type (a, Ts)) = Type (a, Same.map typ Ts) + | typ T = f T; in typ end; fun map_types_same f = @@ -62,72 +48,64 @@ fun term (Const (a, T)) = Const (a, f T) | term (Free (a, T)) = Free (a, f T) | term (Var (v, T)) = Var (v, f T) - | term (Bound _) = raise SAME + | term (Bound _) = raise Same.SAME | term (Abs (x, T, t)) = - (Abs (x, f T, term t handle SAME => t) - handle SAME => Abs (x, T, term t)) - | term (t $ u) = (term t $ (term u handle SAME => u) handle SAME => t $ term u); + (Abs (x, f T, Same.commit term t) + handle Same.SAME => Abs (x, T, term t)) + | term (t $ u) = (term t $ Same.commit term u handle Same.SAME => t $ term u); in term end; fun map_aterms_same f = let fun term (Abs (x, T, t)) = Abs (x, T, term t) - | term (t $ u) = (term t $ (term u handle SAME => u) handle SAME => t $ term u) + | term (t $ u) = (term t $ Same.commit term u handle Same.SAME => t $ term u) | term a = f a; in term end; -in - -val map_atypsT_option = option_fn o map_atypsT_same o same_fn; -val map_atyps_option = option_fn o map_types_same o map_atypsT_same o same_fn; -val map_types_option = option_fn o map_types_same o same_fn; -val map_aterms_option = option_fn o map_aterms_same o same_fn; - -end; +val map_atypsT_option = Same.capture o map_atypsT_same o Same.function; +val map_atyps_option = Same.capture o map_types_same o map_atypsT_same o Same.function; +val map_types_option = Same.capture o map_types_same o Same.function; +val map_aterms_option = Same.capture o map_aterms_same o Same.function; (* generalization of fixed variables *) local -fun generalizeT_same [] _ _ = raise SAME +fun generalizeT_same [] _ _ = raise Same.SAME | generalizeT_same tfrees idx ty = let - fun gen_typ (Type (a, Ts)) = Type (a, gen_typs Ts) - | gen_typ (TFree (a, S)) = + fun gen (Type (a, Ts)) = Type (a, Same.map gen Ts) + | gen (TFree (a, S)) = if member (op =) tfrees a then TVar ((a, idx), S) - else raise SAME - | gen_typ _ = raise SAME - and gen_typs (T :: Ts) = - (gen_typ T :: (gen_typs Ts handle SAME => Ts) - handle SAME => T :: gen_typs Ts) - | gen_typs [] = raise SAME; - in gen_typ ty end; + else raise Same.SAME + | gen _ = raise Same.SAME; + in gen ty end; -fun generalize_same ([], []) _ _ = raise SAME +fun generalize_same ([], []) _ _ = raise Same.SAME | generalize_same (tfrees, frees) idx tm = let val genT = generalizeT_same tfrees idx; fun gen (Free (x, T)) = if member (op =) frees x then - Var (Name.clean_index (x, idx), genT T handle SAME => T) + Var (Name.clean_index (x, idx), Same.commit genT T) else Free (x, genT T) | gen (Var (xi, T)) = Var (xi, genT T) | gen (Const (c, T)) = Const (c, genT T) - | gen (Bound _) = raise SAME + | gen (Bound _) = raise Same.SAME | gen (Abs (x, T, t)) = - (Abs (x, genT T, gen t handle SAME => t) - handle SAME => Abs (x, T, gen t)) - | gen (t $ u) = (gen t $ (gen u handle SAME => u) handle SAME => t $ gen u); + (Abs (x, genT T, Same.commit gen t) + handle Same.SAME => Abs (x, T, gen t)) + | gen (t $ u) = (gen t $ Same.commit gen u handle Same.SAME => t $ gen u); in gen tm end; in -fun generalize names i tm = generalize_same names i tm handle SAME => tm; -fun generalizeT names i ty = generalizeT_same names i ty handle SAME => ty; +fun generalize names i tm = generalize_same names i tm handle Same.SAME => tm; +fun generalizeT names i ty = generalizeT_same names i ty handle Same.SAME => ty; -fun generalize_option names i tm = SOME (generalize_same names i tm) handle SAME => NONE; -fun generalizeT_option names i ty = SOME (generalizeT_same names i ty) handle SAME => NONE; +fun generalize_option names i tm = SOME (generalize_same names i tm) handle Same.SAME => NONE; +fun generalizeT_option names i ty = SOME (generalizeT_same names i ty) handle Same.SAME => NONE; end; @@ -148,12 +126,12 @@ | subst_typ (TVar ((a, i), S)) = (case AList.lookup Term.eq_tvar instT ((a, i), S) of SOME (T, j) => (maxify j; T) - | NONE => (maxify i; raise SAME)) - | subst_typ _ = raise SAME + | NONE => (maxify i; raise Same.SAME)) + | subst_typ _ = raise Same.SAME and subst_typs (T :: Ts) = - (subst_typ T :: (subst_typs Ts handle SAME => Ts) - handle SAME => T :: subst_typs Ts) - | subst_typs [] = raise SAME; + (subst_typ T :: Same.commit subst_typs Ts + handle Same.SAME => T :: subst_typs Ts) + | subst_typs [] = raise Same.SAME; in subst_typ ty end; fun instantiate_same maxidx (instT, inst) tm = @@ -164,43 +142,43 @@ fun subst (Const (c, T)) = Const (c, substT T) | subst (Free (x, T)) = Free (x, substT T) | subst (Var ((x, i), T)) = - let val (T', same) = (substT T, false) handle SAME => (T, true) in + let val (T', same) = (substT T, false) handle Same.SAME => (T, true) in (case AList.lookup Term.eq_var inst ((x, i), T') of SOME (t, j) => (maxify j; t) - | NONE => (maxify i; if same then raise SAME else Var ((x, i), T'))) + | NONE => (maxify i; if same then raise Same.SAME else Var ((x, i), T'))) end - | subst (Bound _) = raise SAME + | subst (Bound _) = raise Same.SAME | subst (Abs (x, T, t)) = - (Abs (x, substT T, subst t handle SAME => t) - handle SAME => Abs (x, T, subst t)) - | subst (t $ u) = (subst t $ (subst u handle SAME => u) handle SAME => t $ subst u); + (Abs (x, substT T, Same.commit subst t) + handle Same.SAME => Abs (x, T, subst t)) + | subst (t $ u) = (subst t $ Same.commit subst u handle Same.SAME => t $ subst u); in subst tm end; in fun instantiateT_maxidx instT ty i = let val maxidx = ref i - in (instantiateT_same maxidx instT ty handle SAME => ty, ! maxidx) end; + in (instantiateT_same maxidx instT ty handle Same.SAME => ty, ! maxidx) end; fun instantiate_maxidx insts tm i = let val maxidx = ref i - in (instantiate_same maxidx insts tm handle SAME => tm, ! maxidx) end; + in (instantiate_same maxidx insts tm handle Same.SAME => tm, ! maxidx) end; fun instantiateT [] ty = ty | instantiateT instT ty = - (instantiateT_same (ref ~1) (no_indexes1 instT) ty handle SAME => ty); + (instantiateT_same (ref ~1) (no_indexes1 instT) ty handle Same.SAME => ty); fun instantiate ([], []) tm = tm | instantiate insts tm = - (instantiate_same (ref ~1) (no_indexes2 insts) tm handle SAME => tm); + (instantiate_same (ref ~1) (no_indexes2 insts) tm handle Same.SAME => tm); fun instantiateT_option [] _ = NONE | instantiateT_option instT ty = - (SOME (instantiateT_same (ref ~1) (no_indexes1 instT) ty) handle SAME => NONE); + (SOME (instantiateT_same (ref ~1) (no_indexes1 instT) ty) handle Same.SAME => NONE); fun instantiate_option ([], []) _ = NONE | instantiate_option insts tm = - (SOME (instantiate_same (ref ~1) (no_indexes2 insts) tm) handle SAME => NONE); + (SOME (instantiate_same (ref ~1) (no_indexes2 insts) tm) handle Same.SAME => NONE); end; diff -r 63686057cbe8 -r abda97d2deea src/Pure/thm.ML --- a/src/Pure/thm.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/thm.ML Thu Jul 23 21:13:21 2009 +0200 @@ -141,12 +141,11 @@ val bicompose: bool -> bool * thm * int -> int -> thm -> thm Seq.seq val biresolution: bool -> (bool * thm) list -> int -> thm -> thm Seq.seq val rename_boundvars: term -> term -> thm -> thm - val future: thm future -> cterm -> thm - val pending_groups: thm -> Task_Queue.group list -> Task_Queue.group list - val status_of: thm -> {oracle: bool, unfinished: bool, failed: bool} + val join_proofs: thm list -> unit val proof_body_of: thm -> proof_body val proof_of: thm -> proof - val join_proof: thm -> unit + val status_of: thm -> {oracle: bool, unfinished: bool, failed: bool} + val future: thm future -> cterm -> thm val get_name: thm -> string val put_name: string -> thm -> thm val extern_oracles: theory -> xstring list @@ -317,7 +316,7 @@ (Ctyp {T = TVar ((a, i), S), thy_ref = Theory.check_thy thy, maxidx = i, sorts = sorts}, Ctyp {T = T, thy_ref = Theory.check_thy thy, maxidx = maxidx2, sorts = sorts}); fun mk_ctinst ((x, i), (T, t)) = - let val T = Envir.typ_subst_TVars Tinsts T in + let val T = Envir.subst_type Tinsts T in (Cterm {t = Var ((x, i), T), T = T, thy_ref = Theory.check_thy thy, maxidx = i, sorts = sorts}, Cterm {t = t, T = T, thy_ref = Theory.check_thy thy, maxidx = maxidx2, sorts = sorts}) @@ -345,9 +344,7 @@ tpairs: (term * term) list, (*flex-flex pairs*) prop: term} (*conclusion*) and deriv = Deriv of - {max_promise: serial, - open_promises: (serial * thm future) OrdList.T, - promises: (serial * thm future) OrdList.T, + {promises: (serial * thm future) OrdList.T, body: Pt.proof_body}; type conv = cterm -> thm; @@ -504,11 +501,10 @@ (** derivations **) -fun make_deriv max_promise open_promises promises oracles thms proof = - Deriv {max_promise = max_promise, open_promises = open_promises, promises = promises, - body = PBody {oracles = oracles, thms = thms, proof = proof}}; +fun make_deriv promises oracles thms proof = + Deriv {promises = promises, body = PBody {oracles = oracles, thms = thms, proof = proof}}; -val empty_deriv = make_deriv ~1 [] [] [] [] Pt.MinProof; +val empty_deriv = make_deriv [] [] [] Pt.MinProof; (* inference rules *) @@ -516,13 +512,9 @@ fun promise_ord ((i, _), (j, _)) = int_ord (j, i); fun deriv_rule2 f - (Deriv {max_promise = max1, open_promises = open_ps1, promises = ps1, - body = PBody {oracles = oras1, thms = thms1, proof = prf1}}) - (Deriv {max_promise = max2, open_promises = open_ps2, promises = ps2, - body = PBody {oracles = oras2, thms = thms2, proof = prf2}}) = + (Deriv {promises = ps1, body = PBody {oracles = oras1, thms = thms1, proof = prf1}}) + (Deriv {promises = ps2, body = PBody {oracles = oras2, thms = thms2, proof = prf2}}) = let - val max = Int.max (max1, max2); - val open_ps = OrdList.union promise_ord open_ps1 open_ps2; val ps = OrdList.union promise_ord ps1 ps2; val oras = Pt.merge_oracles oras1 oras2; val thms = Pt.merge_thms thms1 thms2; @@ -532,10 +524,10 @@ | 1 => MinProof | 0 => MinProof | i => error ("Illegal level of detail for proof objects: " ^ string_of_int i)); - in make_deriv max open_ps ps oras thms prf end; + in make_deriv ps oras thms prf end; fun deriv_rule1 f = deriv_rule2 (K f) empty_deriv; -fun deriv_rule0 prf = deriv_rule1 I (make_deriv ~1 [] [] [] [] prf); +fun deriv_rule0 prf = deriv_rule1 I (make_deriv [] [] [] prf); @@ -1232,7 +1224,7 @@ if i < 0 then raise THM ("negative increment", 0, [thm]) else if i = 0 then thm else - Thm (deriv_rule1 (Pt.map_proof_terms (Logic.incr_indexes ([], i)) (Logic.incr_tvar i)) der, + Thm (deriv_rule1 (Pt.incr_indexes i) der, {thy_ref = thy_ref, tags = [], maxidx = maxidx + i, @@ -1247,12 +1239,12 @@ val Thm (der, {thy_ref, maxidx, shyps, hyps, prop, ...}) = state; val thy = Theory.deref thy_ref; val (tpairs, Bs, Bi, C) = dest_state (state, i); - fun newth n (env as Envir.Envir {maxidx, ...}, tpairs) = + fun newth n (env, tpairs) = Thm (deriv_rule1 ((if Envir.is_empty env then I else (Pt.norm_proof' env)) o Pt.assumption_proof Bs Bi n) der, {tags = [], - maxidx = maxidx, + maxidx = Envir.maxidx_of env, shyps = Envir.insert_sorts env shyps, hyps = hyps, tpairs = @@ -1465,15 +1457,15 @@ (*Faster normalization: skip assumptions that were lifted over*) fun norm_term_skip env 0 t = Envir.norm_term env t - | norm_term_skip env n (Const("all",_)$Abs(a,T,t)) = - let val Envir.Envir{iTs, ...} = env - val T' = Envir.typ_subst_TVars iTs T - (*Must instantiate types of parameters because they are flattened; - this could be a NEW parameter*) - in Term.all T' $ Abs(a, T', norm_term_skip env n t) end - | norm_term_skip env n (Const("==>", _) $ A $ B) = - Logic.mk_implies (A, norm_term_skip env (n-1) B) - | norm_term_skip env n t = error"norm_term_skip: too few assumptions??"; + | norm_term_skip env n (Const ("all", _) $ Abs (a, T, t)) = + let + val T' = Envir.subst_type (Envir.type_env env) T + (*Must instantiate types of parameters because they are flattened; + this could be a NEW parameter*) + in Term.all T' $ Abs (a, T', norm_term_skip env n t) end + | norm_term_skip env n (Const ("==>", _) $ A $ B) = + Logic.mk_implies (A, norm_term_skip env (n - 1) B) + | norm_term_skip env n t = error "norm_term_skip: too few assumptions??"; (*Composition of object rule r=(A1...Am/B) with proof state s=(B1...Bn/C) @@ -1495,7 +1487,7 @@ and nlift = Logic.count_prems (strip_all_body Bi) + (if eres_flg then ~1 else 0) val thy = Theory.deref (merge_thys2 state orule); (** Add new theorem with prop = '[| Bs; As |] ==> C' to thq **) - fun addth A (As, oldAs, rder', n) ((env as Envir.Envir {maxidx, ...}, tpairs), thq) = + fun addth A (As, oldAs, rder', n) ((env, tpairs), thq) = let val normt = Envir.norm_term env; (*perform minimal copying here by examining env*) val (ntpairs, normp) = @@ -1520,7 +1512,7 @@ curry op oo (Pt.norm_proof' env)) (Pt.bicompose_proof flatten Bs oldAs As A n (nlift+1))) rder' sder, {tags = [], - maxidx = maxidx, + maxidx = Envir.maxidx_of env, shyps = Envir.insert_sorts env (Sorts.union rshyps sshyps), hyps = union_hyps rhyps shyps, tpairs = ntpairs, @@ -1614,6 +1606,36 @@ (*** Future theorems -- proofs with promises ***) +(* fulfilled proofs *) + +fun raw_body (Thm (Deriv {body, ...}, _)) = body; + +fun fulfill_body (Thm (Deriv {promises, body}, {thy_ref, ...})) = + Pt.fulfill_proof (Theory.deref thy_ref) + (map #1 promises ~~ fulfill_bodies (map #2 promises)) body +and fulfill_bodies futures = map fulfill_body (Exn.release_all (Future.join_results futures)); + +val join_proofs = Pt.join_bodies o map fulfill_body; + +fun proof_body_of thm = (Pt.join_bodies [raw_body thm]; fulfill_body thm); +val proof_of = Pt.proof_of o proof_body_of; + + +(* derivation status *) + +fun status_of (Thm (Deriv {promises, body}, _)) = + let + val ps = map (Future.peek o snd) promises; + val bodies = body :: + map_filter (fn SOME (Exn.Result th) => SOME (raw_body th) | _ => NONE) ps; + val {oracle, unfinished, failed} = Pt.status_of bodies; + in + {oracle = oracle, + unfinished = unfinished orelse exists is_none ps, + failed = failed orelse exists (fn SOME (Exn.Exn _) => true | _ => false) ps} + end; + + (* future rule *) fun future_result i orig_thy orig_shyps orig_prop raw_thm = @@ -1623,12 +1645,13 @@ val _ = Theory.check_thy orig_thy; fun err msg = raise THM ("future_result: " ^ msg, 0, [thm]); - val Thm (Deriv {max_promise, ...}, {shyps, hyps, tpairs, prop, ...}) = thm; + val Thm (Deriv {promises, ...}, {shyps, hyps, tpairs, prop, ...}) = thm; val _ = prop aconv orig_prop orelse err "bad prop"; val _ = null tpairs orelse err "bad tpairs"; val _ = null hyps orelse err "bad hyps"; val _ = Sorts.subset (shyps, orig_shyps) orelse err "bad shyps"; - val _ = max_promise < i orelse err "bad dependencies"; + val _ = forall (fn (j, _) => i <> j) promises orelse err "bad dependencies"; + val _ = fulfill_bodies (map #2 promises); in thm end; fun future future_thm ct = @@ -1639,9 +1662,8 @@ val i = serial (); val future = future_thm |> Future.map (future_result i thy sorts prop); - val promise = (i, future); in - Thm (make_deriv i [promise] [promise] [] [] (Pt.promise_proof thy i prop), + Thm (make_deriv [(i, future)] [] [] (Pt.promise_proof thy i prop), {thy_ref = thy_ref, tags = [], maxidx = maxidx, @@ -1652,57 +1674,21 @@ end; -(* derivation status *) - -fun raw_proof_body_of (Thm (Deriv {body, ...}, _)) = body; -val raw_proof_of = Proofterm.proof_of o raw_proof_body_of; - -fun pending_groups (Thm (Deriv {open_promises, ...}, _)) = - fold (insert Task_Queue.eq_group o Future.group_of o #2) open_promises; - -fun status_of (Thm (Deriv {promises, body, ...}, _)) = - let - val ps = map (Future.peek o snd) promises; - val bodies = body :: - map_filter (fn SOME (Exn.Result th) => SOME (raw_proof_body_of th) | _ => NONE) ps; - val {oracle, unfinished, failed} = Pt.status_of bodies; - in - {oracle = oracle, - unfinished = unfinished orelse exists is_none ps, - failed = failed orelse exists (fn SOME (Exn.Exn _) => true | _ => false) ps} - end; - - -(* fulfilled proofs *) - -fun proof_body_of (Thm (Deriv {open_promises, promises, body, ...}, {thy_ref, ...})) = - let - val _ = Exn.release_all (map (Future.join_result o #2) (rev open_promises)); - val ps = map (apsnd (raw_proof_body_of o Future.join)) promises; - in Pt.fulfill_proof (Theory.deref thy_ref) ps body end; - -val proof_of = Proofterm.proof_of o proof_body_of; -val join_proof = ignore o proof_body_of; - - (* closed derivations with official name *) fun get_name thm = - Pt.get_name (hyps_of thm) (prop_of thm) (raw_proof_of thm); + Pt.get_name (hyps_of thm) (prop_of thm) (Pt.proof_of (raw_body thm)); fun put_name name (thm as Thm (der, args)) = let - val Deriv {max_promise, open_promises, promises, body, ...} = der; + val Deriv {promises, body} = der; val {thy_ref, hyps, prop, tpairs, ...} = args; val _ = null tpairs orelse raise THM ("put_name: unsolved flex-flex constraints", 0, [thm]); val ps = map (apsnd (Future.map proof_body_of)) promises; val thy = Theory.deref thy_ref; val (pthm, proof) = Pt.thm_proof thy name hyps prop ps body; - - val open_promises' = open_promises |> filter (fn (_, p) => - (case Future.peek p of SOME (Exn.Result _) => false | _ => true)); - val der' = make_deriv max_promise open_promises' [] [] [pthm] proof; + val der' = make_deriv [] [] [pthm] proof; val _ = Theory.check_thy thy; in Thm (der', args) end; @@ -1718,7 +1704,7 @@ raise THM ("Oracle's result must have type prop: " ^ name, 0, []) else let val (ora, prf) = Pt.oracle_proof name prop in - Thm (make_deriv ~1 [] [] [ora] [] prf, + Thm (make_deriv [] [ora] [] prf, {thy_ref = Theory.merge_refs (thy_ref1, thy_ref2), tags = [], maxidx = maxidx, @@ -1756,5 +1742,5 @@ end; -structure BasicThm: BASIC_THM = Thm; -open BasicThm; +structure Basic_Thm: BASIC_THM = Thm; +open Basic_Thm; diff -r 63686057cbe8 -r abda97d2deea src/Pure/type.ML --- a/src/Pure/type.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/type.ML Thu Jul 23 21:13:21 2009 +0200 @@ -494,12 +494,15 @@ (*equality with respect to a type environment*) -fun eq_type tye (T, T') = +fun equal_type tye (T, T') = (case (devar tye T, devar tye T') of (Type (s, Ts), Type (s', Ts')) => - s = s' andalso ListPair.all (eq_type tye) (Ts, Ts') + s = s' andalso ListPair.all (equal_type tye) (Ts, Ts') | (U, U') => U = U'); +fun eq_type tye = + if Vartab.is_empty tye then op = else equal_type tye; + (** extend and merge type signatures **) diff -r 63686057cbe8 -r abda97d2deea src/Pure/type_infer.ML --- a/src/Pure/type_infer.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/type_infer.ML Thu Jul 23 21:13:21 2009 +0200 @@ -40,7 +40,9 @@ fun param i (x, S) = TVar (("?" ^ x, i), S); val paramify_vars = - perhaps (Term_Subst.map_atypsT_option (fn TVar ((x, i), S) => SOME (param i (x, S)) | _ => NONE)); + Same.commit + (Term_Subst.map_atypsT_same + (fn TVar ((x, i), S) => (param i (x, S)) | _ => raise Same.SAME)); val paramify_dummies = let @@ -71,13 +73,12 @@ (** pretyps and preterms **) -(*links to parameters may get instantiated, anything else is rigid*) +(*parameters may get instantiated, anything else is rigid*) datatype pretyp = PType of string * pretyp list | PTFree of string * sort | PTVar of indexname * sort | - Param of sort | - Link of pretyp ref; + Param of int * sort; datatype preterm = PConst of string * pretyp | @@ -91,11 +92,11 @@ (* utils *) -val mk_param = Link o ref o Param; - -fun deref (T as Link (ref (Param _))) = T - | deref (Link (ref T)) = deref T - | deref T = T; +fun deref tye (T as Param (i, S)) = + (case Inttab.lookup tye i of + NONE => T + | SOME U => deref tye U) + | deref tye T = T; fun fold_pretyps f (PConst (_, T)) x = f T x | fold_pretyps f (PFree (_, T)) x = f T x @@ -111,46 +112,50 @@ (* pretyp_of *) -fun pretyp_of is_para typ params = +fun pretyp_of is_para typ params_idx = let - val params' = fold_atyps + val (params', idx) = fold_atyps (fn TVar (xi as (x, _), S) => - (fn ps => + (fn ps_idx as (ps, idx) => if is_para xi andalso not (Vartab.defined ps xi) - then Vartab.update (xi, mk_param S) ps else ps) - | _ => I) typ params; + then (Vartab.update (xi, Param (idx, S)) ps, idx + 1) else ps_idx) + | _ => I) typ params_idx; - fun pre_of (TVar (v as (xi, _))) = + fun pre_of (TVar (v as (xi, _))) idx = (case Vartab.lookup params' xi of NONE => PTVar v - | SOME p => p) - | pre_of (TFree ("'_dummy_", S)) = mk_param S - | pre_of (TFree v) = PTFree v - | pre_of (T as Type (a, Ts)) = - if T = dummyT then mk_param [] - else PType (a, map pre_of Ts); - in (pre_of typ, params') end; + | SOME p => p, idx) + | pre_of (TFree ("'_dummy_", S)) idx = (Param (idx, S), idx + 1) + | pre_of (TFree v) idx = (PTFree v, idx) + | pre_of (T as Type (a, Ts)) idx = + if T = dummyT then (Param (idx, []), idx + 1) + else + let val (Ts', idx') = fold_map pre_of Ts idx + in (PType (a, Ts'), idx') end; + + val (ptyp, idx') = pre_of typ idx; + in (ptyp, (params', idx')) end; (* preterm_of *) -fun preterm_of const_type is_para tm (vparams, params) = +fun preterm_of const_type is_para tm (vparams, params, idx) = let - fun add_vparm xi ps = + fun add_vparm xi (ps_idx as (ps, idx)) = if not (Vartab.defined ps xi) then - Vartab.update (xi, mk_param []) ps - else ps; + (Vartab.update (xi, Param (idx, [])) ps, idx + 1) + else ps_idx; - val vparams' = fold_aterms + val (vparams', idx') = fold_aterms (fn Var (_, Type ("_polymorphic_", _)) => I | Var (xi, _) => add_vparm xi | Free (x, _) => add_vparm (x, ~1) | _ => I) - tm vparams; + tm (vparams, idx); fun var_param xi = the (Vartab.lookup vparams' xi); val preT_of = pretyp_of is_para; - fun polyT_of T = fst (pretyp_of (K true) T Vartab.empty); + fun polyT_of T idx = apsnd snd (pretyp_of (K true) T (Vartab.empty, idx)); fun constraint T t ps = if T = dummyT then (t, ps) @@ -158,29 +163,33 @@ let val (T', ps') = preT_of T ps in (Constraint (t, T'), ps') end; - fun pre_of (Const (c, T)) ps = + fun pre_of (Const (c, T)) (ps, idx) = (case const_type c of - SOME U => constraint T (PConst (c, polyT_of U)) ps + SOME U => + let val (pU, idx') = polyT_of U idx + in constraint T (PConst (c, pU)) (ps, idx') end | NONE => raise TYPE ("No such constant: " ^ quote c, [], [])) - | pre_of (Var (xi, Type ("_polymorphic_", [T]))) ps = (PVar (xi, polyT_of T), ps) - | pre_of (Var (xi, T)) ps = constraint T (PVar (xi, var_param xi)) ps - | pre_of (Free (x, T)) ps = constraint T (PFree (x, var_param (x, ~1))) ps - | pre_of (Const ("_type_constraint_", Type ("fun", [T, _])) $ t) ps = - pre_of t ps |-> constraint T - | pre_of (Bound i) ps = (PBound i, ps) - | pre_of (Abs (x, T, t)) ps = + | pre_of (Var (xi, Type ("_polymorphic_", [T]))) (ps, idx) = + let val (pT, idx') = polyT_of T idx + in (PVar (xi, pT), (ps, idx')) end + | pre_of (Var (xi, T)) ps_idx = constraint T (PVar (xi, var_param xi)) ps_idx + | pre_of (Free (x, T)) ps_idx = constraint T (PFree (x, var_param (x, ~1))) ps_idx + | pre_of (Const ("_type_constraint_", Type ("fun", [T, _])) $ t) ps_idx = + pre_of t ps_idx |-> constraint T + | pre_of (Bound i) ps_idx = (PBound i, ps_idx) + | pre_of (Abs (x, T, t)) ps_idx = let - val (T', ps') = preT_of T ps; - val (t', ps'') = pre_of t ps'; - in (PAbs (x, T', t'), ps'') end - | pre_of (t $ u) ps = + val (T', ps_idx') = preT_of T ps_idx; + val (t', ps_idx'') = pre_of t ps_idx'; + in (PAbs (x, T', t'), ps_idx'') end + | pre_of (t $ u) ps_idx = let - val (t', ps') = pre_of t ps; - val (u', ps'') = pre_of u ps'; - in (PAppl (t', u'), ps'') end; + val (t', ps_idx') = pre_of t ps_idx; + val (u', ps_idx'') = pre_of u ps_idx'; + in (PAppl (t', u'), ps_idx'') end; - val (tm', params') = pre_of tm params; - in (tm', (vparams', params')) end; + val (tm', (params', idx'')) = pre_of tm (params, idx'); + in (tm', (vparams', params', idx'')) end; @@ -188,62 +197,64 @@ (* add_parms *) -fun add_parmsT (PType (_, Ts)) rs = fold add_parmsT Ts rs - | add_parmsT (Link (r as ref (Param _))) rs = insert (op =) r rs - | add_parmsT (Link (ref T)) rs = add_parmsT T rs - | add_parmsT _ rs = rs; +fun add_parmsT tye T = + (case deref tye T of + PType (_, Ts) => fold (add_parmsT tye) Ts + | Param (i, _) => insert (op =) i + | _ => I); -val add_parms = fold_pretyps add_parmsT; +fun add_parms tye = fold_pretyps (add_parmsT tye); (* add_names *) -fun add_namesT (PType (_, Ts)) = fold add_namesT Ts - | add_namesT (PTFree (x, _)) = Name.declare x - | add_namesT (PTVar ((x, _), _)) = Name.declare x - | add_namesT (Link (ref T)) = add_namesT T - | add_namesT (Param _) = I; +fun add_namesT tye T = + (case deref tye T of + PType (_, Ts) => fold (add_namesT tye) Ts + | PTFree (x, _) => Name.declare x + | PTVar ((x, _), _) => Name.declare x + | Param _ => I); -val add_names = fold_pretyps add_namesT; +fun add_names tye = fold_pretyps (add_namesT tye); (* simple_typ/term_of *) -(*deref links, fail on params*) -fun simple_typ_of (PType (a, Ts)) = Type (a, map simple_typ_of Ts) - | simple_typ_of (PTFree v) = TFree v - | simple_typ_of (PTVar v) = TVar v - | simple_typ_of (Link (ref T)) = simple_typ_of T - | simple_typ_of (Param _) = sys_error "simple_typ_of: illegal Param"; +fun simple_typ_of tye f T = + (case deref tye T of + PType (a, Ts) => Type (a, map (simple_typ_of tye f) Ts) + | PTFree v => TFree v + | PTVar v => TVar v + | Param (i, S) => TVar (f i, S)); (*convert types, drop constraints*) -fun simple_term_of (PConst (c, T)) = Const (c, simple_typ_of T) - | simple_term_of (PFree (x, T)) = Free (x, simple_typ_of T) - | simple_term_of (PVar (xi, T)) = Var (xi, simple_typ_of T) - | simple_term_of (PBound i) = Bound i - | simple_term_of (PAbs (x, T, t)) = Abs (x, simple_typ_of T, simple_term_of t) - | simple_term_of (PAppl (t, u)) = simple_term_of t $ simple_term_of u - | simple_term_of (Constraint (t, _)) = simple_term_of t; +fun simple_term_of tye f (PConst (c, T)) = Const (c, simple_typ_of tye f T) + | simple_term_of tye f (PFree (x, T)) = Free (x, simple_typ_of tye f T) + | simple_term_of tye f (PVar (xi, T)) = Var (xi, simple_typ_of tye f T) + | simple_term_of tye f (PBound i) = Bound i + | simple_term_of tye f (PAbs (x, T, t)) = + Abs (x, simple_typ_of tye f T, simple_term_of tye f t) + | simple_term_of tye f (PAppl (t, u)) = + simple_term_of tye f t $ simple_term_of tye f u + | simple_term_of tye f (Constraint (t, _)) = simple_term_of tye f t; -(* typs_terms_of *) (*DESTRUCTIVE*) +(* typs_terms_of *) -fun typs_terms_of used maxidx (Ts, ts) = +fun typs_terms_of tye used maxidx (Ts, ts) = let - fun elim (r as ref (Param S), x) = r := PTVar ((x, maxidx + 1), S) - | elim _ = (); - - val used' = fold add_names ts (fold add_namesT Ts used); - val parms = rev (fold add_parms ts (fold add_parmsT Ts [])); + val used' = fold (add_names tye) ts (fold (add_namesT tye) Ts used); + val parms = rev (fold (add_parms tye) ts (fold (add_parmsT tye) Ts [])); val names = Name.invents used' ("?" ^ Name.aT) (length parms); - val _ = ListPair.app elim (parms, names); - in (map simple_typ_of Ts, map simple_term_of ts) end; + val tab = Inttab.make (parms ~~ names); + fun f i = (the (Inttab.lookup tab i), maxidx + 1); + in (map (simple_typ_of tye f) Ts, map (simple_term_of tye f) ts) end; -(** order-sorted unification of types **) (*DESTRUCTIVE*) +(** order-sorted unification of types **) -exception NO_UNIFIER of string; +exception NO_UNIFIER of string * pretyp Inttab.table; fun unify pp tsig = let @@ -254,49 +265,54 @@ "Variable " ^ x ^ "::" ^ Pretty.string_of_sort pp S' ^ " not of sort " ^ Pretty.string_of_sort pp S; - fun meet (_, []) = () - | meet (Link (r as (ref (Param S'))), S) = - if Type.subsort tsig (S', S) then () - else r := mk_param (Type.inter_sort tsig (S', S)) - | meet (Link (ref T), S) = meet (T, S) - | meet (PType (a, Ts), S) = - ListPair.app meet (Ts, Type.arity_sorts pp tsig a S - handle ERROR msg => raise NO_UNIFIER msg) - | meet (PTFree (x, S'), S) = - if Type.subsort tsig (S', S) then () - else raise NO_UNIFIER (not_of_sort x S' S) - | meet (PTVar (xi, S'), S) = - if Type.subsort tsig (S', S) then () - else raise NO_UNIFIER (not_of_sort (Term.string_of_vname xi) S' S) - | meet (Param _, _) = sys_error "meet"; + fun meet (_, []) tye_idx = tye_idx + | meet (Param (i, S'), S) (tye_idx as (tye, idx)) = + if Type.subsort tsig (S', S) then tye_idx + else (Inttab.update_new (i, + Param (idx, Type.inter_sort tsig (S', S))) tye, idx + 1) + | meet (PType (a, Ts), S) (tye_idx as (tye, _)) = + meets (Ts, Type.arity_sorts pp tsig a S + handle ERROR msg => raise NO_UNIFIER (msg, tye)) tye_idx + | meet (PTFree (x, S'), S) (tye_idx as (tye, _)) = + if Type.subsort tsig (S', S) then tye_idx + else raise NO_UNIFIER (not_of_sort x S' S, tye) + | meet (PTVar (xi, S'), S) (tye_idx as (tye, _)) = + if Type.subsort tsig (S', S) then tye_idx + else raise NO_UNIFIER (not_of_sort (Term.string_of_vname xi) S' S, tye) + and meets (T :: Ts, S :: Ss) (tye_idx as (tye, _)) = + meets (Ts, Ss) (meet (deref tye T, S) tye_idx) + | meets _ tye_idx = tye_idx; (* occurs check and assigment *) - fun occurs_check r (Link (r' as ref T)) = - if r = r' then raise NO_UNIFIER "Occurs check!" - else occurs_check r T - | occurs_check r (PType (_, Ts)) = List.app (occurs_check r) Ts - | occurs_check _ _ = (); + fun occurs_check tye i (Param (i', S)) = + if i = i' then raise NO_UNIFIER ("Occurs check!", tye) + else + (case Inttab.lookup tye i' of + NONE => () + | SOME T => occurs_check tye i T) + | occurs_check tye i (PType (_, Ts)) = List.app (occurs_check tye i) Ts + | occurs_check _ _ _ = (); - fun assign r T S = - (case deref T of - T' as Link (r' as ref (Param _)) => - if r = r' then () else (meet (T', S); r := T') - | T' => (occurs_check r T'; meet (T', S); r := T')); + fun assign i (T as Param (i', _)) S (tye_idx as (tye, idx)) = + if i = i' then tye_idx + else meet (T, S) (Inttab.update_new (i, T) tye, idx) + | assign i T S (tye, idx) = + (occurs_check tye i T; meet (T, S) (Inttab.update_new (i, T) tye, idx)); (* unification *) - fun unif (Link (r as ref (Param S)), T) = assign r T S - | unif (T, Link (r as ref (Param S))) = assign r T S - | unif (Link (ref T), U) = unif (T, U) - | unif (T, Link (ref U)) = unif (T, U) - | unif (PType (a, Ts), PType (b, Us)) = + fun unif (T1, T2) (tye_idx as (tye, idx)) = + (case (deref tye T1, deref tye T2) of + (Param (i, S), T) => assign i T S tye_idx + | (T, Param (i, S)) => assign i T S tye_idx + | (PType (a, Ts), PType (b, Us)) => if a <> b then - raise NO_UNIFIER ("Clash of types " ^ quote a ^ " and " ^ quote b) - else ListPair.app unif (Ts, Us) - | unif (T, U) = if T = U then () else raise NO_UNIFIER ""; + raise NO_UNIFIER ("Clash of types " ^ quote a ^ " and " ^ quote b, tye) + else fold unif (Ts ~~ Us) tye_idx + | (T, U) => if T = U then tye_idx else raise NO_UNIFIER ("", tye)); in unif end; @@ -318,7 +334,7 @@ ""]; -(* infer *) (*DESTRUCTIVE*) +(* infer *) fun infer pp tsig = let @@ -327,9 +343,9 @@ fun unif_failed msg = "Type unification failed" ^ (if msg = "" then "" else ": " ^ msg) ^ "\n"; - fun prep_output bs ts Ts = + fun prep_output tye bs ts Ts = let - val (Ts_bTs', ts') = typs_terms_of Name.context ~1 (Ts @ map snd bs, ts); + val (Ts_bTs', ts') = typs_terms_of tye Name.context ~1 (Ts @ map snd bs, ts); val (Ts', Ts'') = chop (length Ts) Ts_bTs'; fun prep t = let val xs = rev (Term.variant_frees t (rev (map fst bs ~~ Ts''))) @@ -339,9 +355,9 @@ fun err_loose i = raise TYPE ("Loose bound variable: B." ^ string_of_int i, [], []); - fun err_appl msg bs t T u U = + fun err_appl msg tye bs t T u U = let - val ([t', u'], [T', U']) = prep_output bs [t, u] [T, U]; + val ([t', u'], [T', U']) = prep_output tye bs [t, u] [T, U]; val why = (case T' of Type ("fun", _) => "Incompatible operand type" @@ -349,9 +365,9 @@ val text = unif_failed msg ^ cat_lines (appl_error pp why t' T' u' U'); in raise TYPE (text, [T', U'], [t', u']) end; - fun err_constraint msg bs t T U = + fun err_constraint msg tye bs t T U = let - val ([t'], [T', U']) = prep_output bs [t] [T, U]; + val ([t'], [T', U']) = prep_output tye bs [t] [T, U]; val text = cat_lines [unif_failed msg, "Cannot meet type constraint:", "", @@ -367,23 +383,28 @@ val unif = unify pp tsig; - fun inf _ (PConst (_, T)) = T - | inf _ (PFree (_, T)) = T - | inf _ (PVar (_, T)) = T - | inf bs (PBound i) = snd (nth bs i handle Subscript => err_loose i) - | inf bs (PAbs (x, T, t)) = PType ("fun", [T, inf ((x, T) :: bs) t]) - | inf bs (PAppl (t, u)) = + fun inf _ (PConst (_, T)) tye_idx = (T, tye_idx) + | inf _ (PFree (_, T)) tye_idx = (T, tye_idx) + | inf _ (PVar (_, T)) tye_idx = (T, tye_idx) + | inf bs (PBound i) tye_idx = + (snd (nth bs i handle Subscript => err_loose i), tye_idx) + | inf bs (PAbs (x, T, t)) tye_idx = + let val (U, tye_idx') = inf ((x, T) :: bs) t tye_idx + in (PType ("fun", [T, U]), tye_idx') end + | inf bs (PAppl (t, u)) tye_idx = let - val T = inf bs t; - val U = inf bs u; - val V = mk_param []; + val (T, tye_idx') = inf bs t tye_idx; + val (U, (tye, idx)) = inf bs u tye_idx'; + val V = Param (idx, []); val U_to_V = PType ("fun", [U, V]); - val _ = unif (U_to_V, T) handle NO_UNIFIER msg => err_appl msg bs t T u U; - in V end - | inf bs (Constraint (t, U)) = - let val T = inf bs t in - unif (T, U) handle NO_UNIFIER msg => err_constraint msg bs t T U; - T + val tye_idx'' = unif (U_to_V, T) (tye, idx + 1) + handle NO_UNIFIER (msg, tye') => err_appl msg tye' bs t T u U; + in (V, tye_idx'') end + | inf bs (Constraint (t, U)) tye_idx = + let val (T, tye_idx') = inf bs t tye_idx in + (T, + unif (T, U) tye_idx' + handle NO_UNIFIER (msg, tye) => err_constraint msg tye bs t T U) end; in inf [] end; @@ -402,11 +423,12 @@ (*convert to preterms*) val ts = burrow_types check_typs raw_ts; - val (ts', (vps, ps)) = - fold_map (preterm_of const_type is_param o constrain_vars) ts (Vartab.empty, Vartab.empty); + val (ts', (_, _, idx)) = + fold_map (preterm_of const_type is_param o constrain_vars) ts + (Vartab.empty, Vartab.empty, 0); (*do type inference*) - val _ = List.app (ignore o infer pp tsig) ts'; - in #2 (typs_terms_of used maxidx ([], ts')) end; + val (tye, _) = fold (snd oo infer pp tsig) ts' (Inttab.empty, idx); + in #2 (typs_terms_of tye used maxidx ([], ts')) end; end; diff -r 63686057cbe8 -r abda97d2deea src/Pure/unify.ML --- a/src/Pure/unify.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Pure/unify.ML Thu Jul 23 21:13:21 2009 +0200 @@ -52,36 +52,48 @@ type dpair = binderlist * term * term; -fun body_type(Envir.Envir{iTs,...}) = -let fun bT(Type("fun",[_,T])) = bT T - | bT(T as TVar ixnS) = (case Type.lookup iTs ixnS of - NONE => T | SOME(T') => bT T') - | bT T = T -in bT end; +fun body_type env = + let + val tyenv = Envir.type_env env; + fun bT (Type ("fun", [_, T])) = bT T + | bT (T as TVar v) = + (case Type.lookup tyenv v of + NONE => T + | SOME T' => bT T') + | bT T = T; + in bT end; -fun binder_types(Envir.Envir{iTs,...}) = -let fun bTs(Type("fun",[T,U])) = T :: bTs U - | bTs(T as TVar ixnS) = (case Type.lookup iTs ixnS of - NONE => [] | SOME(T') => bTs T') - | bTs _ = [] -in bTs end; +fun binder_types env = + let + val tyenv = Envir.type_env env; + fun bTs (Type ("fun", [T, U])) = T :: bTs U + | bTs (T as TVar v) = + (case Type.lookup tyenv v of + NONE => [] + | SOME T' => bTs T') + | bTs _ = []; + in bTs end; fun strip_type env T = (binder_types env T, body_type env T); fun fastype env (Ts, t) = Envir.fastype env (map snd Ts) t; -(*Eta normal form*) -fun eta_norm(env as Envir.Envir{iTs,...}) = - let fun etif (Type("fun",[T,U]), t) = - Abs("", T, etif(U, incr_boundvars 1 t $ Bound 0)) - | etif (TVar ixnS, t) = - (case Type.lookup iTs ixnS of - NONE => t | SOME(T) => etif(T,t)) - | etif (_,t) = t; - fun eta_nm (rbinder, Abs(a,T,body)) = - Abs(a, T, eta_nm ((a,T)::rbinder, body)) - | eta_nm (rbinder, t) = etif(fastype env (rbinder,t), t) +(* eta normal form *) + +fun eta_norm env = + let + val tyenv = Envir.type_env env; + fun etif (Type ("fun", [T, U]), t) = + Abs ("", T, etif (U, incr_boundvars 1 t $ Bound 0)) + | etif (TVar v, t) = + (case Type.lookup tyenv v of + NONE => t + | SOME T => etif (T, t)) + | etif (_, t) = t; + fun eta_nm (rbinder, Abs (a, T, body)) = + Abs (a, T, eta_nm ((a, T) :: rbinder, body)) + | eta_nm (rbinder, t) = etif (fastype env (rbinder, t), t); in eta_nm end; @@ -186,11 +198,14 @@ exception ASSIGN; (*Raised if not an assignment*) -fun unify_types thy (T,U, env as Envir.Envir{asol,iTs,maxidx}) = - if T=U then env - else let val (iTs',maxidx') = Sign.typ_unify thy (U, T) (iTs, maxidx) - in Envir.Envir{asol=asol,maxidx=maxidx',iTs=iTs'} end - handle Type.TUNIFY => raise CANTUNIFY; +fun unify_types thy (T, U, env) = + if T = U then env + else + let + val Envir.Envir {maxidx, tenv, tyenv} = env; + val (tyenv', maxidx') = Sign.typ_unify thy (U, T) (tyenv, maxidx); + in Envir.Envir {maxidx = maxidx', tenv = tenv, tyenv = tyenv'} end + handle Type.TUNIFY => raise CANTUNIFY; fun test_unify_types thy (args as (T,U,_)) = let val str_of = Syntax.string_of_typ_global thy; @@ -636,9 +651,9 @@ (*Pattern matching*) -fun first_order_matchers thy pairs (Envir.Envir {asol = tenv, iTs = tyenv, maxidx}) = +fun first_order_matchers thy pairs (Envir.Envir {maxidx, tenv, tyenv}) = let val (tyenv', tenv') = fold (Pattern.first_order_match thy) pairs (tyenv, tenv) - in Seq.single (Envir.Envir {asol = tenv', iTs = tyenv', maxidx = maxidx}) end + in Seq.single (Envir.Envir {maxidx = maxidx, tenv = tenv', tyenv = tyenv'}) end handle Pattern.MATCH => Seq.empty; (*General matching -- keeps variables disjoint*) @@ -661,10 +676,12 @@ Term.map_aterms (fn t as Var ((x, i), T) => if i > maxidx then Var ((x, i - offset), T) else t | t => t); - fun norm_tvar (Envir.Envir {iTs = tyenv, ...}) ((x, i), S) = - ((x, i - offset), (S, decr_indexesT (Envir.norm_type tyenv (TVar ((x, i), S))))); - fun norm_var (env as Envir.Envir {iTs = tyenv, ...}) ((x, i), T) = + fun norm_tvar env ((x, i), S) = + ((x, i - offset), + (S, decr_indexesT (Envir.norm_type (Envir.type_env env) (TVar ((x, i), S))))); + fun norm_var env ((x, i), T) = let + val tyenv = Envir.type_env env; val T' = Envir.norm_type tyenv T; val t' = Envir.norm_term env (Var ((x, i), T')); in ((x, i - offset), (decr_indexesT T', decr_indexes t')) end; @@ -672,8 +689,8 @@ fun result env = if Envir.above env maxidx then (* FIXME proper handling of generated vars!? *) SOME (Envir.Envir {maxidx = maxidx, - iTs = Vartab.make (map (norm_tvar env) pat_tvars), - asol = Vartab.make (map (norm_var env) pat_vars)}) + tyenv = Vartab.make (map (norm_tvar env) pat_tvars), + tenv = Vartab.make (map (norm_var env) pat_vars)}) else NONE; val empty = Envir.empty maxidx'; diff -r 63686057cbe8 -r abda97d2deea src/Sequents/prover.ML --- a/src/Sequents/prover.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Sequents/prover.ML Thu Jul 23 21:13:21 2009 +0200 @@ -27,7 +27,8 @@ fun warn_duplicates [] = [] | warn_duplicates dups = - (warning (cat_lines ("Ignoring duplicate theorems:" :: map Display.string_of_thm dups)); + (warning (cat_lines ("Ignoring duplicate theorems:" :: + map Display.string_of_thm_without_context dups)); dups); fun (Pack(safes,unsafes)) add_safes ths = @@ -50,8 +51,9 @@ fun print_pack (Pack(safes,unsafes)) = - (writeln "Safe rules:"; Display.print_thms safes; - writeln "Unsafe rules:"; Display.print_thms unsafes); + writeln (cat_lines + (["Safe rules:"] @ map Display.string_of_thm_without_context safes @ + ["Unsafe rules:"] @ map Display.string_of_thm_without_context unsafes)); (*Returns the list of all formulas in the sequent*) fun forms_of_seq (Const("SeqO'",_) $ P $ u) = P :: forms_of_seq u diff -r 63686057cbe8 -r abda97d2deea src/Sequents/simpdata.ML --- a/src/Sequents/simpdata.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Sequents/simpdata.ML Thu Jul 23 21:13:21 2009 +0200 @@ -40,7 +40,7 @@ | (Const("Not",_)$_) => th RS @{thm iff_reflection_F} | _ => th RS @{thm iff_reflection_T}) | _ => error ("addsimps: unable to use theorem\n" ^ - Display.string_of_thm th)); + Display.string_of_thm_without_context th)); (*Replace premises x=y, X<->Y by X==Y*) val mk_meta_prems = diff -r 63686057cbe8 -r abda97d2deea src/Tools/Code/code_ml.ML --- a/src/Tools/Code/code_ml.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Tools/Code/code_ml.ML Thu Jul 23 21:13:21 2009 +0200 @@ -966,7 +966,7 @@ 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 Code_Thingol.eval thy I postproc evaluator t end; + in Code_Thingol.eval thy postproc evaluator t end; (* instrumentalization by antiquotation *) diff -r 63686057cbe8 -r abda97d2deea src/Tools/Code/code_preproc.ML --- a/src/Tools/Code/code_preproc.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Tools/Code/code_preproc.ML Thu Jul 23 21:13:21 2009 +0200 @@ -9,7 +9,7 @@ sig val map_pre: (simpset -> simpset) -> theory -> theory val map_post: (simpset -> simpset) -> theory -> theory - val add_inline: thm -> theory -> theory + val add_unfold: thm -> theory -> theory val add_functrans: string * (theory -> (thm * bool) list -> (thm * bool) list option) -> theory -> theory val del_functrans: string -> theory -> theory val simple_functrans: (theory -> thm list -> thm list option) @@ -23,9 +23,10 @@ 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) + val resubst_triv_consts: theory -> term -> term + val eval_conv: theory -> (code_algebra -> code_graph -> (string * sort) list -> term -> cterm -> thm) -> cterm -> thm - val eval: theory -> (sort -> sort) -> ((term -> term) -> 'a -> 'a) + val eval: theory -> ((term -> term) -> 'a -> 'a) -> (code_algebra -> code_graph -> (string * sort) list -> term -> 'a) -> term -> 'a val setup: theory -> theory @@ -77,14 +78,24 @@ |> Code.purge_data |> (Code_Preproc_Data.map o map_thmproc) f; -val map_pre = map_data o apfst o apfst; -val map_post = map_data o apfst o apsnd; +val map_pre_post = map_data o apfst; +val map_pre = map_pre_post o apfst; +val map_post = map_pre_post o apsnd; -val add_inline = map_pre o MetaSimplifier.add_simp; -val del_inline = map_pre o MetaSimplifier.del_simp; +val add_unfold = map_pre o MetaSimplifier.add_simp; +val del_unfold = map_pre o MetaSimplifier.del_simp; val add_post = map_post o MetaSimplifier.add_simp; val del_post = map_post o MetaSimplifier.del_simp; - + +fun add_unfold_post raw_thm thy = + let + val thm = LocalDefs.meta_rewrite_rule (ProofContext.init thy) raw_thm; + val thm_sym = Thm.symmetric thm; + in + thy |> map_pre_post (fn (pre, post) => + (pre |> MetaSimplifier.add_simp thm, post |> MetaSimplifier.del_simp thm_sym)) + end; + fun add_functrans (name, f) = (map_data o apsnd) (AList.update (op =) (name, (serial (), f))); @@ -151,7 +162,10 @@ |> rhs_conv (Simplifier.rewrite post) end; -fun postprocess_term thy = term_of_conv thy (postprocess_conv thy); +fun resubst_triv_consts thy = map_aterms (fn t as Const (c, ty) => Const (Code.resubst_alias thy c, ty) + | t => t); + +fun postprocess_term thy = term_of_conv thy (postprocess_conv thy) #> resubst_triv_consts thy; fun print_codeproc thy = let @@ -200,7 +214,7 @@ |> map (fn (s, thms) => (Pretty.block o Pretty.fbreaks) ( Pretty.str s - :: map (Display.pretty_thm o fst) thms + :: map (Display.pretty_thm_global thy o fst) thms )) |> Pretty.chunks; @@ -485,7 +499,7 @@ Abs (v, prepare_sorts_typ prep_sort ty, prepare_sorts prep_sort t) | prepare_sorts _ (t as Bound _) = t; -fun gen_eval thy cterm_of conclude_evaluation prep_sort evaluator proto_ct = +fun gen_eval thy cterm_of conclude_evaluation evaluator proto_ct = let val pp = Syntax.pp_global thy; val ct = cterm_of proto_ct; @@ -497,8 +511,10 @@ 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 add_triv_classes = curry (Sorts.inter_sort (Sign.classes_of thy)) + (Code.triv_classes thy); + val t'' = prepare_sorts add_triv_classes t'; val (algebra', eqngr') = obtain thy consts [t'']; in conclude_evaluation (evaluator algebra' eqngr' vs t'' ct') thm end; @@ -513,12 +529,12 @@ 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]) + ^ (cat_lines o map (Display.string_of_thm_global thy)) [thm1, thm2, thm3]) end; in gen_eval thy I conclude_evaluation end; -fun eval thy prep_sort postproc evaluator = gen_eval thy (Thm.cterm_of thy) - (K o postproc (postprocess_term thy)) prep_sort (simple_evaluator evaluator); +fun eval thy postproc evaluator = gen_eval thy (Thm.cterm_of thy) + (K o postproc (postprocess_term thy)) (simple_evaluator evaluator); (** setup **) @@ -526,16 +542,19 @@ val setup = let fun mk_attribute f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I); - fun add_del_attribute_parser (add, del) = + fun add_del_attribute_parser add del = Attrib.add_del (mk_attribute add) (mk_attribute del); + fun both f g thm = f thm #> g thm; in - Attrib.setup @{binding code_inline} (add_del_attribute_parser (add_inline, del_inline)) + Attrib.setup @{binding code_unfold} (add_del_attribute_parser + (both Codegen.add_unfold add_unfold) (both Codegen.del_unfold del_unfold)) + "preprocessing equations for code generators" + #> Attrib.setup @{binding code_inline} (add_del_attribute_parser add_unfold del_unfold) "preprocessing equations for code generator" - #> Attrib.setup @{binding code_post} (add_del_attribute_parser (add_post, del_post)) + #> Attrib.setup @{binding code_post} (add_del_attribute_parser add_post del_post) "postprocessing equations for code generator" - #> Attrib.setup @{binding code_unfold} (Scan.succeed (mk_attribute - (fn thm => Codegen.add_unfold thm #> add_inline thm))) - "preprocessing equations for code generators" + #> Attrib.setup @{binding code_unfold_post} (Scan.succeed (mk_attribute add_unfold_post)) + "pre- and postprocessing equations for code generator" end; val _ = diff -r 63686057cbe8 -r abda97d2deea src/Tools/Code/code_printer.ML --- a/src/Tools/Code/code_printer.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Tools/Code/code_printer.ML Thu Jul 23 21:13:21 2009 +0200 @@ -82,7 +82,7 @@ open Code_Thingol; -fun nerror thm s = error (s ^ ",\nin equation " ^ Display.string_of_thm thm); +fun nerror thm s = error (s ^ ",\nin equation " ^ Display.string_of_thm_without_context thm); (** assembling text pieces **) diff -r 63686057cbe8 -r abda97d2deea src/Tools/Code/code_thingol.ML --- a/src/Tools/Code/code_thingol.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Tools/Code/code_thingol.ML Thu Jul 23 21:13:21 2009 +0200 @@ -82,10 +82,10 @@ val read_const_exprs: theory -> string list -> string list * string list val consts_program: theory -> string list -> string list * (naming * program) val cached_program: theory -> naming * program - val eval_conv: theory -> (sort -> sort) + val eval_conv: theory -> (naming -> program -> ((string * sort) list * typscheme) * iterm -> string list -> cterm -> thm) -> cterm -> thm - val eval: theory -> (sort -> sort) -> ((term -> term) -> 'a -> 'a) + val eval: theory -> ((term -> term) -> 'a -> 'a) -> (naming -> program -> ((string * sort) list * typscheme) * iterm -> string list -> 'a) -> term -> 'a end; @@ -469,7 +469,7 @@ let val err_class = Sorts.class_error (Syntax.pp_global thy) e; val err_thm = case thm - of SOME thm => "\n(in code equation " ^ Display.string_of_thm thm ^ ")" | NONE => ""; + of SOME thm => "\n(in code equation " ^ Display.string_of_thm_global thy thm ^ ")" | NONE => ""; val err_typ = "Type " ^ Syntax.string_of_typ_global thy ty ^ " not of sort " ^ Syntax.string_of_sort_global thy sort; in error ("Wellsortedness error" ^ err_thm ^ ":\n" ^ err_typ ^ "\n" ^ err_class) end; @@ -803,8 +803,8 @@ val vs'' = map (fn (v, _) => (v, (the o AList.lookup (op =) vs o prefix "'") v)) vs'; in evaluator naming program ((vs'', (vs', ty')), t') deps end; -fun eval_conv thy prep_sort = Code_Preproc.eval_conv thy prep_sort o base_evaluator thy; -fun eval thy prep_sort postproc = Code_Preproc.eval thy prep_sort postproc o base_evaluator thy; +fun eval_conv thy = Code_Preproc.eval_conv thy o base_evaluator thy; +fun eval thy postproc = Code_Preproc.eval thy postproc o base_evaluator thy; (** diagnostic commands **) diff -r 63686057cbe8 -r abda97d2deea src/Tools/Compute_Oracle/linker.ML --- a/src/Tools/Compute_Oracle/linker.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Tools/Compute_Oracle/linker.ML Thu Jul 23 21:13:21 2009 +0200 @@ -341,7 +341,7 @@ let val (newsubsts, instances) = Linker.add_instances thy instances monocs val _ = if not (null newsubsts) then changed := true else () - val newpats = map (fn subst => Envir.subst_TVars subst p) newsubsts + val newpats = map (fn subst => Envir.subst_term_types subst p) newsubsts in PolyPattern (p, instances, instancepats@newpats) end diff -r 63686057cbe8 -r abda97d2deea src/Tools/coherent.ML --- a/src/Tools/coherent.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Tools/coherent.ML Thu Jul 23 21:13:21 2009 +0200 @@ -131,7 +131,7 @@ let val seq = Seq.of_list rules |> Seq.maps (fn (th, ps, cs) => valid_conj ctxt facts empty_env ps |> Seq.maps (fn (env as (tye, _), is) => let val cs' = map (fn (Ts, ts) => - (map (Envir.typ_subst_TVars tye) Ts, map (Envir.subst_vars env) ts)) cs + (map (Envir.subst_type tye) Ts, map (Envir.subst_term env) ts)) cs in inst_extra_vars ctxt dom cs' |> Seq.map_filter (fn (inst, cs'') => @@ -177,14 +177,14 @@ fun thm_of_cl_prf thy goal asms (ClPrf (th, (tye, env), insts, is, prfs)) = let val _ = message (fn () => space_implode "\n" - ("asms:" :: map Display.string_of_thm asms) ^ "\n\n"); + ("asms:" :: map (Display.string_of_thm_global thy) asms) ^ "\n\n"); val th' = Drule.implies_elim_list (Thm.instantiate (map (fn (ixn, (S, T)) => (Thm.ctyp_of thy (TVar ((ixn, S))), Thm.ctyp_of thy T)) (Vartab.dest tye), map (fn (ixn, (T, t)) => - (Thm.cterm_of thy (Var (ixn, Envir.typ_subst_TVars tye T)), + (Thm.cterm_of thy (Var (ixn, Envir.subst_type tye T)), Thm.cterm_of thy t)) (Vartab.dest env) @ map (fn (ixnT, t) => (Thm.cterm_of thy (Var ixnT), Thm.cterm_of thy t)) insts) th) diff -r 63686057cbe8 -r abda97d2deea src/Tools/eqsubst.ML --- a/src/Tools/eqsubst.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Tools/eqsubst.ML Thu Jul 23 21:13:21 2009 +0200 @@ -127,7 +127,7 @@ (* changes object "=" to meta "==" which prepares a given rewrite rule *) fun prep_meta_eq ctxt = - Simplifier.mksimps (Simplifier.local_simpset_of ctxt) #> map Drule.zero_var_indexes; + Simplifier.mksimps (simpset_of ctxt) #> map Drule.zero_var_indexes; (* a type abriviation for match information *) @@ -231,9 +231,9 @@ or should I be using them somehow? *) fun mk_insts env = (Vartab.dest (Envir.type_env env), - Envir.alist_of env); - val initenv = Envir.Envir {asol = Vartab.empty, - iTs = typinsttab, maxidx = ix2}; + Vartab.dest (Envir.term_env env)); + val initenv = + Envir.Envir {maxidx = ix2, tenv = Vartab.empty, tyenv = typinsttab}; val useq = Unify.smash_unifiers thry [a] initenv handle UnequalLengths => Seq.empty | Term.TERM _ => Seq.empty; diff -r 63686057cbe8 -r abda97d2deea src/Tools/induct.ML --- a/src/Tools/induct.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Tools/induct.ML Thu Jul 23 21:13:21 2009 +0200 @@ -124,7 +124,7 @@ fun pretty_rules ctxt kind rs = let val thms = map snd (Item_Net.content rs) - in Pretty.big_list kind (map (ProofContext.pretty_thm ctxt) thms) end; + in Pretty.big_list kind (map (Display.pretty_thm ctxt) thms) end; (* context data *) @@ -423,14 +423,15 @@ local -fun dest_env thy (env as Envir.Envir {iTs, ...}) = +fun dest_env thy env = let val cert = Thm.cterm_of thy; val certT = Thm.ctyp_of thy; - val pairs = Envir.alist_of env; + val pairs = Vartab.dest (Envir.term_env env); + val types = Vartab.dest (Envir.type_env env); val ts = map (cert o Envir.norm_term env o #2 o #2) pairs; val xs = map2 (curry (cert o Var)) (map #1 pairs) (map (#T o Thm.rep_cterm) ts); - in (map (fn (xi, (S, T)) => (certT (TVar (xi, S)), certT T)) (Vartab.dest iTs), xs ~~ ts) end; + in (map (fn (xi, (S, T)) => (certT (TVar (xi, S)), certT T)) types, xs ~~ ts) end; in @@ -450,8 +451,7 @@ val rule' = Thm.incr_indexes (maxidx + 1) rule; val concl = Logic.strip_assums_concl goal; in - Unify.smash_unifiers thy [(Thm.concl_of rule', concl)] - (Envir.empty (#maxidx (Thm.rep_thm rule'))) + Unify.smash_unifiers thy [(Thm.concl_of rule', concl)] (Envir.empty (Thm.maxidx_of rule')) |> Seq.map (fn env => Drule.instantiate (dest_env thy env) rule') end end handle Subscript => Seq.empty; diff -r 63686057cbe8 -r abda97d2deea src/Tools/nbe.ML --- a/src/Tools/nbe.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/Tools/nbe.ML Thu Jul 23 21:13:21 2009 +0200 @@ -439,9 +439,6 @@ fun normalize thy naming program ((vs0, (vs, ty)), t) deps = let - fun subst_const f = map_aterms (fn t as Term.Const (c, ty) => Term.Const (f c, ty) - | t => t); - val resubst_triv_consts = subst_const (Code.resubst_alias thy); val ty' = typ_of_itype program vs0 ty; fun type_infer t = singleton (TypeInfer.infer_types (Syntax.pp_global thy) (Sign.tsig_of thy) I @@ -453,8 +450,8 @@ val string_of_term = setmp show_types true (Syntax.string_of_term_global thy); in compile_eval thy naming program (vs, t) deps + |> Code_Preproc.resubst_triv_consts thy |> tracing (fn t => "Normalized:\n" ^ string_of_term t) - |> resubst_triv_consts |> type_infer |> tracing (fn t => "Types inferred:\n" ^ string_of_term t) |> check_tvars @@ -463,9 +460,6 @@ (* evaluation oracle *) -fun add_triv_classes thy = curry (Sorts.inter_sort (Sign.classes_of thy)) - (Code.triv_classes thy); - fun mk_equals thy lhs raw_rhs = let val ty = Thm.typ_of (Thm.ctyp_of_term lhs); @@ -506,9 +500,9 @@ val norm_conv = no_frees_conv (fn ct => let val thy = Thm.theory_of_cterm ct; - in Code_Thingol.eval_conv thy (add_triv_classes thy) (norm_oracle thy) ct end); + in Code_Thingol.eval_conv thy (norm_oracle thy) ct end); -fun norm thy = no_frees_rew (Code_Thingol.eval thy (add_triv_classes thy) I (normalize thy)); +fun norm thy = no_frees_rew (Code_Thingol.eval thy I (normalize thy)); (* evaluation command *) diff -r 63686057cbe8 -r abda97d2deea src/ZF/Datatype_ZF.thy --- a/src/ZF/Datatype_ZF.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/ZF/Datatype_ZF.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1,8 +1,6 @@ (* Title: ZF/Datatype.thy - ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1997 University of Cambridge - *) header{*Datatype and CoDatatype Definitions*} @@ -103,7 +101,7 @@ handle Match => NONE; - val conv = Simplifier.simproc (the_context ()) "data_free" ["(x::i) = y"] proc; + val conv = Simplifier.simproc @{theory} "data_free" ["(x::i) = y"] proc; end; diff -r 63686057cbe8 -r abda97d2deea src/ZF/IntDiv_ZF.thy --- a/src/ZF/IntDiv_ZF.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/ZF/IntDiv_ZF.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1732,7 +1732,7 @@ (if ~b | #0 $<= integ_of w then integ_of v zdiv (integ_of w) else (integ_of v $+ #1) zdiv (integ_of w))" - apply (simp_tac (simpset_of Int.thy add: zadd_assoc integ_of_BIT) + apply (simp_tac (global_simpset_of Int.thy add: zadd_assoc integ_of_BIT) apply (simp (no_asm_simp) del: bin_arith_extra_simps@bin_rel_simps add: zdiv_zmult_zmult1 pos_zdiv_mult_2 lemma neg_zdiv_mult_2) done @@ -1778,7 +1778,7 @@ then #2 $* (integ_of v zmod integ_of w) $+ #1 else #2 $* ((integ_of v $+ #1) zmod integ_of w) - #1 else #2 $* (integ_of v zmod integ_of w))" - apply (simp_tac (simpset_of Int.thy add: zadd_assoc integ_of_BIT) + apply (simp_tac (global_simpset_of Int.thy add: zadd_assoc integ_of_BIT) apply (simp (no_asm_simp) del: bin_arith_extra_simps@bin_rel_simps add: zmod_zmult_zmult1 pos_zmod_mult_2 lemma neg_zmod_mult_2) done diff -r 63686057cbe8 -r abda97d2deea src/ZF/OrdQuant.thy --- a/src/ZF/OrdQuant.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/ZF/OrdQuant.thy Thu Jul 23 21:13:21 2009 +0200 @@ -1,5 +1,4 @@ (* Title: ZF/AC/OrdQuant.thy - ID: $Id$ Authors: Krzysztof Grabczewski and L C Paulson *) @@ -382,9 +381,9 @@ in -val defREX_regroup = Simplifier.simproc (the_context ()) +val defREX_regroup = Simplifier.simproc @{theory} "defined REX" ["EX x[M]. P(x) & Q(x)"] rearrange_bex; -val defRALL_regroup = Simplifier.simproc (the_context ()) +val defRALL_regroup = Simplifier.simproc @{theory} "defined RALL" ["ALL x[M]. P(x) --> Q(x)"] rearrange_ball; end; diff -r 63686057cbe8 -r abda97d2deea src/ZF/Tools/ind_cases.ML --- a/src/ZF/Tools/ind_cases.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/ZF/Tools/ind_cases.ML Thu Jul 23 21:13:21 2009 +0200 @@ -47,8 +47,7 @@ fun inductive_cases args thy = let - val read_prop = Syntax.read_prop (ProofContext.init thy); - val mk_cases = smart_cases thy (Simplifier.simpset_of thy) read_prop; + val mk_cases = smart_cases thy (global_simpset_of thy) (Syntax.read_prop_global thy); val facts = args |> map (fn ((name, srcs), props) => ((name, map (Attrib.attribute thy) srcs), map (Thm.no_attributes o single o mk_cases) props)); @@ -62,7 +61,7 @@ (Scan.lift (Scan.repeat1 Args.name_source) >> (fn props => fn ctxt => props - |> map (smart_cases (ProofContext.theory_of ctxt) (local_simpset_of ctxt) (Syntax.read_prop ctxt)) + |> map (smart_cases (ProofContext.theory_of ctxt) (simpset_of ctxt) (Syntax.read_prop ctxt)) |> Method.erule 0)) "dynamic case analysis on sets"; diff -r 63686057cbe8 -r abda97d2deea src/ZF/Tools/inductive_package.ML --- a/src/ZF/Tools/inductive_package.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/ZF/Tools/inductive_package.ML Thu Jul 23 21:13:21 2009 +0200 @@ -247,8 +247,7 @@ (intr_tms, map intro_tacsf (mk_disj_rls (length intr_tms))) |> ListPair.map (fn (t, tacs) => Goal.prove_global thy1 [] [] t - (fn _ => EVERY (rewrite_goals_tac part_rec_defs :: tacs))) - handle MetaSimplifier.SIMPLIFIER (msg, thm) => (Display.print_thm thm; error msg); + (fn _ => EVERY (rewrite_goals_tac part_rec_defs :: tacs))); (********) val dummy = writeln " Proving the elimination rule..."; @@ -318,11 +317,12 @@ val ind_prems = map (induct_prem (map (rpair pred) rec_tms)) intr_tms; - val dummy = if !Ind_Syntax.trace then - (writeln "ind_prems = "; - List.app (writeln o Syntax.string_of_term ctxt1) ind_prems; - writeln "raw_induct = "; Display.print_thm raw_induct) - else (); + val dummy = + if ! Ind_Syntax.trace then + writeln (cat_lines + (["ind_prems:"] @ map (Syntax.string_of_term ctxt1) ind_prems @ + ["raw_induct:", Display.string_of_thm ctxt1 raw_induct])) + else (); (*We use a MINIMAL simpset. Even FOL_ss contains too many simpules. @@ -351,9 +351,10 @@ ORELSE' bound_hyp_subst_tac)), ind_tac (rev (map (rewrite_rule part_rec_defs) prems)) (length prems)]); - val dummy = if !Ind_Syntax.trace then - (writeln "quant_induct = "; Display.print_thm quant_induct) - else (); + val dummy = + if ! Ind_Syntax.trace then + writeln ("quant_induct:\n" ^ Display.string_of_thm ctxt1 quant_induct) + else (); (*** Prove the simultaneous induction rule ***) @@ -426,9 +427,10 @@ REPEAT (rewrite_goals_tac [Pr.split_eq] THEN lemma_tac 1)])) else (writeln " [ No mutual induction rule needed ]"; @{thm TrueI}); - val dummy = if !Ind_Syntax.trace then - (writeln "lemma = "; Display.print_thm lemma) - else (); + val dummy = + if ! Ind_Syntax.trace then + writeln ("lemma: " ^ Display.string_of_thm ctxt1 lemma) + else (); (*Mutual induction follows by freeness of Inl/Inr.*) diff -r 63686057cbe8 -r abda97d2deea src/ZF/Tools/typechk.ML --- a/src/ZF/Tools/typechk.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/ZF/Tools/typechk.ML Thu Jul 23 21:13:21 2009 +0200 @@ -27,7 +27,8 @@ fun add_rule th (tcs as TC {rules, net}) = if member Thm.eq_thm_prop rules th then - (warning ("Ignoring duplicate type-checking rule\n" ^ Display.string_of_thm th); tcs) + (warning ("Ignoring duplicate type-checking rule\n" ^ + Display.string_of_thm_without_context th); tcs) else TC {rules = th :: rules, net = Net.insert_term (K false) (Thm.concl_of th, th) net}; @@ -36,7 +37,8 @@ if member Thm.eq_thm_prop rules th then TC {net = Net.delete_term Thm.eq_thm_prop (Thm.concl_of th, th) net, rules = remove Thm.eq_thm_prop th rules} - else (warning ("No such type-checking rule\n" ^ Display.string_of_thm th); tcs); + else (warning ("No such type-checking rule\n" ^ + Display.string_of_thm_without_context th); tcs); (* generic data *) @@ -60,7 +62,7 @@ fun print_tcset ctxt = let val TC {rules, ...} = tcset_of ctxt in Pretty.writeln (Pretty.big_list "type-checking rules:" - (map (ProofContext.pretty_thm ctxt) rules)) + (map (Display.pretty_thm ctxt) rules)) end; diff -r 63686057cbe8 -r abda97d2deea src/ZF/UNITY/Constrains.thy --- a/src/ZF/UNITY/Constrains.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/ZF/UNITY/Constrains.thy Thu Jul 23 21:13:21 2009 +0200 @@ -471,7 +471,7 @@ (*proves "co" properties when the program is specified*) fun constrains_tac ctxt = - let val css as (cs, ss) = local_clasimpset_of ctxt in + let val css as (cs, ss) = clasimpset_of ctxt in SELECT_GOAL (EVERY [REPEAT (Always_Int_tac 1), REPEAT (etac @{thm Always_ConstrainsI} 1 @@ -494,7 +494,7 @@ (*For proving invariants*) fun always_tac ctxt i = - rtac @{thm AlwaysI} i THEN force_tac (local_clasimpset_of ctxt) i THEN constrains_tac ctxt i; + rtac @{thm AlwaysI} i THEN force_tac (clasimpset_of ctxt) i THEN constrains_tac ctxt i; *} setup Program_Defs.setup diff -r 63686057cbe8 -r abda97d2deea src/ZF/UNITY/SubstAx.thy --- a/src/ZF/UNITY/SubstAx.thy Thu Jul 23 21:12:57 2009 +0200 +++ b/src/ZF/UNITY/SubstAx.thy Thu Jul 23 21:13:21 2009 +0200 @@ -351,7 +351,7 @@ ML {* (*proves "ensures/leadsTo" properties when the program is specified*) fun ensures_tac ctxt sact = - let val css as (cs, ss) = local_clasimpset_of ctxt in + let val css as (cs, ss) = clasimpset_of ctxt in SELECT_GOAL (EVERY [REPEAT (Always_Int_tac 1), etac @{thm Always_LeadsTo_Basis} 1 diff -r 63686057cbe8 -r abda97d2deea src/ZF/int_arith.ML --- a/src/ZF/int_arith.ML Thu Jul 23 21:12:57 2009 +0200 +++ b/src/ZF/int_arith.ML Thu Jul 23 21:13:21 2009 +0200 @@ -145,7 +145,7 @@ val numeral_simp_ss = ZF_ss addsimps add_0s @ bin_simps @ tc_rules @ intifys fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) - THEN ALLGOALS (asm_simp_tac (local_simpset_of (Simplifier.the_context ss))) + THEN ALLGOALS (asm_simp_tac (simpset_of (Simplifier.the_context ss))) val simplify_meta_eq = ArithData.simplify_meta_eq (add_0s @ mult_1s) end;