merged, using src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML and src/HOL/Tools/Sledgehammer/sledgehammer_run.ML from 347c3b0cab44;
authorwenzelm
Mon, 11 Nov 2013 17:44:21 +0100
changeset 5575750199af40c27
parent 55756 9d3c7a04a65e
parent 55750 347c3b0cab44
child 55758 27246f8b2dac
merged, using src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML and src/HOL/Tools/Sledgehammer/sledgehammer_run.ML from 347c3b0cab44;
CONTRIBUTORS
NEWS
src/Doc/manual.bib
src/HOL/BNF/Tools/bnf_fp_rec_sugar.ML
src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML
src/HOL/Library/Abstract_Rat.thy
src/HOL/Library/Glbs.thy
src/HOL/Library/Univ_Poly.thy
src/HOL/Lubs.thy
     1.1 --- a/CONTRIBUTORS	Mon Nov 11 17:34:44 2013 +0100
     1.2 +++ b/CONTRIBUTORS	Mon Nov 11 17:44:21 2013 +0100
     1.3 @@ -3,6 +3,10 @@
     1.4  who is listed as an author in one of the source files of this Isabelle
     1.5  distribution.
     1.6  
     1.7 +Contributions to this Isabelle version
     1.8 +--------------------------------------
     1.9 +
    1.10 +
    1.11  Contributions to Isabelle2013-1
    1.12  -------------------------------
    1.13  
     2.1 --- a/NEWS	Mon Nov 11 17:34:44 2013 +0100
     2.2 +++ b/NEWS	Mon Nov 11 17:44:21 2013 +0100
     2.3 @@ -1,6 +1,67 @@
     2.4  Isabelle NEWS -- history user-relevant changes
     2.5  ==============================================
     2.6  
     2.7 +New in this Isabelle version
     2.8 +----------------------------
     2.9 +
    2.10 +*** HOL ***
    2.11 +
    2.12 +* Qualified constant names Wellfounded.acc, Wellfounded.accp.
    2.13 +INCOMPATIBILITY.
    2.14 +
    2.15 +* Fact generalization and consolidation:
    2.16 +    neq_one_mod_two, mod_2_not_eq_zero_eq_one_int ~> not_mod_2_eq_0_eq_1
    2.17 +INCOMPATIBILITY.
    2.18 +
    2.19 +* Purely algebraic definition of even.  Fact generalization and consolidation:
    2.20 +    nat_even_iff_2_dvd, int_even_iff_2_dvd ~> even_iff_2_dvd
    2.21 +    even_zero_(nat|int) ~> even_zero
    2.22 +INCOMPATIBILITY.
    2.23 +
    2.24 +* Elimination of fact duplicates:
    2.25 +    equals_zero_I ~> minus_unique
    2.26 +    diff_eq_0_iff_eq ~> right_minus_eq
    2.27 +INCOMPATIBILITY.
    2.28 +
    2.29 +* Fact name consolidation:
    2.30 +    diff_def, diff_minus, ab_diff_minus ~> diff_conv_add_uminus
    2.31 +    minus_le_self_iff ~> neg_less_eq_nonneg
    2.32 +    le_minus_self_iff ~> less_eq_neg_nonpos
    2.33 +    neg_less_nonneg ~> neg_less_pos
    2.34 +    less_minus_self_iff ~> less_neg_neg [simp]
    2.35 +INCOMPATIBILITY.
    2.36 +
    2.37 +* More simplification rules on unary and binary minus:
    2.38 +add_diff_cancel, add_diff_cancel_left, add_le_same_cancel1,
    2.39 +add_le_same_cancel2, add_less_same_cancel1, add_less_same_cancel2,
    2.40 +add_minus_cancel, diff_add_cancel, le_add_same_cancel1,
    2.41 +le_add_same_cancel2, less_add_same_cancel1, less_add_same_cancel2,
    2.42 +minus_add_cancel, uminus_add_conv_diff.  These correspondingly
    2.43 +have been taken away from fact collections algebra_simps and
    2.44 +field_simps.  INCOMPATIBILITY.
    2.45 +
    2.46 +To restore proofs, the following patterns are helpful:
    2.47 +
    2.48 +a) Arbitrary failing proof not involving "diff_def":
    2.49 +Consider simplification with algebra_simps or field_simps.
    2.50 +
    2.51 +b) Lifting rules from addition to subtraction:
    2.52 +Try with "using <rule for addition> of [… "- _" …]" by simp".
    2.53 +
    2.54 +c) Simplification with "diff_def": just drop "diff_def".
    2.55 +Consider simplification with algebra_simps or field_simps;
    2.56 +or the brute way with
    2.57 +"simp add: diff_conv_add_uminus del: add_uminus_conv_diff".
    2.58 +
    2.59 +* SUP and INF generalized to conditionally_complete_lattice
    2.60 +
    2.61 +* Theory Lubs moved HOL image to HOL-Library. It is replaced by
    2.62 +Conditionally_Complete_Lattices.   INCOMPATIBILITY.
    2.63 +
    2.64 +* Introduce bdd_above and bdd_below in Conditionally_Complete_Lattices, use them
    2.65 +instead of explicitly stating boundedness of sets.
    2.66 +
    2.67 +
    2.68  New in Isabelle2013-1 (November 2013)
    2.69  -------------------------------------
    2.70  
     3.1 --- a/src/Doc/Datatypes/Datatypes.thy	Mon Nov 11 17:34:44 2013 +0100
     3.2 +++ b/src/Doc/Datatypes/Datatypes.thy	Mon Nov 11 17:44:21 2013 +0100
     3.3 @@ -9,21 +9,8 @@
     3.4  
     3.5  theory Datatypes
     3.6  imports Setup
     3.7 -keywords
     3.8 -  "primcorec_notyet" :: thy_decl
     3.9  begin
    3.10  
    3.11 -(*<*)
    3.12 -(* FIXME: Temporary setup until "primcorec" and "primcorecursive" are fully implemented. *)
    3.13 -ML_command {*
    3.14 -fun add_dummy_cmd _ _ lthy = lthy;
    3.15 -
    3.16 -val _ = Outer_Syntax.local_theory @{command_spec "primcorec_notyet"} ""
    3.17 -  (Parse.fixes -- Parse_Spec.where_alt_specs >> uncurry add_dummy_cmd);
    3.18 -*}
    3.19 -(*>*)
    3.20 -
    3.21 -
    3.22  section {* Introduction
    3.23    \label{sec:introduction} *}
    3.24  
    3.25 @@ -54,17 +41,19 @@
    3.26  
    3.27  text {*
    3.28  \noindent
    3.29 -The package also provides some convenience, notably automatically generated
    3.30 -discriminators and selectors.
    3.31 -
    3.32 -In addition to plain inductive datatypes, the new package supports coinductive
    3.33 -datatypes, or \emph{codatatypes}, which may have infinite values. For example,
    3.34 -the following command introduces the type of lazy lists, which comprises both
    3.35 -finite and infinite values:
    3.36 +Furthermore, the package provides a lot of convenience, including automatically
    3.37 +generated discriminators, selectors, and relators as well as a wealth of
    3.38 +properties about them.
    3.39 +
    3.40 +In addition to inductive datatypes, the new package supports coinductive
    3.41 +datatypes, or \emph{codatatypes}, which allow infinite values. For example, the
    3.42 +following command introduces the type of lazy lists, which comprises both finite
    3.43 +and infinite values:
    3.44  *}
    3.45  
    3.46  (*<*)
    3.47      locale early
    3.48 +    locale late
    3.49  (*>*)
    3.50      codatatype (*<*)(in early) (*>*)'a llist = LNil | LCons 'a "'a llist"
    3.51  
    3.52 @@ -80,10 +69,10 @@
    3.53      codatatype (*<*)(in early) (*>*)'a tree\<^sub>i\<^sub>i = Node\<^sub>i\<^sub>i 'a "'a tree\<^sub>i\<^sub>i llist"
    3.54  
    3.55  text {*
    3.56 -The first two tree types allow only finite branches, whereas the last two allow
    3.57 -branches of infinite length. Orthogonally, the nodes in the first and third
    3.58 -types have finite branching, whereas those of the second and fourth may have
    3.59 -infinitely many direct subtrees.
    3.60 +The first two tree types allow only paths of finite length, whereas the last two
    3.61 +allow infinite paths. Orthogonally, the nodes in the first and third types have
    3.62 +finitely many direct subtrees, whereas those of the second and fourth may have
    3.63 +infinite branching.
    3.64  
    3.65  To use the package, it is necessary to import the @{theory BNF} theory, which
    3.66  can be precompiled into the \texttt{HOL-BNF} image. The following commands show
    3.67 @@ -152,15 +141,15 @@
    3.68  
    3.69  
    3.70  \newbox\boxA
    3.71 -\setbox\boxA=\hbox{\texttt{nospam}}
    3.72 -
    3.73 -\newcommand\authoremaili{\texttt{blan{\color{white}nospam}\kern-\wd\boxA{}chette@\allowbreak
    3.74 +\setbox\boxA=\hbox{\texttt{NOSPAM}}
    3.75 +
    3.76 +\newcommand\authoremaili{\texttt{blan{\color{white}NOSPAM}\kern-\wd\boxA{}chette@\allowbreak
    3.77  in.\allowbreak tum.\allowbreak de}}
    3.78 -\newcommand\authoremailii{\texttt{lore{\color{white}nospam}\kern-\wd\boxA{}nz.panny@\allowbreak
    3.79 +\newcommand\authoremailii{\texttt{lore{\color{white}NOSPAM}\kern-\wd\boxA{}nz.panny@\allowbreak
    3.80  \allowbreak tum.\allowbreak de}}
    3.81 -\newcommand\authoremailiii{\texttt{pope{\color{white}nospam}\kern-\wd\boxA{}scua@\allowbreak
    3.82 +\newcommand\authoremailiii{\texttt{pope{\color{white}NOSPAM}\kern-\wd\boxA{}scua@\allowbreak
    3.83  in.\allowbreak tum.\allowbreak de}}
    3.84 -\newcommand\authoremailiv{\texttt{tray{\color{white}nospam}\kern-\wd\boxA{}tel@\allowbreak
    3.85 +\newcommand\authoremailiv{\texttt{tray{\color{white}NOSPAM}\kern-\wd\boxA{}tel@\allowbreak
    3.86  in.\allowbreak tum.\allowbreak de}}
    3.87  
    3.88  The commands @{command datatype_new} and @{command primrec_new} are expected to
    3.89 @@ -171,13 +160,6 @@
    3.90  Comments and bug reports concerning either the tool or this tutorial should be
    3.91  directed to the authors at \authoremaili, \authoremailii, \authoremailiii,
    3.92  and \authoremailiv.
    3.93 -
    3.94 -\begin{framed}
    3.95 -\noindent
    3.96 -\textbf{Warning:}\enskip This tutorial and the package it describes are under
    3.97 -construction. Please forgive their appearance. Should you have suggestions
    3.98 -or comments regarding either, please let the authors know.
    3.99 -\end{framed}
   3.100  *}
   3.101  
   3.102  
   3.103 @@ -195,7 +177,7 @@
   3.104  text {*
   3.105  Datatypes are illustrated through concrete examples featuring different flavors
   3.106  of recursion. More examples can be found in the directory
   3.107 -\verb|~~/src/HOL/BNF/Examples|.
   3.108 +\verb|~~/src/HOL/|\allowbreak\verb|BNF/Examples|.
   3.109  *}
   3.110  
   3.111  subsubsection {* Nonrecursive Types
   3.112 @@ -260,7 +242,8 @@
   3.113  
   3.114  text {*
   3.115  \noindent
   3.116 -Lists were shown in the introduction. Terminated lists are a variant:
   3.117 +Lists were shown in the introduction. Terminated lists are a variant that
   3.118 +stores a value of type @{typ 'b} at the very end:
   3.119  *}
   3.120  
   3.121      datatype_new (*<*)(in early) (*>*)('a, 'b) tlist = TNil 'b | TCons 'a "('a, 'b) tlist"
   3.122 @@ -310,7 +293,7 @@
   3.123  Not all nestings are admissible. For example, this command will fail:
   3.124  *}
   3.125  
   3.126 -    datatype_new 'a wrong = Wrong (*<*)'a
   3.127 +    datatype_new 'a wrong = W1 | W2 (*<*)'a
   3.128      typ (*>*)"'a wrong \<Rightarrow> 'a"
   3.129  
   3.130  text {*
   3.131 @@ -321,7 +304,7 @@
   3.132  *}
   3.133  
   3.134      datatype_new ('a, 'b) fn = Fn "'a \<Rightarrow> 'b"
   3.135 -    datatype_new 'a also_wrong = Also_Wrong (*<*)'a
   3.136 +    datatype_new 'a also_wrong = W1 | W2 (*<*)'a
   3.137      typ (*>*)"('a also_wrong, 'a) fn"
   3.138  
   3.139  text {*
   3.140 @@ -344,20 +327,30 @@
   3.141  datatype_new} and @{command codatatype} commands.
   3.142  Section~\ref{sec:registering-bounded-natural-functors} explains how to register
   3.143  arbitrary type constructors as BNFs.
   3.144 +
   3.145 +Here is another example that fails:
   3.146  *}
   3.147  
   3.148 -
   3.149 -subsubsection {* Custom Names and Syntaxes
   3.150 -  \label{sssec:datatype-custom-names-and-syntaxes} *}
   3.151 +    datatype_new 'a pow_list = PNil 'a (*<*)'a
   3.152 +    datatype_new 'a pow_list' = PNil' 'a (*>*)| PCons "('a * 'a) pow_list"
   3.153 +
   3.154 +text {*
   3.155 +\noindent
   3.156 +This one features a different flavor of nesting, where the recursive call in the
   3.157 +type specification occurs around (rather than inside) another type constructor.
   3.158 +*}
   3.159 +
   3.160 +subsubsection {* Auxiliary Constants and Properties
   3.161 +  \label{sssec:datatype-auxiliary-constants-and-properties} *}
   3.162  
   3.163  text {*
   3.164  The @{command datatype_new} command introduces various constants in addition to
   3.165  the constructors. With each datatype are associated set functions, a map
   3.166  function, a relator, discriminators, and selectors, all of which can be given
   3.167 -custom names. In the example below, the traditional names
   3.168 -@{text set}, @{text map}, @{text list_all2}, @{text null}, @{text hd}, and
   3.169 -@{text tl} override the default names @{text list_set}, @{text list_map}, @{text
   3.170 -list_rel}, @{text is_Nil}, @{text un_Cons1}, and @{text un_Cons2}:
   3.171 +custom names. In the example below, the familiar names @{text null}, @{text hd},
   3.172 +@{text tl}, @{text set}, @{text map}, and @{text list_all2}, override the
   3.173 +default names @{text is_Nil}, @{text un_Cons1}, @{text un_Cons2},
   3.174 +@{text list_set}, @{text list_map}, and @{text list_rel}:
   3.175  *}
   3.176  
   3.177  (*<*)
   3.178 @@ -380,14 +373,34 @@
   3.179  
   3.180  text {*
   3.181  \noindent
   3.182 -The command introduces a discriminator @{const null} and a pair of selectors
   3.183 -@{const hd} and @{const tl} characterized as follows:
   3.184 +
   3.185 +\begin{tabular}{@ {}ll@ {}}
   3.186 +Constructors: &
   3.187 +  @{text "Nil \<Colon> 'a list"} \\
   3.188 +&
   3.189 +  @{text "Cons \<Colon> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list"} \\
   3.190 +Discriminator: &
   3.191 +  @{text "null \<Colon> 'a list \<Rightarrow> bool"} \\
   3.192 +Selectors: &
   3.193 +  @{text "hd \<Colon> 'a list \<Rightarrow> 'a"} \\
   3.194 +&
   3.195 +  @{text "tl \<Colon> 'a list \<Rightarrow> 'a list"} \\
   3.196 +Set function: &
   3.197 +  @{text "set \<Colon> 'a list \<Rightarrow> 'a set"} \\
   3.198 +Map function: &
   3.199 +  @{text "map \<Colon> ('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list"} \\
   3.200 +Relator: &
   3.201 +  @{text "list_all2 \<Colon> ('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> bool"}
   3.202 +\end{tabular}
   3.203 +
   3.204 +The discriminator @{const null} and the selectors @{const hd} and @{const tl}
   3.205 +are characterized as follows:
   3.206  %
   3.207  \[@{thm list.collapse(1)[of xs, no_vars]}
   3.208    \qquad @{thm list.collapse(2)[of xs, no_vars]}\]
   3.209  %
   3.210 -For two-constructor datatypes, a single discriminator constant suffices. The
   3.211 -discriminator associated with @{const Cons} is simply
   3.212 +For two-constructor datatypes, a single discriminator constant is sufficient.
   3.213 +The discriminator associated with @{const Cons} is simply
   3.214  @{term "\<lambda>xs. \<not> null xs"}.
   3.215  
   3.216  The @{text defaults} clause following the @{const Nil} constructor specifies a
   3.217 @@ -589,6 +602,10 @@
   3.218  or the function type. In principle, it should be possible to support old-style
   3.219  datatypes as well, but the command does not support this yet (and there is
   3.220  currently no way to register old-style datatypes as new-style datatypes).
   3.221 +
   3.222 +\item The recursor produced for types that recurse through functions has a
   3.223 +different signature than with the old package. This makes it impossible to use
   3.224 +the old \keyw{primrec} command.
   3.225  \end{itemize}
   3.226  
   3.227  An alternative to @{command datatype_new_compat} is to use the old package's
   3.228 @@ -636,7 +653,7 @@
   3.229  \noindent
   3.230  The case combinator, discriminators, and selectors are collectively called
   3.231  \emph{destructors}. The prefix ``@{text "t."}'' is an optional component of the
   3.232 -name and is normally hidden. 
   3.233 +name and is normally hidden.
   3.234  *}
   3.235  
   3.236  
   3.237 @@ -798,6 +815,10 @@
   3.238  
   3.239  \end{description}
   3.240  \end{indentblock}
   3.241 +
   3.242 +\noindent
   3.243 +In addition, equational versions of @{text t.disc} are registered with the @{text "[code]"}
   3.244 +attribute.
   3.245  *}
   3.246  
   3.247  
   3.248 @@ -818,16 +839,20 @@
   3.249  @{thm list.map(1)[no_vars]} \\
   3.250  @{thm list.map(2)[no_vars]}
   3.251  
   3.252 -\item[@{text "t."}\hthm{rel\_inject} @{text "[simp, code]"}\rm:] ~ \\
   3.253 +\item[@{text "t."}\hthm{rel\_inject} @{text "[simp]"}\rm:] ~ \\
   3.254  @{thm list.rel_inject(1)[no_vars]} \\
   3.255  @{thm list.rel_inject(2)[no_vars]}
   3.256  
   3.257 -\item[@{text "t."}\hthm{rel\_distinct} @{text "[simp, code]"}\rm:] ~ \\
   3.258 +\item[@{text "t."}\hthm{rel\_distinct} @{text "[simp]"}\rm:] ~ \\
   3.259  @{thm list.rel_distinct(1)[no_vars]} \\
   3.260  @{thm list.rel_distinct(2)[no_vars]}
   3.261  
   3.262  \end{description}
   3.263  \end{indentblock}
   3.264 +
   3.265 +\noindent
   3.266 +In addition, equational versions of @{text t.rel_inject} and @{text
   3.267 +rel_distinct} are registered with the @{text "[code]"} attribute.
   3.268  *}
   3.269  
   3.270  
   3.271 @@ -890,17 +915,18 @@
   3.272  to register new-style datatypes as old-style datatypes.
   3.273  
   3.274  \item \emph{The recursor @{text "t_rec"} has a different signature for nested
   3.275 -recursive datatypes.} In the old package, nested recursion was internally
   3.276 -reduced to mutual recursion. This reduction was visible in the type of the
   3.277 -recursor, used by \keyw{primrec}. In the new package, nested recursion is
   3.278 -handled in a more modular fashion. The old-style recursor can be generated on
   3.279 -demand using @{command primrec_new}, as explained in
   3.280 +recursive datatypes.} In the old package, nested recursion through non-functions
   3.281 +was internally reduced to mutual recursion. This reduction was visible in the
   3.282 +type of the recursor, used by \keyw{primrec}. Recursion through functions was
   3.283 +handled specially. In the new package, nested recursion (for functions and
   3.284 +non-functions) is handled in a more modular fashion. The old-style recursor can
   3.285 +be generated on demand using @{command primrec_new}, as explained in
   3.286  Section~\ref{sssec:primrec-nested-as-mutual-recursion}, if the recursion is via
   3.287  new-style datatypes.
   3.288  
   3.289 -\item \emph{Accordingly, the induction principle is different for nested
   3.290 -recursive datatypes.} Again, the old-style induction principle can be generated
   3.291 -on demand using @{command primrec_new}, as explained in
   3.292 +\item \emph{Accordingly, the induction rule is different for nested recursive
   3.293 +datatypes.} Again, the old-style induction rule can be generated on demand using
   3.294 +@{command primrec_new}, as explained in
   3.295  Section~\ref{sssec:primrec-nested-as-mutual-recursion}, if the recursion is via
   3.296  new-style datatypes.
   3.297  
   3.298 @@ -940,9 +966,9 @@
   3.299    \label{sec:defining-recursive-functions} *}
   3.300  
   3.301  text {*
   3.302 -Recursive functions over datatypes can be specified using @{command
   3.303 -primrec_new}, which supports primitive recursion, or using the more general
   3.304 -\keyw{fun} and \keyw{function} commands. Here, the focus is on @{command
   3.305 +Recursive functions over datatypes can be specified using the @{command
   3.306 +primrec_new} command, which supports primitive recursion, or using the more
   3.307 +general \keyw{fun} and \keyw{function} commands. Here, the focus is on @{command
   3.308  primrec_new}; the other two commands are described in a separate tutorial
   3.309  \cite{isabelle-function}.
   3.310  
   3.311 @@ -1026,9 +1052,10 @@
   3.312  
   3.313  text {*
   3.314  \noindent
   3.315 -The next example is not primitive recursive, but it can be defined easily using
   3.316 -\keyw{fun}. The @{command datatype_new_compat} command is needed to register
   3.317 -new-style datatypes for use with \keyw{fun} and \keyw{function}
   3.318 +The next example is defined using \keyw{fun} to escape the syntactic
   3.319 +restrictions imposed on primitive recursive functions. The
   3.320 +@{command datatype_new_compat} command is needed to register new-style datatypes
   3.321 +for use with \keyw{fun} and \keyw{function}
   3.322  (Section~\ref{sssec:datatype-new-compat}):
   3.323  *}
   3.324  
   3.325 @@ -1124,28 +1151,51 @@
   3.326  (@{text \<Rightarrow>}) is simply composition (@{text "op \<circ>"}):
   3.327  *}
   3.328  
   3.329 -    primrec_new (*<*)(in early) (*>*)ftree_map :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where
   3.330 -      "ftree_map f (FTLeaf x) = FTLeaf (f x)" |
   3.331 -      "ftree_map f (FTNode g) = FTNode (ftree_map f \<circ> g)"
   3.332 +    primrec_new (*<*)(in early) (*>*)relabel_ft :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where
   3.333 +      "relabel_ft f (FTLeaf x) = FTLeaf (f x)" |
   3.334 +      "relabel_ft f (FTNode g) = FTNode (relabel_ft f \<circ> g)"
   3.335  
   3.336  text {*
   3.337  \noindent
   3.338 -(No such map function is defined by the package because the type
   3.339 -variable @{typ 'a} is dead in @{typ "'a ftree"}.)
   3.340 -
   3.341 -Using \keyw{fun} or \keyw{function}, recursion through functions can be
   3.342 -expressed using $\lambda$-expressions and function application rather
   3.343 -than through composition. For example:
   3.344 +For convenience, recursion through functions can also be expressed using
   3.345 +$\lambda$-abstractions and function application rather than through composition.
   3.346 +For example:
   3.347  *}
   3.348  
   3.349 -    datatype_new_compat ftree
   3.350 +    primrec_new relabel_ft :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where
   3.351 +      "relabel_ft f (FTLeaf x) = FTLeaf (f x)" |
   3.352 +      "relabel_ft f (FTNode g) = FTNode (\<lambda>x. relabel_ft f (g x))"
   3.353  
   3.354  text {* \blankline *}
   3.355  
   3.356 -    function ftree_map :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where
   3.357 -      "ftree_map f (FTLeaf x) = FTLeaf (f x)" |
   3.358 -      "ftree_map f (FTNode g) = FTNode (\<lambda>x. ftree_map f (g x))"
   3.359 -    by auto (metis ftree.exhaust)
   3.360 +    primrec_new subtree_ft :: "'a \<Rightarrow> 'a ftree \<Rightarrow> 'a ftree" where
   3.361 +      "subtree_ft x (FTNode g) = g x"
   3.362 +
   3.363 +text {*
   3.364 +\noindent
   3.365 +For recursion through curried $n$-ary functions, $n$ applications of
   3.366 +@{term "op \<circ>"} are necessary. The examples below illustrate the case where
   3.367 +$n = 2$:
   3.368 +*}
   3.369 +
   3.370 +    datatype_new 'a ftree2 = FTLeaf2 'a | FTNode2 "'a \<Rightarrow> 'a \<Rightarrow> 'a ftree2"
   3.371 +
   3.372 +text {* \blankline *}
   3.373 +
   3.374 +    primrec_new (*<*)(in early) (*>*)relabel_ft2 :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree2 \<Rightarrow> 'a ftree2" where
   3.375 +      "relabel_ft2 f (FTLeaf2 x) = FTLeaf2 (f x)" |
   3.376 +      "relabel_ft2 f (FTNode2 g) = FTNode2 (op \<circ> (op \<circ> (relabel_ft2 f)) g)"
   3.377 +
   3.378 +text {* \blankline *}
   3.379 +
   3.380 +    primrec_new relabel_ft2 :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a ftree2 \<Rightarrow> 'a ftree2" where
   3.381 +      "relabel_ft2 f (FTLeaf2 x) = FTLeaf2 (f x)" |
   3.382 +      "relabel_ft2 f (FTNode2 g) = FTNode2 (\<lambda>x y. relabel_ft2 f (g x y))"
   3.383 +
   3.384 +text {* \blankline *}
   3.385 +
   3.386 +    primrec_new subtree_ft2 :: "'a \<Rightarrow> 'a \<Rightarrow> 'a ftree2 \<Rightarrow> 'a ftree2" where
   3.387 +      "subtree_ft2 x y (FTNode2 g) = g x y"
   3.388  
   3.389  
   3.390  subsubsection {* Nested-as-Mutual Recursion
   3.391 @@ -1177,12 +1227,12 @@
   3.392  
   3.393  text {*
   3.394  \noindent
   3.395 -Appropriate induction principles are generated under the names
   3.396 +Appropriate induction rules are generated as
   3.397  @{thm [source] at\<^sub>f\<^sub>f.induct},
   3.398  @{thm [source] ats\<^sub>f\<^sub>f.induct}, and
   3.399 -@{thm [source] at\<^sub>f\<^sub>f_ats\<^sub>f\<^sub>f.induct}.
   3.400 -
   3.401 -%%% TODO: Add recursors.
   3.402 +@{thm [source] at\<^sub>f\<^sub>f_ats\<^sub>f\<^sub>f.induct}. The
   3.403 +induction rules and the underlying recursors are generated on a per-need basis
   3.404 +and are kept in a cache to speed up subsequent definitions.
   3.405  
   3.406  Here is a second example:
   3.407  *}
   3.408 @@ -1340,7 +1390,7 @@
   3.409  \begin{itemize}
   3.410  \setlength{\itemsep}{0pt}
   3.411  
   3.412 -\item \emph{Theorems sometimes have different names.}
   3.413 +\item \emph{Some theorems have different names.}
   3.414  For $m > 1$ mutually recursive functions,
   3.415  @{text "f\<^sub>1_\<dots>_f\<^sub>m.simps"} has been broken down into separate
   3.416  subcollections @{text "f\<^sub>i.simps"}.
   3.417 @@ -1415,7 +1465,7 @@
   3.418  text {*
   3.419  \noindent
   3.420  Notice that the @{const cont} selector is associated with both @{const Skip}
   3.421 -and @{const Choice}.
   3.422 +and @{const Action}.
   3.423  *}
   3.424  
   3.425  
   3.426 @@ -1606,10 +1656,10 @@
   3.427    \label{sec:defining-corecursive-functions} *}
   3.428  
   3.429  text {*
   3.430 -Corecursive functions can be specified using @{command primcorec} and
   3.431 -@{command primcorecursive}, which support primitive corecursion, or using the
   3.432 -more general \keyw{partial\_function} command. Here, the focus is on
   3.433 -the former two. More examples can be found in the directory
   3.434 +Corecursive functions can be specified using the @{command primcorec} and
   3.435 +\keyw{prim\-corec\-ursive} commands, which support primitive corecursion, or
   3.436 +using the more general \keyw{partial\_function} command. Here, the focus is on
   3.437 +the first two. More examples can be found in the directory
   3.438  \verb|~~/src/HOL/BNF/Examples|.
   3.439  
   3.440  Whereas recursive functions consume datatypes one constructor at a time,
   3.441 @@ -1630,7 +1680,7 @@
   3.442  This style is popular in the coalgebraic literature.
   3.443  
   3.444  \item The \emph{constructor view} specifies $f$ by equations of the form
   3.445 -\[@{text "\<dots> \<Longrightarrow> f x\<^sub>1 \<dots> x\<^sub>n = C \<dots>"}\]
   3.446 +\[@{text "\<dots> \<Longrightarrow> f x\<^sub>1 \<dots> x\<^sub>n = C\<^sub>j \<dots>"}\]
   3.447  This style is often more concise than the previous one.
   3.448  
   3.449  \item The \emph{code view} specifies $f$ by a single equation of the form
   3.450 @@ -1643,14 +1693,6 @@
   3.451  All three styles are available as input syntax. Whichever syntax is chosen,
   3.452  characteristic theorems for all three styles are generated.
   3.453  
   3.454 -\begin{framed}
   3.455 -\noindent
   3.456 -\textbf{Warning:}\enskip The @{command primcorec} and @{command primcorecursive}
   3.457 -commands are under development. Some of the functionality described here is
   3.458 -vaporware. An alternative is to define corecursive functions directly using the
   3.459 -generated @{text t_unfold} or @{text t_corec} combinators.
   3.460 -\end{framed}
   3.461 -
   3.462  %%% TODO: partial_function? E.g. for defining tail recursive function on lazy
   3.463  %%% lists (cf. terminal0 in TLList.thy)
   3.464  *}
   3.465 @@ -1668,11 +1710,6 @@
   3.466  present the same examples expressed using the constructor and destructor views.
   3.467  *}
   3.468  
   3.469 -(*<*)
   3.470 -    locale code_view
   3.471 -    begin
   3.472 -(*>*)
   3.473 -
   3.474  subsubsection {* Simple Corecursion
   3.475    \label{sssec:primcorec-simple-corecursion} *}
   3.476  
   3.477 @@ -1683,19 +1720,19 @@
   3.478  *}
   3.479  
   3.480      primcorec literate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a llist" where
   3.481 -      "literate f x = LCons x (literate f (f x))"
   3.482 +      "literate g x = LCons x (literate g (g x))"
   3.483  
   3.484  text {* \blankline *}
   3.485  
   3.486      primcorec siterate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a stream" where
   3.487 -      "siterate f x = SCons x (siterate f (f x))"
   3.488 +      "siterate g x = SCons x (siterate g (g x))"
   3.489  
   3.490  text {*
   3.491  \noindent
   3.492  The constructor ensures that progress is made---i.e., the function is
   3.493  \emph{productive}. The above functions compute the infinite lazy list or stream
   3.494 -@{text "[x, f x, f (f x), \<dots>]"}. Productivity guarantees that prefixes
   3.495 -@{text "[x, f x, f (f x), \<dots>, (f ^^ k) x]"} of arbitrary finite length
   3.496 +@{text "[x, g x, g (g x), \<dots>]"}. Productivity guarantees that prefixes
   3.497 +@{text "[x, g x, g (g x), \<dots>, (g ^^ k) x]"} of arbitrary finite length
   3.498  @{text k} can be computed by unfolding the code equation a finite number of
   3.499  times.
   3.500  
   3.501 @@ -1714,7 +1751,7 @@
   3.502  appear around constructors that guard corecursive calls:
   3.503  *}
   3.504  
   3.505 -    primcorec_notyet lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where
   3.506 +    primcorec lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where
   3.507        "lappend xs ys =
   3.508           (case xs of
   3.509              LNil \<Rightarrow> ys
   3.510 @@ -1735,7 +1772,7 @@
   3.511  pseudorandom seed (@{text n}):
   3.512  *}
   3.513  
   3.514 -    primcorec_notyet
   3.515 +    primcorec
   3.516        random_process :: "'a stream \<Rightarrow> (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> 'a process"
   3.517      where
   3.518        "random_process s f n =
   3.519 @@ -1780,43 +1817,71 @@
   3.520  The next pair of examples generalize the @{const literate} and @{const siterate}
   3.521  functions (Section~\ref{sssec:primcorec-nested-corecursion}) to possibly
   3.522  infinite trees in which subnodes are organized either as a lazy list (@{text
   3.523 -tree\<^sub>i\<^sub>i}) or as a finite set (@{text tree\<^sub>i\<^sub>s}):
   3.524 +tree\<^sub>i\<^sub>i}) or as a finite set (@{text tree\<^sub>i\<^sub>s}). They rely on the map functions of
   3.525 +the nesting type constructors to lift the corecursive calls:
   3.526  *}
   3.527  
   3.528      primcorec iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" where
   3.529 -      "iterate\<^sub>i\<^sub>i f x = Node\<^sub>i\<^sub>i x (lmap (iterate\<^sub>i\<^sub>i f) (f x))"
   3.530 +      "iterate\<^sub>i\<^sub>i g x = Node\<^sub>i\<^sub>i x (lmap (iterate\<^sub>i\<^sub>i g) (g x))"
   3.531  
   3.532  text {* \blankline *}
   3.533  
   3.534      primcorec iterate\<^sub>i\<^sub>s :: "('a \<Rightarrow> 'a fset) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>s" where
   3.535 -      "iterate\<^sub>i\<^sub>s f x = Node\<^sub>i\<^sub>s x (fimage (iterate\<^sub>i\<^sub>s f) (f x))"
   3.536 +      "iterate\<^sub>i\<^sub>s g x = Node\<^sub>i\<^sub>s x (fimage (iterate\<^sub>i\<^sub>s g) (g x))"
   3.537  
   3.538  text {*
   3.539  \noindent
   3.540 -Deterministic finite automata (DFAs) are traditionally defined as 5-tuples
   3.541 -@{text "(Q, \<Sigma>, \<delta>, q\<^sub>0, F)"}, where @{text Q} is a finite set of states,
   3.542 +Both examples follow the usual format for constructor arguments associated
   3.543 +with nested recursive occurrences of the datatype. Consider
   3.544 +@{const iterate\<^sub>i\<^sub>i}. The term @{term "g x"} constructs an @{typ "'a llist"}
   3.545 +value, which is turned into an @{typ "'a tree\<^sub>i\<^sub>i llist"} value using
   3.546 +@{const lmap}.
   3.547 +
   3.548 +This format may sometimes feel artificial. The following function constructs
   3.549 +a tree with a single, infinite branch from a stream:
   3.550 +*}
   3.551 +
   3.552 +    primcorec tree\<^sub>i\<^sub>i_of_stream :: "'a stream \<Rightarrow> 'a tree\<^sub>i\<^sub>i" where
   3.553 +      "tree\<^sub>i\<^sub>i_of_stream s =
   3.554 +         Node\<^sub>i\<^sub>i (shd s) (lmap tree\<^sub>i\<^sub>i_of_stream (LCons (stl s) LNil))"
   3.555 +
   3.556 +text {*
   3.557 +\noindent
   3.558 +Fortunately, it is easy to prove the following lemma, where the corecursive call
   3.559 +is moved inside the lazy list constructor, thereby eliminating the need for
   3.560 +@{const lmap}:
   3.561 +*}
   3.562 +
   3.563 +    lemma tree\<^sub>i\<^sub>i_of_stream_alt:
   3.564 +      "tree\<^sub>i\<^sub>i_of_stream s = Node\<^sub>i\<^sub>i (shd s) (LCons (tree\<^sub>i\<^sub>i_of_stream (stl s)) LNil)"
   3.565 +    by (subst tree\<^sub>i\<^sub>i_of_stream.code) simp
   3.566 +
   3.567 +text {*
   3.568 +The next example illustrates corecursion through functions, which is a bit
   3.569 +special. Deterministic finite automata (DFAs) are traditionally defined as
   3.570 +5-tuples @{text "(Q, \<Sigma>, \<delta>, q\<^sub>0, F)"}, where @{text Q} is a finite set of states,
   3.571  @{text \<Sigma>} is a finite alphabet, @{text \<delta>} is a transition function, @{text q\<^sub>0}
   3.572  is an initial state, and @{text F} is a set of final states. The following
   3.573  function translates a DFA into a @{type state_machine}:
   3.574  *}
   3.575  
   3.576 -    primcorec (*<*)(in early) (*>*)
   3.577 -      sm_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> 'a state_machine"
   3.578 +    primcorec
   3.579 +      (*<*)(in early) (*>*)sm_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> 'a state_machine"
   3.580      where
   3.581 -      "sm_of_dfa \<delta> F q = State_Machine (q \<in> F) (sm_of_dfa \<delta> F o \<delta> q)"
   3.582 +      "sm_of_dfa \<delta> F q = State_Machine (q \<in> F) (sm_of_dfa \<delta> F \<circ> \<delta> q)"
   3.583  
   3.584  text {*
   3.585  \noindent
   3.586  The map function for the function type (@{text \<Rightarrow>}) is composition
   3.587 -(@{text "op \<circ>"}). For convenience, corecursion through functions can be
   3.588 -expressed using $\lambda$-expressions and function application rather
   3.589 +(@{text "op \<circ>"}). For convenience, corecursion through functions can
   3.590 +also be expressed using $\lambda$-abstractions and function application rather
   3.591  than through composition. For example:
   3.592  *}
   3.593  
   3.594      primcorec
   3.595        sm_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> 'a state_machine"
   3.596      where
   3.597 -      "sm_of_dfa \<delta> F q = State_Machine (q \<in> F) (sm_of_dfa \<delta> F o \<delta> q)"
   3.598 +      "sm_of_dfa \<delta> F q = State_Machine (q \<in> F) (\<lambda>a. sm_of_dfa \<delta> F (\<delta> q a))"
   3.599  
   3.600  text {* \blankline *}
   3.601  
   3.602 @@ -1833,9 +1898,32 @@
   3.603      primcorec
   3.604        or_sm :: "'a state_machine \<Rightarrow> 'a state_machine \<Rightarrow> 'a state_machine"
   3.605      where
   3.606 -      "or_sm M N =
   3.607 -         State_Machine (accept M \<or> accept N)
   3.608 -           (\<lambda>a. or_sm (trans M a) (trans N a))"
   3.609 +      "or_sm M N = State_Machine (accept M \<or> accept N)
   3.610 +         (\<lambda>a. or_sm (trans M a) (trans N a))"
   3.611 +
   3.612 +text {*
   3.613 +\noindent
   3.614 +For recursion through curried $n$-ary functions, $n$ applications of
   3.615 +@{term "op \<circ>"} are necessary. The examples below illustrate the case where
   3.616 +$n = 2$:
   3.617 +*}
   3.618 +
   3.619 +    codatatype ('a, 'b) state_machine2 =
   3.620 +      State_Machine2 (accept2: bool) (trans2: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) state_machine2")
   3.621 +
   3.622 +text {* \blankline *}
   3.623 +
   3.624 +    primcorec
   3.625 +      (*<*)(in early) (*>*)sm2_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> ('a, 'b) state_machine2"
   3.626 +    where
   3.627 +      "sm2_of_dfa \<delta> F q = State_Machine2 (q \<in> F) (op \<circ> (op \<circ> (sm2_of_dfa \<delta> F)) (\<delta> q))"
   3.628 +
   3.629 +text {* \blankline *}
   3.630 +
   3.631 +    primcorec
   3.632 +      sm2_of_dfa :: "('q \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'q) \<Rightarrow> 'q set \<Rightarrow> 'q \<Rightarrow> ('a, 'b) state_machine2"
   3.633 +    where
   3.634 +      "sm2_of_dfa \<delta> F q = State_Machine2 (q \<in> F) (\<lambda>a b. sm2_of_dfa \<delta> F (\<delta> q a b))"
   3.635  
   3.636  
   3.637  subsubsection {* Nested-as-Mutual Corecursion
   3.638 @@ -1848,15 +1936,31 @@
   3.639  pretend that nested codatatypes are mutually corecursive. For example:
   3.640  *}
   3.641  
   3.642 -    primcorec_notyet
   3.643 +(*<*)
   3.644 +    context late
   3.645 +    begin
   3.646 +(*>*)
   3.647 +    primcorec
   3.648        iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" and
   3.649        iterates\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a llist \<Rightarrow> 'a tree\<^sub>i\<^sub>i llist"
   3.650      where
   3.651 -      "iterate\<^sub>i\<^sub>i f x = Node\<^sub>i\<^sub>i x (iterates\<^sub>i\<^sub>i f (f x))" |
   3.652 -      "iterates\<^sub>i\<^sub>i f xs =
   3.653 +      "iterate\<^sub>i\<^sub>i g x = Node\<^sub>i\<^sub>i x (iterates\<^sub>i\<^sub>i g (g x))" |
   3.654 +      "iterates\<^sub>i\<^sub>i g xs =
   3.655           (case xs of
   3.656              LNil \<Rightarrow> LNil
   3.657 -          | LCons x xs' \<Rightarrow> LCons (iterate\<^sub>i\<^sub>i f x) (iterates\<^sub>i\<^sub>i f xs'))"
   3.658 +          | LCons x xs' \<Rightarrow> LCons (iterate\<^sub>i\<^sub>i g x) (iterates\<^sub>i\<^sub>i g xs'))"
   3.659 +
   3.660 +text {*
   3.661 +\noindent
   3.662 +Coinduction rules are generated as
   3.663 +@{thm [source] iterate\<^sub>i\<^sub>i.coinduct},
   3.664 +@{thm [source] iterates\<^sub>i\<^sub>i.coinduct}, and
   3.665 +@{thm [source] iterate\<^sub>i\<^sub>i_iterates\<^sub>i\<^sub>i.coinduct}
   3.666 +and analogously for @{text strong_coinduct}. These rules and the
   3.667 +underlying corecursors are generated on a per-need basis and are kept in a cache
   3.668 +to speed up subsequent definitions.
   3.669 +*}
   3.670 +
   3.671  (*<*)
   3.672      end
   3.673  (*>*)
   3.674 @@ -1866,7 +1970,7 @@
   3.675    \label{ssec:primrec-constructor-view} *}
   3.676  
   3.677  (*<*)
   3.678 -    locale ctr_view = code_view
   3.679 +    locale ctr_view
   3.680      begin
   3.681  (*>*)
   3.682  
   3.683 @@ -1937,7 +2041,7 @@
   3.684    \label{ssec:primrec-destructor-view} *}
   3.685  
   3.686  (*<*)
   3.687 -    locale dest_view
   3.688 +    locale dtr_view
   3.689      begin
   3.690  (*>*)
   3.691  
   3.692 @@ -1951,13 +2055,13 @@
   3.693      primcorec literate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a llist" where
   3.694        "\<not> lnull (literate _ x)" |
   3.695        "lhd (literate _ x) = x" |
   3.696 -      "ltl (literate f x) = literate f (f x)"
   3.697 +      "ltl (literate g x) = literate g (g x)"
   3.698  
   3.699  text {* \blankline *}
   3.700  
   3.701      primcorec siterate :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a stream" where
   3.702        "shd (siterate _ x) = x" |
   3.703 -      "stl (siterate f x) = siterate f (f x)"
   3.704 +      "stl (siterate g x) = siterate g (g x)"
   3.705  
   3.706  text {* \blankline *}
   3.707  
   3.708 @@ -1993,6 +2097,9 @@
   3.709  (*<*)
   3.710      end
   3.711  
   3.712 +    locale dtr_view2
   3.713 +    begin
   3.714 +
   3.715      primcorec lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist" where
   3.716        "lnull xs \<Longrightarrow> lnull ys \<Longrightarrow> lnull (lappend xs ys)" |
   3.717  (*>*)
   3.718 @@ -2000,8 +2107,6 @@
   3.719  (*<*) |
   3.720        "lhd (lappend xs ys) = lhd (if lnull xs then ys else xs)" |
   3.721        "ltl (lappend xs ys) = (if xs = LNil then ltl ys else lappend (ltl xs) ys)"
   3.722 -
   3.723 -    context dest_view begin
   3.724  (*>*)
   3.725  
   3.726  text {*
   3.727 @@ -2044,8 +2149,8 @@
   3.728  text {* \blankline *}
   3.729  
   3.730      primcorec iterate\<^sub>i\<^sub>i :: "('a \<Rightarrow> 'a llist) \<Rightarrow> 'a \<Rightarrow> 'a tree\<^sub>i\<^sub>i" where
   3.731 -      "lbl\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i f x) = x" |
   3.732 -      "sub\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i f x) = lmap (iterate\<^sub>i\<^sub>i f) (f x)"
   3.733 +      "lbl\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i g x) = x" |
   3.734 +      "sub\<^sub>i\<^sub>i (iterate\<^sub>i\<^sub>i g x) = lmap (iterate\<^sub>i\<^sub>i g) (g x)"
   3.735  (*<*)
   3.736      end
   3.737  (*>*)
   3.738 @@ -2149,13 +2254,30 @@
   3.739  
   3.740  @{rail "
   3.741    @@{command bnf} target? (name ':')? term \\
   3.742 -    term_list term term_list term?
   3.743 +    term_list term term_list? term?
   3.744    ;
   3.745    X_list: '[' (X + ',') ']'
   3.746  "}
   3.747  *}
   3.748  
   3.749  
   3.750 +(* NOTYET
   3.751 +subsubsection {* \keyw{bnf\_decl}
   3.752 +  \label{sssec:bnf-decl} *}
   3.753 +
   3.754 +text {*
   3.755 +%%% TODO: use command_def once the command is available
   3.756 +\begin{matharray}{rcl}
   3.757 +  @{text "bnf_decl"} & : & @{text "local_theory \<rightarrow> local_theory"}
   3.758 +\end{matharray}
   3.759 +
   3.760 +@{rail "
   3.761 +  @@{command bnf} target? dt_name
   3.762 +"}
   3.763 +*}
   3.764 +*)
   3.765 +
   3.766 +
   3.767  subsubsection {* \keyw{print\_bnfs}
   3.768    \label{sssec:print-bnfs} *}
   3.769  
   3.770 @@ -2307,8 +2429,9 @@
   3.771  suggested major simplifications to the internal constructions, much of which has
   3.772  yet to be implemented. Florian Haftmann and Christian Urban provided general
   3.773  advice on Isabelle and package writing. Stefan Milius and Lutz Schr\"oder
   3.774 -found an elegant proof to eliminate one of the BNF assumptions. Christian
   3.775 -Sternagel suggested many textual improvements to this tutorial.
   3.776 +found an elegant proof to eliminate one of the BNF assumptions. Andreas
   3.777 +Lochbihler and Christian Sternagel suggested many textual improvements to this
   3.778 +tutorial.
   3.779  *}
   3.780  
   3.781  end
     4.1 --- a/src/Doc/Datatypes/document/root.tex	Mon Nov 11 17:34:44 2013 +0100
     4.2 +++ b/src/Doc/Datatypes/document/root.tex	Mon Nov 11 17:44:21 2013 +0100
     4.3 @@ -58,10 +58,10 @@
     4.4  
     4.5  \begin{abstract}
     4.6  \noindent
     4.7 -This tutorial describes how to use the new package for defining datatypes and
     4.8 -codatatypes in Isabelle/HOL. The package provides five main commands:
     4.9 +This tutorial describes the new package for defining datatypes and codatatypes
    4.10 +in Isabelle/HOL. The package provides four main commands:
    4.11  \keyw{datatype\_new}, \keyw{codatatype}, \keyw{primrec\_new},
    4.12 -\keyw{primcorecursive}, and \keyw{primcorec}. The commands suffixed by
    4.13 +and \keyw{primcorec}. The commands suffixed by
    4.14  \keyw{\_new} are intended to subsume, and eventually replace, the corresponding
    4.15  commands from the old datatype package.
    4.16  \end{abstract}
     5.1 --- a/src/Doc/Functions/Functions.thy	Mon Nov 11 17:34:44 2013 +0100
     5.2 +++ b/src/Doc/Functions/Functions.thy	Mon Nov 11 17:44:21 2013 +0100
     5.3 @@ -1003,13 +1003,13 @@
     5.4    recursive calls. In general, there is one introduction rule for each
     5.5    recursive call.
     5.6  
     5.7 -  The predicate @{term "accp findzero_rel"} is the accessible part of
     5.8 +  The predicate @{term "Wellfounded.accp findzero_rel"} is the accessible part of
     5.9    that relation. An argument belongs to the accessible part, if it can
    5.10    be reached in a finite number of steps (cf.~its definition in @{text
    5.11    "Wellfounded.thy"}).
    5.12  
    5.13    Since the domain predicate is just an abbreviation, you can use
    5.14 -  lemmas for @{const accp} and @{const findzero_rel} directly. Some
    5.15 +  lemmas for @{const Wellfounded.accp} and @{const findzero_rel} directly. Some
    5.16    lemmas which are occasionally useful are @{thm [source] accpI}, @{thm [source]
    5.17    accp_downward}, and of course the introduction and elimination rules
    5.18    for the recursion relation @{thm [source] "findzero_rel.intros"} and @{thm
     6.1 --- a/src/Doc/Nitpick/document/root.tex	Mon Nov 11 17:34:44 2013 +0100
     6.2 +++ b/src/Doc/Nitpick/document/root.tex	Mon Nov 11 17:44:21 2013 +0100
     6.3 @@ -1965,6 +1965,8 @@
     6.4  \texttt{.kki}, \texttt{.cnf}, \texttt{.out}, and
     6.5  \texttt{.err}; you may safely remove them after Nitpick has run.
     6.6  
     6.7 +\textbf{Warning:} This option is not thread-safe. Use at your own risks.
     6.8 +
     6.9  \nopagebreak
    6.10  {\small See also \textit{debug} (\S\ref{output-format}).}
    6.11  \end{enum}
    6.12 @@ -2794,11 +2796,12 @@
    6.13  \subsection{Registering Coinductive Datatypes}
    6.14  \label{registering-coinductive-datatypes}
    6.15  
    6.16 +Coinductive datatypes defined using the \textbf{codatatype} command that do not
    6.17 +involve nested recursion through non-codatatypes are supported by Nitpick.
    6.18  If you have defined a custom coinductive datatype, you can tell Nitpick about
    6.19 -it, so that it can use an efficient Kodkod axiomatization similar to the one it
    6.20 -uses for lazy lists. The interface for registering and unregistering coinductive
    6.21 -datatypes consists of the following pair of functions defined in the
    6.22 -\textit{Nitpick\_HOL} structure:
    6.23 +it, so that it can use an efficient Kodkod axiomatization. The interface for
    6.24 +registering and unregistering coinductive datatypes consists of the following
    6.25 +pair of functions defined in the \textit{Nitpick\_HOL} structure:
    6.26  
    6.27  \prew
    6.28  $\textbf{val}\,~\textit{register\_codatatype\/} : {}$ \\
    6.29 @@ -2886,6 +2889,12 @@
    6.30  \item[\labelitemi] Nitpick produces spurious counterexamples when invoked after a
    6.31  \textbf{guess} command in a structured proof.
    6.32  
    6.33 +\item[\labelitemi] Datatypes defined using \textbf{datatype\_new} are not
    6.34 +supported.
    6.35 +
    6.36 +\item[\labelitemi] Codatatypes defined using \textbf{codatatype} that
    6.37 +involve nested recursion through non-codatatypes are not supported.
    6.38 +
    6.39  \item[\labelitemi] The \textit{nitpick\_xxx} attributes and the
    6.40  \textit{Nitpick\_xxx.register\_yyy} functions can cause havoc if used
    6.41  improperly.
     7.1 --- a/src/Doc/ProgProve/Bool_nat_list.thy	Mon Nov 11 17:34:44 2013 +0100
     7.2 +++ b/src/Doc/ProgProve/Bool_nat_list.thy	Mon Nov 11 17:44:21 2013 +0100
     7.3 @@ -422,10 +422,16 @@
     7.4  \subsection{Exercises}
     7.5  
     7.6  \begin{exercise}
     7.7 +Use the \isacom{value} command to evaluate the following expressions:
     7.8 +@{term[source] "1 + (2::nat)"}, @{term[source] "1 + (2::int)"},
     7.9 +@{term[source] "1 - (2::nat)"} and @{term[source] "1 - (2::int)"}.
    7.10 +\end{exercise}
    7.11 +
    7.12 +\begin{exercise}
    7.13  Start from the definition of @{const add} given above.
    7.14 -Prove it is associative (@{prop"add (add m n) p = add m (add n p)"})
    7.15 -and commutative (@{prop"add m n = add n m"}). Define a recursive function
    7.16 -@{text double} @{text"::"} @{typ"nat \<Rightarrow> nat"} and prove that @{prop"double m = add m m"}.
    7.17 +Prove that @{const add} it is associative and commutative.
    7.18 +Define a recursive function @{text double} @{text"::"} @{typ"nat \<Rightarrow> nat"}
    7.19 +and prove @{prop"double m = add m m"}.
    7.20  \end{exercise}
    7.21  
    7.22  \begin{exercise}
    7.23 @@ -436,11 +442,15 @@
    7.24  
    7.25  \begin{exercise}
    7.26  Define a recursive function @{text "snoc ::"} @{typ"'a list \<Rightarrow> 'a \<Rightarrow> 'a list"}
    7.27 -that appends an element to the end of a list. Do not use the predefined append
    7.28 -operator @{text"@"}. With the help of @{text snoc} define a recursive function
    7.29 -@{text "reverse ::"} @{typ"'a list \<Rightarrow> 'a list"} that reverses a list. Do not
    7.30 -use the predefined function @{const rev}.
    7.31 -Prove @{prop"reverse(reverse xs) = xs"}.
    7.32 +that appends an element to the end of a list. With the help of @{text snoc}
    7.33 +define a recursive function @{text "reverse ::"} @{typ"'a list \<Rightarrow> 'a list"}
    7.34 +that reverses a list. Prove @{prop"reverse(reverse xs) = xs"}.
    7.35 +\end{exercise}
    7.36 +
    7.37 +\begin{exercise}
    7.38 +Define a recursive function @{text "sum ::"} @{typ"nat \<Rightarrow> nat"} such that
    7.39 +\mbox{@{text"sum n"}} @{text"="} @{text"0 + ... + n"} and prove
    7.40 +@{prop" sum(n::nat) = n * (n+1) div 2"}.
    7.41  \end{exercise}
    7.42  *}
    7.43  (*<*)
     8.1 --- a/src/Doc/ProgProve/Isar.thy	Mon Nov 11 17:34:44 2013 +0100
     8.2 +++ b/src/Doc/ProgProve/Isar.thy	Mon Nov 11 17:44:21 2013 +0100
     8.3 @@ -595,10 +595,10 @@
     8.4  \exercise
     8.5  Give a readable, structured proof of the following lemma:
     8.6  *}
     8.7 -lemma assumes T: "\<forall> x y. T x y \<or> T y x"
     8.8 -  and A: "\<forall> x y. A x y \<and> A y x \<longrightarrow> x = y"
     8.9 -  and TA: "\<forall> x y. T x y \<longrightarrow> A x y" and "A x y"
    8.10 -shows "T x y"
    8.11 +lemma assumes T: "\<forall>x y. T x y \<or> T y x"
    8.12 +  and A: "\<forall>x y. A x y \<and> A y x \<longrightarrow> x = y"
    8.13 +  and TA: "\<forall>x y. T x y \<longrightarrow> A x y" and "A x y"
    8.14 +  shows "T x y"
    8.15  (*<*)oops(*>*)
    8.16  text{*
    8.17  \endexercise
    8.18 @@ -612,10 +612,11 @@
    8.19  text{*
    8.20  Hint: There are predefined functions @{const_typ take} and @{const_typ drop}
    8.21  such that @{text"take k [x\<^sub>1,\<dots>] = [x\<^sub>1,\<dots>,x\<^sub>k]"} and
    8.22 -@{text"drop k [x\<^sub>1,\<dots>] = [x\<^bsub>k+1\<^esub>,\<dots>]"}. Let @{text simp} and especially
    8.23 -sledgehammer find and apply the relevant @{const take} and @{const drop} lemmas for you.
    8.24 +@{text"drop k [x\<^sub>1,\<dots>] = [x\<^bsub>k+1\<^esub>,\<dots>]"}. Let sledgehammer find and apply
    8.25 +the relevant @{const take} and @{const drop} lemmas for you.
    8.26  \endexercise
    8.27  
    8.28 +
    8.29  \section{Case Analysis and Induction}
    8.30  
    8.31  \subsection{Datatype Case Analysis}
    8.32 @@ -1075,45 +1076,38 @@
    8.33  @{text induct} method.
    8.34  \end{warn}
    8.35  
    8.36 +
    8.37  \subsection{Exercises}
    8.38  
    8.39 +
    8.40 +\exercise
    8.41 +Give a structured proof by rule inversion:
    8.42 +*}
    8.43 +
    8.44 +lemma assumes a: "ev(Suc(Suc n))" shows "ev n"
    8.45 +(*<*)oops(*>*)
    8.46 +
    8.47 +text{*
    8.48 +\endexercise
    8.49 +
    8.50 +\begin{exercise}
    8.51 +Give a structured proof of @{prop "\<not> ev(Suc(Suc(Suc 0)))"}
    8.52 +by rule inversions. If there are no cases to be proved you can close
    8.53 +a proof immediateley with \isacom{qed}.
    8.54 +\end{exercise}
    8.55 +
    8.56 +\begin{exercise}
    8.57 +Recall predicate @{text star} from \autoref{sec:star} and @{text iter}
    8.58 +from Exercise~\ref{exe:iter}. Prove @{prop "iter r n x y \<Longrightarrow> star r x y"}
    8.59 +in a structured style, do not just sledgehammer each case of the
    8.60 +required induction.
    8.61 +\end{exercise}
    8.62 +
    8.63  \begin{exercise}
    8.64  Define a recursive function @{text "elems ::"} @{typ"'a list \<Rightarrow> 'a set"}
    8.65  and prove @{prop "x : elems xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> elems ys"}.
    8.66  \end{exercise}
    8.67 -
    8.68 -\begin{exercise}
    8.69 -A context-free grammar can be seen as an inductive definition where each
    8.70 -nonterminal $A$ is an inductively defined predicate on lists of terminal
    8.71 -symbols: $A(w)$ mans
    8.72 -that $w$ is in the language generated by $A$. For example, the production $S
    8.73 -\to a S b$ can be viewed as the implication @{prop"S w \<Longrightarrow> S (a # w @ [b])"}
    8.74 -where @{text a} and @{text b} are constructors of some datatype of terminal
    8.75 -symbols: \isacom{datatype} @{text"tsymbs = a | b | \<dots>"}
    8.76 -
    8.77 -Define the two grammars
    8.78 -\[
    8.79 -\begin{array}{r@ {\quad}c@ {\quad}l}
    8.80 -S &\to& \varepsilon \quad\mid\quad a~S~b \quad\mid\quad S~S \\
    8.81 -T &\to& \varepsilon \quad\mid\quad T~a~T~b
    8.82 -\end{array}
    8.83 -\]
    8.84 -($\varepsilon$ is the empty word)
    8.85 -as two inductive predicates and prove @{prop"S w \<longleftrightarrow> T w"}.
    8.86 -\end{exercise}
    8.87 -
    8.88  *}
    8.89 -(*
    8.90 -lemma "\<not> ev(Suc(Suc(Suc 0)))"
    8.91 -proof
    8.92 -  assume "ev(Suc(Suc(Suc 0)))"
    8.93 -  then show False
    8.94 -  proof cases
    8.95 -    case evSS
    8.96 -    from `ev(Suc 0)` show False by cases
    8.97 -  qed
    8.98 -qed
    8.99 -*)
   8.100  
   8.101  (*<*)
   8.102  end
     9.1 --- a/src/Doc/ProgProve/Logic.thy	Mon Nov 11 17:34:44 2013 +0100
     9.2 +++ b/src/Doc/ProgProve/Logic.thy	Mon Nov 11 17:44:21 2013 +0100
     9.3 @@ -141,6 +141,28 @@
     9.4  See \cite{Nipkow-Main} for the wealth of further predefined functions in theory
     9.5  @{theory Main}.
     9.6  
     9.7 +
     9.8 +\subsection{Exercises}
     9.9 +
    9.10 +\exercise
    9.11 +Start from the data type of binary trees defined earlier:
    9.12 +*}
    9.13 +
    9.14 +datatype 'a tree = Tip | Node "'a tree" 'a "'a tree"
    9.15 +
    9.16 +text{*
    9.17 +Define a function @{text "set ::"} @{typ "'a tree \<Rightarrow> 'a set"}
    9.18 +that returns the elements in a tree and a function
    9.19 +@{text "ord ::"} @{typ "int tree \<Rightarrow> bool"}
    9.20 +the tests if an @{typ "int tree"} is ordered.
    9.21 +
    9.22 +Define a function @{text ins} that inserts an element into an ordered @{typ "int tree"}
    9.23 +while maintaining the order of the tree. If the element is already in the tree, the
    9.24 +same tree should be returned. Prove correctness of @{text ins}:
    9.25 +@{prop "set(ins x t) = {x} \<union> set t"} and @{prop "ord t \<Longrightarrow> ord(ins i t)"}.
    9.26 +\endexercise
    9.27 +
    9.28 +
    9.29  \section{Proof Automation}
    9.30  
    9.31  So far we have only seen @{text simp} and @{text auto}: Both perform
    9.32 @@ -459,12 +481,12 @@
    9.33  text{* In this particular example we could have backchained with
    9.34  @{thm[source] Suc_leD}, too, but because the premise is more complicated than the conclusion this can easily lead to nontermination.
    9.35  
    9.36 -\subsection{Finding Theorems}
    9.37 -
    9.38 -Command \isacom{find{\isacharunderscorekeyword}theorems} searches for specific theorems in the current
    9.39 -theory. Search criteria include pattern matching on terms and on names.
    9.40 -For details see the Isabelle/Isar Reference Manual~\cite{IsarRef}.
    9.41 -\bigskip
    9.42 +%\subsection{Finding Theorems}
    9.43 +%
    9.44 +%Command \isacom{find{\isacharunderscorekeyword}theorems} searches for specific theorems in the current
    9.45 +%theory. Search criteria include pattern matching on terms and on names.
    9.46 +%For details see the Isabelle/Isar Reference Manual~\cite{IsarRef}.
    9.47 +%\bigskip
    9.48  
    9.49  \begin{warn}
    9.50  To ease readability we will drop the question marks
    9.51 @@ -708,8 +730,8 @@
    9.52  apply(rename_tac u x y)
    9.53  defer
    9.54  (*>*)
    9.55 -txt{* The induction is over @{prop"star r x y"} and we try to prove
    9.56 -\mbox{@{prop"star r y z \<Longrightarrow> star r x z"}},
    9.57 +txt{* The induction is over @{prop"star r x y"} (the first matching assumption)
    9.58 +and we try to prove \mbox{@{prop"star r y z \<Longrightarrow> star r x z"}},
    9.59  which we abbreviate by @{prop"P x y"}. These are our two subgoals:
    9.60  @{subgoals[display,indent=0]}
    9.61  The first one is @{prop"P x x"}, the result of case @{thm[source]refl},
    9.62 @@ -764,6 +786,95 @@
    9.63  conditions}. In rule inductions, these side-conditions appear as additional
    9.64  assumptions. The \isacom{for} clause seen in the definition of the reflexive
    9.65  transitive closure merely simplifies the form of the induction rule.
    9.66 +
    9.67 +
    9.68 +\subsection{Exercises}
    9.69 +
    9.70 +\begin{exercise}
    9.71 +Formalise the following definition of palindromes
    9.72 +\begin{itemize}
    9.73 +\item The empty list and a singleton list are palindromes.
    9.74 +\item If @{text xs} is a palindrome, so is @{term "a # xs @ [a]"}.
    9.75 +\end{itemize}
    9.76 +as an inductive predicate @{text "palindrome ::"} @{typ "'a list \<Rightarrow> bool"}
    9.77 +and prove that @{prop "rev xs = xs"} if @{text xs} is a palindrome.
    9.78 +\end{exercise}
    9.79 +
    9.80 +\exercise
    9.81 +We could also have defined @{const star} as follows:
    9.82 +*}
    9.83 +
    9.84 +inductive star' :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" for r where
    9.85 +refl': "star' r x x" |
    9.86 +step': "star' r x y \<Longrightarrow> r y z \<Longrightarrow> star' r x z"
    9.87 +
    9.88 +text{*
    9.89 +The single @{text r} step is performer after rather than before the @{text star'}
    9.90 +steps. Prove @{prop "star' r x y \<Longrightarrow> star r x y"} and
    9.91 +@{prop "star r x y \<Longrightarrow> star r' x y"}. You may need lemmas.
    9.92 +Note that rule induction fails
    9.93 +if the assumption about the inductive predicate is not the first assumption.
    9.94 +\endexercise
    9.95 +
    9.96 +\begin{exercise}\label{exe:iter}
    9.97 +Analogous to @{const star}, give an inductive definition of the @{text n}-fold iteration
    9.98 +of a relation @{text r}: @{term "iter r n x y"} should hold if there are @{text x\<^sub>0}, \dots, @{text x\<^sub>n}
    9.99 +such that @{prop"x = x\<^sub>0"}, @{prop"x\<^sub>n = y"} and @{text"r x\<^bsub>i\<^esub> x\<^bsub>i+1\<^esub>"} for
   9.100 +all @{prop"i < n"}. Correct and prove the following claim:
   9.101 +@{prop"star r x y \<Longrightarrow> iter r n x y"}.
   9.102 +\end{exercise}
   9.103 +
   9.104 +\begin{exercise}
   9.105 +A context-free grammar can be seen as an inductive definition where each
   9.106 +nonterminal $A$ is an inductively defined predicate on lists of terminal
   9.107 +symbols: $A(w)$ mans that $w$ is in the language generated by $A$.
   9.108 +For example, the production $S \to a S b$ can be viewed as the implication
   9.109 +@{prop"S w \<Longrightarrow> S (a # w @ [b])"} where @{text a} and @{text b} are terminal symbols,
   9.110 +i.e., elements of some alphabet. The alphabet can be defined like this:
   9.111 +\isacom{datatype} @{text"alpha = a | b | \<dots>"}
   9.112 +
   9.113 +Define the two grammars (where $\varepsilon$ is the empty word)
   9.114 +\[
   9.115 +\begin{array}{r@ {\quad}c@ {\quad}l}
   9.116 +S &\to& \varepsilon \quad\mid\quad aSb \quad\mid\quad SS \\
   9.117 +T &\to& \varepsilon \quad\mid\quad TaTb
   9.118 +\end{array}
   9.119 +\]
   9.120 +as two inductive predicates.
   9.121 +If you think of @{text a} and @{text b} as ``@{text "("}'' and  ``@{text ")"}'',
   9.122 +the grammars defines strings of balanced parentheses.
   9.123 +Prove @{prop"T w \<Longrightarrow> S w"} and @{prop "S w \<Longrightarrow> T w"} separately and conclude
   9.124 +@{prop "S w = T w"}.
   9.125 +\end{exercise}
   9.126 +
   9.127 +\ifsem
   9.128 +\begin{exercise}
   9.129 +In \autoref{sec:AExp} we defined a recursive evaluation function
   9.130 +@{text "aval :: aexp \<Rightarrow> state \<Rightarrow> val"}.
   9.131 +Define an inductive evaluation predicate
   9.132 +@{text "aval_rel :: aexp \<Rightarrow> state \<Rightarrow> val \<Rightarrow> bool"}
   9.133 +and prove that it agrees with the recursive function:
   9.134 +@{prop "aval_rel a s v \<Longrightarrow> aval a s = v"}, 
   9.135 +@{prop "aval a s = v \<Longrightarrow> aval_rel a s v"} and thus
   9.136 +\noquotes{@{prop [source] "aval_rel a s v \<longleftrightarrow> aval a s = v"}}.
   9.137 +\end{exercise}
   9.138 +
   9.139 +\begin{exercise}
   9.140 +Consider the stack machine from Chapter~3
   9.141 +and recall the concept of \concept{stack underflow}
   9.142 +from Exercise~\ref{exe:stack-underflow}.
   9.143 +Define an inductive predicate
   9.144 +@{text "ok :: nat \<Rightarrow> instr list \<Rightarrow> nat \<Rightarrow> bool"}
   9.145 +such that @{text "ok n is n'"} means that with any initial stack of length
   9.146 +@{text n} the instructions @{text "is"} can be executed
   9.147 +without stack underflow and that the final stack has length @{text n'}.
   9.148 +Prove that @{text ok} correctly computes the final stack size
   9.149 +@{prop[display] "\<lbrakk>ok n is n'; length stk = n\<rbrakk> \<Longrightarrow> length (exec is s stk) = n'"}
   9.150 +and that instruction sequences generated by @{text comp}
   9.151 +cannot cause stack underflow: \ @{text "ok n (comp a) ?"} \ for
   9.152 +some suitable value of @{text "?"}.
   9.153 +\end{exercise}
   9.154 +\fi
   9.155  *}
   9.156  (*<*)
   9.157  end
    10.1 --- a/src/Doc/ProgProve/Types_and_funs.thy	Mon Nov 11 17:34:44 2013 +0100
    10.2 +++ b/src/Doc/ProgProve/Types_and_funs.thy	Mon Nov 11 17:44:21 2013 +0100
    10.3 @@ -156,7 +156,7 @@
    10.4  
    10.5  fun div2 :: "nat \<Rightarrow> nat" where
    10.6  "div2 0 = 0" |
    10.7 -"div2 (Suc 0) = Suc 0" |
    10.8 +"div2 (Suc 0) = 0" |
    10.9  "div2 (Suc(Suc n)) = Suc(div2 n)"
   10.10  
   10.11  text{* does not just define @{const div2} but also proves a
   10.12 @@ -200,6 +200,34 @@
   10.13  But note that the induction rule does not mention @{text f} at all,
   10.14  except in its name, and is applicable independently of @{text f}.
   10.15  
   10.16 +
   10.17 +\subsection{Exercises}
   10.18 +
   10.19 +\begin{exercise}
   10.20 +Starting from the type @{text "'a tree"} defined in the text, define
   10.21 +a function @{text "contents ::"} @{typ "'a tree \<Rightarrow> 'a list"}
   10.22 +that collects all values in a tree in a list, in any order,
   10.23 +without removing duplicates.
   10.24 +Then define a function @{text "treesum ::"} @{typ "nat tree \<Rightarrow> nat"}
   10.25 +that sums up all values in a tree of natural numbers
   10.26 +and prove @{prop "treesum t = listsum(contents t)"}.
   10.27 +\end{exercise}
   10.28 +
   10.29 +\begin{exercise}
   10.30 +Define a new type @{text "'a tree2"} of binary trees where values are also
   10.31 +stored in the leaves of the tree.  Also reformulate the
   10.32 +@{const mirror} function accordingly. Define two functions
   10.33 +@{text "pre_order"} and @{text "post_order"} of type @{text "'a tree2 \<Rightarrow> 'a list"}
   10.34 +that traverse a tree and collect all stored values in the respective order in
   10.35 +a list. Prove @{prop "pre_order (mirror t) = rev (post_order t)"}.
   10.36 +\end{exercise}
   10.37 +
   10.38 +\begin{exercise}
   10.39 +Prove that @{const div2} defined above divides every number by @{text 2},
   10.40 +not just those of the form @{text"n+n"}: @{prop "div2 n = n div 2"}.
   10.41 +\end{exercise}
   10.42 +
   10.43 +
   10.44  \section{Induction Heuristics}
   10.45  
   10.46  We have already noted that theorems about recursive functions are proved by
   10.47 @@ -307,6 +335,18 @@
   10.48  matters in some cases. The variables that need to be quantified are typically
   10.49  those that change in recursive calls.
   10.50  
   10.51 +
   10.52 +\subsection{Exercises}
   10.53 +
   10.54 +\begin{exercise}
   10.55 +Write a tail-recursive variant of the @{text add} function on @{typ nat}:
   10.56 +@{term "itadd :: nat \<Rightarrow> nat \<Rightarrow> nat"}.
   10.57 +Tail-recursive means that in the recursive case, @{text itadd} needs to call
   10.58 +itself directly: \mbox{@{term"itadd (Suc m) n"}} @{text"= itadd \<dots>"}.
   10.59 +Prove @{prop "itadd m n = add m n"}.
   10.60 +\end{exercise}
   10.61 +
   10.62 +
   10.63  \section{Simplification}
   10.64  
   10.65  So far we have talked a lot about simplifying terms without explaining the concept. \concept{Simplification} means
   10.66 @@ -485,6 +525,31 @@
   10.67  
   10.68  \subsection{Exercises}
   10.69  
   10.70 +\exercise\label{exe:tree0}
   10.71 +Define a datatype @{text tree0} of binary tree skeletons which do not store
   10.72 +any information, neither in the inner nodes nor in the leaves.
   10.73 +Define a function @{text "nodes :: tree0 \<Rightarrow> nat"} that counts the total number
   10.74 +all nodes (inner nodes and leaves) in such a tree.
   10.75 +Consider the following recursive function:
   10.76 +*}
   10.77 +(*<*)
   10.78 +datatype tree0 = Tip | Node tree0 tree0
   10.79 +(*>*)
   10.80 +fun explode :: "nat \<Rightarrow> tree0 \<Rightarrow> tree0" where
   10.81 +"explode 0 t = t" |
   10.82 +"explode (Suc n) t = explode n (Node t t)"
   10.83 +
   10.84 +text {*
   10.85 +Find an equation expressing the size of a tree after exploding it
   10.86 +(\noquotes{@{term [source] "nodes (explode n t)"}}) as a function
   10.87 +of @{term "nodes t"} and @{text n}. Prove your equation.
   10.88 +You may use the usual arithmetic operators including the exponentiation
   10.89 +operator ``@{text"^"}''. For example, \noquotes{@{prop [source] "2 ^ 2 = 4"}}.
   10.90 +
   10.91 +Hint: simplifying with the list of theorems @{thm[source] algebra_simps}
   10.92 +takes care of common algebraic properties of the arithmetic operators.
   10.93 +\endexercise
   10.94 +
   10.95  \exercise
   10.96  Define arithmetic expressions in one variable over integers (type @{typ int})
   10.97  as a data type:
   10.98 @@ -506,8 +571,7 @@
   10.99  that transforms an expression into a polynomial. This may require auxiliary
  10.100  functions. Prove that @{text coeffs} preserves the value of the expression:
  10.101  \mbox{@{prop"evalp (coeffs e) x = eval e x"}.}
  10.102 -Hint: simplifying with @{thm[source] algebra_simps} takes care of
  10.103 -common algebraic properties of @{text "+"} and @{text "*"}.
  10.104 +Hint: consider the hint in \autoref{exe:tree0}.
  10.105  \endexercise
  10.106  *}
  10.107  (*<*)
    11.1 --- a/src/Doc/Sledgehammer/document/root.tex	Mon Nov 11 17:34:44 2013 +0100
    11.2 +++ b/src/Doc/Sledgehammer/document/root.tex	Mon Nov 11 17:44:21 2013 +0100
    11.3 @@ -121,8 +121,8 @@
    11.4  
    11.5  For Isabelle/jEdit users, Sledgehammer provides an automatic mode that can be
    11.6  enabled via the ``Auto Sledgehammer'' option under ``Plugins > Plugin Options >
    11.7 -Isabelle > General.'' In this mode, Sledgehammer is run on every newly entered
    11.8 -theorem.
    11.9 +Isabelle > General.'' In this mode, a reduced version of Sledgehammer is run on
   11.10 +every newly entered theorem for a few seconds.
   11.11  
   11.12  \newbox\boxA
   11.13  \setbox\boxA=\hbox{\texttt{NOSPAM}}
   11.14 @@ -719,12 +719,16 @@
   11.15  If you use Isabelle/jEdit, Sledgehammer also provides an automatic mode that can
   11.16  be enabled via the ``Auto Sledgehammer'' option under ``Plugins > Plugin Options
   11.17  > Isabelle > General.'' For automatic runs, only the first prover set using
   11.18 -\textit{provers} (\S\ref{mode-of-operation}) is considered, fewer facts are
   11.19 -passed to the prover, \textit{slice} (\S\ref{mode-of-operation}) is disabled,
   11.20 -\textit{strict} (\S\ref{problem-encoding}) is enabled, \textit{verbose}
   11.21 -(\S\ref{output-format}) and \textit{debug} (\S\ref{output-format}) are disabled,
   11.22 -and \textit{timeout} (\S\ref{timeouts}) is superseded by the ``Auto Time Limit''
   11.23 -option in jEdit. Sledgehammer's output is also more concise.
   11.24 +\textit{provers} (\S\ref{mode-of-operation}) is considered (typically E),
   11.25 +\textit{slice} (\S\ref{mode-of-operation}) is disabled,
   11.26 +\textit{minimize} (\S\ref{mode-of-operation}) is disabled, fewer facts are
   11.27 +passed to the prover, \textit{fact\_filter} (\S\ref{relevance-filter}) is set to
   11.28 +\textit{mepo}, \textit{strict} (\S\ref{problem-encoding}) is enabled,
   11.29 +\textit{verbose} (\S\ref{output-format}) and \textit{debug}
   11.30 +(\S\ref{output-format}) are disabled, \textit{preplay\_timeout}
   11.31 +(\S\ref{timeouts}) is set to 0, and \textit{timeout} (\S\ref{timeouts}) is
   11.32 +superseded by the ``Auto Time Limit'' option in jEdit. Sledgehammer's output is
   11.33 +also more concise.
   11.34  
   11.35  \subsection{Metis}
   11.36  
   11.37 @@ -999,8 +1003,7 @@
   11.38  number of facts. For SMT solvers, several slices are tried with the same options
   11.39  each time but fewer and fewer facts. According to benchmarks with a timeout of
   11.40  30 seconds, slicing is a valuable optimization, and you should probably leave it
   11.41 -enabled unless you are conducting experiments. This option is implicitly
   11.42 -disabled for (short) automatic runs.
   11.43 +enabled unless you are conducting experiments.
   11.44  
   11.45  \nopagebreak
   11.46  {\small See also \textit{verbose} (\S\ref{output-format}).}
   11.47 @@ -1035,6 +1038,8 @@
   11.48  simultaneously. The files are identified by the prefixes \texttt{prob\_} and
   11.49  \texttt{mash\_}; you may safely remove them after Sledgehammer has run.
   11.50  
   11.51 +\textbf{Warning:} This option is not thread-safe. Use at your own risks.
   11.52 +
   11.53  \nopagebreak
   11.54  {\small See also \textit{debug} (\S\ref{output-format}).}
   11.55  \end{enum}
   11.56 @@ -1282,14 +1287,12 @@
   11.57  
   11.58  \opfalse{verbose}{quiet}
   11.59  Specifies whether the \textbf{sledgehammer} command should explain what it does.
   11.60 -This option is implicitly disabled for automatic runs.
   11.61  
   11.62  \opfalse{debug}{no\_debug}
   11.63  Specifies whether Sledgehammer should display additional debugging information
   11.64  beyond what \textit{verbose} already displays. Enabling \textit{debug} also
   11.65  enables \textit{verbose} and \textit{blocking} (\S\ref{mode-of-operation})
   11.66 -behind the scenes. The \textit{debug} option is implicitly disabled for
   11.67 -automatic runs.
   11.68 +behind the scenes.
   11.69  
   11.70  \nopagebreak
   11.71  {\small See also \textit{spy} (\S\ref{mode-of-operation}) and
   11.72 @@ -1349,8 +1352,6 @@
   11.73  \opdefault{timeout}{float\_or\_none}{\upshape 30}
   11.74  Specifies the maximum number of seconds that the automatic provers should spend
   11.75  searching for a proof. This excludes problem preparation and is a soft limit.
   11.76 -For automatic runs, the ``Auto Time Limit'' option under ``Plugins > Plugin
   11.77 -Options > Isabelle > General'' is used instead.
   11.78  
   11.79  \opdefault{preplay\_timeout}{float\_or\_none}{\upshape 3}
   11.80  Specifies the maximum number of seconds that \textit{metis} or \textit{smt}
    12.1 --- a/src/Doc/Tutorial/document/sets.tex	Mon Nov 11 17:34:44 2013 +0100
    12.2 +++ b/src/Doc/Tutorial/document/sets.tex	Mon Nov 11 17:44:21 2013 +0100
    12.3 @@ -660,8 +660,8 @@
    12.4  \textbf{Composition} of relations (the infix \sdx{O}) is also
    12.5  available: 
    12.6  \begin{isabelle}
    12.7 -r\ O\ s\ \isasymequiv\ \isacharbraceleft(x,z).\ \isasymexists y.\ (x,y)\ \isasymin\ s\ \isasymand\ (y,z)\ \isasymin\ r\isacharbraceright
    12.8 -\rulenamedx{rel_comp_def}
    12.9 +r\ O\ s\ = \isacharbraceleft(x,z).\ \isasymexists y.\ (x,y)\ \isasymin\ s\ \isasymand\ (y,z)\ \isasymin\ r\isacharbraceright
   12.10 +\rulenamedx{relcomp_unfold}
   12.11  \end{isabelle}
   12.12  %
   12.13  This is one of the many lemmas proved about these concepts: 
   12.14 @@ -677,7 +677,7 @@
   12.15  \isasymlbrakk r\isacharprime\ \isasymsubseteq\ r;\ s\isacharprime\
   12.16  \isasymsubseteq\ s\isasymrbrakk\ \isasymLongrightarrow\ r\isacharprime\ O\
   12.17  s\isacharprime\ \isasymsubseteq\ r\ O\ s%
   12.18 -\rulename{rel_comp_mono}
   12.19 +\rulename{relcomp_mono}
   12.20  \end{isabelle}
   12.21  
   12.22  \indexbold{converse!of a relation}%
   12.23 @@ -695,7 +695,7 @@
   12.24  Here is a typical law proved about converse and composition: 
   12.25  \begin{isabelle}
   12.26  (r\ O\ s)\isasyminverse\ =\ s\isasyminverse\ O\ r\isasyminverse
   12.27 -\rulename{converse_rel_comp}
   12.28 +\rulename{converse_relcomp}
   12.29  \end{isabelle}
   12.30  
   12.31  \indexbold{image!under a relation}%
    13.1 --- a/src/Doc/manual.bib	Mon Nov 11 17:34:44 2013 +0100
    13.2 +++ b/src/Doc/manual.bib	Mon Nov 11 17:44:21 2013 +0100
    13.3 @@ -926,7 +926,7 @@
    13.4    note = "\url{https://github.com/frelindb/agsyHOL}"}
    13.5  
    13.6  @incollection{lochbihler-2010,
    13.7 -  title = "Coinduction",
    13.8 +  title = "Coinductive",
    13.9    author = "Andreas Lochbihler",
   13.10    booktitle = "The Archive of Formal Proofs",
   13.11    editor = "Gerwin Klein and Tobias Nipkow and Lawrence C. Paulson",
    14.1 --- a/src/HOL/ATP.thy	Mon Nov 11 17:34:44 2013 +0100
    14.2 +++ b/src/HOL/ATP.thy	Mon Nov 11 17:44:21 2013 +0100
    14.3 @@ -18,34 +18,34 @@
    14.4  
    14.5  subsection {* Higher-order reasoning helpers *}
    14.6  
    14.7 -definition fFalse :: bool where [no_atp]:
    14.8 +definition fFalse :: bool where
    14.9  "fFalse \<longleftrightarrow> False"
   14.10  
   14.11 -definition fTrue :: bool where [no_atp]:
   14.12 +definition fTrue :: bool where
   14.13  "fTrue \<longleftrightarrow> True"
   14.14  
   14.15 -definition fNot :: "bool \<Rightarrow> bool" where [no_atp]:
   14.16 +definition fNot :: "bool \<Rightarrow> bool" where
   14.17  "fNot P \<longleftrightarrow> \<not> P"
   14.18  
   14.19 -definition fComp :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" where [no_atp]:
   14.20 +definition fComp :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" where
   14.21  "fComp P = (\<lambda>x. \<not> P x)"
   14.22  
   14.23 -definition fconj :: "bool \<Rightarrow> bool \<Rightarrow> bool" where [no_atp]:
   14.24 +definition fconj :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
   14.25  "fconj P Q \<longleftrightarrow> P \<and> Q"
   14.26  
   14.27 -definition fdisj :: "bool \<Rightarrow> bool \<Rightarrow> bool" where [no_atp]:
   14.28 +definition fdisj :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
   14.29  "fdisj P Q \<longleftrightarrow> P \<or> Q"
   14.30  
   14.31 -definition fimplies :: "bool \<Rightarrow> bool \<Rightarrow> bool" where [no_atp]:
   14.32 +definition fimplies :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
   14.33  "fimplies P Q \<longleftrightarrow> (P \<longrightarrow> Q)"
   14.34  
   14.35 -definition fequal :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where [no_atp]:
   14.36 +definition fequal :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
   14.37  "fequal x y \<longleftrightarrow> (x = y)"
   14.38  
   14.39 -definition fAll :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where [no_atp]:
   14.40 +definition fAll :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where
   14.41  "fAll P \<longleftrightarrow> All P"
   14.42  
   14.43 -definition fEx :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where [no_atp]:
   14.44 +definition fEx :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where
   14.45  "fEx P \<longleftrightarrow> Ex P"
   14.46  
   14.47  lemma fTrue_ne_fFalse: "fFalse \<noteq> fTrue"
    15.1 --- a/src/HOL/Archimedean_Field.thy	Mon Nov 11 17:34:44 2013 +0100
    15.2 +++ b/src/HOL/Archimedean_Field.thy	Mon Nov 11 17:44:21 2013 +0100
    15.3 @@ -129,12 +129,8 @@
    15.4    fix y z assume
    15.5      "of_int y \<le> x \<and> x < of_int (y + 1)"
    15.6      "of_int z \<le> x \<and> x < of_int (z + 1)"
    15.7 -  then have
    15.8 -    "of_int y \<le> x" "x < of_int (y + 1)"
    15.9 -    "of_int z \<le> x" "x < of_int (z + 1)"
   15.10 -    by simp_all
   15.11 -  from le_less_trans [OF `of_int y \<le> x` `x < of_int (z + 1)`]
   15.12 -       le_less_trans [OF `of_int z \<le> x` `x < of_int (y + 1)`]
   15.13 +  with le_less_trans [of "of_int y" "x" "of_int (z + 1)"]
   15.14 +       le_less_trans [of "of_int z" "x" "of_int (y + 1)"]
   15.15    show "y = z" by (simp del: of_int_add)
   15.16  qed
   15.17  
    16.1 --- a/src/HOL/BNF/BNF_FP_Base.thy	Mon Nov 11 17:34:44 2013 +0100
    16.2 +++ b/src/HOL/BNF/BNF_FP_Base.thy	Mon Nov 11 17:44:21 2013 +0100
    16.3 @@ -172,7 +172,5 @@
    16.4  ML_file "Tools/bnf_fp_n2m.ML"
    16.5  ML_file "Tools/bnf_fp_n2m_sugar.ML"
    16.6  ML_file "Tools/bnf_fp_rec_sugar_util.ML"
    16.7 -ML_file "Tools/bnf_fp_rec_sugar_tactics.ML"
    16.8 -ML_file "Tools/bnf_fp_rec_sugar.ML"
    16.9  
   16.10  end
    17.1 --- a/src/HOL/BNF/BNF_GFP.thy	Mon Nov 11 17:34:44 2013 +0100
    17.2 +++ b/src/HOL/BNF/BNF_GFP.thy	Mon Nov 11 17:44:21 2013 +0100
    17.3 @@ -308,6 +308,8 @@
    17.4  lemma fun_rel_image2p: "(fun_rel R (image2p f g R)) f g"
    17.5    unfolding fun_rel_def image2p_def by auto
    17.6  
    17.7 +ML_file "Tools/bnf_gfp_rec_sugar_tactics.ML"
    17.8 +ML_file "Tools/bnf_gfp_rec_sugar.ML"
    17.9  ML_file "Tools/bnf_gfp_util.ML"
   17.10  ML_file "Tools/bnf_gfp_tactics.ML"
   17.11  ML_file "Tools/bnf_gfp.ML"
    18.1 --- a/src/HOL/BNF/BNF_LFP.thy	Mon Nov 11 17:34:44 2013 +0100
    18.2 +++ b/src/HOL/BNF/BNF_LFP.thy	Mon Nov 11 17:44:21 2013 +0100
    18.3 @@ -230,6 +230,7 @@
    18.4  lemma predicate2D_vimage2p: "\<lbrakk>R \<le> vimage2p f g S; R x y\<rbrakk> \<Longrightarrow> S (f x) (g y)"
    18.5    unfolding vimage2p_def by auto
    18.6  
    18.7 +ML_file "Tools/bnf_lfp_rec_sugar.ML"
    18.8  ML_file "Tools/bnf_lfp_util.ML"
    18.9  ML_file "Tools/bnf_lfp_tactics.ML"
   18.10  ML_file "Tools/bnf_lfp.ML"
    19.1 --- a/src/HOL/BNF/Basic_BNFs.thy	Mon Nov 11 17:34:44 2013 +0100
    19.2 +++ b/src/HOL/BNF/Basic_BNFs.thy	Mon Nov 11 17:44:21 2013 +0100
    19.3 @@ -27,15 +27,14 @@
    19.4  lemma wpull_Grp_def: "wpull A B1 B2 f1 f2 p1 p2 \<longleftrightarrow> Grp B1 f1 OO (Grp B2 f2)\<inverse>\<inverse> \<le> (Grp A p1)\<inverse>\<inverse> OO Grp A p2"
    19.5    unfolding wpull_def Grp_def by auto
    19.6  
    19.7 -bnf ID: "id :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" ["\<lambda>x. {x}"] "\<lambda>_:: 'a. natLeq" ["id :: 'a \<Rightarrow> 'a"]
    19.8 +bnf ID: "id :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" ["\<lambda>x. {x}"] "\<lambda>_:: 'a. natLeq"
    19.9    "id :: ('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
   19.10  apply (auto simp: Grp_def fun_eq_iff relcompp.simps natLeq_card_order natLeq_cinfinite)
   19.11  apply (rule ordLess_imp_ordLeq[OF finite_ordLess_infinite[OF _ natLeq_Well_order]])
   19.12  apply (auto simp add: Field_card_of Field_natLeq card_of_well_order_on)[3]
   19.13  done
   19.14  
   19.15 -bnf DEADID: "id :: 'a \<Rightarrow> 'a" [] "\<lambda>_:: 'a. natLeq +c |UNIV :: 'a set|" ["SOME x :: 'a. True"]
   19.16 -  "op = :: 'a \<Rightarrow> 'a \<Rightarrow> bool"
   19.17 +bnf DEADID: "id :: 'a \<Rightarrow> 'a" [] "\<lambda>_:: 'a. natLeq +c |UNIV :: 'a set|" "op = :: 'a \<Rightarrow> 'a \<Rightarrow> bool"
   19.18  by (auto simp add: wpull_Grp_def Grp_def
   19.19    card_order_csum natLeq_card_order card_of_card_order_on
   19.20    cinfinite_csum natLeq_cinfinite)
   19.21 @@ -148,7 +147,7 @@
   19.22  
   19.23  lemmas prod_set_defs = fsts_def[abs_def] snds_def[abs_def]
   19.24  
   19.25 -bnf map_pair [fsts, snds] "\<lambda>_::'a \<times> 'b. natLeq" [Pair] prod_rel
   19.26 +bnf map_pair [fsts, snds] "\<lambda>_::'a \<times> 'b. natLeq" prod_rel
   19.27  proof (unfold prod_set_defs)
   19.28    show "map_pair id id = id" by (rule map_pair.id)
   19.29  next
   19.30 @@ -193,7 +192,7 @@
   19.31          Grp {x. {fst x} \<subseteq> Collect (split R) \<and> {snd x} \<subseteq> Collect (split S)} (map_pair snd snd)"
   19.32    unfolding prod_set_defs prod_rel_def Grp_def relcompp.simps conversep.simps fun_eq_iff
   19.33    by auto
   19.34 -qed simp+
   19.35 +qed
   19.36  
   19.37  (* Categorical version of pullback: *)
   19.38  lemma wpull_cat:
   19.39 @@ -231,7 +230,7 @@
   19.40    ultimately show ?thesis using card_of_ordLeq by fast
   19.41  qed
   19.42  
   19.43 -bnf "op \<circ>" [range] "\<lambda>_:: 'a \<Rightarrow> 'b. natLeq +c |UNIV :: 'a set|" ["%c x::'b::type. c::'a::type"]
   19.44 +bnf "op \<circ>" [range] "\<lambda>_:: 'a \<Rightarrow> 'b. natLeq +c |UNIV :: 'a set|"
   19.45    "fun_rel op ="
   19.46  proof
   19.47    fix f show "id \<circ> f = id f" by simp
   19.48 @@ -278,6 +277,6 @@
   19.49           Grp {x. range x \<subseteq> Collect (split R)} (op \<circ> snd)"
   19.50    unfolding fun_rel_def Grp_def fun_eq_iff relcompp.simps conversep.simps  subset_iff image_iff
   19.51    by auto (force, metis pair_collapse)
   19.52 -qed auto
   19.53 +qed
   19.54  
   19.55  end
    20.1 --- a/src/HOL/BNF/Examples/Misc_Codatatype.thy	Mon Nov 11 17:34:44 2013 +0100
    20.2 +++ b/src/HOL/BNF/Examples/Misc_Codatatype.thy	Mon Nov 11 17:44:21 2013 +0100
    20.3 @@ -19,9 +19,9 @@
    20.4  
    20.5  codatatype simple'' = X1'' nat int | X2''
    20.6  
    20.7 -codatatype 'a stream = Stream 'a "'a stream"
    20.8 +codatatype 'a stream = Stream (shd: 'a) (stl: "'a stream")
    20.9  
   20.10 -codatatype 'a mylist = MyNil | MyCons 'a "'a mylist"
   20.11 +codatatype 'a mylist = MyNil | MyCons (myhd: 'a) (mytl: "'a mylist")
   20.12  
   20.13  codatatype ('b, 'c, 'd, 'e) some_passive =
   20.14    SP1 "('b, 'c, 'd, 'e) some_passive" | SP2 'b | SP3 'c | SP4 'd | SP5 'e
    21.1 --- a/src/HOL/BNF/Examples/Misc_Datatype.thy	Mon Nov 11 17:34:44 2013 +0100
    21.2 +++ b/src/HOL/BNF/Examples/Misc_Datatype.thy	Mon Nov 11 17:44:21 2013 +0100
    21.3 @@ -19,7 +19,7 @@
    21.4  
    21.5  datatype_new simple'' = X1'' nat int | X2''
    21.6  
    21.7 -datatype_new 'a mylist = MyNil | MyCons 'a "'a mylist"
    21.8 +datatype_new 'a mylist = MyNil | MyCons (myhd: 'a) (mytl: "'a mylist")
    21.9  
   21.10  datatype_new ('b, 'c, 'd, 'e) some_passive =
   21.11    SP1 "('b, 'c, 'd, 'e) some_passive" | SP2 'b | SP3 'c | SP4 'd | SP5 'e
    22.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    22.2 +++ b/src/HOL/BNF/Examples/Misc_Primcorec.thy	Mon Nov 11 17:44:21 2013 +0100
    22.3 @@ -0,0 +1,112 @@
    22.4 +(*  Title:      HOL/BNF/Examples/Misc_Primcorec.thy
    22.5 +    Author:     Jasmin Blanchette, TU Muenchen
    22.6 +    Copyright   2013
    22.7 +
    22.8 +Miscellaneous primitive corecursive function definitions.
    22.9 +*)
   22.10 +
   22.11 +header {* Miscellaneous Primitive Corecursive Function Definitions *}
   22.12 +
   22.13 +theory Misc_Primcorec
   22.14 +imports Misc_Codatatype
   22.15 +begin
   22.16 +
   22.17 +primcorec simple_of_bools :: "bool \<Rightarrow> bool \<Rightarrow> simple" where
   22.18 +  "simple_of_bools b b' = (if b then if b' then X1 else X2 else if b' then X3 else X4)"
   22.19 +
   22.20 +primcorec simple'_of_bools :: "bool \<Rightarrow> bool \<Rightarrow> simple'" where
   22.21 +  "simple'_of_bools b b' =
   22.22 +     (if b then if b' then X1' () else X2' () else if b' then X3' () else X4' ())"
   22.23 +
   22.24 +primcorec inc_simple'' :: "nat \<Rightarrow> simple'' \<Rightarrow> simple''" where
   22.25 +  "inc_simple'' k s = (case s of X1'' n i \<Rightarrow> X1'' (n + k) (i + int k) | X2'' \<Rightarrow> X2'')"
   22.26 +
   22.27 +primcorec sinterleave :: "'a stream \<Rightarrow> 'a stream \<Rightarrow> 'a stream" where
   22.28 +  "sinterleave s s' = Stream (shd s) (sinterleave s' (stl s))"
   22.29 +
   22.30 +primcorec myapp :: "'a mylist \<Rightarrow> 'a mylist \<Rightarrow> 'a mylist" where
   22.31 +  "myapp xs ys =
   22.32 +     (if xs = MyNil then ys
   22.33 +      else if ys = MyNil then xs
   22.34 +      else MyCons (myhd xs) (myapp (mytl xs) ys))"
   22.35 +
   22.36 +primcorec shuffle_sp :: "('a, 'b, 'c, 'd) some_passive \<Rightarrow> ('d, 'a, 'b, 'c) some_passive" where
   22.37 +  "shuffle_sp sp =
   22.38 +     (case sp of
   22.39 +       SP1 sp' \<Rightarrow> SP1 (shuffle_sp sp')
   22.40 +     | SP2 a \<Rightarrow> SP3 a
   22.41 +     | SP3 b \<Rightarrow> SP4 b
   22.42 +     | SP4 c \<Rightarrow> SP5 c
   22.43 +     | SP5 d \<Rightarrow> SP2 d)"
   22.44 +
   22.45 +primcorec rename_lam :: "(string \<Rightarrow> string) \<Rightarrow> lambda \<Rightarrow> lambda" where
   22.46 +  "rename_lam f l =
   22.47 +     (case l of
   22.48 +       Var s \<Rightarrow> Var (f s)
   22.49 +     | App l l' \<Rightarrow> App (rename_lam f l) (rename_lam f l')
   22.50 +     | Abs s l \<Rightarrow> Abs (f s) (rename_lam f l)
   22.51 +     | Let SL l \<Rightarrow> Let (fimage (map_pair f (rename_lam f)) SL) (rename_lam f l))"
   22.52 +
   22.53 +primcorec
   22.54 +  j1_sum :: "('a\<Colon>{zero,one,plus}) \<Rightarrow> 'a J1" and
   22.55 +  j2_sum :: "'a \<Rightarrow> 'a J2"
   22.56 +where
   22.57 +  "n = 0 \<Longrightarrow> is_J11 (j1_sum n)" |
   22.58 +  "un_J111 (j1_sum _) = 0" |
   22.59 +  "un_J112 (j1_sum _) = j1_sum 0" |
   22.60 +  "un_J121 (j1_sum n) = n + 1" |
   22.61 +  "un_J122 (j1_sum n) = j2_sum (n + 1)" |
   22.62 +  "n = 0 \<Longrightarrow> is_J21 (j2_sum n)" |
   22.63 +  "un_J221 (j2_sum n) = j1_sum (n + 1)" |
   22.64 +  "un_J222 (j2_sum n) = j2_sum (n + 1)"
   22.65 +
   22.66 +primcorec forest_of_mylist :: "'a tree mylist \<Rightarrow> 'a forest" where
   22.67 +  "forest_of_mylist ts =
   22.68 +     (case ts of
   22.69 +       MyNil \<Rightarrow> FNil
   22.70 +     | MyCons t ts \<Rightarrow> FCons t (forest_of_mylist ts))"
   22.71 +
   22.72 +primcorec mylist_of_forest :: "'a forest \<Rightarrow> 'a tree mylist" where
   22.73 +  "mylist_of_forest f =
   22.74 +     (case f of
   22.75 +       FNil \<Rightarrow> MyNil
   22.76 +     | FCons t ts \<Rightarrow> MyCons t (mylist_of_forest ts))"
   22.77 +
   22.78 +primcorec semi_stream :: "'a stream \<Rightarrow> 'a stream" where
   22.79 +  "semi_stream s = Stream (shd s) (semi_stream (stl (stl s)))"
   22.80 +
   22.81 +primcorec
   22.82 +  tree'_of_stream :: "'a stream \<Rightarrow> 'a tree'" and
   22.83 +  branch_of_stream :: "'a stream \<Rightarrow> 'a branch"
   22.84 +where
   22.85 +  "tree'_of_stream s =
   22.86 +     TNode' (branch_of_stream (semi_stream s)) (branch_of_stream (semi_stream (stl s)))" |
   22.87 +  "branch_of_stream s = (case s of Stream h t \<Rightarrow> Branch h (tree'_of_stream t))"
   22.88 +
   22.89 +primcorec
   22.90 +  freeze_exp :: "('b \<Rightarrow> 'a) \<Rightarrow> ('a, 'b) exp \<Rightarrow> ('a, 'b) exp" and
   22.91 +  freeze_trm :: "('b \<Rightarrow> 'a) \<Rightarrow> ('a, 'b) trm \<Rightarrow> ('a, 'b) trm" and
   22.92 +  freeze_factor :: "('b \<Rightarrow> 'a) \<Rightarrow> ('a, 'b) factor \<Rightarrow> ('a, 'b) factor"
   22.93 +where
   22.94 +  "freeze_exp g e =
   22.95 +     (case e of
   22.96 +       Term t \<Rightarrow> Term (freeze_trm g t)
   22.97 +     | Sum t e \<Rightarrow> Sum (freeze_trm g t) (freeze_exp g e))" |
   22.98 +  "freeze_trm g t =
   22.99 +     (case t of
  22.100 +       Factor f \<Rightarrow> Factor (freeze_factor g f)
  22.101 +     | Prod f t \<Rightarrow> Prod (freeze_factor g f) (freeze_trm g t))" |
  22.102 +  "freeze_factor g f =
  22.103 +     (case f of
  22.104 +       C a \<Rightarrow> C a
  22.105 +     | V b \<Rightarrow> C (g b)
  22.106 +     | Paren e \<Rightarrow> Paren (freeze_exp g e))"
  22.107 +
  22.108 +primcorec poly_unity :: "'a poly_unit" where
  22.109 +  "poly_unity = U (\<lambda>_. poly_unity)"
  22.110 +
  22.111 +primcorec build_cps :: "('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> bool stream) \<Rightarrow> 'a \<Rightarrow> bool stream \<Rightarrow> 'a cps" where
  22.112 +  "shd b \<Longrightarrow> build_cps f g a b = CPS1 a" |
  22.113 +  "_ \<Longrightarrow> build_cps f g a b = CPS2 (\<lambda>a. build_cps f g (f a) (g a))"
  22.114 +
  22.115 +end
    23.1 --- a/src/HOL/BNF/Examples/Misc_Primrec.thy	Mon Nov 11 17:34:44 2013 +0100
    23.2 +++ b/src/HOL/BNF/Examples/Misc_Primrec.thy	Mon Nov 11 17:44:21 2013 +0100
    23.3 @@ -14,7 +14,7 @@
    23.4  primrec_new nat_of_simple :: "simple \<Rightarrow> nat" where
    23.5    "nat_of_simple X1 = 1" |
    23.6    "nat_of_simple X2 = 2" |
    23.7 -  "nat_of_simple X3 = 2" |
    23.8 +  "nat_of_simple X3 = 3" |
    23.9    "nat_of_simple X4 = 4"
   23.10  
   23.11  primrec_new simple_of_simple' :: "simple' \<Rightarrow> simple" where
    24.1 --- a/src/HOL/BNF/Tools/bnf_comp_tactics.ML	Mon Nov 11 17:34:44 2013 +0100
    24.2 +++ b/src/HOL/BNF/Tools/bnf_comp_tactics.ML	Mon Nov 11 17:44:21 2013 +0100
    24.3 @@ -164,10 +164,9 @@
    24.4  fun mk_comp_wit_tac ctxt Gwit_thms collect_set_map Fwit_thms =
    24.5    ALLGOALS (dtac @{thm in_Union_o_assoc}) THEN
    24.6    unfold_thms_tac ctxt (collect_set_map :: comp_wit_thms) THEN
    24.7 -  REPEAT_DETERM (
    24.8 -    atac 1 ORELSE
    24.9 -    REPEAT_DETERM (eresolve_tac @{thms UnionE UnE imageE} 1) THEN
   24.10 -    (TRY o dresolve_tac Gwit_thms THEN'
   24.11 +  REPEAT_DETERM ((atac ORELSE'
   24.12 +    REPEAT_DETERM o eresolve_tac @{thms UnionE UnE} THEN'
   24.13 +    etac imageE THEN' TRY o dresolve_tac Gwit_thms THEN'
   24.14      (etac FalseE ORELSE'
   24.15      hyp_subst_tac ctxt THEN'
   24.16      dresolve_tac Fwit_thms THEN'
    25.1 --- a/src/HOL/BNF/Tools/bnf_def.ML	Mon Nov 11 17:34:44 2013 +0100
    25.2 +++ b/src/HOL/BNF/Tools/bnf_def.ML	Mon Nov 11 17:44:21 2013 +0100
    25.3 @@ -77,14 +77,20 @@
    25.4    val wit_thms_of_bnf: bnf -> thm list
    25.5    val wit_thmss_of_bnf: bnf -> thm list list
    25.6  
    25.7 +  val mk_map: int -> typ list -> typ list -> term -> term
    25.8 +  val mk_rel: int -> typ list -> typ list -> term -> term
    25.9 +  val build_map: Proof.context -> (typ * typ -> term) -> typ * typ -> term
   25.10 +  val build_rel: Proof.context -> (typ * typ -> term) -> typ * typ -> term
   25.11 +  val flatten_type_args_of_bnf: bnf -> 'a -> 'a list -> 'a list
   25.12 +  val map_flattened_map_args: Proof.context -> string -> (term list -> 'a list) -> term list ->
   25.13 +    'a list
   25.14 +
   25.15    val mk_witness: int list * term -> thm list -> nonemptiness_witness
   25.16    val minimize_wits: (''a list * 'b) list -> (''a list * 'b) list
   25.17    val wits_of_bnf: bnf -> nonemptiness_witness list
   25.18  
   25.19    val zip_axioms: 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list
   25.20  
   25.21 -  val flatten_type_args_of_bnf: bnf -> 'a -> 'a list -> 'a list
   25.22 -
   25.23    datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline
   25.24    datatype fact_policy = Dont_Note | Note_Some | Note_All
   25.25  
   25.26 @@ -447,7 +453,6 @@
   25.27    #> Option.map (morph_bnf (Morphism.thm_morphism (Thm.transfer (Proof_Context.theory_of ctxt))));
   25.28  
   25.29  
   25.30 -
   25.31  (* Utilities *)
   25.32  
   25.33  fun normalize_set insts instA set =
   25.34 @@ -487,6 +492,46 @@
   25.35         else minimize ((I, wit) :: done) todo;
   25.36   in minimize [] wits end;
   25.37  
   25.38 +fun mk_map live Ts Us t =
   25.39 +  let val (Type (_, Ts0), Type (_, Us0)) = strip_typeN (live + 1) (fastype_of t) |>> List.last in
   25.40 +    Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
   25.41 +  end;
   25.42 +
   25.43 +fun mk_rel live Ts Us t =
   25.44 +  let val [Type (_, Ts0), Type (_, Us0)] = binder_types (snd (strip_typeN live (fastype_of t))) in
   25.45 +    Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
   25.46 +  end;
   25.47 +
   25.48 +fun build_map_or_rel mk const of_bnf dest ctxt build_simple =
   25.49 +  let
   25.50 +    fun build (TU as (T, U)) =
   25.51 +      if T = U then
   25.52 +        const T
   25.53 +      else
   25.54 +        (case TU of
   25.55 +          (Type (s, Ts), Type (s', Us)) =>
   25.56 +          if s = s' then
   25.57 +            let
   25.58 +              val bnf = the (bnf_of ctxt s);
   25.59 +              val live = live_of_bnf bnf;
   25.60 +              val mapx = mk live Ts Us (of_bnf bnf);
   25.61 +              val TUs' = map dest (fst (strip_typeN live (fastype_of mapx)));
   25.62 +            in Term.list_comb (mapx, map build TUs') end
   25.63 +          else
   25.64 +            build_simple TU
   25.65 +        | _ => build_simple TU);
   25.66 +  in build end;
   25.67 +
   25.68 +val build_map = build_map_or_rel mk_map HOLogic.id_const map_of_bnf dest_funT;
   25.69 +val build_rel = build_map_or_rel mk_rel HOLogic.eq_const rel_of_bnf dest_pred2T;
   25.70 +
   25.71 +fun map_flattened_map_args ctxt s map_args fs =
   25.72 +  let
   25.73 +    val flat_fs = flatten_type_args_of_bnf (the (bnf_of ctxt s)) Term.dummy fs;
   25.74 +    val flat_fs' = map_args flat_fs;
   25.75 +  in
   25.76 +    permute_like (op aconv) flat_fs fs flat_fs'
   25.77 +  end;
   25.78  
   25.79  
   25.80  (* Names *)
   25.81 @@ -612,14 +657,12 @@
   25.82      val fact_policy = mk_fact_policy no_defs_lthy;
   25.83      val bnf_b = qualify raw_bnf_b;
   25.84      val live = length raw_sets;
   25.85 -    val nwits = length raw_wits;
   25.86  
   25.87      val map_rhs = prep_term no_defs_lthy raw_map;
   25.88      val set_rhss = map (prep_term no_defs_lthy) raw_sets;
   25.89      val (bd_rhsT, bd_rhs) = (case prep_term no_defs_lthy raw_bd_Abs of
   25.90        Abs (_, T, t) => (T, t)
   25.91      | _ => error "Bad bound constant");
   25.92 -    val wit_rhss = map (prep_term no_defs_lthy) raw_wits;
   25.93  
   25.94      fun err T =
   25.95        error ("Trying to register the type " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^
   25.96 @@ -633,7 +676,7 @@
   25.97          | T => err T)
   25.98        else (bnf_b, Local_Theory.full_name no_defs_lthy bnf_b);
   25.99  
  25.100 -    val def_qualify = Binding.qualify false (Binding.name_of bnf_b);
  25.101 +    val def_qualify = Binding.conceal o Binding.qualify false (Binding.name_of bnf_b);
  25.102  
  25.103      fun mk_suffix_binding suf = Binding.suffix_name ("_" ^ suf) bnf_b;
  25.104  
  25.105 @@ -672,21 +715,14 @@
  25.106            else map (fn i => set_name i (fn () => mk_suffix_binding (mk_setN i))) (1 upto live);
  25.107        in bs ~~ set_rhss end;
  25.108      val bd_bind_def = (fn () => def_qualify (mk_suffix_binding bdN), bd_rhs);
  25.109 -    val wit_binds_defs =
  25.110 -      let
  25.111 -        val bs = if nwits = 1 then [fn () => def_qualify (mk_suffix_binding witN)]
  25.112 -          else map (fn i => fn () => def_qualify (mk_suffix_binding (mk_witN i))) (1 upto nwits);
  25.113 -      in bs ~~ wit_rhss end;
  25.114  
  25.115 -    val (((((bnf_map_term, raw_map_def),
  25.116 +    val ((((bnf_map_term, raw_map_def),
  25.117        (bnf_set_terms, raw_set_defs)),
  25.118 -      (bnf_bd_term, raw_bd_def)),
  25.119 -      (bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) =
  25.120 +      (bnf_bd_term, raw_bd_def)), (lthy, lthy_old)) =
  25.121          no_defs_lthy
  25.122          |> maybe_define true map_bind_def
  25.123          ||>> apfst split_list o fold_map (maybe_define true) set_binds_defs
  25.124          ||>> maybe_define true bd_bind_def
  25.125 -        ||>> apfst split_list o fold_map (maybe_define true) wit_binds_defs
  25.126          ||> `(maybe_restore no_defs_lthy);
  25.127  
  25.128      val phi = Proof_Context.export_morphism lthy_old lthy;
  25.129 @@ -694,7 +730,6 @@
  25.130      val bnf_map_def = Morphism.thm phi raw_map_def;
  25.131      val bnf_set_defs = map (Morphism.thm phi) raw_set_defs;
  25.132      val bnf_bd_def = Morphism.thm phi raw_bd_def;
  25.133 -    val bnf_wit_defs = map (Morphism.thm phi) raw_wit_defs;
  25.134  
  25.135      val bnf_map = Morphism.term phi bnf_map_term;
  25.136  
  25.137 @@ -713,7 +748,6 @@
  25.138      val bdT = Morphism.typ phi bd_rhsT;
  25.139      val bnf_bd =
  25.140        Term.subst_TVars (Term.add_tvar_namesT bdT [] ~~ CA_params) (Morphism.term phi bnf_bd_term);
  25.141 -    val bnf_wits = map (normalize_wit CA_params CA alphas o Morphism.term phi) bnf_wit_terms;
  25.142  
  25.143      (*TODO: assert Ds = (TVars of bnf_map) \ (alphas @ betas) as sets*)
  25.144      val deads = (case Ds_opt of
  25.145 @@ -770,7 +804,6 @@
  25.146      val bnf_sets_As = map (mk_bnf_t As') bnf_sets;
  25.147      val bnf_sets_Bs = map (mk_bnf_t Bs') bnf_sets;
  25.148      val bnf_bd_As = mk_bnf_t As' bnf_bd;
  25.149 -    val bnf_wit_As = map (apsnd (mk_bnf_t As')) bnf_wits;
  25.150  
  25.151      val pre_names_lthy = lthy;
  25.152      val ((((((((((((((((((((((((fs, gs), hs), x), y), zs), ys), As),
  25.153 @@ -827,9 +860,23 @@
  25.154        (fn () => def_qualify (if Binding.is_empty rel_b then mk_suffix_binding relN else rel_b),
  25.155           rel_rhs);
  25.156  
  25.157 -    val ((bnf_rel_term, raw_rel_def), (lthy, lthy_old)) =
  25.158 +    val wit_rhss =
  25.159 +      if null raw_wits then
  25.160 +        [fold_rev Term.absdummy As' (Term.list_comb (bnf_map_AsAs,
  25.161 +          map2 (fn T => fn i => Term.absdummy T (Bound i)) As' (live downto 1)) $
  25.162 +          Const (@{const_name undefined}, CA'))]
  25.163 +      else map (prep_term no_defs_lthy) raw_wits;
  25.164 +    val nwits = length wit_rhss;
  25.165 +    val wit_binds_defs =
  25.166 +      let
  25.167 +        val bs = if nwits = 1 then [fn () => def_qualify (mk_suffix_binding witN)]
  25.168 +          else map (fn i => fn () => def_qualify (mk_suffix_binding (mk_witN i))) (1 upto nwits);
  25.169 +      in bs ~~ wit_rhss end;
  25.170 +
  25.171 +    val (((bnf_rel_term, raw_rel_def), (bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) =
  25.172        lthy
  25.173        |> maybe_define (is_some raw_rel_opt) rel_bind_def
  25.174 +      ||>> apfst split_list o fold_map (maybe_define (not (null raw_wits))) wit_binds_defs
  25.175        ||> `(maybe_restore lthy);
  25.176  
  25.177      val phi = Proof_Context.export_morphism lthy_old lthy;
  25.178 @@ -841,11 +888,9 @@
  25.179      val rel = mk_bnf_rel pred2RTs CA' CB';
  25.180      val relAsAs = mk_bnf_rel self_pred2RTs CA' CA';
  25.181  
  25.182 -    val _ = case no_reflexive (raw_map_def :: raw_set_defs @ [raw_bd_def] @
  25.183 -        raw_wit_defs @ [raw_rel_def]) of
  25.184 -        [] => ()
  25.185 -      | defs => Proof_Display.print_consts true lthy_old (K false)
  25.186 -          (map (dest_Free o fst o Logic.dest_equals o prop_of) defs);
  25.187 +    val bnf_wit_defs = map (Morphism.thm phi) raw_wit_defs;
  25.188 +    val bnf_wits = map (normalize_wit CA_params CA alphas o Morphism.term phi) bnf_wit_terms;
  25.189 +    val bnf_wit_As = map (apsnd (mk_bnf_t As')) bnf_wits;
  25.190  
  25.191      val map_id0_goal =
  25.192        let val bnf_map_app_id = Term.list_comb (bnf_map_AsAs, map HOLogic.id_const As') in
  25.193 @@ -945,11 +990,14 @@
  25.194          map wit_goal (0 upto live - 1)
  25.195        end;
  25.196  
  25.197 -    val wit_goalss = map mk_wit_goals bnf_wit_As;
  25.198 +    val trivial_wit_tac = mk_trivial_wit_tac bnf_wit_defs;
  25.199  
  25.200 -    fun after_qed thms lthy =
  25.201 +    val wit_goalss =
  25.202 +      (if null raw_wits then SOME trivial_wit_tac else NONE, map mk_wit_goals bnf_wit_As);
  25.203 +
  25.204 +    fun after_qed mk_wit_thms thms lthy =
  25.205        let
  25.206 -        val (axioms, wit_thms) = apfst (mk_axioms live) (chop (length goals) thms);
  25.207 +        val (axioms, nontriv_wit_thms) = apfst (mk_axioms live) (chop (length goals) thms);
  25.208  
  25.209          val bd_Card_order = #bd_card_order axioms RS @{thm conjunct2[OF card_order_on_Card_order]};
  25.210          val bd_Cinfinite = @{thm conjI} OF [#bd_cinfinite axioms, bd_Card_order];
  25.211 @@ -1022,6 +1070,9 @@
  25.212  
  25.213          val set_map = map (fn thm => Lazy.lazy (fn () => mk_set_map thm)) (#set_map0 axioms);
  25.214  
  25.215 +        val wit_thms =
  25.216 +          if null nontriv_wit_thms then mk_wit_thms (map Lazy.force set_map) else nontriv_wit_thms;
  25.217 +
  25.218          fun mk_in_bd () =
  25.219            let
  25.220              val bdT = fst (dest_relT (fastype_of bnf_bd_As));
  25.221 @@ -1265,35 +1316,45 @@
  25.222    (bnf, Local_Theory.declaration {syntax = false, pervasive = true}
  25.223      (fn phi => Data.map (Symtab.default (key, morph_bnf phi bnf))) lthy);
  25.224  
  25.225 -(* TODO: Once the invariant "nwits > 0" holds, remove "mk_conjunction_balanced'" and "rtac TrueI"
  25.226 -   below *)
  25.227 -fun mk_conjunction_balanced' [] = @{prop True}
  25.228 -  | mk_conjunction_balanced' ts = Logic.mk_conjunction_balanced ts;
  25.229 -
  25.230  fun bnf_def const_policy fact_policy qualify tacs wit_tac Ds map_b rel_b set_bs =
  25.231 -  (fn (_, goals, wit_goalss, after_qed, lthy, one_step_defs) =>
  25.232 +  (fn (_, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, one_step_defs) =>
  25.233    let
  25.234 -    val wits_tac =
  25.235 -      K (TRYALL Goal.conjunction_tac) THEN' K (TRYALL (rtac TrueI)) THEN'
  25.236 -      mk_unfold_thms_then_tac lthy one_step_defs wit_tac;
  25.237 -    val wit_goals = map mk_conjunction_balanced' wit_goalss;
  25.238 -    val wit_thms =
  25.239 -      Goal.prove_sorry lthy [] [] (mk_conjunction_balanced' wit_goals) wits_tac
  25.240 -      |> Conjunction.elim_balanced (length wit_goals)
  25.241 -      |> map2 (Conjunction.elim_balanced o length) wit_goalss
  25.242 -      |> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));
  25.243 +    fun mk_wits_tac set_maps =
  25.244 +      K (TRYALL Goal.conjunction_tac) THEN'
  25.245 +      (case triv_tac_opt of
  25.246 +        SOME tac => tac set_maps
  25.247 +      | NONE => mk_unfold_thms_then_tac lthy one_step_defs wit_tac);
  25.248 +    val wit_goals = map Logic.mk_conjunction_balanced wit_goalss;
  25.249 +    fun mk_wit_thms set_maps =
  25.250 +      Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals) (mk_wits_tac set_maps)
  25.251 +        |> Conjunction.elim_balanced (length wit_goals)
  25.252 +        |> map2 (Conjunction.elim_balanced o length) wit_goalss
  25.253 +        |> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));
  25.254    in
  25.255      map2 (Thm.close_derivation oo Goal.prove_sorry lthy [] [])
  25.256        goals (map (mk_unfold_thms_then_tac lthy one_step_defs) tacs)
  25.257 -    |> (fn thms => after_qed (map single thms @ wit_thms) lthy)
  25.258 +    |> (fn thms => after_qed mk_wit_thms (map single thms) lthy)
  25.259    end) oo prepare_def const_policy fact_policy qualify (K I) Ds map_b rel_b set_bs;
  25.260  
  25.261 -val bnf_cmd = (fn (key, goals, wit_goals, after_qed, lthy, defs) =>
  25.262 -  Proof.unfolding ([[(defs, [])]])
  25.263 -    (Proof.theorem NONE (snd o register_bnf key oo after_qed)
  25.264 -      (map (single o rpair []) goals @ map (map (rpair [])) wit_goals) lthy)) oo
  25.265 -  prepare_def Do_Inline (user_policy Note_Some) I Syntax.read_term NONE Binding.empty Binding.empty
  25.266 -    [];
  25.267 +val bnf_cmd = (fn (key, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, defs) =>
  25.268 +  let
  25.269 +    val wit_goals = map Logic.mk_conjunction_balanced wit_goalss;
  25.270 +    fun mk_triv_wit_thms tac set_maps =
  25.271 +      Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals)
  25.272 +        (K (TRYALL Goal.conjunction_tac) THEN' tac set_maps)
  25.273 +        |> Conjunction.elim_balanced (length wit_goals)
  25.274 +        |> map2 (Conjunction.elim_balanced o length) wit_goalss
  25.275 +        |> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));
  25.276 +    val (mk_wit_thms, nontriv_wit_goals) = 
  25.277 +      (case triv_tac_opt of
  25.278 +        NONE => (fn _ => [], map (map (rpair [])) wit_goalss)
  25.279 +      | SOME tac => (mk_triv_wit_thms tac, []));
  25.280 +  in
  25.281 +    Proof.unfolding ([[(defs, [])]])
  25.282 +      (Proof.theorem NONE (snd o register_bnf key oo after_qed mk_wit_thms)
  25.283 +        (map (single o rpair []) goals @ nontriv_wit_goals) lthy)
  25.284 +  end) oo prepare_def Do_Inline (user_policy Note_Some) I Syntax.read_term NONE
  25.285 +    Binding.empty Binding.empty [];
  25.286  
  25.287  fun print_bnfs ctxt =
  25.288    let
  25.289 @@ -1330,7 +1391,9 @@
  25.290      "register a type as a bounded natural functor"
  25.291      ((parse_opt_binding_colon -- Parse.term --
  25.292         (@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) -- Parse.term --
  25.293 -       (@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) -- Scan.option Parse.term)
  25.294 +       (Scan.option ((@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}))
  25.295 +         >> the_default []) --
  25.296 +       Scan.option Parse.term)
  25.297         >> bnf_cmd);
  25.298  
  25.299  end;
    26.1 --- a/src/HOL/BNF/Tools/bnf_def_tactics.ML	Mon Nov 11 17:34:44 2013 +0100
    26.2 +++ b/src/HOL/BNF/Tools/bnf_def_tactics.ML	Mon Nov 11 17:44:21 2013 +0100
    26.3 @@ -31,7 +31,10 @@
    26.4      {prems: thm list, context: Proof.context} -> tactic
    26.5  
    26.6    val mk_in_bd_tac: int -> thm -> thm -> thm -> thm -> thm list -> thm list -> thm -> thm -> thm ->
    26.7 -    thm -> {prems: 'a, context: Proof.context} -> tactic
    26.8 +    thm -> {prems: thm list, context: Proof.context} -> tactic
    26.9 +
   26.10 +  val mk_trivial_wit_tac: thm list -> thm list -> {prems: thm list, context: Proof.context} ->
   26.11 +    tactic
   26.12  end;
   26.13  
   26.14  structure BNF_Def_Tactics : BNF_DEF_TACTICS =
   26.15 @@ -302,4 +305,8 @@
   26.16             map_comp RS sym, map_id])] 1
   26.17    end;
   26.18  
   26.19 +fun mk_trivial_wit_tac wit_defs set_maps {context = ctxt, prems = _} =
   26.20 +  unfold_thms_tac ctxt wit_defs THEN HEADGOAL (EVERY' (map (fn thm =>
   26.21 +    dtac (thm RS equalityD1 RS set_mp) THEN' etac imageE THEN' atac) set_maps)) THEN ALLGOALS atac;
   26.22 +
   26.23  end;
    27.1 --- a/src/HOL/BNF/Tools/bnf_fp_def_sugar.ML	Mon Nov 11 17:34:44 2013 +0100
    27.2 +++ b/src/HOL/BNF/Tools/bnf_fp_def_sugar.ML	Mon Nov 11 17:44:21 2013 +0100
    27.3 @@ -25,7 +25,9 @@
    27.4       sel_co_iterssss: thm list list list list};
    27.5  
    27.6    val of_fp_sugar: (fp_sugar -> 'a list) -> fp_sugar -> 'a
    27.7 +  val eq_fp_sugar: fp_sugar * fp_sugar -> bool
    27.8    val morph_fp_sugar: morphism -> fp_sugar -> fp_sugar
    27.9 +  val transfer_fp_sugar: Proof.context -> fp_sugar -> fp_sugar
   27.10    val fp_sugar_of: Proof.context -> string -> fp_sugar option
   27.11    val fp_sugars_of: Proof.context -> fp_sugar list
   27.12  
   27.13 @@ -39,17 +41,14 @@
   27.14      'a list
   27.15    val mk_co_iter: theory -> BNF_FP_Util.fp_kind -> typ -> typ list -> term -> term
   27.16    val nesty_bnfs: Proof.context -> typ list list list -> typ list -> BNF_Def.bnf list
   27.17 -  val mk_map: int -> typ list -> typ list -> term -> term
   27.18 -  val mk_rel: int -> typ list -> typ list -> term -> term
   27.19 -  val build_map: local_theory -> (typ * typ -> term) -> typ * typ -> term
   27.20 -  val build_rel: local_theory -> (typ * typ -> term) -> typ * typ -> term
   27.21 -  val dest_map: Proof.context -> string -> term -> term * term list
   27.22 -  val dest_ctr: Proof.context -> string -> term -> term * term list
   27.23  
   27.24    type lfp_sugar_thms =
   27.25      (thm list * thm * Args.src list)
   27.26      * (thm list list * thm list list * Args.src list)
   27.27  
   27.28 +  val morph_lfp_sugar_thms: morphism -> lfp_sugar_thms -> lfp_sugar_thms
   27.29 +  val transfer_lfp_sugar_thms: Proof.context -> lfp_sugar_thms -> lfp_sugar_thms
   27.30 +
   27.31    type gfp_sugar_thms =
   27.32      ((thm list * thm) list * Args.src list)
   27.33      * (thm list list * thm list list * Args.src list)
   27.34 @@ -57,6 +56,9 @@
   27.35      * (thm list list * thm list list * Args.src list)
   27.36      * (thm list list list * thm list list list * Args.src list)
   27.37  
   27.38 +  val morph_gfp_sugar_thms: morphism -> gfp_sugar_thms -> gfp_sugar_thms
   27.39 +  val transfer_gfp_sugar_thms: Proof.context -> gfp_sugar_thms -> gfp_sugar_thms
   27.40 +
   27.41    val mk_co_iters_prelims: BNF_FP_Util.fp_kind -> typ list list list -> typ list -> typ list ->
   27.42      int list -> int list list -> term list list -> Proof.context ->
   27.43      (term list list
   27.44 @@ -87,8 +89,9 @@
   27.45      string * term list * term list list * ((term list list * term list list list)
   27.46        * (typ list * typ list list)) list ->
   27.47      thm -> thm list -> thm list -> thm list list -> BNF_Def.bnf list -> typ list -> typ list ->
   27.48 -    int list list -> int list list -> int list -> thm list list -> Ctr_Sugar.ctr_sugar list ->
   27.49 -    term list list -> thm list list -> (thm list -> thm list) -> local_theory -> gfp_sugar_thms
   27.50 +    typ list -> typ list list list -> int list list -> int list list -> int list -> thm list list ->
   27.51 +    Ctr_Sugar.ctr_sugar list -> term list list -> thm list list -> (thm list -> thm list) ->
   27.52 +    local_theory -> gfp_sugar_thms
   27.53    val co_datatypes: BNF_FP_Util.fp_kind -> (mixfix list -> binding list -> binding list ->
   27.54        binding list list -> binding list -> (string * sort) list -> typ list * typ list list ->
   27.55        BNF_Def.bnf list -> local_theory -> BNF_FP_Util.fp_result * local_theory) ->
   27.56 @@ -207,8 +210,8 @@
   27.57  val id_def = @{thm id_def};
   27.58  val mp_conj = @{thm mp_conj};
   27.59  
   27.60 -val nitpick_attrs = @{attributes [nitpick_simp]};
   27.61 -val code_nitpick_simp_attrs = Code.add_default_eqn_attrib :: nitpick_attrs;
   27.62 +val nitpicksimp_attrs = @{attributes [nitpick_simp]};
   27.63 +val code_nitpicksimp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs;
   27.64  val simp_attrs = @{attributes [simp]};
   27.65  
   27.66  fun tvar_subst thy Ts Us =
   27.67 @@ -232,7 +235,9 @@
   27.68    | flat_corec_preds_predsss_gettersss (p :: ps) (qss :: qsss) (fss :: fsss) =
   27.69      p :: flat_corec_predss_getterss qss fss @ flat_corec_preds_predsss_gettersss ps qsss fsss;
   27.70  
   27.71 -fun mk_tupled_fun x f xs = HOLogic.tupled_lambda x (Term.list_comb (f, xs));
   27.72 +fun mk_tupled_fun x f xs =
   27.73 +  if xs = [x] then f else HOLogic.tupled_lambda x (Term.list_comb (f, xs));
   27.74 +
   27.75  fun mk_uncurried2_fun f xss =
   27.76    mk_tupled_fun (HOLogic.mk_tuple (map HOLogic.mk_tuple xss)) f (flat_rec_arg_args xss);
   27.77  
   27.78 @@ -287,66 +292,6 @@
   27.79    | unzip_corecT _ (Type (@{type_name sum}, Ts)) = Ts
   27.80    | unzip_corecT _ T = [T];
   27.81  
   27.82 -fun mk_map live Ts Us t =
   27.83 -  let val (Type (_, Ts0), Type (_, Us0)) = strip_typeN (live + 1) (fastype_of t) |>> List.last in
   27.84 -    Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
   27.85 -  end;
   27.86 -
   27.87 -fun mk_rel live Ts Us t =
   27.88 -  let val [Type (_, Ts0), Type (_, Us0)] = binder_types (snd (strip_typeN live (fastype_of t))) in
   27.89 -    Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
   27.90 -  end;
   27.91 -
   27.92 -local
   27.93 -
   27.94 -fun build_map_or_rel mk const of_bnf dest lthy build_simple =
   27.95 -  let
   27.96 -    fun build (TU as (T, U)) =
   27.97 -      if T = U then
   27.98 -        const T
   27.99 -      else
  27.100 -        (case TU of
  27.101 -          (Type (s, Ts), Type (s', Us)) =>
  27.102 -          if s = s' then
  27.103 -            let
  27.104 -              val bnf = the (bnf_of lthy s);
  27.105 -              val live = live_of_bnf bnf;
  27.106 -              val mapx = mk live Ts Us (of_bnf bnf);
  27.107 -              val TUs' = map dest (fst (strip_typeN live (fastype_of mapx)));
  27.108 -            in Term.list_comb (mapx, map build TUs') end
  27.109 -          else
  27.110 -            build_simple TU
  27.111 -        | _ => build_simple TU);
  27.112 -  in build end;
  27.113 -
  27.114 -in
  27.115 -
  27.116 -val build_map = build_map_or_rel mk_map HOLogic.id_const map_of_bnf dest_funT;
  27.117 -val build_rel = build_map_or_rel mk_rel HOLogic.eq_const rel_of_bnf dest_pred2T;
  27.118 -
  27.119 -end;
  27.120 -
  27.121 -val dummy_var_name = "?f"
  27.122 -
  27.123 -fun mk_map_pattern ctxt s =
  27.124 -  let
  27.125 -    val bnf = the (bnf_of ctxt s);
  27.126 -    val mapx = map_of_bnf bnf;
  27.127 -    val live = live_of_bnf bnf;
  27.128 -    val (f_Ts, _) = strip_typeN live (fastype_of mapx);
  27.129 -    val fs = map_index (fn (i, T) => Var ((dummy_var_name, i), T)) f_Ts;
  27.130 -  in
  27.131 -    (mapx, betapplys (mapx, fs))
  27.132 -  end;
  27.133 -
  27.134 -fun dest_map ctxt s call =
  27.135 -  let
  27.136 -    val (map0, pat) = mk_map_pattern ctxt s;
  27.137 -    val (_, tenv) = fo_match ctxt call pat;
  27.138 -  in
  27.139 -    (map0, Vartab.fold_rev (fn (_, (_, f)) => cons f) tenv [])
  27.140 -  end;
  27.141 -
  27.142  fun liveness_of_fp_bnf n bnf =
  27.143    (case T_of_bnf bnf of
  27.144      Type (_, Ts) => map (not o member (op =) (deads_of_bnf bnf)) Ts
  27.145 @@ -388,12 +333,19 @@
  27.146  fun nesty_bnfs ctxt ctr_Tsss Us =
  27.147    map_filter (bnf_of ctxt) (fold (fold (fold (add_nesty_bnf_names Us))) ctr_Tsss []);
  27.148  
  27.149 -fun indexify proj xs f p = f (find_index (curry op = (proj p)) xs) p;
  27.150 +fun indexify proj xs f p = f (find_index (curry (op =) (proj p)) xs) p;
  27.151  
  27.152  type lfp_sugar_thms =
  27.153    (thm list * thm * Args.src list)
  27.154    * (thm list list * thm list list * Args.src list)
  27.155  
  27.156 +fun morph_lfp_sugar_thms phi ((inducts, induct, induct_attrs), (foldss, recss, iter_attrs)) =
  27.157 +  ((map (Morphism.thm phi) inducts, Morphism.thm phi induct, induct_attrs),
  27.158 +   (map (map (Morphism.thm phi)) foldss, map (map (Morphism.thm phi)) recss, iter_attrs));
  27.159 +
  27.160 +val transfer_lfp_sugar_thms =
  27.161 +  morph_lfp_sugar_thms o Morphism.thm_morphism o Thm.transfer o Proof_Context.theory_of;
  27.162 +
  27.163  type gfp_sugar_thms =
  27.164    ((thm list * thm) list * Args.src list)
  27.165    * (thm list list * thm list list * Args.src list)
  27.166 @@ -401,6 +353,23 @@
  27.167    * (thm list list * thm list list * Args.src list)
  27.168    * (thm list list list * thm list list list * Args.src list);
  27.169  
  27.170 +fun morph_gfp_sugar_thms phi ((coinducts_pairs, coinduct_attrs),
  27.171 +    (unfoldss, corecss, coiter_attrs), (disc_unfoldss, disc_corecss, disc_iter_attrs),
  27.172 +    (disc_unfold_iffss, disc_corec_iffss, disc_iter_iff_attrs),
  27.173 +    (sel_unfoldsss, sel_corecsss, sel_iter_attrs)) =
  27.174 +  ((map (apfst (map (Morphism.thm phi)) o apsnd (Morphism.thm phi)) coinducts_pairs,
  27.175 +    coinduct_attrs),
  27.176 +   (map (map (Morphism.thm phi)) unfoldss, map (map (Morphism.thm phi)) corecss, coiter_attrs),
  27.177 +   (map (map (Morphism.thm phi)) disc_unfoldss, map (map (Morphism.thm phi)) disc_corecss,
  27.178 +    disc_iter_attrs),
  27.179 +   (map (map (Morphism.thm phi)) disc_unfold_iffss, map (map (Morphism.thm phi)) disc_corec_iffss,
  27.180 +    disc_iter_iff_attrs),
  27.181 +   (map (map (map (Morphism.thm phi))) sel_unfoldsss,
  27.182 +    map (map (map (Morphism.thm phi))) sel_corecsss, sel_iter_attrs));
  27.183 +
  27.184 +val transfer_gfp_sugar_thms =
  27.185 +  morph_gfp_sugar_thms o Morphism.thm_morphism o Thm.transfer o Proof_Context.theory_of;
  27.186 +
  27.187  fun mk_iter_fun_arg_types0 n ms = map2 dest_tupleT ms o dest_sumTN_balanced n o domain_type;
  27.188  
  27.189  fun mk_iter_fun_arg_types ctr_Tsss ns mss =
  27.190 @@ -430,7 +399,7 @@
  27.191          ns mss ctr_Tsss ctor_iter_fun_Tss;
  27.192  
  27.193      val z_Tsss' = map (map flat_rec_arg_args) z_Tssss;
  27.194 -    val h_Tss = map2 (map2 (curry op --->)) z_Tsss' Css;
  27.195 +    val h_Tss = map2 (map2 (curry (op --->))) z_Tsss' Css;
  27.196  
  27.197      val hss = map2 (map2 retype_free) h_Tss gss;
  27.198      val zssss_hd = map2 (map2 (map2 (retype_free o hd))) z_Tssss ysss;
  27.199 @@ -452,7 +421,7 @@
  27.200      val f_sum_prod_Ts = map range_type fun_Ts;
  27.201      val f_prod_Tss = map2 dest_sumTN_balanced ns f_sum_prod_Ts;
  27.202      val f_Tsss = map2 (map2 (dest_tupleT o length)) ctr_Tsss' f_prod_Tss;
  27.203 -    val f_Tssss = map3 (fn C => map2 (map2 (map (curry op --> C) oo unzip_corecT)))
  27.204 +    val f_Tssss = map3 (fn C => map2 (map2 (map (curry (op -->) C) oo unzip_corecT)))
  27.205        Cs ctr_Tsss' f_Tsss;
  27.206      val q_Tssss = map (map (map (fn [_] => [] | [_, T] => [mk_pred1T (domain_type T)]))) f_Tssss;
  27.207    in
  27.208 @@ -577,7 +546,7 @@
  27.209  
  27.210      fun generate_iter suf (f_Tss, _, fss, xssss) ctor_iter =
  27.211        let
  27.212 -        val res_T = fold_rev (curry op --->) f_Tss fpT_to_C;
  27.213 +        val res_T = fold_rev (curry (op --->)) f_Tss fpT_to_C;
  27.214          val b = mk_binding suf;
  27.215          val spec =
  27.216            mk_Trueprop_eq (lists_bmoc fss (Free (Binding.name_of b, res_T)),
  27.217 @@ -596,7 +565,7 @@
  27.218  
  27.219      fun generate_coiter suf ((pfss, cqfsss), (f_sum_prod_Ts, pf_Tss)) dtor_coiter =
  27.220        let
  27.221 -        val res_T = fold_rev (curry op --->) pf_Tss C_to_fpT;
  27.222 +        val res_T = fold_rev (curry (op --->)) pf_Tss C_to_fpT;
  27.223          val b = mk_binding suf;
  27.224          val spec =
  27.225            mk_Trueprop_eq (lists_bmoc pfss (Free (Binding.name_of b, res_T)),
  27.226 @@ -645,7 +614,7 @@
  27.227          val lives = lives_of_bnf bnf;
  27.228          val sets = sets_of_bnf bnf;
  27.229          fun mk_set U =
  27.230 -          (case find_index (curry op = U) lives of
  27.231 +          (case find_index (curry (op =) U) lives of
  27.232              ~1 => Term.dummy
  27.233            | i => nth sets i);
  27.234        in
  27.235 @@ -662,7 +631,7 @@
  27.236            end;
  27.237  
  27.238          fun mk_raw_prem_prems _ (x as Free (_, Type _)) (X as TFree _) =
  27.239 -            [([], (find_index (curry op = X) Xs + 1, x))]
  27.240 +            [([], (find_index (curry (op =) X) Xs + 1, x))]
  27.241            | mk_raw_prem_prems names_lthy (x as Free (s, Type (T_name, Ts0))) (Type (_, Xs_Ts0)) =
  27.242              (case AList.lookup (op =) setss_nested T_name of
  27.243                NONE => []
  27.244 @@ -702,7 +671,7 @@
  27.245  
  27.246          val goal =
  27.247            Library.foldr (Logic.list_implies o apfst (map mk_prem)) (raw_premss,
  27.248 -            HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry op $) ps us)));
  27.249 +            HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry (op $)) ps us)));
  27.250  
  27.251          val kksss = map (map (map (fst o snd) o #2)) raw_premss;
  27.252  
  27.253 @@ -763,13 +732,13 @@
  27.254      val rec_thmss = mk_iter_thmss rec_args_types recs rec_defs (map co_rec_of ctor_iter_thmss);
  27.255    in
  27.256      ((induct_thms, induct_thm, [induct_case_names_attr]),
  27.257 -     (fold_thmss, rec_thmss, code_nitpick_simp_attrs @ simp_attrs))
  27.258 +     (fold_thmss, rec_thmss, code_nitpicksimp_attrs @ simp_attrs))
  27.259    end;
  27.260  
  27.261  fun derive_coinduct_coiters_thms_for_types pre_bnfs (z, cs, cpss,
  27.262        coiters_args_types as [((pgss, crgsss), _), ((phss, cshsss), _)])
  27.263 -    dtor_coinduct dtor_injects dtor_ctors dtor_coiter_thmss nesting_bnfs fpTs Cs kss mss ns
  27.264 -    ctr_defss (ctr_sugars : ctr_sugar list) coiterss coiter_defss export_args lthy =
  27.265 +    dtor_coinduct dtor_injects dtor_ctors dtor_coiter_thmss nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss
  27.266 +    mss ns ctr_defss (ctr_sugars : ctr_sugar list) coiterss coiter_defss export_args lthy =
  27.267    let
  27.268      fun mk_ctor_dtor_coiter_thm dtor_inject dtor_ctor coiter =
  27.269        iffD1 OF [dtor_inject, trans OF [coiter, dtor_ctor RS sym]];
  27.270 @@ -821,40 +790,29 @@
  27.271            map4 (fn u => fn v => fn uvr => fn uv_eq =>
  27.272              fold_rev Term.lambda [u, v] (HOLogic.mk_disj (uvr, uv_eq))) us vs uvrs uv_eqs;
  27.273  
  27.274 -        (* TODO: generalize (cf. "build_map") *)
  27.275 -        fun build_rel rs' T =
  27.276 -          (case find_index (curry op = T) fpTs of
  27.277 -            ~1 =>
  27.278 -            if exists_subtype_in fpTs T then
  27.279 -              let
  27.280 -                val Type (s, Ts) = T
  27.281 -                val bnf = the (bnf_of lthy s);
  27.282 -                val live = live_of_bnf bnf;
  27.283 -                val rel = mk_rel live Ts Ts (rel_of_bnf bnf);
  27.284 -                val Ts' = map domain_type (fst (strip_typeN live (fastype_of rel)));
  27.285 -              in Term.list_comb (rel, map (build_rel rs') Ts') end
  27.286 -            else
  27.287 -              HOLogic.eq_const T
  27.288 -          | kk => nth rs' kk);
  27.289 +        fun build_the_rel rs' T Xs_T =
  27.290 +          build_rel lthy (fn (_, X) => nth rs' (find_index (curry (op =) X) Xs)) (T, Xs_T)
  27.291 +          |> Term.subst_atomic_types (Xs ~~ fpTs);
  27.292  
  27.293 -        fun build_rel_app rs' usel vsel = fold rapp [usel, vsel] (build_rel rs' (fastype_of usel));
  27.294 +        fun build_rel_app rs' usel vsel Xs_T =
  27.295 +          fold rapp [usel, vsel] (build_the_rel rs' (fastype_of usel) Xs_T);
  27.296  
  27.297 -        fun mk_prem_ctr_concls rs' n k udisc usels vdisc vsels =
  27.298 +        fun mk_prem_ctr_concls rs' n k udisc usels vdisc vsels ctrXs_Ts =
  27.299            (if k = n then [] else [HOLogic.mk_eq (udisc, vdisc)]) @
  27.300            (if null usels then
  27.301               []
  27.302             else
  27.303               [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [udisc, vdisc],
  27.304 -                Library.foldr1 HOLogic.mk_conj (map2 (build_rel_app rs') usels vsels))]);
  27.305 +                Library.foldr1 HOLogic.mk_conj (map3 (build_rel_app rs') usels vsels ctrXs_Ts))]);
  27.306  
  27.307 -        fun mk_prem_concl rs' n udiscs uselss vdiscs vselss =
  27.308 -          Library.foldr1 HOLogic.mk_conj
  27.309 -            (flat (map5 (mk_prem_ctr_concls rs' n) (1 upto n) udiscs uselss vdiscs vselss))
  27.310 +        fun mk_prem_concl rs' n udiscs uselss vdiscs vselss ctrXs_Tss =
  27.311 +          Library.foldr1 HOLogic.mk_conj (flat (map6 (mk_prem_ctr_concls rs' n)
  27.312 +            (1 upto n) udiscs uselss vdiscs vselss ctrXs_Tss))
  27.313            handle List.Empty => @{term True};
  27.314  
  27.315 -        fun mk_prem rs' uvr u v n udiscs uselss vdiscs vselss =
  27.316 +        fun mk_prem rs' uvr u v n udiscs uselss vdiscs vselss ctrXs_Tss =
  27.317            fold_rev Logic.all [u, v] (Logic.mk_implies (HOLogic.mk_Trueprop uvr,
  27.318 -            HOLogic.mk_Trueprop (mk_prem_concl rs' n udiscs uselss vdiscs vselss)));
  27.319 +            HOLogic.mk_Trueprop (mk_prem_concl rs' n udiscs uselss vdiscs vselss ctrXs_Tss)));
  27.320  
  27.321          val concl =
  27.322            HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
  27.323 @@ -862,8 +820,8 @@
  27.324                 uvrs us vs));
  27.325  
  27.326          fun mk_goal rs' =
  27.327 -          Logic.list_implies (map8 (mk_prem rs') uvrs us vs ns udiscss uselsss vdiscss vselsss,
  27.328 -            concl);
  27.329 +          Logic.list_implies (map9 (mk_prem rs') uvrs us vs ns udiscss uselsss vdiscss vselsss
  27.330 +            ctrXs_Tsss, concl);
  27.331  
  27.332          val goals = map mk_goal [rs, strong_rs];
  27.333  
  27.334 @@ -1024,7 +982,7 @@
  27.335        coinduct_consumes_attr :: coinduct_case_names_attr :: coinduct_case_concl_attrs;
  27.336    in
  27.337      ((coinduct_thms_pairs, coinduct_case_attrs),
  27.338 -     (unfold_thmss, corec_thmss, code_nitpick_simp_attrs),
  27.339 +     (unfold_thmss, corec_thmss, code_nitpicksimp_attrs),
  27.340       (disc_unfold_thmss, disc_corec_thmss, []),
  27.341       (disc_unfold_iff_thmss, disc_corec_iff_thmss, simp_attrs),
  27.342       (sel_unfold_thmsss, sel_corec_thmsss, simp_attrs))
  27.343 @@ -1074,7 +1032,7 @@
  27.344  
  27.345      val qsoty = quote o Syntax.string_of_typ fake_lthy;
  27.346  
  27.347 -    val _ = (case duplicates (op =) unsorted_As of [] => ()
  27.348 +    val _ = (case Library.duplicates (op =) unsorted_As of [] => ()
  27.349        | A :: _ => error ("Duplicate type parameter " ^ qsoty A ^ " in " ^ co_prefix fp ^
  27.350            "datatype specification"));
  27.351  
  27.352 @@ -1087,7 +1045,7 @@
  27.353  
  27.354      val mixfixes = map mixfix_of specs;
  27.355  
  27.356 -    val _ = (case duplicates Binding.eq_name fp_bs of [] => ()
  27.357 +    val _ = (case Library.duplicates Binding.eq_name fp_bs of [] => ()
  27.358        | b :: _ => error ("Duplicate type name declaration " ^ quote (Binding.name_of b)));
  27.359  
  27.360      val ctr_specss = map ctr_specs_of specs;
  27.361 @@ -1380,15 +1338,22 @@
  27.362                val (rel_distinct_thms, _) =
  27.363                  join_halves n half_rel_distinct_thmss other_half_rel_distinct_thmss;
  27.364  
  27.365 +              val anonymous_notes =
  27.366 +                [(map (fn th => th RS @{thm eq_False[THEN iffD2]}) rel_distinct_thms,
  27.367 +                  code_nitpicksimp_attrs),
  27.368 +                 (map2 (fn th => fn 0 => th RS @{thm eq_True[THEN iffD2]} | _ => th)
  27.369 +                    rel_inject_thms ms, code_nitpicksimp_attrs)]
  27.370 +                |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
  27.371 +
  27.372                val notes =
  27.373 -                [(mapN, map_thms, code_nitpick_simp_attrs @ simp_attrs),
  27.374 -                 (rel_distinctN, rel_distinct_thms, code_nitpick_simp_attrs @ simp_attrs),
  27.375 -                 (rel_injectN, rel_inject_thms, code_nitpick_simp_attrs @ simp_attrs),
  27.376 -                 (setN, flat set_thmss, code_nitpick_simp_attrs @ simp_attrs)]
  27.377 +                [(mapN, map_thms, code_nitpicksimp_attrs @ simp_attrs),
  27.378 +                 (rel_distinctN, rel_distinct_thms, simp_attrs),
  27.379 +                 (rel_injectN, rel_inject_thms, simp_attrs),
  27.380 +                 (setN, flat set_thmss, code_nitpicksimp_attrs @ simp_attrs)]
  27.381                  |> massage_simple_notes fp_b_name;
  27.382              in
  27.383                (((map_thms, rel_inject_thms, rel_distinct_thms, set_thmss), ctr_sugar),
  27.384 -               lthy |> Local_Theory.notes notes |> snd)
  27.385 +               lthy |> Local_Theory.notes (anonymous_notes @ notes) |> snd)
  27.386              end;
  27.387  
  27.388          fun mk_binding suf = qualify false fp_b_name (Binding.suffix_name ("_" ^ suf) fp_b);
  27.389 @@ -1457,8 +1422,9 @@
  27.390               (disc_unfold_iff_thmss, disc_corec_iff_thmss, disc_coiter_iff_attrs),
  27.391               (sel_unfold_thmsss, sel_corec_thmsss, sel_coiter_attrs)) =
  27.392            derive_coinduct_coiters_thms_for_types pre_bnfs (the coiters_args_types) xtor_co_induct
  27.393 -            dtor_injects dtor_ctors xtor_co_iter_thmss nesting_bnfs fpTs Cs kss mss ns ctr_defss
  27.394 -            ctr_sugars coiterss coiter_defss (Proof_Context.export lthy' no_defs_lthy) lthy;
  27.395 +            dtor_injects dtor_ctors xtor_co_iter_thmss nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss ns
  27.396 +            ctr_defss ctr_sugars coiterss coiter_defss (Proof_Context.export lthy' no_defs_lthy)
  27.397 +            lthy;
  27.398  
  27.399          val sel_unfold_thmss = map flat sel_unfold_thmsss;
  27.400          val sel_corec_thmss = map flat sel_corec_thmsss;
  27.401 @@ -1496,6 +1462,12 @@
  27.402             (unfoldN, unfold_thmss, K coiter_attrs)]
  27.403            |> massage_multi_notes;
  27.404  
  27.405 +        fun is_codatatype (Type (s, _)) =
  27.406 +            (case fp_sugar_of lthy s of SOME {fp = Greatest_FP, ...} => true | _ => false)
  27.407 +          | is_codatatype _ = false;
  27.408 +
  27.409 +        val nitpick_supported = forall (is_codatatype o T_of_bnf) nested_bnfs;
  27.410 +
  27.411          fun register_nitpick fpT ({ctrs, casex, ...} : ctr_sugar) =
  27.412            Nitpick_HOL.register_codatatype fpT (fst (dest_Const casex))
  27.413              (map (dest_Const o mk_ctr As) ctrs)
  27.414 @@ -1507,7 +1479,7 @@
  27.415            ctr_sugars coiterss mapss [coinduct_thm, strong_coinduct_thm]
  27.416            (transpose [unfold_thmss, corec_thmss]) (transpose [disc_unfold_thmss, disc_corec_thmss])
  27.417            (transpose [sel_unfold_thmsss, sel_corec_thmsss])
  27.418 -        |> fold2 register_nitpick fpTs ctr_sugars
  27.419 +        |> nitpick_supported ? fold2 register_nitpick fpTs ctr_sugars
  27.420        end;
  27.421  
  27.422      val lthy'' = lthy'
    28.1 --- a/src/HOL/BNF/Tools/bnf_fp_def_sugar_tactics.ML	Mon Nov 11 17:34:44 2013 +0100
    28.2 +++ b/src/HOL/BNF/Tools/bnf_fp_def_sugar_tactics.ML	Mon Nov 11 17:44:21 2013 +0100
    28.3 @@ -151,12 +151,17 @@
    28.4    (atac ORELSE' REPEAT o etac conjE THEN'
    28.5       full_simp_tac
    28.6         (ss_only (@{thm prod.inject} :: no_refl discs @ rel_eqs @ more_simp_thms) ctxt) THEN'
    28.7 -     REPEAT o etac conjE THEN_MAYBE' REPEAT o hyp_subst_tac ctxt THEN' REPEAT o rtac conjI THEN'
    28.8 -     REPEAT o (rtac refl ORELSE' atac));
    28.9 +     REPEAT o etac conjE THEN_MAYBE' REPEAT o hyp_subst_tac ctxt THEN'
   28.10 +     REPEAT o (resolve_tac [refl, conjI] ORELSE' atac));
   28.11  
   28.12  fun mk_coinduct_distinct_ctrs_tac ctxt discs discs' =
   28.13 -  hyp_subst_tac ctxt THEN' REPEAT o etac conjE THEN'
   28.14 -  full_simp_tac (ss_only (refl :: no_refl (union Thm.eq_thm discs discs') @ basic_simp_thms) ctxt);
   28.15 +  let
   28.16 +    val discs'' = map (perhaps (try (fn th => th RS @{thm notnotD}))) (discs @ discs')
   28.17 +      |> distinct Thm.eq_thm_prop;
   28.18 +  in
   28.19 +    hyp_subst_tac ctxt THEN' REPEAT o etac conjE THEN'
   28.20 +    full_simp_tac (ss_only (refl :: no_refl discs'' @ basic_simp_thms) ctxt)
   28.21 +  end;
   28.22  
   28.23  fun mk_coinduct_discharge_prem_tac ctxt rel_eqs' nn kk n pre_rel_def dtor_ctor exhaust ctr_defs
   28.24      discss selss =
    29.1 --- a/src/HOL/BNF/Tools/bnf_fp_n2m.ML	Mon Nov 11 17:34:44 2013 +0100
    29.2 +++ b/src/HOL/BNF/Tools/bnf_fp_n2m.ML	Mon Nov 11 17:44:21 2013 +0100
    29.3 @@ -23,7 +23,7 @@
    29.4  open BNF_FP_N2M_Tactics
    29.5  
    29.6  fun force_typ ctxt T =
    29.7 -  map_types Type_Infer.paramify_vars 
    29.8 +  map_types Type_Infer.paramify_vars
    29.9    #> Type.constraint T
   29.10    #> Syntax.check_term ctxt
   29.11    #> singleton (Variable.polymorphic ctxt);
   29.12 @@ -99,10 +99,6 @@
   29.13      val fp_nesty_bnfss = fp_bnfs :: nesty_bnfss;
   29.14      val fp_nesty_bnfs = distinct eq_bnf (flat fp_nesty_bnfss);
   29.15  
   29.16 -    fun abstract t =
   29.17 -      let val Ts = Term.add_frees t [];
   29.18 -      in fold_rev Term.absfree (filter (member op = Ts) phis') t end;
   29.19 -
   29.20      val rels =
   29.21        let
   29.22          fun find_rel T As Bs = fp_nesty_bnfss
   29.23 @@ -121,10 +117,11 @@
   29.24                in
   29.25                  Term.list_comb (rel, rels)
   29.26                end
   29.27 -          | mk_rel (T as TFree _) _ = nth phis (find_index (curry op = T) As)
   29.28 +          | mk_rel (T as TFree _) _ = (nth phis (find_index (curry op = T) As)
   29.29 +              handle General.Subscript => HOLogic.eq_const T)
   29.30            | mk_rel _ _ = raise Fail "fpTs contains schematic type variables";
   29.31        in
   29.32 -        map2 (abstract oo mk_rel) fpTs fpTs'
   29.33 +        map2 (fold_rev Term.absfree phis' oo mk_rel) fpTs fpTs'
   29.34        end;
   29.35  
   29.36      val pre_rels = map2 (fn Ds => mk_rel_of_bnf Ds (As @ fpTs) (Bs @ fpTs')) Dss bnfs;
   29.37 @@ -224,7 +221,7 @@
   29.38          fun mk_s TU' =
   29.39            let
   29.40              val i = find_index (fn T => co_alg_argT TU' = T) Xs;
   29.41 -            val sF = co_alg_funT TU'; 
   29.42 +            val sF = co_alg_funT TU';
   29.43              val F = nth iter_preTs i;
   29.44              val s = nth iter_strs i;
   29.45            in
   29.46 @@ -238,7 +235,7 @@
   29.47                    |> force_typ names_lthy smapT
   29.48                    |> hidden_to_unit;
   29.49                  val smap_argTs = strip_typeN live (fastype_of smap) |> fst;
   29.50 -                fun mk_smap_arg TU =              
   29.51 +                fun mk_smap_arg TU =
   29.52                    (if domain_type TU = range_type TU then
   29.53                      HOLogic.id_const (domain_type TU)
   29.54                    else if is_rec then
   29.55 @@ -265,7 +262,7 @@
   29.56        in
   29.57          (case b_opt of
   29.58            NONE => ((t, Drule.dummy_thm), lthy)
   29.59 -        | SOME b => Local_Theory.define ((b, NoSyn), ((Thm.def_binding b, []), 
   29.60 +        | SOME b => Local_Theory.define ((b, NoSyn), ((Binding.conceal (Thm.def_binding b), []),
   29.61              fold_rev Term.absfree (if is_rec then rec_strs' else fold_strs') t)) lthy |>> apsnd snd)
   29.62        end;
   29.63  
   29.64 @@ -376,6 +373,6 @@
   29.65         |> morph_fp_result (Morphism.term_morphism (singleton (Variable.polymorphic lthy))));
   29.66    in
   29.67      (fp_res, lthy)
   29.68 -  end
   29.69 +  end;
   29.70  
   29.71  end;
    30.1 --- a/src/HOL/BNF/Tools/bnf_fp_n2m_sugar.ML	Mon Nov 11 17:34:44 2013 +0100
    30.2 +++ b/src/HOL/BNF/Tools/bnf_fp_n2m_sugar.ML	Mon Nov 11 17:44:21 2013 +0100
    30.3 @@ -7,14 +7,16 @@
    30.4  
    30.5  signature BNF_FP_N2M_SUGAR =
    30.6  sig
    30.7 -  val mutualize_fp_sugars: bool -> BNF_FP_Util.fp_kind -> binding list -> typ list ->
    30.8 -    (term -> int list) -> term list list list list -> BNF_FP_Def_Sugar.fp_sugar list ->
    30.9 -    local_theory ->
   30.10 +  val unfold_let: term -> term
   30.11 +  val dest_map: Proof.context -> string -> term -> term * term list
   30.12 +
   30.13 +  val mutualize_fp_sugars: BNF_FP_Util.fp_kind -> binding list -> typ list -> (term -> int list) ->
   30.14 +    term list list list list -> BNF_FP_Def_Sugar.fp_sugar list -> local_theory ->
   30.15      (BNF_FP_Def_Sugar.fp_sugar list
   30.16       * (BNF_FP_Def_Sugar.lfp_sugar_thms option * BNF_FP_Def_Sugar.gfp_sugar_thms option))
   30.17      * local_theory
   30.18 -  val pad_and_indexify_calls: BNF_FP_Def_Sugar.fp_sugar list -> int ->
   30.19 -    (term * term list list) list list -> term list list list list
   30.20 +  val indexify_callsss: BNF_FP_Def_Sugar.fp_sugar -> (term * term list list) list ->
   30.21 +    term list list list
   30.22    val nested_to_mutual_fps: BNF_FP_Util.fp_kind -> binding list -> typ list -> (term -> int list) ->
   30.23      (term * term list list) list list -> local_theory ->
   30.24      (typ list * int list * BNF_FP_Def_Sugar.fp_sugar list
   30.25 @@ -34,171 +36,246 @@
   30.26  
   30.27  val n2mN = "n2m_"
   30.28  
   30.29 +type n2m_sugar = fp_sugar list * (lfp_sugar_thms option * gfp_sugar_thms option);
   30.30 +
   30.31 +structure Data = Generic_Data
   30.32 +(
   30.33 +  type T = n2m_sugar Typtab.table;
   30.34 +  val empty = Typtab.empty;
   30.35 +  val extend = I;
   30.36 +  val merge = Typtab.merge (eq_fst (eq_list eq_fp_sugar));
   30.37 +);
   30.38 +
   30.39 +fun morph_n2m_sugar phi (fp_sugars, (lfp_sugar_thms_opt, gfp_sugar_thms_opt)) =
   30.40 +  (map (morph_fp_sugar phi) fp_sugars,
   30.41 +   (Option.map (morph_lfp_sugar_thms phi) lfp_sugar_thms_opt,
   30.42 +    Option.map (morph_gfp_sugar_thms phi) gfp_sugar_thms_opt));
   30.43 +
   30.44 +val transfer_n2m_sugar =
   30.45 +  morph_n2m_sugar o Morphism.thm_morphism o Thm.transfer o Proof_Context.theory_of;
   30.46 +
   30.47 +fun n2m_sugar_of ctxt =
   30.48 +  Typtab.lookup (Data.get (Context.Proof ctxt))
   30.49 +  #> Option.map (transfer_n2m_sugar ctxt);
   30.50 +
   30.51 +fun register_n2m_sugar key n2m_sugar =
   30.52 +  Local_Theory.declaration {syntax = false, pervasive = false}
   30.53 +    (fn phi => Data.map (Typtab.default (key, morph_n2m_sugar phi n2m_sugar)));
   30.54 +
   30.55 +fun unfold_let (Const (@{const_name Let}, _) $ arg1 $ arg2) = unfold_let (betapply (arg2, arg1))
   30.56 +  | unfold_let (Const (@{const_name prod_case}, _) $ t) =
   30.57 +    (case unfold_let t of
   30.58 +      t' as Abs (s1, T1, Abs (s2, T2, _)) =>
   30.59 +      let
   30.60 +        val x = (s1 ^ s2, Term.maxidx_of_term t + 1);
   30.61 +        val v = Var (x, HOLogic.mk_prodT (T1, T2));
   30.62 +      in
   30.63 +        lambda v (unfold_let (betapplys (t', [HOLogic.mk_fst v, HOLogic.mk_snd v])))
   30.64 +      end
   30.65 +    | _ => t)
   30.66 +  | unfold_let (t $ u) = betapply (unfold_let t, unfold_let u)
   30.67 +  | unfold_let (Abs (s, T, t)) = Abs (s, T, unfold_let t)
   30.68 +  | unfold_let t = t;
   30.69 +
   30.70 +fun mk_map_pattern ctxt s =
   30.71 +  let
   30.72 +    val bnf = the (bnf_of ctxt s);
   30.73 +    val mapx = map_of_bnf bnf;
   30.74 +    val live = live_of_bnf bnf;
   30.75 +    val (f_Ts, _) = strip_typeN live (fastype_of mapx);
   30.76 +    val fs = map_index (fn (i, T) => Var (("?f", i), T)) f_Ts;
   30.77 +  in
   30.78 +    (mapx, betapplys (mapx, fs))
   30.79 +  end;
   30.80 +
   30.81 +fun dest_map ctxt s call =
   30.82 +  let
   30.83 +    val (map0, pat) = mk_map_pattern ctxt s;
   30.84 +    val (_, tenv) = fo_match ctxt call pat;
   30.85 +  in
   30.86 +    (map0, Vartab.fold_rev (fn (_, (_, f)) => cons f) tenv [])
   30.87 +  end;
   30.88 +
   30.89 +fun dest_abs_or_applied_map _ _ (Abs (_, _, t)) = (Term.dummy, [t])
   30.90 +  | dest_abs_or_applied_map ctxt s (t1 $ _) = dest_map ctxt s t1;
   30.91 +
   30.92 +fun map_partition f xs =
   30.93 +  fold_rev (fn x => fn (ys, (good, bad)) =>
   30.94 +      case f x of SOME y => (y :: ys, (x :: good, bad)) | NONE => (ys, (good, x :: bad)))
   30.95 +    xs ([], ([], []));
   30.96 +
   30.97 +fun key_of_fp_eqs fp fpTs fp_eqs =
   30.98 +  Type (fp_case fp "l" "g", fpTs @ maps (fn (x, T) => [TFree x, T]) fp_eqs);
   30.99 +
  30.100  (* TODO: test with sort constraints on As *)
  30.101 -(* TODO: use right sorting order for "fp_sort" w.r.t. original BNFs (?) -- treat new variables
  30.102 -   as deads? *)
  30.103 -fun mutualize_fp_sugars mutualize fp bs fpTs get_indices callssss fp_sugars0 no_defs_lthy0 =
  30.104 -  if mutualize orelse has_duplicates (op =) fpTs then
  30.105 -    let
  30.106 -      val thy = Proof_Context.theory_of no_defs_lthy0;
  30.107 +fun mutualize_fp_sugars fp bs fpTs get_indices callssss fp_sugars0 no_defs_lthy0 =
  30.108 +  let
  30.109 +    val thy = Proof_Context.theory_of no_defs_lthy0;
  30.110  
  30.111 -      val qsotm = quote o Syntax.string_of_term no_defs_lthy0;
  30.112 +    val qsotm = quote o Syntax.string_of_term no_defs_lthy0;
  30.113  
  30.114 -      fun heterogeneous_call t = error ("Heterogeneous recursive call: " ^ qsotm t);
  30.115 -      fun incompatible_calls t1 t2 =
  30.116 -        error ("Incompatible recursive calls: " ^ qsotm t1 ^ " vs. " ^ qsotm t2);
  30.117 +    fun incompatible_calls t1 t2 =
  30.118 +      error ("Incompatible " ^ co_prefix fp ^ "recursive calls: " ^ qsotm t1 ^ " vs. " ^ qsotm t2);
  30.119  
  30.120 -      val b_names = map Binding.name_of bs;
  30.121 -      val fp_b_names = map base_name_of_typ fpTs;
  30.122 +    val b_names = map Binding.name_of bs;
  30.123 +    val fp_b_names = map base_name_of_typ fpTs;
  30.124  
  30.125 -      val nn = length fpTs;
  30.126 +    val nn = length fpTs;
  30.127  
  30.128 -      fun target_ctr_sugar_of_fp_sugar fpT ({T, index, ctr_sugars, ...} : fp_sugar) =
  30.129 -        let
  30.130 -          val rho = Vartab.fold (cons o apsnd snd) (Sign.typ_match thy (T, fpT) Vartab.empty) [];
  30.131 -          val phi = Morphism.term_morphism (Term.subst_TVars rho);
  30.132 -        in
  30.133 -          morph_ctr_sugar phi (nth ctr_sugars index)
  30.134 -        end;
  30.135 +    fun target_ctr_sugar_of_fp_sugar fpT ({T, index, ctr_sugars, ...} : fp_sugar) =
  30.136 +      let
  30.137 +        val rho = Vartab.fold (cons o apsnd snd) (Sign.typ_match thy (T, fpT) Vartab.empty) [];
  30.138 +        val phi = Morphism.term_morphism (Term.subst_TVars rho);
  30.139 +      in
  30.140 +        morph_ctr_sugar phi (nth ctr_sugars index)
  30.141 +      end;
  30.142  
  30.143 -      val ctr_defss = map (of_fp_sugar #ctr_defss) fp_sugars0;
  30.144 -      val mapss = map (of_fp_sugar #mapss) fp_sugars0;
  30.145 -      val ctr_sugars0 = map2 target_ctr_sugar_of_fp_sugar fpTs fp_sugars0;
  30.146 +    val ctr_defss = map (of_fp_sugar #ctr_defss) fp_sugars0;
  30.147 +    val mapss = map (of_fp_sugar #mapss) fp_sugars0;
  30.148 +    val ctr_sugars0 = map2 target_ctr_sugar_of_fp_sugar fpTs fp_sugars0;
  30.149  
  30.150 -      val ctrss = map #ctrs ctr_sugars0;
  30.151 -      val ctr_Tss = map (map fastype_of) ctrss;
  30.152 +    val ctrss = map #ctrs ctr_sugars0;
  30.153 +    val ctr_Tss = map (map fastype_of) ctrss;
  30.154  
  30.155 -      val As' = fold (fold Term.add_tfreesT) ctr_Tss [];
  30.156 -      val As = map TFree As';
  30.157 +    val As' = fold (fold Term.add_tfreesT) ctr_Tss [];
  30.158 +    val As = map TFree As';
  30.159  
  30.160 -      val ((Cs, Xs), no_defs_lthy) =
  30.161 -        no_defs_lthy0
  30.162 -        |> fold Variable.declare_typ As
  30.163 -        |> mk_TFrees nn
  30.164 -        ||>> variant_tfrees fp_b_names;
  30.165 +    val ((Cs, Xs), no_defs_lthy) =
  30.166 +      no_defs_lthy0
  30.167 +      |> fold Variable.declare_typ As
  30.168 +      |> mk_TFrees nn
  30.169 +      ||>> variant_tfrees fp_b_names;
  30.170  
  30.171 -      fun freeze_fp_default (T as Type (s, Ts)) =
  30.172 -          (case find_index (curry (op =) T) fpTs of
  30.173 -            ~1 => Type (s, map freeze_fp_default Ts)
  30.174 -          | kk => nth Xs kk)
  30.175 -        | freeze_fp_default T = T;
  30.176 +    fun check_call_dead live_call call =
  30.177 +      if null (get_indices call) then () else incompatible_calls live_call call;
  30.178  
  30.179 -      fun get_indices_checked call =
  30.180 -        (case get_indices call of
  30.181 -          _ :: _ :: _ => heterogeneous_call call
  30.182 -        | kks => kks);
  30.183 +    fun freeze_fpTs_simple (T as Type (s, Ts)) =
  30.184 +        (case find_index (curry (op =) T) fpTs of
  30.185 +          ~1 => Type (s, map freeze_fpTs_simple Ts)
  30.186 +        | kk => nth Xs kk)
  30.187 +      | freeze_fpTs_simple T = T;
  30.188  
  30.189 -      fun freeze_fp calls (T as Type (s, Ts)) =
  30.190 -          (case map_filter (try (snd o dest_map no_defs_lthy s)) calls of
  30.191 -            [] =>
  30.192 -            (case union (op = o pairself fst)
  30.193 -                (maps (fn call => map (rpair call) (get_indices_checked call)) calls) [] of
  30.194 -              [] => freeze_fp_default T
  30.195 -            | [(kk, _)] => nth Xs kk
  30.196 -            | (_, call1) :: (_, call2) :: _ => incompatible_calls call1 call2)
  30.197 -          | callss =>
  30.198 -            Type (s, map2 freeze_fp (flatten_type_args_of_bnf (the (bnf_of no_defs_lthy s)) []
  30.199 -              (transpose callss)) Ts))
  30.200 -        | freeze_fp _ T = T;
  30.201 +    fun freeze_fpTs_map (callss, (live_call :: _, dead_calls)) s Ts =
  30.202 +      (List.app (check_call_dead live_call) dead_calls;
  30.203 +       Type (s, map2 freeze_fpTs (flatten_type_args_of_bnf (the (bnf_of no_defs_lthy s)) []
  30.204 +         (transpose callss)) Ts))
  30.205 +    and freeze_fpTs calls (T as Type (s, Ts)) =
  30.206 +        (case map_partition (try (snd o dest_map no_defs_lthy s)) calls of
  30.207 +          ([], _) =>
  30.208 +          (case map_partition (try (snd o dest_abs_or_applied_map no_defs_lthy s)) calls of
  30.209 +            ([], _) => freeze_fpTs_simple T
  30.210 +          | callsp => freeze_fpTs_map callsp s Ts)
  30.211 +        | callsp => freeze_fpTs_map callsp s Ts)
  30.212 +      | freeze_fpTs _ T = T;
  30.213  
  30.214 -      val ctr_Tsss = map (map binder_types) ctr_Tss;
  30.215 -      val ctrXs_Tsss = map2 (map2 (map2 freeze_fp)) callssss ctr_Tsss;
  30.216 -      val ctrXs_sum_prod_Ts = map (mk_sumTN_balanced o map HOLogic.mk_tupleT) ctrXs_Tsss;
  30.217 -      val Ts = map (body_type o hd) ctr_Tss;
  30.218 +    val ctr_Tsss = map (map binder_types) ctr_Tss;
  30.219 +    val ctrXs_Tsss = map2 (map2 (map2 freeze_fpTs)) callssss ctr_Tsss;
  30.220 +    val ctrXs_sum_prod_Ts = map (mk_sumTN_balanced o map HOLogic.mk_tupleT) ctrXs_Tsss;
  30.221 +    val Ts = map (body_type o hd) ctr_Tss;
  30.222  
  30.223 -      val ns = map length ctr_Tsss;
  30.224 -      val kss = map (fn n => 1 upto n) ns;
  30.225 -      val mss = map (map length) ctr_Tsss;
  30.226 +    val ns = map length ctr_Tsss;
  30.227 +    val kss = map (fn n => 1 upto n) ns;
  30.228 +    val mss = map (map length) ctr_Tsss;
  30.229  
  30.230 -      val fp_eqs = map dest_TFree Xs ~~ ctrXs_sum_prod_Ts;
  30.231 +    val fp_eqs = map dest_TFree Xs ~~ ctrXs_sum_prod_Ts;
  30.232 +    val key = key_of_fp_eqs fp fpTs fp_eqs;
  30.233 +  in
  30.234 +    (case n2m_sugar_of no_defs_lthy key of
  30.235 +      SOME n2m_sugar => (n2m_sugar, no_defs_lthy)
  30.236 +    | NONE =>
  30.237 +      let
  30.238 +        val base_fp_names = Name.variant_list [] fp_b_names;
  30.239 +        val fp_bs = map2 (fn b_name => fn base_fp_name =>
  30.240 +            Binding.qualify true b_name (Binding.name (n2mN ^ base_fp_name)))
  30.241 +          b_names base_fp_names;
  30.242  
  30.243 -      val base_fp_names = Name.variant_list [] fp_b_names;
  30.244 -      val fp_bs = map2 (fn b_name => fn base_fp_name =>
  30.245 -          Binding.qualify true b_name (Binding.name (n2mN ^ base_fp_name)))
  30.246 -        b_names base_fp_names;
  30.247 +        val (pre_bnfs, (fp_res as {xtor_co_iterss = xtor_co_iterss0, xtor_co_induct, dtor_injects,
  30.248 +               dtor_ctors, xtor_co_iter_thmss, ...}, lthy)) =
  30.249 +          fp_bnf (construct_mutualized_fp fp fpTs fp_sugars0) fp_bs As' fp_eqs no_defs_lthy;
  30.250  
  30.251 -      val (pre_bnfs, (fp_res as {xtor_co_iterss = xtor_co_iterss0, xtor_co_induct,
  30.252 -             dtor_injects, dtor_ctors, xtor_co_iter_thmss, ...}, lthy)) =
  30.253 -        fp_bnf (construct_mutualized_fp fp fpTs fp_sugars0) fp_bs As' fp_eqs no_defs_lthy;
  30.254 +        val nesting_bnfs = nesty_bnfs lthy ctrXs_Tsss As;
  30.255 +        val nested_bnfs = nesty_bnfs lthy ctrXs_Tsss Xs;
  30.256  
  30.257 -      val nesting_bnfs = nesty_bnfs lthy ctrXs_Tsss As;
  30.258 -      val nested_bnfs = nesty_bnfs lthy ctrXs_Tsss Xs;
  30.259 +        val ((xtor_co_iterss, iters_args_types, coiters_args_types), _) =
  30.260 +          mk_co_iters_prelims fp ctr_Tsss fpTs Cs ns mss xtor_co_iterss0 lthy;
  30.261  
  30.262 -      val ((xtor_co_iterss, iters_args_types, coiters_args_types), _) =
  30.263 -        mk_co_iters_prelims fp ctr_Tsss fpTs Cs ns mss xtor_co_iterss0 lthy;
  30.264 +        fun mk_binding b suf = Binding.suffix_name ("_" ^ suf) b;
  30.265  
  30.266 -      fun mk_binding b suf = Binding.suffix_name ("_" ^ suf) b;
  30.267 +        val ((co_iterss, co_iter_defss), lthy) =
  30.268 +          fold_map2 (fn b =>
  30.269 +            (if fp = Least_FP then define_iters [foldN, recN] (the iters_args_types)
  30.270 +             else define_coiters [unfoldN, corecN] (the coiters_args_types))
  30.271 +              (mk_binding b) fpTs Cs) fp_bs xtor_co_iterss lthy
  30.272 +          |>> split_list;
  30.273  
  30.274 -      val ((co_iterss, co_iter_defss), lthy) =
  30.275 -        fold_map2 (fn b =>
  30.276 -          (if fp = Least_FP then define_iters [foldN, recN] (the iters_args_types)
  30.277 -           else define_coiters [unfoldN, corecN] (the coiters_args_types))
  30.278 -            (mk_binding b) fpTs Cs) fp_bs xtor_co_iterss lthy
  30.279 -        |>> split_list;
  30.280 +        val rho = tvar_subst thy Ts fpTs;
  30.281 +        val ctr_sugar_phi = Morphism.compose (Morphism.typ_morphism (Term.typ_subst_TVars rho))
  30.282 +            (Morphism.term_morphism (Term.subst_TVars rho));
  30.283 +        val inst_ctr_sugar = morph_ctr_sugar ctr_sugar_phi;
  30.284  
  30.285 -      val rho = tvar_subst thy Ts fpTs;
  30.286 -      val ctr_sugar_phi =
  30.287 -        Morphism.compose (Morphism.typ_morphism (Term.typ_subst_TVars rho))
  30.288 -          (Morphism.term_morphism (Term.subst_TVars rho));
  30.289 -      val inst_ctr_sugar = morph_ctr_sugar ctr_sugar_phi;
  30.290 +        val ctr_sugars = map inst_ctr_sugar ctr_sugars0;
  30.291  
  30.292 -      val ctr_sugars = map inst_ctr_sugar ctr_sugars0;
  30.293 +        val ((co_inducts, un_fold_thmss, co_rec_thmss, disc_unfold_thmss, disc_corec_thmss,
  30.294 +              sel_unfold_thmsss, sel_corec_thmsss), fp_sugar_thms) =
  30.295 +          if fp = Least_FP then
  30.296 +            derive_induct_iters_thms_for_types pre_bnfs (the iters_args_types) xtor_co_induct
  30.297 +              xtor_co_iter_thmss nesting_bnfs nested_bnfs fpTs Cs Xs ctrXs_Tsss ctrss ctr_defss
  30.298 +              co_iterss co_iter_defss lthy
  30.299 +            |> `(fn ((_, induct, _), (fold_thmss, rec_thmss, _)) =>
  30.300 +              ([induct], fold_thmss, rec_thmss, [], [], [], []))
  30.301 +            ||> (fn info => (SOME info, NONE))
  30.302 +          else
  30.303 +            derive_coinduct_coiters_thms_for_types pre_bnfs (the coiters_args_types) xtor_co_induct
  30.304 +              dtor_injects dtor_ctors xtor_co_iter_thmss nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss
  30.305 +              ns ctr_defss ctr_sugars co_iterss co_iter_defss
  30.306 +              (Proof_Context.export lthy no_defs_lthy) lthy
  30.307 +            |> `(fn ((coinduct_thms_pairs, _), (unfold_thmss, corec_thmss, _),
  30.308 +                    (disc_unfold_thmss, disc_corec_thmss, _), _,
  30.309 +                    (sel_unfold_thmsss, sel_corec_thmsss, _)) =>
  30.310 +              (map snd coinduct_thms_pairs, unfold_thmss, corec_thmss, disc_unfold_thmss,
  30.311 +               disc_corec_thmss, sel_unfold_thmsss, sel_corec_thmsss))
  30.312 +            ||> (fn info => (NONE, SOME info));
  30.313  
  30.314 -      val ((co_inducts, un_fold_thmss, co_rec_thmss, disc_unfold_thmss, disc_corec_thmss,
  30.315 -            sel_unfold_thmsss, sel_corec_thmsss), fp_sugar_thms) =
  30.316 -        if fp = Least_FP then
  30.317 -          derive_induct_iters_thms_for_types pre_bnfs (the iters_args_types) xtor_co_induct
  30.318 -            xtor_co_iter_thmss nesting_bnfs nested_bnfs fpTs Cs Xs ctrXs_Tsss ctrss ctr_defss
  30.319 -            co_iterss co_iter_defss lthy
  30.320 -          |> `(fn ((_, induct, _), (fold_thmss, rec_thmss, _)) =>
  30.321 -            ([induct], fold_thmss, rec_thmss, [], [], [], []))
  30.322 -          ||> (fn info => (SOME info, NONE))
  30.323 -        else
  30.324 -          derive_coinduct_coiters_thms_for_types pre_bnfs (the coiters_args_types) xtor_co_induct
  30.325 -            dtor_injects dtor_ctors xtor_co_iter_thmss nesting_bnfs fpTs Cs kss mss ns ctr_defss
  30.326 -            ctr_sugars co_iterss co_iter_defss (Proof_Context.export lthy no_defs_lthy) lthy
  30.327 -          |> `(fn ((coinduct_thms_pairs, _), (unfold_thmss, corec_thmss, _),
  30.328 -                  (disc_unfold_thmss, disc_corec_thmss, _), _,
  30.329 -                  (sel_unfold_thmsss, sel_corec_thmsss, _)) =>
  30.330 -            (map snd coinduct_thms_pairs, unfold_thmss, corec_thmss, disc_unfold_thmss,
  30.331 -             disc_corec_thmss, sel_unfold_thmsss, sel_corec_thmsss))
  30.332 -          ||> (fn info => (NONE, SOME info));
  30.333 +        val phi = Proof_Context.export_morphism no_defs_lthy no_defs_lthy0;
  30.334  
  30.335 -      val phi = Proof_Context.export_morphism no_defs_lthy no_defs_lthy0;
  30.336 +        fun mk_target_fp_sugar (kk, T) =
  30.337 +          {T = T, fp = fp, index = kk, pre_bnfs = pre_bnfs, nested_bnfs = nested_bnfs,
  30.338 +           nesting_bnfs = nesting_bnfs, fp_res = fp_res, ctr_defss = ctr_defss,
  30.339 +           ctr_sugars = ctr_sugars, co_iterss = co_iterss, mapss = mapss, co_inducts = co_inducts,
  30.340 +           co_iter_thmsss = transpose [un_fold_thmss, co_rec_thmss],
  30.341 +           disc_co_itersss = transpose [disc_unfold_thmss, disc_corec_thmss],
  30.342 +           sel_co_iterssss = transpose [sel_unfold_thmsss, sel_corec_thmsss]}
  30.343 +          |> morph_fp_sugar phi;
  30.344  
  30.345 -      fun mk_target_fp_sugar (kk, T) =
  30.346 -        {T = T, fp = fp, index = kk, pre_bnfs = pre_bnfs, nested_bnfs = nested_bnfs,
  30.347 -         nesting_bnfs = nesting_bnfs, fp_res = fp_res, ctr_defss = ctr_defss,
  30.348 -         ctr_sugars = ctr_sugars, co_iterss = co_iterss, mapss = mapss, co_inducts = co_inducts,
  30.349 -         co_iter_thmsss = transpose [un_fold_thmss, co_rec_thmss],
  30.350 -         disc_co_itersss = transpose [disc_unfold_thmss, disc_corec_thmss],
  30.351 -         sel_co_iterssss = transpose [sel_unfold_thmsss, sel_corec_thmsss]}
  30.352 -        |> morph_fp_sugar phi;
  30.353 -    in
  30.354 -      ((map_index mk_target_fp_sugar fpTs, fp_sugar_thms), lthy)
  30.355 -    end
  30.356 -  else
  30.357 -    (* TODO: reorder hypotheses and predicates in (co)induction rules? *)
  30.358 -    ((fp_sugars0, (NONE, NONE)), no_defs_lthy0);
  30.359 +        val n2m_sugar = (map_index mk_target_fp_sugar fpTs, fp_sugar_thms);
  30.360 +      in
  30.361 +        (n2m_sugar, lthy |> register_n2m_sugar key n2m_sugar)
  30.362 +      end)
  30.363 +  end;
  30.364  
  30.365  fun indexify_callsss fp_sugar callsss =
  30.366    let
  30.367      val {ctrs, ...} = of_fp_sugar #ctr_sugars fp_sugar;
  30.368 -    fun do_ctr ctr =
  30.369 +    fun indexify_ctr ctr =
  30.370        (case AList.lookup Term.aconv_untyped callsss ctr of
  30.371          NONE => replicate (num_binder_types (fastype_of ctr)) []
  30.372 -      | SOME callss => map (map Envir.beta_eta_contract) callss);
  30.373 +      | SOME callss => map (map (Envir.beta_eta_contract o unfold_let)) callss);
  30.374    in
  30.375 -    map do_ctr ctrs
  30.376 +    map indexify_ctr ctrs
  30.377    end;
  30.378  
  30.379 -fun pad_and_indexify_calls fp_sugars0 = map2 indexify_callsss fp_sugars0 oo pad_list [];
  30.380 +fun retypargs tyargs (Type (s, _)) = Type (s, tyargs);
  30.381 +
  30.382 +fun fold_subtype_pairs f (T as Type (s, Ts), U as Type (s', Us)) =
  30.383 +    f (T, U) #> (if s = s' then fold (fold_subtype_pairs f) (Ts ~~ Us) else I)
  30.384 +  | fold_subtype_pairs f TU = f TU;
  30.385  
  30.386  fun nested_to_mutual_fps fp actual_bs actual_Ts get_indices actual_callssss0 lthy =
  30.387    let
  30.388      val qsoty = quote o Syntax.string_of_typ lthy;
  30.389      val qsotys = space_implode " or " o map qsoty;
  30.390  
  30.391 +    fun duplicate_datatype T = error (qsoty T ^ " is not mutually recursive with itself");
  30.392      fun not_co_datatype0 T = error (qsoty T ^ " is not a " ^ co_prefix fp ^ "datatype");
  30.393      fun not_co_datatype (T as Type (s, _)) =
  30.394          if fp = Least_FP andalso
  30.395 @@ -208,32 +285,80 @@
  30.396            not_co_datatype0 T
  30.397        | not_co_datatype T = not_co_datatype0 T;
  30.398      fun not_mutually_nested_rec Ts1 Ts2 =
  30.399 -      error (qsotys Ts1 ^ " is neither mutually recursive with nor nested recursive via " ^
  30.400 -        qsotys Ts2);
  30.401 +      error (qsotys Ts1 ^ " is neither mutually recursive with " ^ qsotys Ts2 ^
  30.402 +        " nor nested recursive via " ^ qsotys Ts2);
  30.403  
  30.404 -    val perm_actual_Ts as Type (_, ty_args0) :: _ =
  30.405 -      sort (int_ord o pairself Term.size_of_typ) actual_Ts;
  30.406 +    val _ = (case Library.duplicates (op =) actual_Ts of [] => () | T :: _ => duplicate_datatype T);
  30.407  
  30.408 -    fun check_enrich_with_mutuals _ [] = []
  30.409 -      | check_enrich_with_mutuals seen ((T as Type (T_name, ty_args)) :: Ts) =
  30.410 -        (case fp_sugar_of lthy T_name of
  30.411 -          SOME ({fp = fp', fp_res = {Ts = Ts', ...}, ...}) =>
  30.412 -          if fp = fp' then
  30.413 +    val perm_actual_Ts =
  30.414 +      sort (prod_ord int_ord Term_Ord.typ_ord o pairself (`Term.size_of_typ)) actual_Ts;
  30.415 +
  30.416 +    fun the_ctrs_of (Type (s, Ts)) = map (mk_ctr Ts) (#ctrs (the (ctr_sugar_of lthy s)));
  30.417 +
  30.418 +    fun the_fp_sugar_of (T as Type (T_name, _)) =
  30.419 +      (case fp_sugar_of lthy T_name of
  30.420 +        SOME (fp_sugar as {fp = fp', ...}) => if fp = fp' then fp_sugar else not_co_datatype T
  30.421 +      | NONE => not_co_datatype T);
  30.422 +
  30.423 +    fun gen_rhss_in gen_Ts rho subTs =
  30.424 +      let
  30.425 +        fun maybe_insert (T, Type (_, gen_tyargs)) =
  30.426 +            if member (op =) subTs T then insert (op =) gen_tyargs else I
  30.427 +          | maybe_insert _ = I;
  30.428 +
  30.429 +        val ctrs = maps the_ctrs_of gen_Ts;
  30.430 +        val gen_ctr_Ts = maps (binder_types o fastype_of) ctrs;
  30.431 +        val ctr_Ts = map (Term.typ_subst_atomic rho) gen_ctr_Ts;
  30.432 +      in
  30.433 +        fold (fold_subtype_pairs maybe_insert) (ctr_Ts ~~ gen_ctr_Ts) []
  30.434 +      end;
  30.435 +
  30.436 +    fun gather_types _ _ num_groups seen gen_seen [] = (num_groups, seen, gen_seen)
  30.437 +      | gather_types lthy rho num_groups seen gen_seen ((T as Type (_, tyargs)) :: Ts) =
  30.438 +        let
  30.439 +          val {fp_res = {Ts = mutual_Ts0, ...}, ...} = the_fp_sugar_of T;
  30.440 +          val mutual_Ts = map (retypargs tyargs) mutual_Ts0;
  30.441 +
  30.442 +          val _ = seen = [] orelse exists (exists_subtype_in seen) mutual_Ts orelse
  30.443 +            not_mutually_nested_rec mutual_Ts seen;
  30.444 +
  30.445 +          fun fresh_tyargs () =
  30.446              let
  30.447 -              val mutual_Ts = map (fn Type (s, _) => Type (s, ty_args)) Ts';
  30.448 -              val _ =
  30.449 -                seen = [] orelse exists (exists_subtype_in seen) mutual_Ts orelse
  30.450 -                not_mutually_nested_rec mutual_Ts seen;
  30.451 -              val (seen', Ts') = List.partition (member (op =) mutual_Ts) Ts;
  30.452 +              (* The name "'z" is unlikely to clash with the context, yielding more cache hits. *)
  30.453 +              val (gen_tyargs, lthy') =
  30.454 +                variant_tfrees (replicate (length tyargs) "z") lthy
  30.455 +                |>> map Logic.varifyT_global;
  30.456 +              val rho' = (gen_tyargs ~~ tyargs) @ rho;
  30.457              in
  30.458 -              mutual_Ts @ check_enrich_with_mutuals (seen @ T :: seen') Ts'
  30.459 -            end
  30.460 -          else
  30.461 -            not_co_datatype T
  30.462 -        | NONE => not_co_datatype T)
  30.463 -      | check_enrich_with_mutuals _ (T :: _) = not_co_datatype T;
  30.464 +              (rho', gen_tyargs, gen_seen, lthy')
  30.465 +            end;
  30.466  
  30.467 -    val perm_Ts = check_enrich_with_mutuals [] perm_actual_Ts;
  30.468 +          val (rho', gen_tyargs, gen_seen', lthy') =
  30.469 +            if exists (exists_subtype_in seen) mutual_Ts then
  30.470 +              (case gen_rhss_in gen_seen rho mutual_Ts of
  30.471 +                [] => fresh_tyargs ()
  30.472 +              | gen_tyargss as gen_tyargs :: gen_tyargss_tl =>
  30.473 +                let
  30.474 +                  val unify_pairs = split_list (maps (curry (op ~~) gen_tyargs) gen_tyargss_tl);
  30.475 +                  val mgu = Type.raw_unifys unify_pairs Vartab.empty;
  30.476 +                  val gen_tyargs' = map (Envir.subst_type mgu) gen_tyargs;
  30.477 +                  val gen_seen' = map (Envir.subst_type mgu) gen_seen;
  30.478 +                in
  30.479 +                  (rho, gen_tyargs', gen_seen', lthy)
  30.480 +                end)
  30.481 +            else
  30.482 +              fresh_tyargs ();
  30.483 +
  30.484 +          val gen_mutual_Ts = map (retypargs gen_tyargs) mutual_Ts0;
  30.485 +          val Ts' = filter_out (member (op =) mutual_Ts) Ts;
  30.486 +        in
  30.487 +          gather_types lthy' rho' (num_groups + 1) (seen @ mutual_Ts) (gen_seen' @ gen_mutual_Ts)
  30.488 +            Ts'
  30.489 +        end
  30.490 +      | gather_types _ _ _ _ _ (T :: _) = not_co_datatype T;
  30.491 +
  30.492 +    val (num_groups, perm_Ts, perm_gen_Ts) = gather_types lthy [] 0 [] [] perm_actual_Ts;
  30.493 +    val perm_frozen_gen_Ts = map Logic.unvarifyT_global perm_gen_Ts;
  30.494  
  30.495      val missing_Ts = perm_Ts |> subtract (op =) actual_Ts;
  30.496      val Ts = actual_Ts @ missing_Ts;
  30.497 @@ -241,6 +366,8 @@
  30.498      val nn = length Ts;
  30.499      val kks = 0 upto nn - 1;
  30.500  
  30.501 +    val callssss0 = pad_list [] nn actual_callssss0;
  30.502 +
  30.503      val common_name = mk_common_name (map Binding.name_of actual_bs);
  30.504      val bs = pad_list (Binding.name common_name) nn actual_bs;
  30.505  
  30.506 @@ -249,16 +376,19 @@
  30.507  
  30.508      val perm_bs = permute bs;
  30.509      val perm_kks = permute kks;
  30.510 +    val perm_callssss0 = permute callssss0;
  30.511      val perm_fp_sugars0 = map (the o fp_sugar_of lthy o fst o dest_Type) perm_Ts;
  30.512  
  30.513 -    val mutualize = exists (fn Type (_, ty_args) => ty_args <> ty_args0) Ts;
  30.514 -    val perm_callssss = pad_and_indexify_calls perm_fp_sugars0 nn actual_callssss0;
  30.515 +    val perm_callssss = map2 indexify_callsss perm_fp_sugars0 perm_callssss0;
  30.516  
  30.517      val get_perm_indices = map (fn kk => find_index (curry (op =) kk) perm_kks) o get_indices;
  30.518  
  30.519      val ((perm_fp_sugars, fp_sugar_thms), lthy) =
  30.520 -      mutualize_fp_sugars mutualize fp perm_bs perm_Ts get_perm_indices perm_callssss
  30.521 -        perm_fp_sugars0 lthy;
  30.522 +      if num_groups > 1 then
  30.523 +        mutualize_fp_sugars fp perm_bs perm_frozen_gen_Ts get_perm_indices perm_callssss
  30.524 +          perm_fp_sugars0 lthy
  30.525 +      else
  30.526 +        ((perm_fp_sugars0, (NONE, NONE)), lthy);
  30.527  
  30.528      val fp_sugars = unpermute perm_fp_sugars;
  30.529    in
    31.1 --- a/src/HOL/BNF/Tools/bnf_fp_rec_sugar.ML	Mon Nov 11 17:34:44 2013 +0100
    31.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    31.3 @@ -1,986 +0,0 @@
    31.4 -(*  Title:      HOL/BNF/Tools/bnf_fp_rec_sugar.ML
    31.5 -    Author:     Lorenz Panny, TU Muenchen
    31.6 -    Copyright   2013
    31.7 -
    31.8 -Recursor and corecursor sugar.
    31.9 -*)
   31.10 -
   31.11 -signature BNF_FP_REC_SUGAR =
   31.12 -sig
   31.13 -  val add_primrec: (binding * typ option * mixfix) list ->
   31.14 -    (Attrib.binding * term) list -> local_theory -> (term list * thm list list) * local_theory
   31.15 -  val add_primrec_cmd: (binding * string option * mixfix) list ->
   31.16 -    (Attrib.binding * string) list -> local_theory -> (term list * thm list list) * local_theory
   31.17 -  val add_primrec_global: (binding * typ option * mixfix) list ->
   31.18 -    (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
   31.19 -  val add_primrec_overloaded: (string * (string * typ) * bool) list ->
   31.20 -    (binding * typ option * mixfix) list ->
   31.21 -    (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
   31.22 -  val add_primrec_simple: ((binding * typ) * mixfix) list -> term list ->
   31.23 -    local_theory -> (string list * (term list * (int list list * thm list list))) * local_theory
   31.24 -  val add_primcorecursive_cmd: bool ->
   31.25 -    (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
   31.26 -    Proof.context -> Proof.state
   31.27 -  val add_primcorec_cmd: bool ->
   31.28 -    (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
   31.29 -    local_theory -> local_theory
   31.30 -end;
   31.31 -
   31.32 -structure BNF_FP_Rec_Sugar : BNF_FP_REC_SUGAR =
   31.33 -struct
   31.34 -
   31.35 -open BNF_Util
   31.36 -open BNF_FP_Util
   31.37 -open BNF_FP_Rec_Sugar_Util
   31.38 -open BNF_FP_Rec_Sugar_Tactics
   31.39 -
   31.40 -val codeN = "code"
   31.41 -val ctrN = "ctr"
   31.42 -val discN = "disc"
   31.43 -val selN = "sel"
   31.44 -
   31.45 -val nitpick_attrs = @{attributes [nitpick_simp]};
   31.46 -val simp_attrs = @{attributes [simp]};
   31.47 -val code_nitpick_attrs = Code.add_default_eqn_attrib :: nitpick_attrs;
   31.48 -val code_nitpick_simp_attrs = Code.add_default_eqn_attrib :: nitpick_attrs @ simp_attrs;
   31.49 -
   31.50 -exception Primrec_Error of string * term list;
   31.51 -
   31.52 -fun primrec_error str = raise Primrec_Error (str, []);
   31.53 -fun primrec_error_eqn str eqn = raise Primrec_Error (str, [eqn]);
   31.54 -fun primrec_error_eqns str eqns = raise Primrec_Error (str, eqns);
   31.55 -
   31.56 -fun finds eq = fold_map (fn x => List.partition (curry eq x) #>> pair x);
   31.57 -
   31.58 -val free_name = try (fn Free (v, _) => v);
   31.59 -val const_name = try (fn Const (v, _) => v);
   31.60 -val undef_const = Const (@{const_name undefined}, dummyT);
   31.61 -
   31.62 -fun permute_args n t = list_comb (t, map Bound (0 :: (n downto 1)))
   31.63 -  |> fold (K (Term.abs (Name.uu, dummyT))) (0 upto n);
   31.64 -val abs_tuple = HOLogic.tupled_lambda o HOLogic.mk_tuple;
   31.65 -fun drop_All t = subst_bounds (strip_qnt_vars @{const_name all} t |> map Free |> rev,
   31.66 -  strip_qnt_body @{const_name all} t)
   31.67 -fun abstract vs =
   31.68 -  let fun a n (t $ u) = a n t $ a n u
   31.69 -        | a n (Abs (v, T, b)) = Abs (v, T, a (n + 1) b)
   31.70 -        | a n t = let val idx = find_index (equal t) vs in
   31.71 -            if idx < 0 then t else Bound (n + idx) end
   31.72 -  in a 0 end;
   31.73 -fun mk_prod1 Ts (t, u) = HOLogic.pair_const (fastype_of1 (Ts, t)) (fastype_of1 (Ts, u)) $ t $ u;
   31.74 -fun mk_tuple1 Ts = the_default HOLogic.unit o try (foldr1 (mk_prod1 Ts));
   31.75 -
   31.76 -fun get_indices fixes t = map (fst #>> Binding.name_of #> Free) fixes
   31.77 -  |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE)
   31.78 -  |> map_filter I;
   31.79 -
   31.80 -
   31.81 -(* Primrec *)
   31.82 -
   31.83 -type eqn_data = {
   31.84 -  fun_name: string,
   31.85 -  rec_type: typ,
   31.86 -  ctr: term,
   31.87 -  ctr_args: term list,
   31.88 -  left_args: term list,
   31.89 -  right_args: term list,
   31.90 -  res_type: typ,
   31.91 -  rhs_term: term,
   31.92 -  user_eqn: term
   31.93 -};
   31.94 -
   31.95 -fun dissect_eqn lthy fun_names eqn' =
   31.96 -  let
   31.97 -    val eqn = drop_All eqn' |> HOLogic.dest_Trueprop
   31.98 -      handle TERM _ =>
   31.99 -        primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
  31.100 -    val (lhs, rhs) = HOLogic.dest_eq eqn
  31.101 -        handle TERM _ =>
  31.102 -          primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
  31.103 -    val (fun_name, args) = strip_comb lhs
  31.104 -      |>> (fn x => if is_Free x then fst (dest_Free x)
  31.105 -          else primrec_error_eqn "malformed function equation (does not start with free)" eqn);
  31.106 -    val (left_args, rest) = take_prefix is_Free args;
  31.107 -    val (nonfrees, right_args) = take_suffix is_Free rest;
  31.108 -    val num_nonfrees = length nonfrees;
  31.109 -    val _ = num_nonfrees = 1 orelse if num_nonfrees = 0 then
  31.110 -      primrec_error_eqn "constructor pattern missing in left-hand side" eqn else
  31.111 -      primrec_error_eqn "more than one non-variable argument in left-hand side" eqn;
  31.112 -    val _ = member (op =) fun_names fun_name orelse
  31.113 -      primrec_error_eqn "malformed function equation (does not start with function name)" eqn
  31.114 -
  31.115 -    val (ctr, ctr_args) = strip_comb (the_single nonfrees);
  31.116 -    val _ = try (num_binder_types o fastype_of) ctr = SOME (length ctr_args) orelse
  31.117 -      primrec_error_eqn "partially applied constructor in pattern" eqn;
  31.118 -    val _ = let val d = duplicates (op =) (left_args @ ctr_args @ right_args) in null d orelse
  31.119 -      primrec_error_eqn ("duplicate variable \"" ^ Syntax.string_of_term lthy (hd d) ^
  31.120 -        "\" in left-hand side") eqn end;
  31.121 -    val _ = forall is_Free ctr_args orelse
  31.122 -      primrec_error_eqn "non-primitive pattern in left-hand side" eqn;
  31.123 -    val _ =
  31.124 -      let val b = fold_aterms (fn x as Free (v, _) =>
  31.125 -        if (not (member (op =) (left_args @ ctr_args @ right_args) x) andalso
  31.126 -        not (member (op =) fun_names v) andalso
  31.127 -        not (Variable.is_fixed lthy v)) then cons x else I | _ => I) rhs []
  31.128 -      in
  31.129 -        null b orelse
  31.130 -        primrec_error_eqn ("extra variable(s) in right-hand side: " ^
  31.131 -          commas (map (Syntax.string_of_term lthy) b)) eqn
  31.132 -      end;
  31.133 -  in
  31.134 -    {fun_name = fun_name,
  31.135 -     rec_type = body_type (type_of ctr),
  31.136 -     ctr = ctr,
  31.137 -     ctr_args = ctr_args,
  31.138 -     left_args = left_args,
  31.139 -     right_args = right_args,
  31.140 -     res_type = map fastype_of (left_args @ right_args) ---> fastype_of rhs,
  31.141 -     rhs_term = rhs,
  31.142 -     user_eqn = eqn'}
  31.143 -  end;
  31.144 -
  31.145 -fun rewrite_map_arg get_ctr_pos rec_type res_type =
  31.146 -  let
  31.147 -    val pT = HOLogic.mk_prodT (rec_type, res_type);
  31.148 -
  31.149 -    val maybe_suc = Option.map (fn x => x + 1);
  31.150 -    fun subst d (t as Bound d') = t |> d = SOME d' ? curry (op $) (fst_const pT)
  31.151 -      | subst d (Abs (v, T, b)) = Abs (v, if d = SOME ~1 then pT else T, subst (maybe_suc d) b)
  31.152 -      | subst d t =
  31.153 -        let
  31.154 -          val (u, vs) = strip_comb t;
  31.155 -          val ctr_pos = try (get_ctr_pos o the) (free_name u) |> the_default ~1;
  31.156 -        in
  31.157 -          if ctr_pos >= 0 then
  31.158 -            if d = SOME ~1 andalso length vs = ctr_pos then
  31.159 -              list_comb (permute_args ctr_pos (snd_const pT), vs)
  31.160 -            else if length vs > ctr_pos andalso is_some d
  31.161 -                andalso d = try (fn Bound n => n) (nth vs ctr_pos) then
  31.162 -              list_comb (snd_const pT $ nth vs ctr_pos, map (subst d) (nth_drop ctr_pos vs))
  31.163 -            else
  31.164 -              primrec_error_eqn ("recursive call not directly applied to constructor argument") t
  31.165 -          else if d = SOME ~1 andalso const_name u = SOME @{const_name comp} then
  31.166 -            list_comb (map_types (K dummyT) u, map2 subst [NONE, d] vs)
  31.167 -          else
  31.168 -            list_comb (u, map (subst (d |> d = SOME ~1 ? K NONE)) vs)
  31.169 -        end
  31.170 -  in
  31.171 -    subst (SOME ~1)
  31.172 -  end;
  31.173 -
  31.174 -fun subst_rec_calls lthy get_ctr_pos has_call ctr_args direct_calls indirect_calls t =
  31.175 -  let
  31.176 -    fun subst bound_Ts (Abs (v, T, b)) = Abs (v, T, subst (T :: bound_Ts) b)
  31.177 -      | subst bound_Ts (t as g' $ y) =
  31.178 -        let
  31.179 -          val maybe_direct_y' = AList.lookup (op =) direct_calls y;
  31.180 -          val maybe_indirect_y' = AList.lookup (op =) indirect_calls y;
  31.181 -          val (g, g_args) = strip_comb g';
  31.182 -          val ctr_pos = try (get_ctr_pos o the) (free_name g) |> the_default ~1;
  31.183 -          val _ = ctr_pos < 0 orelse length g_args >= ctr_pos orelse
  31.184 -            primrec_error_eqn "too few arguments in recursive call" t;
  31.185 -        in
  31.186 -          if not (member (op =) ctr_args y) then
  31.187 -            pairself (subst bound_Ts) (g', y) |> (op $)
  31.188 -          else if ctr_pos >= 0 then
  31.189 -            list_comb (the maybe_direct_y', g_args)
  31.190 -          else if is_some maybe_indirect_y' then
  31.191 -            (if has_call g' then t else y)
  31.192 -            |> massage_indirect_rec_call lthy has_call
  31.193 -              (rewrite_map_arg get_ctr_pos) bound_Ts y (the maybe_indirect_y')
  31.194 -            |> (if has_call g' then I else curry (op $) g')
  31.195 -          else
  31.196 -            t
  31.197 -        end
  31.198 -      | subst _ t = t
  31.199 -  in
  31.200 -    subst [] t
  31.201 -    |> tap (fn u => has_call u andalso (* FIXME detect this case earlier *)
  31.202 -      primrec_error_eqn "recursive call not directly applied to constructor argument" t)
  31.203 -  end;
  31.204 -
  31.205 -fun build_rec_arg lthy (funs_data : eqn_data list list) has_call (ctr_spec : rec_ctr_spec)
  31.206 -    (maybe_eqn_data : eqn_data option) =
  31.207 -  if is_none maybe_eqn_data then undef_const else
  31.208 -    let
  31.209 -      val eqn_data = the maybe_eqn_data;
  31.210 -      val t = #rhs_term eqn_data;
  31.211 -      val ctr_args = #ctr_args eqn_data;
  31.212 -
  31.213 -      val calls = #calls ctr_spec;
  31.214 -      val n_args = fold (curry (op +) o (fn Direct_Rec _ => 2 | _ => 1)) calls 0;
  31.215 -
  31.216 -      val no_calls' = tag_list 0 calls
  31.217 -        |> map_filter (try (apsnd (fn No_Rec n => n | Direct_Rec (n, _) => n)));
  31.218 -      val direct_calls' = tag_list 0 calls
  31.219 -        |> map_filter (try (apsnd (fn Direct_Rec (_, n) => n)));
  31.220 -      val indirect_calls' = tag_list 0 calls
  31.221 -        |> map_filter (try (apsnd (fn Indirect_Rec n => n)));
  31.222 -
  31.223 -      fun make_direct_type _ = dummyT; (* FIXME? *)
  31.224 -
  31.225 -      val rec_res_type_list = map (fn (x :: _) => (#rec_type x, #res_type x)) funs_data;
  31.226 -
  31.227 -      fun make_indirect_type (Type (Tname, Ts)) = Type (Tname, Ts |> map (fn T =>
  31.228 -        let val maybe_res_type = AList.lookup (op =) rec_res_type_list T in
  31.229 -          if is_some maybe_res_type
  31.230 -          then HOLogic.mk_prodT (T, the maybe_res_type)
  31.231 -          else make_indirect_type T end))
  31.232 -        | make_indirect_type T = T;
  31.233 -
  31.234 -      val args = replicate n_args ("", dummyT)
  31.235 -        |> Term.rename_wrt_term t
  31.236 -        |> map Free
  31.237 -        |> fold (fn (ctr_arg_idx, arg_idx) =>
  31.238 -            nth_map arg_idx (K (nth ctr_args ctr_arg_idx)))
  31.239 -          no_calls'
  31.240 -        |> fold (fn (ctr_arg_idx, arg_idx) =>
  31.241 -            nth_map arg_idx (K (nth ctr_args ctr_arg_idx |> map_types make_direct_type)))
  31.242 -          direct_calls'
  31.243 -        |> fold (fn (ctr_arg_idx, arg_idx) =>
  31.244 -            nth_map arg_idx (K (nth ctr_args ctr_arg_idx |> map_types make_indirect_type)))
  31.245 -          indirect_calls';
  31.246 -
  31.247 -      val fun_name_ctr_pos_list =
  31.248 -        map (fn (x :: _) => (#fun_name x, length (#left_args x))) funs_data;
  31.249 -      val get_ctr_pos = try (the o AList.lookup (op =) fun_name_ctr_pos_list) #> the_default ~1;
  31.250 -      val direct_calls = map (apfst (nth ctr_args) o apsnd (nth args)) direct_calls';
  31.251 -      val indirect_calls = map (apfst (nth ctr_args) o apsnd (nth args)) indirect_calls';
  31.252 -
  31.253 -      val abstractions = args @ #left_args eqn_data @ #right_args eqn_data;
  31.254 -    in
  31.255 -      t
  31.256 -      |> subst_rec_calls lthy get_ctr_pos has_call ctr_args direct_calls indirect_calls
  31.257 -      |> fold_rev lambda abstractions
  31.258 -    end;
  31.259 -
  31.260 -fun build_defs lthy bs mxs (funs_data : eqn_data list list) (rec_specs : rec_spec list) has_call =
  31.261 -  let
  31.262 -    val n_funs = length funs_data;
  31.263 -
  31.264 -    val ctr_spec_eqn_data_list' =
  31.265 -      (take n_funs rec_specs |> map #ctr_specs) ~~ funs_data
  31.266 -      |> maps (uncurry (finds (fn (x, y) => #ctr x = #ctr y))
  31.267 -          ##> (fn x => null x orelse
  31.268 -            primrec_error_eqns "excess equations in definition" (map #rhs_term x)) #> fst);
  31.269 -    val _ = ctr_spec_eqn_data_list' |> map (fn (_, x) => length x <= 1 orelse
  31.270 -      primrec_error_eqns ("multiple equations for constructor") (map #user_eqn x));
  31.271 -
  31.272 -    val ctr_spec_eqn_data_list =
  31.273 -      ctr_spec_eqn_data_list' @ (drop n_funs rec_specs |> maps #ctr_specs |> map (rpair []));
  31.274 -
  31.275 -    val recs = take n_funs rec_specs |> map #recx;
  31.276 -    val rec_args = ctr_spec_eqn_data_list
  31.277 -      |> sort ((op <) o pairself (#offset o fst) |> make_ord)
  31.278 -      |> map (uncurry (build_rec_arg lthy funs_data has_call) o apsnd (try the_single));
  31.279 -    val ctr_poss = map (fn x =>
  31.280 -      if length (distinct ((op =) o pairself (length o #left_args)) x) <> 1 then
  31.281 -        primrec_error ("inconstant constructor pattern position for function " ^
  31.282 -          quote (#fun_name (hd x)))
  31.283 -      else
  31.284 -        hd x |> #left_args |> length) funs_data;
  31.285 -  in
  31.286 -    (recs, ctr_poss)
  31.287 -    |-> map2 (fn recx => fn ctr_pos => list_comb (recx, rec_args) |> permute_args ctr_pos)
  31.288 -    |> Syntax.check_terms lthy
  31.289 -    |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.map_name Thm.def_name b, []), t))) bs mxs
  31.290 -  end;
  31.291 -
  31.292 -fun find_rec_calls has_call (eqn_data : eqn_data) =
  31.293 -  let
  31.294 -    fun find (Abs (_, _, b)) ctr_arg = find b ctr_arg
  31.295 -      | find (t as _ $ _) ctr_arg =
  31.296 -        let
  31.297 -          val (f', args') = strip_comb t;
  31.298 -          val n = find_index (equal ctr_arg) args';
  31.299 -        in
  31.300 -          if n < 0 then
  31.301 -            find f' ctr_arg @ maps (fn x => find x ctr_arg) args'
  31.302 -          else
  31.303 -            let val (f, args) = chop n args' |>> curry list_comb f' in
  31.304 -              if has_call f then
  31.305 -                f :: maps (fn x => find x ctr_arg) args
  31.306 -              else
  31.307 -                find f ctr_arg @ maps (fn x => find x ctr_arg) args
  31.308 -            end
  31.309 -        end
  31.310 -      | find _ _ = [];
  31.311 -  in
  31.312 -    map (find (#rhs_term eqn_data)) (#ctr_args eqn_data)
  31.313 -    |> (fn [] => NONE | callss => SOME (#ctr eqn_data, callss))
  31.314 -  end;
  31.315 -
  31.316 -fun prepare_primrec fixes specs lthy =
  31.317 -  let
  31.318 -    val (bs, mxs) = map_split (apfst fst) fixes;
  31.319 -    val fun_names = map Binding.name_of bs;
  31.320 -    val eqns_data = map (dissect_eqn lthy fun_names) specs;
  31.321 -    val funs_data = eqns_data
  31.322 -      |> partition_eq ((op =) o pairself #fun_name)
  31.323 -      |> finds (fn (x, y) => x = #fun_name (hd y)) fun_names |> fst
  31.324 -      |> map (fn (x, y) => the_single y handle List.Empty =>
  31.325 -          primrec_error ("missing equations for function " ^ quote x));
  31.326 -
  31.327 -    val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
  31.328 -    val arg_Ts = map (#rec_type o hd) funs_data;
  31.329 -    val res_Ts = map (#res_type o hd) funs_data;
  31.330 -    val callssss = funs_data
  31.331 -      |> map (partition_eq ((op =) o pairself #ctr))
  31.332 -      |> map (maps (map_filter (find_rec_calls has_call)));
  31.333 -
  31.334 -    val ((n2m, rec_specs, _, induct_thm, induct_thms), lthy') =
  31.335 -      rec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
  31.336 -
  31.337 -    val actual_nn = length funs_data;
  31.338 -
  31.339 -    val _ = let val ctrs = (maps (map #ctr o #ctr_specs) rec_specs) in
  31.340 -      map (fn {ctr, user_eqn, ...} => member (op =) ctrs ctr orelse
  31.341 -        primrec_error_eqn ("argument " ^ quote (Syntax.string_of_term lthy' ctr) ^
  31.342 -          " is not a constructor in left-hand side") user_eqn) eqns_data end;
  31.343 -
  31.344 -    val defs = build_defs lthy' bs mxs funs_data rec_specs has_call;
  31.345 -
  31.346 -    fun prove lthy def_thms' ({ctr_specs, nested_map_idents, nested_map_comps, ...} : rec_spec)
  31.347 -        (fun_data : eqn_data list) =
  31.348 -      let
  31.349 -        val def_thms = map (snd o snd) def_thms';
  31.350 -        val simp_thmss = finds (fn (x, y) => #ctr x = #ctr y) fun_data ctr_specs
  31.351 -          |> fst
  31.352 -          |> map_filter (try (fn (x, [y]) =>
  31.353 -            (#user_eqn x, length (#left_args x) + length (#right_args x), #rec_thm y)))
  31.354 -          |> map (fn (user_eqn, num_extra_args, rec_thm) =>
  31.355 -            mk_primrec_tac lthy num_extra_args nested_map_idents nested_map_comps def_thms rec_thm
  31.356 -            |> K |> Goal.prove lthy [] [] user_eqn);
  31.357 -        val poss = find_indices (fn (x, y) => #ctr x = #ctr y) fun_data eqns_data;
  31.358 -      in
  31.359 -        (poss, simp_thmss)
  31.360 -      end;
  31.361 -
  31.362 -    val notes =
  31.363 -      (if n2m then map2 (fn name => fn thm =>
  31.364 -        (name, inductN, [thm], [])) fun_names (take actual_nn induct_thms) else [])
  31.365 -      |> map (fn (prefix, thmN, thms, attrs) =>
  31.366 -        ((Binding.qualify true prefix (Binding.name thmN), attrs), [(thms, [])]));
  31.367 -
  31.368 -    val common_name = mk_common_name fun_names;
  31.369 -
  31.370 -    val common_notes =
  31.371 -      (if n2m then [(inductN, [induct_thm], [])] else [])
  31.372 -      |> map (fn (thmN, thms, attrs) =>
  31.373 -        ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
  31.374 -  in
  31.375 -    (((fun_names, defs),
  31.376 -      fn lthy => fn defs =>
  31.377 -        split_list (map2 (prove lthy defs) (take actual_nn rec_specs) funs_data)),
  31.378 -      lthy' |> Local_Theory.notes (notes @ common_notes) |> snd)
  31.379 -  end;
  31.380 -
  31.381 -(* primrec definition *)
  31.382 -
  31.383 -fun add_primrec_simple fixes ts lthy =
  31.384 -  let
  31.385 -    val (((names, defs), prove), lthy) = prepare_primrec fixes ts lthy
  31.386 -      handle ERROR str => primrec_error str;
  31.387 -  in
  31.388 -    lthy
  31.389 -    |> fold_map Local_Theory.define defs
  31.390 -    |-> (fn defs => `(fn lthy => (names, (map fst defs, prove lthy defs))))
  31.391 -  end
  31.392 -  handle Primrec_Error (str, eqns) =>
  31.393 -    if null eqns
  31.394 -    then error ("primrec_new error:\n  " ^ str)
  31.395 -    else error ("primrec_new error:\n  " ^ str ^ "\nin\n  " ^
  31.396 -      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
  31.397 -
  31.398 -local
  31.399 -
  31.400 -fun gen_primrec prep_spec (raw_fixes : (binding * 'a option * mixfix) list) raw_spec lthy =
  31.401 -  let
  31.402 -    val d = duplicates (op =) (map (Binding.name_of o #1) raw_fixes)
  31.403 -    val _ = null d orelse primrec_error ("duplicate function name(s): " ^ commas d);
  31.404 -
  31.405 -    val (fixes, specs) = fst (prep_spec raw_fixes raw_spec lthy);
  31.406 -
  31.407 -    val mk_notes =
  31.408 -      flat ooo map3 (fn poss => fn prefix => fn thms =>
  31.409 -        let
  31.410 -          val (bs, attrss) = map_split (fst o nth specs) poss;
  31.411 -          val notes =
  31.412 -            map3 (fn b => fn attrs => fn thm =>
  31.413 -              ((Binding.qualify false prefix b, code_nitpick_simp_attrs @ attrs), [([thm], [])]))
  31.414 -            bs attrss thms;
  31.415 -        in
  31.416 -          ((Binding.qualify true prefix (Binding.name simpsN), []), [(thms, [])]) :: notes
  31.417 -        end);
  31.418 -  in
  31.419 -    lthy
  31.420 -    |> add_primrec_simple fixes (map snd specs)
  31.421 -    |-> (fn (names, (ts, (posss, simpss))) =>
  31.422 -      Spec_Rules.add Spec_Rules.Equational (ts, flat simpss)
  31.423 -      #> Local_Theory.notes (mk_notes posss names simpss)
  31.424 -      #>> pair ts o map snd)
  31.425 -  end;
  31.426 -
  31.427 -in
  31.428 -
  31.429 -val add_primrec = gen_primrec Specification.check_spec;
  31.430 -val add_primrec_cmd = gen_primrec Specification.read_spec;
  31.431 -
  31.432 -end;
  31.433 -
  31.434 -fun add_primrec_global fixes specs thy =
  31.435 -  let
  31.436 -    val lthy = Named_Target.theory_init thy;
  31.437 -    val ((ts, simps), lthy') = add_primrec fixes specs lthy;
  31.438 -    val simps' = burrow (Proof_Context.export lthy' lthy) simps;
  31.439 -  in ((ts, simps'), Local_Theory.exit_global lthy') end;
  31.440 -
  31.441 -fun add_primrec_overloaded ops fixes specs thy =
  31.442 -  let
  31.443 -    val lthy = Overloading.overloading ops thy;
  31.444 -    val ((ts, simps), lthy') = add_primrec fixes specs lthy;
  31.445 -    val simps' = burrow (Proof_Context.export lthy' lthy) simps;
  31.446 -  in ((ts, simps'), Local_Theory.exit_global lthy') end;
  31.447 -
  31.448 -
  31.449 -
  31.450 -(* Primcorec *)
  31.451 -
  31.452 -type co_eqn_data_disc = {
  31.453 -  fun_name: string,
  31.454 -  fun_T: typ,
  31.455 -  fun_args: term list,
  31.456 -  ctr: term,
  31.457 -  ctr_no: int, (*###*)
  31.458 -  disc: term,
  31.459 -  prems: term list,
  31.460 -  auto_gen: bool,
  31.461 -  user_eqn: term
  31.462 -};
  31.463 -
  31.464 -type co_eqn_data_sel = {
  31.465 -  fun_name: string,
  31.466 -  fun_T: typ,
  31.467 -  fun_args: term list,
  31.468 -  ctr: term,
  31.469 -  sel: term,
  31.470 -  rhs_term: term,
  31.471 -  user_eqn: term
  31.472 -};
  31.473 -
  31.474 -datatype co_eqn_data =
  31.475 -  Disc of co_eqn_data_disc |
  31.476 -  Sel of co_eqn_data_sel;
  31.477 -
  31.478 -fun co_dissect_eqn_disc sequential fun_names (corec_specs : corec_spec list) prems' concl
  31.479 -    matchedsss =
  31.480 -  let
  31.481 -    fun find_subterm p = let (* FIXME \<exists>? *)
  31.482 -      fun f (t as u $ v) = if p t then SOME t else merge_options (f u, f v)
  31.483 -        | f t = if p t then SOME t else NONE
  31.484 -      in f end;
  31.485 -
  31.486 -    val applied_fun = concl
  31.487 -      |> find_subterm (member ((op =) o apsnd SOME) fun_names o try (fst o dest_Free o head_of))
  31.488 -      |> the
  31.489 -      handle Option.Option => primrec_error_eqn "malformed discriminator equation" concl;
  31.490 -    val ((fun_name, fun_T), fun_args) = strip_comb applied_fun |>> dest_Free;
  31.491 -    val {ctr_specs, ...} = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name);
  31.492 -
  31.493 -    val discs = map #disc ctr_specs;
  31.494 -    val ctrs = map #ctr ctr_specs;
  31.495 -    val not_disc = head_of concl = @{term Not};
  31.496 -    val _ = not_disc andalso length ctrs <> 2 andalso
  31.497 -      primrec_error_eqn "\<not>ed discriminator for a type with \<noteq> 2 constructors" concl;
  31.498 -    val disc = find_subterm (member (op =) discs o head_of) concl;
  31.499 -    val eq_ctr0 = concl |> perhaps (try (HOLogic.dest_not)) |> try (HOLogic.dest_eq #> snd)
  31.500 -        |> (fn SOME t => let val n = find_index (equal t) ctrs in
  31.501 -          if n >= 0 then SOME n else NONE end | _ => NONE);
  31.502 -    val _ = is_some disc orelse is_some eq_ctr0 orelse
  31.503 -      primrec_error_eqn "no discriminator in equation" concl;
  31.504 -    val ctr_no' =
  31.505 -      if is_none disc then the eq_ctr0 else find_index (equal (head_of (the disc))) discs;
  31.506 -    val ctr_no = if not_disc then 1 - ctr_no' else ctr_no';
  31.507 -    val ctr = #ctr (nth ctr_specs ctr_no);
  31.508 -
  31.509 -    val catch_all = try (fst o dest_Free o the_single) prems' = SOME Name.uu_;
  31.510 -    val matchedss = AList.lookup (op =) matchedsss fun_name |> the_default [];
  31.511 -    val prems = map (abstract (List.rev fun_args)) prems';
  31.512 -    val real_prems =
  31.513 -      (if catch_all orelse sequential then maps negate_disj matchedss else []) @
  31.514 -      (if catch_all then [] else prems);
  31.515 -
  31.516 -    val matchedsss' = AList.delete (op =) fun_name matchedsss
  31.517 -      |> cons (fun_name, if sequential then matchedss @ [prems] else matchedss @ [real_prems]);
  31.518 -
  31.519 -    val user_eqn =
  31.520 -      (real_prems, betapply (#disc (nth ctr_specs ctr_no), applied_fun))
  31.521 -      |>> map HOLogic.mk_Trueprop ||> HOLogic.mk_Trueprop
  31.522 -      |> Logic.list_implies;
  31.523 -  in
  31.524 -    (Disc {
  31.525 -      fun_name = fun_name,
  31.526 -      fun_T = fun_T,
  31.527 -      fun_args = fun_args,
  31.528 -      ctr = ctr,
  31.529 -      ctr_no = ctr_no,
  31.530 -      disc = #disc (nth ctr_specs ctr_no),
  31.531 -      prems = real_prems,
  31.532 -      auto_gen = catch_all,
  31.533 -      user_eqn = user_eqn
  31.534 -    }, matchedsss')
  31.535 -  end;
  31.536 -
  31.537 -fun co_dissect_eqn_sel fun_names (corec_specs : corec_spec list) eqn' of_spec eqn =
  31.538 -  let
  31.539 -    val (lhs, rhs) = HOLogic.dest_eq eqn
  31.540 -      handle TERM _ =>
  31.541 -        primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn;
  31.542 -    val sel = head_of lhs;
  31.543 -    val ((fun_name, fun_T), fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst dest_Free
  31.544 -      handle TERM _ =>
  31.545 -        primrec_error_eqn "malformed selector argument in left-hand side" eqn;
  31.546 -    val corec_spec = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name)
  31.547 -      handle Option.Option => primrec_error_eqn "malformed selector argument in left-hand side" eqn;
  31.548 -    val ctr_spec =
  31.549 -      if is_some of_spec
  31.550 -      then the (find_first (equal (the of_spec) o #ctr) (#ctr_specs corec_spec))
  31.551 -      else #ctr_specs corec_spec |> filter (exists (equal sel) o #sels) |> the_single
  31.552 -        handle List.Empty => primrec_error_eqn "ambiguous selector - use \"of\"" eqn;
  31.553 -    val user_eqn = drop_All eqn';
  31.554 -  in
  31.555 -    Sel {
  31.556 -      fun_name = fun_name,
  31.557 -      fun_T = fun_T,
  31.558 -      fun_args = fun_args,
  31.559 -      ctr = #ctr ctr_spec,
  31.560 -      sel = sel,
  31.561 -      rhs_term = rhs,
  31.562 -      user_eqn = user_eqn
  31.563 -    }
  31.564 -  end;
  31.565 -
  31.566 -fun co_dissect_eqn_ctr sequential fun_names (corec_specs : corec_spec list) eqn' imp_prems imp_rhs
  31.567 -    matchedsss =
  31.568 -  let
  31.569 -    val (lhs, rhs) = HOLogic.dest_eq imp_rhs;
  31.570 -    val fun_name = head_of lhs |> fst o dest_Free;
  31.571 -    val {ctr_specs, ...} = the (AList.lookup (op =) (fun_names ~~ corec_specs) fun_name);
  31.572 -    val (ctr, ctr_args) = strip_comb rhs;
  31.573 -    val {disc, sels, ...} = the (find_first (equal ctr o #ctr) ctr_specs)
  31.574 -      handle Option.Option => primrec_error_eqn "not a constructor" ctr;
  31.575 -
  31.576 -    val disc_imp_rhs = betapply (disc, lhs);
  31.577 -    val (maybe_eqn_data_disc, matchedsss') = if length ctr_specs = 1
  31.578 -      then (NONE, matchedsss)
  31.579 -      else apfst SOME (co_dissect_eqn_disc
  31.580 -          sequential fun_names corec_specs imp_prems disc_imp_rhs matchedsss);
  31.581 -
  31.582 -    val sel_imp_rhss = (sels ~~ ctr_args)
  31.583 -      |> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg));
  31.584 -
  31.585 -(*
  31.586 -val _ = tracing ("reduced\n    " ^ Syntax.string_of_term @{context} imp_rhs ^ "\nto\n    \<cdot> " ^
  31.587 - (is_some maybe_eqn_data_disc ? K (Syntax.string_of_term @{context} disc_imp_rhs ^ "\n    \<cdot> ")) "" ^
  31.588 - space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) sel_imp_rhss));
  31.589 -*)
  31.590 -
  31.591 -    val eqns_data_sel =
  31.592 -      map (co_dissect_eqn_sel fun_names corec_specs eqn' (SOME ctr)) sel_imp_rhss;
  31.593 -  in
  31.594 -    (the_list maybe_eqn_data_disc @ eqns_data_sel, matchedsss')
  31.595 -  end;
  31.596 -
  31.597 -fun co_dissect_eqn sequential fun_names (corec_specs : corec_spec list) eqn' of_spec matchedsss =
  31.598 -  let
  31.599 -    val eqn = drop_All eqn'
  31.600 -      handle TERM _ => primrec_error_eqn "malformed function equation" eqn';
  31.601 -    val (imp_prems, imp_rhs) = Logic.strip_horn eqn
  31.602 -      |> apfst (map HOLogic.dest_Trueprop) o apsnd HOLogic.dest_Trueprop;
  31.603 -
  31.604 -    val head = imp_rhs
  31.605 -      |> perhaps (try HOLogic.dest_not) |> perhaps (try (fst o HOLogic.dest_eq))
  31.606 -      |> head_of;
  31.607 -
  31.608 -    val maybe_rhs = imp_rhs |> perhaps (try (HOLogic.dest_not)) |> try (snd o HOLogic.dest_eq);
  31.609 -
  31.610 -    val discs = maps #ctr_specs corec_specs |> map #disc;
  31.611 -    val sels = maps #ctr_specs corec_specs |> maps #sels;
  31.612 -    val ctrs = maps #ctr_specs corec_specs |> map #ctr;
  31.613 -  in
  31.614 -    if member (op =) discs head orelse
  31.615 -      is_some maybe_rhs andalso
  31.616 -        member (op =) (filter (null o binder_types o fastype_of) ctrs) (the maybe_rhs) then
  31.617 -      co_dissect_eqn_disc sequential fun_names corec_specs imp_prems imp_rhs matchedsss
  31.618 -      |>> single
  31.619 -    else if member (op =) sels head then
  31.620 -      ([co_dissect_eqn_sel fun_names corec_specs eqn' of_spec imp_rhs], matchedsss)
  31.621 -    else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) then
  31.622 -      co_dissect_eqn_ctr sequential fun_names corec_specs eqn' imp_prems imp_rhs matchedsss
  31.623 -    else
  31.624 -      primrec_error_eqn "malformed function equation" eqn
  31.625 -  end;
  31.626 -
  31.627 -fun build_corec_arg_disc (ctr_specs : corec_ctr_spec list)
  31.628 -    ({fun_args, ctr_no, prems, ...} : co_eqn_data_disc) =
  31.629 -  if is_none (#pred (nth ctr_specs ctr_no)) then I else
  31.630 -    mk_conjs prems
  31.631 -    |> curry subst_bounds (List.rev fun_args)
  31.632 -    |> HOLogic.tupled_lambda (HOLogic.mk_tuple fun_args)
  31.633 -    |> K |> nth_map (the (#pred (nth ctr_specs ctr_no)));
  31.634 -
  31.635 -fun build_corec_arg_no_call (sel_eqns : co_eqn_data_sel list) sel =
  31.636 -  find_first (equal sel o #sel) sel_eqns
  31.637 -  |> try (fn SOME {fun_args, rhs_term, ...} => abs_tuple fun_args rhs_term)
  31.638 -  |> the_default undef_const
  31.639 -  |> K;
  31.640 -
  31.641 -fun build_corec_args_direct_call lthy has_call (sel_eqns : co_eqn_data_sel list) sel =
  31.642 -  let
  31.643 -    val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns;
  31.644 -  in
  31.645 -    if is_none maybe_sel_eqn then (I, I, I) else
  31.646 -    let
  31.647 -      val {fun_args, rhs_term, ... } = the maybe_sel_eqn;
  31.648 -      fun rewrite_q _ t = if has_call t then @{term False} else @{term True};
  31.649 -      fun rewrite_g _ t = if has_call t then undef_const else t;
  31.650 -      fun rewrite_h bound_Ts t =
  31.651 -        if has_call t then mk_tuple1 bound_Ts (snd (strip_comb t)) else undef_const;
  31.652 -      fun massage f t = massage_direct_corec_call lthy has_call f [] rhs_term |> abs_tuple fun_args;
  31.653 -    in
  31.654 -      (massage rewrite_q,
  31.655 -       massage rewrite_g,
  31.656 -       massage rewrite_h)
  31.657 -    end
  31.658 -  end;
  31.659 -
  31.660 -fun build_corec_arg_indirect_call lthy has_call (sel_eqns : co_eqn_data_sel list) sel =
  31.661 -  let
  31.662 -    val maybe_sel_eqn = find_first (equal sel o #sel) sel_eqns;
  31.663 -  in
  31.664 -    if is_none maybe_sel_eqn then I else
  31.665 -    let
  31.666 -      val {fun_args, rhs_term, ...} = the maybe_sel_eqn;
  31.667 -      fun rewrite bound_Ts U T (Abs (v, V, b)) = Abs (v, V, rewrite (V :: bound_Ts) U T b)
  31.668 -        | rewrite bound_Ts U T (t as _ $ _) =
  31.669 -          let val (u, vs) = strip_comb t in
  31.670 -            if is_Free u andalso has_call u then
  31.671 -              Inr_const U T $ mk_tuple1 bound_Ts vs
  31.672 -            else if try (fst o dest_Const) u = SOME @{const_name prod_case} then
  31.673 -              map (rewrite bound_Ts U T) vs |> chop 1 |>> HOLogic.mk_split o the_single |> list_comb
  31.674 -            else
  31.675 -              list_comb (rewrite bound_Ts U T u, map (rewrite bound_Ts U T) vs)
  31.676 -          end
  31.677 -        | rewrite _ U T t =
  31.678 -          if is_Free t andalso has_call t then Inr_const U T $ HOLogic.unit else t;
  31.679 -      fun massage t =
  31.680 -        massage_indirect_corec_call lthy has_call rewrite [] (range_type (fastype_of t)) rhs_term
  31.681 -        |> abs_tuple fun_args;
  31.682 -    in
  31.683 -      massage
  31.684 -    end
  31.685 -  end;
  31.686 -
  31.687 -fun build_corec_args_sel lthy has_call (all_sel_eqns : co_eqn_data_sel list)
  31.688 -    (ctr_spec : corec_ctr_spec) =
  31.689 -  let val sel_eqns = filter (equal (#ctr ctr_spec) o #ctr) all_sel_eqns in
  31.690 -    if null sel_eqns then I else
  31.691 -      let
  31.692 -        val sel_call_list = #sels ctr_spec ~~ #calls ctr_spec;
  31.693 -
  31.694 -        val no_calls' = map_filter (try (apsnd (fn No_Corec n => n))) sel_call_list;
  31.695 -        val direct_calls' = map_filter (try (apsnd (fn Direct_Corec n => n))) sel_call_list;
  31.696 -        val indirect_calls' = map_filter (try (apsnd (fn Indirect_Corec n => n))) sel_call_list;
  31.697 -      in
  31.698 -        I
  31.699 -        #> fold (fn (sel, n) => nth_map n (build_corec_arg_no_call sel_eqns sel)) no_calls'
  31.700 -        #> fold (fn (sel, (q, g, h)) =>
  31.701 -          let val (fq, fg, fh) = build_corec_args_direct_call lthy has_call sel_eqns sel in
  31.702 -            nth_map q fq o nth_map g fg o nth_map h fh end) direct_calls'
  31.703 -        #> fold (fn (sel, n) => nth_map n
  31.704 -          (build_corec_arg_indirect_call lthy has_call sel_eqns sel)) indirect_calls'
  31.705 -      end
  31.706 -  end;
  31.707 -
  31.708 -fun co_build_defs lthy bs mxs has_call arg_Tss (corec_specs : corec_spec list)
  31.709 -    (disc_eqnss : co_eqn_data_disc list list) (sel_eqnss : co_eqn_data_sel list list) =
  31.710 -  let
  31.711 -    val corec_specs' = take (length bs) corec_specs;
  31.712 -    val corecs = map #corec corec_specs';
  31.713 -    val ctr_specss = map #ctr_specs corec_specs';
  31.714 -    val corec_args = hd corecs
  31.715 -      |> fst o split_last o binder_types o fastype_of
  31.716 -      |> map (Const o pair @{const_name undefined})
  31.717 -      |> fold2 (fold o build_corec_arg_disc) ctr_specss disc_eqnss
  31.718 -      |> fold2 (fold o build_corec_args_sel lthy has_call) sel_eqnss ctr_specss;
  31.719 -    fun currys [] t = t
  31.720 -      | currys Ts t = t $ mk_tuple1 (List.rev Ts) (map Bound (length Ts - 1 downto 0))
  31.721 -          |> fold_rev (Term.abs o pair Name.uu) Ts;
  31.722 -
  31.723 -(*
  31.724 -val _ = tracing ("corecursor arguments:\n    \<cdot> " ^
  31.725 - space_implode "\n    \<cdot> " (map (Syntax.string_of_term lthy) corec_args));
  31.726 -*)
  31.727 -
  31.728 -    val exclss' =
  31.729 -      disc_eqnss
  31.730 -      |> map (map (fn x => (#fun_args x, #ctr_no x, #prems x, #auto_gen x))
  31.731 -        #> fst o (fn xs => fold_map (fn x => fn ys => ((x, ys), ys @ [x])) xs [])
  31.732 -        #> maps (uncurry (map o pair)
  31.733 -          #> map (fn ((fun_args, c, x, a), (_, c', y, a')) =>
  31.734 -              ((c, c', a orelse a'), (x, s_not (mk_conjs y)))
  31.735 -            ||> apfst (map HOLogic.mk_Trueprop) o apsnd HOLogic.mk_Trueprop
  31.736 -            ||> Logic.list_implies
  31.737 -            ||> curry Logic.list_all (map dest_Free fun_args))))
  31.738 -  in
  31.739 -    map (list_comb o rpair corec_args) corecs
  31.740 -    |> map2 (fn Ts => fn t => if length Ts = 0 then t $ HOLogic.unit else t) arg_Tss
  31.741 -    |> map2 currys arg_Tss
  31.742 -    |> Syntax.check_terms lthy
  31.743 -    |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.map_name Thm.def_name b, []), t))) bs mxs
  31.744 -    |> rpair exclss'
  31.745 -  end;
  31.746 -
  31.747 -fun mk_real_disc_eqns fun_binding arg_Ts ({ctr_specs, ...} : corec_spec)
  31.748 -    (sel_eqns : co_eqn_data_sel list) (disc_eqns : co_eqn_data_disc list) =
  31.749 -  if length disc_eqns <> length ctr_specs - 1 then disc_eqns else
  31.750 -    let
  31.751 -      val n = 0 upto length ctr_specs
  31.752 -        |> the o find_first (fn idx => not (exists (equal idx o #ctr_no) disc_eqns));
  31.753 -      val fun_args = (try (#fun_args o hd) disc_eqns, try (#fun_args o hd) sel_eqns)
  31.754 -        |> the_default (map (curry Free Name.uu) arg_Ts) o merge_options;
  31.755 -      val extra_disc_eqn = {
  31.756 -        fun_name = Binding.name_of fun_binding,
  31.757 -        fun_T = arg_Ts ---> body_type (fastype_of (#ctr (hd ctr_specs))),
  31.758 -        fun_args = fun_args,
  31.759 -        ctr = #ctr (nth ctr_specs n),
  31.760 -        ctr_no = n,
  31.761 -        disc = #disc (nth ctr_specs n),
  31.762 -        prems = maps (negate_conj o #prems) disc_eqns,
  31.763 -        auto_gen = true,
  31.764 -        user_eqn = undef_const};
  31.765 -    in
  31.766 -      chop n disc_eqns ||> cons extra_disc_eqn |> (op @)
  31.767 -    end;
  31.768 -
  31.769 -fun add_primcorec simple sequential fixes specs of_specs lthy =
  31.770 -  let
  31.771 -    val (bs, mxs) = map_split (apfst fst) fixes;
  31.772 -    val (arg_Ts, res_Ts) = map (strip_type o snd o fst #>> HOLogic.mk_tupleT) fixes |> split_list;
  31.773 -
  31.774 -    val callssss = []; (* FIXME *)
  31.775 -
  31.776 -    val ((n2m, corec_specs', _, coinduct_thm, strong_coinduct_thm, coinduct_thms,
  31.777 -          strong_coinduct_thms), lthy') =
  31.778 -      corec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
  31.779 -
  31.780 -    val actual_nn = length bs;
  31.781 -    val fun_names = map Binding.name_of bs;
  31.782 -    val corec_specs = take actual_nn corec_specs'; (*###*)
  31.783 -
  31.784 -    val eqns_data =
  31.785 -      fold_map2 (co_dissect_eqn sequential fun_names corec_specs) (map snd specs) of_specs []
  31.786 -      |> flat o fst;
  31.787 -
  31.788 -    val disc_eqnss' = map_filter (try (fn Disc x => x)) eqns_data
  31.789 -      |> partition_eq ((op =) o pairself #fun_name)
  31.790 -      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
  31.791 -      |> map (sort ((op <) o pairself #ctr_no |> make_ord) o flat o snd);
  31.792 -    val _ = disc_eqnss' |> map (fn x =>
  31.793 -      let val d = duplicates ((op =) o pairself #ctr_no) x in null d orelse
  31.794 -        primrec_error_eqns "excess discriminator equations in definition"
  31.795 -          (maps (fn t => filter (equal (#ctr_no t) o #ctr_no) x) d |> map #user_eqn) end);
  31.796 -
  31.797 -    val sel_eqnss = map_filter (try (fn Sel x => x)) eqns_data
  31.798 -      |> partition_eq ((op =) o pairself #fun_name)
  31.799 -      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
  31.800 -      |> map (flat o snd);
  31.801 -
  31.802 -    val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
  31.803 -    val arg_Tss = map (binder_types o snd o fst) fixes;
  31.804 -    val disc_eqnss = map5 mk_real_disc_eqns bs arg_Tss corec_specs sel_eqnss disc_eqnss';
  31.805 -    val (defs, exclss') =
  31.806 -      co_build_defs lthy' bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss;
  31.807 -
  31.808 -    fun excl_tac (c, c', a) =
  31.809 -      if a orelse c = c' orelse sequential then
  31.810 -        SOME (K (HEADGOAL (mk_primcorec_assumption_tac lthy [])))
  31.811 -      else if simple then
  31.812 -        SOME (K (auto_tac lthy))
  31.813 -      else
  31.814 -        NONE;
  31.815 -
  31.816 -(*
  31.817 -val _ = tracing ("exclusiveness properties:\n    \<cdot> " ^
  31.818 - space_implode "\n    \<cdot> " (maps (map (Syntax.string_of_term lthy o snd)) exclss'));
  31.819 -*)
  31.820 -
  31.821 -    val exclss'' = exclss' |> map (map (fn (idx, t) =>
  31.822 -      (idx, (Option.map (Goal.prove lthy [] [] t) (excl_tac idx), t))));
  31.823 -    val taut_thmss = map (map (apsnd (the o fst)) o filter (is_some o fst o snd)) exclss'';
  31.824 -    val (obligation_idxss, obligationss) = exclss''
  31.825 -      |> map (map (apsnd (rpair [] o snd)) o filter (is_none o fst o snd))
  31.826 -      |> split_list o map split_list;
  31.827 -
  31.828 -    fun prove thmss' def_thms' lthy =
  31.829 -      let
  31.830 -        val def_thms = map (snd o snd) def_thms';
  31.831 -
  31.832 -        val exclss' = map (op ~~) (obligation_idxss ~~ thmss');
  31.833 -        fun mk_exclsss excls n =
  31.834 -          (excls, map (fn k => replicate k [TrueI] @ replicate (n - k) []) (0 upto n - 1))
  31.835 -          |-> fold (fn ((c, c', _), thm) => nth_map c (nth_map c' (K [thm])));
  31.836 -        val exclssss = (exclss' ~~ taut_thmss |> map (op @), fun_names ~~ corec_specs)
  31.837 -          |-> map2 (fn excls => fn (_, {ctr_specs, ...}) => mk_exclsss excls (length ctr_specs));
  31.838 -
  31.839 -        fun prove_disc ({ctr_specs, ...} : corec_spec) exclsss
  31.840 -            ({fun_name, fun_T, fun_args, ctr_no, prems, ...} : co_eqn_data_disc) =
  31.841 -          if Term.aconv_untyped (#disc (nth ctr_specs ctr_no), @{term "\<lambda>x. x = x"}) then [] else
  31.842 -            let
  31.843 -              val {disc_corec, ...} = nth ctr_specs ctr_no;
  31.844 -              val k = 1 + ctr_no;
  31.845 -              val m = length prems;
  31.846 -              val t =
  31.847 -                list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
  31.848 -                |> curry betapply (#disc (nth ctr_specs ctr_no)) (*###*)
  31.849 -                |> HOLogic.mk_Trueprop
  31.850 -                |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
  31.851 -                |> curry Logic.list_all (map dest_Free fun_args);
  31.852 -            in
  31.853 -              mk_primcorec_disc_tac lthy def_thms disc_corec k m exclsss
  31.854 -              |> K |> Goal.prove lthy [] [] t
  31.855 -              |> pair (#disc (nth ctr_specs ctr_no))
  31.856 -              |> single
  31.857 -            end;
  31.858 -
  31.859 -        fun prove_sel ({nested_maps, nested_map_idents, nested_map_comps, ctr_specs, ...}
  31.860 -            : corec_spec) (disc_eqns : co_eqn_data_disc list) exclsss
  31.861 -            ({fun_name, fun_T, fun_args, ctr, sel, rhs_term, ...} : co_eqn_data_sel) =
  31.862 -          let
  31.863 -            val SOME ctr_spec = find_first (equal ctr o #ctr) ctr_specs;
  31.864 -            val ctr_no = find_index (equal ctr o #ctr) ctr_specs;
  31.865 -            val prems = the_default (maps (negate_conj o #prems) disc_eqns)
  31.866 -                (find_first (equal ctr_no o #ctr_no) disc_eqns |> Option.map #prems);
  31.867 -            val sel_corec = find_index (equal sel) (#sels ctr_spec)
  31.868 -              |> nth (#sel_corecs ctr_spec);
  31.869 -            val k = 1 + ctr_no;
  31.870 -            val m = length prems;
  31.871 -            val t =
  31.872 -              list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
  31.873 -              |> curry betapply sel
  31.874 -              |> rpair (abstract (List.rev fun_args) rhs_term)
  31.875 -              |> HOLogic.mk_Trueprop o HOLogic.mk_eq
  31.876 -              |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
  31.877 -              |> curry Logic.list_all (map dest_Free fun_args);
  31.878 -            val (distincts, _, sel_splits, sel_split_asms) = case_thms_of_term lthy [] rhs_term;
  31.879 -          in
  31.880 -            mk_primcorec_sel_tac lthy def_thms distincts sel_splits sel_split_asms nested_maps
  31.881 -              nested_map_idents nested_map_comps sel_corec k m exclsss
  31.882 -            |> K |> Goal.prove lthy [] [] t
  31.883 -            |> pair sel
  31.884 -          end;
  31.885 -
  31.886 -        fun prove_ctr disc_alist sel_alist (disc_eqns : co_eqn_data_disc list)
  31.887 -            (sel_eqns : co_eqn_data_sel list) ({ctr, disc, sels, collapse, ...} : corec_ctr_spec) =
  31.888 -          if not (exists (equal ctr o #ctr) disc_eqns)
  31.889 -              andalso not (exists (equal ctr o #ctr) sel_eqns)
  31.890 -            orelse (* don't try to prove theorems when some sel_eqns are missing *)
  31.891 -              filter (equal ctr o #ctr) sel_eqns
  31.892 -              |> fst o finds ((op =) o apsnd #sel) sels
  31.893 -              |> exists (null o snd)
  31.894 -          then [] else
  31.895 -            let
  31.896 -              val (fun_name, fun_T, fun_args, prems) =
  31.897 -                (find_first (equal ctr o #ctr) disc_eqns, find_first (equal ctr o #ctr) sel_eqns)
  31.898 -                |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #prems x))
  31.899 -                ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, []))
  31.900 -                |> the o merge_options;
  31.901 -              val m = length prems;
  31.902 -              val t = filter (equal ctr o #ctr) sel_eqns
  31.903 -                |> fst o finds ((op =) o apsnd #sel) sels
  31.904 -                |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x)) #-> abstract)
  31.905 -                |> curry list_comb ctr
  31.906 -                |> curry HOLogic.mk_eq (list_comb (Free (fun_name, fun_T),
  31.907 -                  map Bound (length fun_args - 1 downto 0)))
  31.908 -                |> HOLogic.mk_Trueprop
  31.909 -                |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
  31.910 -                |> curry Logic.list_all (map dest_Free fun_args);
  31.911 -              val maybe_disc_thm = AList.lookup (op =) disc_alist disc;
  31.912 -              val sel_thms = map snd (filter (member (op =) sels o fst) sel_alist);
  31.913 -            in
  31.914 -              mk_primcorec_ctr_of_dtr_tac lthy m collapse maybe_disc_thm sel_thms
  31.915 -              |> K |> Goal.prove lthy [] [] t
  31.916 -              |> single
  31.917 -            end;
  31.918 -
  31.919 -        val disc_alists = map3 (maps oo prove_disc) corec_specs exclssss disc_eqnss;
  31.920 -        val sel_alists = map4 (map ooo prove_sel) corec_specs disc_eqnss exclssss sel_eqnss;
  31.921 -
  31.922 -        val disc_thmss = map (map snd) disc_alists;
  31.923 -        val sel_thmss = map (map snd) sel_alists;
  31.924 -        val ctr_thmss = map5 (maps oooo prove_ctr) disc_alists sel_alists disc_eqnss sel_eqnss
  31.925 -          (map #ctr_specs corec_specs);
  31.926 -
  31.927 -        val simp_thmss = map2 append disc_thmss sel_thmss
  31.928 -
  31.929 -        val common_name = mk_common_name fun_names;
  31.930 -
  31.931 -        val notes =
  31.932 -          [(coinductN, map (if n2m then single else K []) coinduct_thms, []),
  31.933 -           (codeN, ctr_thmss(*FIXME*), code_nitpick_attrs),
  31.934 -           (ctrN, ctr_thmss, []),
  31.935 -           (discN, disc_thmss, simp_attrs),
  31.936 -           (selN, sel_thmss, simp_attrs),
  31.937 -           (simpsN, simp_thmss, []),
  31.938 -           (strong_coinductN, map (if n2m then single else K []) strong_coinduct_thms, [])]
  31.939 -          |> maps (fn (thmN, thmss, attrs) =>
  31.940 -            map2 (fn fun_name => fn thms =>
  31.941 -                ((Binding.qualify true fun_name (Binding.name thmN), attrs), [(thms, [])]))
  31.942 -              fun_names (take actual_nn thmss))
  31.943 -          |> filter_out (null o fst o hd o snd);
  31.944 -
  31.945 -        val common_notes =
  31.946 -          [(coinductN, if n2m then [coinduct_thm] else [], []),
  31.947 -           (strong_coinductN, if n2m then [strong_coinduct_thm] else [], [])]
  31.948 -          |> filter_out (null o #2)
  31.949 -          |> map (fn (thmN, thms, attrs) =>
  31.950 -            ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
  31.951 -      in
  31.952 -        lthy |> Local_Theory.notes (notes @ common_notes) |> snd
  31.953 -      end;
  31.954 -
  31.955 -    fun after_qed thmss' = fold_map Local_Theory.define defs #-> prove thmss';
  31.956 -
  31.957 -    val _ = if not simple orelse forall null obligationss then () else
  31.958 -      primrec_error "need exclusiveness proofs - use primcorecursive instead of primcorec";
  31.959 -  in
  31.960 -    if simple then
  31.961 -      lthy'
  31.962 -      |> after_qed (map (fn [] => []) obligationss)
  31.963 -      |> pair NONE o SOME
  31.964 -    else
  31.965 -      lthy'
  31.966 -      |> Proof.theorem NONE after_qed obligationss
  31.967 -      |> Proof.refine (Method.primitive_text I)
  31.968 -      |> Seq.hd
  31.969 -      |> rpair NONE o SOME
  31.970 -  end;
  31.971 -
  31.972 -fun add_primcorec_ursive_cmd simple seq (raw_fixes, raw_specs') lthy =
  31.973 -  let
  31.974 -    val (raw_specs, of_specs) = split_list raw_specs' ||> map (Option.map (Syntax.read_term lthy));
  31.975 -    val ((fixes, specs), _) = Specification.read_spec raw_fixes raw_specs lthy;
  31.976 -  in
  31.977 -    add_primcorec simple seq fixes specs of_specs lthy
  31.978 -    handle ERROR str => primrec_error str
  31.979 -  end
  31.980 -  handle Primrec_Error (str, eqns) =>
  31.981 -    if null eqns
  31.982 -    then error ("primcorec error:\n  " ^ str)
  31.983 -    else error ("primcorec error:\n  " ^ str ^ "\nin\n  " ^
  31.984 -      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
  31.985 -
  31.986 -val add_primcorecursive_cmd = (the o fst) ooo add_primcorec_ursive_cmd false;
  31.987 -val add_primcorec_cmd = (the o snd) ooo add_primcorec_ursive_cmd true;
  31.988 -
  31.989 -end;
    32.1 --- a/src/HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML	Mon Nov 11 17:34:44 2013 +0100
    32.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    32.3 @@ -1,116 +0,0 @@
    32.4 -(*  Title:      HOL/BNF/Tools/bnf_fp_rec_sugar_tactics.ML
    32.5 -    Author:     Jasmin Blanchette, TU Muenchen
    32.6 -    Copyright   2013
    32.7 -
    32.8 -Tactics for recursor and corecursor sugar.
    32.9 -*)
   32.10 -
   32.11 -signature BNF_FP_REC_SUGAR_TACTICS =
   32.12 -sig
   32.13 -  val mk_primcorec_assumption_tac: Proof.context -> thm list -> int -> tactic
   32.14 -  val mk_primcorec_code_of_raw_code_tac: thm list -> thm -> tactic
   32.15 -  val mk_primcorec_ctr_of_dtr_tac: Proof.context -> int -> thm -> thm option -> thm list -> tactic
   32.16 -  val mk_primcorec_disc_tac: Proof.context -> thm list -> thm -> int -> int -> thm list list list ->
   32.17 -    tactic
   32.18 -  val mk_primcorec_raw_code_of_ctr_tac: Proof.context -> thm list -> thm list -> thm list ->
   32.19 -    thm list -> int list -> thm list -> tactic
   32.20 -  val mk_primcorec_sel_tac: Proof.context -> thm list -> thm list -> thm list -> thm list ->
   32.21 -    thm list -> thm list -> thm list -> thm -> int -> int -> thm list list list -> tactic
   32.22 -  val mk_primrec_tac: Proof.context -> int -> thm list -> thm list -> thm list -> thm -> tactic
   32.23 -end;
   32.24 -
   32.25 -structure BNF_FP_Rec_Sugar_Tactics : BNF_FP_REC_SUGAR_TACTICS =
   32.26 -struct
   32.27 -
   32.28 -open BNF_Util
   32.29 -open BNF_Tactics
   32.30 -
   32.31 -val falseEs = @{thms not_TrueE FalseE};
   32.32 -val neq_eq_eq_contradict = @{thm neq_eq_eq_contradict};
   32.33 -val split_if = @{thm split_if};
   32.34 -val split_if_asm = @{thm split_if_asm};
   32.35 -val split_connectI = @{thms allI impI conjI};
   32.36 -
   32.37 -fun mk_primrec_tac ctxt num_extra_args map_idents map_comps fun_defs recx =
   32.38 -  unfold_thms_tac ctxt fun_defs THEN
   32.39 -  HEADGOAL (rtac (funpow num_extra_args (fn thm => thm RS fun_cong) recx RS trans)) THEN
   32.40 -  unfold_thms_tac ctxt (@{thms id_def split o_def fst_conv snd_conv} @ map_comps @ map_idents) THEN
   32.41 -  HEADGOAL (rtac refl);
   32.42 -
   32.43 -fun mk_primcorec_assumption_tac ctxt discIs =
   32.44 -  SELECT_GOAL (unfold_thms_tac ctxt
   32.45 -      @{thms not_not not_False_eq_True de_Morgan_conj de_Morgan_disj} THEN
   32.46 -    SOLVE (HEADGOAL (REPEAT o (rtac refl ORELSE' atac ORELSE' etac conjE ORELSE'
   32.47 -    resolve_tac @{thms TrueI conjI disjI1 disjI2} ORELSE'
   32.48 -    dresolve_tac discIs THEN' atac ORELSE'
   32.49 -    etac notE THEN' atac ORELSE'
   32.50 -    etac disjE))));
   32.51 -
   32.52 -fun mk_primcorec_same_case_tac m =
   32.53 -  HEADGOAL (if m = 0 then rtac TrueI
   32.54 -    else REPEAT_DETERM_N (m - 1) o (rtac conjI THEN' atac) THEN' atac);
   32.55 -
   32.56 -fun mk_primcorec_different_case_tac ctxt excl =
   32.57 -  unfold_thms_tac ctxt @{thms not_not not_False_eq_True not_True_eq_False} THEN
   32.58 -  HEADGOAL (rtac excl THEN_ALL_NEW mk_primcorec_assumption_tac ctxt []);
   32.59 -
   32.60 -fun mk_primcorec_cases_tac ctxt k m exclsss =
   32.61 -  let val n = length exclsss in
   32.62 -    EVERY (map (fn [] => if k = n then all_tac else mk_primcorec_same_case_tac m
   32.63 -        | [excl] => mk_primcorec_different_case_tac ctxt excl)
   32.64 -      (take k (nth exclsss (k - 1))))
   32.65 -  end;
   32.66 -
   32.67 -fun mk_primcorec_prelude ctxt defs thm =
   32.68 -  unfold_thms_tac ctxt defs THEN HEADGOAL (rtac thm) THEN
   32.69 -  unfold_thms_tac ctxt @{thms Let_def split};
   32.70 -
   32.71 -fun mk_primcorec_disc_tac ctxt defs disc_corec k m exclsss =
   32.72 -  mk_primcorec_prelude ctxt defs disc_corec THEN mk_primcorec_cases_tac ctxt k m exclsss;
   32.73 -
   32.74 -fun mk_primcorec_sel_tac ctxt defs distincts splits split_asms maps map_idents map_comps f_sel k m
   32.75 -    exclsss =
   32.76 -  mk_primcorec_prelude ctxt defs (f_sel RS trans) THEN
   32.77 -  mk_primcorec_cases_tac ctxt k m exclsss THEN
   32.78 -  HEADGOAL (REPEAT_DETERM o (rtac refl ORELSE' rtac ext ORELSE'
   32.79 -    eresolve_tac falseEs ORELSE'
   32.80 -    resolve_tac split_connectI ORELSE'
   32.81 -    Splitter.split_asm_tac (split_if_asm :: split_asms) ORELSE'
   32.82 -    Splitter.split_tac (split_if :: splits) ORELSE'
   32.83 -    eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE'
   32.84 -    etac notE THEN' atac ORELSE'
   32.85 -    (CHANGED o SELECT_GOAL (unfold_thms_tac ctxt
   32.86 -      (@{thms id_apply o_def split_def sum.cases} @ maps @ map_comps @ map_idents)))));
   32.87 -
   32.88 -fun mk_primcorec_ctr_of_dtr_tac ctxt m collapse maybe_disc_f sel_fs =
   32.89 -  HEADGOAL (rtac ((if null sel_fs then collapse else collapse RS sym) RS trans) THEN'
   32.90 -    (the_default (K all_tac) (Option.map rtac maybe_disc_f)) THEN' REPEAT_DETERM_N m o atac) THEN
   32.91 -  unfold_thms_tac ctxt sel_fs THEN HEADGOAL (rtac refl);
   32.92 -
   32.93 -(* TODO: reduce code duplication with selector tactic above *)
   32.94 -fun mk_primcorec_raw_code_of_ctr_single_tac ctxt distincts discIs splits split_asms m f_ctr =
   32.95 -  HEADGOAL (REPEAT o (resolve_tac split_connectI ORELSE' split_tac (split_if :: splits))) THEN
   32.96 -  mk_primcorec_prelude ctxt [] (f_ctr RS trans) THEN
   32.97 -  HEADGOAL ((REPEAT_DETERM_N m o mk_primcorec_assumption_tac ctxt discIs) THEN'
   32.98 -    SELECT_GOAL (SOLVE (HEADGOAL (REPEAT_DETERM o
   32.99 -    (rtac refl ORELSE' atac ORELSE'
  32.100 -     resolve_tac split_connectI ORELSE'
  32.101 -     Splitter.split_asm_tac (split_if_asm :: split_asms) ORELSE'
  32.102 -     Splitter.split_tac (split_if :: splits) ORELSE'
  32.103 -     mk_primcorec_assumption_tac ctxt discIs ORELSE'
  32.104 -     eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE'
  32.105 -     (TRY o dresolve_tac discIs) THEN' etac notE THEN' atac)))));
  32.106 -
  32.107 -fun mk_primcorec_raw_code_of_ctr_tac ctxt distincts discIs splits split_asms ms ctr_thms =
  32.108 -  EVERY (map2 (mk_primcorec_raw_code_of_ctr_single_tac ctxt distincts discIs splits split_asms)
  32.109 -    ms ctr_thms);
  32.110 -
  32.111 -fun mk_primcorec_code_of_raw_code_tac splits raw =
  32.112 -  HEADGOAL (rtac raw ORELSE' rtac (raw RS trans) THEN' REPEAT_DETERM o
  32.113 -    (rtac refl ORELSE'
  32.114 -     (TRY o rtac sym) THEN' atac ORELSE'
  32.115 -     resolve_tac split_connectI ORELSE'
  32.116 -     Splitter.split_tac (split_if :: splits) ORELSE'
  32.117 -     etac notE THEN' atac));
  32.118 -
  32.119 -end;
    33.1 --- a/src/HOL/BNF/Tools/bnf_fp_rec_sugar_util.ML	Mon Nov 11 17:34:44 2013 +0100
    33.2 +++ b/src/HOL/BNF/Tools/bnf_fp_rec_sugar_util.ML	Mon Nov 11 17:44:21 2013 +0100
    33.3 @@ -8,410 +8,26 @@
    33.4  
    33.5  signature BNF_FP_REC_SUGAR_UTIL =
    33.6  sig
    33.7 -  datatype rec_call =
    33.8 -    No_Rec of int |
    33.9 -    Direct_Rec of int (*before*) * int (*after*) |
   33.10 -    Indirect_Rec of int
   33.11 +  val indexed: 'a list -> int -> int list * int
   33.12 +  val indexedd: 'a list list -> int -> int list list * int
   33.13 +  val indexeddd: ''a list list list -> int -> int list list list * int
   33.14 +  val indexedddd: 'a list list list list -> int -> int list list list list * int
   33.15 +  val find_index_eq: ''a list -> ''a -> int
   33.16 +  val finds: ('a * 'b -> bool) -> 'a list -> 'b list -> ('a * 'b list) list * 'b list
   33.17  
   33.18 -  datatype corec_call =
   33.19 -    Dummy_No_Corec of int |
   33.20 -    No_Corec of int |
   33.21 -    Direct_Corec of int (*stop?*) * int (*end*) * int (*continue*) |
   33.22 -    Indirect_Corec of int
   33.23 +  val drop_All: term -> term
   33.24  
   33.25 -  type rec_ctr_spec =
   33.26 -    {ctr: term,
   33.27 -     offset: int,
   33.28 -     calls: rec_call list,
   33.29 -     rec_thm: thm}
   33.30 +  val mk_partial_compN: int -> typ -> term -> term
   33.31 +  val mk_partial_comp: typ -> typ -> term -> term
   33.32 +  val mk_compN: int -> typ list -> term * term -> term
   33.33 +  val mk_comp: typ list -> term * term -> term
   33.34  
   33.35 -  type corec_ctr_spec =
   33.36 -    {ctr: term,
   33.37 -     disc: term,
   33.38 -     sels: term list,
   33.39 -     pred: int option,
   33.40 -     calls: corec_call list,
   33.41 -     discI: thm,
   33.42 -     sel_thms: thm list,
   33.43 -     collapse: thm,
   33.44 -     corec_thm: thm,
   33.45 -     disc_corec: thm,
   33.46 -     sel_corecs: thm list}
   33.47 -
   33.48 -  type rec_spec =
   33.49 -    {recx: term,
   33.50 -     nested_map_idents: thm list,
   33.51 -     nested_map_comps: thm list,
   33.52 -     ctr_specs: rec_ctr_spec list}
   33.53 -
   33.54 -  type corec_spec =
   33.55 -    {corec: term,
   33.56 -     nested_maps: thm list,
   33.57 -     nested_map_idents: thm list,
   33.58 -     nested_map_comps: thm list,
   33.59 -     ctr_specs: corec_ctr_spec list}
   33.60 -
   33.61 -  val s_not: term -> term
   33.62 -  val mk_conjs: term list -> term
   33.63 -  val mk_disjs: term list -> term
   33.64 -  val s_not_disj: term -> term list
   33.65 -  val negate_conj: term list -> term list
   33.66 -  val negate_disj: term list -> term list
   33.67 -
   33.68 -  val massage_indirect_rec_call: Proof.context -> (term -> bool) -> (typ -> typ -> term -> term) ->
   33.69 -    typ list -> term -> term -> term -> term
   33.70 -  val massage_direct_corec_call: Proof.context -> (term -> bool) -> (typ list -> term -> term) ->
   33.71 -    typ list -> term -> term
   33.72 -  val massage_indirect_corec_call: Proof.context -> (term -> bool) ->
   33.73 -    (typ list -> typ -> typ -> term -> term) -> typ list -> typ -> term -> term
   33.74 -  val expand_corec_code_rhs: Proof.context -> (term -> bool) -> typ list -> term -> term
   33.75 -  val massage_corec_code_rhs: Proof.context -> (typ list -> term -> term list -> term) ->
   33.76 -    typ list -> term -> term
   33.77 -  val fold_rev_corec_code_rhs: Proof.context -> (term list -> term -> term list -> 'a -> 'a) ->
   33.78 -    typ list -> term -> 'a -> 'a
   33.79 -  val case_thms_of_term: Proof.context -> typ list -> term ->
   33.80 -    thm list * thm list * thm list * thm list
   33.81 -
   33.82 -  val rec_specs_of: binding list -> typ list -> typ list -> (term -> int list) ->
   33.83 -    ((term * term list list) list) list -> local_theory ->
   33.84 -    (bool * rec_spec list * typ list * thm * thm list) * local_theory
   33.85 -  val corec_specs_of: binding list -> typ list -> typ list -> (term -> int list) ->
   33.86 -    ((term * term list list) list) list -> local_theory ->
   33.87 -    (bool * corec_spec list * typ list * thm * thm * thm list * thm list) * local_theory
   33.88 +  val get_indices: ((binding * typ) * 'a) list -> term -> int list
   33.89  end;
   33.90  
   33.91  structure BNF_FP_Rec_Sugar_Util : BNF_FP_REC_SUGAR_UTIL =
   33.92  struct
   33.93  
   33.94 -open Ctr_Sugar
   33.95 -open BNF_Util
   33.96 -open BNF_Def
   33.97 -open BNF_FP_Util
   33.98 -open BNF_FP_Def_Sugar
   33.99 -open BNF_FP_N2M_Sugar
  33.100 -
  33.101 -datatype rec_call =
  33.102 -  No_Rec of int |
  33.103 -  Direct_Rec of int * int |
  33.104 -  Indirect_Rec of int;
  33.105 -
  33.106 -datatype corec_call =
  33.107 -  Dummy_No_Corec of int |
  33.108 -  No_Corec of int |
  33.109 -  Direct_Corec of int * int * int |
  33.110 -  Indirect_Corec of int;
  33.111 -
  33.112 -type rec_ctr_spec =
  33.113 -  {ctr: term,
  33.114 -   offset: int,
  33.115 -   calls: rec_call list,
  33.116 -   rec_thm: thm};
  33.117 -
  33.118 -type corec_ctr_spec =
  33.119 -  {ctr: term,
  33.120 -   disc: term,
  33.121 -   sels: term list,
  33.122 -   pred: int option,
  33.123 -   calls: corec_call list,
  33.124 -   discI: thm,
  33.125 -   sel_thms: thm list,
  33.126 -   collapse: thm,
  33.127 -   corec_thm: thm,
  33.128 -   disc_corec: thm,
  33.129 -   sel_corecs: thm list};
  33.130 -
  33.131 -type rec_spec =
  33.132 -  {recx: term,
  33.133 -   nested_map_idents: thm list,
  33.134 -   nested_map_comps: thm list,
  33.135 -   ctr_specs: rec_ctr_spec list};
  33.136 -
  33.137 -type corec_spec =
  33.138 -  {corec: term,
  33.139 -   nested_maps: thm list,
  33.140 -   nested_map_idents: thm list,
  33.141 -   nested_map_comps: thm list,
  33.142 -   ctr_specs: corec_ctr_spec list};
  33.143 -
  33.144 -val id_def = @{thm id_def};
  33.145 -
  33.146 -exception AINT_NO_MAP of term;
  33.147 -
  33.148 -fun ill_formed_rec_call ctxt t =
  33.149 -  error ("Ill-formed recursive call: " ^ quote (Syntax.string_of_term ctxt t));
  33.150 -fun ill_formed_corec_call ctxt t =
  33.151 -  error ("Ill-formed corecursive call: " ^ quote (Syntax.string_of_term ctxt t));
  33.152 -fun invalid_map ctxt t =
  33.153 -  error ("Invalid map function in " ^ quote (Syntax.string_of_term ctxt t));
  33.154 -fun unexpected_rec_call ctxt t =
  33.155 -  error ("Unexpected recursive call: " ^ quote (Syntax.string_of_term ctxt t));
  33.156 -fun unexpected_corec_call ctxt t =
  33.157 -  error ("Unexpected corecursive call: " ^ quote (Syntax.string_of_term ctxt t));
  33.158 -
  33.159 -fun s_not @{const True} = @{const False}
  33.160 -  | s_not @{const False} = @{const True}
  33.161 -  | s_not (@{const Not} $ t) = t
  33.162 -  | s_not t = HOLogic.mk_not t
  33.163 -
  33.164 -val mk_conjs = try (foldr1 HOLogic.mk_conj) #> the_default @{const True};
  33.165 -val mk_disjs = try (foldr1 HOLogic.mk_disj) #> the_default @{const False};
  33.166 -
  33.167 -val s_not_disj = map s_not o HOLogic.disjuncts;
  33.168 -
  33.169 -fun negate_conj [t] = s_not_disj t
  33.170 -  | negate_conj ts = [mk_disjs (map s_not ts)];
  33.171 -
  33.172 -fun negate_disj [t] = s_not_disj t
  33.173 -  | negate_disj ts = [mk_disjs (map (mk_conjs o s_not_disj) ts)];
  33.174 -
  33.175 -fun factor_out_types ctxt massage destU U T =
  33.176 -  (case try destU U of
  33.177 -    SOME (U1, U2) => if U1 = T then massage T U2 else invalid_map ctxt
  33.178 -  | NONE => invalid_map ctxt);
  33.179 -
  33.180 -fun map_flattened_map_args ctxt s map_args fs =
  33.181 -  let
  33.182 -    val flat_fs = flatten_type_args_of_bnf (the (bnf_of ctxt s)) Term.dummy fs;
  33.183 -    val flat_fs' = map_args flat_fs;
  33.184 -  in
  33.185 -    permute_like (op aconv) flat_fs fs flat_fs'
  33.186 -  end;
  33.187 -
  33.188 -fun massage_indirect_rec_call ctxt has_call raw_massage_fun bound_Ts y y' =
  33.189 -  let
  33.190 -    val typof = curry fastype_of1 bound_Ts;
  33.191 -    val build_map_fst = build_map ctxt (fst_const o fst);
  33.192 -
  33.193 -    val yT = typof y;
  33.194 -    val yU = typof y';
  33.195 -
  33.196 -    fun y_of_y' () = build_map_fst (yU, yT) $ y';
  33.197 -    val elim_y = Term.map_aterms (fn t => if t = y then y_of_y' () else t);
  33.198 -
  33.199 -    fun massage_direct_fun U T t =
  33.200 -      if has_call t then factor_out_types ctxt raw_massage_fun HOLogic.dest_prodT U T t
  33.201 -      else HOLogic.mk_comp (t, build_map_fst (U, T));
  33.202 -
  33.203 -    fun massage_map (Type (_, Us)) (Type (s, Ts)) t =
  33.204 -        (case try (dest_map ctxt s) t of
  33.205 -          SOME (map0, fs) =>
  33.206 -          let
  33.207 -            val Type (_, ran_Ts) = range_type (typof t);
  33.208 -            val map' = mk_map (length fs) Us ran_Ts map0;
  33.209 -            val fs' = map_flattened_map_args ctxt s (map3 massage_map_or_map_arg Us Ts) fs;
  33.210 -          in
  33.211 -            Term.list_comb (map', fs')
  33.212 -          end
  33.213 -        | NONE => raise AINT_NO_MAP t)
  33.214 -      | massage_map _ _ t = raise AINT_NO_MAP t
  33.215 -    and massage_map_or_map_arg U T t =
  33.216 -      if T = U then
  33.217 -        if has_call t then unexpected_rec_call ctxt t else t
  33.218 -      else
  33.219 -        massage_map U T t
  33.220 -        handle AINT_NO_MAP _ => massage_direct_fun U T t;
  33.221 -
  33.222 -    fun massage_call (t as t1 $ t2) =
  33.223 -        if t2 = y then
  33.224 -          massage_map yU yT (elim_y t1) $ y'
  33.225 -          handle AINT_NO_MAP t' => invalid_map ctxt t'
  33.226 -        else
  33.227 -          ill_formed_rec_call ctxt t
  33.228 -      | massage_call t = if t = y then y_of_y' () else ill_formed_rec_call ctxt t;
  33.229 -  in
  33.230 -    massage_call
  33.231 -  end;
  33.232 -
  33.233 -fun fold_rev_let_if_case ctxt f bound_Ts t =
  33.234 -  let
  33.235 -    val thy = Proof_Context.theory_of ctxt;
  33.236 -
  33.237 -    fun fld conds t =
  33.238 -      (case Term.strip_comb t of
  33.239 -        (Const (@{const_name Let}, _), [arg1, arg2]) => fld conds (betapply (arg2, arg1))
  33.240 -      | (Const (@{const_name If}, _), [cond, then_branch, else_branch]) =>
  33.241 -        fld (conds @ HOLogic.conjuncts cond) then_branch
  33.242 -        o fld (conds @ s_not_disj cond) else_branch
  33.243 -      | (Const (c, _), args as _ :: _ :: _) =>
  33.244 -        let val n = num_binder_types (Sign.the_const_type thy c) - 1 in
  33.245 -          if n >= 0 andalso n < length args then
  33.246 -            (case fastype_of1 (bound_Ts, nth args n) of
  33.247 -              Type (s, Ts) =>
  33.248 -              (case dest_case ctxt s Ts t of
  33.249 -                NONE => apsnd (f conds t)
  33.250 -              | SOME (conds', branches) =>
  33.251 -                apfst (cons s) o fold_rev (uncurry fld)
  33.252 -                  (map (append conds o HOLogic.conjuncts) conds' ~~ branches))
  33.253 -            | _ => apsnd (f conds t))
  33.254 -          else
  33.255 -            apsnd (f conds t)
  33.256 -        end
  33.257 -      | _ => apsnd (f conds t))
  33.258 -  in
  33.259 -    fld [] t o pair []
  33.260 -  end;
  33.261 -
  33.262 -fun case_of ctxt = ctr_sugar_of ctxt #> Option.map (fst o dest_Const o #casex);
  33.263 -
  33.264 -fun massage_let_if_case ctxt has_call massage_leaf =
  33.265 -  let
  33.266 -    val thy = Proof_Context.theory_of ctxt;
  33.267 -
  33.268 -    fun check_no_call t = if has_call t then unexpected_corec_call ctxt t else ();
  33.269 -
  33.270 -    fun massage_abs bound_Ts (Abs (s, T, t)) = Abs (s, T, massage_abs (T :: bound_Ts) t)
  33.271 -      | massage_abs bound_Ts t = massage_rec bound_Ts t
  33.272 -    and massage_rec bound_Ts t =
  33.273 -      let val typof = curry fastype_of1 bound_Ts in
  33.274 -        (case Term.strip_comb t of
  33.275 -          (Const (@{const_name Let}, _), [arg1, arg2]) =>
  33.276 -          massage_rec bound_Ts (betapply (arg2, arg1))
  33.277 -        | (Const (@{const_name If}, _), obj :: (branches as [_, _])) =>
  33.278 -          let val branches' = map (massage_rec bound_Ts) branches in
  33.279 -            Term.list_comb (If_const (typof (hd branches')) $ tap check_no_call obj, branches')
  33.280 -          end
  33.281 -        | (Const (c, _), args as _ :: _ :: _) =>
  33.282 -          let
  33.283 -            val gen_T = Sign.the_const_type thy c;
  33.284 -            val (gen_branch_Ts, gen_body_fun_T) = strip_fun_type gen_T;
  33.285 -            val n = length gen_branch_Ts;
  33.286 -          in
  33.287 -            if n < length args then
  33.288 -              (case gen_body_fun_T of
  33.289 -                Type (_, [Type (T_name, _), _]) =>
  33.290 -                if case_of ctxt T_name = SOME c then
  33.291 -                  let
  33.292 -                    val (branches, obj_leftovers) = chop n args;
  33.293 -                    val branches' = map (massage_abs bound_Ts o Envir.eta_long bound_Ts) branches;
  33.294 -                    val branch_Ts' = map typof branches';
  33.295 -                    val casex' = Const (c, branch_Ts' ---> map typof obj_leftovers --->
  33.296 -                      snd (strip_typeN (num_binder_types (hd gen_branch_Ts)) (hd branch_Ts')));
  33.297 -                  in
  33.298 -                    Term.list_comb (casex', branches' @ tap (List.app check_no_call) obj_leftovers)
  33.299 -                  end
  33.300 -                else
  33.301 -                  massage_leaf bound_Ts t
  33.302 -              | _ => massage_leaf bound_Ts t)
  33.303 -            else
  33.304 -              massage_leaf bound_Ts t
  33.305 -          end
  33.306 -        | _ => massage_leaf bound_Ts t)
  33.307 -      end
  33.308 -  in
  33.309 -    massage_rec
  33.310 -  end;
  33.311 -
  33.312 -val massage_direct_corec_call = massage_let_if_case;
  33.313 -
  33.314 -fun curried_type (Type (@{type_name fun}, [Type (@{type_name prod}, Ts), T])) = Ts ---> T;
  33.315 -
  33.316 -fun massage_indirect_corec_call ctxt has_call raw_massage_call bound_Ts U t =
  33.317 -  let
  33.318 -    val build_map_Inl = build_map ctxt (uncurry Inl_const o dest_sumT o snd)
  33.319 -
  33.320 -    fun massage_direct_call bound_Ts U T t =
  33.321 -      if has_call t then factor_out_types ctxt (raw_massage_call bound_Ts) dest_sumT U T t
  33.322 -      else build_map_Inl (T, U) $ t;
  33.323 -
  33.324 -    fun massage_direct_fun bound_Ts U T t =
  33.325 -      let
  33.326 -        val var = Var ((Name.uu, Term.maxidx_of_term t + 1),
  33.327 -          domain_type (fastype_of1 (bound_Ts, t)));
  33.328 -      in
  33.329 -        Term.lambda var (massage_direct_call bound_Ts U T (t $ var))
  33.330 -      end;
  33.331 -
  33.332 -    fun massage_map bound_Ts (Type (_, Us)) (Type (s, Ts)) t =
  33.333 -        (case try (dest_map ctxt s) t of
  33.334 -          SOME (map0, fs) =>
  33.335 -          let
  33.336 -            val Type (_, dom_Ts) = domain_type (fastype_of1 (bound_Ts, t));
  33.337 -            val map' = mk_map (length fs) dom_Ts Us map0;
  33.338 -            val fs' =
  33.339 -              map_flattened_map_args ctxt s (map3 (massage_map_or_map_arg bound_Ts) Us Ts) fs;
  33.340 -          in
  33.341 -            Term.list_comb (map', fs')
  33.342 -          end
  33.343 -        | NONE => raise AINT_NO_MAP t)
  33.344 -      | massage_map _ _ _ t = raise AINT_NO_MAP t
  33.345 -    and massage_map_or_map_arg bound_Ts U T t =
  33.346 -      if T = U then
  33.347 -        if has_call t then unexpected_corec_call ctxt t else t
  33.348 -      else
  33.349 -        massage_map bound_Ts U T t
  33.350 -        handle AINT_NO_MAP _ => massage_direct_fun bound_Ts U T t;
  33.351 -
  33.352 -    fun massage_call bound_Ts U T =
  33.353 -      massage_let_if_case ctxt has_call (fn bound_Ts => fn t =>
  33.354 -        if has_call t then
  33.355 -          (case U of
  33.356 -            Type (s, Us) =>
  33.357 -            (case try (dest_ctr ctxt s) t of
  33.358 -              SOME (f, args) =>
  33.359 -              let
  33.360 -                val typof = curry fastype_of1 bound_Ts;
  33.361 -                val f' = mk_ctr Us f
  33.362 -                val f'_T = typof f';
  33.363 -                val arg_Ts = map typof args;
  33.364 -              in
  33.365 -                Term.list_comb (f', map3 (massage_call bound_Ts) (binder_types f'_T) arg_Ts args)
  33.366 -              end
  33.367 -            | NONE =>
  33.368 -              (case t of
  33.369 -                Const (@{const_name prod_case}, _) $ t' =>
  33.370 -                let
  33.371 -                  val U' = curried_type U;
  33.372 -                  val T' = curried_type T;
  33.373 -                in
  33.374 -                  Const (@{const_name prod_case}, U' --> U) $ massage_call bound_Ts U' T' t'
  33.375 -                end
  33.376 -              | t1 $ t2 =>
  33.377 -                (if has_call t2 then
  33.378 -                  massage_direct_call bound_Ts U T t
  33.379 -                else
  33.380 -                  massage_map bound_Ts U T t1 $ t2
  33.381 -                  handle AINT_NO_MAP _ => massage_direct_call bound_Ts U T t)
  33.382 -              | Abs (s, T', t') =>
  33.383 -                Abs (s, T', massage_call (T' :: bound_Ts) (range_type U) (range_type T) t')
  33.384 -              | _ => massage_direct_call bound_Ts U T t))
  33.385 -          | _ => ill_formed_corec_call ctxt t)
  33.386 -        else
  33.387 -          build_map_Inl (T, U) $ t) bound_Ts;
  33.388 -
  33.389 -    val T = fastype_of1 (bound_Ts, t);
  33.390 -  in
  33.391 -    if has_call t then massage_call bound_Ts U T t else build_map_Inl (T, U) $ t
  33.392 -  end;
  33.393 -
  33.394 -fun expand_ctr_term ctxt s Ts t =
  33.395 -  (case ctr_sugar_of ctxt s of
  33.396 -    SOME {ctrs, casex, ...} =>
  33.397 -    Term.list_comb (mk_case Ts (Type (s, Ts)) casex, map (mk_ctr Ts) ctrs) $ t
  33.398 -  | NONE => raise Fail "expand_ctr_term");
  33.399 -
  33.400 -fun expand_corec_code_rhs ctxt has_call bound_Ts t =
  33.401 -  (case fastype_of1 (bound_Ts, t) of
  33.402 -    Type (s, Ts) =>
  33.403 -    massage_let_if_case ctxt has_call (fn _ => fn t =>
  33.404 -      if can (dest_ctr ctxt s) t then t else expand_ctr_term ctxt s Ts t) bound_Ts t
  33.405 -  | _ => raise Fail "expand_corec_code_rhs");
  33.406 -
  33.407 -fun massage_corec_code_rhs ctxt massage_ctr =
  33.408 -  massage_let_if_case ctxt (K false)
  33.409 -    (fn bound_Ts => uncurry (massage_ctr bound_Ts) o Term.strip_comb);
  33.410 -
  33.411 -fun fold_rev_corec_code_rhs ctxt f =
  33.412 -  snd ooo fold_rev_let_if_case ctxt (fn conds => uncurry (f conds) o Term.strip_comb);
  33.413 -
  33.414 -fun case_thms_of_term ctxt bound_Ts t =
  33.415 -  let
  33.416 -    val (caseT_names, _) = fold_rev_let_if_case ctxt (K (K I)) bound_Ts t ();
  33.417 -    val ctr_sugars = map (the o ctr_sugar_of ctxt) caseT_names;
  33.418 -  in
  33.419 -    (maps #distincts ctr_sugars, maps #discIs ctr_sugars, maps #sel_splits ctr_sugars,
  33.420 -     maps #sel_split_asms ctr_sugars)
  33.421 -  end;
  33.422 -
  33.423  fun indexed xs h = let val h' = h + length xs in (h upto h' - 1, h') end;
  33.424  fun indexedd xss = fold_map indexed xss;
  33.425  fun indexeddd xsss = fold_map indexedd xsss;
  33.426 @@ -419,205 +35,32 @@
  33.427  
  33.428  fun find_index_eq hs h = find_index (curry (op =) h) hs;
  33.429  
  33.430 -(*FIXME: remove special cases for products and sum once they are registered as datatypes*)
  33.431 -fun map_thms_of_typ ctxt (Type (s, _)) =
  33.432 -    if s = @{type_name prod} then
  33.433 -      @{thms map_pair_simp}
  33.434 -    else if s = @{type_name sum} then
  33.435 -      @{thms sum_map.simps}
  33.436 -    else
  33.437 -      (case fp_sugar_of ctxt s of
  33.438 -        SOME {index, mapss, ...} => nth mapss index
  33.439 -      | NONE => [])
  33.440 -  | map_thms_of_typ _ _ = [];
  33.441 +fun finds eq = fold_map (fn x => List.partition (curry eq x) #>> pair x);
  33.442  
  33.443 -fun rec_specs_of bs arg_Ts res_Ts get_indices callssss0 lthy =
  33.444 -  let
  33.445 -    val thy = Proof_Context.theory_of lthy;
  33.446 +fun drop_All t =
  33.447 +  subst_bounds (strip_qnt_vars @{const_name all} t |> map Free |> rev,
  33.448 +    strip_qnt_body @{const_name all} t);
  33.449  
  33.450 -    val ((missing_arg_Ts, perm0_kks,
  33.451 -          fp_sugars as {nested_bnfs, fp_res = {xtor_co_iterss = ctor_iters1 :: _, ...},
  33.452 -            co_inducts = [induct_thm], ...} :: _, (lfp_sugar_thms, _)), lthy') =
  33.453 -      nested_to_mutual_fps Least_FP bs arg_Ts get_indices callssss0 lthy;
  33.454 -
  33.455 -    val perm_fp_sugars = sort (int_ord o pairself #index) fp_sugars;
  33.456 -
  33.457 -    val indices = map #index fp_sugars;
  33.458 -    val perm_indices = map #index perm_fp_sugars;
  33.459 -
  33.460 -    val perm_ctrss = map (#ctrs o of_fp_sugar #ctr_sugars) perm_fp_sugars;
  33.461 -    val perm_ctr_Tsss = map (map (binder_types o fastype_of)) perm_ctrss;
  33.462 -    val perm_fpTs = map (body_type o fastype_of o hd) perm_ctrss;
  33.463 -
  33.464 -    val nn0 = length arg_Ts;
  33.465 -    val nn = length perm_fpTs;
  33.466 -    val kks = 0 upto nn - 1;
  33.467 -    val perm_ns = map length perm_ctr_Tsss;
  33.468 -    val perm_mss = map (map length) perm_ctr_Tsss;
  33.469 -
  33.470 -    val perm_Cs = map (body_type o fastype_of o co_rec_of o of_fp_sugar (#xtor_co_iterss o #fp_res))
  33.471 -      perm_fp_sugars;
  33.472 -    val perm_fun_arg_Tssss =
  33.473 -      mk_iter_fun_arg_types perm_ctr_Tsss perm_ns perm_mss (co_rec_of ctor_iters1);
  33.474 -
  33.475 -    fun unpermute0 perm0_xs = permute_like (op =) perm0_kks kks perm0_xs;
  33.476 -    fun unpermute perm_xs = permute_like (op =) perm_indices indices perm_xs;
  33.477 -
  33.478 -    val induct_thms = unpermute0 (conj_dests nn induct_thm);
  33.479 -
  33.480 -    val fpTs = unpermute perm_fpTs;
  33.481 -    val Cs = unpermute perm_Cs;
  33.482 -
  33.483 -    val As_rho = tvar_subst thy (take nn0 fpTs) arg_Ts;
  33.484 -    val Cs_rho = map (fst o dest_TVar) Cs ~~ pad_list HOLogic.unitT nn res_Ts;
  33.485 -
  33.486 -    val substA = Term.subst_TVars As_rho;
  33.487 -    val substAT = Term.typ_subst_TVars As_rho;
  33.488 -    val substCT = Term.typ_subst_TVars Cs_rho;
  33.489 -
  33.490 -    val perm_Cs' = map substCT perm_Cs;
  33.491 -
  33.492 -    fun offset_of_ctr 0 _ = 0
  33.493 -      | offset_of_ctr n (({ctrs, ...} : ctr_sugar) :: ctr_sugars) =
  33.494 -        length ctrs + offset_of_ctr (n - 1) ctr_sugars;
  33.495 -
  33.496 -    fun call_of [i] [T] = (if exists_subtype_in Cs T then Indirect_Rec else No_Rec) i
  33.497 -      | call_of [i, i'] _ = Direct_Rec (i, i');
  33.498 -
  33.499 -    fun mk_ctr_spec ctr offset fun_arg_Tss rec_thm =
  33.500 -      let
  33.501 -        val (fun_arg_hss, _) = indexedd fun_arg_Tss 0;
  33.502 -        val fun_arg_hs = flat_rec_arg_args fun_arg_hss;
  33.503 -        val fun_arg_iss = map (map (find_index_eq fun_arg_hs)) fun_arg_hss;
  33.504 -      in
  33.505 -        {ctr = substA ctr, offset = offset, calls = map2 call_of fun_arg_iss fun_arg_Tss,
  33.506 -         rec_thm = rec_thm}
  33.507 -      end;
  33.508 -
  33.509 -    fun mk_ctr_specs index (ctr_sugars : ctr_sugar list) iter_thmsss =
  33.510 -      let
  33.511 -        val ctrs = #ctrs (nth ctr_sugars index);
  33.512 -        val rec_thmss = co_rec_of (nth iter_thmsss index);
  33.513 -        val k = offset_of_ctr index ctr_sugars;
  33.514 -        val n = length ctrs;
  33.515 -      in
  33.516 -        map4 mk_ctr_spec ctrs (k upto k + n - 1) (nth perm_fun_arg_Tssss index) rec_thmss
  33.517 -      end;
  33.518 -
  33.519 -    fun mk_spec ({T, index, ctr_sugars, co_iterss = iterss, co_iter_thmsss = iter_thmsss, ...}
  33.520 -      : fp_sugar) =
  33.521 -      {recx = mk_co_iter thy Least_FP (substAT T) perm_Cs' (co_rec_of (nth iterss index)),
  33.522 -       nested_map_idents = map (unfold_thms lthy [id_def] o map_id0_of_bnf) nested_bnfs,
  33.523 -       nested_map_comps = map map_comp_of_bnf nested_bnfs,
  33.524 -       ctr_specs = mk_ctr_specs index ctr_sugars iter_thmsss};
  33.525 -  in
  33.526 -    ((is_some lfp_sugar_thms, map mk_spec fp_sugars, missing_arg_Ts, induct_thm, induct_thms),
  33.527 -     lthy')
  33.528 +fun mk_partial_comp gT fT g =
  33.529 +  let val T = domain_type fT --> range_type gT in
  33.530 +    Const (@{const_name Fun.comp}, gT --> fT --> T) $ g
  33.531    end;
  33.532  
  33.533 -fun corec_specs_of bs arg_Ts res_Ts get_indices callssss0 lthy =
  33.534 -  let
  33.535 -    val thy = Proof_Context.theory_of lthy;
  33.536 +fun mk_partial_compN 0 _ g = g
  33.537 +  | mk_partial_compN n fT g =
  33.538 +    let val g' = mk_partial_compN (n - 1) (range_type fT) g in
  33.539 +      mk_partial_comp (fastype_of g') fT g'
  33.540 +    end;
  33.541  
  33.542 -    val ((missing_res_Ts, perm0_kks,
  33.543 -          fp_sugars as {nested_bnfs, fp_res = {xtor_co_iterss = dtor_coiters1 :: _, ...},
  33.544 -            co_inducts = coinduct_thms, ...} :: _, (_, gfp_sugar_thms)), lthy') =
  33.545 -      nested_to_mutual_fps Greatest_FP bs res_Ts get_indices callssss0 lthy;
  33.546 -
  33.547 -    val perm_fp_sugars = sort (int_ord o pairself #index) fp_sugars;
  33.548 -
  33.549 -    val indices = map #index fp_sugars;
  33.550 -    val perm_indices = map #index perm_fp_sugars;
  33.551 -
  33.552 -    val perm_ctrss = map (#ctrs o of_fp_sugar #ctr_sugars) perm_fp_sugars;
  33.553 -    val perm_ctr_Tsss = map (map (binder_types o fastype_of)) perm_ctrss;
  33.554 -    val perm_fpTs = map (body_type o fastype_of o hd) perm_ctrss;
  33.555 -
  33.556 -    val nn0 = length res_Ts;
  33.557 -    val nn = length perm_fpTs;
  33.558 -    val kks = 0 upto nn - 1;
  33.559 -    val perm_ns = map length perm_ctr_Tsss;
  33.560 -
  33.561 -    val perm_Cs = map (domain_type o body_fun_type o fastype_of o co_rec_of o
  33.562 -      of_fp_sugar (#xtor_co_iterss o #fp_res)) perm_fp_sugars;
  33.563 -    val (perm_p_Tss, (perm_q_Tssss, _, perm_f_Tssss, _)) =
  33.564 -      mk_coiter_fun_arg_types perm_ctr_Tsss perm_Cs perm_ns (co_rec_of dtor_coiters1);
  33.565 -
  33.566 -    val (perm_p_hss, h) = indexedd perm_p_Tss 0;
  33.567 -    val (perm_q_hssss, h') = indexedddd perm_q_Tssss h;
  33.568 -    val (perm_f_hssss, _) = indexedddd perm_f_Tssss h';
  33.569 -
  33.570 -    val fun_arg_hs =
  33.571 -      flat (map3 flat_corec_preds_predsss_gettersss perm_p_hss perm_q_hssss perm_f_hssss);
  33.572 -
  33.573 -    fun unpermute0 perm0_xs = permute_like (op =) perm0_kks kks perm0_xs;
  33.574 -    fun unpermute perm_xs = permute_like (op =) perm_indices indices perm_xs;
  33.575 -
  33.576 -    val coinduct_thmss = map (unpermute0 o conj_dests nn) coinduct_thms;
  33.577 -
  33.578 -    val p_iss = map (map (find_index_eq fun_arg_hs)) (unpermute perm_p_hss);
  33.579 -    val q_issss = map (map (map (map (find_index_eq fun_arg_hs)))) (unpermute perm_q_hssss);
  33.580 -    val f_issss = map (map (map (map (find_index_eq fun_arg_hs)))) (unpermute perm_f_hssss);
  33.581 -
  33.582 -    val f_Tssss = unpermute perm_f_Tssss;
  33.583 -    val fpTs = unpermute perm_fpTs;
  33.584 -    val Cs = unpermute perm_Cs;
  33.585 -
  33.586 -    val As_rho = tvar_subst thy (take nn0 fpTs) res_Ts;
  33.587 -    val Cs_rho = map (fst o dest_TVar) Cs ~~ pad_list HOLogic.unitT nn arg_Ts;
  33.588 -
  33.589 -    val substA = Term.subst_TVars As_rho;
  33.590 -    val substAT = Term.typ_subst_TVars As_rho;
  33.591 -    val substCT = Term.typ_subst_TVars Cs_rho;
  33.592 -
  33.593 -    val perm_Cs' = map substCT perm_Cs;
  33.594 -
  33.595 -    fun call_of nullary [] [g_i] [Type (@{type_name fun}, [_, T])] =
  33.596 -        (if exists_subtype_in Cs T then Indirect_Corec
  33.597 -         else if nullary then Dummy_No_Corec
  33.598 -         else No_Corec) g_i
  33.599 -      | call_of _ [q_i] [g_i, g_i'] _ = Direct_Corec (q_i, g_i, g_i');
  33.600 -
  33.601 -    fun mk_ctr_spec ctr disc sels p_ho q_iss f_iss f_Tss discI sel_thms collapse corec_thm
  33.602 -        disc_corec sel_corecs =
  33.603 -      let val nullary = not (can dest_funT (fastype_of ctr)) in
  33.604 -        {ctr = substA ctr, disc = substA disc, sels = map substA sels, pred = p_ho,
  33.605 -         calls = map3 (call_of nullary) q_iss f_iss f_Tss, discI = discI, sel_thms = sel_thms,
  33.606 -         collapse = collapse, corec_thm = corec_thm, disc_corec = disc_corec,
  33.607 -         sel_corecs = sel_corecs}
  33.608 -      end;
  33.609 -
  33.610 -    fun mk_ctr_specs index (ctr_sugars : ctr_sugar list) p_is q_isss f_isss f_Tsss
  33.611 -        coiter_thmsss disc_coitersss sel_coiterssss =
  33.612 -      let
  33.613 -        val ctrs = #ctrs (nth ctr_sugars index);
  33.614 -        val discs = #discs (nth ctr_sugars index);
  33.615 -        val selss = #selss (nth ctr_sugars index);
  33.616 -        val p_ios = map SOME p_is @ [NONE];
  33.617 -        val discIs = #discIs (nth ctr_sugars index);
  33.618 -        val sel_thmss = #sel_thmss (nth ctr_sugars index);
  33.619 -        val collapses = #collapses (nth ctr_sugars index);
  33.620 -        val corec_thms = co_rec_of (nth coiter_thmsss index);
  33.621 -        val disc_corecs = co_rec_of (nth disc_coitersss index);
  33.622 -        val sel_corecss = co_rec_of (nth sel_coiterssss index);
  33.623 -      in
  33.624 -        map13 mk_ctr_spec ctrs discs selss p_ios q_isss f_isss f_Tsss discIs sel_thmss collapses
  33.625 -          corec_thms disc_corecs sel_corecss
  33.626 -      end;
  33.627 -
  33.628 -    fun mk_spec ({T, index, ctr_sugars, co_iterss = coiterss, co_iter_thmsss = coiter_thmsss,
  33.629 -          disc_co_itersss = disc_coitersss, sel_co_iterssss = sel_coiterssss, ...} : fp_sugar)
  33.630 -        p_is q_isss f_isss f_Tsss =
  33.631 -      {corec = mk_co_iter thy Greatest_FP (substAT T) perm_Cs' (co_rec_of (nth coiterss index)),
  33.632 -       nested_maps = maps (map_thms_of_typ lthy o T_of_bnf) nested_bnfs,
  33.633 -       nested_map_idents = map (unfold_thms lthy [id_def] o map_id0_of_bnf) nested_bnfs,
  33.634 -       nested_map_comps = map map_comp_of_bnf nested_bnfs,
  33.635 -       ctr_specs = mk_ctr_specs index ctr_sugars p_is q_isss f_isss f_Tsss coiter_thmsss
  33.636 -         disc_coitersss sel_coiterssss};
  33.637 -  in
  33.638 -    ((is_some gfp_sugar_thms, map5 mk_spec fp_sugars p_iss q_issss f_issss f_Tssss, missing_res_Ts,
  33.639 -      co_induct_of coinduct_thms, strong_co_induct_of coinduct_thms, co_induct_of coinduct_thmss,
  33.640 -      strong_co_induct_of coinduct_thmss), lthy')
  33.641 +fun mk_compN n bound_Ts (g, f) =
  33.642 +  let val typof = curry fastype_of1 bound_Ts in
  33.643 +    mk_partial_compN n (typof f) g $ f
  33.644    end;
  33.645  
  33.646 +val mk_comp = mk_compN 1;
  33.647 +
  33.648 +fun get_indices fixes t = map (fst #>> Binding.name_of #> Free) fixes
  33.649 +  |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE)
  33.650 +  |> map_filter I;
  33.651 +
  33.652  end;
    34.1 --- a/src/HOL/BNF/Tools/bnf_gfp.ML	Mon Nov 11 17:34:44 2013 +0100
    34.2 +++ b/src/HOL/BNF/Tools/bnf_gfp.ML	Mon Nov 11 17:44:21 2013 +0100
    34.3 @@ -23,7 +23,7 @@
    34.4  open BNF_Comp
    34.5  open BNF_FP_Util
    34.6  open BNF_FP_Def_Sugar
    34.7 -open BNF_FP_Rec_Sugar
    34.8 +open BNF_GFP_Rec_Sugar
    34.9  open BNF_GFP_Util
   34.10  open BNF_GFP_Tactics
   34.11  
   34.12 @@ -2744,8 +2744,8 @@
   34.13                ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
   34.14              bs thmss)
   34.15        in
   34.16 -       (timer, Jbnfs, (folded_dtor_map_o_thms, folded_dtor_map_thms), folded_dtor_set_thmss',
   34.17 -         dtor_set_induct_thms, dtor_Jrel_thms, Jbnf_common_notes @ Jbnf_notes, lthy)
   34.18 +        (timer, Jbnfs, (folded_dtor_map_o_thms, folded_dtor_map_thms), folded_dtor_set_thmss',
   34.19 +          dtor_set_induct_thms, dtor_Jrel_thms, Jbnf_common_notes @ Jbnf_notes, lthy)
   34.20        end;
   34.21  
   34.22        val dtor_unfold_o_map_thms = mk_xtor_un_fold_o_map_thms Greatest_FP false m
    35.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    35.2 +++ b/src/HOL/BNF/Tools/bnf_gfp_rec_sugar.ML	Mon Nov 11 17:44:21 2013 +0100
    35.3 @@ -0,0 +1,1150 @@
    35.4 +(*  Title:      HOL/BNF/Tools/bnf_gfp_rec_sugar.ML
    35.5 +    Author:     Lorenz Panny, TU Muenchen
    35.6 +    Author:     Jasmin Blanchette, TU Muenchen
    35.7 +    Copyright   2013
    35.8 +
    35.9 +Corecursor sugar.
   35.10 +*)
   35.11 +
   35.12 +signature BNF_GFP_REC_SUGAR =
   35.13 +sig
   35.14 +  val add_primcorecursive_cmd: bool ->
   35.15 +    (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
   35.16 +    Proof.context -> Proof.state
   35.17 +  val add_primcorec_cmd: bool ->
   35.18 +    (binding * string option * mixfix) list * ((Attrib.binding * string) * string option) list ->
   35.19 +    local_theory -> local_theory
   35.20 +end;
   35.21 +
   35.22 +structure BNF_GFP_Rec_Sugar : BNF_GFP_REC_SUGAR =
   35.23 +struct
   35.24 +
   35.25 +open Ctr_Sugar
   35.26 +open BNF_Util
   35.27 +open BNF_Def
   35.28 +open BNF_FP_Util
   35.29 +open BNF_FP_Def_Sugar
   35.30 +open BNF_FP_N2M_Sugar
   35.31 +open BNF_FP_Rec_Sugar_Util
   35.32 +open BNF_GFP_Rec_Sugar_Tactics
   35.33 +
   35.34 +val codeN = "code"
   35.35 +val ctrN = "ctr"
   35.36 +val discN = "disc"
   35.37 +val selN = "sel"
   35.38 +
   35.39 +val nitpicksimp_attrs = @{attributes [nitpick_simp]};
   35.40 +val simp_attrs = @{attributes [simp]};
   35.41 +val code_nitpicksimp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs;
   35.42 +
   35.43 +exception Primcorec_Error of string * term list;
   35.44 +
   35.45 +fun primcorec_error str = raise Primcorec_Error (str, []);
   35.46 +fun primcorec_error_eqn str eqn = raise Primcorec_Error (str, [eqn]);
   35.47 +fun primcorec_error_eqns str eqns = raise Primcorec_Error (str, eqns);
   35.48 +
   35.49 +datatype corec_call =
   35.50 +  Dummy_No_Corec of int |
   35.51 +  No_Corec of int |
   35.52 +  Mutual_Corec of int * int * int |
   35.53 +  Nested_Corec of int;
   35.54 +
   35.55 +type basic_corec_ctr_spec =
   35.56 +  {ctr: term,
   35.57 +   disc: term,
   35.58 +   sels: term list};
   35.59 +
   35.60 +type corec_ctr_spec =
   35.61 +  {ctr: term,
   35.62 +   disc: term,
   35.63 +   sels: term list,
   35.64 +   pred: int option,
   35.65 +   calls: corec_call list,
   35.66 +   discI: thm,
   35.67 +   sel_thms: thm list,
   35.68 +   collapse: thm,
   35.69 +   corec_thm: thm,
   35.70 +   disc_corec: thm,
   35.71 +   sel_corecs: thm list};
   35.72 +
   35.73 +type corec_spec =
   35.74 +  {corec: term,
   35.75 +   nested_map_idents: thm list,
   35.76 +   nested_map_comps: thm list,
   35.77 +   ctr_specs: corec_ctr_spec list};
   35.78 +
   35.79 +exception AINT_NO_MAP of term;
   35.80 +
   35.81 +fun not_codatatype ctxt T =
   35.82 +  error ("Not a codatatype: " ^ Syntax.string_of_typ ctxt T);
   35.83 +fun ill_formed_corec_call ctxt t =
   35.84 +  error ("Ill-formed corecursive call: " ^ quote (Syntax.string_of_term ctxt t));
   35.85 +fun invalid_map ctxt t =
   35.86 +  error ("Invalid map function in " ^ quote (Syntax.string_of_term ctxt t));
   35.87 +fun unexpected_corec_call ctxt t =
   35.88 +  error ("Unexpected corecursive call: " ^ quote (Syntax.string_of_term ctxt t));
   35.89 +
   35.90 +val mk_conjs = try (foldr1 HOLogic.mk_conj) #> the_default @{const True};
   35.91 +val mk_disjs = try (foldr1 HOLogic.mk_disj) #> the_default @{const False};
   35.92 +
   35.93 +val conjuncts_s = filter_out (curry (op =) @{const True}) o HOLogic.conjuncts;
   35.94 +
   35.95 +fun s_not @{const True} = @{const False}
   35.96 +  | s_not @{const False} = @{const True}
   35.97 +  | s_not (@{const Not} $ t) = t
   35.98 +  | s_not (@{const conj} $ t $ u) = @{const disj} $ s_not t $ s_not u
   35.99 +  | s_not (@{const disj} $ t $ u) = @{const conj} $ s_not t $ s_not u
  35.100 +  | s_not t = @{const Not} $ t;
  35.101 +
  35.102 +val s_not_conj = conjuncts_s o s_not o mk_conjs;
  35.103 +
  35.104 +fun propagate_unit_pos u cs = if member (op aconv) cs u then [@{const False}] else cs;
  35.105 +
  35.106 +fun propagate_unit_neg not_u cs = remove (op aconv) not_u cs;
  35.107 +
  35.108 +fun propagate_units css =
  35.109 +  (case List.partition (can the_single) css of
  35.110 +     ([], _) => css
  35.111 +   | ([u] :: uss, css') =>
  35.112 +     [u] :: propagate_units (map (propagate_unit_neg (s_not u))
  35.113 +       (map (propagate_unit_pos u) (uss @ css'))));
  35.114 +
  35.115 +fun s_conjs cs =
  35.116 +  if member (op aconv) cs @{const False} then @{const False}
  35.117 +  else mk_conjs (remove (op aconv) @{const True} cs);
  35.118 +
  35.119 +fun s_disjs ds =
  35.120 +  if member (op aconv) ds @{const True} then @{const True}
  35.121 +  else mk_disjs (remove (op aconv) @{const False} ds);
  35.122 +
  35.123 +fun s_dnf css0 =
  35.124 +  let val css = propagate_units css0 in
  35.125 +    if null css then
  35.126 +      [@{const False}]
  35.127 +    else if exists null css then
  35.128 +      []
  35.129 +    else
  35.130 +      map (fn c :: cs => (c, cs)) css
  35.131 +      |> AList.coalesce (op =)
  35.132 +      |> map (fn (c, css) => c :: s_dnf css)
  35.133 +      |> (fn [cs] => cs | css => [s_disjs (map s_conjs css)])
  35.134 +  end;
  35.135 +
  35.136 +fun fold_rev_let_if_case ctxt f bound_Ts t =
  35.137 +  let
  35.138 +    val thy = Proof_Context.theory_of ctxt;
  35.139 +
  35.140 +    fun fld conds t =
  35.141 +      (case Term.strip_comb t of
  35.142 +        (Const (@{const_name Let}, _), [_, _]) => fld conds (unfold_let t)
  35.143 +      | (Const (@{const_name If}, _), [cond, then_branch, else_branch]) =>
  35.144 +        fld (conds @ conjuncts_s cond) then_branch o fld (conds @ s_not_conj [cond]) else_branch
  35.145 +      | (Const (c, _), args as _ :: _ :: _) =>
  35.146 +        let val n = num_binder_types (Sign.the_const_type thy c) - 1 in
  35.147 +          if n >= 0 andalso n < length args then
  35.148 +            (case fastype_of1 (bound_Ts, nth args n) of
  35.149 +              Type (s, Ts) =>
  35.150 +              (case dest_case ctxt s Ts t of
  35.151 +                NONE => apsnd (f conds t)
  35.152 +              | SOME (conds', branches) =>
  35.153 +                apfst (cons s) o fold_rev (uncurry fld)
  35.154 +                  (map (append conds o conjuncts_s) conds' ~~ branches))
  35.155 +            | _ => apsnd (f conds t))
  35.156 +          else
  35.157 +            apsnd (f conds t)
  35.158 +        end
  35.159 +      | _ => apsnd (f conds t))
  35.160 +  in
  35.161 +    fld [] t o pair []
  35.162 +  end;
  35.163 +
  35.164 +fun case_of ctxt = ctr_sugar_of ctxt #> Option.map (fst o dest_Const o #casex);
  35.165 +
  35.166 +fun massage_let_if_case ctxt has_call massage_leaf =
  35.167 +  let
  35.168 +    val thy = Proof_Context.theory_of ctxt;
  35.169 +
  35.170 +    fun check_no_call t = if has_call t then unexpected_corec_call ctxt t else ();
  35.171 +
  35.172 +    fun massage_abs bound_Ts 0 t = massage_rec bound_Ts t
  35.173 +      | massage_abs bound_Ts m (Abs (s, T, t)) = Abs (s, T, massage_abs (T :: bound_Ts) (m - 1) t)
  35.174 +      | massage_abs bound_Ts m t =
  35.175 +        let val T = domain_type (fastype_of1 (bound_Ts, t)) in
  35.176 +          Abs (Name.uu, T, massage_abs (T :: bound_Ts) (m - 1) (incr_boundvars 1 t $ Bound 0))
  35.177 +        end
  35.178 +    and massage_rec bound_Ts t =
  35.179 +      let val typof = curry fastype_of1 bound_Ts in
  35.180 +        (case Term.strip_comb t of
  35.181 +          (Const (@{const_name Let}, _), [_, _]) => massage_rec bound_Ts (unfold_let t)
  35.182 +        | (Const (@{const_name If}, _), obj :: (branches as [_, _])) =>
  35.183 +          let val branches' = map (massage_rec bound_Ts) branches in
  35.184 +            Term.list_comb (If_const (typof (hd branches')) $ tap check_no_call obj, branches')
  35.185 +          end
  35.186 +        | (Const (c, _), args as _ :: _ :: _) =>
  35.187 +          (case try strip_fun_type (Sign.the_const_type thy c) of
  35.188 +            SOME (gen_branch_Ts, gen_body_fun_T) =>
  35.189 +            let
  35.190 +              val gen_branch_ms = map num_binder_types gen_branch_Ts;
  35.191 +              val n = length gen_branch_ms;
  35.192 +            in
  35.193 +              if n < length args then
  35.194 +                (case gen_body_fun_T of
  35.195 +                  Type (_, [Type (T_name, _), _]) =>
  35.196 +                  if case_of ctxt T_name = SOME c then
  35.197 +                    let
  35.198 +                      val (branches, obj_leftovers) = chop n args;
  35.199 +                      val branches' = map2 (massage_abs bound_Ts) gen_branch_ms branches;
  35.200 +                      val branch_Ts' = map typof branches';
  35.201 +                      val body_T' = snd (strip_typeN (hd gen_branch_ms) (hd branch_Ts'));
  35.202 +                      val casex' = Const (c, branch_Ts' ---> map typof obj_leftovers ---> body_T');
  35.203 +                    in
  35.204 +                      Term.list_comb (casex',
  35.205 +                        branches' @ tap (List.app check_no_call) obj_leftovers)
  35.206 +                    end
  35.207 +                  else
  35.208 +                    massage_leaf bound_Ts t
  35.209 +                | _ => massage_leaf bound_Ts t)
  35.210 +              else
  35.211 +                massage_leaf bound_Ts t
  35.212 +            end
  35.213 +          | NONE => massage_leaf bound_Ts t)
  35.214 +        | _ => massage_leaf bound_Ts t)
  35.215 +      end
  35.216 +  in
  35.217 +    massage_rec
  35.218 +  end;
  35.219 +
  35.220 +val massage_mutual_corec_call = massage_let_if_case;
  35.221 +
  35.222 +fun curried_type (Type (@{type_name fun}, [Type (@{type_name prod}, Ts), T])) = Ts ---> T;
  35.223 +
  35.224 +fun massage_nested_corec_call ctxt has_call raw_massage_call bound_Ts U t =
  35.225 +  let
  35.226 +    fun check_no_call t = if has_call t then unexpected_corec_call ctxt t else ();
  35.227 +
  35.228 +    val build_map_Inl = build_map ctxt (uncurry Inl_const o dest_sumT o snd);
  35.229 +
  35.230 +    fun massage_mutual_call bound_Ts U T t =
  35.231 +      if has_call t then
  35.232 +        (case try dest_sumT U of
  35.233 +          SOME (U1, U2) => if U1 = T then raw_massage_call bound_Ts T U2 t else invalid_map ctxt t
  35.234 +        | NONE => invalid_map ctxt t)
  35.235 +      else
  35.236 +        build_map_Inl (T, U) $ t;
  35.237 +
  35.238 +    fun massage_mutual_fun bound_Ts U T t =
  35.239 +      (case t of
  35.240 +        Const (@{const_name comp}, _) $ t1 $ t2 =>
  35.241 +        mk_comp bound_Ts (massage_mutual_fun bound_Ts U T t1, tap check_no_call t2)
  35.242 +      | _ =>
  35.243 +        let
  35.244 +          val var = Var ((Name.uu, Term.maxidx_of_term t + 1),
  35.245 +            domain_type (fastype_of1 (bound_Ts, t)));
  35.246 +        in
  35.247 +          Term.lambda var (massage_mutual_call bound_Ts U T (t $ var))
  35.248 +        end);
  35.249 +
  35.250 +    fun massage_map bound_Ts (Type (_, Us)) (Type (s, Ts)) t =
  35.251 +        (case try (dest_map ctxt s) t of
  35.252 +          SOME (map0, fs) =>
  35.253 +          let
  35.254 +            val Type (_, dom_Ts) = domain_type (fastype_of1 (bound_Ts, t));
  35.255 +            val map' = mk_map (length fs) dom_Ts Us map0;
  35.256 +            val fs' =
  35.257 +              map_flattened_map_args ctxt s (map3 (massage_map_or_map_arg bound_Ts) Us Ts) fs;
  35.258 +          in
  35.259 +            Term.list_comb (map', fs')
  35.260 +          end
  35.261 +        | NONE => raise AINT_NO_MAP t)
  35.262 +      | massage_map _ _ _ t = raise AINT_NO_MAP t
  35.263 +    and massage_map_or_map_arg bound_Ts U T t =
  35.264 +      if T = U then
  35.265 +        tap check_no_call t
  35.266 +      else
  35.267 +        massage_map bound_Ts U T t
  35.268 +        handle AINT_NO_MAP _ => massage_mutual_fun bound_Ts U T t;
  35.269 +
  35.270 +    fun massage_call bound_Ts U T =
  35.271 +      massage_let_if_case ctxt has_call (fn bound_Ts => fn t =>
  35.272 +        if has_call t then
  35.273 +          (case t of
  35.274 +            Const (@{const_name prod_case}, _) $ t' =>
  35.275 +            let
  35.276 +              val U' = curried_type U;
  35.277 +              val T' = curried_type T;
  35.278 +            in
  35.279 +              Const (@{const_name prod_case}, U' --> U) $ massage_call bound_Ts U' T' t'
  35.280 +            end
  35.281 +          | t1 $ t2 =>
  35.282 +            (if has_call t2 then
  35.283 +              massage_mutual_call bound_Ts U T t
  35.284 +            else
  35.285 +              massage_map bound_Ts U T t1 $ t2
  35.286 +              handle AINT_NO_MAP _ => massage_mutual_call bound_Ts U T t)
  35.287 +          | Abs (s, T', t') =>
  35.288 +            Abs (s, T', massage_call (T' :: bound_Ts) (range_type U) (range_type T) t')
  35.289 +          | _ => massage_mutual_call bound_Ts U T t)
  35.290 +        else
  35.291 +          build_map_Inl (T, U) $ t) bound_Ts;
  35.292 +
  35.293 +    val T = fastype_of1 (bound_Ts, t);
  35.294 +  in
  35.295 +    if has_call t then massage_call bound_Ts U T t else build_map_Inl (T, U) $ t
  35.296 +  end;
  35.297 +
  35.298 +val fold_rev_corec_call = fold_rev_let_if_case;
  35.299 +
  35.300 +fun expand_to_ctr_term ctxt s Ts t =
  35.301 +  (case ctr_sugar_of ctxt s of
  35.302 +    SOME {ctrs, casex, ...} =>
  35.303 +    Term.list_comb (mk_case Ts (Type (s, Ts)) casex, map (mk_ctr Ts) ctrs) $ t
  35.304 +  | NONE => raise Fail "expand_to_ctr_term");
  35.305 +
  35.306 +fun expand_corec_code_rhs ctxt has_call bound_Ts t =
  35.307 +  (case fastype_of1 (bound_Ts, t) of
  35.308 +    Type (s, Ts) =>
  35.309 +    massage_let_if_case ctxt has_call (fn _ => fn t =>
  35.310 +      if can (dest_ctr ctxt s) t then t else expand_to_ctr_term ctxt s Ts t) bound_Ts t
  35.311 +  | _ => raise Fail "expand_corec_code_rhs");
  35.312 +
  35.313 +fun massage_corec_code_rhs ctxt massage_ctr =
  35.314 +  massage_let_if_case ctxt (K false)
  35.315 +    (fn bound_Ts => uncurry (massage_ctr bound_Ts) o Term.strip_comb);
  35.316 +
  35.317 +fun fold_rev_corec_code_rhs ctxt f =
  35.318 +  snd ooo fold_rev_let_if_case ctxt (fn conds => uncurry (f conds) o Term.strip_comb);
  35.319 +
  35.320 +fun case_thms_of_term ctxt bound_Ts t =
  35.321 +  let
  35.322 +    val (caseT_names, _) = fold_rev_let_if_case ctxt (K (K I)) bound_Ts t ();
  35.323 +    val ctr_sugars = map (the o ctr_sugar_of ctxt) caseT_names;
  35.324 +  in
  35.325 +    (maps #distincts ctr_sugars, maps #discIs ctr_sugars, maps #sel_splits ctr_sugars,
  35.326 +     maps #sel_split_asms ctr_sugars)
  35.327 +  end;
  35.328 +
  35.329 +fun basic_corec_specs_of ctxt res_T =
  35.330 +  (case res_T of
  35.331 +    Type (T_name, _) =>
  35.332 +    (case Ctr_Sugar.ctr_sugar_of ctxt T_name of
  35.333 +      NONE => not_codatatype ctxt res_T
  35.334 +    | SOME {ctrs, discs, selss, ...} =>
  35.335 +      let
  35.336 +        val thy = Proof_Context.theory_of ctxt;
  35.337 +
  35.338 +        val gfpT = body_type (fastype_of (hd ctrs));
  35.339 +        val As_rho = tvar_subst thy [gfpT] [res_T];
  35.340 +        val substA = Term.subst_TVars As_rho;
  35.341 +
  35.342 +        fun mk_spec ctr disc sels = {ctr = substA ctr, disc = substA disc, sels = map substA sels};
  35.343 +      in
  35.344 +        map3 mk_spec ctrs discs selss
  35.345 +      end)
  35.346 +  | _ => not_codatatype ctxt res_T);
  35.347 +
  35.348 +fun corec_specs_of bs arg_Ts res_Ts get_indices callssss0 lthy =
  35.349 +  let
  35.350 +    val thy = Proof_Context.theory_of lthy;
  35.351 +
  35.352 +    val ((missing_res_Ts, perm0_kks,
  35.353 +          fp_sugars as {nested_bnfs, fp_res = {xtor_co_iterss = dtor_coiters1 :: _, ...},
  35.354 +            co_inducts = coinduct_thms, ...} :: _, (_, gfp_sugar_thms)), lthy') =
  35.355 +      nested_to_mutual_fps Greatest_FP bs res_Ts get_indices callssss0 lthy;
  35.356 +
  35.357 +    val perm_fp_sugars = sort (int_ord o pairself #index) fp_sugars;
  35.358 +
  35.359 +    val indices = map #index fp_sugars;
  35.360 +    val perm_indices = map #index perm_fp_sugars;
  35.361 +
  35.362 +    val perm_ctrss = map (#ctrs o of_fp_sugar #ctr_sugars) perm_fp_sugars;
  35.363 +    val perm_ctr_Tsss = map (map (binder_types o fastype_of)) perm_ctrss;
  35.364 +    val perm_gfpTs = map (body_type o fastype_of o hd) perm_ctrss;
  35.365 +
  35.366 +    val nn0 = length res_Ts;
  35.367 +    val nn = length perm_gfpTs;
  35.368 +    val kks = 0 upto nn - 1;
  35.369 +    val perm_ns = map length perm_ctr_Tsss;
  35.370 +
  35.371 +    val perm_Cs = map (domain_type o body_fun_type o fastype_of o co_rec_of o
  35.372 +      of_fp_sugar (#xtor_co_iterss o #fp_res)) perm_fp_sugars;
  35.373 +    val (perm_p_Tss, (perm_q_Tssss, _, perm_f_Tssss, _)) =
  35.374 +      mk_coiter_fun_arg_types perm_ctr_Tsss perm_Cs perm_ns (co_rec_of dtor_coiters1);
  35.375 +
  35.376 +    val (perm_p_hss, h) = indexedd perm_p_Tss 0;
  35.377 +    val (perm_q_hssss, h') = indexedddd perm_q_Tssss h;
  35.378 +    val (perm_f_hssss, _) = indexedddd perm_f_Tssss h';
  35.379 +
  35.380 +    val fun_arg_hs =
  35.381 +      flat (map3 flat_corec_preds_predsss_gettersss perm_p_hss perm_q_hssss perm_f_hssss);
  35.382 +
  35.383 +    fun unpermute0 perm0_xs = permute_like (op =) perm0_kks kks perm0_xs;
  35.384 +    fun unpermute perm_xs = permute_like (op =) perm_indices indices perm_xs;
  35.385 +
  35.386 +    val coinduct_thmss = map (unpermute0 o conj_dests nn) coinduct_thms;
  35.387 +
  35.388 +    val p_iss = map (map (find_index_eq fun_arg_hs)) (unpermute perm_p_hss);
  35.389 +    val q_issss = map (map (map (map (find_index_eq fun_arg_hs)))) (unpermute perm_q_hssss);
  35.390 +    val f_issss = map (map (map (map (find_index_eq fun_arg_hs)))) (unpermute perm_f_hssss);
  35.391 +
  35.392 +    val f_Tssss = unpermute perm_f_Tssss;
  35.393 +    val gfpTs = unpermute perm_gfpTs;
  35.394 +    val Cs = unpermute perm_Cs;
  35.395 +
  35.396 +    val As_rho = tvar_subst thy (take nn0 gfpTs) res_Ts;
  35.397 +    val Cs_rho = map (fst o dest_TVar) Cs ~~ pad_list HOLogic.unitT nn arg_Ts;
  35.398 +
  35.399 +    val substA = Term.subst_TVars As_rho;
  35.400 +    val substAT = Term.typ_subst_TVars As_rho;
  35.401 +    val substCT = Term.typ_subst_TVars Cs_rho;
  35.402 +
  35.403 +    val perm_Cs' = map substCT perm_Cs;
  35.404 +
  35.405 +    fun call_of nullary [] [g_i] [Type (@{type_name fun}, [_, T])] =
  35.406 +        (if exists_subtype_in Cs T then Nested_Corec
  35.407 +         else if nullary then Dummy_No_Corec
  35.408 +         else No_Corec) g_i
  35.409 +      | call_of _ [q_i] [g_i, g_i'] _ = Mutual_Corec (q_i, g_i, g_i');
  35.410 +
  35.411 +    fun mk_ctr_spec ctr disc sels p_ho q_iss f_iss f_Tss discI sel_thms collapse corec_thm
  35.412 +        disc_corec sel_corecs =
  35.413 +      let val nullary = not (can dest_funT (fastype_of ctr)) in
  35.414 +        {ctr = substA ctr, disc = substA disc, sels = map substA sels, pred = p_ho,
  35.415 +         calls = map3 (call_of nullary) q_iss f_iss f_Tss, discI = discI, sel_thms = sel_thms,
  35.416 +         collapse = collapse, corec_thm = corec_thm, disc_corec = disc_corec,
  35.417 +         sel_corecs = sel_corecs}
  35.418 +      end;
  35.419 +
  35.420 +    fun mk_ctr_specs index (ctr_sugars : ctr_sugar list) p_is q_isss f_isss f_Tsss coiter_thmsss
  35.421 +        disc_coitersss sel_coiterssss =
  35.422 +      let
  35.423 +        val ctrs = #ctrs (nth ctr_sugars index);
  35.424 +        val discs = #discs (nth ctr_sugars index);
  35.425 +        val selss = #selss (nth ctr_sugars index);
  35.426 +        val p_ios = map SOME p_is @ [NONE];
  35.427 +        val discIs = #discIs (nth ctr_sugars index);
  35.428 +        val sel_thmss = #sel_thmss (nth ctr_sugars index);
  35.429 +        val collapses = #collapses (nth ctr_sugars index);
  35.430 +        val corec_thms = co_rec_of (nth coiter_thmsss index);
  35.431 +        val disc_corecs = co_rec_of (nth disc_coitersss index);
  35.432 +        val sel_corecss = co_rec_of (nth sel_coiterssss index);
  35.433 +      in
  35.434 +        map13 mk_ctr_spec ctrs discs selss p_ios q_isss f_isss f_Tsss discIs sel_thmss collapses
  35.435 +          corec_thms disc_corecs sel_corecss
  35.436 +      end;
  35.437 +
  35.438 +    fun mk_spec ({T, index, ctr_sugars, co_iterss = coiterss, co_iter_thmsss = coiter_thmsss,
  35.439 +          disc_co_itersss = disc_coitersss, sel_co_iterssss = sel_coiterssss, ...} : fp_sugar)
  35.440 +        p_is q_isss f_isss f_Tsss =
  35.441 +      {corec = mk_co_iter thy Greatest_FP (substAT T) perm_Cs' (co_rec_of (nth coiterss index)),
  35.442 +       nested_map_idents = map (unfold_thms lthy @{thms id_def} o map_id0_of_bnf) nested_bnfs,
  35.443 +       nested_map_comps = map map_comp_of_bnf nested_bnfs,
  35.444 +       ctr_specs = mk_ctr_specs index ctr_sugars p_is q_isss f_isss f_Tsss coiter_thmsss
  35.445 +         disc_coitersss sel_coiterssss};
  35.446 +  in
  35.447 +    ((is_some gfp_sugar_thms, map5 mk_spec fp_sugars p_iss q_issss f_issss f_Tssss, missing_res_Ts,
  35.448 +      co_induct_of coinduct_thms, strong_co_induct_of coinduct_thms, co_induct_of coinduct_thmss,
  35.449 +      strong_co_induct_of coinduct_thmss), lthy')
  35.450 +  end;
  35.451 +
  35.452 +val undef_const = Const (@{const_name undefined}, dummyT);
  35.453 +
  35.454 +val abs_tuple = HOLogic.tupled_lambda o HOLogic.mk_tuple;
  35.455 +fun abstract vs =
  35.456 +  let fun a n (t $ u) = a n t $ a n u
  35.457 +        | a n (Abs (v, T, b)) = Abs (v, T, a (n + 1) b)
  35.458 +        | a n t = let val idx = find_index (equal t) vs in
  35.459 +            if idx < 0 then t else Bound (n + idx) end
  35.460 +  in a 0 end;
  35.461 +
  35.462 +fun mk_prod1 bound_Ts (t, u) =
  35.463 +  HOLogic.pair_const (fastype_of1 (bound_Ts, t)) (fastype_of1 (bound_Ts, u)) $ t $ u;
  35.464 +fun mk_tuple1 bound_Ts = the_default HOLogic.unit o try (foldr1 (mk_prod1 bound_Ts));
  35.465 +
  35.466 +type coeqn_data_disc = {
  35.467 +  fun_name: string,
  35.468 +  fun_T: typ,
  35.469 +  fun_args: term list,
  35.470 +  ctr: term,
  35.471 +  ctr_no: int, (*###*)
  35.472 +  disc: term,
  35.473 +  prems: term list,
  35.474 +  auto_gen: bool,
  35.475 +  maybe_ctr_rhs: term option,
  35.476 +  maybe_code_rhs: term option,
  35.477 +  user_eqn: term
  35.478 +};
  35.479 +
  35.480 +type coeqn_data_sel = {
  35.481 +  fun_name: string,
  35.482 +  fun_T: typ,
  35.483 +  fun_args: term list,
  35.484 +  ctr: term,
  35.485 +  sel: term,
  35.486 +  rhs_term: term,
  35.487 +  user_eqn: term
  35.488 +};
  35.489 +
  35.490 +datatype coeqn_data =
  35.491 +  Disc of coeqn_data_disc |
  35.492 +  Sel of coeqn_data_sel;
  35.493 +
  35.494 +fun dissect_coeqn_disc seq fun_names (basic_ctr_specss : basic_corec_ctr_spec list list)
  35.495 +    maybe_ctr_rhs maybe_code_rhs prems' concl matchedsss =
  35.496 +  let
  35.497 +    fun find_subterm p =
  35.498 +      let (* FIXME \<exists>? *)
  35.499 +        fun find (t as u $ v) = if p t then SOME t else merge_options (find u, find v)
  35.500 +          | find t = if p t then SOME t else NONE;
  35.501 +      in find end;
  35.502 +
  35.503 +    val applied_fun = concl
  35.504 +      |> find_subterm (member ((op =) o apsnd SOME) fun_names o try (fst o dest_Free o head_of))
  35.505 +      |> the
  35.506 +      handle Option.Option => primcorec_error_eqn "malformed discriminator formula" concl;
  35.507 +    val ((fun_name, fun_T), fun_args) = strip_comb applied_fun |>> dest_Free;
  35.508 +    val SOME basic_ctr_specs = AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name;
  35.509 +
  35.510 +    val discs = map #disc basic_ctr_specs;
  35.511 +    val ctrs = map #ctr basic_ctr_specs;
  35.512 +    val not_disc = head_of concl = @{term Not};
  35.513 +    val _ = not_disc andalso length ctrs <> 2 andalso
  35.514 +      primcorec_error_eqn "negated discriminator for a type with \<noteq> 2 constructors" concl;
  35.515 +    val disc' = find_subterm (member (op =) discs o head_of) concl;
  35.516 +    val eq_ctr0 = concl |> perhaps (try HOLogic.dest_not) |> try (HOLogic.dest_eq #> snd)
  35.517 +        |> (fn SOME t => let val n = find_index (equal t) ctrs in
  35.518 +          if n >= 0 then SOME n else NONE end | _ => NONE);
  35.519 +    val _ = is_some disc' orelse is_some eq_ctr0 orelse
  35.520 +      primcorec_error_eqn "no discriminator in equation" concl;
  35.521 +    val ctr_no' =
  35.522 +      if is_none disc' then the eq_ctr0 else find_index (equal (head_of (the disc'))) discs;
  35.523 +    val ctr_no = if not_disc then 1 - ctr_no' else ctr_no';
  35.524 +    val {ctr, disc, ...} = nth basic_ctr_specs ctr_no;
  35.525 +
  35.526 +    val catch_all = try (fst o dest_Free o the_single) prems' = SOME Name.uu_;
  35.527 +    val matchedss = AList.lookup (op =) matchedsss fun_name |> the_default [];
  35.528 +    val prems = map (abstract (List.rev fun_args)) prems';
  35.529 +    val real_prems =
  35.530 +      (if catch_all orelse seq then maps s_not_conj matchedss else []) @
  35.531 +      (if catch_all then [] else prems);
  35.532 +
  35.533 +    val matchedsss' = AList.delete (op =) fun_name matchedsss
  35.534 +      |> cons (fun_name, if seq then matchedss @ [prems] else matchedss @ [real_prems]);
  35.535 +
  35.536 +    val user_eqn =
  35.537 +      (real_prems, concl)
  35.538 +      |>> map HOLogic.mk_Trueprop ||> HOLogic.mk_Trueprop o abstract (List.rev fun_args)
  35.539 +      |> curry Logic.list_all (map dest_Free fun_args) o Logic.list_implies;
  35.540 +  in
  35.541 +    (Disc {
  35.542 +      fun_name = fun_name,
  35.543 +      fun_T = fun_T,
  35.544 +      fun_args = fun_args,
  35.545 +      ctr = ctr,
  35.546 +      ctr_no = ctr_no,
  35.547 +      disc = disc,
  35.548 +      prems = real_prems,
  35.549 +      auto_gen = catch_all,
  35.550 +      maybe_ctr_rhs = maybe_ctr_rhs,
  35.551 +      maybe_code_rhs = maybe_code_rhs,
  35.552 +      user_eqn = user_eqn
  35.553 +    }, matchedsss')
  35.554 +  end;
  35.555 +
  35.556 +fun dissect_coeqn_sel fun_names (basic_ctr_specss : basic_corec_ctr_spec list list) eqn'
  35.557 +    maybe_of_spec eqn =
  35.558 +  let
  35.559 +    val (lhs, rhs) = HOLogic.dest_eq eqn
  35.560 +      handle TERM _ =>
  35.561 +        primcorec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn;
  35.562 +    val sel = head_of lhs;
  35.563 +    val ((fun_name, fun_T), fun_args) = dest_comb lhs |> snd |> strip_comb |> apfst dest_Free
  35.564 +      handle TERM _ =>
  35.565 +        primcorec_error_eqn "malformed selector argument in left-hand side" eqn;
  35.566 +    val basic_ctr_specs = the (AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name)
  35.567 +      handle Option.Option =>
  35.568 +        primcorec_error_eqn "malformed selector argument in left-hand side" eqn;
  35.569 +    val {ctr, ...} =
  35.570 +      (case maybe_of_spec of
  35.571 +        SOME of_spec => the (find_first (equal of_spec o #ctr) basic_ctr_specs)
  35.572 +      | NONE => filter (exists (equal sel) o #sels) basic_ctr_specs |> the_single
  35.573 +          handle List.Empty => primcorec_error_eqn "ambiguous selector - use \"of\"" eqn);
  35.574 +    val user_eqn = drop_All eqn';
  35.575 +  in
  35.576 +    Sel {
  35.577 +      fun_name = fun_name,
  35.578 +      fun_T = fun_T,
  35.579 +      fun_args = fun_args,
  35.580 +      ctr = ctr,
  35.581 +      sel = sel,
  35.582 +      rhs_term = rhs,
  35.583 +      user_eqn = user_eqn
  35.584 +    }
  35.585 +  end;
  35.586 +
  35.587 +fun dissect_coeqn_ctr seq fun_names (basic_ctr_specss : basic_corec_ctr_spec list list) eqn'
  35.588 +    maybe_code_rhs prems concl matchedsss =
  35.589 +  let
  35.590 +    val (lhs, rhs) = HOLogic.dest_eq concl;
  35.591 +    val (fun_name, fun_args) = strip_comb lhs |>> fst o dest_Free;
  35.592 +    val SOME basic_ctr_specs = AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name;
  35.593 +    val (ctr, ctr_args) = strip_comb (unfold_let rhs);
  35.594 +    val {disc, sels, ...} = the (find_first (equal ctr o #ctr) basic_ctr_specs)
  35.595 +      handle Option.Option => primcorec_error_eqn "not a constructor" ctr;
  35.596 +
  35.597 +    val disc_concl = betapply (disc, lhs);
  35.598 +    val (maybe_eqn_data_disc, matchedsss') = if length basic_ctr_specs = 1
  35.599 +      then (NONE, matchedsss)
  35.600 +      else apfst SOME (dissect_coeqn_disc seq fun_names basic_ctr_specss
  35.601 +          (SOME (abstract (List.rev fun_args) rhs)) maybe_code_rhs prems disc_concl matchedsss);
  35.602 +
  35.603 +    val sel_concls = sels ~~ ctr_args
  35.604 +      |> map (fn (sel, ctr_arg) => HOLogic.mk_eq (betapply (sel, lhs), ctr_arg));
  35.605 +
  35.606 +(*
  35.607 +val _ = tracing ("reduced\n    " ^ Syntax.string_of_term @{context} concl ^ "\nto\n    \<cdot> " ^
  35.608 + (is_some maybe_eqn_data_disc ? K (Syntax.string_of_term @{context} disc_concl ^ "\n    \<cdot> ")) "" ^
  35.609 + space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) sel_concls) ^
  35.610 + "\nfor premise(s)\n    \<cdot> " ^
  35.611 + space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) prems));
  35.612 +*)
  35.613 +
  35.614 +    val eqns_data_sel =
  35.615 +      map (dissect_coeqn_sel fun_names basic_ctr_specss eqn' (SOME ctr)) sel_concls;
  35.616 +  in
  35.617 +    (the_list maybe_eqn_data_disc @ eqns_data_sel, matchedsss')
  35.618 +  end;
  35.619 +
  35.620 +fun dissect_coeqn_code lthy has_call fun_names basic_ctr_specss eqn' concl matchedsss =
  35.621 +  let
  35.622 +    val (lhs, (rhs', rhs)) = HOLogic.dest_eq concl ||> `(expand_corec_code_rhs lthy has_call []);
  35.623 +    val (fun_name, fun_args) = strip_comb lhs |>> fst o dest_Free;
  35.624 +    val SOME basic_ctr_specs = AList.lookup (op =) (fun_names ~~ basic_ctr_specss) fun_name;
  35.625 +
  35.626 +    val cond_ctrs = fold_rev_corec_code_rhs lthy (fn cs => fn ctr => fn _ =>
  35.627 +        if member ((op =) o apsnd #ctr) basic_ctr_specs ctr
  35.628 +        then cons (ctr, cs)
  35.629 +        else primcorec_error_eqn "not a constructor" ctr) [] rhs' []
  35.630 +      |> AList.group (op =);
  35.631 +
  35.632 +    val ctr_premss = (case cond_ctrs of [_] => [[]] | _ => map (s_dnf o snd) cond_ctrs);
  35.633 +    val ctr_concls = cond_ctrs |> map (fn (ctr, _) =>
  35.634 +        binder_types (fastype_of ctr)
  35.635 +        |> map_index (fn (n, T) => massage_corec_code_rhs lthy (fn _ => fn ctr' => fn args =>
  35.636 +          if ctr' = ctr then nth args n else Const (@{const_name undefined}, T)) [] rhs')
  35.637 +        |> curry list_comb ctr
  35.638 +        |> curry HOLogic.mk_eq lhs);
  35.639 +  in
  35.640 +    fold_map2 (dissect_coeqn_ctr false fun_names basic_ctr_specss eqn'
  35.641 +        (SOME (abstract (List.rev fun_args) rhs)))
  35.642 +      ctr_premss ctr_concls matchedsss
  35.643 +  end;
  35.644 +
  35.645 +fun dissect_coeqn lthy seq has_call fun_names (basic_ctr_specss : basic_corec_ctr_spec list list)
  35.646 +    eqn' maybe_of_spec matchedsss =
  35.647 +  let
  35.648 +    val eqn = drop_All eqn'
  35.649 +      handle TERM _ => primcorec_error_eqn "malformed function equation" eqn';
  35.650 +    val (prems, concl) = Logic.strip_horn eqn
  35.651 +      |> apfst (map HOLogic.dest_Trueprop) o apsnd HOLogic.dest_Trueprop;
  35.652 +
  35.653 +    val head = concl
  35.654 +      |> perhaps (try HOLogic.dest_not) |> perhaps (try (fst o HOLogic.dest_eq))
  35.655 +      |> head_of;
  35.656 +
  35.657 +    val maybe_rhs = concl |> perhaps (try HOLogic.dest_not) |> try (snd o HOLogic.dest_eq);
  35.658 +
  35.659 +    val discs = maps (map #disc) basic_ctr_specss;
  35.660 +    val sels = maps (maps #sels) basic_ctr_specss;
  35.661 +    val ctrs = maps (map #ctr) basic_ctr_specss;
  35.662 +  in
  35.663 +    if member (op =) discs head orelse
  35.664 +      is_some maybe_rhs andalso
  35.665 +        member (op =) (filter (null o binder_types o fastype_of) ctrs) (the maybe_rhs) then
  35.666 +      dissect_coeqn_disc seq fun_names basic_ctr_specss NONE NONE prems concl matchedsss
  35.667 +      |>> single
  35.668 +    else if member (op =) sels head then
  35.669 +      ([dissect_coeqn_sel fun_names basic_ctr_specss eqn' maybe_of_spec concl], matchedsss)
  35.670 +    else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) andalso
  35.671 +      member (op =) ctrs (head_of (unfold_let (the maybe_rhs))) then
  35.672 +      dissect_coeqn_ctr seq fun_names basic_ctr_specss eqn' NONE prems concl matchedsss
  35.673 +    else if is_Free head andalso member (op =) fun_names (fst (dest_Free head)) andalso
  35.674 +      null prems then
  35.675 +      dissect_coeqn_code lthy has_call fun_names basic_ctr_specss eqn' concl matchedsss
  35.676 +      |>> flat
  35.677 +    else
  35.678 +      primcorec_error_eqn "malformed function equation" eqn
  35.679 +  end;
  35.680 +
  35.681 +fun build_corec_arg_disc (ctr_specs : corec_ctr_spec list)
  35.682 +    ({fun_args, ctr_no, prems, ...} : coeqn_data_disc) =
  35.683 +  if is_none (#pred (nth ctr_specs ctr_no)) then I else
  35.684 +    s_conjs prems
  35.685 +    |> curry subst_bounds (List.rev fun_args)
  35.686 +    |> HOLogic.tupled_lambda (HOLogic.mk_tuple fun_args)
  35.687 +    |> K |> nth_map (the (#pred (nth ctr_specs ctr_no)));
  35.688 +
  35.689 +fun build_corec_arg_no_call (sel_eqns : coeqn_data_sel list) sel =
  35.690 +  find_first (equal sel o #sel) sel_eqns
  35.691 +  |> try (fn SOME {fun_args, rhs_term, ...} => abs_tuple fun_args rhs_term)
  35.692 +  |> the_default undef_const
  35.693 +  |> K;
  35.694 +
  35.695 +fun build_corec_args_mutual_call lthy has_call (sel_eqns : coeqn_data_sel list) sel =
  35.696 +  (case find_first (equal sel o #sel) sel_eqns of
  35.697 +    NONE => (I, I, I)
  35.698 +  | SOME {fun_args, rhs_term, ... } =>
  35.699 +    let
  35.700 +      val bound_Ts = List.rev (map fastype_of fun_args);
  35.701 +      fun rewrite_stop _ t = if has_call t then @{term False} else @{term True};
  35.702 +      fun rewrite_end _ t = if has_call t then undef_const else t;
  35.703 +      fun rewrite_cont bound_Ts t =
  35.704 +        if has_call t then mk_tuple1 bound_Ts (snd (strip_comb t)) else undef_const;
  35.705 +      fun massage f _ = massage_mutual_corec_call lthy has_call f bound_Ts rhs_term
  35.706 +        |> abs_tuple fun_args;
  35.707 +    in
  35.708 +      (massage rewrite_stop, massage rewrite_end, massage rewrite_cont)
  35.709 +    end);
  35.710 +
  35.711 +fun build_corec_arg_nested_call lthy has_call (sel_eqns : coeqn_data_sel list) sel =
  35.712 +  (case find_first (equal sel o #sel) sel_eqns of
  35.713 +    NONE => I
  35.714 +  | SOME {fun_args, rhs_term, ...} =>
  35.715 +    let
  35.716 +      val bound_Ts = List.rev (map fastype_of fun_args);
  35.717 +      fun rewrite bound_Ts U T (Abs (v, V, b)) = Abs (v, V, rewrite (V :: bound_Ts) U T b)
  35.718 +        | rewrite bound_Ts U T (t as _ $ _) =
  35.719 +          let val (u, vs) = strip_comb t in
  35.720 +            if is_Free u andalso has_call u then
  35.721 +              Inr_const U T $ mk_tuple1 bound_Ts vs
  35.722 +            else if try (fst o dest_Const) u = SOME @{const_name prod_case} then
  35.723 +              map (rewrite bound_Ts U T) vs |> chop 1 |>> HOLogic.mk_split o the_single |> list_comb
  35.724 +            else
  35.725 +              list_comb (rewrite bound_Ts U T u, map (rewrite bound_Ts U T) vs)
  35.726 +          end
  35.727 +        | rewrite _ U T t =
  35.728 +          if is_Free t andalso has_call t then Inr_const U T $ HOLogic.unit else t;
  35.729 +      fun massage t =
  35.730 +        rhs_term
  35.731 +        |> massage_nested_corec_call lthy has_call rewrite bound_Ts (range_type (fastype_of t))
  35.732 +        |> abs_tuple fun_args;
  35.733 +    in
  35.734 +      massage
  35.735 +    end);
  35.736 +
  35.737 +fun build_corec_args_sel lthy has_call (all_sel_eqns : coeqn_data_sel list)
  35.738 +    (ctr_spec : corec_ctr_spec) =
  35.739 +  (case filter (equal (#ctr ctr_spec) o #ctr) all_sel_eqns of
  35.740 +    [] => I
  35.741 +  | sel_eqns =>
  35.742 +    let
  35.743 +      val sel_call_list = #sels ctr_spec ~~ #calls ctr_spec;
  35.744 +      val no_calls' = map_filter (try (apsnd (fn No_Corec n => n))) sel_call_list;
  35.745 +      val mutual_calls' = map_filter (try (apsnd (fn Mutual_Corec n => n))) sel_call_list;
  35.746 +      val nested_calls' = map_filter (try (apsnd (fn Nested_Corec n => n))) sel_call_list;
  35.747 +    in
  35.748 +      I
  35.749 +      #> fold (fn (sel, n) => nth_map n (build_corec_arg_no_call sel_eqns sel)) no_calls'
  35.750 +      #> fold (fn (sel, (q, g, h)) =>
  35.751 +        let val (fq, fg, fh) = build_corec_args_mutual_call lthy has_call sel_eqns sel in
  35.752 +          nth_map q fq o nth_map g fg o nth_map h fh end) mutual_calls'
  35.753 +      #> fold (fn (sel, n) => nth_map n
  35.754 +        (build_corec_arg_nested_call lthy has_call sel_eqns sel)) nested_calls'
  35.755 +    end);
  35.756 +
  35.757 +fun build_codefs lthy bs mxs has_call arg_Tss (corec_specs : corec_spec list)
  35.758 +    (disc_eqnss : coeqn_data_disc list list) (sel_eqnss : coeqn_data_sel list list) =
  35.759 +  let
  35.760 +    val corecs = map #corec corec_specs;
  35.761 +    val ctr_specss = map #ctr_specs corec_specs;
  35.762 +    val corec_args = hd corecs
  35.763 +      |> fst o split_last o binder_types o fastype_of
  35.764 +      |> map (Const o pair @{const_name undefined})
  35.765 +      |> fold2 (fold o build_corec_arg_disc) ctr_specss disc_eqnss
  35.766 +      |> fold2 (fold o build_corec_args_sel lthy has_call) sel_eqnss ctr_specss;
  35.767 +    fun currys [] t = t
  35.768 +      | currys Ts t = t $ mk_tuple1 (List.rev Ts) (map Bound (length Ts - 1 downto 0))
  35.769 +          |> fold_rev (Term.abs o pair Name.uu) Ts;
  35.770 +
  35.771 +(*
  35.772 +val _ = tracing ("corecursor arguments:\n    \<cdot> " ^
  35.773 + space_implode "\n    \<cdot> " (map (Syntax.string_of_term lthy) corec_args));
  35.774 +*)
  35.775 +
  35.776 +    val exclss' =
  35.777 +      disc_eqnss
  35.778 +      |> map (map (fn x => (#fun_args x, #ctr_no x, #prems x, #auto_gen x))
  35.779 +        #> fst o (fn xs => fold_map (fn x => fn ys => ((x, ys), ys @ [x])) xs [])
  35.780 +        #> maps (uncurry (map o pair)
  35.781 +          #> map (fn ((fun_args, c, x, a), (_, c', y, a')) =>
  35.782 +              ((c, c', a orelse a'), (x, s_not (s_conjs y)))
  35.783 +            ||> apfst (map HOLogic.mk_Trueprop) o apsnd HOLogic.mk_Trueprop
  35.784 +            ||> Logic.list_implies
  35.785 +            ||> curry Logic.list_all (map dest_Free fun_args))))
  35.786 +  in
  35.787 +    map (list_comb o rpair corec_args) corecs
  35.788 +    |> map2 (fn Ts => fn t => if length Ts = 0 then t $ HOLogic.unit else t) arg_Tss
  35.789 +    |> map2 currys arg_Tss
  35.790 +    |> Syntax.check_terms lthy
  35.791 +    |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.conceal (Thm.def_binding b), []), t)))
  35.792 +      bs mxs
  35.793 +    |> rpair exclss'
  35.794 +  end;
  35.795 +
  35.796 +fun mk_real_disc_eqns fun_binding arg_Ts ({ctr_specs, ...} : corec_spec)
  35.797 +    (sel_eqns : coeqn_data_sel list) (disc_eqns : coeqn_data_disc list) =
  35.798 +  if length disc_eqns <> length ctr_specs - 1 then disc_eqns else
  35.799 +    let
  35.800 +      val n = 0 upto length ctr_specs
  35.801 +        |> the o find_first (fn idx => not (exists (equal idx o #ctr_no) disc_eqns));
  35.802 +      val fun_args = (try (#fun_args o hd) disc_eqns, try (#fun_args o hd) sel_eqns)
  35.803 +        |> the_default (map (curry Free Name.uu) arg_Ts) o merge_options;
  35.804 +      val extra_disc_eqn = {
  35.805 +        fun_name = Binding.name_of fun_binding,
  35.806 +        fun_T = arg_Ts ---> body_type (fastype_of (#ctr (hd ctr_specs))),
  35.807 +        fun_args = fun_args,
  35.808 +        ctr = #ctr (nth ctr_specs n),
  35.809 +        ctr_no = n,
  35.810 +        disc = #disc (nth ctr_specs n),
  35.811 +        prems = maps (s_not_conj o #prems) disc_eqns,
  35.812 +        auto_gen = true,
  35.813 +        maybe_ctr_rhs = NONE,
  35.814 +        maybe_code_rhs = NONE,
  35.815 +        user_eqn = undef_const};
  35.816 +    in
  35.817 +      chop n disc_eqns ||> cons extra_disc_eqn |> (op @)
  35.818 +    end;
  35.819 +
  35.820 +fun find_corec_calls ctxt has_call basic_ctr_specs ({ctr, sel, rhs_term, ...} : coeqn_data_sel) =
  35.821 +  let
  35.822 +    val sel_no = find_first (equal ctr o #ctr) basic_ctr_specs
  35.823 +      |> find_index (equal sel) o #sels o the;
  35.824 +    fun find t = if has_call t then snd (fold_rev_corec_call ctxt (K cons) [] t []) else [];
  35.825 +  in
  35.826 +    find rhs_term
  35.827 +    |> K |> nth_map sel_no |> AList.map_entry (op =) ctr
  35.828 +  end;
  35.829 +
  35.830 +fun add_primcorec_ursive maybe_tac seq fixes specs maybe_of_specs lthy =
  35.831 +  let
  35.832 +    val thy = Proof_Context.theory_of lthy;
  35.833 +
  35.834 +    val (bs, mxs) = map_split (apfst fst) fixes;
  35.835 +    val (arg_Ts, res_Ts) = map (strip_type o snd o fst #>> HOLogic.mk_tupleT) fixes |> split_list;
  35.836 +
  35.837 +    val _ = (case filter_out (fn (_, T) => Sign.of_sort thy (T, HOLogic.typeS)) (bs ~~ arg_Ts) of
  35.838 +        [] => ()
  35.839 +      | (b, _) :: _ => primcorec_error ("type of " ^ Binding.print b ^ " contains top sort"));
  35.840 +
  35.841 +    val fun_names = map Binding.name_of bs;
  35.842 +    val basic_ctr_specss = map (basic_corec_specs_of lthy) res_Ts;
  35.843 +    val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
  35.844 +    val eqns_data =
  35.845 +      fold_map2 (dissect_coeqn lthy seq has_call fun_names basic_ctr_specss) (map snd specs)
  35.846 +        maybe_of_specs []
  35.847 +      |> flat o fst;
  35.848 +
  35.849 +    val callssss =
  35.850 +      map_filter (try (fn Sel x => x)) eqns_data
  35.851 +      |> partition_eq ((op =) o pairself #fun_name)
  35.852 +      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
  35.853 +      |> map (flat o snd)
  35.854 +      |> map2 (fold o find_corec_calls lthy has_call) basic_ctr_specss
  35.855 +      |> map2 (curry (op |>)) (map (map (fn {ctr, sels, ...} =>
  35.856 +        (ctr, map (K []) sels))) basic_ctr_specss);
  35.857 +
  35.858 +(*
  35.859 +val _ = tracing ("callssss = " ^ @{make_string} callssss);
  35.860 +*)
  35.861 +
  35.862 +    val ((n2m, corec_specs', _, coinduct_thm, strong_coinduct_thm, coinduct_thms,
  35.863 +          strong_coinduct_thms), lthy') =
  35.864 +      corec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
  35.865 +    val actual_nn = length bs;
  35.866 +    val corec_specs = take actual_nn corec_specs'; (*###*)
  35.867 +    val ctr_specss = map #ctr_specs corec_specs;
  35.868 +
  35.869 +    val disc_eqnss' = map_filter (try (fn Disc x => x)) eqns_data
  35.870 +      |> partition_eq ((op =) o pairself #fun_name)
  35.871 +      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
  35.872 +      |> map (sort ((op <) o pairself #ctr_no |> make_ord) o flat o snd);
  35.873 +    val _ = disc_eqnss' |> map (fn x =>
  35.874 +      let val d = duplicates ((op =) o pairself #ctr_no) x in null d orelse
  35.875 +        primcorec_error_eqns "excess discriminator formula in definition"
  35.876 +          (maps (fn t => filter (equal (#ctr_no t) o #ctr_no) x) d |> map #user_eqn) end);
  35.877 +
  35.878 +    val sel_eqnss = map_filter (try (fn Sel x => x)) eqns_data
  35.879 +      |> partition_eq ((op =) o pairself #fun_name)
  35.880 +      |> fst o finds (fn (x, ({fun_name, ...} :: _)) => x = fun_name) fun_names
  35.881 +      |> map (flat o snd);
  35.882 +
  35.883 +    val arg_Tss = map (binder_types o snd o fst) fixes;
  35.884 +    val disc_eqnss = map5 mk_real_disc_eqns bs arg_Tss corec_specs sel_eqnss disc_eqnss';
  35.885 +    val (defs, exclss') =
  35.886 +      build_codefs lthy' bs mxs has_call arg_Tss corec_specs disc_eqnss sel_eqnss;
  35.887 +
  35.888 +    fun excl_tac (c, c', a) =
  35.889 +      if a orelse c = c' orelse seq then SOME (K (HEADGOAL (mk_primcorec_assumption_tac lthy [])))
  35.890 +      else maybe_tac;
  35.891 +
  35.892 +(*
  35.893 +val _ = tracing ("exclusiveness properties:\n    \<cdot> " ^
  35.894 + space_implode "\n    \<cdot> " (maps (map (Syntax.string_of_term lthy o snd)) exclss'));
  35.895 +*)
  35.896 +
  35.897 +    val exclss'' = exclss' |> map (map (fn (idx, t) =>
  35.898 +      (idx, (Option.map (Goal.prove lthy [] [] t #> Thm.close_derivation) (excl_tac idx), t))));
  35.899 +    val taut_thmss = map (map (apsnd (the o fst)) o filter (is_some o fst o snd)) exclss'';
  35.900 +    val (goal_idxss, goalss) = exclss''
  35.901 +      |> map (map (apsnd (rpair [] o snd)) o filter (is_none o fst o snd))
  35.902 +      |> split_list o map split_list;
  35.903 +
  35.904 +    fun prove thmss' def_thms' lthy =
  35.905 +      let
  35.906 +        val def_thms = map (snd o snd) def_thms';
  35.907 +
  35.908 +        val exclss' = map (op ~~) (goal_idxss ~~ thmss');
  35.909 +        fun mk_exclsss excls n =
  35.910 +          (excls, map (fn k => replicate k [TrueI] @ replicate (n - k) []) (0 upto n - 1))
  35.911 +          |-> fold (fn ((c, c', _), thm) => nth_map c (nth_map c' (K [thm])));
  35.912 +        val exclssss = (exclss' ~~ taut_thmss |> map (op @), fun_names ~~ corec_specs)
  35.913 +          |-> map2 (fn excls => fn (_, {ctr_specs, ...}) => mk_exclsss excls (length ctr_specs));
  35.914 +
  35.915 +        fun prove_disc ({ctr_specs, ...} : corec_spec) exclsss
  35.916 +            ({fun_name, fun_T, fun_args, ctr_no, prems, ...} : coeqn_data_disc) =
  35.917 +          if Term.aconv_untyped (#disc (nth ctr_specs ctr_no), @{term "\<lambda>x. x = x"}) then
  35.918 +            []
  35.919 +          else
  35.920 +            let
  35.921 +              val {disc_corec, ...} = nth ctr_specs ctr_no;
  35.922 +              val k = 1 + ctr_no;
  35.923 +              val m = length prems;
  35.924 +              val t =
  35.925 +                list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
  35.926 +                |> curry betapply (#disc (nth ctr_specs ctr_no)) (*###*)
  35.927 +                |> HOLogic.mk_Trueprop
  35.928 +                |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
  35.929 +                |> curry Logic.list_all (map dest_Free fun_args);
  35.930 +            in
  35.931 +              if prems = [@{term False}] then [] else
  35.932 +              mk_primcorec_disc_tac lthy def_thms disc_corec k m exclsss
  35.933 +              |> K |> Goal.prove lthy [] [] t
  35.934 +              |> Thm.close_derivation
  35.935 +              |> pair (#disc (nth ctr_specs ctr_no))
  35.936 +              |> single
  35.937 +            end;
  35.938 +
  35.939 +        fun prove_sel ({nested_map_idents, nested_map_comps, ctr_specs, ...} : corec_spec)
  35.940 +            (disc_eqns : coeqn_data_disc list) exclsss
  35.941 +            ({fun_name, fun_T, fun_args, ctr, sel, rhs_term, ...} : coeqn_data_sel) =
  35.942 +          let
  35.943 +            val SOME ctr_spec = find_first (equal ctr o #ctr) ctr_specs;
  35.944 +            val ctr_no = find_index (equal ctr o #ctr) ctr_specs;
  35.945 +            val prems = the_default (maps (s_not_conj o #prems) disc_eqns)
  35.946 +                (find_first (equal ctr_no o #ctr_no) disc_eqns |> Option.map #prems);
  35.947 +            val sel_corec = find_index (equal sel) (#sels ctr_spec)
  35.948 +              |> nth (#sel_corecs ctr_spec);
  35.949 +            val k = 1 + ctr_no;
  35.950 +            val m = length prems;
  35.951 +            val t =
  35.952 +              list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0))
  35.953 +              |> curry betapply sel
  35.954 +              |> rpair (abstract (List.rev fun_args) rhs_term)
  35.955 +              |> HOLogic.mk_Trueprop o HOLogic.mk_eq
  35.956 +              |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
  35.957 +              |> curry Logic.list_all (map dest_Free fun_args);
  35.958 +            val (distincts, _, sel_splits, sel_split_asms) = case_thms_of_term lthy [] rhs_term;
  35.959 +          in
  35.960 +            mk_primcorec_sel_tac lthy def_thms distincts sel_splits sel_split_asms nested_map_idents
  35.961 +              nested_map_comps sel_corec k m exclsss
  35.962 +            |> K |> Goal.prove lthy [] [] t
  35.963 +            |> Thm.close_derivation
  35.964 +            |> pair sel
  35.965 +          end;
  35.966 +
  35.967 +        fun prove_ctr disc_alist sel_alist (disc_eqns : coeqn_data_disc list)
  35.968 +            (sel_eqns : coeqn_data_sel list) ({ctr, disc, sels, collapse, ...} : corec_ctr_spec) =
  35.969 +          (* don't try to prove theorems when some sel_eqns are missing *)
  35.970 +          if not (exists (equal ctr o #ctr) disc_eqns)
  35.971 +              andalso not (exists (equal ctr o #ctr) sel_eqns)
  35.972 +            orelse
  35.973 +              filter (equal ctr o #ctr) sel_eqns
  35.974 +              |> fst o finds ((op =) o apsnd #sel) sels
  35.975 +              |> exists (null o snd)
  35.976 +          then [] else
  35.977 +            let
  35.978 +              val (fun_name, fun_T, fun_args, prems, maybe_rhs) =
  35.979 +                (find_first (equal ctr o #ctr) disc_eqns, find_first (equal ctr o #ctr) sel_eqns)
  35.980 +                |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #prems x,
  35.981 +                  #maybe_ctr_rhs x))
  35.982 +                ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, [], NONE))
  35.983 +                |> the o merge_options;
  35.984 +              val m = length prems;
  35.985 +              val t = (if is_some maybe_rhs then the maybe_rhs else
  35.986 +                  filter (equal ctr o #ctr) sel_eqns
  35.987 +                  |> fst o finds ((op =) o apsnd #sel) sels
  35.988 +                  |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x)) #-> abstract)
  35.989 +                  |> curry list_comb ctr)
  35.990 +                |> curry HOLogic.mk_eq (list_comb (Free (fun_name, fun_T),
  35.991 +                  map Bound (length fun_args - 1 downto 0)))
  35.992 +                |> HOLogic.mk_Trueprop
  35.993 +                |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
  35.994 +                |> curry Logic.list_all (map dest_Free fun_args);
  35.995 +              val maybe_disc_thm = AList.lookup (op =) disc_alist disc;
  35.996 +              val sel_thms = map snd (filter (member (op =) sels o fst) sel_alist);
  35.997 +            in
  35.998 +              if prems = [@{term False}] then [] else
  35.999 +                mk_primcorec_ctr_of_dtr_tac lthy m collapse maybe_disc_thm sel_thms
 35.1000 +                |> K |> Goal.prove lthy [] [] t
 35.1001 +                |> Thm.close_derivation
 35.1002 +                |> pair ctr
 35.1003 +                |> single
 35.1004 +            end;
 35.1005 +
 35.1006 +        fun prove_code disc_eqns sel_eqns ctr_alist ctr_specs =
 35.1007 +          let
 35.1008 +            val (fun_name, fun_T, fun_args, maybe_rhs) =
 35.1009 +              (find_first (member (op =) (map #ctr ctr_specs) o #ctr) disc_eqns,
 35.1010 +               find_first (member (op =) (map #ctr ctr_specs) o #ctr) sel_eqns)
 35.1011 +              |>> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, #maybe_code_rhs x))
 35.1012 +              ||> Option.map (fn x => (#fun_name x, #fun_T x, #fun_args x, NONE))
 35.1013 +              |> the o merge_options;
 35.1014 +
 35.1015 +            val bound_Ts = List.rev (map fastype_of fun_args);
 35.1016 +
 35.1017 +            val lhs = list_comb (Free (fun_name, fun_T), map Bound (length fun_args - 1 downto 0));
 35.1018 +            val maybe_rhs_info =
 35.1019 +              (case maybe_rhs of
 35.1020 +                SOME rhs =>
 35.1021 +                let
 35.1022 +                  val raw_rhs = expand_corec_code_rhs lthy has_call bound_Ts rhs;
 35.1023 +                  val cond_ctrs =
 35.1024 +                    fold_rev_corec_code_rhs lthy (K oo (cons oo pair)) bound_Ts raw_rhs [];
 35.1025 +                  val ctr_thms = map (the o AList.lookup (op =) ctr_alist o snd) cond_ctrs;
 35.1026 +                in SOME (rhs, raw_rhs, ctr_thms) end
 35.1027 +              | NONE =>
 35.1028 +                let
 35.1029 +                  fun prove_code_ctr {ctr, sels, ...} =
 35.1030 +                    if not (exists (equal ctr o fst) ctr_alist) then NONE else
 35.1031 +                      let
 35.1032 +                        val prems = find_first (equal ctr o #ctr) disc_eqns
 35.1033 +                          |> Option.map #prems |> the_default [];
 35.1034 +                        val t =
 35.1035 +                          filter (equal ctr o #ctr) sel_eqns
 35.1036 +                          |> fst o finds ((op =) o apsnd #sel) sels
 35.1037 +                          |> map (snd #> (fn [x] => (List.rev (#fun_args x), #rhs_term x))
 35.1038 +                            #-> abstract)
 35.1039 +                          |> curry list_comb ctr;
 35.1040 +                      in
 35.1041 +                        SOME (prems, t)
 35.1042 +                      end;
 35.1043 +                  val maybe_ctr_conds_argss = map prove_code_ctr ctr_specs;
 35.1044 +                in
 35.1045 +                  if exists is_none maybe_ctr_conds_argss then NONE else
 35.1046 +                    let
 35.1047 +                      val rhs = fold_rev (fn SOME (prems, u) => fn t => mk_If (s_conjs prems) u t)
 35.1048 +                        maybe_ctr_conds_argss
 35.1049 +                        (Const (@{const_name Code.abort}, @{typ String.literal} -->
 35.1050 +                            (@{typ unit} --> body_type fun_T) --> body_type fun_T) $
 35.1051 +                          HOLogic.mk_literal fun_name $
 35.1052 +                          absdummy @{typ unit} (incr_boundvars 1 lhs));
 35.1053 +                    in SOME (rhs, rhs, map snd ctr_alist) end
 35.1054 +                end);
 35.1055 +          in
 35.1056 +            (case maybe_rhs_info of
 35.1057 +              NONE => []
 35.1058 +            | SOME (rhs, raw_rhs, ctr_thms) =>
 35.1059 +              let
 35.1060 +                val ms = map (Logic.count_prems o prop_of) ctr_thms;
 35.1061 +                val (raw_t, t) = (raw_rhs, rhs)
 35.1062 +                  |> pairself
 35.1063 +                    (curry HOLogic.mk_eq (list_comb (Free (fun_name, fun_T),
 35.1064 +                      map Bound (length fun_args - 1 downto 0)))
 35.1065 +                    #> HOLogic.mk_Trueprop
 35.1066 +                    #> curry Logic.list_all (map dest_Free fun_args));
 35.1067 +                val (distincts, discIs, sel_splits, sel_split_asms) =
 35.1068 +                  case_thms_of_term lthy bound_Ts raw_rhs;
 35.1069 +
 35.1070 +                val raw_code_thm = mk_primcorec_raw_code_of_ctr_tac lthy distincts discIs sel_splits
 35.1071 +                    sel_split_asms ms ctr_thms
 35.1072 +                  |> K |> Goal.prove lthy [] [] raw_t
 35.1073 +                  |> Thm.close_derivation;
 35.1074 +              in
 35.1075 +                mk_primcorec_code_of_raw_code_tac lthy distincts sel_splits raw_code_thm
 35.1076 +                |> K |> Goal.prove lthy [] [] t
 35.1077 +                |> Thm.close_derivation
 35.1078 +                |> single
 35.1079 +              end)
 35.1080 +          end;
 35.1081 +
 35.1082 +        val disc_alists = map3 (maps oo prove_disc) corec_specs exclssss disc_eqnss;
 35.1083 +        val sel_alists = map4 (map ooo prove_sel) corec_specs disc_eqnss exclssss sel_eqnss;
 35.1084 +        val disc_thmss = map (map snd) disc_alists;
 35.1085 +        val sel_thmss = map (map snd) sel_alists;
 35.1086 +
 35.1087 +        val ctr_alists = map5 (maps oooo prove_ctr) disc_alists sel_alists disc_eqnss sel_eqnss
 35.1088 +          ctr_specss;
 35.1089 +        val ctr_thmss = map (map snd) ctr_alists;
 35.1090 +
 35.1091 +        val code_thmss = map4 prove_code disc_eqnss sel_eqnss ctr_alists ctr_specss;
 35.1092 +
 35.1093 +        val simp_thmss = map2 append disc_thmss sel_thmss
 35.1094 +
 35.1095 +        val common_name = mk_common_name fun_names;
 35.1096 +
 35.1097 +        val notes =
 35.1098 +          [(coinductN, map (if n2m then single else K []) coinduct_thms, []),
 35.1099 +           (codeN, code_thmss, code_nitpicksimp_attrs),
 35.1100 +           (ctrN, ctr_thmss, []),
 35.1101 +           (discN, disc_thmss, simp_attrs),
 35.1102 +           (selN, sel_thmss, simp_attrs),
 35.1103 +           (simpsN, simp_thmss, []),
 35.1104 +           (strong_coinductN, map (if n2m then single else K []) strong_coinduct_thms, [])]
 35.1105 +          |> maps (fn (thmN, thmss, attrs) =>
 35.1106 +            map2 (fn fun_name => fn thms =>
 35.1107 +                ((Binding.qualify true fun_name (Binding.name thmN), attrs), [(thms, [])]))
 35.1108 +              fun_names (take actual_nn thmss))
 35.1109 +          |> filter_out (null o fst o hd o snd);
 35.1110 +
 35.1111 +        val common_notes =
 35.1112 +          [(coinductN, if n2m then [coinduct_thm] else [], []),
 35.1113 +           (strong_coinductN, if n2m then [strong_coinduct_thm] else [], [])]
 35.1114 +          |> filter_out (null o #2)
 35.1115 +          |> map (fn (thmN, thms, attrs) =>
 35.1116 +            ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
 35.1117 +      in
 35.1118 +        lthy |> Local_Theory.notes (notes @ common_notes) |> snd
 35.1119 +      end;
 35.1120 +
 35.1121 +    fun after_qed thmss' = fold_map Local_Theory.define defs #-> prove thmss';
 35.1122 +  in
 35.1123 +    (goalss, after_qed, lthy')
 35.1124 +  end;
 35.1125 +
 35.1126 +fun add_primcorec_ursive_cmd maybe_tac seq (raw_fixes, raw_specs') lthy =
 35.1127 +  let
 35.1128 +    val (raw_specs, maybe_of_specs) =
 35.1129 +      split_list raw_specs' ||> map (Option.map (Syntax.read_term lthy));
 35.1130 +    val ((fixes, specs), _) = Specification.read_spec raw_fixes raw_specs lthy;
 35.1131 +  in
 35.1132 +    add_primcorec_ursive maybe_tac seq fixes specs maybe_of_specs lthy
 35.1133 +    handle ERROR str => primcorec_error str
 35.1134 +  end
 35.1135 +  handle Primcorec_Error (str, eqns) =>
 35.1136 +    if null eqns
 35.1137 +    then error ("primcorec error:\n  " ^ str)
 35.1138 +    else error ("primcorec error:\n  " ^ str ^ "\nin\n  " ^
 35.1139 +      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
 35.1140 +
 35.1141 +val add_primcorecursive_cmd = (fn (goalss, after_qed, lthy) =>
 35.1142 +  lthy
 35.1143 +  |> Proof.theorem NONE after_qed goalss
 35.1144 +  |> Proof.refine (Method.primitive_text I)
 35.1145 +  |> Seq.hd) ooo add_primcorec_ursive_cmd NONE;
 35.1146 +
 35.1147 +val add_primcorec_cmd = (fn (goalss, after_qed, lthy) =>
 35.1148 +  lthy
 35.1149 +  |> after_qed (map (fn [] => []
 35.1150 +      | _ => primcorec_error "need exclusiveness proofs - use primcorecursive instead of primcorec")
 35.1151 +    goalss)) ooo add_primcorec_ursive_cmd (SOME (fn {context = ctxt, ...} => auto_tac ctxt));
 35.1152 +
 35.1153 +end;
    36.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    36.2 +++ b/src/HOL/BNF/Tools/bnf_gfp_rec_sugar_tactics.ML	Mon Nov 11 17:44:21 2013 +0100
    36.3 @@ -0,0 +1,135 @@
    36.4 +(*  Title:      HOL/BNF/Tools/bnf_gfp_rec_sugar_tactics.ML
    36.5 +    Author:     Jasmin Blanchette, TU Muenchen
    36.6 +    Copyright   2013
    36.7 +
    36.8 +Tactics for corecursor sugar.
    36.9 +*)
   36.10 +
   36.11 +signature BNF_GFP_REC_SUGAR_TACTICS =
   36.12 +sig
   36.13 +  val mk_primcorec_assumption_tac: Proof.context -> thm list -> int -> tactic
   36.14 +  val mk_primcorec_code_of_raw_code_tac: Proof.context -> thm list -> thm list -> thm -> tactic
   36.15 +  val mk_primcorec_ctr_of_dtr_tac: Proof.context -> int -> thm -> thm option -> thm list -> tactic
   36.16 +  val mk_primcorec_disc_tac: Proof.context -> thm list -> thm -> int -> int -> thm list list list ->
   36.17 +    tactic
   36.18 +  val mk_primcorec_raw_code_of_ctr_tac: Proof.context -> thm list -> thm list -> thm list ->
   36.19 +    thm list -> int list -> thm list -> tactic
   36.20 +  val mk_primcorec_sel_tac: Proof.context -> thm list -> thm list -> thm list -> thm list ->
   36.21 +    thm list -> thm list -> thm -> int -> int -> thm list list list -> tactic
   36.22 +end;
   36.23 +
   36.24 +structure BNF_GFP_Rec_Sugar_Tactics : BNF_GFP_REC_SUGAR_TACTICS =
   36.25 +struct
   36.26 +
   36.27 +open BNF_Util
   36.28 +open BNF_Tactics
   36.29 +
   36.30 +val falseEs = @{thms not_TrueE FalseE};
   36.31 +val Let_def = @{thm Let_def};
   36.32 +val neq_eq_eq_contradict = @{thm neq_eq_eq_contradict};
   36.33 +val split_if = @{thm split_if};
   36.34 +val split_if_asm = @{thm split_if_asm};
   36.35 +val split_connectI = @{thms allI impI conjI};
   36.36 +
   36.37 +fun mk_primcorec_assumption_tac ctxt discIs =
   36.38 +  SELECT_GOAL (unfold_thms_tac ctxt
   36.39 +      @{thms not_not not_False_eq_True not_True_eq_False de_Morgan_conj de_Morgan_disj} THEN
   36.40 +    SOLVE (HEADGOAL (REPEAT o (rtac refl ORELSE' atac ORELSE' etac conjE ORELSE'
   36.41 +    eresolve_tac falseEs ORELSE'
   36.42 +    resolve_tac @{thms TrueI conjI disjI1 disjI2} ORELSE'
   36.43 +    dresolve_tac discIs THEN' atac ORELSE'
   36.44 +    etac notE THEN' atac ORELSE'
   36.45 +    etac disjE))));
   36.46 +
   36.47 +fun mk_primcorec_same_case_tac m =
   36.48 +  HEADGOAL (if m = 0 then rtac TrueI
   36.49 +    else REPEAT_DETERM_N (m - 1) o (rtac conjI THEN' atac) THEN' atac);
   36.50 +
   36.51 +fun mk_primcorec_different_case_tac ctxt m excl =
   36.52 +  HEADGOAL (if m = 0 then mk_primcorec_assumption_tac ctxt []
   36.53 +    else dtac excl THEN' (REPEAT_DETERM_N (m - 1) o atac) THEN' mk_primcorec_assumption_tac ctxt []);
   36.54 +
   36.55 +fun mk_primcorec_cases_tac ctxt k m exclsss =
   36.56 +  let val n = length exclsss in
   36.57 +    EVERY (map (fn [] => if k = n then all_tac else mk_primcorec_same_case_tac m
   36.58 +        | [excl] => mk_primcorec_different_case_tac ctxt m excl)
   36.59 +      (take k (nth exclsss (k - 1))))
   36.60 +  end;
   36.61 +
   36.62 +fun mk_primcorec_prelude ctxt defs thm =
   36.63 +  unfold_thms_tac ctxt defs THEN HEADGOAL (rtac thm) THEN
   36.64 +  unfold_thms_tac ctxt @{thms Let_def split};
   36.65 +
   36.66 +fun mk_primcorec_disc_tac ctxt defs disc_corec k m exclsss =
   36.67 +  mk_primcorec_prelude ctxt defs disc_corec THEN mk_primcorec_cases_tac ctxt k m exclsss;
   36.68 +
   36.69 +fun mk_primcorec_sel_tac ctxt defs distincts splits split_asms map_idents map_comps f_sel k m
   36.70 +    exclsss =
   36.71 +  mk_primcorec_prelude ctxt defs (f_sel RS trans) THEN
   36.72 +  mk_primcorec_cases_tac ctxt k m exclsss THEN
   36.73 +  HEADGOAL (REPEAT_DETERM o (rtac refl ORELSE' rtac ext ORELSE'
   36.74 +    eresolve_tac falseEs ORELSE'
   36.75 +    resolve_tac split_connectI ORELSE'
   36.76 +    Splitter.split_asm_tac (split_if_asm :: split_asms) ORELSE'
   36.77 +    Splitter.split_tac (split_if :: splits) ORELSE'
   36.78 +    eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac ORELSE'
   36.79 +    etac notE THEN' atac ORELSE'
   36.80 +    (CHANGED o SELECT_GOAL (unfold_thms_tac ctxt
   36.81 +      (@{thms id_def o_def split_def sum.cases} @ map_comps @ map_idents)))));
   36.82 +
   36.83 +fun mk_primcorec_ctr_of_dtr_tac ctxt m collapse maybe_disc_f sel_fs =
   36.84 +  HEADGOAL (rtac ((if null sel_fs then collapse else collapse RS sym) RS trans) THEN'
   36.85 +    (the_default (K all_tac) (Option.map rtac maybe_disc_f)) THEN' REPEAT_DETERM_N m o atac) THEN
   36.86 +  unfold_thms_tac ctxt (Let_def :: sel_fs) THEN HEADGOAL (rtac refl);
   36.87 +
   36.88 +fun inst_split_eq ctxt split =
   36.89 +  (case prop_of split of
   36.90 +    @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ (Var (_, Type (_, [T, _])) $ _) $ _) =>
   36.91 +    let
   36.92 +      val s = Name.uu;
   36.93 +      val eq = Abs (Name.uu, T, HOLogic.mk_eq (Free (s, T), Bound 0));
   36.94 +      val split' = Drule.instantiate' [] [SOME (certify ctxt eq)] split;
   36.95 +    in
   36.96 +      Thm.generalize ([], [s]) (Thm.maxidx_of split' + 1) split'
   36.97 +    end
   36.98 +  | _ => split);
   36.99 +
  36.100 +fun distinct_in_prems_tac distincts =
  36.101 +  eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac;
  36.102 +
  36.103 +(* TODO: reduce code duplication with selector tactic above *)
  36.104 +fun mk_primcorec_raw_code_of_ctr_single_tac ctxt distincts discIs splits split_asms m f_ctr =
  36.105 +  let
  36.106 +    val splits' =
  36.107 +      map (fn th => th RS iffD2) (@{thm split_if_eq2} :: map (inst_split_eq ctxt) splits)
  36.108 +  in
  36.109 +    HEADGOAL (REPEAT o (resolve_tac (splits' @ split_connectI))) THEN
  36.110 +    mk_primcorec_prelude ctxt [] (f_ctr RS trans) THEN
  36.111 +    HEADGOAL ((REPEAT_DETERM_N m o mk_primcorec_assumption_tac ctxt discIs) THEN'
  36.112 +      SELECT_GOAL (SOLVE (HEADGOAL (REPEAT_DETERM o
  36.113 +      (rtac refl ORELSE' atac ORELSE'
  36.114 +       resolve_tac (@{thm Code.abort_def} :: split_connectI) ORELSE'
  36.115 +       Splitter.split_tac (split_if :: splits) ORELSE'
  36.116 +       Splitter.split_asm_tac (split_if_asm :: split_asms) ORELSE'
  36.117 +       mk_primcorec_assumption_tac ctxt discIs ORELSE'
  36.118 +       distinct_in_prems_tac distincts ORELSE'
  36.119 +       (TRY o dresolve_tac discIs) THEN' etac notE THEN' atac)))))
  36.120 +  end;
  36.121 +
  36.122 +fun mk_primcorec_raw_code_of_ctr_tac ctxt distincts discIs splits split_asms ms f_ctrs =
  36.123 +  EVERY (map2 (mk_primcorec_raw_code_of_ctr_single_tac ctxt distincts discIs splits split_asms) ms
  36.124 +    f_ctrs) THEN
  36.125 +  IF_UNSOLVED (unfold_thms_tac ctxt @{thms Code.abort_def} THEN
  36.126 +    HEADGOAL (REPEAT_DETERM o resolve_tac (refl :: split_connectI)));
  36.127 +
  36.128 +fun mk_primcorec_code_of_raw_code_tac ctxt distincts splits raw =
  36.129 +  HEADGOAL (rtac raw ORELSE' rtac (raw RS trans) THEN'
  36.130 +    SELECT_GOAL (unfold_thms_tac ctxt [Let_def]) THEN' REPEAT_DETERM o
  36.131 +    (rtac refl ORELSE' atac ORELSE'
  36.132 +     resolve_tac split_connectI ORELSE'
  36.133 +     Splitter.split_tac (split_if :: splits) ORELSE'
  36.134 +     distinct_in_prems_tac distincts ORELSE'
  36.135 +     rtac sym THEN' atac ORELSE'
  36.136 +     etac notE THEN' atac));
  36.137 +
  36.138 +end;
    37.1 --- a/src/HOL/BNF/Tools/bnf_lfp.ML	Mon Nov 11 17:34:44 2013 +0100
    37.2 +++ b/src/HOL/BNF/Tools/bnf_lfp.ML	Mon Nov 11 17:44:21 2013 +0100
    37.3 @@ -22,7 +22,7 @@
    37.4  open BNF_Comp
    37.5  open BNF_FP_Util
    37.6  open BNF_FP_Def_Sugar
    37.7 -open BNF_FP_Rec_Sugar
    37.8 +open BNF_LFP_Rec_Sugar
    37.9  open BNF_LFP_Util
   37.10  open BNF_LFP_Tactics
   37.11  
    38.1 --- a/src/HOL/BNF/Tools/bnf_lfp_compat.ML	Mon Nov 11 17:34:44 2013 +0100
    38.2 +++ b/src/HOL/BNF/Tools/bnf_lfp_compat.ML	Mon Nov 11 17:44:21 2013 +0100
    38.3 @@ -57,8 +57,10 @@
    38.4      val fpTs as fpT1 :: _ = map (fn s => Type (s, As)) fpT_names';
    38.5  
    38.6      fun add_nested_types_of (T as Type (s, _)) seen =
    38.7 -      if member (op =) seen T orelse s = @{type_name fun} then
    38.8 +      if member (op =) seen T then
    38.9          seen
   38.10 +      else if s = @{type_name fun} then
   38.11 +        (warning "Partial support for recursion through functions -- 'primrec' will fail"; seen)
   38.12        else
   38.13          (case try lfp_sugar_of s of
   38.14            SOME ({T = T0, fp_res = {Ts = mutual_Ts0, ...}, ctr_sugars, ...}) =>
   38.15 @@ -91,11 +93,13 @@
   38.16      val nn = length Ts;
   38.17      val get_indices = K [];
   38.18      val fp_sugars0 = if nn = 1 then [fp_sugar0] else map (lfp_sugar_of o fst o dest_Type) Ts;
   38.19 -    val callssss = pad_and_indexify_calls fp_sugars0 nn [];
   38.20 -    val has_nested = nn > nn_fp;
   38.21 +    val callssss = map (fn fp_sugar0 => indexify_callsss fp_sugar0 []) fp_sugars0;
   38.22  
   38.23      val ((fp_sugars, (lfp_sugar_thms, _)), lthy) =
   38.24 -      mutualize_fp_sugars has_nested Least_FP compat_bs Ts get_indices callssss fp_sugars0 lthy;
   38.25 +      if nn > nn_fp then
   38.26 +        mutualize_fp_sugars Least_FP compat_bs Ts get_indices callssss fp_sugars0 lthy
   38.27 +      else
   38.28 +        ((fp_sugars0, (NONE, NONE)), lthy);
   38.29  
   38.30      val {ctr_sugars, co_inducts = [induct], co_iterss, co_iter_thmsss = iter_thmsss, ...} :: _ =
   38.31        fp_sugars;
    39.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    39.2 +++ b/src/HOL/BNF/Tools/bnf_lfp_rec_sugar.ML	Mon Nov 11 17:44:21 2013 +0100
    39.3 @@ -0,0 +1,604 @@
    39.4 +(*  Title:      HOL/BNF/Tools/bnf_lfp_rec_sugar.ML
    39.5 +    Author:     Lorenz Panny, TU Muenchen
    39.6 +    Author:     Jasmin Blanchette, TU Muenchen
    39.7 +    Copyright   2013
    39.8 +
    39.9 +Recursor sugar.
   39.10 +*)
   39.11 +
   39.12 +signature BNF_LFP_REC_SUGAR =
   39.13 +sig
   39.14 +  val add_primrec: (binding * typ option * mixfix) list ->
   39.15 +    (Attrib.binding * term) list -> local_theory -> (term list * thm list list) * local_theory
   39.16 +  val add_primrec_cmd: (binding * string option * mixfix) list ->
   39.17 +    (Attrib.binding * string) list -> local_theory -> (term list * thm list list) * local_theory
   39.18 +  val add_primrec_global: (binding * typ option * mixfix) list ->
   39.19 +    (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
   39.20 +  val add_primrec_overloaded: (string * (string * typ) * bool) list ->
   39.21 +    (binding * typ option * mixfix) list ->
   39.22 +    (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
   39.23 +  val add_primrec_simple: ((binding * typ) * mixfix) list -> term list ->
   39.24 +    local_theory -> (string list * (term list * (int list list * thm list list))) * local_theory
   39.25 +end;
   39.26 +
   39.27 +structure BNF_LFP_Rec_Sugar : BNF_LFP_REC_SUGAR =
   39.28 +struct
   39.29 +
   39.30 +open Ctr_Sugar
   39.31 +open BNF_Util
   39.32 +open BNF_Tactics
   39.33 +open BNF_Def
   39.34 +open BNF_FP_Util
   39.35 +open BNF_FP_Def_Sugar
   39.36 +open BNF_FP_N2M_Sugar
   39.37 +open BNF_FP_Rec_Sugar_Util
   39.38 +
   39.39 +val nitpicksimp_attrs = @{attributes [nitpick_simp]};
   39.40 +val simp_attrs = @{attributes [simp]};
   39.41 +val code_nitpicksimp_simp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs @ simp_attrs;
   39.42 +
   39.43 +exception Primrec_Error of string * term list;
   39.44 +
   39.45 +fun primrec_error str = raise Primrec_Error (str, []);
   39.46 +fun primrec_error_eqn str eqn = raise Primrec_Error (str, [eqn]);
   39.47 +fun primrec_error_eqns str eqns = raise Primrec_Error (str, eqns);
   39.48 +
   39.49 +datatype rec_call =
   39.50 +  No_Rec of int * typ |
   39.51 +  Mutual_Rec of (int * typ) * (int * typ) |
   39.52 +  Nested_Rec of int * typ;
   39.53 +
   39.54 +type rec_ctr_spec =
   39.55 +  {ctr: term,
   39.56 +   offset: int,
   39.57 +   calls: rec_call list,
   39.58 +   rec_thm: thm};
   39.59 +
   39.60 +type rec_spec =
   39.61 +  {recx: term,
   39.62 +   nested_map_idents: thm list,
   39.63 +   nested_map_comps: thm list,
   39.64 +   ctr_specs: rec_ctr_spec list};
   39.65 +
   39.66 +exception AINT_NO_MAP of term;
   39.67 +
   39.68 +fun ill_formed_rec_call ctxt t =
   39.69 +  error ("Ill-formed recursive call: " ^ quote (Syntax.string_of_term ctxt t));
   39.70 +fun invalid_map ctxt t =
   39.71 +  error ("Invalid map function in " ^ quote (Syntax.string_of_term ctxt t));
   39.72 +fun unexpected_rec_call ctxt t =
   39.73 +  error ("Unexpected recursive call: " ^ quote (Syntax.string_of_term ctxt t));
   39.74 +
   39.75 +fun massage_nested_rec_call ctxt has_call raw_massage_fun bound_Ts y y' =
   39.76 +  let
   39.77 +    fun check_no_call t = if has_call t then unexpected_rec_call ctxt t else ();
   39.78 +
   39.79 +    val typof = curry fastype_of1 bound_Ts;
   39.80 +    val build_map_fst = build_map ctxt (fst_const o fst);
   39.81 +
   39.82 +    val yT = typof y;
   39.83 +    val yU = typof y';
   39.84 +
   39.85 +    fun y_of_y' () = build_map_fst (yU, yT) $ y';
   39.86 +    val elim_y = Term.map_aterms (fn t => if t = y then y_of_y' () else t);
   39.87 +
   39.88 +    fun massage_mutual_fun U T t =
   39.89 +      (case t of
   39.90 +        Const (@{const_name comp}, _) $ t1 $ t2 =>
   39.91 +        mk_comp bound_Ts (tap check_no_call t1, massage_mutual_fun U T t2)
   39.92 +      | _ =>
   39.93 +        if has_call t then
   39.94 +          (case try HOLogic.dest_prodT U of
   39.95 +            SOME (U1, U2) => if U1 = T then raw_massage_fun T U2 t else invalid_map ctxt t
   39.96 +          | NONE => invalid_map ctxt t)
   39.97 +        else
   39.98 +          mk_comp bound_Ts (t, build_map_fst (U, T)));
   39.99 +
  39.100 +    fun massage_map (Type (_, Us)) (Type (s, Ts)) t =
  39.101 +        (case try (dest_map ctxt s) t of
  39.102 +          SOME (map0, fs) =>
  39.103 +          let
  39.104 +            val Type (_, ran_Ts) = range_type (typof t);
  39.105 +            val map' = mk_map (length fs) Us ran_Ts map0;
  39.106 +            val fs' = map_flattened_map_args ctxt s (map3 massage_map_or_map_arg Us Ts) fs;
  39.107 +          in
  39.108 +            Term.list_comb (map', fs')
  39.109 +          end
  39.110 +        | NONE => raise AINT_NO_MAP t)
  39.111 +      | massage_map _ _ t = raise AINT_NO_MAP t
  39.112 +    and massage_map_or_map_arg U T t =
  39.113 +      if T = U then
  39.114 +        tap check_no_call t
  39.115 +      else
  39.116 +        massage_map U T t
  39.117 +        handle AINT_NO_MAP _ => massage_mutual_fun U T t;
  39.118 +
  39.119 +    fun massage_call (t as t1 $ t2) =
  39.120 +        if has_call t then
  39.121 +          if t2 = y then
  39.122 +            massage_map yU yT (elim_y t1) $ y'
  39.123 +            handle AINT_NO_MAP t' => invalid_map ctxt t'
  39.124 +          else
  39.125 +            let val (g, xs) = Term.strip_comb t2 in
  39.126 +              if g = y then
  39.127 +                if exists has_call xs then unexpected_rec_call ctxt t2
  39.128 +                else Term.list_comb (massage_call (mk_compN (length xs) bound_Ts (t1, y)), xs)
  39.129 +              else
  39.130 +                ill_formed_rec_call ctxt t
  39.131 +            end
  39.132 +        else
  39.133 +          elim_y t
  39.134 +      | massage_call t = if t = y then y_of_y' () else ill_formed_rec_call ctxt t;
  39.135 +  in
  39.136 +    massage_call
  39.137 +  end;
  39.138 +
  39.139 +fun rec_specs_of bs arg_Ts res_Ts get_indices callssss0 lthy =
  39.140 +  let
  39.141 +    val thy = Proof_Context.theory_of lthy;
  39.142 +
  39.143 +    val ((missing_arg_Ts, perm0_kks,
  39.144 +          fp_sugars as {nested_bnfs, fp_res = {xtor_co_iterss = ctor_iters1 :: _, ...},
  39.145 +            co_inducts = [induct_thm], ...} :: _, (lfp_sugar_thms, _)), lthy') =
  39.146 +      nested_to_mutual_fps Least_FP bs arg_Ts get_indices callssss0 lthy;
  39.147 +
  39.148 +    val perm_fp_sugars = sort (int_ord o pairself #index) fp_sugars;
  39.149 +
  39.150 +    val indices = map #index fp_sugars;
  39.151 +    val perm_indices = map #index perm_fp_sugars;
  39.152 +
  39.153 +    val perm_ctrss = map (#ctrs o of_fp_sugar #ctr_sugars) perm_fp_sugars;
  39.154 +    val perm_ctr_Tsss = map (map (binder_types o fastype_of)) perm_ctrss;
  39.155 +    val perm_lfpTs = map (body_type o fastype_of o hd) perm_ctrss;
  39.156 +
  39.157 +    val nn0 = length arg_Ts;
  39.158 +    val nn = length perm_lfpTs;
  39.159 +    val kks = 0 upto nn - 1;
  39.160 +    val perm_ns = map length perm_ctr_Tsss;
  39.161 +    val perm_mss = map (map length) perm_ctr_Tsss;
  39.162 +
  39.163 +    val perm_Cs = map (body_type o fastype_of o co_rec_of o of_fp_sugar (#xtor_co_iterss o #fp_res))
  39.164 +      perm_fp_sugars;
  39.165 +    val perm_fun_arg_Tssss =
  39.166 +      mk_iter_fun_arg_types perm_ctr_Tsss perm_ns perm_mss (co_rec_of ctor_iters1);
  39.167 +
  39.168 +    fun unpermute0 perm0_xs = permute_like (op =) perm0_kks kks perm0_xs;
  39.169 +    fun unpermute perm_xs = permute_like (op =) perm_indices indices perm_xs;
  39.170 +
  39.171 +    val induct_thms = unpermute0 (conj_dests nn induct_thm);
  39.172 +
  39.173 +    val lfpTs = unpermute perm_lfpTs;
  39.174 +    val Cs = unpermute perm_Cs;
  39.175 +
  39.176 +    val As_rho = tvar_subst thy (take nn0 lfpTs) arg_Ts;
  39.177 +    val Cs_rho = map (fst o dest_TVar) Cs ~~ pad_list HOLogic.unitT nn res_Ts;
  39.178 +
  39.179 +    val substA = Term.subst_TVars As_rho;
  39.180 +    val substAT = Term.typ_subst_TVars As_rho;
  39.181 +    val substCT = Term.typ_subst_TVars Cs_rho;
  39.182 +    val substACT = substAT o substCT;
  39.183 +
  39.184 +    val perm_Cs' = map substCT perm_Cs;
  39.185 +
  39.186 +    fun offset_of_ctr 0 _ = 0
  39.187 +      | offset_of_ctr n (({ctrs, ...} : ctr_sugar) :: ctr_sugars) =
  39.188 +        length ctrs + offset_of_ctr (n - 1) ctr_sugars;
  39.189 +
  39.190 +    fun call_of [i] [T] = (if exists_subtype_in Cs T then Nested_Rec else No_Rec) (i, substACT T)
  39.191 +      | call_of [i, i'] [T, T'] = Mutual_Rec ((i, substACT T), (i', substACT T'));
  39.192 +
  39.193 +    fun mk_ctr_spec ctr offset fun_arg_Tss rec_thm =
  39.194 +      let
  39.195 +        val (fun_arg_hss, _) = indexedd fun_arg_Tss 0;
  39.196 +        val fun_arg_hs = flat_rec_arg_args fun_arg_hss;
  39.197 +        val fun_arg_iss = map (map (find_index_eq fun_arg_hs)) fun_arg_hss;
  39.198 +      in
  39.199 +        {ctr = substA ctr, offset = offset, calls = map2 call_of fun_arg_iss fun_arg_Tss,
  39.200 +         rec_thm = rec_thm}
  39.201 +      end;
  39.202 +
  39.203 +    fun mk_ctr_specs index (ctr_sugars : ctr_sugar list) iter_thmsss =
  39.204 +      let
  39.205 +        val ctrs = #ctrs (nth ctr_sugars index);
  39.206 +        val rec_thms = co_rec_of (nth iter_thmsss index);
  39.207 +        val k = offset_of_ctr index ctr_sugars;
  39.208 +        val n = length ctrs;
  39.209 +      in
  39.210 +        map4 mk_ctr_spec ctrs (k upto k + n - 1) (nth perm_fun_arg_Tssss index) rec_thms
  39.211 +      end;
  39.212 +
  39.213 +    fun mk_spec ({T, index, ctr_sugars, co_iterss = iterss, co_iter_thmsss = iter_thmsss, ...}
  39.214 +      : fp_sugar) =
  39.215 +      {recx = mk_co_iter thy Least_FP (substAT T) perm_Cs' (co_rec_of (nth iterss index)),
  39.216 +       nested_map_idents = map (unfold_thms lthy @{thms id_def} o map_id0_of_bnf) nested_bnfs,
  39.217 +       nested_map_comps = map map_comp_of_bnf nested_bnfs,
  39.218 +       ctr_specs = mk_ctr_specs index ctr_sugars iter_thmsss};
  39.219 +  in
  39.220 +    ((is_some lfp_sugar_thms, map mk_spec fp_sugars, missing_arg_Ts, induct_thm, induct_thms),
  39.221 +     lthy')
  39.222 +  end;
  39.223 +
  39.224 +val undef_const = Const (@{const_name undefined}, dummyT);
  39.225 +
  39.226 +fun permute_args n t =
  39.227 +  list_comb (t, map Bound (0 :: (n downto 1))) |> fold (K (Term.abs (Name.uu, dummyT))) (0 upto n);
  39.228 +
  39.229 +type eqn_data = {
  39.230 +  fun_name: string,
  39.231 +  rec_type: typ,
  39.232 +  ctr: term,
  39.233 +  ctr_args: term list,
  39.234 +  left_args: term list,
  39.235 +  right_args: term list,
  39.236 +  res_type: typ,
  39.237 +  rhs_term: term,
  39.238 +  user_eqn: term
  39.239 +};
  39.240 +
  39.241 +fun dissect_eqn lthy fun_names eqn' =
  39.242 +  let
  39.243 +    val eqn = drop_All eqn' |> HOLogic.dest_Trueprop
  39.244 +      handle TERM _ =>
  39.245 +        primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
  39.246 +    val (lhs, rhs) = HOLogic.dest_eq eqn
  39.247 +        handle TERM _ =>
  39.248 +          primrec_error_eqn "malformed function equation (expected \"lhs = rhs\")" eqn';
  39.249 +    val (fun_name, args) = strip_comb lhs
  39.250 +      |>> (fn x => if is_Free x then fst (dest_Free x)
  39.251 +          else primrec_error_eqn "malformed function equation (does not start with free)" eqn);
  39.252 +    val (left_args, rest) = take_prefix is_Free args;
  39.253 +    val (nonfrees, right_args) = take_suffix is_Free rest;
  39.254 +    val num_nonfrees = length nonfrees;
  39.255 +    val _ = num_nonfrees = 1 orelse if num_nonfrees = 0 then
  39.256 +      primrec_error_eqn "constructor pattern missing in left-hand side" eqn else
  39.257 +      primrec_error_eqn "more than one non-variable argument in left-hand side" eqn;
  39.258 +    val _ = member (op =) fun_names fun_name orelse
  39.259 +      primrec_error_eqn "malformed function equation (does not start with function name)" eqn
  39.260 +
  39.261 +    val (ctr, ctr_args) = strip_comb (the_single nonfrees);
  39.262 +    val _ = try (num_binder_types o fastype_of) ctr = SOME (length ctr_args) orelse
  39.263 +      primrec_error_eqn "partially applied constructor in pattern" eqn;
  39.264 +    val _ = let val d = duplicates (op =) (left_args @ ctr_args @ right_args) in null d orelse
  39.265 +      primrec_error_eqn ("duplicate variable \"" ^ Syntax.string_of_term lthy (hd d) ^
  39.266 +        "\" in left-hand side") eqn end;
  39.267 +    val _ = forall is_Free ctr_args orelse
  39.268 +      primrec_error_eqn "non-primitive pattern in left-hand side" eqn;
  39.269 +    val _ =
  39.270 +      let val b = fold_aterms (fn x as Free (v, _) =>
  39.271 +        if (not (member (op =) (left_args @ ctr_args @ right_args) x) andalso
  39.272 +        not (member (op =) fun_names v) andalso
  39.273 +        not (Variable.is_fixed lthy v)) then cons x else I | _ => I) rhs []
  39.274 +      in
  39.275 +        null b orelse
  39.276 +        primrec_error_eqn ("extra variable(s) in right-hand side: " ^
  39.277 +          commas (map (Syntax.string_of_term lthy) b)) eqn
  39.278 +      end;
  39.279 +  in
  39.280 +    {fun_name = fun_name,
  39.281 +     rec_type = body_type (type_of ctr),
  39.282 +     ctr = ctr,
  39.283 +     ctr_args = ctr_args,
  39.284 +     left_args = left_args,
  39.285 +     right_args = right_args,
  39.286 +     res_type = map fastype_of (left_args @ right_args) ---> fastype_of rhs,
  39.287 +     rhs_term = rhs,
  39.288 +     user_eqn = eqn'}
  39.289 +  end;
  39.290 +
  39.291 +fun rewrite_map_arg get_ctr_pos rec_type res_type =
  39.292 +  let
  39.293 +    val pT = HOLogic.mk_prodT (rec_type, res_type);
  39.294 +
  39.295 +    val maybe_suc = Option.map (fn x => x + 1);
  39.296 +    fun subst d (t as Bound d') = t |> d = SOME d' ? curry (op $) (fst_const pT)
  39.297 +      | subst d (Abs (v, T, b)) = Abs (v, if d = SOME ~1 then pT else T, subst (maybe_suc d) b)
  39.298 +      | subst d t =
  39.299 +        let
  39.300 +          val (u, vs) = strip_comb t;
  39.301 +          val ctr_pos = try (get_ctr_pos o fst o dest_Free) u |> the_default ~1;
  39.302 +        in
  39.303 +          if ctr_pos >= 0 then
  39.304 +            if d = SOME ~1 andalso length vs = ctr_pos then
  39.305 +              list_comb (permute_args ctr_pos (snd_const pT), vs)
  39.306 +            else if length vs > ctr_pos andalso is_some d
  39.307 +                andalso d = try (fn Bound n => n) (nth vs ctr_pos) then
  39.308 +              list_comb (snd_const pT $ nth vs ctr_pos, map (subst d) (nth_drop ctr_pos vs))
  39.309 +            else
  39.310 +              primrec_error_eqn ("recursive call not directly applied to constructor argument") t
  39.311 +          else
  39.312 +            list_comb (u, map (subst (d |> d = SOME ~1 ? K NONE)) vs)
  39.313 +        end
  39.314 +  in
  39.315 +    subst (SOME ~1)
  39.316 +  end;
  39.317 +
  39.318 +fun subst_rec_calls lthy get_ctr_pos has_call ctr_args mutual_calls nested_calls =
  39.319 +  let
  39.320 +    fun try_nested_rec bound_Ts y t =
  39.321 +      AList.lookup (op =) nested_calls y
  39.322 +      |> Option.map (fn y' =>
  39.323 +        massage_nested_rec_call lthy has_call (rewrite_map_arg get_ctr_pos) bound_Ts y y' t);
  39.324 +
  39.325 +    fun subst bound_Ts (t as g' $ y) =
  39.326 +        let
  39.327 +          fun subst_rec () = subst bound_Ts g' $ subst bound_Ts y;
  39.328 +          val y_head = head_of y;
  39.329 +        in
  39.330 +          if not (member (op =) ctr_args y_head) then
  39.331 +            subst_rec ()
  39.332 +          else
  39.333 +            (case try_nested_rec bound_Ts y_head t of
  39.334 +              SOME t' => t'
  39.335 +            | NONE =>
  39.336 +              let val (g, g_args) = strip_comb g' in
  39.337 +                (case try (get_ctr_pos o fst o dest_Free) g of
  39.338 +                  SOME ctr_pos =>
  39.339 +                  (length g_args >= ctr_pos orelse
  39.340 +                   primrec_error_eqn "too few arguments in recursive call" t;
  39.341 +                   (case AList.lookup (op =) mutual_calls y of
  39.342 +                     SOME y' => list_comb (y', g_args)
  39.343 +                   | NONE => subst_rec ()))
  39.344 +                | NONE => subst_rec ())
  39.345 +              end)
  39.346 +        end
  39.347 +      | subst bound_Ts (Abs (v, T, b)) = Abs (v, T, subst (T :: bound_Ts) b)
  39.348 +      | subst _ t = t
  39.349 +
  39.350 +    fun subst' t =
  39.351 +      if has_call t then
  39.352 +        (* FIXME detect this case earlier? *)
  39.353 +        primrec_error_eqn "recursive call not directly applied to constructor argument" t
  39.354 +      else
  39.355 +        try_nested_rec [] (head_of t) t |> the_default t
  39.356 +  in
  39.357 +    subst' o subst []
  39.358 +  end;
  39.359 +
  39.360 +fun build_rec_arg lthy (funs_data : eqn_data list list) has_call (ctr_spec : rec_ctr_spec)
  39.361 +    (maybe_eqn_data : eqn_data option) =
  39.362 +  (case maybe_eqn_data of
  39.363 +    NONE => undef_const
  39.364 +  | SOME {ctr_args, left_args, right_args, rhs_term = t, ...} =>
  39.365 +    let
  39.366 +      val calls = #calls ctr_spec;
  39.367 +      val n_args = fold (Integer.add o (fn Mutual_Rec _ => 2 | _ => 1)) calls 0;
  39.368 +
  39.369 +      val no_calls' = tag_list 0 calls
  39.370 +        |> map_filter (try (apsnd (fn No_Rec p => p | Mutual_Rec (p, _) => p)));
  39.371 +      val mutual_calls' = tag_list 0 calls
  39.372 +        |> map_filter (try (apsnd (fn Mutual_Rec (_, p) => p)));
  39.373 +      val nested_calls' = tag_list 0 calls
  39.374 +        |> map_filter (try (apsnd (fn Nested_Rec p => p)));
  39.375 +
  39.376 +      val args = replicate n_args ("", dummyT)
  39.377 +        |> Term.rename_wrt_term t
  39.378 +        |> map Free
  39.379 +        |> fold (fn (ctr_arg_idx, (arg_idx, _)) =>
  39.380 +            nth_map arg_idx (K (nth ctr_args ctr_arg_idx)))
  39.381 +          no_calls'
  39.382 +        |> fold (fn (ctr_arg_idx, (arg_idx, T)) =>
  39.383 +            nth_map arg_idx (K (retype_free T (nth ctr_args ctr_arg_idx))))
  39.384 +          mutual_calls'
  39.385 +        |> fold (fn (ctr_arg_idx, (arg_idx, T)) =>
  39.386 +            nth_map arg_idx (K (retype_free T (nth ctr_args ctr_arg_idx))))
  39.387 +          nested_calls';
  39.388 +
  39.389 +      val fun_name_ctr_pos_list =
  39.390 +        map (fn (x :: _) => (#fun_name x, length (#left_args x))) funs_data;
  39.391 +      val get_ctr_pos = try (the o AList.lookup (op =) fun_name_ctr_pos_list) #> the_default ~1;
  39.392 +      val mutual_calls = map (apfst (nth ctr_args) o apsnd (nth args o fst)) mutual_calls';
  39.393 +      val nested_calls = map (apfst (nth ctr_args) o apsnd (nth args o fst)) nested_calls';
  39.394 +    in
  39.395 +      t
  39.396 +      |> subst_rec_calls lthy get_ctr_pos has_call ctr_args mutual_calls nested_calls
  39.397 +      |> fold_rev lambda (args @ left_args @ right_args)
  39.398 +    end);
  39.399 +
  39.400 +fun build_defs lthy bs mxs (funs_data : eqn_data list list) (rec_specs : rec_spec list) has_call =
  39.401 +  let
  39.402 +    val n_funs = length funs_data;
  39.403 +
  39.404 +    val ctr_spec_eqn_data_list' =
  39.405 +      (take n_funs rec_specs |> map #ctr_specs) ~~ funs_data
  39.406 +      |> maps (uncurry (finds (fn (x, y) => #ctr x = #ctr y))
  39.407 +          ##> (fn x => null x orelse
  39.408 +            primrec_error_eqns "excess equations in definition" (map #rhs_term x)) #> fst);
  39.409 +    val _ = ctr_spec_eqn_data_list' |> map (fn (_, x) => length x <= 1 orelse
  39.410 +      primrec_error_eqns ("multiple equations for constructor") (map #user_eqn x));
  39.411 +
  39.412 +    val ctr_spec_eqn_data_list =
  39.413 +      ctr_spec_eqn_data_list' @ (drop n_funs rec_specs |> maps #ctr_specs |> map (rpair []));
  39.414 +
  39.415 +    val recs = take n_funs rec_specs |> map #recx;
  39.416 +    val rec_args = ctr_spec_eqn_data_list
  39.417 +      |> sort ((op <) o pairself (#offset o fst) |> make_ord)
  39.418 +      |> map (uncurry (build_rec_arg lthy funs_data has_call) o apsnd (try the_single));
  39.419 +    val ctr_poss = map (fn x =>
  39.420 +      if length (distinct ((op =) o pairself (length o #left_args)) x) <> 1 then
  39.421 +        primrec_error ("inconstant constructor pattern position for function " ^
  39.422 +          quote (#fun_name (hd x)))
  39.423 +      else
  39.424 +        hd x |> #left_args |> length) funs_data;
  39.425 +  in
  39.426 +    (recs, ctr_poss)
  39.427 +    |-> map2 (fn recx => fn ctr_pos => list_comb (recx, rec_args) |> permute_args ctr_pos)
  39.428 +    |> Syntax.check_terms lthy
  39.429 +    |> map3 (fn b => fn mx => fn t => ((b, mx), ((Binding.conceal (Thm.def_binding b), []), t)))
  39.430 +      bs mxs
  39.431 +  end;
  39.432 +
  39.433 +fun find_rec_calls has_call ({ctr, ctr_args, rhs_term, ...} : eqn_data) =
  39.434 +  let
  39.435 +    fun find bound_Ts (Abs (_, T, b)) ctr_arg = find (T :: bound_Ts) b ctr_arg
  39.436 +      | find bound_Ts (t as _ $ _) ctr_arg =
  39.437 +        let
  39.438 +          val typof = curry fastype_of1 bound_Ts;
  39.439 +          val (f', args') = strip_comb t;
  39.440 +          val n = find_index (equal ctr_arg o head_of) args';
  39.441 +        in
  39.442 +          if n < 0 then
  39.443 +            find bound_Ts f' ctr_arg @ maps (fn x => find bound_Ts x ctr_arg) args'
  39.444 +          else
  39.445 +            let
  39.446 +              val (f, args as arg :: _) = chop n args' |>> curry list_comb f'
  39.447 +              val (arg_head, arg_args) = Term.strip_comb arg;
  39.448 +            in
  39.449 +              if has_call f then
  39.450 +                mk_partial_compN (length arg_args) (typof arg_head) f ::
  39.451 +                maps (fn x => find bound_Ts x ctr_arg) args
  39.452 +              else
  39.453 +                find bound_Ts f ctr_arg @ maps (fn x => find bound_Ts x ctr_arg) args
  39.454 +            end
  39.455 +        end
  39.456 +      | find _ _ _ = [];
  39.457 +  in
  39.458 +    map (find [] rhs_term) ctr_args
  39.459 +    |> (fn [] => NONE | callss => SOME (ctr, callss))
  39.460 +  end;
  39.461 +
  39.462 +fun mk_primrec_tac ctxt num_extra_args map_idents map_comps fun_defs recx =
  39.463 +  unfold_thms_tac ctxt fun_defs THEN
  39.464 +  HEADGOAL (rtac (funpow num_extra_args (fn thm => thm RS fun_cong) recx RS trans)) THEN
  39.465 +  unfold_thms_tac ctxt (@{thms id_def split o_def fst_conv snd_conv} @ map_comps @ map_idents) THEN
  39.466 +  HEADGOAL (rtac refl);
  39.467 +
  39.468 +fun prepare_primrec fixes specs lthy =
  39.469 +  let
  39.470 +    val thy = Proof_Context.theory_of lthy;
  39.471 +
  39.472 +    val (bs, mxs) = map_split (apfst fst) fixes;
  39.473 +    val fun_names = map Binding.name_of bs;
  39.474 +    val eqns_data = map (dissect_eqn lthy fun_names) specs;
  39.475 +    val funs_data = eqns_data
  39.476 +      |> partition_eq ((op =) o pairself #fun_name)
  39.477 +      |> finds (fn (x, y) => x = #fun_name (hd y)) fun_names |> fst
  39.478 +      |> map (fn (x, y) => the_single y handle List.Empty =>
  39.479 +          primrec_error ("missing equations for function " ^ quote x));
  39.480 +
  39.481 +    val has_call = exists_subterm (map (fst #>> Binding.name_of #> Free) fixes |> member (op =));
  39.482 +    val arg_Ts = map (#rec_type o hd) funs_data;
  39.483 +    val res_Ts = map (#res_type o hd) funs_data;
  39.484 +    val callssss = funs_data
  39.485 +      |> map (partition_eq ((op =) o pairself #ctr))
  39.486 +      |> map (maps (map_filter (find_rec_calls has_call)));
  39.487 +
  39.488 +    val _ = (case filter_out (fn (_, T) => Sign.of_sort thy (T, HOLogic.typeS)) (bs ~~ res_Ts) of
  39.489 +        [] => ()
  39.490 +      | (b, _) :: _ => primrec_error ("type of " ^ Binding.print b ^ " contains top sort"));
  39.491 +
  39.492 +    val ((n2m, rec_specs, _, induct_thm, induct_thms), lthy') =
  39.493 +      rec_specs_of bs arg_Ts res_Ts (get_indices fixes) callssss lthy;
  39.494 +
  39.495 +    val actual_nn = length funs_data;
  39.496 +
  39.497 +    val _ = let val ctrs = (maps (map #ctr o #ctr_specs) rec_specs) in
  39.498 +      map (fn {ctr, user_eqn, ...} => member (op =) ctrs ctr orelse
  39.499 +        primrec_error_eqn ("argument " ^ quote (Syntax.string_of_term lthy' ctr) ^
  39.500 +          " is not a constructor in left-hand side") user_eqn) eqns_data end;
  39.501 +
  39.502 +    val defs = build_defs lthy' bs mxs funs_data rec_specs has_call;
  39.503 +
  39.504 +    fun prove lthy def_thms' ({ctr_specs, nested_map_idents, nested_map_comps, ...} : rec_spec)
  39.505 +        (fun_data : eqn_data list) =
  39.506 +      let
  39.507 +        val def_thms = map (snd o snd) def_thms';
  39.508 +        val simp_thmss = finds (fn (x, y) => #ctr x = #ctr y) fun_data ctr_specs
  39.509 +          |> fst
  39.510 +          |> map_filter (try (fn (x, [y]) =>
  39.511 +            (#user_eqn x, length (#left_args x) + length (#right_args x), #rec_thm y)))
  39.512 +          |> map (fn (user_eqn, num_extra_args, rec_thm) =>
  39.513 +            mk_primrec_tac lthy num_extra_args nested_map_idents nested_map_comps def_thms rec_thm
  39.514 +            |> K |> Goal.prove lthy [] [] user_eqn
  39.515 +            |> Thm.close_derivation);
  39.516 +        val poss = find_indices (fn (x, y) => #ctr x = #ctr y) fun_data eqns_data;
  39.517 +      in
  39.518 +        (poss, simp_thmss)
  39.519 +      end;
  39.520 +
  39.521 +    val notes =
  39.522 +      (if n2m then map2 (fn name => fn thm =>
  39.523 +        (name, inductN, [thm], [])) fun_names (take actual_nn induct_thms) else [])
  39.524 +      |> map (fn (prefix, thmN, thms, attrs) =>
  39.525 +        ((Binding.qualify true prefix (Binding.name thmN), attrs), [(thms, [])]));
  39.526 +
  39.527 +    val common_name = mk_common_name fun_names;
  39.528 +
  39.529 +    val common_notes =
  39.530 +      (if n2m then [(inductN, [induct_thm], [])] else [])
  39.531 +      |> map (fn (thmN, thms, attrs) =>
  39.532 +        ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
  39.533 +  in
  39.534 +    (((fun_names, defs),
  39.535 +      fn lthy => fn defs =>
  39.536 +        split_list (map2 (prove lthy defs) (take actual_nn rec_specs) funs_data)),
  39.537 +      lthy' |> Local_Theory.notes (notes @ common_notes) |> snd)
  39.538 +  end;
  39.539 +
  39.540 +(* primrec definition *)
  39.541 +
  39.542 +fun add_primrec_simple fixes ts lthy =
  39.543 +  let
  39.544 +    val (((names, defs), prove), lthy) = prepare_primrec fixes ts lthy
  39.545 +      handle ERROR str => primrec_error str;
  39.546 +  in
  39.547 +    lthy
  39.548 +    |> fold_map Local_Theory.define defs
  39.549 +    |-> (fn defs => `(fn lthy => (names, (map fst defs, prove lthy defs))))
  39.550 +  end
  39.551 +  handle Primrec_Error (str, eqns) =>
  39.552 +    if null eqns
  39.553 +    then error ("primrec_new error:\n  " ^ str)
  39.554 +    else error ("primrec_new error:\n  " ^ str ^ "\nin\n  " ^
  39.555 +      space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
  39.556 +
  39.557 +local
  39.558 +
  39.559 +fun gen_primrec prep_spec (raw_fixes : (binding * 'a option * mixfix) list) raw_spec lthy =
  39.560 +  let
  39.561 +    val d = duplicates (op =) (map (Binding.name_of o #1) raw_fixes)
  39.562 +    val _ = null d orelse primrec_error ("duplicate function name(s): " ^ commas d);
  39.563 +
  39.564 +    val (fixes, specs) = fst (prep_spec raw_fixes raw_spec lthy);
  39.565 +
  39.566 +    val mk_notes =
  39.567 +      flat ooo map3 (fn poss => fn prefix => fn thms =>
  39.568 +        let
  39.569 +          val (bs, attrss) = map_split (fst o nth specs) poss;
  39.570 +          val notes =
  39.571 +            map3 (fn b => fn attrs => fn thm =>
  39.572 +              ((Binding.qualify false prefix b, code_nitpicksimp_simp_attrs @ attrs), [([thm], [])]))
  39.573 +            bs attrss thms;
  39.574 +        in
  39.575 +          ((Binding.qualify true prefix (Binding.name simpsN), []), [(thms, [])]) :: notes
  39.576 +        end);
  39.577 +  in
  39.578 +    lthy
  39.579 +    |> add_primrec_simple fixes (map snd specs)
  39.580 +    |-> (fn (names, (ts, (posss, simpss))) =>
  39.581 +      Spec_Rules.add Spec_Rules.Equational (ts, flat simpss)
  39.582 +      #> Local_Theory.notes (mk_notes posss names simpss)
  39.583 +      #>> pair ts o map snd)
  39.584 +  end;
  39.585 +
  39.586 +in
  39.587 +
  39.588 +val add_primrec = gen_primrec Specification.check_spec;
  39.589 +val add_primrec_cmd = gen_primrec Specification.read_spec;
  39.590 +
  39.591 +end;
  39.592 +
  39.593 +fun add_primrec_global fixes specs thy =
  39.594 +  let
  39.595 +    val lthy = Named_Target.theory_init thy;
  39.596 +    val ((ts, simps), lthy') = add_primrec fixes specs lthy;
  39.597 +    val simps' = burrow (Proof_Context.export lthy' lthy) simps;
  39.598 +  in ((ts, simps'), Local_Theory.exit_global lthy') end;
  39.599 +
  39.600 +fun add_primrec_overloaded ops fixes specs thy =
  39.601 +  let
  39.602 +    val lthy = Overloading.overloading ops thy;
  39.603 +    val ((ts, simps), lthy') = add_primrec fixes specs lthy;
  39.604 +    val simps' = burrow (Proof_Context.export lthy' lthy) simps;
  39.605 +  in ((ts, simps'), Local_Theory.exit_global lthy') end;
  39.606 +
  39.607 +end;
    40.1 --- a/src/HOL/BNF/Tools/ctr_sugar.ML	Mon Nov 11 17:34:44 2013 +0100
    40.2 +++ b/src/HOL/BNF/Tools/ctr_sugar.ML	Mon Nov 11 17:44:21 2013 +0100
    40.3 @@ -33,6 +33,7 @@
    40.4       case_conv_ifs: thm list};
    40.5  
    40.6    val morph_ctr_sugar: morphism -> ctr_sugar -> ctr_sugar
    40.7 +  val transfer_ctr_sugar: Proof.context -> ctr_sugar -> ctr_sugar
    40.8    val ctr_sugar_of: Proof.context -> string -> ctr_sugar option
    40.9    val ctr_sugars_of: Proof.context -> ctr_sugar list
   40.10  
   40.11 @@ -174,10 +175,11 @@
   40.12  val dest_attrs = @{attributes [dest]};
   40.13  val safe_elim_attrs = @{attributes [elim!]};
   40.14  val iff_attrs = @{attributes [iff]};
   40.15 -val induct_simp_attrs = @{attributes [induct_simp]};
   40.16 -val nitpick_attrs = @{attributes [nitpick_simp]};
   40.17 +val inductsimp_attrs = @{attributes [induct_simp]};
   40.18 +val nitpicksimp_attrs = @{attributes [nitpick_simp]};
   40.19  val simp_attrs = @{attributes [simp]};
   40.20 -val code_nitpick_simp_simp_attrs = Code.add_default_eqn_attrib :: nitpick_attrs @ simp_attrs;
   40.21 +val code_nitpicksimp_attrs = Code.add_default_eqn_attrib :: nitpicksimp_attrs;
   40.22 +val code_nitpicksimp_simp_attrs = code_nitpicksimp_attrs @ simp_attrs;
   40.23  
   40.24  fun unflat_lookup eq xs ys = map (fn xs' => permute_like eq xs xs' ys);
   40.25  
   40.26 @@ -391,7 +393,8 @@
   40.27           Term.lambda w (Library.foldr1 HOLogic.mk_disj (map3 mk_case_disj xctrs xfs xss)));
   40.28  
   40.29      val ((raw_case, (_, raw_case_def)), (lthy', lthy)) = no_defs_lthy
   40.30 -      |> Local_Theory.define ((case_binding, NoSyn), ((Thm.def_binding case_binding, []), case_rhs))
   40.31 +      |> Local_Theory.define ((case_binding, NoSyn),
   40.32 +        ((Binding.conceal (Thm.def_binding case_binding), []), case_rhs))
   40.33        ||> `Local_Theory.restore;
   40.34  
   40.35      val phi = Proof_Context.export_morphism lthy lthy';
   40.36 @@ -869,8 +872,15 @@
   40.37          val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases));
   40.38          val cases_type_attr = Attrib.internal (K (Induct.cases_type fcT_name));
   40.39  
   40.40 +        val anonymous_notes =
   40.41 +          [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs),
   40.42 +           (map (fn th => th RS @{thm eq_False[THEN iffD2]}
   40.43 +              handle THM _ => th RS @{thm eq_True[THEN iffD2]}) nontriv_disc_thms,
   40.44 +            code_nitpicksimp_attrs)]
   40.45 +          |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
   40.46 +
   40.47          val notes =
   40.48 -          [(caseN, case_thms, code_nitpick_simp_simp_attrs),
   40.49 +          [(caseN, case_thms, code_nitpicksimp_simp_attrs),
   40.50             (case_congN, [case_cong_thm], []),
   40.51             (case_conv_ifN, case_conv_if_thms, []),
   40.52             (collapseN, safe_collapse_thms, simp_attrs),
   40.53 @@ -878,12 +888,12 @@
   40.54             (discIN, nontriv_discI_thms, []),
   40.55             (disc_excludeN, disc_exclude_thms, dest_attrs),
   40.56             (disc_exhaustN, disc_exhaust_thms, [exhaust_case_names_attr]),
   40.57 -           (distinctN, distinct_thms, simp_attrs @ induct_simp_attrs),
   40.58 +           (distinctN, distinct_thms, simp_attrs @ inductsimp_attrs),
   40.59             (exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]),
   40.60             (expandN, expand_thms, []),
   40.61 -           (injectN, inject_thms, iff_attrs @ induct_simp_attrs),
   40.62 +           (injectN, inject_thms, iff_attrs @ inductsimp_attrs),
   40.63             (nchotomyN, [nchotomy_thm], []),
   40.64 -           (selN, all_sel_thms, code_nitpick_simp_simp_attrs),
   40.65 +           (selN, all_sel_thms, code_nitpicksimp_simp_attrs),
   40.66             (sel_exhaustN, sel_exhaust_thms, [exhaust_case_names_attr]),
   40.67             (sel_splitN, sel_split_thms, []),
   40.68             (sel_split_asmN, sel_split_asm_thms, []),
   40.69 @@ -895,10 +905,6 @@
   40.70            |> map (fn (thmN, thms, attrs) =>
   40.71              ((qualify true (Binding.name thmN), attrs), [(thms, [])]));
   40.72  
   40.73 -        val notes' =
   40.74 -          [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs)]
   40.75 -          |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
   40.76 -
   40.77          val ctr_sugar =
   40.78            {ctrs = ctrs, casex = casex, discs = discs, selss = selss, exhaust = exhaust_thm,
   40.79             nchotomy = nchotomy_thm, injects = inject_thms, distincts = distinct_thms,
   40.80 @@ -915,7 +921,7 @@
   40.81              (Local_Theory.declaration {syntax = false, pervasive = true}
   40.82                 (fn phi => Case_Translation.register
   40.83                    (Morphism.term phi casex) (map (Morphism.term phi) ctrs)))
   40.84 -         |> Local_Theory.notes (notes' @ notes) |> snd
   40.85 +         |> Local_Theory.notes (anonymous_notes @ notes) |> snd
   40.86           |> register_ctr_sugar fcT_name ctr_sugar)
   40.87        end;
   40.88    in
    41.1 --- a/src/HOL/Big_Operators.thy	Mon Nov 11 17:34:44 2013 +0100
    41.2 +++ b/src/HOL/Big_Operators.thy	Mon Nov 11 17:44:21 2013 +0100
    41.3 @@ -696,11 +696,7 @@
    41.4  lemma setsum_subtractf:
    41.5    "setsum (%x. ((f x)::'a::ab_group_add) - g x) A =
    41.6      setsum f A - setsum g A"
    41.7 -proof (cases "finite A")
    41.8 -  case True thus ?thesis by (simp add: diff_minus setsum_addf setsum_negf)
    41.9 -next
   41.10 -  case False thus ?thesis by simp
   41.11 -qed
   41.12 +  using setsum_addf [of f "- g" A] by (simp add: setsum_negf)
   41.13  
   41.14  lemma setsum_nonneg:
   41.15    assumes nn: "\<forall>x\<in>A. (0::'a::{ordered_ab_semigroup_add,comm_monoid_add}) \<le> f x"
   41.16 @@ -1999,35 +1995,35 @@
   41.17    assumes fin_nonempty: "finite A" "A \<noteq> {}"
   41.18  begin
   41.19  
   41.20 -lemma Min_ge_iff [simp, no_atp]:
   41.21 +lemma Min_ge_iff [simp]:
   41.22    "x \<le> Min A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"
   41.23    using fin_nonempty by (fact Min.bounded_iff)
   41.24  
   41.25 -lemma Max_le_iff [simp, no_atp]:
   41.26 +lemma Max_le_iff [simp]:
   41.27    "Max A \<le> x \<longleftrightarrow> (\<forall>a\<in>A. a \<le> x)"
   41.28    using fin_nonempty by (fact Max.bounded_iff)
   41.29  
   41.30 -lemma Min_gr_iff [simp, no_atp]:
   41.31 +lemma Min_gr_iff [simp]:
   41.32    "x < Min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)"
   41.33    using fin_nonempty  by (induct rule: finite_ne_induct) simp_all
   41.34  
   41.35 -lemma Max_less_iff [simp, no_atp]:
   41.36 +lemma Max_less_iff [simp]:
   41.37    "Max A < x \<longleftrightarrow> (\<forall>a\<in>A. a < x)"
   41.38    using fin_nonempty by (induct rule: finite_ne_induct) simp_all
   41.39  
   41.40 -lemma Min_le_iff [no_atp]:
   41.41 +lemma Min_le_iff:
   41.42    "Min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)"
   41.43    using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: min_le_iff_disj)
   41.44  
   41.45 -lemma Max_ge_iff [no_atp]:
   41.46 +lemma Max_ge_iff:
   41.47    "x \<le> Max A \<longleftrightarrow> (\<exists>a\<in>A. x \<le> a)"
   41.48    using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: le_max_iff_disj)
   41.49  
   41.50 -lemma Min_less_iff [no_atp]:
   41.51 +lemma Min_less_iff:
   41.52    "Min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)"
   41.53    using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: min_less_iff_disj)
   41.54  
   41.55 -lemma Max_gr_iff [no_atp]:
   41.56 +lemma Max_gr_iff:
   41.57    "x < Max A \<longleftrightarrow> (\<exists>a\<in>A. x < a)"
   41.58    using fin_nonempty by (induct rule: finite_ne_induct) (simp_all add: less_max_iff_disj)
   41.59  
    42.1 --- a/src/HOL/Codegenerator_Test/Generate_Binary_Nat.thy	Mon Nov 11 17:34:44 2013 +0100
    42.2 +++ b/src/HOL/Codegenerator_Test/Generate_Binary_Nat.thy	Mon Nov 11 17:44:21 2013 +0100
    42.3 @@ -16,6 +16,20 @@
    42.4    by a corresponding @{text export_code} command.
    42.5  *}
    42.6  
    42.7 -export_code _ checking SML OCaml? Haskell? Scala
    42.8 +text {* Formal joining of hierarchy of implicit definitions in Scala *}
    42.9 +
   42.10 +class semiring_numeral_even_odd = semiring_numeral_div + even_odd
   42.11 +
   42.12 +instance nat :: semiring_numeral_even_odd ..
   42.13 +
   42.14 +definition semiring_numeral_even_odd :: "'a itself \<Rightarrow> 'a::semiring_numeral_even_odd"
   42.15 +where
   42.16 +  "semiring_numeral_even_odd TYPE('a) = undefined"
   42.17 +
   42.18 +definition semiring_numeral_even_odd_nat :: "nat itself \<Rightarrow> nat"
   42.19 +where
   42.20 +  "semiring_numeral_even_odd_nat = semiring_numeral_even_odd"
   42.21 +
   42.22 +export_code _ checking  SML OCaml? Haskell? Scala
   42.23  
   42.24  end
    43.1 --- a/src/HOL/Codegenerator_Test/Generate_Efficient_Datastructures.thy	Mon Nov 11 17:34:44 2013 +0100
    43.2 +++ b/src/HOL/Codegenerator_Test/Generate_Efficient_Datastructures.thy	Mon Nov 11 17:44:21 2013 +0100
    43.3 @@ -26,7 +26,7 @@
    43.4    "pred_of_set = pred_of_set" ..
    43.5  
    43.6  lemma [code, code del]:
    43.7 -  "acc = acc" ..
    43.8 +  "Wellfounded.acc = Wellfounded.acc" ..
    43.9  
   43.10  lemma [code, code del]:
   43.11    "Cardinality.card' = Cardinality.card'" ..
    44.1 --- a/src/HOL/Complete_Lattices.thy	Mon Nov 11 17:34:44 2013 +0100
    44.2 +++ b/src/HOL/Complete_Lattices.thy	Mon Nov 11 17:44:21 2013 +0100
    44.3 @@ -15,10 +15,66 @@
    44.4  
    44.5  class Inf =
    44.6    fixes Inf :: "'a set \<Rightarrow> 'a" ("\<Sqinter>_" [900] 900)
    44.7 +begin
    44.8 +
    44.9 +definition INFI :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
   44.10 +  INF_def: "INFI A f = \<Sqinter>(f ` A)"
   44.11 +
   44.12 +lemma INF_image [simp]: "INFI (f`A) g = INFI A (\<lambda>x. g (f x))"
   44.13 +  by (simp add: INF_def image_image)
   44.14 +
   44.15 +lemma INF_cong: "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> INFI A C = INFI B D"
   44.16 +  by (simp add: INF_def image_def)
   44.17 +
   44.18 +end
   44.19  
   44.20  class Sup =
   44.21    fixes Sup :: "'a set \<Rightarrow> 'a" ("\<Squnion>_" [900] 900)
   44.22 +begin
   44.23  
   44.24 +definition SUPR :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
   44.25 +  SUP_def: "SUPR A f = \<Squnion>(f ` A)"
   44.26 +
   44.27 +lemma SUP_image [simp]: "SUPR (f`A) g = SUPR A (%x. g (f x))"
   44.28 +  by (simp add: SUP_def image_image)
   44.29 +
   44.30 +lemma SUP_cong: "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> SUPR A C = SUPR B D"
   44.31 +  by (simp add: SUP_def image_def)
   44.32 +
   44.33 +end
   44.34 +
   44.35 +text {*
   44.36 +  Note: must use names @{const INFI} and @{const SUPR} here instead of
   44.37 +  @{text INF} and @{text SUP} to allow the following syntax coexist
   44.38 +  with the plain constant names.
   44.39 +*}
   44.40 +
   44.41 +syntax
   44.42 +  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3INF _./ _)" [0, 10] 10)
   44.43 +  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3INF _:_./ _)" [0, 0, 10] 10)
   44.44 +  "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3SUP _./ _)" [0, 10] 10)
   44.45 +  "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3SUP _:_./ _)" [0, 0, 10] 10)
   44.46 +
   44.47 +syntax (xsymbols)
   44.48 +  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
   44.49 +  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
   44.50 +  "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
   44.51 +  "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
   44.52 +
   44.53 +translations
   44.54 +  "INF x y. B"   == "INF x. INF y. B"
   44.55 +  "INF x. B"     == "CONST INFI CONST UNIV (%x. B)"
   44.56 +  "INF x. B"     == "INF x:CONST UNIV. B"
   44.57 +  "INF x:A. B"   == "CONST INFI A (%x. B)"
   44.58 +  "SUP x y. B"   == "SUP x. SUP y. B"
   44.59 +  "SUP x. B"     == "CONST SUPR CONST UNIV (%x. B)"
   44.60 +  "SUP x. B"     == "SUP x:CONST UNIV. B"
   44.61 +  "SUP x:A. B"   == "CONST SUPR A (%x. B)"
   44.62 +
   44.63 +print_translation {*
   44.64 +  [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INFI} @{syntax_const "_INF"},
   44.65 +    Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax SUPR} @{syntax_const "_SUP"}]
   44.66 +*} -- {* to avoid eta-contraction of body *}
   44.67  
   44.68  subsection {* Abstract complete lattices *}
   44.69  
   44.70 @@ -49,59 +105,17 @@
   44.71      (unfold_locales, (fact Inf_empty Sup_empty
   44.72          Sup_upper Sup_least Inf_lower Inf_greatest)+)
   44.73  
   44.74 -definition INFI :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
   44.75 -  INF_def: "INFI A f = \<Sqinter>(f ` A)"
   44.76 -
   44.77 -definition SUPR :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
   44.78 -  SUP_def: "SUPR A f = \<Squnion>(f ` A)"
   44.79 -
   44.80 -text {*
   44.81 -  Note: must use names @{const INFI} and @{const SUPR} here instead of
   44.82 -  @{text INF} and @{text SUP} to allow the following syntax coexist
   44.83 -  with the plain constant names.
   44.84 -*}
   44.85 -
   44.86  end
   44.87  
   44.88 -syntax
   44.89 -  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3INF _./ _)" [0, 10] 10)
   44.90 -  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3INF _:_./ _)" [0, 0, 10] 10)
   44.91 -  "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3SUP _./ _)" [0, 10] 10)
   44.92 -  "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3SUP _:_./ _)" [0, 0, 10] 10)
   44.93 -
   44.94 -syntax (xsymbols)
   44.95 -  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
   44.96 -  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
   44.97 -  "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
   44.98 -  "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
   44.99 -
  44.100 -translations
  44.101 -  "INF x y. B"   == "INF x. INF y. B"
  44.102 -  "INF x. B"     == "CONST INFI CONST UNIV (%x. B)"
  44.103 -  "INF x. B"     == "INF x:CONST UNIV. B"
  44.104 -  "INF x:A. B"   == "CONST INFI A (%x. B)"
  44.105 -  "SUP x y. B"   == "SUP x. SUP y. B"
  44.106 -  "SUP x. B"     == "CONST SUPR CONST UNIV (%x. B)"
  44.107 -  "SUP x. B"     == "SUP x:CONST UNIV. B"
  44.108 -  "SUP x:A. B"   == "CONST SUPR A (%x. B)"
  44.109 -
  44.110 -print_translation {*
  44.111 -  [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INFI} @{syntax_const "_INF"},
  44.112 -    Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax SUPR} @{syntax_const "_SUP"}]
  44.113 -*} -- {* to avoid eta-contraction of body *}
  44.114 -
  44.115  context complete_lattice
  44.116  begin
  44.117  
  44.118 -lemma INF_foundation_dual [no_atp]:
  44.119 -  "complete_lattice.SUPR Inf = INFI"
  44.120 -  by (simp add: fun_eq_iff INF_def
  44.121 -    complete_lattice.SUP_def [OF dual_complete_lattice])
  44.122 +lemma INF_foundation_dual:
  44.123 +  "Sup.SUPR Inf = INFI"
  44.124 +  by (simp add: fun_eq_iff INF_def Sup.SUP_def)
  44.125  
  44.126 -lemma SUP_foundation_dual [no_atp]:
  44.127 -  "complete_lattice.INFI Sup = SUPR"
  44.128 -  by (simp add: fun_eq_iff SUP_def
  44.129 -    complete_lattice.INF_def [OF dual_complete_lattice])
  44.130 +lemma SUP_foundation_dual:
  44.131 +  "Inf.INFI Sup = SUPR" by (simp add: fun_eq_iff SUP_def Inf.INF_def)
  44.132  
  44.133  lemma Sup_eqI:
  44.134    "(\<And>y. y \<in> A \<Longrightarrow> y \<le> x) \<Longrightarrow> (\<And>y. (\<And>z. z \<in> A \<Longrightarrow> z \<le> y) \<Longrightarrow> x \<le> y) \<Longrightarrow> \<Squnion>A = x"
  44.135 @@ -181,12 +195,6 @@
  44.136    "\<Squnion>UNIV = \<top>"
  44.137    by (auto intro!: antisym Sup_upper)
  44.138  
  44.139 -lemma INF_image [simp]: "(\<Sqinter>x\<in>f`A. g x) = (\<Sqinter>x\<in>A. g (f x))"
  44.140 -  by (simp add: INF_def image_image)
  44.141 -
  44.142 -lemma SUP_image [simp]: "(\<Squnion>x\<in>f`A. g x) = (\<Squnion>x\<in>A. g (f x))"
  44.143 -  by (simp add: SUP_def image_image)
  44.144 -
  44.145  lemma Inf_Sup: "\<Sqinter>A = \<Squnion>{b. \<forall>a \<in> A. b \<sqsubseteq> a}"
  44.146    by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least)
  44.147  
  44.148 @@ -199,14 +207,6 @@
  44.149  lemma Sup_subset_mono: "A \<subseteq> B \<Longrightarrow> \<Squnion>A \<sqsubseteq> \<Squnion>B"
  44.150    by (auto intro: Sup_least Sup_upper)
  44.151  
  44.152 -lemma INF_cong:
  44.153 -  "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> (\<Sqinter>x\<in>A. C x) = (\<Sqinter>x\<in>B. D x)"
  44.154 -  by (simp add: INF_def image_def)
  44.155 -
  44.156 -lemma SUP_cong:
  44.157 -  "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> (\<Squnion>x\<in>A. C x) = (\<Squnion>x\<in>B. D x)"
  44.158 -  by (simp add: SUP_def image_def)
  44.159 -
  44.160  lemma Inf_mono:
  44.161    assumes "\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<sqsubseteq> b"
  44.162    shows "\<Sqinter>A \<sqsubseteq> \<Sqinter>B"
  44.163 @@ -306,7 +306,7 @@
  44.164    show "?R \<le> ?L" by (rule SUP_least) (auto intro: le_supI1 le_supI2 SUP_upper)
  44.165  qed
  44.166  
  44.167 -lemma Inf_top_conv [simp, no_atp]:
  44.168 +lemma Inf_top_conv [simp]:
  44.169    "\<Sqinter>A = \<top> \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)"
  44.170    "\<top> = \<Sqinter>A \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)"
  44.171  proof -
  44.172 @@ -333,7 +333,7 @@
  44.173   "\<top> = (\<Sqinter>x\<in>A. B x) \<longleftrightarrow> (\<forall>x\<in>A. B x = \<top>)"
  44.174    by (auto simp add: INF_def)
  44.175  
  44.176 -lemma Sup_bot_conv [simp, no_atp]:
  44.177 +lemma Sup_bot_conv [simp]:
  44.178    "\<Squnion>A = \<bottom> \<longleftrightarrow> (\<forall>x\<in>A. x = \<bottom>)" (is ?P)
  44.179    "\<bottom> = \<Squnion>A \<longleftrightarrow> (\<forall>x\<in>A. x = \<bottom>)" (is ?Q)
  44.180    using dual_complete_lattice
  44.181 @@ -769,7 +769,7 @@
  44.182      by (simp add: Inf_set_def image_def)
  44.183  qed
  44.184  
  44.185 -lemma Inter_iff [simp,no_atp]: "A \<in> \<Inter>C \<longleftrightarrow> (\<forall>X\<in>C. A \<in> X)"
  44.186 +lemma Inter_iff [simp]: "A \<in> \<Inter>C \<longleftrightarrow> (\<forall>X\<in>C. A \<in> X)"
  44.187    by (unfold Inter_eq) blast
  44.188  
  44.189  lemma InterI [intro!]: "(\<And>X. X \<in> C \<Longrightarrow> A \<in> X) \<Longrightarrow> A \<in> \<Inter>C"
  44.190 @@ -814,7 +814,7 @@
  44.191  lemma Inter_Un_distrib: "\<Inter>(A \<union> B) = \<Inter>A \<inter> \<Inter>B"
  44.192    by (fact Inf_union_distrib)
  44.193  
  44.194 -lemma Inter_UNIV_conv [simp, no_atp]:
  44.195 +lemma Inter_UNIV_conv [simp]:
  44.196    "\<Inter>A = UNIV \<longleftrightarrow> (\<forall>x\<in>A. x = UNIV)"
  44.197    "UNIV = \<Inter>A \<longleftrightarrow> (\<forall>x\<in>A. x = UNIV)"
  44.198    by (fact Inf_top_conv)+
  44.199 @@ -952,7 +952,7 @@
  44.200      by (simp add: Sup_set_def image_def)
  44.201  qed
  44.202  
  44.203 -lemma Union_iff [simp, no_atp]:
  44.204 +lemma Union_iff [simp]:
  44.205    "A \<in> \<Union>C \<longleftrightarrow> (\<exists>X\<in>C. A\<in>X)"
  44.206    by (unfold Union_eq) blast
  44.207  
  44.208 @@ -987,10 +987,10 @@
  44.209  lemma Union_Int_subset: "\<Union>(A \<inter> B) \<subseteq> \<Union>A \<inter> \<Union>B"
  44.210    by (fact Sup_inter_less_eq)
  44.211  
  44.212 -lemma Union_empty_conv [no_atp]: "(\<Union>A = {}) \<longleftrightarrow> (\<forall>x\<in>A. x = {})"
  44.213 +lemma Union_empty_conv: "(\<Union>A = {}) \<longleftrightarrow> (\<forall>x\<in>A. x = {})"
  44.214    by (fact Sup_bot_conv) (* already simp *)
  44.215  
  44.216 -lemma empty_Union_conv [no_atp]: "({} = \<Union>A) \<longleftrightarrow> (\<forall>x\<in>A. x = {})"
  44.217 +lemma empty_Union_conv: "({} = \<Union>A) \<longleftrightarrow> (\<forall>x\<in>A. x = {})"
  44.218    by (fact Sup_bot_conv) (* already simp *)
  44.219  
  44.220  lemma subset_Pow_Union: "A \<subseteq> Pow (\<Union>A)"
  44.221 @@ -1044,7 +1044,7 @@
  44.222    [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax UNION} @{syntax_const "_UNION"}]
  44.223  *} -- {* to avoid eta-contraction of body *}
  44.224  
  44.225 -lemma UNION_eq [no_atp]:
  44.226 +lemma UNION_eq:
  44.227    "(\<Union>x\<in>A. B x) = {y. \<exists>x\<in>A. y \<in> B x}"
  44.228    by (auto simp add: SUP_def)
  44.229  
  44.230 @@ -1088,13 +1088,13 @@
  44.231  lemma UN_least: "(\<And>x. x \<in> A \<Longrightarrow> B x \<subseteq> C) \<Longrightarrow> (\<Union>x\<in>A. B x) \<subseteq> C"
  44.232    by (fact SUP_least)
  44.233  
  44.234 -lemma Collect_bex_eq [no_atp]: "{x. \<exists>y\<in>A. P x y} = (\<Union>y\<in>A. {x. P x y})"
  44.235 +lemma Collect_bex_eq: "{x. \<exists>y\<in>A. P x y} = (\<Union>y\<in>A. {x. P x y})"
  44.236    by blast
  44.237  
  44.238  lemma UN_insert_distrib: "u \<in> A \<Longrightarrow> (\<Union>x\<in>A. insert a (B x)) = insert a (\<Union>x\<in>A. B x)"
  44.239    by blast
  44.240  
  44.241 -lemma UN_empty [no_atp]: "(\<Union>x\<in>{}. B x) = {}"
  44.242 +lemma UN_empty: "(\<Union>x\<in>{}. B x) = {}"
  44.243    by (fact SUP_empty)
  44.244  
  44.245  lemma UN_empty2: "(\<Union>x\<in>A. {}) = {}"
  44.246 @@ -1126,7 +1126,7 @@
  44.247    "(\<Union>x\<in>A. B x) = {} \<longleftrightarrow> (\<forall>x\<in>A. B x = {})"
  44.248    by (fact SUP_bot_conv)+ (* already simp *)
  44.249  
  44.250 -lemma Collect_ex_eq [no_atp]: "{x. \<exists>y. P x y} = (\<Union>y. {x. P x y})"
  44.251 +lemma Collect_ex_eq: "{x. \<exists>y. P x y} = (\<Union>y. {x. P x y})"
  44.252    by blast
  44.253  
  44.254  lemma ball_UN: "(\<forall>z \<in> UNION A B. P z) \<longleftrightarrow> (\<forall>x\<in>A. \<forall>z \<in> B x. P z)"
  44.255 @@ -1248,7 +1248,7 @@
  44.256    "\<And>A B f. (\<Inter>x\<in>f`A. B x) = (\<Inter>a\<in>A. B (f a))"
  44.257    by auto
  44.258  
  44.259 -lemma UN_ball_bex_simps [simp, no_atp]:
  44.260 +lemma UN_ball_bex_simps [simp]:
  44.261    "\<And>A P. (\<forall>x\<in>\<Union>A. P x) \<longleftrightarrow> (\<forall>y\<in>A. \<forall>x\<in>y. P x)"
  44.262    "\<And>A B P. (\<forall>x\<in>UNION A B. P x) = (\<forall>a\<in>A. \<forall>x\<in> B a. P x)"
  44.263    "\<And>A P. (\<exists>x\<in>\<Union>A. P x) \<longleftrightarrow> (\<exists>y\<in>A. \<exists>x\<in>y. P x)"
    45.1 --- a/src/HOL/Complex.thy	Mon Nov 11 17:34:44 2013 +0100
    45.2 +++ b/src/HOL/Complex.thy	Mon Nov 11 17:44:21 2013 +0100
    45.3 @@ -587,7 +587,7 @@
    45.4    by (simp add: cis_def)
    45.5  
    45.6  lemma cis_divide: "cis a / cis b = cis (a - b)"
    45.7 -  by (simp add: complex_divide_def cis_mult diff_minus)
    45.8 +  by (simp add: complex_divide_def cis_mult)
    45.9  
   45.10  lemma cos_n_Re_cis_pow_n: "cos (real n * a) = Re(cis a ^ n)"
   45.11    by (auto simp add: DeMoivre)
    46.1 --- a/src/HOL/Conditionally_Complete_Lattices.thy	Mon Nov 11 17:34:44 2013 +0100
    46.2 +++ b/src/HOL/Conditionally_Complete_Lattices.thy	Mon Nov 11 17:44:21 2013 +0100
    46.3 @@ -1,20 +1,161 @@
    46.4  (*  Title:      HOL/Conditionally_Complete_Lattices.thy
    46.5      Author:     Amine Chaieb and L C Paulson, University of Cambridge
    46.6      Author:     Johannes Hölzl, TU München
    46.7 +    Author:     Luke S. Serafin, Carnegie Mellon University
    46.8  *)
    46.9  
   46.10  header {* Conditionally-complete Lattices *}
   46.11  
   46.12  theory Conditionally_Complete_Lattices
   46.13 -imports Main Lubs
   46.14 +imports Main
   46.15  begin
   46.16  
   46.17 -lemma Sup_fin_eq_Max: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup_fin X = Max X"
   46.18 +lemma (in linorder) Sup_fin_eq_Max: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup_fin X = Max X"
   46.19    by (induct X rule: finite_ne_induct) (simp_all add: sup_max)
   46.20  
   46.21 -lemma Inf_fin_eq_Min: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf_fin X = Min X"
   46.22 +lemma (in linorder) Inf_fin_eq_Min: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf_fin X = Min X"
   46.23    by (induct X rule: finite_ne_induct) (simp_all add: inf_min)
   46.24  
   46.25 +context preorder
   46.26 +begin
   46.27 +
   46.28 +definition "bdd_above A \<longleftrightarrow> (\<exists>M. \<forall>x \<in> A. x \<le> M)"
   46.29 +definition "bdd_below A \<longleftrightarrow> (\<exists>m. \<forall>x \<in> A. m \<le> x)"
   46.30 +
   46.31 +lemma bdd_aboveI[intro]: "(\<And>x. x \<in> A \<Longrightarrow> x \<le> M) \<Longrightarrow> bdd_above A"
   46.32 +  by (auto simp: bdd_above_def)
   46.33 +
   46.34 +lemma bdd_belowI[intro]: "(\<And>x. x \<in> A \<Longrightarrow> m \<le> x) \<Longrightarrow> bdd_below A"
   46.35 +  by (auto simp: bdd_below_def)
   46.36 +
   46.37 +lemma bdd_aboveI2: "(\<And>x. x \<in> A \<Longrightarrow> f x \<le> M) \<Longrightarrow> bdd_above (f`A)"
   46.38 +  by force
   46.39 +
   46.40 +lemma bdd_belowI2: "(\<And>x. x \<in> A \<Longrightarrow> m \<le> f x) \<Longrightarrow> bdd_below (f`A)"
   46.41 +  by force
   46.42 +
   46.43 +lemma bdd_above_empty [simp, intro]: "bdd_above {}"
   46.44 +  unfolding bdd_above_def by auto
   46.45 +
   46.46 +lemma bdd_below_empty [simp, intro]: "bdd_below {}"
   46.47 +  unfolding bdd_below_def by auto
   46.48 +
   46.49 +lemma bdd_above_mono: "bdd_above B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> bdd_above A"
   46.50 +  by (metis (full_types) bdd_above_def order_class.le_neq_trans psubsetD)
   46.51 +
   46.52 +lemma bdd_below_mono: "bdd_below B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> bdd_below A"
   46.53 +  by (metis bdd_below_def order_class.le_neq_trans psubsetD)
   46.54 +
   46.55 +lemma bdd_above_Int1 [simp]: "bdd_above A \<Longrightarrow> bdd_above (A \<inter> B)"
   46.56 +  using bdd_above_mono by auto
   46.57 +
   46.58 +lemma bdd_above_Int2 [simp]: "bdd_above B \<Longrightarrow> bdd_above (A \<inter> B)"
   46.59 +  using bdd_above_mono by auto
   46.60 +
   46.61 +lemma bdd_below_Int1 [simp]: "bdd_below A \<Longrightarrow> bdd_below (A \<inter> B)"
   46.62 +  using bdd_below_mono by auto
   46.63 +
   46.64 +lemma bdd_below_Int2 [simp]: "bdd_below B \<Longrightarrow> bdd_below (A \<inter> B)"
   46.65 +  using bdd_below_mono by auto
   46.66 +
   46.67 +lemma bdd_above_Ioo [simp, intro]: "bdd_above {a <..< b}"
   46.68 +  by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le)
   46.69 +
   46.70 +lemma bdd_above_Ico [simp, intro]: "bdd_above {a ..< b}"
   46.71 +  by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le)
   46.72 +
   46.73 +lemma bdd_above_Iio [simp, intro]: "bdd_above {..< b}"
   46.74 +  by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
   46.75 +
   46.76 +lemma bdd_above_Ioc [simp, intro]: "bdd_above {a <.. b}"
   46.77 +  by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
   46.78 +
   46.79 +lemma bdd_above_Icc [simp, intro]: "bdd_above {a .. b}"
   46.80 +  by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
   46.81 +
   46.82 +lemma bdd_above_Iic [simp, intro]: "bdd_above {.. b}"
   46.83 +  by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
   46.84 +
   46.85 +lemma bdd_below_Ioo [simp, intro]: "bdd_below {a <..< b}"
   46.86 +  by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le)
   46.87 +
   46.88 +lemma bdd_below_Ioc [simp, intro]: "bdd_below {a <.. b}"
   46.89 +  by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le)
   46.90 +
   46.91 +lemma bdd_below_Ioi [simp, intro]: "bdd_below {a <..}"
   46.92 +  by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
   46.93 +
   46.94 +lemma bdd_below_Ico [simp, intro]: "bdd_below {a ..< b}"
   46.95 +  by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
   46.96 +
   46.97 +lemma bdd_below_Icc [simp, intro]: "bdd_below {a .. b}"
   46.98 +  by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
   46.99 +
  46.100 +lemma bdd_below_Ici [simp, intro]: "bdd_below {a ..}"
  46.101 +  by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
  46.102 +
  46.103 +end
  46.104 +
  46.105 +lemma (in order_top) bdd_above_top[simp, intro!]: "bdd_above A"
  46.106 +  by (rule bdd_aboveI[of _ top]) simp
  46.107 +
  46.108 +lemma (in order_bot) bdd_above_bot[simp, intro!]: "bdd_below A"
  46.109 +  by (rule bdd_belowI[of _ bot]) simp
  46.110 +
  46.111 +lemma bdd_above_uminus[simp]:
  46.112 +  fixes X :: "'a::ordered_ab_group_add set"
  46.113 +  shows "bdd_above (uminus ` X) \<longleftrightarrow> bdd_below X"
  46.114 +  by (auto simp: bdd_above_def bdd_below_def intro: le_imp_neg_le) (metis le_imp_neg_le minus_minus)
  46.115 +
  46.116 +lemma bdd_below_uminus[simp]:
  46.117 +  fixes X :: "'a::ordered_ab_group_add set"
  46.118 +  shows"bdd_below (uminus ` X) \<longleftrightarrow> bdd_above X"
  46.119 +  by (auto simp: bdd_above_def bdd_below_def intro: le_imp_neg_le) (metis le_imp_neg_le minus_minus)
  46.120 +
  46.121 +context lattice
  46.122 +begin
  46.123 +
  46.124 +lemma bdd_above_insert [simp]: "bdd_above (insert a A) = bdd_above A"
  46.125 +  by (auto simp: bdd_above_def intro: le_supI2 sup_ge1)
  46.126 +
  46.127 +lemma bdd_below_insert [simp]: "bdd_below (insert a A) = bdd_below A"
  46.128 +  by (auto simp: bdd_below_def intro: le_infI2 inf_le1)
  46.129 +
  46.130 +lemma bdd_finite [simp]:
  46.131 +  assumes "finite A" shows bdd_above_finite: "bdd_above A" and bdd_below_finite: "bdd_below A"
  46.132 +  using assms by (induct rule: finite_induct, auto)
  46.133 +
  46.134 +lemma bdd_above_Un [simp]: "bdd_above (A \<union> B) = (bdd_above A \<and> bdd_above B)"
  46.135 +proof
  46.136 +  assume "bdd_above (A \<union> B)"
  46.137 +  thus "bdd_above A \<and> bdd_above B" unfolding bdd_above_def by auto
  46.138 +next
  46.139 +  assume "bdd_above A \<and> bdd_above B"
  46.140 +  then obtain a b where "\<forall>x\<in>A. x \<le> a" "\<forall>x\<in>B. x \<le> b" unfolding bdd_above_def by auto
  46.141 +  hence "\<forall>x \<in> A \<union> B. x \<le> sup a b" by (auto intro: Un_iff le_supI1 le_supI2)
  46.142 +  thus "bdd_above (A \<union> B)" unfolding bdd_above_def ..
  46.143 +qed
  46.144 +
  46.145 +lemma bdd_below_Un [simp]: "bdd_below (A \<union> B) = (bdd_below A \<and> bdd_below B)"
  46.146 +proof
  46.147 +  assume "bdd_below (A \<union> B)"
  46.148 +  thus "bdd_below A \<and> bdd_below B" unfolding bdd_below_def by auto
  46.149 +next
  46.150 +  assume "bdd_below A \<and> bdd_below B"
  46.151 +  then obtain a b where "\<forall>x\<in>A. a \<le> x" "\<forall>x\<in>B. b \<le> x" unfolding bdd_below_def by auto
  46.152 +  hence "\<forall>x \<in> A \<union> B. inf a b \<le> x" by (auto intro: Un_iff le_infI1 le_infI2)
  46.153 +  thus "bdd_below (A \<union> B)" unfolding bdd_below_def ..
  46.154 +qed
  46.155 +
  46.156 +lemma bdd_above_sup[simp]: "bdd_above ((\<lambda>x. sup (f x) (g x)) ` A) \<longleftrightarrow> bdd_above (f`A) \<and> bdd_above (g`A)"
  46.157 +  by (auto simp: bdd_above_def intro: le_supI1 le_supI2)
  46.158 +
  46.159 +lemma bdd_below_inf[simp]: "bdd_below ((\<lambda>x. inf (f x) (g x)) ` A) \<longleftrightarrow> bdd_below (f`A) \<and> bdd_below (g`A)"
  46.160 +  by (auto simp: bdd_below_def intro: le_infI1 le_infI2)
  46.161 +
  46.162 +end
  46.163 +
  46.164 +
  46.165  text {*
  46.166  
  46.167  To avoid name classes with the @{class complete_lattice}-class we prefix @{const Sup} and
  46.168 @@ -23,46 +164,42 @@
  46.169  *}
  46.170  
  46.171  class conditionally_complete_lattice = lattice + Sup + Inf +
  46.172 -  assumes cInf_lower: "x \<in> X \<Longrightarrow> (\<And>a. a \<in> X \<Longrightarrow> z \<le> a) \<Longrightarrow> Inf X \<le> x"
  46.173 +  assumes cInf_lower: "x \<in> X \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> x"
  46.174      and cInf_greatest: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> z \<le> Inf X"
  46.175 -  assumes cSup_upper: "x \<in> X \<Longrightarrow> (\<And>a. a \<in> X \<Longrightarrow> a \<le> z) \<Longrightarrow> x \<le> Sup X"
  46.176 +  assumes cSup_upper: "x \<in> X \<Longrightarrow> bdd_above X \<Longrightarrow> x \<le> Sup X"
  46.177      and cSup_least: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X \<le> z"
  46.178  begin
  46.179  
  46.180 -lemma cSup_eq_maximum: (*REAL_SUP_MAX in HOL4*)
  46.181 -  "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X = z"
  46.182 -  by (blast intro: antisym cSup_upper cSup_least)
  46.183 +lemma cSup_upper2: "x \<in> X \<Longrightarrow> y \<le> x \<Longrightarrow> bdd_above X \<Longrightarrow> y \<le> Sup X"
  46.184 +  by (metis cSup_upper order_trans)
  46.185  
  46.186 -lemma cInf_eq_minimum: (*REAL_INF_MIN in HOL4*)
  46.187 -  "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X = z"
  46.188 -  by (intro antisym cInf_lower[of z X z] cInf_greatest[of X z]) auto
  46.189 +lemma cInf_lower2: "x \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> y"
  46.190 +  by (metis cInf_lower order_trans)
  46.191  
  46.192 -lemma cSup_le_iff: "S \<noteq> {} \<Longrightarrow> (\<And>a. a \<in> S \<Longrightarrow> a \<le> z) \<Longrightarrow> Sup S \<le> a \<longleftrightarrow> (\<forall>x\<in>S. x \<le> a)"
  46.193 +lemma cSup_mono: "B \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> (\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. b \<le> a) \<Longrightarrow> Sup B \<le> Sup A"
  46.194 +  by (metis cSup_least cSup_upper2)
  46.195 +
  46.196 +lemma cInf_mono: "B \<noteq> {} \<Longrightarrow> bdd_below A \<Longrightarrow> (\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<le> b) \<Longrightarrow> Inf A \<le> Inf B"
  46.197 +  by (metis cInf_greatest cInf_lower2)
  46.198 +
  46.199 +lemma cSup_subset_mono: "A \<noteq> {} \<Longrightarrow> bdd_above B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Sup A \<le> Sup B"
  46.200 +  by (metis cSup_least cSup_upper subsetD)
  46.201 +
  46.202 +lemma cInf_superset_mono: "A \<noteq> {} \<Longrightarrow> bdd_below B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Inf B \<le> Inf A"
  46.203 +  by (metis cInf_greatest cInf_lower subsetD)
  46.204 +
  46.205 +lemma cSup_eq_maximum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X = z"
  46.206 +  by (intro antisym cSup_upper[of z X] cSup_least[of X z]) auto
  46.207 +
  46.208 +lemma cInf_eq_minimum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X = z"
  46.209 +  by (intro antisym cInf_lower[of z X] cInf_greatest[of X z]) auto
  46.210 +
  46.211 +lemma cSup_le_iff: "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> Sup S \<le> a \<longleftrightarrow> (\<forall>x\<in>S. x \<le> a)"
  46.212    by (metis order_trans cSup_upper cSup_least)
  46.213  
  46.214 -lemma le_cInf_iff: "S \<noteq> {} \<Longrightarrow> (\<And>a. a \<in> S \<Longrightarrow> z \<le> a) \<Longrightarrow> a \<le> Inf S \<longleftrightarrow> (\<forall>x\<in>S. a \<le> x)"
  46.215 +lemma le_cInf_iff: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> a \<le> Inf S \<longleftrightarrow> (\<forall>x\<in>S. a \<le> x)"
  46.216    by (metis order_trans cInf_lower cInf_greatest)
  46.217  
  46.218 -lemma cSup_singleton [simp]: "Sup {x} = x"
  46.219 -  by (intro cSup_eq_maximum) auto
  46.220 -
  46.221 -lemma cInf_singleton [simp]: "Inf {x} = x"
  46.222 -  by (intro cInf_eq_minimum) auto
  46.223 -
  46.224 -lemma cSup_upper2: (*REAL_IMP_LE_SUP in HOL4*)
  46.225 -  "x \<in> X \<Longrightarrow> y \<le> x \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> y \<le> Sup X"
  46.226 -  by (metis cSup_upper order_trans)
  46.227 - 
  46.228 -lemma cInf_lower2:
  46.229 -  "x \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X \<le> y"
  46.230 -  by (metis cInf_lower order_trans)
  46.231 -
  46.232 -lemma cSup_upper_EX: "x \<in> X \<Longrightarrow> \<exists>z. \<forall>x. x \<in> X \<longrightarrow> x \<le> z \<Longrightarrow> x \<le> Sup X"
  46.233 -  by (blast intro: cSup_upper)
  46.234 -
  46.235 -lemma cInf_lower_EX:  "x \<in> X \<Longrightarrow> \<exists>z. \<forall>x. x \<in> X \<longrightarrow> z \<le> x \<Longrightarrow> Inf X \<le> x"
  46.236 -  by (blast intro: cInf_lower)
  46.237 -
  46.238  lemma cSup_eq_non_empty:
  46.239    assumes 1: "X \<noteq> {}"
  46.240    assumes 2: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
  46.241 @@ -77,67 +214,47 @@
  46.242    shows "Inf X = a"
  46.243    by (intro 3 1 antisym cInf_greatest) (auto intro: 2 1 cInf_lower)
  46.244  
  46.245 -lemma cInf_cSup: "S \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf S = Sup {x. \<forall>s\<in>S. x \<le> s}"
  46.246 -  by (rule cInf_eq_non_empty) (auto intro: cSup_upper cSup_least)
  46.247 +lemma cInf_cSup: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> Inf S = Sup {x. \<forall>s\<in>S. x \<le> s}"
  46.248 +  by (rule cInf_eq_non_empty) (auto intro!: cSup_upper cSup_least simp: bdd_below_def)
  46.249  
  46.250 -lemma cSup_cInf: "S \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup S = Inf {x. \<forall>s\<in>S. s \<le> x}"
  46.251 -  by (rule cSup_eq_non_empty) (auto intro: cInf_lower cInf_greatest)
  46.252 +lemma cSup_cInf: "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> Sup S = Inf {x. \<forall>s\<in>S. s \<le> x}"
  46.253 +  by (rule cSup_eq_non_empty) (auto intro!: cInf_lower cInf_greatest simp: bdd_above_def)
  46.254  
  46.255 -lemma cSup_insert: 
  46.256 -  assumes x: "X \<noteq> {}"
  46.257 -      and z: "\<And>x. x \<in> X \<Longrightarrow> x \<le> z"
  46.258 -  shows "Sup (insert a X) = sup a (Sup X)"
  46.259 -proof (intro cSup_eq_non_empty)
  46.260 -  fix y assume "\<And>x. x \<in> insert a X \<Longrightarrow> x \<le> y" with x show "sup a (Sup X) \<le> y" by (auto intro: cSup_least)
  46.261 -qed (auto intro: le_supI2 z cSup_upper)
  46.262 +lemma cSup_insert: "X \<noteq> {} \<Longrightarrow> bdd_above X \<Longrightarrow> Sup (insert a X) = sup a (Sup X)"
  46.263 +  by (intro cSup_eq_non_empty) (auto intro: le_supI2 cSup_upper cSup_least)
  46.264  
  46.265 -lemma cInf_insert: 
  46.266 -  assumes x: "X \<noteq> {}"
  46.267 -      and z: "\<And>x. x \<in> X \<Longrightarrow> z \<le> x"
  46.268 -  shows "Inf (insert a X) = inf a (Inf X)"
  46.269 -proof (intro cInf_eq_non_empty)
  46.270 -  fix y assume "\<And>x. x \<in> insert a X \<Longrightarrow> y \<le> x" with x show "y \<le> inf a (Inf X)" by (auto intro: cInf_greatest)
  46.271 -qed (auto intro: le_infI2 z cInf_lower)
  46.272 +lemma cInf_insert: "X \<noteq> {} \<Longrightarrow> bdd_below X \<Longrightarrow> Inf (insert a X) = inf a (Inf X)"
  46.273 +  by (intro cInf_eq_non_empty) (auto intro: le_infI2 cInf_lower cInf_greatest)
  46.274  
  46.275 -lemma cSup_insert_If: 
  46.276 -  "(\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup (insert a X) = (if X = {} then a else sup a (Sup X))"
  46.277 -  using cSup_insert[of X z] by simp
  46.278 +lemma cSup_singleton [simp]: "Sup {x} = x"
  46.279 +  by (intro cSup_eq_maximum) auto
  46.280  
  46.281 -lemma cInf_insert_if: 
  46.282 -  "(\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf (insert a X) = (if X = {} then a else inf a (Inf X))"
  46.283 -  using cInf_insert[of X z] by simp
  46.284 +lemma cInf_singleton [simp]: "Inf {x} = x"
  46.285 +  by (intro cInf_eq_minimum) auto
  46.286 +
  46.287 +lemma cSup_insert_If:  "bdd_above X \<Longrightarrow> Sup (insert a X) = (if X = {} then a else sup a (Sup X))"
  46.288 +  using cSup_insert[of X] by simp
  46.289 +
  46.290 +lemma cInf_insert_If: "bdd_below X \<Longrightarrow> Inf (insert a X) = (if X = {} then a else inf a (Inf X))"
  46.291 +  using cInf_insert[of X] by simp
  46.292  
  46.293  lemma le_cSup_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> x \<le> Sup X"
  46.294  proof (induct X arbitrary: x rule: finite_induct)
  46.295    case (insert x X y) then show ?case
  46.296 -    apply (cases "X = {}")
  46.297 -    apply simp
  46.298 -    apply (subst cSup_insert[of _ "Sup X"])
  46.299 -    apply (auto intro: le_supI2)
  46.300 -    done
  46.301 +    by (cases "X = {}") (auto simp: cSup_insert intro: le_supI2)
  46.302  qed simp
  46.303  
  46.304  lemma cInf_le_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> Inf X \<le> x"
  46.305  proof (induct X arbitrary: x rule: finite_induct)
  46.306    case (insert x X y) then show ?case
  46.307 -    apply (cases "X = {}")
  46.308 -    apply simp
  46.309 -    apply (subst cInf_insert[of _ "Inf X"])
  46.310 -    apply (auto intro: le_infI2)
  46.311 -    done
  46.312 +    by (cases "X = {}") (auto simp: cInf_insert intro: le_infI2)
  46.313  qed simp
  46.314  
  46.315  lemma cSup_eq_Sup_fin: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Sup_fin X"
  46.316 -proof (induct X rule: finite_ne_induct)
  46.317 -  case (insert x X) then show ?case
  46.318 -    using cSup_insert[of X "Sup_fin X" x] le_cSup_finite[of X] by simp
  46.319 -qed simp
  46.320 +  by (induct X rule: finite_ne_induct) (simp_all add: cSup_insert)
  46.321  
  46.322  lemma cInf_eq_Inf_fin: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Inf_fin X"
  46.323 -proof (induct X rule: finite_ne_induct)
  46.324 -  case (insert x X) then show ?case
  46.325 -    using cInf_insert[of X "Inf_fin X" x] cInf_le_finite[of X] by simp
  46.326 -qed simp
  46.327 +  by (induct X rule: finite_ne_induct) (simp_all add: cInf_insert)
  46.328  
  46.329  lemma cSup_atMost[simp]: "Sup {..x} = x"
  46.330    by (auto intro!: cSup_eq_maximum)
  46.331 @@ -157,16 +274,91 @@
  46.332  lemma cInf_atLeastAtMost[simp]: "y \<le> x \<Longrightarrow> Inf {y..x} = y"
  46.333    by (auto intro!: cInf_eq_minimum)
  46.334  
  46.335 +lemma cINF_lower: "bdd_below (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> INFI A f \<le> f x"
  46.336 +  unfolding INF_def by (rule cInf_lower) auto
  46.337 +
  46.338 +lemma cINF_greatest: "A \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> m \<le> f x) \<Longrightarrow> m \<le> INFI A f"
  46.339 +  unfolding INF_def by (rule cInf_greatest) auto
  46.340 +
  46.341 +lemma cSUP_upper: "x \<in> A \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> f x \<le> SUPR A f"
  46.342 +  unfolding SUP_def by (rule cSup_upper) auto
  46.343 +
  46.344 +lemma cSUP_least: "A \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<le> M) \<Longrightarrow> SUPR A f \<le> M"
  46.345 +  unfolding SUP_def by (rule cSup_least) auto
  46.346 +
  46.347 +lemma cINF_lower2: "bdd_below (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> f x \<le> u \<Longrightarrow> INFI A f \<le> u"
  46.348 +  by (auto intro: cINF_lower assms order_trans)
  46.349 +
  46.350 +lemma cSUP_upper2: "bdd_above (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> u \<le> f x \<Longrightarrow> u \<le> SUPR A f"
  46.351 +  by (auto intro: cSUP_upper assms order_trans)
  46.352 +
  46.353 +lemma cSUP_const: "A \<noteq> {} \<Longrightarrow> (SUP x:A. c) = c"
  46.354 +  by (intro antisym cSUP_least) (auto intro: cSUP_upper)
  46.355 +
  46.356 +lemma cINF_const: "A \<noteq> {} \<Longrightarrow> (INF x:A. c) = c"
  46.357 +  by (intro antisym cINF_greatest) (auto intro: cINF_lower)
  46.358 +
  46.359 +lemma le_cINF_iff: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> u \<le> INFI A f \<longleftrightarrow> (\<forall>x\<in>A. u \<le> f x)"
  46.360 +  by (metis cINF_greatest cINF_lower assms order_trans)
  46.361 +
  46.362 +lemma cSUP_le_iff: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> SUPR A f \<le> u \<longleftrightarrow> (\<forall>x\<in>A. f x \<le> u)"
  46.363 +  by (metis cSUP_least cSUP_upper assms order_trans)
  46.364 +
  46.365 +lemma less_cINF_D: "bdd_below (f`A) \<Longrightarrow> y < (INF i:A. f i) \<Longrightarrow> i \<in> A \<Longrightarrow> y < f i"
  46.366 +  by (metis cINF_lower less_le_trans)
  46.367 +
  46.368 +lemma cSUP_lessD: "bdd_above (f`A) \<Longrightarrow> (SUP i:A. f i) < y \<Longrightarrow> i \<in> A \<Longrightarrow> f i < y"
  46.369 +  by (metis cSUP_upper le_less_trans)
  46.370 +
  46.371 +lemma cINF_insert: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> INFI (insert a A) f = inf (f a) (INFI A f)"
  46.372 +  by (metis INF_def cInf_insert assms empty_is_image image_insert)
  46.373 +
  46.374 +lemma cSUP_insert: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> SUPR (insert a A) f = sup (f a) (SUPR A f)"
  46.375 +  by (metis SUP_def cSup_insert assms empty_is_image image_insert)
  46.376 +
  46.377 +lemma cINF_mono: "B \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> (\<And>m. m \<in> B \<Longrightarrow> \<exists>n\<in>A. f n \<le> g m) \<Longrightarrow> INFI A f \<le> INFI B g"
  46.378 +  unfolding INF_def by (auto intro: cInf_mono)
  46.379 +
  46.380 +lemma cSUP_mono: "A \<noteq> {} \<Longrightarrow> bdd_above (g ` B) \<Longrightarrow> (\<And>n. n \<in> A \<Longrightarrow> \<exists>m\<in>B. f n \<le> g m) \<Longrightarrow> SUPR A f \<le> SUPR B g"
  46.381 +  unfolding SUP_def by (auto intro: cSup_mono)
  46.382 +
  46.383 +lemma cINF_superset_mono: "A \<noteq> {} \<Longrightarrow> bdd_below (g ` B) \<Longrightarrow> A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> g x \<le> f x) \<Longrightarrow> INFI B g \<le> INFI A f"
  46.384 +  by (rule cINF_mono) auto
  46.385 +
  46.386 +lemma cSUP_subset_mono: "A \<noteq> {} \<Longrightarrow> bdd_above (g ` B) \<Longrightarrow> A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> f x \<le> g x) \<Longrightarrow> SUPR A f \<le> SUPR B g"
  46.387 +  by (rule cSUP_mono) auto
  46.388 +
  46.389 +lemma less_eq_cInf_inter: "bdd_below A \<Longrightarrow> bdd_below B \<Longrightarrow> A \<inter> B \<noteq> {} \<Longrightarrow> inf (Inf A) (Inf B) \<le> Inf (A \<inter> B)"
  46.390 +  by (metis cInf_superset_mono lattice_class.inf_sup_ord(1) le_infI1)
  46.391 +
  46.392 +lemma cSup_inter_less_eq: "bdd_above A \<Longrightarrow> bdd_above B \<Longrightarrow> A \<inter> B \<noteq> {} \<Longrightarrow> Sup (A \<inter> B) \<le> sup (Sup A) (Sup B) "
  46.393 +  by (metis cSup_subset_mono lattice_class.inf_sup_ord(1) le_supI1)
  46.394 +
  46.395 +lemma cInf_union_distrib: "A \<noteq> {} \<Longrightarrow> bdd_below A \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_below B \<Longrightarrow> Inf (A \<union> B) = inf (Inf A) (Inf B)"
  46.396 +  by (intro antisym le_infI cInf_greatest cInf_lower) (auto intro: le_infI1 le_infI2 cInf_lower)
  46.397 +
  46.398 +lemma cINF_union: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_below (f`B) \<Longrightarrow> INFI (A \<union> B) f = inf (INFI A f) (INFI B f)"
  46.399 +  unfolding INF_def using assms by (auto simp add: image_Un intro: cInf_union_distrib)
  46.400 +
  46.401 +lemma cSup_union_distrib: "A \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_above B \<Longrightarrow> Sup (A \<union> B) = sup (Sup A) (Sup B)"
  46.402 +  by (intro antisym le_supI cSup_least cSup_upper) (auto intro: le_supI1 le_supI2 cSup_upper)
  46.403 +
  46.404 +lemma cSUP_union: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_above (f`B) \<Longrightarrow> SUPR (A \<union> B) f = sup (SUPR A f) (SUPR B f)"
  46.405 +  unfolding SUP_def by (auto simp add: image_Un intro: cSup_union_distrib)
  46.406 +
  46.407 +lemma cINF_inf_distrib: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> bdd_below (g`A) \<Longrightarrow> inf (INFI A f) (INFI A g) = (INF a:A. inf (f a) (g a))"
  46.408 +  by (intro antisym le_infI cINF_greatest cINF_lower2)
  46.409 +     (auto intro: le_infI1 le_infI2 cINF_greatest cINF_lower le_infI)
  46.410 +
  46.411 +lemma SUP_sup_distrib: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> bdd_above (g`A) \<Longrightarrow> sup (SUPR A f) (SUPR A g) = (SUP a:A. sup (f a) (g a))"
  46.412 +  by (intro antisym le_supI cSUP_least cSUP_upper2)
  46.413 +     (auto intro: le_supI1 le_supI2 cSUP_least cSUP_upper le_supI)
  46.414 +
  46.415  end
  46.416  
  46.417  instance complete_lattice \<subseteq> conditionally_complete_lattice
  46.418    by default (auto intro: Sup_upper Sup_least Inf_lower Inf_greatest)
  46.419  
  46.420 -lemma isLub_cSup: 
  46.421 -  "(S::'a :: conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> (\<exists>b. S *<= b) \<Longrightarrow> isLub UNIV S (Sup S)"
  46.422 -  by  (auto simp add: isLub_def setle_def leastP_def isUb_def
  46.423 -            intro!: setgeI intro: cSup_upper cSup_least)
  46.424 -
  46.425  lemma cSup_eq:
  46.426    fixes a :: "'a :: {conditionally_complete_lattice, no_bot}"
  46.427    assumes upper: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
  46.428 @@ -185,33 +377,33 @@
  46.429    assume "X = {}" with gt_ex[of a] least show ?thesis by (auto simp: less_le_not_le)
  46.430  qed (intro cInf_eq_non_empty assms)
  46.431  
  46.432 -lemma cSup_le: "(S::'a::conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> S *<= b \<Longrightarrow> Sup S \<le> b"
  46.433 -  by (metis cSup_least setle_def)
  46.434 -
  46.435 -lemma cInf_ge: "(S::'a :: conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> b <=* S \<Longrightarrow> Inf S \<ge> b"
  46.436 -  by (metis cInf_greatest setge_def)
  46.437 -
  46.438  class conditionally_complete_linorder = conditionally_complete_lattice + linorder
  46.439  begin
  46.440  
  46.441  lemma less_cSup_iff : (*REAL_SUP_LE in HOL4*)
  46.442 -  "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> y < Sup X \<longleftrightarrow> (\<exists>x\<in>X. y < x)"
  46.443 +  "X \<noteq> {} \<Longrightarrow> bdd_above X \<Longrightarrow> y < Sup X \<longleftrightarrow> (\<exists>x\<in>X. y < x)"
  46.444    by (rule iffI) (metis cSup_least not_less, metis cSup_upper less_le_trans)
  46.445  
  46.446 -lemma cInf_less_iff: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X < y \<longleftrightarrow> (\<exists>x\<in>X. x < y)"
  46.447 +lemma cInf_less_iff: "X \<noteq> {} \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X < y \<longleftrightarrow> (\<exists>x\<in>X. x < y)"
  46.448    by (rule iffI) (metis cInf_greatest not_less, metis cInf_lower le_less_trans)
  46.449  
  46.450 +lemma cINF_less_iff: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> (INF i:A. f i) < a \<longleftrightarrow> (\<exists>x\<in>A. f x < a)"
  46.451 +  unfolding INF_def using cInf_less_iff[of "f`A"] by auto
  46.452 +
  46.453 +lemma less_cSUP_iff: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> a < (SUP i:A. f i) \<longleftrightarrow> (\<exists>x\<in>A. a < f x)"
  46.454 +  unfolding SUP_def using less_cSup_iff[of "f`A"] by auto
  46.455 +
  46.456  lemma less_cSupE:
  46.457    assumes "y < Sup X" "X \<noteq> {}" obtains x where "x \<in> X" "y < x"
  46.458    by (metis cSup_least assms not_le that)
  46.459  
  46.460  lemma less_cSupD:
  46.461    "X \<noteq> {} \<Longrightarrow> z < Sup X \<Longrightarrow> \<exists>x\<in>X. z < x"
  46.462 -  by (metis less_cSup_iff not_leE)
  46.463 +  by (metis less_cSup_iff not_leE bdd_above_def)
  46.464  
  46.465  lemma cInf_lessD:
  46.466    "X \<noteq> {} \<Longrightarrow> Inf X < z \<Longrightarrow> \<exists>x\<in>X. x < z"
  46.467 -  by (metis cInf_less_iff not_leE)
  46.468 +  by (metis cInf_less_iff not_leE bdd_below_def)
  46.469  
  46.470  lemma complete_interval:
  46.471    assumes "a < b" and "P a" and "\<not> P b"
  46.472 @@ -219,7 +411,7 @@
  46.473               (\<forall>d. (\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x) \<longrightarrow> d \<le> c)"
  46.474  proof (rule exI [where x = "Sup {d. \<forall>x. a \<le> x & x < d --> P x}"], auto)
  46.475    show "a \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
  46.476 -    by (rule cSup_upper [where z=b], auto)
  46.477 +    by (rule cSup_upper, auto simp: bdd_above_def)
  46.478         (metis `a < b` `\<not> P b` linear less_le)
  46.479  next
  46.480    show "Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c} \<le> b"
  46.481 @@ -240,12 +432,36 @@
  46.482    fix d
  46.483      assume 0: "\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x"
  46.484      thus "d \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
  46.485 -      by (rule_tac z="b" in cSup_upper, auto) 
  46.486 +      by (rule_tac cSup_upper, auto simp: bdd_above_def)
  46.487           (metis `a<b` `~ P b` linear less_le)
  46.488  qed
  46.489  
  46.490  end
  46.491  
  46.492 +lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Max X"
  46.493 +  using cSup_eq_Sup_fin[of X] Sup_fin_eq_Max[of X] by simp
  46.494 +
  46.495 +lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Min X"
  46.496 +  using cInf_eq_Inf_fin[of X] Inf_fin_eq_Min[of X] by simp
  46.497 +
  46.498 +lemma cSup_lessThan[simp]: "Sup {..<x::'a::{conditionally_complete_linorder, no_bot, dense_linorder}} = x"
  46.499 +  by (auto intro!: cSup_eq_non_empty intro: dense_le)
  46.500 +
  46.501 +lemma cSup_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Sup {y<..<x::'a::{conditionally_complete_linorder, no_bot, dense_linorder}} = x"
  46.502 +  by (auto intro!: cSup_eq intro: dense_le_bounded)
  46.503 +
  46.504 +lemma cSup_atLeastLessThan[simp]: "y < x \<Longrightarrow> Sup {y..<x::'a::{conditionally_complete_linorder, no_bot, dense_linorder}} = x"
  46.505 +  by (auto intro!: cSup_eq intro: dense_le_bounded)
  46.506 +
  46.507 +lemma cInf_greaterThan[simp]: "Inf {x::'a::{conditionally_complete_linorder, no_top, dense_linorder} <..} = x"
  46.508 +  by (auto intro!: cInf_eq intro: dense_ge)
  46.509 +
  46.510 +lemma cInf_greaterThanAtMost[simp]: "y < x \<Longrightarrow> Inf {y<..x::'a::{conditionally_complete_linorder, no_top, dense_linorder}} = y"
  46.511 +  by (auto intro!: cInf_eq intro: dense_ge_bounded)
  46.512 +
  46.513 +lemma cInf_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Inf {y<..<x::'a::{conditionally_complete_linorder, no_top, dense_linorder}} = y"
  46.514 +  by (auto intro!: cInf_eq intro: dense_ge_bounded)
  46.515 +
  46.516  class linear_continuum = conditionally_complete_linorder + dense_linorder +
  46.517    assumes UNIV_not_singleton: "\<exists>a b::'a. a \<noteq> b"
  46.518  begin
  46.519 @@ -255,50 +471,92 @@
  46.520  
  46.521  end
  46.522  
  46.523 -lemma cSup_bounds:
  46.524 -  fixes S :: "'a :: conditionally_complete_lattice set"
  46.525 -  assumes Se: "S \<noteq> {}" and l: "a <=* S" and u: "S *<= b"
  46.526 -  shows "a \<le> Sup S \<and> Sup S \<le> b"
  46.527 -proof-
  46.528 -  from isLub_cSup[OF Se] u have lub: "isLub UNIV S (Sup S)" by blast
  46.529 -  hence b: "Sup S \<le> b" using u 
  46.530 -    by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def) 
  46.531 -  from Se obtain y where y: "y \<in> S" by blast
  46.532 -  from lub l have "a \<le> Sup S"
  46.533 -    by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def)
  46.534 -       (metis le_iff_sup le_sup_iff y)
  46.535 -  with b show ?thesis by blast
  46.536 +instantiation nat :: conditionally_complete_linorder
  46.537 +begin
  46.538 +
  46.539 +definition "Sup (X::nat set) = Max X"
  46.540 +definition "Inf (X::nat set) = (LEAST n. n \<in> X)"
  46.541 +
  46.542 +lemma bdd_above_nat: "bdd_above X \<longleftrightarrow> finite (X::nat set)"
  46.543 +proof
  46.544 +  assume "bdd_above X"
  46.545 +  then obtain z where "X \<subseteq> {.. z}"
  46.546 +    by (auto simp: bdd_above_def)
  46.547 +  then show "finite X"
  46.548 +    by (rule finite_subset) simp
  46.549 +qed simp
  46.550 +
  46.551 +instance
  46.552 +proof
  46.553 +  fix x :: nat and X :: "nat set"
  46.554 +  { assume "x \<in> X" "bdd_below X" then show "Inf X \<le> x"
  46.555 +      by (simp add: Inf_nat_def Least_le) }
  46.556 +  { assume "X \<noteq> {}" "\<And>y. y \<in> X \<Longrightarrow> x \<le> y" then show "x \<le> Inf X"
  46.557 +      unfolding Inf_nat_def ex_in_conv[symmetric] by (rule LeastI2_ex) }
  46.558 +  { assume "x \<in> X" "bdd_above X" then show "x \<le> Sup X"
  46.559 +      by (simp add: Sup_nat_def bdd_above_nat) }
  46.560 +  { assume "X \<noteq> {}" "\<And>y. y \<in> X \<Longrightarrow> y \<le> x" 
  46.561 +    moreover then have "bdd_above X"
  46.562 +      by (auto simp: bdd_above_def)
  46.563 +    ultimately show "Sup X \<le> x"
  46.564 +      by (simp add: Sup_nat_def bdd_above_nat) }
  46.565  qed
  46.566 +end
  46.567  
  46.568 +instantiation int :: conditionally_complete_linorder
  46.569 +begin
  46.570  
  46.571 -lemma cSup_unique: "(S::'a :: {conditionally_complete_linorder, no_bot} set) *<= b \<Longrightarrow> (\<forall>b'<b. \<exists>x\<in>S. b' < x) \<Longrightarrow> Sup S = b"
  46.572 -  by (rule cSup_eq) (auto simp: not_le[symmetric] setle_def)
  46.573 +definition "Sup (X::int set) = (THE x. x \<in> X \<and> (\<forall>y\<in>X. y \<le> x))"
  46.574 +definition "Inf (X::int set) = - (Sup (uminus ` X))"
  46.575  
  46.576 -lemma cInf_unique: "b <=* (S::'a :: {conditionally_complete_linorder, no_top} set) \<Longrightarrow> (\<forall>b'>b. \<exists>x\<in>S. b' > x) \<Longrightarrow> Inf S = b"
  46.577 -  by (rule cInf_eq) (auto simp: not_le[symmetric] setge_def)
  46.578 +instance
  46.579 +proof
  46.580 +  { fix x :: int and X :: "int set" assume "X \<noteq> {}" "bdd_above X"
  46.581 +    then obtain x y where "X \<subseteq> {..y}" "x \<in> X"
  46.582 +      by (auto simp: bdd_above_def)
  46.583 +    then have *: "finite (X \<inter> {x..y})" "X \<inter> {x..y} \<noteq> {}" and "x \<le> y"
  46.584 +      by (auto simp: subset_eq)
  46.585 +    have "\<exists>!x\<in>X. (\<forall>y\<in>X. y \<le> x)"
  46.586 +    proof
  46.587 +      { fix z assume "z \<in> X"
  46.588 +        have "z \<le> Max (X \<inter> {x..y})"
  46.589 +        proof cases
  46.590 +          assume "x \<le> z" with `z \<in> X` `X \<subseteq> {..y}` *(1) show ?thesis
  46.591 +            by (auto intro!: Max_ge)
  46.592 +        next
  46.593 +          assume "\<not> x \<le> z"
  46.594 +          then have "z < x" by simp
  46.595 +          also have "x \<le> Max (X \<inter> {x..y})"
  46.596 +            using `x \<in> X` *(1) `x \<le> y` by (intro Max_ge) auto
  46.597 +          finally show ?thesis by simp
  46.598 +        qed }
  46.599 +      note le = this
  46.600 +      with Max_in[OF *] show ex: "Max (X \<inter> {x..y}) \<in> X \<and> (\<forall>z\<in>X. z \<le> Max (X \<inter> {x..y}))" by auto
  46.601  
  46.602 -lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Max X"
  46.603 -  using cSup_eq_Sup_fin[of X] Sup_fin_eq_Max[of X] by simp
  46.604 +      fix z assume *: "z \<in> X \<and> (\<forall>y\<in>X. y \<le> z)"
  46.605 +      with le have "z \<le> Max (X \<inter> {x..y})"
  46.606 +        by auto
  46.607 +      moreover have "Max (X \<inter> {x..y}) \<le> z"
  46.608 +        using * ex by auto
  46.609 +      ultimately show "z = Max (X \<inter> {x..y})"
  46.610 +        by auto
  46.611 +    qed
  46.612 +    then have "Sup X \<in> X \<and> (\<forall>y\<in>X. y \<le> Sup X)"
  46.613 +      unfolding Sup_int_def by (rule theI') }
  46.614 +  note Sup_int = this
  46.615  
  46.616 -lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Min X"
  46.617 -  using cInf_eq_Inf_fin[of X] Inf_fin_eq_Min[of X] by simp
  46.618 +  { fix x :: int and X :: "int set" assume "x \<in> X" "bdd_above X" then show "x \<le> Sup X"
  46.619 +      using Sup_int[of X] by auto }
  46.620 +  note le_Sup = this
  46.621 +  { fix x :: int and X :: "int set" assume "X \<noteq> {}" "\<And>y. y \<in> X \<Longrightarrow> y \<le> x" then show "Sup X \<le> x"
  46.622 +      using Sup_int[of X] by (auto simp: bdd_above_def) }
  46.623 +  note Sup_le = this
  46.624  
  46.625 -lemma cSup_lessThan[simp]: "Sup {..<x::'a::{conditionally_complete_linorder, unbounded_dense_linorder}} = x"
  46.626 -  by (auto intro!: cSup_eq_non_empty intro: dense_le)
  46.627 -
  46.628 -lemma cSup_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Sup {y<..<x::'a::{conditionally_complete_linorder, unbounded_dense_linorder}} = x"
  46.629 -  by (auto intro!: cSup_eq intro: dense_le_bounded)
  46.630 -
  46.631 -lemma cSup_atLeastLessThan[simp]: "y < x \<Longrightarrow> Sup {y..<x::'a::{conditionally_complete_linorder, unbounded_dense_linorder}} = x"
  46.632 -  by (auto intro!: cSup_eq intro: dense_le_bounded)
  46.633 -
  46.634 -lemma cInf_greaterThan[simp]: "Inf {x::'a::{conditionally_complete_linorder, unbounded_dense_linorder} <..} = x"
  46.635 -  by (auto intro!: cInf_eq intro: dense_ge)
  46.636 -
  46.637 -lemma cInf_greaterThanAtMost[simp]: "y < x \<Longrightarrow> Inf {y<..x::'a::{conditionally_complete_linorder, unbounded_dense_linorder}} = y"
  46.638 -  by (auto intro!: cInf_eq intro: dense_ge_bounded)
  46.639 -
  46.640 -lemma cInf_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Inf {y<..<x::'a::{conditionally_complete_linorder, unbounded_dense_linorder}} = y"
  46.641 -  by (auto intro!: cInf_eq intro: dense_ge_bounded)
  46.642 +  { fix x :: int and X :: "int set" assume "x \<in> X" "bdd_below X" then show "Inf X \<le> x"
  46.643 +      using le_Sup[of "-x" "uminus ` X"] by (auto simp: Inf_int_def) }
  46.644 +  { fix x :: int and X :: "int set" assume "X \<noteq> {}" "\<And>y. y \<in> X \<Longrightarrow> x \<le> y" then show "x \<le> Inf X"
  46.645 +      using Sup_le[of "uminus ` X" "-x"] by (force simp: Inf_int_def) }
  46.646 +qed
  46.647 +end
  46.648  
  46.649  end
    47.1 --- a/src/HOL/Decision_Procs/Approximation.thy	Mon Nov 11 17:34:44 2013 +0100
    47.2 +++ b/src/HOL/Decision_Procs/Approximation.thy	Mon Nov 11 17:44:21 2013 +0100
    47.3 @@ -29,7 +29,7 @@
    47.4    have shift_pow: "\<And>i. - (x * ((-1)^i * a (Suc i) * x ^ i)) = (-1)^(Suc i) * a (Suc i) * x ^ (Suc i)"
    47.5      by auto
    47.6    show ?thesis
    47.7 -    unfolding setsum_right_distrib shift_pow diff_minus setsum_negf[symmetric]
    47.8 +    unfolding setsum_right_distrib shift_pow uminus_add_conv_diff [symmetric] setsum_negf[symmetric]
    47.9      setsum_head_upt_Suc[OF zero_less_Suc]
   47.10      setsum_reindex[OF inj_Suc, unfolded comp_def, symmetric, of "\<lambda> n. (-1)^n  *a n * x^n"] by auto
   47.11  qed
   47.12 @@ -132,14 +132,7 @@
   47.13  lemma get_odd[simp]: "odd (get_odd n)" unfolding get_odd_def by (cases "odd n", auto)
   47.14  lemma get_even[simp]: "even (get_even n)" unfolding get_even_def by (cases "even n", auto)
   47.15  lemma get_odd_ex: "\<exists> k. Suc k = get_odd n \<and> odd (Suc k)"
   47.16 -proof (cases "odd n")
   47.17 -  case True hence "0 < n" by (rule odd_pos)
   47.18 -  from gr0_implies_Suc[OF this] obtain k where "Suc k = n" by auto
   47.19 -  thus ?thesis unfolding get_odd_def if_P[OF True] using True[unfolded `Suc k = n`[symmetric]] by blast
   47.20 -next
   47.21 -  case False hence "odd (Suc n)" by auto
   47.22 -  thus ?thesis unfolding get_odd_def if_not_P[OF False] by blast
   47.23 -qed
   47.24 +  by (auto simp: get_odd_def odd_pos intro!: exI[of _ "n - 1"])
   47.25  
   47.26  lemma get_even_double: "\<exists>i. get_even n = 2 * i" using get_even[unfolded even_mult_two_ex] .
   47.27  lemma get_odd_double: "\<exists>i. get_odd n = 2 * i + 1" using get_odd[unfolded odd_Suc_mult_two_ex] by auto
   47.28 @@ -151,47 +144,9 @@
   47.29                        else if u < 0         then (u ^ n, l ^ n)
   47.30                                              else (0, (max (-l) u) ^ n))"
   47.31  
   47.32 -lemma float_power_bnds: fixes x :: real
   47.33 -  assumes "(l1, u1) = float_power_bnds n l u" and "x \<in> {l .. u}"
   47.34 -  shows "x ^ n \<in> {l1..u1}"
   47.35 -proof (cases "even n")
   47.36 -  case True
   47.37 -  show ?thesis
   47.38 -  proof (cases "0 < l")
   47.39 -    case True hence "odd n \<or> 0 < l" and "0 \<le> real l" by auto
   47.40 -    have u1: "u1 = u ^ n" and l1: "l1 = l ^ n" using assms
   47.41 -      unfolding float_power_bnds_def if_P[OF `odd n \<or> 0 < l`] by auto
   47.42 -    have "real l ^ n \<le> x ^ n" and "x ^ n \<le> real u ^ n " using `0 \<le> real l` assms
   47.43 -      by (auto simp: power_mono)
   47.44 -    thus ?thesis using assms `0 < l` unfolding l1 u1 by auto
   47.45 -  next
   47.46 -    case False hence P: "\<not> (odd n \<or> 0 < l)" using `even n` by auto
   47.47 -    show ?thesis
   47.48 -    proof (cases "u < 0")
   47.49 -      case True hence "0 \<le> - real u" and "- real u \<le> - x" and "0 \<le> - x" and "-x \<le> - real l" using assms  by auto
   47.50 -      hence "real u ^ n \<le> x ^ n" and "x ^ n \<le> real l ^ n" using power_mono[of  "-x" "-real l" n] power_mono[of "-real u" "-x" n]
   47.51 -        unfolding power_minus_even[OF `even n`] by auto
   47.52 -      moreover have u1: "u1 = l ^ n" and l1: "l1 = u ^ n" using assms unfolding float_power_bnds_def if_not_P[OF P] if_P[OF True] by auto
   47.53 -      ultimately show ?thesis by auto
   47.54 -    next
   47.55 -      case False
   47.56 -      have "\<bar>x\<bar> \<le> real (max (-l) u)"
   47.57 -      proof (cases "-l \<le> u")
   47.58 -        case True thus ?thesis unfolding max_def if_P[OF True] using assms by auto
   47.59 -      next
   47.60 -        case False thus ?thesis unfolding max_def if_not_P[OF False] using assms by auto
   47.61 -      qed
   47.62 -      hence x_abs: "\<bar>x\<bar> \<le> \<bar>real (max (-l) u)\<bar>" by auto
   47.63 -      have u1: "u1 = (max (-l) u) ^ n" and l1: "l1 = 0" using assms unfolding float_power_bnds_def if_not_P[OF P] if_not_P[OF False] by auto
   47.64 -      show ?thesis unfolding atLeastAtMost_iff l1 u1 using zero_le_even_power[OF `even n`] power_mono_even[OF `even n` x_abs] by auto
   47.65 -    qed
   47.66 -  qed
   47.67 -next
   47.68 -  case False hence "odd n \<or> 0 < l" by auto
   47.69 -  have u1: "u1 = u ^ n" and l1: "l1 = l ^ n" using assms unfolding float_power_bnds_def if_P[OF `odd n \<or> 0 < l`] by auto
   47.70 -  have "real l ^ n \<le> x ^ n" and "x ^ n \<le> real u ^ n " using assms unfolding atLeastAtMost_iff using power_mono_odd[OF False] by auto
   47.71 -  thus ?thesis unfolding atLeastAtMost_iff l1 u1 less_float_def by auto
   47.72 -qed
   47.73 +lemma float_power_bnds: "(l1, u1) = float_power_bnds n l u \<Longrightarrow> x \<in> {l .. u} \<Longrightarrow> (x::real) ^ n \<in> {l1..u1}"
   47.74 +  by (auto simp: float_power_bnds_def max_def split: split_if_asm
   47.75 +           intro: power_mono_odd power_mono power_mono_even zero_le_even_power)
   47.76  
   47.77  lemma bnds_power: "\<forall> (x::real) l u. (l1, u1) = float_power_bnds n l u \<and> x \<in> {l .. u} \<longrightarrow> l1 \<le> x ^ n \<and> x ^ n \<le> u1"
   47.78    using float_power_bnds by auto
   47.79 @@ -244,7 +199,7 @@
   47.80  qed
   47.81  
   47.82  lemma sqrt_iteration_bound: assumes "0 < real x"
   47.83 -  shows "sqrt x < (sqrt_iteration prec n x)"
   47.84 +  shows "sqrt x < sqrt_iteration prec n x"
   47.85  proof (induct n)
   47.86    case 0
   47.87    show ?case
   47.88 @@ -356,20 +311,8 @@
   47.89    note ub = this
   47.90  
   47.91    show ?thesis
   47.92 -  proof (cases "0 < x")
   47.93 -    case True with lb ub show ?thesis by auto
   47.94 -  next case False show ?thesis
   47.95 -  proof (cases "real x = 0")
   47.96 -    case True thus ?thesis
   47.97 -      by (auto simp add: lb_sqrt.simps ub_sqrt.simps)
   47.98 -  next
   47.99 -    case False with `\<not> 0 < x` have "x < 0" and "0 < -x"
  47.100 -      by auto
  47.101 -
  47.102 -    with `\<not> 0 < x`
  47.103 -    show ?thesis using lb[OF `0 < -x`] ub[OF `0 < -x`]
  47.104 -      by (auto simp add: real_sqrt_minus lb_sqrt.simps ub_sqrt.simps)
  47.105 -  qed qed
  47.106 +    using lb[of "-x"] ub[of "-x"] lb[of x] ub[of x]
  47.107 +    by (auto simp add: lb_sqrt.simps ub_sqrt.simps real_sqrt_minus)
  47.108  qed
  47.109  
  47.110  lemma bnds_sqrt: "\<forall> (x::real) lx ux. (l, u) = (lb_sqrt prec lx, ub_sqrt prec ux) \<and> x \<in> {lx .. ux} \<longrightarrow> l \<le> sqrt x \<and> sqrt x \<le> u"
  47.111 @@ -412,8 +355,8 @@
  47.112    assumes "0 \<le> real x" "real x \<le> 1" and "even n"
  47.113    shows "arctan x \<in> {(x * lb_arctan_horner prec n 1 (x * x)) .. (x * ub_arctan_horner prec (Suc n) 1 (x * x))}"
  47.114  proof -
  47.115 -  let "?c i" = "-1^i * (1 / (i * 2 + (1::nat)) * real x ^ (i * 2 + 1))"
  47.116 -  let "?S n" = "\<Sum> i=0..<n. ?c i"
  47.117 +  let ?c = "\<lambda>i. -1^i * (1 / (i * 2 + (1::nat)) * real x ^ (i * 2 + 1))"
  47.118 +  let ?S = "\<lambda>n. \<Sum> i=0..<n. ?c i"
  47.119  
  47.120    have "0 \<le> real (x * x)" by auto
  47.121    from `even n` obtain m where "2 * m = n" unfolding even_mult_two_ex by auto
  47.122 @@ -457,30 +400,11 @@
  47.123  
  47.124  lemma arctan_0_1_bounds: assumes "0 \<le> real x" "real x \<le> 1"
  47.125    shows "arctan x \<in> {(x * lb_arctan_horner prec (get_even n) 1 (x * x)) .. (x * ub_arctan_horner prec (get_odd n) 1 (x * x))}"
  47.126 -proof (cases "even n")
  47.127 -  case True
  47.128 -  obtain n' where "Suc n' = get_odd n" and "odd (Suc n')" using get_odd_ex by auto
  47.129 -  hence "even n'" unfolding even_Suc by auto
  47.130 -  have "arctan x \<le> x * ub_arctan_horner prec (get_odd n) 1 (x * x)"
  47.131 -    unfolding `Suc n' = get_odd n`[symmetric] using arctan_0_1_bounds'[OF `0 \<le> real x` `real x \<le> 1` `even n'`] by auto
  47.132 -  moreover
  47.133 -  have "x * lb_arctan_horner prec (get_even n) 1 (x * x) \<le> arctan x"
  47.134 -    unfolding get_even_def if_P[OF True] using arctan_0_1_bounds'[OF `0 \<le> real x` `real x \<le> 1` `even n`] by auto
  47.135 -  ultimately show ?thesis by auto
  47.136 -next
  47.137 -  case False hence "0 < n" by (rule odd_pos)
  47.138 -  from gr0_implies_Suc[OF this] obtain n' where "n = Suc n'" ..
  47.139 -  from False[unfolded this even_Suc]
  47.140 -  have "even n'" and "even (Suc (Suc n'))" by auto
  47.141 -  have "get_odd n = Suc n'" unfolding get_odd_def if_P[OF False] using `n = Suc n'` .
  47.142 -
  47.143 -  have "arctan x \<le> x * ub_arctan_horner prec (get_odd n) 1 (x * x)"
  47.144 -    unfolding `get_odd n = Suc n'` using arctan_0_1_bounds'[OF `0 \<le> real x` `real x \<le> 1` `even n'`] by auto
  47.145 -  moreover
  47.146 -  have "(x * lb_arctan_horner prec (get_even n) 1 (x * x)) \<le> arctan x"
  47.147 -    unfolding get_even_def if_not_P[OF False] unfolding `n = Suc n'` using arctan_0_1_bounds'[OF `0 \<le> real x` `real x \<le> 1` `even (Suc (Suc n'))`] by auto
  47.148 -  ultimately show ?thesis by auto
  47.149 -qed
  47.150 +  using
  47.151 +    arctan_0_1_bounds'[OF assms, of n prec]
  47.152 +    arctan_0_1_bounds'[OF assms, of "n + 1" prec]
  47.153 +    arctan_0_1_bounds'[OF assms, of "n - 1" prec]
  47.154 +  by (auto simp: get_even_def get_odd_def odd_pos simp del: ub_arctan_horner.simps lb_arctan_horner.simps)
  47.155  
  47.156  subsection "Compute \<pi>"
  47.157  
  47.158 @@ -530,16 +454,11 @@
  47.159      finally have "?k * lb_arctan_horner prec (get_even n) 1 (?k * ?k) \<le> arctan (1 / k)" .
  47.160    } note lb_arctan = this
  47.161  
  47.162 -  have "pi \<le> ub_pi n"
  47.163 -    unfolding ub_pi_def machin_pi Let_def unfolding Float_num
  47.164 -    using lb_arctan[of 239] ub_arctan[of 5] powr_realpow[of 2 2]
  47.165 -    by (auto intro!: mult_left_mono add_mono simp add: diff_minus)
  47.166 -  moreover
  47.167 -  have "lb_pi n \<le> pi"
  47.168 -    unfolding lb_pi_def machin_pi Let_def Float_num
  47.169 -    using lb_arctan[of 5] ub_arctan[of 239] powr_realpow[of 2 2]
  47.170 -    by (auto intro!: mult_left_mono add_mono simp add: diff_minus)
  47.171 -  ultimately show ?thesis by auto
  47.172 +  have "pi \<le> ub_pi n \<and> lb_pi n \<le> pi"
  47.173 +    unfolding lb_pi_def ub_pi_def machin_pi Let_def unfolding Float_num
  47.174 +    using lb_arctan[of 5] ub_arctan[of 239] lb_arctan[of 239] ub_arctan[of 5] powr_realpow[of 2 2]
  47.175 +    by (auto intro!: mult_left_mono add_mono simp add: uminus_add_conv_diff [symmetric] simp del: uminus_add_conv_diff)
  47.176 +  then show ?thesis by auto
  47.177  qed
  47.178  
  47.179  subsection "Compute arcus tangens in the entire domain"
  47.180 @@ -1208,8 +1127,8 @@
  47.181      using x unfolding k[symmetric]
  47.182      by (cases "k = 0")
  47.183         (auto intro!: add_mono
  47.184 -                simp add: diff_minus k[symmetric]
  47.185 -                simp del: float_of_numeral)
  47.186 +                simp add: k [symmetric] uminus_add_conv_diff [symmetric]
  47.187 +                simp del: float_of_numeral uminus_add_conv_diff)
  47.188    note lx = this[THEN conjunct1] and ux = this[THEN conjunct2]
  47.189    hence lx_less_ux: "?lx \<le> real ?ux" by (rule order_trans)
  47.190  
  47.191 @@ -1223,7 +1142,7 @@
  47.192      also have "\<dots> \<le> cos (x + (-k) * (2 * pi))"
  47.193        using cos_monotone_minus_pi_0'[OF pi_lx lx x_le_0]
  47.194        by (simp only: uminus_float.rep_eq real_of_int_minus
  47.195 -        cos_minus diff_minus mult_minus_left)
  47.196 +        cos_minus mult_minus_left) simp
  47.197      finally have "(lb_cos prec (- ?lx)) \<le> cos x"
  47.198        unfolding cos_periodic_int . }
  47.199    note negative_lx = this
  47.200 @@ -1236,7 +1155,7 @@
  47.201      have "cos (x + (-k) * (2 * pi)) \<le> cos ?lx"
  47.202        using cos_monotone_0_pi'[OF lx_0 lx pi_x]
  47.203        by (simp only: real_of_int_minus
  47.204 -        cos_minus diff_minus mult_minus_left)
  47.205 +        cos_minus mult_minus_left) simp
  47.206      also have "\<dots> \<le> (ub_cos prec ?lx)"
  47.207        using lb_cos[OF lx_0 pi_lx] by simp
  47.208      finally have "cos x \<le> (ub_cos prec ?lx)"
  47.209 @@ -1251,7 +1170,7 @@
  47.210      have "cos (x + (-k) * (2 * pi)) \<le> cos (real (- ?ux))"
  47.211        using cos_monotone_minus_pi_0'[OF pi_x ux ux_0]
  47.212        by (simp only: uminus_float.rep_eq real_of_int_minus
  47.213 -          cos_minus diff_minus mult_minus_left)
  47.214 +          cos_minus mult_minus_left) simp
  47.215      also have "\<dots> \<le> (ub_cos prec (- ?ux))"
  47.216        using lb_cos_minus[OF pi_ux ux_0, of prec] by simp
  47.217      finally have "cos x \<le> (ub_cos prec (- ?ux))"
  47.218 @@ -1268,7 +1187,7 @@
  47.219      also have "\<dots> \<le> cos (x + (-k) * (2 * pi))"
  47.220        using cos_monotone_0_pi'[OF x_ge_0 ux pi_ux]
  47.221        by (simp only: real_of_int_minus
  47.222 -        cos_minus diff_minus mult_minus_left)
  47.223 +        cos_minus mult_minus_left) simp
  47.224      finally have "(lb_cos prec ?ux) \<le> cos x"
  47.225        unfolding cos_periodic_int . }
  47.226    note positive_ux = this
  47.227 @@ -1343,7 +1262,7 @@
  47.228        also have "\<dots> \<le> cos ((?ux - 2 * ?lpi))"
  47.229          using cos_monotone_minus_pi_0'[OF pi_x x_le_ux ux_0]
  47.230          by (simp only: minus_float.rep_eq real_of_int_minus real_of_one minus_one[symmetric]
  47.231 -            diff_minus mult_minus_left mult_1_left)
  47.232 +            mult_minus_left mult_1_left) simp
  47.233        also have "\<dots> = cos ((- (?ux - 2 * ?lpi)))"
  47.234          unfolding uminus_float.rep_eq cos_minus ..
  47.235        also have "\<dots> \<le> (ub_cos prec (- (?ux - 2 * ?lpi)))"
  47.236 @@ -1387,7 +1306,7 @@
  47.237        also have "\<dots> \<le> cos ((?lx + 2 * ?lpi))"
  47.238          using cos_monotone_0_pi'[OF lx_0 lx_le_x pi_x]
  47.239          by (simp only: minus_float.rep_eq real_of_int_minus real_of_one
  47.240 -          minus_one[symmetric] diff_minus mult_minus_left mult_1_left)
  47.241 +          minus_one[symmetric] mult_minus_left mult_1_left) simp
  47.242        also have "\<dots> \<le> (ub_cos prec (?lx + 2 * ?lpi))"
  47.243          using lb_cos[OF lx_0 pi_lx] by simp
  47.244        finally show ?thesis unfolding u by (simp add: real_of_float_max)
  47.245 @@ -1808,10 +1727,8 @@
  47.246    done
  47.247  
  47.248  lemma Float_pos_eq_mantissa_pos:  "Float m e > 0 \<longleftrightarrow> m > 0"
  47.249 -  apply (auto simp add: zero_less_mult_iff zero_float_def powr_gt_zero[of 2 "exponent x"] dest: less_zeroE)
  47.250    using powr_gt_zero[of 2 "e"]
  47.251 -  apply simp
  47.252 -  done
  47.253 +  by (auto simp add: zero_less_mult_iff zero_float_def simp del: powr_gt_zero dest: less_zeroE)
  47.254  
  47.255  lemma Float_representation_aux:
  47.256    fixes m e
  47.257 @@ -2164,12 +2081,12 @@
  47.258    unfolding divide_inverse interpret_floatarith.simps ..
  47.259  
  47.260  lemma interpret_floatarith_diff: "interpret_floatarith (Add a (Minus b)) vs = (interpret_floatarith a vs) - (interpret_floatarith b vs)"
  47.261 -  unfolding diff_minus interpret_floatarith.simps ..
  47.262 +  unfolding interpret_floatarith.simps by simp
  47.263  
  47.264  lemma interpret_floatarith_sin: "interpret_floatarith (Cos (Add (Mult Pi (Num (Float 1 -1))) (Minus a))) vs =
  47.265    sin (interpret_floatarith a vs)"
  47.266    unfolding sin_cos_eq interpret_floatarith.simps
  47.267 -            interpret_floatarith_divide interpret_floatarith_diff diff_minus
  47.268 +            interpret_floatarith_divide interpret_floatarith_diff
  47.269    by auto
  47.270  
  47.271  lemma interpret_floatarith_tan:
  47.272 @@ -3192,7 +3109,7 @@
  47.273  
  47.274    from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x'
  47.275    have "ly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]"
  47.276 -    by (auto simp add: diff_minus)
  47.277 +    by auto
  47.278    from order_less_le_trans[OF _ this, of 0] `0 < ly`
  47.279    show ?thesis by auto
  47.280  qed
  47.281 @@ -3214,7 +3131,7 @@
  47.282  
  47.283    from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x'
  47.284    have "ly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]"
  47.285 -    by (auto simp add: diff_minus)
  47.286 +    by auto
  47.287    from order_trans[OF _ this, of 0] `0 \<le> ly`
  47.288    show ?thesis by auto
  47.289  qed
    48.1 --- a/src/HOL/Decision_Procs/Cooper.thy	Mon Nov 11 17:34:44 2013 +0100
    48.2 +++ b/src/HOL/Decision_Procs/Cooper.thy	Mon Nov 11 17:44:21 2013 +0100
    48.3 @@ -1400,9 +1400,8 @@
    48.4    also have "\<dots> = (j dvd (- (c*x - ?e)))"
    48.5      by (simp only: dvd_minus_iff)
    48.6    also have "\<dots> = (j dvd (c* (- x)) + ?e)"
    48.7 -    apply (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] diff_minus add_ac minus_add_distrib)
    48.8 -    apply (simp add: algebra_simps)
    48.9 -    done
   48.10 +    by (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] add_ac minus_add_distrib)
   48.11 +      (simp add: algebra_simps)
   48.12    also have "\<dots> = Ifm bbs ((- x)#bs) (Dvd j (CN 0 c e))"
   48.13      using numbound0_I[OF nb, where bs="bs" and b="x" and b'="- x"] by simp
   48.14    finally show ?case .
   48.15 @@ -1413,9 +1412,8 @@
   48.16    also have "\<dots> = (j dvd (- (c*x - ?e)))"
   48.17      by (simp only: dvd_minus_iff)
   48.18    also have "\<dots> = (j dvd (c* (- x)) + ?e)"
   48.19 -    apply (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] diff_minus add_ac minus_add_distrib)
   48.20 -    apply (simp add: algebra_simps)
   48.21 -    done
   48.22 +    by (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] add_ac minus_add_distrib)
   48.23 +      (simp add: algebra_simps)
   48.24    also have "\<dots> = Ifm bbs ((- x)#bs) (Dvd j (CN 0 c e))"
   48.25      using numbound0_I[OF nb, where bs="bs" and b="x" and b'="- x"] by simp
   48.26    finally show ?case by simp
    49.1 --- a/src/HOL/Decision_Procs/MIR.thy	Mon Nov 11 17:34:44 2013 +0100
    49.2 +++ b/src/HOL/Decision_Procs/MIR.thy	Mon Nov 11 17:44:21 2013 +0100
    49.3 @@ -1727,7 +1727,7 @@
    49.4    {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt a))" 
    49.5        by (simp add: nb Let_def split_def isint_Floor isint_neg)
    49.6      have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def)
    49.7 -    also have "\<dots> = (?I (?l (Lt a)))" apply (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) by (simp add: Ia cp cnz Let_def split_def diff_minus)
    49.8 +    also have "\<dots> = (?I (?l (Lt a)))" apply (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) by (simp add: Ia cp cnz Let_def split_def)
    49.9      finally have ?case using l by simp}
   49.10    moreover
   49.11    {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt a))" 
   49.12 @@ -1752,13 +1752,13 @@
   49.13    {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Le a))" 
   49.14        by (simp add: nb Let_def split_def isint_Floor isint_neg)
   49.15      have "?I (Le a) = (real (?c * i) + (?N ?r) \<le> 0)" using Ia by (simp add: Let_def split_def)
   49.16 -    also have "\<dots> = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_minus)
   49.17 +    also have "\<dots> = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def)
   49.18      finally have ?case using l by simp}
   49.19    moreover
   49.20    {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Le a))" 
   49.21        by (simp add: nb Let_def split_def isint_Floor isint_neg)
   49.22      have "?I (Le a) = (real (?c * i) + (?N ?r) \<le> 0)" using Ia by (simp add: Let_def split_def)
   49.23 -    also from cn cnz have "\<dots> = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_minus[symmetric] add_ac ,arith)
   49.24 +    also from cn cnz have "\<dots> = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def add_ac, arith)
   49.25      finally have ?case using l by simp}
   49.26    ultimately show ?case by blast
   49.27  next
   49.28 @@ -1777,13 +1777,13 @@
   49.29    {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Gt a))" 
   49.30        by (simp add: nb Let_def split_def isint_Floor isint_neg)
   49.31      have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def)
   49.32 -    also have "\<dots> = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_minus)
   49.33 +    also have "\<dots> = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def)
   49.34      finally have ?case using l by simp}
   49.35    moreover
   49.36    {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Gt a))" 
   49.37        by (simp add: nb Let_def split_def isint_Floor isint_neg)
   49.38      have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def)
   49.39 -    also from cn cnz have "\<dots> = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_minus[symmetric] add_ac, arith)
   49.40 +    also from cn cnz have "\<dots> = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def add_ac, arith)
   49.41      finally have ?case using l by simp}
   49.42    ultimately show ?case by blast
   49.43  next
   49.44 @@ -1802,13 +1802,13 @@
   49.45    {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Ge a))" 
   49.46        by (simp add: nb Let_def split_def isint_Floor isint_neg)
   49.47      have "?I (Ge a) = (real (?c * i) + (?N ?r) \<ge> 0)" using Ia by (simp add: Let_def split_def)
   49.48 -    also have "\<dots> = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_minus)
   49.49 +    also have "\<dots> = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def)
   49.50      finally have ?case using l by simp}
   49.51    moreover
   49.52    {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Ge a))" 
   49.53        by (simp add: nb Let_def split_def isint_Floor isint_neg)
   49.54      have "?I (Ge a) = (real (?c * i) + (?N ?r) \<ge> 0)" using Ia by (simp add: Let_def split_def)
   49.55 -    also from cn cnz have "\<dots> = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_minus[symmetric] add_ac, arith)
   49.56 +    also from cn cnz have "\<dots> = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def add_ac, arith)
   49.57      finally have ?case using l by simp}
   49.58    ultimately show ?case by blast
   49.59  next
   49.60 @@ -3125,7 +3125,8 @@
   49.61      hence pid: "c*i + ?fe \<le> c*d" by (simp only: real_of_int_le_iff)
   49.62      with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + ?fe = j1" by auto
   49.63      hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - ?N i e + real j1" 
   49.64 -      by (simp only: real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] algebra_simps)
   49.65 +      by (simp only: real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff])
   49.66 +        (simp add: algebra_simps)
   49.67      with nob  have ?case by blast }
   49.68    ultimately show ?case by blast
   49.69  next
   49.70 @@ -3148,11 +3149,12 @@
   49.71      hence pid: "c*i + 1 + ?fe \<le> c*d" by (simp only: real_of_int_le_iff)
   49.72      with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + 1+ ?fe = j1" by auto
   49.73      hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) + 1= - ?N i e + real j1"
   49.74 -      by (simp only: real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] algebra_simps real_of_one) 
   49.75 +      by (simp only: real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] real_of_one) 
   49.76 +        (simp add: algebra_simps)
   49.77      hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = (- ?N i e + real j1) - 1"
   49.78        by (simp only: algebra_simps)
   49.79          hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - 1 - ?N i e + real j1"
   49.80 -          by (simp only: add_ac diff_minus)
   49.81 +          by (simp add: algebra_simps minus_one [symmetric] del: minus_one)
   49.82      with nob  have ?case by blast }
   49.83    ultimately show ?case by blast
   49.84  next
   49.85 @@ -3477,10 +3479,7 @@
   49.86    qed
   49.87  next
   49.88    case (3 a b) then show ?case
   49.89 -    apply auto
   49.90 -    apply (erule_tac x = "(aa, aaa, ba)" in ballE) apply simp_all
   49.91 -    apply (erule_tac x = "(ab, ac, baa)" in ballE) apply simp_all
   49.92 -    done
   49.93 +    by auto
   49.94  qed (auto simp add: Let_def split_def algebra_simps)
   49.95  
   49.96  lemma real_in_int_intervals: 
   49.97 @@ -3615,7 +3614,7 @@
   49.98        by(simp only: myle[of _ "real n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real n *x + ?N s - ?N (Floor s)"]) 
   49.99      hence "\<exists> j\<in> {n .. 0}. 0 \<ge> - (real n *x + ?N s - ?N (Floor s) - real j) \<and> - (real n *x + ?N s - ?N (Floor s) - real (j+1)) > 0" by (simp only: th1[rule_format] th2[rule_format])
  49.100      hence "\<exists> j\<in> {n.. 0}. ?I (?p (p,n,s) j)"
  49.101 -      using pns by (simp add: fp_def nn diff_minus add_ac mult_ac
  49.102 +      using pns by (simp add: fp_def nn algebra_simps
  49.103          del: diff_less_0_iff_less diff_le_0_iff_le) 
  49.104      then obtain "j" where j_def: "j\<in> {n .. 0} \<and> ?I (?p (p,n,s) j)" by blast
  49.105      hence "\<exists>x \<in> {?p (p,n,s) j |j. n\<le> j \<and> j \<le> 0 }. ?I x" by auto
  49.106 @@ -4832,7 +4831,7 @@
  49.107    shows "(Ifm bs (E p)) = (\<exists> (i::int). Ifm (real i#bs) (E (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (exsplit p))))" (is "?lhs = ?rhs")
  49.108  proof-
  49.109    have "?rhs = (\<exists> (i::int). \<exists> x. 0\<le> x \<and> x < 1 \<and> Ifm (x#(real i)#bs) (exsplit p))"
  49.110 -    by (simp add: myless[of _ "1"] myless[of _ "0"] add_ac diff_minus)
  49.111 +    by (simp add: myless[of _ "1"] myless[of _ "0"] add_ac)
  49.112    also have "\<dots> = (\<exists> (i::int). \<exists> x. 0\<le> x \<and> x < 1 \<and> Ifm ((real i + x) #bs) p)"
  49.113      by (simp only: exsplit[OF qf] add_ac)
  49.114    also have "\<dots> = (\<exists> x. Ifm (x#bs) p)" 
  49.115 @@ -5196,7 +5195,7 @@
  49.116    hence "\<forall> j\<in> set ?js. bound0 (subst0 (C j) ?smq)" 
  49.117      by (auto simp only: subst0_bound0[OF qfmq])
  49.118    hence th: "\<forall> j\<in> set ?js. bound0 (simpfm (subst0 (C j) ?smq))"
  49.119 -    by (auto simp add: simpfm_bound0)
  49.120 +    by auto
  49.121    from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp 
  49.122    from Bn jsnb have "\<forall> (b,j) \<in> set ?bjs. numbound0 (Add b (C j))"
  49.123      by simp
    50.1 --- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Mon Nov 11 17:34:44 2013 +0100
    50.2 +++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Mon Nov 11 17:44:21 2013 +0100
    50.3 @@ -1959,7 +1959,7 @@
    50.4        by (simp add: field_simps)
    50.5      have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Eq (CNP 0 a r))" by (simp only: th)
    50.6      also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r = 0" 
    50.7 -      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"])
    50.8 +      by (simp add: r [of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
    50.9      also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) =0 "
   50.10        using c d mult_cancel_left[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp
   50.11      also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r =0" 
   50.12 @@ -2041,7 +2041,7 @@
   50.13        by (simp add: field_simps)
   50.14      have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (NEq (CNP 0 a r))" by (simp only: th)
   50.15      also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r \<noteq> 0" 
   50.16 -      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"])
   50.17 +      by (simp add: r [of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
   50.18      also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) \<noteq> 0 "
   50.19        using c d mult_cancel_left[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp
   50.20      also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r \<noteq> 0" 
   50.21 @@ -2106,7 +2106,7 @@
   50.22        by (simp add: field_simps)
   50.23      have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
   50.24      also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r < 0" 
   50.25 -      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"])
   50.26 +      by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
   50.27      also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) < 0"
   50.28        
   50.29        using dc' dc'' mult_less_cancel_left_disj[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp
   50.30 @@ -2127,7 +2127,7 @@
   50.31        by (simp add: field_simps)
   50.32      have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
   50.33      also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r < 0" 
   50.34 -      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"])
   50.35 +      by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
   50.36  
   50.37      also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) > 0"
   50.38        
   50.39 @@ -2251,7 +2251,7 @@
   50.40        by (simp add: field_simps)
   50.41      have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
   50.42      also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r <= 0" 
   50.43 -      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"])
   50.44 +      by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
   50.45      also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) <= 0"
   50.46        
   50.47        using dc' dc'' mult_le_cancel_left[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp
   50.48 @@ -2272,7 +2272,7 @@
   50.49        by (simp add: field_simps)
   50.50      have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
   50.51      also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r <= 0" 
   50.52 -      by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"])
   50.53 +      by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
   50.54  
   50.55      also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) >= 0"
   50.56        
   50.57 @@ -2356,8 +2356,11 @@
   50.58  
   50.59  lemma msubst_I: assumes lp: "islin p" and nc: "isnpoly c" and nd: "isnpoly d"
   50.60    shows "Ifm vs (x#bs) (msubst p ((c,t),(d,s))) = Ifm vs (((- Itm vs (x#bs) t /  Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /2)#bs) p"
   50.61 -  using lp
   50.62 -by (induct p rule: islin.induct, auto simp add: tmbound0_I[where b="(- (Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>) + - (Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>)) / 2" and b'=x and bs = bs and vs=vs] bound0_I[where b="(- (Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>) + - (Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>)) / 2" and b'=x and bs = bs and vs=vs] msubsteq msubstneq msubstlt[OF nc nd] msubstle[OF nc nd])
   50.63 +  using lp by (induct p rule: islin.induct)
   50.64 +    (auto simp add: tmbound0_I
   50.65 +    [where b = "(- (Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>) - (Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>)) / 2"
   50.66 +      and b' = x and bs = bs and vs = vs]
   50.67 +    msubsteq msubstneq msubstlt [OF nc nd] msubstle [OF nc nd])
   50.68  
   50.69  lemma msubst_nb: assumes lp: "islin p" and t: "tmbound0 t" and s: "tmbound0 s"
   50.70    shows "bound0 (msubst p ((c,t),(d,s)))"
   50.71 @@ -2429,7 +2432,7 @@
   50.72    with evaldjf_bound0[of ?Up "(simpfm o (msubst (simpfm p)))"]
   50.73    have "bound0 (evaldjf (simpfm o (msubst (simpfm p))) ?Up)" by blast
   50.74    with mp_nb pp_nb 
   50.75 -  have th1: "bound0 (disj ?mp (disj ?pp (evaldjf (simpfm o (msubst ?q)) ?Up )))" by (simp add: disj_nb)
   50.76 +  have th1: "bound0 (disj ?mp (disj ?pp (evaldjf (simpfm o (msubst ?q)) ?Up )))" by simp
   50.77    from decr0_qf[OF th1] have thqf: "qfree (ferrack p)" by (simp add: ferrack_def Let_def)
   50.78    have "?lhs \<longleftrightarrow> (\<exists>x. Ifm vs (x#bs) ?q)" by simp
   50.79    also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> (\<exists>(c, t)\<in>set ?U. \<exists>(d, s)\<in>set ?U. ?I (msubst (simpfm p) ((c, t), d, s)))" using fr_eq_msubst[OF lq, of vs bs x] by simp
   50.80 @@ -2612,7 +2615,7 @@
   50.81  lemma msubst2_nb: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" and lp: "islin p" and tnb: "tmbound0 t"
   50.82    shows "bound0 (msubst2 p c t)"
   50.83  using lp tnb
   50.84 -by (simp add: msubst2_def msubstneg_nb msubstpos_nb conj_nb disj_nb lt_nb simpfm_bound0)
   50.85 +by (simp add: msubst2_def msubstneg_nb msubstpos_nb lt_nb simpfm_bound0)
   50.86  
   50.87  lemma mult_minus2_left: "-2 * (x::'a::comm_ring_1) = - (2 * x)"
   50.88    by simp
   50.89 @@ -2666,8 +2669,8 @@
   50.90          using H(3) by (auto simp add: msubst2_def lt[OF stupid(1)]  lt[OF stupid(2)] zero_less_mult_iff mult_less_0_iff)
   50.91        from msubst2[OF lp nn nn'(1), of x bs ] H(3) nn'
   50.92        have "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p"
   50.93 -        apply (simp add: add_divide_distrib mult_minus2_left)
   50.94 -        by (simp add: mult_commute)}
   50.95 +        by (simp add: add_divide_distrib diff_divide_distrib mult_minus2_left mult_commute)
   50.96 +    }
   50.97      moreover
   50.98      {fix c t d s assume H: "(c,t) \<in> set (uset p)" "(d,s) \<in> set (uset p)" 
   50.99        "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p"
  50.100 @@ -2675,7 +2678,9 @@
  50.101        hence nn: "isnpoly (C (-2, 1) *\<^sub>p c*\<^sub>p d)" "\<lparr>(C (-2, 1) *\<^sub>p c*\<^sub>p d)\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
  50.102          using H(3,4) by (simp_all add: polymul_norm n2)
  50.103        from msubst2[OF lp nn, of x bs ] H(3,4,5) 
  50.104 -      have "Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))" apply (simp add: add_divide_distrib mult_minus2_left) by (simp add: mult_commute)}
  50.105 +      have "Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))"
  50.106 +        by (simp add: diff_divide_distrib add_divide_distrib mult_minus2_left mult_commute)
  50.107 +    }
  50.108      ultimately show ?thesis by blast
  50.109    qed
  50.110    from fr_eq2[OF lp, of vs bs x] show ?thesis
    51.1 --- a/src/HOL/Decision_Procs/Polynomial_List.thy	Mon Nov 11 17:34:44 2013 +0100
    51.2 +++ b/src/HOL/Decision_Procs/Polynomial_List.thy	Mon Nov 11 17:44:21 2013 +0100
    51.3 @@ -2,371 +2,379 @@
    51.4      Author:     Amine Chaieb
    51.5  *)
    51.6  
    51.7 -header {* Univariate Polynomials as Lists *}
    51.8 +header {* Univariate Polynomials as lists *}
    51.9  
   51.10  theory Polynomial_List
   51.11 -imports Main
   51.12 +imports Complex_Main
   51.13  begin
   51.14  
   51.15 -text{* Application of polynomial as a real function. *}
   51.16 +text{* Application of polynomial as a function. *}
   51.17  
   51.18 -primrec poly :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a::comm_ring"
   51.19 +primrec (in semiring_0) poly :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a"
   51.20  where
   51.21    poly_Nil:  "poly [] x = 0"
   51.22 -| poly_Cons: "poly (h # t) x = h + x * poly t x"
   51.23 +| poly_Cons: "poly (h#t) x = h + x * poly t x"
   51.24  
   51.25  
   51.26  subsection{*Arithmetic Operations on Polynomials*}
   51.27  
   51.28  text{*addition*}
   51.29 -primrec padd :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a::comm_ring_1 list"  (infixl "+++" 65)
   51.30 +
   51.31 +primrec (in semiring_0) padd :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "+++" 65)
   51.32  where
   51.33    padd_Nil:  "[] +++ l2 = l2"
   51.34 -| padd_Cons: "(h # t) +++ l2 = (if l2 = [] then h # t else (h + hd l2) # (t +++ tl l2))"
   51.35 +| padd_Cons: "(h#t) +++ l2 = (if l2 = [] then h#t else (h + hd l2)#(t +++ tl l2))"
   51.36  
   51.37  text{*Multiplication by a constant*}
   51.38 -primrec cmult :: "'a::comm_ring_1 \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "%*" 70)
   51.39 -where
   51.40 +primrec (in semiring_0) cmult :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "%*" 70) where
   51.41    cmult_Nil:  "c %* [] = []"
   51.42 -| cmult_Cons: "c %* (h # t) = (c * h) # (c %* t)"
   51.43 +| cmult_Cons: "c %* (h#t) = (c * h)#(c %* t)"
   51.44  
   51.45  text{*Multiplication by a polynomial*}
   51.46 -primrec pmult :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a::comm_ring_1 list"  (infixl "***" 70)
   51.47 +primrec (in semiring_0) pmult :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "***" 70)
   51.48  where
   51.49    pmult_Nil:  "[] *** l2 = []"
   51.50 -| pmult_Cons: "(h # t) *** l2 =
   51.51 -    (if t = [] then h %* l2 else (h %* l2) +++ (0 # (t *** l2)))"
   51.52 +| pmult_Cons: "(h#t) *** l2 = (if t = [] then h %* l2
   51.53 +                              else (h %* l2) +++ ((0) # (t *** l2)))"
   51.54  
   51.55  text{*Repeated multiplication by a polynomial*}
   51.56 -primrec mulexp :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a::comm_ring_1 list"
   51.57 -where
   51.58 +primrec (in semiring_0) mulexp :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a  list \<Rightarrow> 'a list" where
   51.59    mulexp_zero:  "mulexp 0 p q = q"
   51.60  | mulexp_Suc:   "mulexp (Suc n) p q = p *** mulexp n p q"
   51.61  
   51.62  text{*Exponential*}
   51.63 -primrec pexp :: "'a list \<Rightarrow> nat \<Rightarrow> 'a::comm_ring_1 list"  (infixl "%^" 80)
   51.64 -where
   51.65 +primrec (in semiring_1) pexp :: "'a list \<Rightarrow> nat \<Rightarrow> 'a list"  (infixl "%^" 80) where
   51.66    pexp_0:   "p %^ 0 = [1]"
   51.67  | pexp_Suc: "p %^ (Suc n) = p *** (p %^ n)"
   51.68  
   51.69  text{*Quotient related value of dividing a polynomial by x + a*}
   51.70  (* Useful for divisor properties in inductive proofs *)
   51.71 -primrec pquot :: "'a list \<Rightarrow> 'a::field \<Rightarrow> 'a list"
   51.72 +primrec (in field) "pquot" :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list"
   51.73  where
   51.74 -  pquot_Nil: "pquot [] a = []"
   51.75 -| pquot_Cons: "pquot (h # t) a =
   51.76 -    (if t = [] then [h] else (inverse a * (h - hd (pquot t a))) # pquot t a)"
   51.77 -
   51.78 +  pquot_Nil:  "pquot [] a= []"
   51.79 +| pquot_Cons: "pquot (h#t) a =
   51.80 +    (if t = [] then [h] else (inverse(a) * (h - hd( pquot t a)))#(pquot t a))"
   51.81  
   51.82  text{*normalization of polynomials (remove extra 0 coeff)*}
   51.83 -primrec pnormalize :: "'a::comm_ring_1 list \<Rightarrow> 'a list"
   51.84 -where
   51.85 +primrec (in semiring_0) pnormalize :: "'a list \<Rightarrow> 'a list" where
   51.86    pnormalize_Nil:  "pnormalize [] = []"
   51.87 -| pnormalize_Cons: "pnormalize (h # p) =
   51.88 -    (if (pnormalize p = []) then (if h = 0 then [] else [h])
   51.89 -     else (h # pnormalize p))"
   51.90 +| pnormalize_Cons: "pnormalize (h#p) =
   51.91 +    (if pnormalize p = [] then (if h = 0 then [] else [h]) else h # pnormalize p)"
   51.92  
   51.93 -definition "pnormal p = ((pnormalize p = p) \<and> p \<noteq> [])"
   51.94 -definition "nonconstant p = (pnormal p \<and> (\<forall>x. p \<noteq> [x]))"
   51.95 +definition (in semiring_0) "pnormal p = ((pnormalize p = p) \<and> p \<noteq> [])"
   51.96 +definition (in semiring_0) "nonconstant p = (pnormal p \<and> (\<forall>x. p \<noteq> [x]))"
   51.97  text{*Other definitions*}
   51.98  
   51.99 -definition poly_minus :: "'a list => ('a :: comm_ring_1) list"  ("-- _" [80] 80)
  51.100 +definition (in ring_1) poly_minus :: "'a list \<Rightarrow> 'a list" ("-- _" [80] 80)
  51.101    where "-- p = (- 1) %* p"
  51.102  
  51.103 -definition divides :: "'a::comm_ring_1 list \<Rightarrow> 'a list \<Rightarrow> bool"  (infixl "divides" 70)
  51.104 -  where "p1 divides p2 = (\<exists>q. poly p2 = poly (p1 *** q))"
  51.105 +definition (in semiring_0) divides :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"  (infixl "divides" 70)
  51.106 +  where "p1 divides p2 = (\<exists>q. poly p2 = poly(p1 *** q))"
  51.107  
  51.108 -definition order :: "'a::comm_ring_1 \<Rightarrow> 'a list \<Rightarrow> nat" --{*order of a polynomial*}
  51.109 -  where "order a p = (SOME n. ([-a, 1] %^ n) divides p & ~ (([-a, 1] %^ Suc n) divides p))"
  51.110 +lemma (in semiring_0) dividesI:
  51.111 +  "poly p2 = poly (p1 *** q) \<Longrightarrow> p1 divides p2"
  51.112 +  by (auto simp add: divides_def)
  51.113  
  51.114 -definition degree :: "'a::comm_ring_1 list \<Rightarrow> nat" --{*degree of a polynomial*}
  51.115 +lemma (in semiring_0) dividesE:
  51.116 +  assumes "p1 divides p2"
  51.117 +  obtains q where "poly p2 = poly (p1 *** q)"
  51.118 +  using assms by (auto simp add: divides_def)
  51.119 +
  51.120 +    --{*order of a polynomial*}
  51.121 +definition (in ring_1) order :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" where
  51.122 +  "order a p = (SOME n. ([-a, 1] %^ n) divides p \<and> ~ (([-a, 1] %^ (Suc n)) divides p))"
  51.123 +
  51.124 +     --{*degree of a polynomial*}
  51.125 +definition (in semiring_0) degree :: "'a list \<Rightarrow> nat"
  51.126    where "degree p = length (pnormalize p) - 1"
  51.127  
  51.128 -definition rsquarefree :: "'a::comm_ring_1 list \<Rightarrow> bool"
  51.129 -  where --{*squarefree polynomials --- NB with respect to real roots only.*}
  51.130 -  "rsquarefree p = (poly p \<noteq> poly [] \<and> (\<forall>a. order a p = 0 \<or> order a p = 1))"
  51.131 +     --{*squarefree polynomials --- NB with respect to real roots only.*}
  51.132 +definition (in ring_1) rsquarefree :: "'a list \<Rightarrow> bool"
  51.133 +  where "rsquarefree p \<longleftrightarrow> poly p \<noteq> poly [] \<and> (\<forall>a. order a p = 0 \<or> order a p = 1)"
  51.134  
  51.135 -lemma padd_Nil2 [simp]: "p +++ [] = p"
  51.136 +context semiring_0
  51.137 +begin
  51.138 +
  51.139 +lemma padd_Nil2[simp]: "p +++ [] = p"
  51.140    by (induct p) auto
  51.141  
  51.142  lemma padd_Cons_Cons: "(h1 # p1) +++ (h2 # p2) = (h1 + h2) # (p1 +++ p2)"
  51.143    by auto
  51.144  
  51.145 -lemma pminus_Nil [simp]: "-- [] = []"
  51.146 +lemma pminus_Nil: "-- [] = []"
  51.147    by (simp add: poly_minus_def)
  51.148  
  51.149 -lemma pmult_singleton: "[h1] *** p1 = h1 %* p1"
  51.150 -  by simp
  51.151 +lemma pmult_singleton: "[h1] *** p1 = h1 %* p1" by simp
  51.152  
  51.153 -lemma poly_ident_mult [simp]: "1 %* t = t"
  51.154 -  by (induct t) auto
  51.155 +end
  51.156  
  51.157 -lemma poly_simple_add_Cons [simp]: "[a] +++ ((0)#t) = (a#t)"
  51.158 +lemma (in semiring_1) poly_ident_mult[simp]: "1 %* t = t" by (induct t) auto
  51.159 +
  51.160 +lemma (in semiring_0) poly_simple_add_Cons[simp]: "[a] +++ ((0)#t) = (a#t)"
  51.161    by simp
  51.162  
  51.163  text{*Handy general properties*}
  51.164  
  51.165 -lemma padd_commut: "b +++ a = a +++ b"
  51.166 -  apply (induct b arbitrary: a)
  51.167 -  apply auto
  51.168 -  apply (rule padd_Cons [THEN ssubst])
  51.169 -  apply (case_tac aa)
  51.170 -  apply auto
  51.171 +lemma (in comm_semiring_0) padd_commut: "b +++ a = a +++ b"
  51.172 +proof (induct b arbitrary: a)
  51.173 +  case Nil
  51.174 +  thus ?case by auto
  51.175 +next
  51.176 +  case (Cons b bs a)
  51.177 +  thus ?case by (cases a) (simp_all add: add_commute)
  51.178 +qed
  51.179 +
  51.180 +lemma (in comm_semiring_0) padd_assoc: "\<forall>b c. (a +++ b) +++ c = a +++ (b +++ c)"
  51.181 +  apply (induct a)
  51.182 +  apply (simp, clarify)
  51.183 +  apply (case_tac b, simp_all add: add_ac)
  51.184    done
  51.185  
  51.186 -lemma padd_assoc: "(a +++ b) +++ c = a +++ (b +++ c)"
  51.187 -  apply (induct a arbitrary: b c)
  51.188 +lemma (in semiring_0) poly_cmult_distr: "a %* ( p +++ q) = (a %* p +++ a %* q)"
  51.189 +  apply (induct p arbitrary: q)
  51.190    apply simp
  51.191 -  apply (case_tac b)
  51.192 -  apply simp_all
  51.193 +  apply (case_tac q, simp_all add: distrib_left)
  51.194    done
  51.195  
  51.196 -lemma poly_cmult_distr: "a %* ( p +++ q) = (a %* p +++ a %* q)"
  51.197 -  apply (induct p arbitrary: q)
  51.198 +lemma (in ring_1) pmult_by_x[simp]: "[0, 1] *** t = ((0)#t)"
  51.199 +  apply (induct t)
  51.200    apply simp
  51.201 -  apply (case_tac q)
  51.202 -  apply (simp_all add: distrib_left)
  51.203 +  apply (auto simp add: padd_commut)
  51.204 +  apply (case_tac t, auto)
  51.205    done
  51.206  
  51.207 -lemma pmult_by_x [simp]: "[0, 1] *** t = ((0)#t)"
  51.208 -  by (induct t) (auto simp add: padd_commut)
  51.209 -
  51.210 -
  51.211  text{*properties of evaluation of polynomials.*}
  51.212  
  51.213 -lemma poly_add: "poly (p1 +++ p2) x = poly p1 x + poly p2 x"
  51.214 -  apply (induct p1 arbitrary: p2)
  51.215 -  apply auto
  51.216 -  apply (case_tac "p2")
  51.217 -  apply (auto simp add: distrib_left)
  51.218 -  done
  51.219 +lemma (in semiring_0) poly_add: "poly (p1 +++ p2) x = poly p1 x + poly p2 x"
  51.220 +proof(induct p1 arbitrary: p2)
  51.221 +  case Nil
  51.222 +  thus ?case by simp
  51.223 +next
  51.224 +  case (Cons a as p2)
  51.225 +  thus ?case
  51.226 +    by (cases p2) (simp_all  add: add_ac distrib_left)
  51.227 +qed
  51.228  
  51.229 -lemma poly_cmult: "poly (c %* p) x = c * poly p x"
  51.230 +lemma (in comm_semiring_0) poly_cmult: "poly (c %* p) x = c * poly p x"
  51.231    apply (induct p)
  51.232 -  apply simp
  51.233 -  apply (cases "x = 0")
  51.234 +  apply (case_tac [2] "x = zero")
  51.235    apply (auto simp add: distrib_left mult_ac)
  51.236    done
  51.237  
  51.238 -lemma poly_minus: "poly (-- p) x = - (poly p x)"
  51.239 -  by (simp add: poly_minus_def poly_cmult)
  51.240 +lemma (in comm_semiring_0) poly_cmult_map: "poly (map (op * c) p) x = c*poly p x"
  51.241 +  by (induct p) (auto simp add: distrib_left mult_ac)
  51.242  
  51.243 -lemma poly_mult: "poly (p1 *** p2) x = poly p1 x * poly p2 x"
  51.244 -  apply (induct p1 arbitrary: p2)
  51.245 -  apply (case_tac p1)
  51.246 -  apply (auto simp add: poly_cmult poly_add distrib_right distrib_left mult_ac)
  51.247 +lemma (in comm_ring_1) poly_minus: "poly (-- p) x = - (poly p x)"
  51.248 +  apply (simp add: poly_minus_def)
  51.249 +  apply (auto simp add: poly_cmult)
  51.250    done
  51.251  
  51.252 -lemma poly_exp: "poly (p %^ n) (x::'a::comm_ring_1) = (poly p x) ^ n"
  51.253 +lemma (in comm_semiring_0) poly_mult: "poly (p1 *** p2) x = poly p1 x * poly p2 x"
  51.254 +proof (induct p1 arbitrary: p2)
  51.255 +  case Nil
  51.256 +  thus ?case by simp
  51.257 +next
  51.258 +  case (Cons a as p2)
  51.259 +  thus ?case by (cases as)
  51.260 +    (simp_all add: poly_cmult poly_add distrib_right distrib_left mult_ac)
  51.261 +qed
  51.262 +
  51.263 +class idom_char_0 = idom + ring_char_0
  51.264 +
  51.265 +subclass (in field_char_0) idom_char_0 ..
  51.266 +
  51.267 +lemma (in comm_ring_1) poly_exp: "poly (p %^ n) x = (poly p x) ^ n"
  51.268    by (induct n) (auto simp add: poly_cmult poly_mult)
  51.269  
  51.270  text{*More Polynomial Evaluation Lemmas*}
  51.271  
  51.272 -lemma poly_add_rzero [simp]: "poly (a +++ []) x = poly a x"
  51.273 +lemma (in semiring_0) poly_add_rzero[simp]: "poly (a +++ []) x = poly a x"
  51.274    by simp
  51.275  
  51.276 -lemma poly_mult_assoc: "poly ((a *** b) *** c) x = poly (a *** (b *** c)) x"
  51.277 +lemma (in comm_semiring_0) poly_mult_assoc: "poly ((a *** b) *** c) x = poly (a *** (b *** c)) x"
  51.278    by (simp add: poly_mult mult_assoc)
  51.279  
  51.280 -lemma poly_mult_Nil2 [simp]: "poly (p *** []) x = 0"
  51.281 +lemma (in semiring_0) poly_mult_Nil2[simp]: "poly (p *** []) x = 0"
  51.282    by (induct p) auto
  51.283  
  51.284 -lemma poly_exp_add: "poly (p %^ (n + d)) x = poly( p %^ n *** p %^ d) x"
  51.285 +lemma (in comm_semiring_1) poly_exp_add: "poly (p %^ (n + d)) x = poly( p %^ n *** p %^ d) x"
  51.286    by (induct n) (auto simp add: poly_mult mult_assoc)
  51.287  
  51.288  subsection{*Key Property: if @{term "f(a) = 0"} then @{term "(x - a)"} divides
  51.289   @{term "p(x)"} *}
  51.290  
  51.291 -lemma poly_linear_rem: "\<exists>q r. h # t = [r] +++ [-a, 1] *** q"
  51.292 -  apply (induct t arbitrary: h)
  51.293 -  apply (rule_tac x = "[]" in exI)
  51.294 -  apply (rule_tac x = h in exI)
  51.295 -  apply simp
  51.296 -  apply (drule_tac x = aa in meta_spec)
  51.297 -  apply safe
  51.298 -  apply (rule_tac x = "r#q" in exI)
  51.299 -  apply (rule_tac x = "a*r + h" in exI)
  51.300 -  apply (case_tac q)
  51.301 -  apply auto
  51.302 -  done
  51.303 +lemma (in comm_ring_1) lemma_poly_linear_rem: "\<forall>h. \<exists>q r. h#t = [r] +++ [-a, 1] *** q"
  51.304 +proof(induct t)
  51.305 +  case Nil
  51.306 +  { fix h have "[h] = [h] +++ [- a, 1] *** []" by simp }
  51.307 +  thus ?case by blast
  51.308 +next
  51.309 +  case (Cons  x xs)
  51.310 +  { fix h
  51.311 +    from Cons.hyps[rule_format, of x]
  51.312 +    obtain q r where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast
  51.313 +    have "h#x#xs = [a*r + h] +++ [-a, 1] *** (r#q)"
  51.314 +      using qr by (cases q) (simp_all add: algebra_simps)
  51.315 +    hence "\<exists>q r. h#x#xs = [r] +++ [-a, 1] *** q" by blast}
  51.316 +  thus ?case by blast
  51.317 +qed
  51.318  
  51.319 -lemma poly_linear_divides: "poly p a = 0 \<longleftrightarrow> p = [] \<or> (\<exists>q. p = [-a, 1] *** q)"
  51.320 -  apply (auto simp add: poly_add poly_cmult distrib_left)
  51.321 -  apply (case_tac p)
  51.322 -  apply simp
  51.323 -  apply (cut_tac h = aa and t = list and a = a in poly_linear_rem)
  51.324 -  apply safe
  51.325 -  apply (case_tac q)
  51.326 -  apply auto
  51.327 -  apply (drule_tac x = "[]" in spec)
  51.328 -  apply simp
  51.329 -  apply (auto simp add: poly_add poly_cmult add_assoc)
  51.330 -  apply (drule_tac x = "aa#lista" in spec)
  51.331 -  apply auto
  51.332 -  done
  51.333 +lemma (in comm_ring_1) poly_linear_rem: "\<exists>q r. h#t = [r] +++ [-a, 1] *** q"
  51.334 +  using lemma_poly_linear_rem [where t = t and a = a] by auto
  51.335  
  51.336 -lemma lemma_poly_length_mult [simp]: "length (k %* p +++  (h # (a %* p))) = Suc (length p)"
  51.337 -  by (induct p arbitrary: h k a) auto
  51.338  
  51.339 -lemma lemma_poly_length_mult2 [simp]: "length (k %* p +++  (h # p)) = Suc (length p)"
  51.340 -  by (induct p arbitrary: h k) auto
  51.341 +lemma (in comm_ring_1) poly_linear_divides: "(poly p a = 0) = ((p = []) | (\<exists>q. p = [-a, 1] *** q))"
  51.342 +proof -
  51.343 +  { assume p: "p = []" hence ?thesis by simp }
  51.344 +  moreover
  51.345 +  {
  51.346 +    fix x xs assume p: "p = x#xs"
  51.347 +    {
  51.348 +      fix q assume "p = [-a, 1] *** q"
  51.349 +      hence "poly p a = 0" by (simp add: poly_add poly_cmult)
  51.350 +    }
  51.351 +    moreover
  51.352 +    { assume p0: "poly p a = 0"
  51.353 +      from poly_linear_rem[of x xs a] obtain q r
  51.354 +      where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast
  51.355 +      have "r = 0" using p0 by (simp only: p qr poly_mult poly_add) simp
  51.356 +      hence "\<exists>q. p = [- a, 1] *** q"
  51.357 +        using p qr
  51.358 +        apply -
  51.359 +        apply (rule exI[where x=q])
  51.360 +        apply auto
  51.361 +        apply (cases q)
  51.362 +        apply auto
  51.363 +        done
  51.364 +    }
  51.365 +    ultimately have ?thesis using p by blast
  51.366 +  }
  51.367 +  ultimately show ?thesis by (cases p) auto
  51.368 +qed
  51.369  
  51.370 -lemma poly_length_mult [simp]: "length([-a, 1] *** q) = Suc (length q)"
  51.371 +lemma (in semiring_0) lemma_poly_length_mult[simp]: "\<forall>h k a. length (k %* p +++  (h # (a %* p))) = Suc (length p)"
  51.372 +  by (induct p) auto
  51.373 +
  51.374 +lemma (in semiring_0) lemma_poly_length_mult2[simp]: "\<forall>h k. length (k %* p +++  (h # p)) = Suc (length p)"
  51.375 +  by (induct p) auto
  51.376 +
  51.377 +lemma (in ring_1) poly_length_mult[simp]: "length([-a,1] *** q) = Suc (length q)"
  51.378    by auto
  51.379  
  51.380 -
  51.381  subsection{*Polynomial length*}
  51.382  
  51.383 -lemma poly_cmult_length [simp]: "length (a %* p) = length p"
  51.384 +lemma (in semiring_0) poly_cmult_length[simp]: "length (a %* p) = length p"
  51.385    by (induct p) auto
  51.386  
  51.387 -lemma poly_add_length:
  51.388 -  "length (p1 +++ p2) = (if (length p1 < length p2) then length p2 else length p1)"
  51.389 -  by (induct p1 arbitrary: p2) auto
  51.390 +lemma (in semiring_0) poly_add_length: "length (p1 +++ p2) = max (length p1) (length p2)"
  51.391 +  by (induct p1 arbitrary: p2) (simp_all, arith)
  51.392  
  51.393 -lemma poly_root_mult_length [simp]: "length ([a, b] *** p) = Suc (length p)"
  51.394 -  by simp
  51.395 +lemma (in semiring_0) poly_root_mult_length[simp]: "length([a,b] *** p) = Suc (length p)"
  51.396 +  by (simp add: poly_add_length)
  51.397  
  51.398 -lemma poly_mult_not_eq_poly_Nil [simp]:
  51.399 -  "poly (p *** q) x \<noteq> poly [] x \<longleftrightarrow> poly p x \<noteq> poly [] x \<and> poly q x \<noteq> poly [] (x::'a::idom)"
  51.400 +lemma (in idom) poly_mult_not_eq_poly_Nil[simp]:
  51.401 +  "poly (p *** q) x \<noteq> poly [] x \<longleftrightarrow> poly p x \<noteq> poly [] x \<and> poly q x \<noteq> poly [] x"
  51.402    by (auto simp add: poly_mult)
  51.403  
  51.404 -lemma poly_mult_eq_zero_disj: "poly (p *** q) (x::'a::idom) = 0 \<longleftrightarrow> poly p x = 0 \<or> poly q x = 0"
  51.405 +lemma (in idom) poly_mult_eq_zero_disj: "poly (p *** q) x = 0 \<longleftrightarrow> poly p x = 0 \<or> poly q x = 0"
  51.406    by (auto simp add: poly_mult)
  51.407  
  51.408  text{*Normalisation Properties*}
  51.409  
  51.410 -lemma poly_normalized_nil: "pnormalize p = [] \<longrightarrow> poly p x = 0"
  51.411 +lemma (in semiring_0) poly_normalized_nil: "(pnormalize p = []) --> (poly p x = 0)"
  51.412    by (induct p) auto
  51.413  
  51.414  text{*A nontrivial polynomial of degree n has no more than n roots*}
  51.415 +lemma (in idom) poly_roots_index_lemma:
  51.416 +   assumes p: "poly p x \<noteq> poly [] x" and n: "length p = n"
  51.417 +  shows "\<exists>i. \<forall>x. poly p x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)"
  51.418 +  using p n
  51.419 +proof (induct n arbitrary: p x)
  51.420 +  case 0
  51.421 +  thus ?case by simp
  51.422 +next
  51.423 +  case (Suc n p x)
  51.424 +  {
  51.425 +    assume C: "\<And>i. \<exists>x. poly p x = 0 \<and> (\<forall>m\<le>Suc n. x \<noteq> i m)"
  51.426 +    from Suc.prems have p0: "poly p x \<noteq> 0" "p\<noteq> []" by auto
  51.427 +    from p0(1)[unfolded poly_linear_divides[of p x]]
  51.428 +    have "\<forall>q. p \<noteq> [- x, 1] *** q" by blast
  51.429 +    from C obtain a where a: "poly p a = 0" by blast
  51.430 +    from a[unfolded poly_linear_divides[of p a]] p0(2)
  51.431 +    obtain q where q: "p = [-a, 1] *** q" by blast
  51.432 +    have lg: "length q = n" using q Suc.prems(2) by simp
  51.433 +    from q p0 have qx: "poly q x \<noteq> poly [] x"
  51.434 +      by (auto simp add: poly_mult poly_add poly_cmult)
  51.435 +    from Suc.hyps[OF qx lg] obtain i where
  51.436 +      i: "\<forall>x. poly q x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)" by blast
  51.437 +    let ?i = "\<lambda>m. if m = Suc n then a else i m"
  51.438 +    from C[of ?i] obtain y where y: "poly p y = 0" "\<forall>m\<le> Suc n. y \<noteq> ?i m"
  51.439 +      by blast
  51.440 +    from y have "y = a \<or> poly q y = 0"
  51.441 +      by (simp only: q poly_mult_eq_zero_disj poly_add) (simp add: algebra_simps)
  51.442 +    with i[rule_format, of y] y(1) y(2) have False
  51.443 +      apply auto
  51.444 +      apply (erule_tac x = "m" in allE)
  51.445 +      apply auto
  51.446 +      done
  51.447 +  }
  51.448 +  thus ?case by blast
  51.449 +qed
  51.450  
  51.451 -lemma poly_roots_index_lemma0 [rule_format]:
  51.452 -   "\<forall>p x. poly p x \<noteq> poly [] x & length p = n
  51.453 -    --> (\<exists>i. \<forall>x. (poly p x = (0::'a::idom)) --> (\<exists>m. (m \<le> n & x = i m)))"
  51.454 -  apply (induct n)
  51.455 -  apply safe
  51.456 -  apply (rule ccontr)
  51.457 -  apply (subgoal_tac "\<exists>a. poly p a = 0")
  51.458 -  apply safe
  51.459 -  apply (drule poly_linear_divides [THEN iffD1])
  51.460 -  apply safe
  51.461 -  apply (drule_tac x = q in spec)
  51.462 -  apply (drule_tac x = x in spec)
  51.463 -  apply (simp del: poly_Nil pmult_Cons)
  51.464 -  apply (erule exE)
  51.465 -  apply (drule_tac x = "%m. if m = Suc n then a else i m" in spec)
  51.466 -  apply safe
  51.467 -  apply (drule poly_mult_eq_zero_disj [THEN iffD1])
  51.468 -  apply safe
  51.469 -  apply (drule_tac x = "Suc (length q)" in spec)
  51.470 -  apply (auto simp add: field_simps)
  51.471 -  apply (drule_tac x = xa in spec)
  51.472 -  apply (clarsimp simp add: field_simps)
  51.473 -  apply (drule_tac x = m in spec)
  51.474 -  apply (auto simp add:field_simps)
  51.475 -  done
  51.476 -lemmas poly_roots_index_lemma1 = conjI [THEN poly_roots_index_lemma0]
  51.477  
  51.478 -lemma poly_roots_index_length0:
  51.479 -  "poly p (x::'a::idom) \<noteq> poly [] x \<Longrightarrow>
  51.480 -    \<exists>i. \<forall>x. (poly p x = 0) \<longrightarrow> (\<exists>n. n \<le> length p & x = i n)"
  51.481 -  by (blast intro: poly_roots_index_lemma1)
  51.482 +lemma (in idom) poly_roots_index_length:
  51.483 +  "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>i. \<forall>x. (poly p x = 0) \<longrightarrow> (\<exists>n. n \<le> length p \<and> x = i n)"
  51.484 +  by (blast intro: poly_roots_index_lemma)
  51.485  
  51.486 -lemma poly_roots_finite_lemma:
  51.487 -  "poly p (x::'a::idom) \<noteq> poly [] x \<Longrightarrow>
  51.488 -    \<exists>N i. \<forall>x. (poly p x = 0) \<longrightarrow> (\<exists>n. (n::nat) < N & x = i n)"
  51.489 -  apply (drule poly_roots_index_length0)
  51.490 -  apply safe
  51.491 +lemma (in idom) poly_roots_finite_lemma1:
  51.492 +  "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>N i. \<forall>x. (poly p x = 0) \<longrightarrow> (\<exists>n. (n::nat) < N \<and> x = i n)"
  51.493 +  apply (drule poly_roots_index_length, safe)
  51.494    apply (rule_tac x = "Suc (length p)" in exI)
  51.495    apply (rule_tac x = i in exI)
  51.496    apply (simp add: less_Suc_eq_le)
  51.497    done
  51.498  
  51.499 -
  51.500 -lemma real_finite_lemma:
  51.501 -  assumes "\<forall>x. P x \<longrightarrow> (\<exists>n. n < length j & x = j!n)"
  51.502 -  shows "finite {(x::'a::idom). P x}"
  51.503 +lemma (in idom) idom_finite_lemma:
  51.504 +  assumes P: "\<forall>x. P x --> (\<exists>n. n < length j \<and> x = j!n)"
  51.505 +  shows "finite {x. P x}"
  51.506  proof -
  51.507    let ?M = "{x. P x}"
  51.508    let ?N = "set j"
  51.509 -  have "?M \<subseteq> ?N" using assms by auto
  51.510 -  then show ?thesis using finite_subset by auto
  51.511 +  have "?M \<subseteq> ?N" using P by auto
  51.512 +  thus ?thesis using finite_subset by auto
  51.513  qed
  51.514  
  51.515 -lemma poly_roots_index_lemma [rule_format]:
  51.516 -  "\<forall>p x. poly p x \<noteq> poly [] x & length p = n
  51.517 -    \<longrightarrow> (\<exists>i. \<forall>x. (poly p x = (0::'a::{idom})) \<longrightarrow> x \<in> set i)"
  51.518 -  apply (induct n)
  51.519 -  apply safe
  51.520 -  apply (rule ccontr)
  51.521 -  apply (subgoal_tac "\<exists>a. poly p a = 0")
  51.522 -  apply safe
  51.523 -  apply (drule poly_linear_divides [THEN iffD1])
  51.524 -  apply safe
  51.525 -  apply (drule_tac x = q in spec)
  51.526 -  apply (drule_tac x = x in spec)
  51.527 -  apply (auto simp del: poly_Nil pmult_Cons)
  51.528 -  apply (drule_tac x = "a#i" in spec)
  51.529 -  apply (auto simp only: poly_mult List.list.size)
  51.530 -  apply (drule_tac x = xa in spec)
  51.531 -  apply (clarsimp simp add: field_simps)
  51.532 +lemma (in idom) poly_roots_finite_lemma2:
  51.533 +  "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>i. \<forall>x. poly p x = 0 \<longrightarrow> x \<in> set i"
  51.534 +  apply (drule poly_roots_index_length, safe)
  51.535 +  apply (rule_tac x="map (\<lambda>n. i n) [0 ..< Suc (length p)]" in exI)
  51.536 +  apply (auto simp add: image_iff)
  51.537 +  apply (erule_tac x="x" in allE, clarsimp)
  51.538 +  apply (case_tac "n = length p")
  51.539 +  apply (auto simp add: order_le_less)
  51.540    done
  51.541  
  51.542 -lemmas poly_roots_index_lemma2 = conjI [THEN poly_roots_index_lemma]
  51.543 -
  51.544 -lemma poly_roots_index_length:
  51.545 -  "poly p (x::'a::idom) \<noteq> poly [] x \<Longrightarrow>
  51.546 -    \<exists>i. \<forall>x. (poly p x = 0) --> x \<in> set i"
  51.547 -  by (blast intro: poly_roots_index_lemma2)
  51.548 -
  51.549 -lemma poly_roots_finite_lemma':
  51.550 -  "poly p (x::'a::idom) \<noteq> poly [] x \<Longrightarrow>
  51.551 -    \<exists>i. \<forall>x. (poly p x = 0) --> x \<in> set i"
  51.552 -  apply (drule poly_roots_index_length)
  51.553 -  apply auto
  51.554 -  done
  51.555 -
  51.556 -lemma UNIV_nat_infinite: "\<not> finite (UNIV :: nat set)"
  51.557 -  unfolding finite_conv_nat_seg_image
  51.558 -proof (auto simp add: set_eq_iff image_iff)
  51.559 -  fix n::nat and f:: "nat \<Rightarrow> nat"
  51.560 -  let ?N = "{i. i < n}"
  51.561 -  let ?fN = "f ` ?N"
  51.562 -  let ?y = "Max ?fN + 1"
  51.563 -  from nat_seg_image_imp_finite[of "?fN" "f" n]
  51.564 -  have thfN: "finite ?fN" by simp
  51.565 -  { assume "n =0" hence "\<exists>x. \<forall>xa<n. x \<noteq> f xa" by auto }
  51.566 -  moreover
  51.567 -  { assume nz: "n \<noteq> 0"
  51.568 -    hence thne: "?fN \<noteq> {}" by auto
  51.569 -    have "\<forall>x\<in> ?fN. Max ?fN \<ge> x" using nz Max_ge_iff[OF thfN thne] by auto
  51.570 -    hence "\<forall>x\<in> ?fN. ?y > x" by (auto simp add: less_Suc_eq_le)
  51.571 -    hence "?y \<notin> ?fN" by auto
  51.572 -    hence "\<exists>x. \<forall>xa<n. x \<noteq> f xa" by auto }
  51.573 -  ultimately show "\<exists>x. \<forall>xa<n. x \<noteq> f xa" by blast
  51.574 +lemma (in ring_char_0) UNIV_ring_char_0_infinte: "\<not> (finite (UNIV:: 'a set))"
  51.575 +proof
  51.576 +  assume F: "finite (UNIV :: 'a set)"
  51.577 +  have "finite (UNIV :: nat set)"
  51.578 +  proof (rule finite_imageD)
  51.579 +    have "of_nat ` UNIV \<subseteq> UNIV" by simp
  51.580 +    then show "finite (of_nat ` UNIV :: 'a set)" using F by (rule finite_subset)
  51.581 +    show "inj (of_nat :: nat \<Rightarrow> 'a)" by (simp add: inj_on_def)
  51.582 +  qed
  51.583 +  with infinite_UNIV_nat show False ..
  51.584  qed
  51.585  
  51.586 -lemma UNIV_ring_char_0_infinte: "\<not> finite (UNIV:: ('a::ring_char_0) set)"
  51.587 +lemma (in idom_char_0) poly_roots_finite: "poly p \<noteq> poly [] \<longleftrightarrow> finite {x. poly p x = 0}"
  51.588  proof
  51.589 -  assume F: "finite (UNIV :: 'a set)"
  51.590 -  have th0: "of_nat ` UNIV \<subseteq> (UNIV :: 'a set)" by simp
  51.591 -  from finite_subset[OF th0 F] have th: "finite (of_nat ` UNIV :: 'a set)" .
  51.592 -  have th': "inj_on (of_nat::nat \<Rightarrow> 'a) UNIV"
  51.593 -    unfolding inj_on_def by auto
  51.594 -  from finite_imageD[OF th th'] UNIV_nat_infinite
  51.595 -  show False by blast
  51.596 -qed
  51.597 -
  51.598 -lemma poly_roots_finite: "poly p \<noteq> poly [] \<longleftrightarrow> finite {x. poly p x = (0::'a::{idom,ring_char_0})}"
  51.599 -proof
  51.600 -  assume "poly p \<noteq> poly []"
  51.601 -  then show "finite {x. poly p x = (0::'a)}"
  51.602 +  assume H: "poly p \<noteq> poly []"
  51.603 +  show "finite {x. poly p x = (0::'a)}"
  51.604 +    using H
  51.605      apply -
  51.606 -    apply (erule contrapos_np)
  51.607 -    apply (rule ext)
  51.608 +    apply (erule contrapos_np, rule ext)
  51.609      apply (rule ccontr)
  51.610 -    apply (clarify dest!: poly_roots_finite_lemma')
  51.611 +    apply (clarify dest!: poly_roots_finite_lemma2)
  51.612      using finite_subset
  51.613    proof -
  51.614      fix x i
  51.615 @@ -377,119 +385,142 @@
  51.616      with finite_subset F show False by auto
  51.617    qed
  51.618  next
  51.619 -  assume "finite {x. poly p x = (0\<Colon>'a)}"
  51.620 -  then show "poly p \<noteq> poly []"
  51.621 -    using UNIV_ring_char_0_infinte by auto
  51.622 +  assume F: "finite {x. poly p x = (0\<Colon>'a)}"
  51.623 +  show "poly p \<noteq> poly []" using F UNIV_ring_char_0_infinte by auto
  51.624  qed
  51.625  
  51.626  text{*Entirety and Cancellation for polynomials*}
  51.627  
  51.628 -lemma poly_entire_lemma:
  51.629 -  "poly (p:: ('a::{idom,ring_char_0}) list) \<noteq> poly [] \<Longrightarrow> poly q \<noteq> poly [] \<Longrightarrow>
  51.630 -    poly (p *** q) \<noteq> poly []"
  51.631 -  by (auto simp add: poly_roots_finite poly_mult Collect_disj_eq)
  51.632 +lemma (in idom_char_0) poly_entire_lemma2:
  51.633 +  assumes p0: "poly p \<noteq> poly []"
  51.634 +    and q0: "poly q \<noteq> poly []"
  51.635 +  shows "poly (p***q) \<noteq> poly []"
  51.636 +proof -
  51.637 +  let ?S = "\<lambda>p. {x. poly p x = 0}"
  51.638 +  have "?S (p *** q) = ?S p \<union> ?S q" by (auto simp add: poly_mult)
  51.639 +  with p0 q0 show ?thesis  unfolding poly_roots_finite by auto
  51.640 +qed
  51.641  
  51.642 -lemma poly_entire:
  51.643 -  "poly (p *** q) = poly ([]::('a::{idom,ring_char_0}) list) \<longleftrightarrow>
  51.644 -    (poly p = poly [] \<or> poly q = poly [])"
  51.645 -  apply (auto dest: fun_cong simp add: poly_entire_lemma poly_mult)
  51.646 -  apply (blast intro: ccontr dest: poly_entire_lemma poly_mult [THEN subst])
  51.647 -  done
  51.648 +lemma (in idom_char_0) poly_entire:
  51.649 +  "poly (p *** q) = poly [] \<longleftrightarrow> poly p = poly [] \<or> poly q = poly []"
  51.650 +  using poly_entire_lemma2[of p q]
  51.651 +  by (auto simp add: fun_eq_iff poly_mult)
  51.652  
  51.653 -lemma poly_entire_neg:
  51.654 -  "poly (p *** q) \<noteq> poly ([]::('a::{idom,ring_char_0}) list) \<longleftrightarrow>
  51.655 -    poly p \<noteq> poly [] \<and> poly q \<noteq> poly []"
  51.656 +lemma (in idom_char_0) poly_entire_neg:
  51.657 +  "poly (p *** q) \<noteq> poly [] \<longleftrightarrow> poly p \<noteq> poly [] \<and> poly q \<noteq> poly []"
  51.658    by (simp add: poly_entire)
  51.659  
  51.660  lemma fun_eq: "f = g \<longleftrightarrow> (\<forall>x. f x = g x)"
  51.661    by auto
  51.662  
  51.663 -lemma poly_add_minus_zero_iff: "poly (p +++ -- q) = poly [] \<longleftrightarrow> poly p = poly q"
  51.664 -  by (auto simp add: field_simps poly_add poly_minus_def fun_eq poly_cmult)
  51.665 +lemma (in comm_ring_1) poly_add_minus_zero_iff:
  51.666 +  "poly (p +++ -- q) = poly [] \<longleftrightarrow> poly p = poly q"
  51.667 +  by (auto simp add: algebra_simps poly_add poly_minus_def fun_eq poly_cmult)
  51.668  
  51.669 -lemma poly_add_minus_mult_eq: "poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))"
  51.670 -  by (auto simp add: poly_add poly_minus_def fun_eq poly_mult poly_cmult distrib_left)
  51.671 +lemma (in comm_ring_1) poly_add_minus_mult_eq:
  51.672 +  "poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))"
  51.673 +  by (auto simp add: poly_add poly_minus_def fun_eq poly_mult poly_cmult algebra_simps)
  51.674  
  51.675 -lemma poly_mult_left_cancel:
  51.676 -  "(poly (p *** q) = poly (p *** r)) =
  51.677 -    (poly p = poly ([]::('a::{idom,ring_char_0}) list) | poly q = poly r)"
  51.678 -  apply (rule_tac p1 = "p *** q" in poly_add_minus_zero_iff [THEN subst])
  51.679 -  apply (auto simp add: poly_add_minus_mult_eq poly_entire poly_add_minus_zero_iff)
  51.680 -  done
  51.681 +subclass (in idom_char_0) comm_ring_1 ..
  51.682  
  51.683 -lemma poly_exp_eq_zero [simp]:
  51.684 -  "poly (p %^ n) = poly ([]::('a::idom) list) \<longleftrightarrow> poly p = poly [] \<and> n \<noteq> 0"
  51.685 +lemma (in idom_char_0) poly_mult_left_cancel:
  51.686 +  "poly (p *** q) = poly (p *** r) \<longleftrightarrow> poly p = poly [] \<or> poly q = poly r"
  51.687 +proof -
  51.688 +  have "poly (p *** q) = poly (p *** r) \<longleftrightarrow> poly (p *** q +++ -- (p *** r)) = poly []"
  51.689 +    by (simp only: poly_add_minus_zero_iff)
  51.690 +  also have "\<dots> \<longleftrightarrow> poly p = poly [] \<or> poly q = poly r"
  51.691 +    by (auto intro: simp add: poly_add_minus_mult_eq poly_entire poly_add_minus_zero_iff)
  51.692 +  finally show ?thesis .
  51.693 +qed
  51.694 +
  51.695 +lemma (in idom) poly_exp_eq_zero[simp]:
  51.696 +  "poly (p %^ n) = poly [] \<longleftrightarrow> poly p = poly [] \<and> n \<noteq> 0"
  51.697    apply (simp only: fun_eq add: HOL.all_simps [symmetric])
  51.698    apply (rule arg_cong [where f = All])
  51.699    apply (rule ext)
  51.700 -  apply (induct_tac n)
  51.701 -  apply (auto simp add: poly_mult)
  51.702 +  apply (induct n)
  51.703 +  apply (auto simp add: poly_exp poly_mult)
  51.704    done
  51.705  
  51.706 -lemma poly_prime_eq_zero [simp]: "poly [a, 1::'a::comm_ring_1] \<noteq> poly []"
  51.707 +lemma (in comm_ring_1) poly_prime_eq_zero[simp]: "poly [a,1] \<noteq> poly []"
  51.708    apply (simp add: fun_eq)
  51.709 -  apply (rule_tac x = "1 - a" in exI)
  51.710 -  apply simp
  51.711 +  apply (rule_tac x = "minus one a" in exI)
  51.712 +  apply (simp add: add_commute [of a])
  51.713    done
  51.714  
  51.715 -lemma poly_exp_prime_eq_zero [simp]: "poly ([a, (1::'a::idom)] %^ n) \<noteq> poly []"
  51.716 +lemma (in idom) poly_exp_prime_eq_zero: "poly ([a, 1] %^ n) \<noteq> poly []"
  51.717    by auto
  51.718  
  51.719  text{*A more constructive notion of polynomials being trivial*}
  51.720  
  51.721 -lemma poly_zero_lemma':
  51.722 -  "poly (h # t) = poly [] \<Longrightarrow> h = (0::'a::{idom,ring_char_0}) & poly t = poly []"
  51.723 +lemma (in idom_char_0) poly_zero_lemma': "poly (h # t) = poly [] \<Longrightarrow> h = 0 \<and> poly t = poly []"
  51.724    apply (simp add: fun_eq)
  51.725 -  apply (case_tac "h = 0")
  51.726 -  apply (drule_tac [2] x = 0 in spec)
  51.727 -  apply auto
  51.728 -  apply (case_tac "poly t = poly []")
  51.729 -  apply simp
  51.730 +  apply (case_tac "h = zero")
  51.731 +  apply (drule_tac [2] x = zero in spec, auto)
  51.732 +  apply (cases "poly t = poly []", simp)
  51.733  proof -
  51.734    fix x
  51.735 -  assume H: "\<forall>x. x = (0\<Colon>'a) \<or> poly t x = (0\<Colon>'a)"  and pnz: "poly t \<noteq> poly []"
  51.736 +  assume H: "\<forall>x. x = (0\<Colon>'a) \<or> poly t x = (0\<Colon>'a)"
  51.737 +    and pnz: "poly t \<noteq> poly []"
  51.738    let ?S = "{x. poly t x = 0}"
  51.739    from H have "\<forall>x. x \<noteq>0 \<longrightarrow> poly t x = 0" by blast
  51.740    hence th: "?S \<supseteq> UNIV - {0}" by auto
  51.741    from poly_roots_finite pnz have th': "finite ?S" by blast
  51.742 -  from finite_subset[OF th th'] UNIV_ring_char_0_infinte[where ?'a = 'a]
  51.743 -  show "poly t x = (0\<Colon>'a)" by simp
  51.744 +  from finite_subset[OF th th'] UNIV_ring_char_0_infinte show "poly t x = (0\<Colon>'a)"
  51.745 +    by simp
  51.746  qed
  51.747  
  51.748 -lemma poly_zero: "poly p = poly [] \<longleftrightarrow> list_all (\<lambda>c. c = (0::'a::{idom,ring_char_0})) p"
  51.749 +lemma (in idom_char_0) poly_zero: "(poly p = poly []) = list_all (%c. c = 0) p"
  51.750    apply (induct p)
  51.751    apply simp
  51.752    apply (rule iffI)
  51.753 -  apply (drule poly_zero_lemma')
  51.754 -  apply auto
  51.755 +  apply (drule poly_zero_lemma', auto)
  51.756    done
  51.757  
  51.758 +lemma (in idom_char_0) poly_0: "list_all (\<lambda>c. c = 0) p \<Longrightarrow> poly p x = 0"
  51.759 +  unfolding poly_zero[symmetric] by simp
  51.760 +
  51.761 +
  51.762  
  51.763  text{*Basics of divisibility.*}
  51.764  
  51.765 -lemma poly_primes: "[a, (1::'a::idom)] divides (p *** q) \<longleftrightarrow> [a, 1] divides p \<or> [a, 1] divides q"
  51.766 +lemma (in idom) poly_primes:
  51.767 +  "[a, 1] divides (p *** q) \<longleftrightarrow> [a, 1] divides p \<or> [a, 1] divides q"
  51.768    apply (auto simp add: divides_def fun_eq poly_mult poly_add poly_cmult distrib_right [symmetric])
  51.769 -  apply (drule_tac x = "-a" in spec)
  51.770 +  apply (drule_tac x = "uminus a" in spec)
  51.771 +  apply (simp add: poly_linear_divides poly_add poly_cmult distrib_right [symmetric])
  51.772 +  apply (cases "p = []")
  51.773 +  apply (rule exI[where x="[]"])
  51.774 +  apply simp
  51.775 +  apply (cases "q = []")
  51.776 +  apply (erule allE[where x="[]"], simp)
  51.777 +
  51.778 +  apply clarsimp
  51.779 +  apply (cases "\<exists>q\<Colon>'a list. p = a %* q +++ ((0\<Colon>'a) # q)")
  51.780 +  apply (clarsimp simp add: poly_add poly_cmult)
  51.781 +  apply (rule_tac x="qa" in exI)
  51.782 +  apply (simp add: distrib_right [symmetric])
  51.783 +  apply clarsimp
  51.784 +
  51.785    apply (auto simp add: poly_linear_divides poly_add poly_cmult distrib_right [symmetric])
  51.786 -  apply (rule_tac x = "qa *** q" in exI)
  51.787 -  apply (rule_tac [2] x = "p *** qa" in exI)
  51.788 +  apply (rule_tac x = "pmult qa q" in exI)
  51.789 +  apply (rule_tac [2] x = "pmult p qa" in exI)
  51.790    apply (auto simp add: poly_add poly_mult poly_cmult mult_ac)
  51.791    done
  51.792  
  51.793 -lemma poly_divides_refl [simp]: "p divides p"
  51.794 +lemma (in comm_semiring_1) poly_divides_refl[simp]: "p divides p"
  51.795    apply (simp add: divides_def)
  51.796 -  apply (rule_tac x = "[1]" in exI)
  51.797 +  apply (rule_tac x = "[one]" in exI)
  51.798    apply (auto simp add: poly_mult fun_eq)
  51.799    done
  51.800  
  51.801 -lemma poly_divides_trans: "p divides q \<Longrightarrow> q divides r \<Longrightarrow> p divides r"
  51.802 -  apply (simp add: divides_def)
  51.803 -  apply safe
  51.804 -  apply (rule_tac x = "qa *** qaa" in exI)
  51.805 +lemma (in comm_semiring_1) poly_divides_trans: "p divides q \<Longrightarrow> q divides r \<Longrightarrow> p divides r"
  51.806 +  apply (simp add: divides_def, safe)
  51.807 +  apply (rule_tac x = "pmult qa qaa" in exI)
  51.808    apply (auto simp add: poly_mult fun_eq mult_assoc)
  51.809    done
  51.810  
  51.811 -lemma poly_divides_exp: "m \<le> n \<Longrightarrow> (p %^ m) divides (p %^ n)"
  51.812 +lemma (in comm_semiring_1) poly_divides_exp: "m \<le> n \<Longrightarrow> (p %^ m) divides (p %^ n)"
  51.813    apply (auto simp add: le_iff_add)
  51.814    apply (induct_tac k)
  51.815    apply (rule_tac [2] poly_divides_trans)
  51.816 @@ -498,34 +529,37 @@
  51.817    apply (auto simp add: poly_mult fun_eq mult_ac)
  51.818    done
  51.819  
  51.820 -lemma poly_exp_divides: "(p %^ n) divides q \<Longrightarrow> m \<le> n \<Longrightarrow> (p %^ m) divides q"
  51.821 +lemma (in comm_semiring_1) poly_exp_divides:
  51.822 +  "(p %^ n) divides q \<Longrightarrow> m \<le> n \<Longrightarrow> (p %^ m) divides q"
  51.823    by (blast intro: poly_divides_exp poly_divides_trans)
  51.824  
  51.825 -lemma poly_divides_add: "p divides q \<Longrightarrow> p divides r \<Longrightarrow> p divides (q +++ r)"
  51.826 -  apply (simp add: divides_def)
  51.827 -  apply auto
  51.828 -  apply (rule_tac x = "qa +++ qaa" in exI)
  51.829 +lemma (in comm_semiring_0) poly_divides_add:
  51.830 +  "p divides q \<Longrightarrow> p divides r \<Longrightarrow> p divides (q +++ r)"
  51.831 +  apply (simp add: divides_def, auto)
  51.832 +  apply (rule_tac x = "padd qa qaa" in exI)
  51.833    apply (auto simp add: poly_add fun_eq poly_mult distrib_left)
  51.834    done
  51.835  
  51.836 -lemma poly_divides_diff: "p divides q \<Longrightarrow> p divides (q +++ r) \<Longrightarrow> p divides r"
  51.837 -  apply (auto simp add: divides_def)
  51.838 -  apply (rule_tac x = "qaa +++ -- qa" in exI)
  51.839 +lemma (in comm_ring_1) poly_divides_diff:
  51.840 +  "p divides q \<Longrightarrow> p divides (q +++ r) \<Longrightarrow> p divides r"
  51.841 +  apply (simp add: divides_def, auto)
  51.842 +  apply (rule_tac x = "padd qaa (poly_minus qa)" in exI)
  51.843    apply (auto simp add: poly_add fun_eq poly_mult poly_minus algebra_simps)
  51.844    done
  51.845  
  51.846 -lemma poly_divides_diff2: "p divides r \<Longrightarrow> p divides (q +++ r) \<Longrightarrow> p divides q"
  51.847 +lemma (in comm_ring_1) poly_divides_diff2:
  51.848 +  "p divides r \<Longrightarrow> p divides (q +++ r) \<Longrightarrow> p divides q"
  51.849    apply (erule poly_divides_diff)
  51.850    apply (auto simp add: poly_add fun_eq poly_mult divides_def add_ac)
  51.851    done
  51.852  
  51.853 -lemma poly_divides_zero: "poly p = poly [] \<Longrightarrow> q divides p"
  51.854 +lemma (in semiring_0) poly_divides_zero: "poly p = poly [] \<Longrightarrow> q divides p"
  51.855    apply (simp add: divides_def)
  51.856 -  apply (rule exI [where x = "[]"])
  51.857 +  apply (rule exI[where x="[]"])
  51.858    apply (auto simp add: fun_eq poly_mult)
  51.859    done
  51.860  
  51.861 -lemma poly_divides_zero2 [simp]: "q divides []"
  51.862 +lemma (in semiring_0) poly_divides_zero2 [simp]: "q divides []"
  51.863    apply (simp add: divides_def)
  51.864    apply (rule_tac x = "[]" in exI)
  51.865    apply (auto simp add: fun_eq)
  51.866 @@ -533,195 +567,256 @@
  51.867  
  51.868  text{*At last, we can consider the order of a root.*}
  51.869  
  51.870 -lemma poly_order_exists_lemma [rule_format]:
  51.871 -  "\<forall>p. length p = d \<longrightarrow> poly p \<noteq> poly [] \<longrightarrow>
  51.872 -    (\<exists>n q. p = mulexp n [-a, (1::'a::{idom,ring_char_0})] q & poly q a \<noteq> 0)"
  51.873 -  apply (induct "d")
  51.874 -  apply (simp add: fun_eq)
  51.875 -  apply safe
  51.876 -  apply (case_tac "poly p a = 0")
  51.877 -  apply (drule_tac poly_linear_divides [THEN iffD1])
  51.878 -  apply safe
  51.879 -  apply (drule_tac x = q in spec)
  51.880 -  apply (drule_tac poly_entire_neg [THEN iffD1])
  51.881 -  apply safe
  51.882 -  apply force
  51.883 -  apply (rule_tac x = "Suc n" in exI)
  51.884 -  apply (rule_tac x = qa in exI)
  51.885 -  apply (simp del: pmult_Cons)
  51.886 -  apply (rule_tac x = 0 in exI)
  51.887 -  apply force
  51.888 -  done
  51.889 +lemma (in idom_char_0) poly_order_exists_lemma:
  51.890 +  assumes lp: "length p = d"
  51.891 +    and p: "poly p \<noteq> poly []"
  51.892 +  shows "\<exists>n q. p = mulexp n [-a, 1] q \<and> poly q a \<noteq> 0"
  51.893 +  using lp p
  51.894 +proof (induct d arbitrary: p)
  51.895 +  case 0
  51.896 +  thus ?case by simp
  51.897 +next
  51.898 +  case (Suc n p)
  51.899 +  show ?case
  51.900 +  proof (cases "poly p a = 0")
  51.901 +    case True
  51.902 +    from Suc.prems have h: "length p = Suc n" "poly p \<noteq> poly []" by auto
  51.903 +    hence pN: "p \<noteq> []" by auto
  51.904 +    from True[unfolded poly_linear_divides] pN obtain q where q: "p = [-a, 1] *** q"
  51.905 +      by blast
  51.906 +    from q h True have qh: "length q = n" "poly q \<noteq> poly []"
  51.907 +      apply -
  51.908 +      apply simp
  51.909 +      apply (simp only: fun_eq)
  51.910 +      apply (rule ccontr)
  51.911 +      apply (simp add: fun_eq poly_add poly_cmult)
  51.912 +      done
  51.913 +    from Suc.hyps[OF qh] obtain m r where mr: "q = mulexp m [-a,1] r" "poly r a \<noteq> 0"
  51.914 +      by blast
  51.915 +    from mr q have "p = mulexp (Suc m) [-a,1] r \<and> poly r a \<noteq> 0" by simp
  51.916 +    then show ?thesis by blast
  51.917 +  next
  51.918 +    case False
  51.919 +    then show ?thesis
  51.920 +      using Suc.prems
  51.921 +      apply simp
  51.922 +      apply (rule exI[where x="0::nat"])
  51.923 +      apply simp
  51.924 +      done
  51.925 +  qed
  51.926 +qed
  51.927 +
  51.928 +
  51.929 +lemma (in comm_semiring_1) poly_mulexp: "poly (mulexp n p q) x = (poly p x) ^ n * poly q x"
  51.930 +  by (induct n) (auto simp add: poly_mult mult_ac)
  51.931 +
  51.932 +lemma (in comm_semiring_1) divides_left_mult:
  51.933 +  assumes d:"(p***q) divides r" shows "p divides r \<and> q divides r"
  51.934 +proof-
  51.935 +  from d obtain t where r:"poly r = poly (p***q *** t)"
  51.936 +    unfolding divides_def by blast
  51.937 +  hence "poly r = poly (p *** (q *** t))"
  51.938 +    "poly r = poly (q *** (p***t))" by(auto simp add: fun_eq poly_mult mult_ac)
  51.939 +  thus ?thesis unfolding divides_def by blast
  51.940 +qed
  51.941 +
  51.942  
  51.943  (* FIXME: Tidy up *)
  51.944 -lemma poly_order_exists:
  51.945 -  "length p = d \<Longrightarrow> poly p \<noteq> poly [] \<Longrightarrow>
  51.946 -    \<exists>n. ([-a, 1] %^ n) divides p \<and> \<not> (([-a, (1::'a::{idom,ring_char_0})] %^ (Suc n)) divides p)"
  51.947 -  apply (drule poly_order_exists_lemma [where a=a])
  51.948 -  apply assumption
  51.949 -  apply clarify
  51.950 -  apply (rule_tac x = n in exI)
  51.951 -  apply safe
  51.952 -  apply (unfold divides_def)
  51.953 -  apply (rule_tac x = q in exI)
  51.954 -  apply (induct_tac n)
  51.955 -  apply simp
  51.956 -  apply (simp add: poly_add poly_cmult poly_mult distrib_left mult_ac)
  51.957 -  apply safe
  51.958 -  apply (subgoal_tac "poly (mulexp n [- a, 1] q) \<noteq> poly ([- a, 1] %^ Suc n *** qa)")
  51.959 -  apply simp
  51.960 -  apply (induct_tac n)
  51.961 -  apply (simp del: pmult_Cons pexp_Suc)
  51.962 -  apply (erule_tac Q = "poly q a = 0" in contrapos_np)
  51.963 -  apply (simp add: poly_add poly_cmult)
  51.964 -  apply (rule pexp_Suc [THEN ssubst])
  51.965 -  apply (rule ccontr)
  51.966 -  apply (simp add: poly_mult_left_cancel poly_mult_assoc del: pmult_Cons pexp_Suc)
  51.967 -  done
  51.968  
  51.969 -lemma poly_one_divides [simp]: "[1] divides p"
  51.970 -  by (auto simp: divides_def)
  51.971 +lemma (in semiring_1) zero_power_iff: "0 ^ n = (if n = 0 then 1 else 0)"
  51.972 +  by (induct n) simp_all
  51.973  
  51.974 -lemma poly_order: "poly p \<noteq> poly [] \<Longrightarrow>
  51.975 -    \<exists>! n. ([-a, (1::'a::{idom,ring_char_0})] %^ n) divides p \<and> \<not> (([-a, 1] %^ Suc n) divides p)"
  51.976 +lemma (in idom_char_0) poly_order_exists:
  51.977 +  assumes "length p = d" and "poly p \<noteq> poly []"
  51.978 +  shows "\<exists>n. [- a, 1] %^ n divides p \<and> \<not> [- a, 1] %^ Suc n divides p"
  51.979 +proof -
  51.980 +  from assms have "\<exists>n q. p = mulexp n [- a, 1] q \<and> poly q a \<noteq> 0"
  51.981 +    by (rule poly_order_exists_lemma)
  51.982 +  then obtain n q where p: "p = mulexp n [- a, 1] q" and "poly q a \<noteq> 0" by blast
  51.983 +  have "[- a, 1] %^ n divides mulexp n [- a, 1] q"
  51.984 +  proof (rule dividesI)
  51.985 +    show "poly (mulexp n [- a, 1] q) = poly ([- a, 1] %^ n *** q)"
  51.986 +      by (induct n) (simp_all add: poly_add poly_cmult poly_mult algebra_simps)
  51.987 +  qed
  51.988 +  moreover have "\<not> [- a, 1] %^ Suc n divides mulexp n [- a, 1] q"
  51.989 +  proof
  51.990 +    assume "[- a, 1] %^ Suc n divides mulexp n [- a, 1] q"
  51.991 +    then obtain m where "poly (mulexp n [- a, 1] q) = poly ([- a, 1] %^ Suc n *** m)"
  51.992 +      by (rule dividesE)
  51.993 +    moreover have "poly (mulexp n [- a, 1] q) \<noteq> poly ([- a, 1] %^ Suc n *** m)"
  51.994 +    proof (induct n)
  51.995 +      case 0 show ?case
  51.996 +      proof (rule ccontr)
  51.997 +        assume "\<not> poly (mulexp 0 [- a, 1] q) \<noteq> poly ([- a, 1] %^ Suc 0 *** m)"
  51.998 +        then have "poly q a = 0"
  51.999 +          by (simp add: poly_add poly_cmult)
 51.1000 +        with `poly q a \<noteq> 0` show False by simp
 51.1001 +      qed
 51.1002 +    next
 51.1003 +      case (Suc n) show ?case
 51.1004 +        by (rule pexp_Suc [THEN ssubst], rule ccontr)
 51.1005 +          (simp add: poly_mult_left_cancel poly_mult_assoc Suc del: pmult_Cons pexp_Suc)
 51.1006 +    qed
 51.1007 +    ultimately show False by simp
 51.1008 +  qed
 51.1009 +  ultimately show ?thesis by (auto simp add: p)
 51.1010 +qed
 51.1011 +
 51.1012 +lemma (in semiring_1) poly_one_divides[simp]: "[1] divides p"
 51.1013 +  by (auto simp add: divides_def)
 51.1014 +
 51.1015 +lemma (in idom_char_0) poly_order:
 51.1016 +  "poly p \<noteq> poly [] \<Longrightarrow> \<exists>!n. ([-a, 1] %^ n) divides p \<and> \<not> (([-a, 1] %^ Suc n) divides p)"
 51.1017    apply (auto intro: poly_order_exists simp add: less_linear simp del: pmult_Cons pexp_Suc)
 51.1018    apply (cut_tac x = y and y = n in less_linear)
 51.1019    apply (drule_tac m = n in poly_exp_divides)
 51.1020    apply (auto dest: Suc_le_eq [THEN iffD2, THEN [2] poly_exp_divides]
 51.1021 -    simp del: pmult_Cons pexp_Suc)
 51.1022 +              simp del: pmult_Cons pexp_Suc)
 51.1023    done
 51.1024  
 51.1025  text{*Order*}
 51.1026  
 51.1027 -lemma some1_equalityD: "n = (SOME n. P n) \<Longrightarrow> EX! n. P n \<Longrightarrow> P n"
 51.1028 +lemma some1_equalityD: "n = (SOME n. P n) \<Longrightarrow> \<exists>!n. P n \<Longrightarrow> P n"
 51.1029    by (blast intro: someI2)
 51.1030  
 51.1031 -lemma order:
 51.1032 -  "(([-a, (1::'a::{idom,ring_char_0})] %^ n) divides p &
 51.1033 -    ~(([-a, 1] %^ (Suc n)) divides p)) =
 51.1034 -    ((n = order a p) & ~(poly p = poly []))"
 51.1035 +lemma (in idom_char_0) order:
 51.1036 +      "(([-a, 1] %^ n) divides p \<and>
 51.1037 +        ~(([-a, 1] %^ (Suc n)) divides p)) =
 51.1038 +        ((n = order a p) \<and> ~(poly p = poly []))"
 51.1039    apply (unfold order_def)
 51.1040    apply (rule iffI)
 51.1041    apply (blast dest: poly_divides_zero intro!: some1_equality [symmetric] poly_order)
 51.1042    apply (blast intro!: poly_order [THEN [2] some1_equalityD])
 51.1043    done
 51.1044  
 51.1045 -lemma order2: "poly p \<noteq> poly [] \<Longrightarrow>
 51.1046 -  ([-a, (1::'a::{idom,ring_char_0})] %^ (order a p)) divides p &
 51.1047 -    ~(([-a, 1] %^ (Suc(order a p))) divides p)"
 51.1048 +lemma (in idom_char_0) order2:
 51.1049 +  "poly p \<noteq> poly [] \<Longrightarrow>
 51.1050 +    ([-a, 1] %^ (order a p)) divides p \<and> \<not> (([-a, 1] %^ (Suc (order a p))) divides p)"
 51.1051    by (simp add: order del: pexp_Suc)
 51.1052  
 51.1053 -lemma order_unique: "poly p \<noteq> poly [] \<Longrightarrow> ([-a, 1] %^ n) divides p \<Longrightarrow>
 51.1054 -  \<not> (([-a, (1::'a::{idom,ring_char_0})] %^ (Suc n)) divides p) \<Longrightarrow> n = order a p"
 51.1055 +lemma (in idom_char_0) order_unique:
 51.1056 +  "poly p \<noteq> poly [] \<Longrightarrow> ([-a, 1] %^ n) divides p \<Longrightarrow> ~(([-a, 1] %^ (Suc n)) divides p) \<Longrightarrow>
 51.1057 +    n = order a p"
 51.1058    using order [of a n p] by auto
 51.1059  
 51.1060 -lemma order_unique_lemma:
 51.1061 -  "(poly p \<noteq> poly [] \<and> ([-a, 1] %^ n) divides p \<and>
 51.1062 -    \<not> (([-a, (1::'a::{idom,ring_char_0})] %^ (Suc n)) divides p)) \<Longrightarrow>
 51.1063 +lemma (in idom_char_0) order_unique_lemma:
 51.1064 +  "poly p \<noteq> poly [] \<and> ([-a, 1] %^ n) divides p \<and> ~(([-a, 1] %^ (Suc n)) divides p) \<Longrightarrow>
 51.1065      n = order a p"
 51.1066    by (blast intro: order_unique)
 51.1067  
 51.1068 -lemma order_poly: "poly p = poly q ==> order a p = order a q"
 51.1069 +lemma (in ring_1) order_poly: "poly p = poly q \<Longrightarrow> order a p = order a q"
 51.1070    by (auto simp add: fun_eq divides_def poly_mult order_def)
 51.1071  
 51.1072 -lemma pexp_one [simp]: "p %^ (Suc 0) = p"
 51.1073 -  by (induct p) simp_all
 51.1074 +lemma (in semiring_1) pexp_one[simp]: "p %^ (Suc 0) = p"
 51.1075 +  by (induct "p") auto
 51.1076  
 51.1077 -lemma lemma_order_root:
 51.1078 -  "0 < n & [- a, 1] %^ n divides p & ~ [- a, 1] %^ (Suc n) divides p \<Longrightarrow> poly p a = 0"
 51.1079 -  apply (induct n arbitrary: p a)
 51.1080 -  apply blast
 51.1081 -  apply (auto simp add: divides_def poly_mult simp del: pmult_Cons)
 51.1082 +lemma (in comm_ring_1) lemma_order_root:
 51.1083 +  "0 < n \<and> [- a, 1] %^ n divides p \<and> ~ [- a, 1] %^ (Suc n) divides p \<Longrightarrow> poly p a = 0"
 51.1084 +  by (induct n arbitrary: a p) (auto simp add: divides_def poly_mult simp del: pmult_Cons)
 51.1085 +
 51.1086 +lemma (in idom_char_0) order_root:
 51.1087 +  "poly p a = 0 \<longleftrightarrow> poly p = poly [] \<or> order a p \<noteq> 0"
 51.1088 +  apply (cases "poly p = poly []")
 51.1089 +  apply auto
 51.1090 +  apply (simp add: poly_linear_divides del: pmult_Cons, safe)
 51.1091 +  apply (drule_tac [!] a = a in order2)
 51.1092 +  apply (rule ccontr)
 51.1093 +  apply (simp add: divides_def poly_mult fun_eq del: pmult_Cons, blast)
 51.1094 +  using neq0_conv
 51.1095 +  apply (blast intro: lemma_order_root)
 51.1096    done
 51.1097  
 51.1098 -lemma order_root: "poly p a = (0::'a::{idom,ring_char_0}) \<longleftrightarrow> poly p = poly [] \<or> order a p \<noteq> 0"
 51.1099 -  apply (cases "poly p = poly []")
 51.1100 -  apply auto
 51.1101 -  apply (simp add: poly_linear_divides del: pmult_Cons)
 51.1102 -  apply safe
 51.1103 -  apply (drule_tac [!] a = a in order2)
 51.1104 -  apply (rule ccontr)
 51.1105 -  apply (simp add: divides_def poly_mult fun_eq del: pmult_Cons)
 51.1106 -  apply blast
 51.1107 -  using neq0_conv apply (blast intro: lemma_order_root)
 51.1108 -  done
 51.1109 -
 51.1110 -lemma order_divides: "([-a, 1::'a::{idom,ring_char_0}] %^ n) divides p \<longleftrightarrow>
 51.1111 -    poly p = poly [] \<or> n \<le> order a p"
 51.1112 +lemma (in idom_char_0) order_divides:
 51.1113 +  "([-a, 1] %^ n) divides p \<longleftrightarrow> poly p = poly [] \<or> n \<le> order a p"
 51.1114    apply (cases "poly p = poly []")
 51.1115    apply auto
 51.1116    apply (simp add: divides_def fun_eq poly_mult)
 51.1117    apply (rule_tac x = "[]" in exI)
 51.1118 -  apply (auto dest!: order2 [where a = a] intro: poly_exp_divides simp del: pexp_Suc)
 51.1119 +  apply (auto dest!: order2 [where a=a] intro: poly_exp_divides simp del: pexp_Suc)
 51.1120    done
 51.1121  
 51.1122 -lemma order_decomp:
 51.1123 -  "poly p \<noteq> poly [] \<Longrightarrow>
 51.1124 -    \<exists>q. poly p = poly (([-a, 1] %^ (order a p)) *** q) \<and>
 51.1125 -      \<not> ([-a, 1::'a::{idom,ring_char_0}] divides q)"
 51.1126 +lemma (in idom_char_0) order_decomp:
 51.1127 +  "poly p \<noteq> poly [] \<Longrightarrow> \<exists>q. poly p = poly (([-a, 1] %^ (order a p)) *** q) \<and> ~([-a, 1] divides q)"
 51.1128    apply (unfold divides_def)
 51.1129    apply (drule order2 [where a = a])
 51.1130 -  apply (simp add: divides_def del: pexp_Suc pmult_Cons)
 51.1131 -  apply safe
 51.1132 -  apply (rule_tac x = q in exI)
 51.1133 -  apply safe
 51.1134 +  apply (simp add: divides_def del: pexp_Suc pmult_Cons, safe)
 51.1135 +  apply (rule_tac x = q in exI, safe)
 51.1136    apply (drule_tac x = qa in spec)
 51.1137    apply (auto simp add: poly_mult fun_eq poly_exp mult_ac simp del: pmult_Cons)
 51.1138    done
 51.1139  
 51.1140  text{*Important composition properties of orders.*}
 51.1141 -
 51.1142 -lemma order_mult: "poly (p *** q) \<noteq> poly [] \<Longrightarrow>
 51.1143 -  order a (p *** q) = order a p + order (a::'a::{idom,ring_char_0}) q"
 51.1144 -  apply (cut_tac a = a and p = "p***q" and n = "order a p + order a q" in order)
 51.1145 +lemma order_mult:
 51.1146 +  "poly (p *** q) \<noteq> poly [] \<Longrightarrow>
 51.1147 +    order a (p *** q) = order a p + order (a::'a::{idom_char_0}) q"
 51.1148 +  apply (cut_tac a = a and p = "p *** q" and n = "order a p + order a q" in order)
 51.1149    apply (auto simp add: poly_entire simp del: pmult_Cons)
 51.1150    apply (drule_tac a = a in order2)+
 51.1151    apply safe
 51.1152 -  apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons)
 51.1153 -  apply safe
 51.1154 +  apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe)
 51.1155    apply (rule_tac x = "qa *** qaa" in exI)
 51.1156    apply (simp add: poly_mult mult_ac del: pmult_Cons)
 51.1157    apply (drule_tac a = a in order_decomp)+
 51.1158    apply safe
 51.1159 -  apply (subgoal_tac "[-a, 1] divides (qa *** qaa) ")
 51.1160 +  apply (subgoal_tac "[-a,1] divides (qa *** qaa) ")
 51.1161    apply (simp add: poly_primes del: pmult_Cons)
 51.1162    apply (auto simp add: divides_def simp del: pmult_Cons)
 51.1163    apply (rule_tac x = qb in exI)
 51.1164 -  apply (subgoal_tac "poly ([-a, 1] %^ (order a p) *** (qa *** qaa)) =
 51.1165 -    poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))")
 51.1166 -  apply (drule poly_mult_left_cancel [THEN iffD1])
 51.1167 -  apply force
 51.1168 -  apply (subgoal_tac "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (qa *** qaa))) =
 51.1169 -    poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))) ")
 51.1170 -  apply (drule poly_mult_left_cancel [THEN iffD1])
 51.1171 -  apply force
 51.1172 +  apply (subgoal_tac "poly ([-a, 1] %^ (order a p) *** (qa *** qaa)) = poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))")
 51.1173 +  apply (drule poly_mult_left_cancel [THEN iffD1], force)
 51.1174 +  apply (subgoal_tac "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (qa *** qaa))) = poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))) ")
 51.1175 +  apply (drule poly_mult_left_cancel [THEN iffD1], force)
 51.1176    apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons)
 51.1177    done
 51.1178  
 51.1179 -lemma order_root2: "poly p \<noteq> poly [] \<Longrightarrow> poly p a = 0 \<longleftrightarrow> order (a::'a::{idom,ring_char_0}) p \<noteq> 0"
 51.1180 +lemma (in idom_char_0) order_mult:
 51.1181 +  assumes "poly (p *** q) \<noteq> poly []"
 51.1182 +  shows "order a (p *** q) = order a p + order a q"
 51.1183 +  using assms
 51.1184 +  apply (cut_tac a = a and p = "pmult p q" and n = "order a p + order a q" in order)
 51.1185 +  apply (auto simp add: poly_entire simp del: pmult_Cons)
 51.1186 +  apply (drule_tac a = a in order2)+
 51.1187 +  apply safe
 51.1188 +  apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe)
 51.1189 +  apply (rule_tac x = "pmult qa qaa" in exI)
 51.1190 +  apply (simp add: poly_mult mult_ac del: pmult_Cons)
 51.1191 +  apply (drule_tac a = a in order_decomp)+
 51.1192 +  apply safe
 51.1193 +  apply (subgoal_tac "[uminus a, one] divides pmult qa qaa")
 51.1194 +  apply (simp add: poly_primes del: pmult_Cons)
 51.1195 +  apply (auto simp add: divides_def simp del: pmult_Cons)
 51.1196 +  apply (rule_tac x = qb in exI)
 51.1197 +  apply (subgoal_tac "poly (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa)) =
 51.1198 +    poly (pmult (pexp [uminus a, one] (?order a p)) (pmult [uminus a, one] qb))")
 51.1199 +  apply (drule poly_mult_left_cancel [THEN iffD1], force)
 51.1200 +  apply (subgoal_tac "poly (pmult (pexp [uminus a, one] (order a q))
 51.1201 +      (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa))) =
 51.1202 +    poly (pmult (pexp [uminus a, one] (order a q))
 51.1203 +      (pmult (pexp [uminus a, one] (order a p)) (pmult [uminus a, one] qb)))")
 51.1204 +  apply (drule poly_mult_left_cancel [THEN iffD1], force)
 51.1205 +  apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons)
 51.1206 +  done
 51.1207 +
 51.1208 +lemma (in idom_char_0) order_root2: "poly p \<noteq> poly [] \<Longrightarrow> poly p a = 0 \<longleftrightarrow> order a p \<noteq> 0"
 51.1209    by (rule order_root [THEN ssubst]) auto
 51.1210  
 51.1211 -lemma pmult_one [simp]: "[1] *** p = p"
 51.1212 -  by auto
 51.1213 +lemma (in semiring_1) pmult_one[simp]: "[1] *** p = p" by auto
 51.1214  
 51.1215 -lemma poly_Nil_zero: "poly [] = poly [0]"
 51.1216 +lemma (in semiring_0) poly_Nil_zero: "poly [] = poly [0]"
 51.1217    by (simp add: fun_eq)
 51.1218  
 51.1219 -lemma rsquarefree_decomp:
 51.1220 -  "rsquarefree p \<Longrightarrow> poly p a = (0::'a::{idom,ring_char_0}) \<Longrightarrow>
 51.1221 +lemma (in idom_char_0) rsquarefree_decomp:
 51.1222 +  "rsquarefree p \<Longrightarrow> poly p a = 0 \<Longrightarrow>
 51.1223      \<exists>q. poly p = poly ([-a, 1] *** q) \<and> poly q a \<noteq> 0"
 51.1224 -  apply (simp add: rsquarefree_def)
 51.1225 -  apply safe
 51.1226 +  apply (simp add: rsquarefree_def, safe)
 51.1227    apply (frule_tac a = a in order_decomp)
 51.1228    apply (drule_tac x = a in spec)
 51.1229    apply (drule_tac a = a in order_root2 [symmetric])
 51.1230    apply (auto simp del: pmult_Cons)
 51.1231 -  apply (rule_tac x = q in exI)
 51.1232 -  apply safe
 51.1233 +  apply (rule_tac x = q in exI, safe)
 51.1234    apply (simp add: poly_mult fun_eq)
 51.1235    apply (drule_tac p1 = q in poly_linear_divides [THEN iffD1])
 51.1236 -  apply (simp add: divides_def del: pmult_Cons)
 51.1237 -  apply safe
 51.1238 +  apply (simp add: divides_def del: pmult_Cons, safe)
 51.1239    apply (drule_tac x = "[]" in spec)
 51.1240    apply (auto simp add: fun_eq)
 51.1241    done
 51.1242 @@ -729,72 +824,222 @@
 51.1243  
 51.1244  text{*Normalization of a polynomial.*}
 51.1245  
 51.1246 -lemma poly_normalize [simp]: "poly (pnormalize p) = poly p"
 51.1247 +lemma (in semiring_0) poly_normalize[simp]: "poly (pnormalize p) = poly p"
 51.1248    by (induct p) (auto simp add: fun_eq)
 51.1249  
 51.1250 -
 51.1251  text{*The degree of a polynomial.*}
 51.1252  
 51.1253 -lemma lemma_degree_zero: "list_all (\<lambda>c. c = 0) p \<longleftrightarrow> pnormalize p = []"
 51.1254 +lemma (in semiring_0) lemma_degree_zero: "list_all (%c. c = 0) p \<longleftrightarrow> pnormalize p = []"
 51.1255    by (induct p) auto
 51.1256  
 51.1257 -lemma degree_zero: "poly p = poly ([] :: 'a::{idom,ring_char_0} list) \<Longrightarrow> degree p = 0"
 51.1258 -  apply (cases "pnormalize p = []")
 51.1259 -  apply (simp add: degree_def)
 51.1260 -  apply (auto simp add: poly_zero lemma_degree_zero)
 51.1261 -  done
 51.1262 +lemma (in idom_char_0) degree_zero:
 51.1263 +  assumes "poly p = poly []"
 51.1264 +  shows "degree p = 0"
 51.1265 +  using assms
 51.1266 +  by (cases "pnormalize p = []") (auto simp add: degree_def poly_zero lemma_degree_zero)
 51.1267  
 51.1268 -lemma pnormalize_sing: "pnormalize [x] = [x] \<longleftrightarrow> x \<noteq> 0"
 51.1269 +lemma (in semiring_0) pnormalize_sing: "(pnormalize [x] = [x]) \<longleftrightarrow> x \<noteq> 0"
 51.1270    by simp
 51.1271  
 51.1272 -lemma pnormalize_pair: "y \<noteq> 0 \<longleftrightarrow> (pnormalize [x, y] = [x, y])"
 51.1273 +lemma (in semiring_0) pnormalize_pair: "y \<noteq> 0 \<longleftrightarrow> (pnormalize [x, y] = [x, y])"
 51.1274    by simp
 51.1275  
 51.1276 -lemma pnormal_cons: "pnormal p \<Longrightarrow> pnormal (c # p)"
 51.1277 +lemma (in semiring_0) pnormal_cons: "pnormal p \<Longrightarrow> pnormal (c#p)"
 51.1278    unfolding pnormal_def by simp
 51.1279  
 51.1280 -lemma pnormal_tail: "p \<noteq> [] \<Longrightarrow> pnormal (c # p) \<Longrightarrow> pnormal p"
 51.1281 -  unfolding pnormal_def
 51.1282 -  apply (cases "pnormalize p = []")
 51.1283 +lemma (in semiring_0) pnormal_tail: "p\<noteq>[] \<Longrightarrow> pnormal (c#p) \<Longrightarrow> pnormal p"
 51.1284 +  unfolding pnormal_def by(auto split: split_if_asm)
 51.1285 +
 51.1286 +
 51.1287 +lemma (in semiring_0) pnormal_last_nonzero: "pnormal p \<Longrightarrow> last p \<noteq> 0"
 51.1288 +  by (induct p) (simp_all add: pnormal_def split: split_if_asm)
 51.1289 +
 51.1290 +lemma (in semiring_0) pnormal_length: "pnormal p \<Longrightarrow> 0 < length p"
 51.1291 +  unfolding pnormal_def length_greater_0_conv by blast
 51.1292 +
 51.1293 +lemma (in semiring_0) pnormal_last_length: "0 < length p \<Longrightarrow> last p \<noteq> 0 \<Longrightarrow> pnormal p"
 51.1294 +  by (induct p) (auto simp: pnormal_def  split: split_if_asm)
 51.1295 +
 51.1296 +
 51.1297 +lemma (in semiring_0) pnormal_id: "pnormal p \<longleftrightarrow> 0 < length p \<and> last p \<noteq> 0"
 51.1298 +  using pnormal_last_length pnormal_length pnormal_last_nonzero by blast
 51.1299 +
 51.1300 +lemma (in idom_char_0) poly_Cons_eq:
 51.1301 +  "poly (c # cs) = poly (d # ds) \<longleftrightarrow> c = d \<and> poly cs = poly ds"
 51.1302 +  (is "?lhs \<longleftrightarrow> ?rhs")
 51.1303 +proof
 51.1304 +  assume eq: ?lhs
 51.1305 +  hence "\<And>x. poly ((c#cs) +++ -- (d#ds)) x = 0"
 51.1306 +    by (simp only: poly_minus poly_add algebra_simps) (simp add: algebra_simps)
 51.1307 +  hence "poly ((c#cs) +++ -- (d#ds)) = poly []" by(simp add: fun_eq_iff)
 51.1308 +  hence "c = d \<and> list_all (\<lambda>x. x=0) ((cs +++ -- ds))"
 51.1309 +    unfolding poly_zero by (simp add: poly_minus_def algebra_simps)
 51.1310 +  hence "c = d \<and> (\<forall>x. poly (cs +++ -- ds) x = 0)"
 51.1311 +    unfolding poly_zero[symmetric] by simp
 51.1312 +  then show ?rhs by (simp add: poly_minus poly_add algebra_simps fun_eq_iff)
 51.1313 +next
 51.1314 +  assume ?rhs
 51.1315 +  then show ?lhs by(simp add:fun_eq_iff)
 51.1316 +qed
 51.1317 +
 51.1318 +lemma (in idom_char_0) pnormalize_unique: "poly p = poly q \<Longrightarrow> pnormalize p = pnormalize q"
 51.1319 +proof (induct q arbitrary: p)
 51.1320 +  case Nil
 51.1321 +  thus ?case by (simp only: poly_zero lemma_degree_zero) simp
 51.1322 +next
 51.1323 +  case (Cons c cs p)
 51.1324 +  thus ?case
 51.1325 +  proof (induct p)
 51.1326 +    case Nil
 51.1327 +    hence "poly [] = poly (c#cs)" by blast
 51.1328 +    then have "poly (c#cs) = poly [] " by simp
 51.1329 +    thus ?case by (simp only: poly_zero lemma_degree_zero) simp
 51.1330 +  next
 51.1331 +    case (Cons d ds)
 51.1332 +    hence eq: "poly (d # ds) = poly (c # cs)" by blast
 51.1333 +    hence eq': "\<And>x. poly (d # ds) x = poly (c # cs) x" by simp
 51.1334 +    hence "poly (d # ds) 0 = poly (c # cs) 0" by blast
 51.1335 +    hence dc: "d = c" by auto
 51.1336 +    with eq have "poly ds = poly cs"
 51.1337 +      unfolding  poly_Cons_eq by simp
 51.1338 +    with Cons.prems have "pnormalize ds = pnormalize cs" by blast
 51.1339 +    with dc show ?case by simp
 51.1340 +  qed
 51.1341 +qed
 51.1342 +
 51.1343 +lemma (in idom_char_0) degree_unique:
 51.1344 +  assumes pq: "poly p = poly q"
 51.1345 +  shows "degree p = degree q"
 51.1346 +  using pnormalize_unique[OF pq] unfolding degree_def by simp
 51.1347 +
 51.1348 +lemma (in semiring_0) pnormalize_length:
 51.1349 +  "length (pnormalize p) \<le> length p" by (induct p) auto
 51.1350 +
 51.1351 +lemma (in semiring_0) last_linear_mul_lemma:
 51.1352 +  "last ((a %* p) +++ (x#(b %* p))) = (if p = [] then x else b * last p)"
 51.1353 +  apply (induct p arbitrary: a x b)
 51.1354    apply auto
 51.1355 -  apply (cases "c = 0")
 51.1356 +  apply (subgoal_tac "padd (cmult aa p) (times b a # cmult b p) \<noteq> []")
 51.1357 +  apply simp
 51.1358 +  apply (induct_tac p)
 51.1359    apply auto
 51.1360    done
 51.1361  
 51.1362 -lemma pnormal_last_nonzero: "pnormal p \<Longrightarrow> last p \<noteq> 0"
 51.1363 -  apply (induct p)
 51.1364 -  apply (auto simp add: pnormal_def)
 51.1365 -  apply (case_tac "pnormalize p = []")
 51.1366 -  apply auto
 51.1367 -  apply (case_tac "a = 0")
 51.1368 -  apply auto
 51.1369 -  done
 51.1370 +lemma (in semiring_1) last_linear_mul:
 51.1371 +  assumes p: "p \<noteq> []"
 51.1372 +  shows "last ([a,1] *** p) = last p"
 51.1373 +proof -
 51.1374 +  from p obtain c cs where cs: "p = c#cs" by (cases p) auto
 51.1375 +  from cs have eq: "[a,1] *** p = (a %* (c#cs)) +++ (0#(1 %* (c#cs)))"
 51.1376 +    by (simp add: poly_cmult_distr)
 51.1377 +  show ?thesis using cs
 51.1378 +    unfolding eq last_linear_mul_lemma by simp
 51.1379 +qed
 51.1380  
 51.1381 -lemma  pnormal_length: "pnormal p \<Longrightarrow> 0 < length p"
 51.1382 -  unfolding pnormal_def length_greater_0_conv by blast
 51.1383 +lemma (in semiring_0) pnormalize_eq: "last p \<noteq> 0 \<Longrightarrow> pnormalize p = p"
 51.1384 +  by (induct p) (auto split: split_if_asm)
 51.1385  
 51.1386 -lemma pnormal_last_length: "0 < length p \<Longrightarrow> last p \<noteq> 0 \<Longrightarrow> pnormal p"
 51.1387 -  apply (induct p)
 51.1388 -  apply auto
 51.1389 -  apply (case_tac "p = []")
 51.1390 -  apply auto
 51.1391 -  apply (simp add: pnormal_def)
 51.1392 -  apply (rule pnormal_cons)
 51.1393 -  apply auto
 51.1394 -  done
 51.1395 +lemma (in semiring_0) last_pnormalize: "pnormalize p \<noteq> [] \<Longrightarrow> last (pnormalize p) \<noteq> 0"
 51.1396 +  by (induct p) auto
 51.1397  
 51.1398 -lemma pnormal_id: "pnormal p \<longleftrightarrow> 0 < length p \<and> last p \<noteq> 0"
 51.1399 -  using pnormal_last_length pnormal_length pnormal_last_nonzero by blast
 51.1400 +lemma (in semiring_0) pnormal_degree: "last p \<noteq> 0 \<Longrightarrow> degree p = length p - 1"
 51.1401 +  using pnormalize_eq[of p] unfolding degree_def by simp
 51.1402 +
 51.1403 +lemma (in semiring_0) poly_Nil_ext: "poly [] = (\<lambda>x. 0)"
 51.1404 +  by (rule ext) simp
 51.1405 +
 51.1406 +lemma (in idom_char_0) linear_mul_degree:
 51.1407 +  assumes p: "poly p \<noteq> poly []"
 51.1408 +  shows "degree ([a,1] *** p) = degree p + 1"
 51.1409 +proof -
 51.1410 +  from p have pnz: "pnormalize p \<noteq> []"
 51.1411 +    unfolding poly_zero lemma_degree_zero .
 51.1412 +
 51.1413 +  from last_linear_mul[OF pnz, of a] last_pnormalize[OF pnz]
 51.1414 +  have l0: "last ([a, 1] *** pnormalize p) \<noteq> 0" by simp
 51.1415 +  from last_pnormalize[OF pnz] last_linear_mul[OF pnz, of a]
 51.1416 +    pnormal_degree[OF l0] pnormal_degree[OF last_pnormalize[OF pnz]] pnz
 51.1417 +
 51.1418 +  have th: "degree ([a,1] *** pnormalize p) = degree (pnormalize p) + 1"
 51.1419 +    by simp
 51.1420 +
 51.1421 +  have eqs: "poly ([a,1] *** pnormalize p) = poly ([a,1] *** p)"
 51.1422 +    by (rule ext) (simp add: poly_mult poly_add poly_cmult)
 51.1423 +  from degree_unique[OF eqs] th
 51.1424 +  show ?thesis by (simp add: degree_unique[OF poly_normalize])
 51.1425 +qed
 51.1426 +
 51.1427 +lemma (in idom_char_0) linear_pow_mul_degree:
 51.1428 +  "degree([a,1] %^n *** p) = (if poly p = poly [] then 0 else degree p + n)"
 51.1429 +proof (induct n arbitrary: a p)
 51.1430 +  case (0 a p)
 51.1431 +  show ?case
 51.1432 +  proof (cases "poly p = poly []")
 51.1433 +    case True
 51.1434 +    then show ?thesis
 51.1435 +      using degree_unique[OF True] by (simp add: degree_def)
 51.1436 +  next
 51.1437 +    case False
 51.1438 +    then show ?thesis by (auto simp add: poly_Nil_ext)
 51.1439 +  qed
 51.1440 +next
 51.1441 +  case (Suc n a p)
 51.1442 +  have eq: "poly ([a,1] %^(Suc n) *** p) = poly ([a,1] %^ n *** ([a,1] *** p))"
 51.1443 +    apply (rule ext)
 51.1444 +    apply (simp add: poly_mult poly_add poly_cmult)
 51.1445 +    apply (simp add: mult_ac add_ac distrib_left)
 51.1446 +    done
 51.1447 +  note deq = degree_unique[OF eq]
 51.1448 +  show ?case
 51.1449 +  proof (cases "poly p = poly []")
 51.1450 +    case True
 51.1451 +    with eq have eq': "poly ([a,1] %^(Suc n) *** p) = poly []"
 51.1452 +      apply -
 51.1453 +      apply (rule ext)
 51.1454 +      apply (simp add: poly_mult poly_cmult poly_add)
 51.1455 +      done
 51.1456 +    from degree_unique[OF eq'] True show ?thesis
 51.1457 +      by (simp add: degree_def)
 51.1458 +  next
 51.1459 +    case False
 51.1460 +    then have ap: "poly ([a,1] *** p) \<noteq> poly []"
 51.1461 +      using poly_mult_not_eq_poly_Nil unfolding poly_entire by auto
 51.1462 +    have eq: "poly ([a,1] %^(Suc n) *** p) = poly ([a,1]%^n *** ([a,1] *** p))"
 51.1463 +      by (rule ext, simp add: poly_mult poly_add poly_exp poly_cmult algebra_simps)
 51.1464 +    from ap have ap': "(poly ([a,1] *** p) = poly []) = False"
 51.1465 +      by blast
 51.1466 +    have th0: "degree ([a,1]%^n *** ([a,1] *** p)) = degree ([a,1] *** p) + n"
 51.1467 +      apply (simp only: Suc.hyps[of a "pmult [a,one] p"] ap')
 51.1468 +      apply simp
 51.1469 +      done
 51.1470 +    from degree_unique[OF eq] ap False th0 linear_mul_degree[OF False, of a]
 51.1471 +    show ?thesis by (auto simp del: poly.simps)
 51.1472 +  qed
 51.1473 +qed
 51.1474 +
 51.1475 +lemma (in idom_char_0) order_degree:
 51.1476 +  assumes p0: "poly p \<noteq> poly []"
 51.1477 +  shows "order a p \<le> degree p"
 51.1478 +proof -
 51.1479 +  from order2[OF p0, unfolded divides_def]
 51.1480 +  obtain q where q: "poly p = poly ([- a, 1]%^ (order a p) *** q)" by blast
 51.1481 +  {
 51.1482 +    assume "poly q = poly []"
 51.1483 +    with q p0 have False by (simp add: poly_mult poly_entire)
 51.1484 +  }
 51.1485 +  with degree_unique[OF q, unfolded linear_pow_mul_degree] show ?thesis
 51.1486 +    by auto
 51.1487 +qed
 51.1488  
 51.1489  text{*Tidier versions of finiteness of roots.*}
 51.1490  
 51.1491 -lemma poly_roots_finite_set:
 51.1492 -  "poly p \<noteq> poly [] \<Longrightarrow> finite {x::'a::{idom,ring_char_0}. poly p x = 0}"
 51.1493 +lemma (in idom_char_0) poly_roots_finite_set:
 51.1494 +  "poly p \<noteq> poly [] \<Longrightarrow> finite {x. poly p x = 0}"
 51.1495    unfolding poly_roots_finite .
 51.1496  
 51.1497  text{*bound for polynomial.*}
 51.1498  
 51.1499 -lemma poly_mono: "abs x \<le> k \<Longrightarrow> abs (poly p (x::'a::{linordered_idom})) \<le> poly (map abs p) k"
 51.1500 +lemma poly_mono: "abs(x) \<le> k \<Longrightarrow> abs(poly p (x::'a::{linordered_idom})) \<le> poly (map abs p) k"
 51.1501    apply (induct p)
 51.1502    apply auto
 51.1503    apply (rule_tac y = "abs a + abs (x * poly p x)" in order_trans)
 51.1504 @@ -802,7 +1047,6 @@
 51.1505    apply (auto intro!: mult_mono simp add: abs_mult)
 51.1506    done
 51.1507  
 51.1508 -lemma poly_Sing: "poly [c] x = c"
 51.1509 -  by simp
 51.1510 +lemma (in semiring_0) poly_Sing: "poly [c] x = c" by simp
 51.1511  
 51.1512  end
    52.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    52.2 +++ b/src/HOL/Decision_Procs/Rat_Pair.thy	Mon Nov 11 17:44:21 2013 +0100
    52.3 @@ -0,0 +1,522 @@
    52.4 +(*  Title:      HOL/Decision_Procs/Rat_Pair.thy
    52.5 +    Author:     Amine Chaieb
    52.6 +*)
    52.7 +
    52.8 +header {* Rational numbers as pairs *}
    52.9 +
   52.10 +theory Rat_Pair
   52.11 +imports Complex_Main
   52.12 +begin
   52.13 +
   52.14 +type_synonym Num = "int \<times> int"
   52.15 +
   52.16 +abbreviation Num0_syn :: Num  ("0\<^sub>N")
   52.17 +  where "0\<^sub>N \<equiv> (0, 0)"
   52.18 +
   52.19 +abbreviation Numi_syn :: "int \<Rightarrow> Num"  ("'((_)')\<^sub>N")
   52.20 +  where "(i)\<^sub>N \<equiv> (i, 1)"
   52.21 +
   52.22 +definition isnormNum :: "Num \<Rightarrow> bool" where
   52.23 +  "isnormNum = (\<lambda>(a,b). (if a = 0 then b = 0 else b > 0 \<and> gcd a b = 1))"
   52.24 +
   52.25 +definition normNum :: "Num \<Rightarrow> Num" where
   52.26 +  "normNum = (\<lambda>(a,b).
   52.27 +    (if a=0 \<or> b = 0 then (0,0) else
   52.28 +      (let g = gcd a b
   52.29 +       in if b > 0 then (a div g, b div g) else (- (a div g), - (b div g)))))"
   52.30 +
   52.31 +declare gcd_dvd1_int[presburger] gcd_dvd2_int[presburger]
   52.32 +
   52.33 +lemma normNum_isnormNum [simp]: "isnormNum (normNum x)"
   52.34 +proof -
   52.35 +  obtain a b where x: "x = (a, b)" by (cases x)
   52.36 +  { assume "a=0 \<or> b = 0" hence ?thesis by (simp add: x normNum_def isnormNum_def) }
   52.37 +  moreover
   52.38 +  { assume anz: "a \<noteq> 0" and bnz: "b \<noteq> 0"
   52.39 +    let ?g = "gcd a b"
   52.40 +    let ?a' = "a div ?g"
   52.41 +    let ?b' = "b div ?g"
   52.42 +    let ?g' = "gcd ?a' ?b'"
   52.43 +    from anz bnz have "?g \<noteq> 0" by simp  with gcd_ge_0_int[of a b]
   52.44 +    have gpos: "?g > 0" by arith
   52.45 +    have gdvd: "?g dvd a" "?g dvd b" by arith+
   52.46 +    from dvd_mult_div_cancel[OF gdvd(1)] dvd_mult_div_cancel[OF gdvd(2)] anz bnz
   52.47 +    have nz': "?a' \<noteq> 0" "?b' \<noteq> 0" by - (rule notI, simp)+
   52.48 +    from anz bnz have stupid: "a \<noteq> 0 \<or> b \<noteq> 0" by arith
   52.49 +    from div_gcd_coprime_int[OF stupid] have gp1: "?g' = 1" .
   52.50 +    from bnz have "b < 0 \<or> b > 0" by arith
   52.51 +    moreover
   52.52 +    { assume b: "b > 0"
   52.53 +      from b have "?b' \<ge> 0"
   52.54 +        by (presburger add: pos_imp_zdiv_nonneg_iff[OF gpos])
   52.55 +      with nz' have b': "?b' > 0" by arith
   52.56 +      from b b' anz bnz nz' gp1 have ?thesis
   52.57 +        by (simp add: x isnormNum_def normNum_def Let_def split_def) }
   52.58 +    moreover {
   52.59 +      assume b: "b < 0"
   52.60 +      { assume b': "?b' \<ge> 0"
   52.61 +        from gpos have th: "?g \<ge> 0" by arith
   52.62 +        from mult_nonneg_nonneg[OF th b'] dvd_mult_div_cancel[OF gdvd(2)]
   52.63 +        have False using b by arith }
   52.64 +      hence b': "?b' < 0" by (presburger add: linorder_not_le[symmetric])
   52.65 +      from anz bnz nz' b b' gp1 have ?thesis
   52.66 +        by (simp add: x isnormNum_def normNum_def Let_def split_def) }
   52.67 +    ultimately have ?thesis by blast
   52.68 +  }
   52.69 +  ultimately show ?thesis by blast
   52.70 +qed
   52.71 +
   52.72 +text {* Arithmetic over Num *}
   52.73 +
   52.74 +definition Nadd :: "Num \<Rightarrow> Num \<Rightarrow> Num"  (infixl "+\<^sub>N" 60) where
   52.75 +  "Nadd = (\<lambda>(a,b) (a',b'). if a = 0 \<or> b = 0 then normNum(a',b')
   52.76 +    else if a'=0 \<or> b' = 0 then normNum(a,b)
   52.77 +    else normNum(a*b' + b*a', b*b'))"
   52.78 +
   52.79 +definition Nmul :: "Num \<Rightarrow> Num \<Rightarrow> Num"  (infixl "*\<^sub>N" 60) where
   52.80 +  "Nmul = (\<lambda>(a,b) (a',b'). let g = gcd (a*a') (b*b')
   52.81 +    in (a*a' div g, b*b' div g))"
   52.82 +
   52.83 +definition Nneg :: "Num \<Rightarrow> Num" ("~\<^sub>N")
   52.84 +  where "Nneg \<equiv> (\<lambda>(a,b). (-a,b))"
   52.85 +
   52.86 +definition Nsub :: "Num \<Rightarrow> Num \<Rightarrow> Num"  (infixl "-\<^sub>N" 60)
   52.87 +  where "Nsub = (\<lambda>a b. a +\<^sub>N ~\<^sub>N b)"
   52.88 +
   52.89 +definition Ninv :: "Num \<Rightarrow> Num"
   52.90 +  where "Ninv = (\<lambda>(a,b). if a < 0 then (-b, \<bar>a\<bar>) else (b,a))"
   52.91 +
   52.92 +definition Ndiv :: "Num \<Rightarrow> Num \<Rightarrow> Num"  (infixl "\<div>\<^sub>N" 60)
   52.93 +  where "Ndiv = (\<lambda>a b. a *\<^sub>N Ninv b)"
   52.94 +
   52.95 +lemma Nneg_normN[simp]: "isnormNum x \<Longrightarrow> isnormNum (~\<^sub>N x)"
   52.96 +  by (simp add: isnormNum_def Nneg_def split_def)
   52.97 +
   52.98 +lemma Nadd_normN[simp]: "isnormNum (x +\<^sub>N y)"
   52.99 +  by (simp add: Nadd_def split_def)
  52.100 +
  52.101 +lemma Nsub_normN[simp]: "\<lbrakk> isnormNum y\<rbrakk> \<Longrightarrow> isnormNum (x -\<^sub>N y)"
  52.102 +  by (simp add: Nsub_def split_def)
  52.103 +
  52.104 +lemma Nmul_normN[simp]:
  52.105 +  assumes xn: "isnormNum x" and yn: "isnormNum y"
  52.106 +  shows "isnormNum (x *\<^sub>N y)"
  52.107 +proof -
  52.108 +  obtain a b where x: "x = (a, b)" by (cases x)
  52.109 +  obtain a' b' where y: "y = (a', b')" by (cases y)
  52.110 +  { assume "a = 0"
  52.111 +    hence ?thesis using xn x y
  52.112 +      by (simp add: isnormNum_def Let_def Nmul_def split_def) }
  52.113 +  moreover
  52.114 +  { assume "a' = 0"
  52.115 +    hence ?thesis using yn x y
  52.116 +      by (simp add: isnormNum_def Let_def Nmul_def split_def) }
  52.117 +  moreover
  52.118 +  { assume a: "a \<noteq>0" and a': "a'\<noteq>0"
  52.119 +    hence bp: "b > 0" "b' > 0" using xn yn x y by (simp_all add: isnormNum_def)
  52.120 +    from mult_pos_pos[OF bp] have "x *\<^sub>N y = normNum (a * a', b * b')"
  52.121 +      using x y a a' bp by (simp add: Nmul_def Let_def split_def normNum_def)
  52.122 +    hence ?thesis by simp }
  52.123 +  ultimately show ?thesis by blast
  52.124 +qed
  52.125 +
  52.126 +lemma Ninv_normN[simp]: "isnormNum x \<Longrightarrow> isnormNum (Ninv x)"
  52.127 +  by (simp add: Ninv_def isnormNum_def split_def)
  52.128 +    (cases "fst x = 0", auto simp add: gcd_commute_int)
  52.129 +
  52.130 +lemma isnormNum_int[simp]:
  52.131 +  "isnormNum 0\<^sub>N" "isnormNum ((1::int)\<^sub>N)" "i \<noteq> 0 \<Longrightarrow> isnormNum (i)\<^sub>N"
  52.132 +  by (simp_all add: isnormNum_def)
  52.133 +
  52.134 +
  52.135 +text {* Relations over Num *}
  52.136 +
  52.137 +definition Nlt0:: "Num \<Rightarrow> bool"  ("0>\<^sub>N")
  52.138 +  where "Nlt0 = (\<lambda>(a,b). a < 0)"
  52.139 +
  52.140 +definition Nle0:: "Num \<Rightarrow> bool"  ("0\<ge>\<^sub>N")
  52.141 +  where "Nle0 = (\<lambda>(a,b). a \<le> 0)"
  52.142 +
  52.143 +definition Ngt0:: "Num \<Rightarrow> bool"  ("0<\<^sub>N")
  52.144 +  where "Ngt0 = (\<lambda>(a,b). a > 0)"
  52.145 +
  52.146 +definition Nge0:: "Num \<Rightarrow> bool"  ("0\<le>\<^sub>N")
  52.147 +  where "Nge0 = (\<lambda>(a,b). a \<ge> 0)"
  52.148 +
  52.149 +definition Nlt :: "Num \<Rightarrow> Num \<Rightarrow> bool"  (infix "<\<^sub>N" 55)
  52.150 +  where "Nlt = (\<lambda>a b. 0>\<^sub>N (a -\<^sub>N b))"
  52.151 +
  52.152 +definition Nle :: "Num \<Rightarrow> Num \<Rightarrow> bool"  (infix "\<le>\<^sub>N" 55)
  52.153 +  where "Nle = (\<lambda>a b. 0\<ge>\<^sub>N (a -\<^sub>N b))"
  52.154 +
  52.155 +definition "INum = (\<lambda>(a,b). of_int a / of_int b)"
  52.156 +
  52.157 +lemma INum_int [simp]: "INum (i)\<^sub>N = ((of_int i) ::'a::field)" "INum 0\<^sub>N = (0::'a::field)"
  52.158 +  by (simp_all add: INum_def)
  52.159 +
  52.160 +lemma isnormNum_unique[simp]:
  52.161 +  assumes na: "isnormNum x" and nb: "isnormNum y"
  52.162 +  shows "((INum x ::'a::{field_char_0, field_inverse_zero}) = INum y) = (x = y)" (is "?lhs = ?rhs")
  52.163 +proof
  52.164 +  obtain a b where x: "x = (a, b)" by (cases x)
  52.165 +  obtain a' b' where y: "y = (a', b')" by (cases y)
  52.166 +  assume H: ?lhs
  52.167 +  { assume "a = 0 \<or> b = 0 \<or> a' = 0 \<or> b' = 0"
  52.168 +    hence ?rhs using na nb H
  52.169 +      by (simp add: x y INum_def split_def isnormNum_def split: split_if_asm) }
  52.170 +  moreover
  52.171 +  { assume az: "a \<noteq> 0" and bz: "b \<noteq> 0" and a'z: "a'\<noteq>0" and b'z: "b'\<noteq>0"
  52.172 +    from az bz a'z b'z na nb have pos: "b > 0" "b' > 0" by (simp_all add: x y isnormNum_def)
  52.173 +    from H bz b'z have eq: "a * b' = a'*b"
  52.174 +      by (simp add: x y INum_def eq_divide_eq divide_eq_eq of_int_mult[symmetric] del: of_int_mult)
  52.175 +    from az a'z na nb have gcd1: "gcd a b = 1" "gcd b a = 1" "gcd a' b' = 1" "gcd b' a' = 1"
  52.176 +      by (simp_all add: x y isnormNum_def add: gcd_commute_int)
  52.177 +    from eq have raw_dvd: "a dvd a' * b" "b dvd b' * a" "a' dvd a * b'" "b' dvd b * a'"
  52.178 +      apply -
  52.179 +      apply algebra
  52.180 +      apply algebra
  52.181 +      apply simp
  52.182 +      apply algebra
  52.183 +      done
  52.184 +    from zdvd_antisym_abs[OF coprime_dvd_mult_int[OF gcd1(2) raw_dvd(2)]
  52.185 +        coprime_dvd_mult_int[OF gcd1(4) raw_dvd(4)]]
  52.186 +      have eq1: "b = b'" using pos by arith
  52.187 +      with eq have "a = a'" using pos by simp
  52.188 +      with eq1 have ?rhs by (simp add: x y) }
  52.189 +  ultimately show ?rhs by blast
  52.190 +next
  52.191 +  assume ?rhs thus ?lhs by simp
  52.192 +qed
  52.193 +
  52.194 +
  52.195 +lemma isnormNum0[simp]:
  52.196 +    "isnormNum x \<Longrightarrow> (INum x = (0::'a::{field_char_0, field_inverse_zero})) = (x = 0\<^sub>N)"
  52.197 +  unfolding INum_int(2)[symmetric]
  52.198 +  by (rule isnormNum_unique) simp_all
  52.199 +
  52.200 +lemma of_int_div_aux: "d ~= 0 ==> ((of_int x)::'a::field_char_0) / (of_int d) =
  52.201 +    of_int (x div d) + (of_int (x mod d)) / ((of_int d)::'a)"
  52.202 +proof -
  52.203 +  assume "d ~= 0"
  52.204 +  let ?t = "of_int (x div d) * ((of_int d)::'a) + of_int(x mod d)"
  52.205 +  let ?f = "\<lambda>x. x / of_int d"
  52.206 +  have "x = (x div d) * d + x mod d"
  52.207 +    by auto
  52.208 +  then have eq: "of_int x = ?t"
  52.209 +    by (simp only: of_int_mult[symmetric] of_int_add [symmetric])
  52.210 +  then have "of_int x / of_int d = ?t / of_int d"
  52.211 +    using cong[OF refl[of ?f] eq] by simp
  52.212 +  then show ?thesis by (simp add: add_divide_distrib algebra_simps `d ~= 0`)
  52.213 +qed
  52.214 +
  52.215 +lemma of_int_div: "(d::int) ~= 0 ==> d dvd n ==>
  52.216 +    (of_int(n div d)::'a::field_char_0) = of_int n / of_int d"
  52.217 +  apply (frule of_int_div_aux [of d n, where ?'a = 'a])
  52.218 +  apply simp
  52.219 +  apply (simp add: dvd_eq_mod_eq_0)
  52.220 +  done
  52.221 +
  52.222 +
  52.223 +lemma normNum[simp]: "INum (normNum x) = (INum x :: 'a::{field_char_0, field_inverse_zero})"
  52.224 +proof -
  52.225 +  obtain a b where x: "x = (a, b)" by (cases x)
  52.226 +  { assume "a = 0 \<or> b = 0"
  52.227 +    hence ?thesis by (simp add: x INum_def normNum_def split_def Let_def) }
  52.228 +  moreover
  52.229 +  { assume a: "a \<noteq> 0" and b: "b \<noteq> 0"
  52.230 +    let ?g = "gcd a b"
  52.231 +    from a b have g: "?g \<noteq> 0"by simp
  52.232 +    from of_int_div[OF g, where ?'a = 'a]
  52.233 +    have ?thesis by (auto simp add: x INum_def normNum_def split_def Let_def) }
  52.234 +  ultimately show ?thesis by blast
  52.235 +qed
  52.236 +
  52.237 +lemma INum_normNum_iff:
  52.238 +  "(INum x ::'a::{field_char_0, field_inverse_zero}) = INum y \<longleftrightarrow> normNum x = normNum y"
  52.239 +  (is "?lhs = ?rhs")
  52.240 +proof -
  52.241 +  have "normNum x = normNum y \<longleftrightarrow> (INum (normNum x) :: 'a) = INum (normNum y)"
  52.242 +    by (simp del: normNum)
  52.243 +  also have "\<dots> = ?lhs" by simp
  52.244 +  finally show ?thesis by simp
  52.245 +qed
  52.246 +
  52.247 +lemma Nadd[simp]: "INum (x +\<^sub>N y) = INum x + (INum y :: 'a :: {field_char_0, field_inverse_zero})"
  52.248 +proof -
  52.249 +  let ?z = "0:: 'a"
  52.250 +  obtain a b where x: "x = (a, b)" by (cases x)
  52.251 +  obtain a' b' where y: "y = (a', b')" by (cases y)
  52.252 +  { assume "a=0 \<or> a'= 0 \<or> b =0 \<or> b' = 0"
  52.253 +    hence ?thesis
  52.254 +      apply (cases "a=0", simp_all add: x y Nadd_def)
  52.255 +      apply (cases "b= 0", simp_all add: INum_def)
  52.256 +       apply (cases "a'= 0", simp_all)
  52.257 +       apply (cases "b'= 0", simp_all)
  52.258 +       done }
  52.259 +  moreover
  52.260 +  { assume aa': "a \<noteq> 0" "a'\<noteq> 0" and bb': "b \<noteq> 0" "b' \<noteq> 0"
  52.261 +    { assume z: "a * b' + b * a' = 0"
  52.262 +      hence "of_int (a*b' + b*a') / (of_int b* of_int b') = ?z" by simp
  52.263 +      hence "of_int b' * of_int a / (of_int b * of_int b') +
  52.264 +          of_int b * of_int a' / (of_int b * of_int b') = ?z"
  52.265 +        by (simp add:add_divide_distrib)
  52.266 +      hence th: "of_int a / of_int b + of_int a' / of_int b' = ?z" using bb' aa'
  52.267 +        by simp
  52.268 +      from z aa' bb' have ?thesis
  52.269 +        by (simp add: x y th Nadd_def normNum_def INum_def split_def) }
  52.270 +    moreover {
  52.271 +      assume z: "a * b' + b * a' \<noteq> 0"
  52.272 +      let ?g = "gcd (a * b' + b * a') (b * b')"
  52.273 +      have gz: "?g \<noteq> 0" using z by simp
  52.274 +      have ?thesis using aa' bb' z gz
  52.275 +        of_int_div[where ?'a = 'a, OF gz gcd_dvd1_int[where x="a * b' + b * a'" and y="b*b'"]]
  52.276 +        of_int_div[where ?'a = 'a, OF gz gcd_dvd2_int[where x="a * b' + b * a'" and y="b*b'"]]
  52.277 +        by (simp add: x y Nadd_def INum_def normNum_def Let_def) (simp add: field_simps)
  52.278 +    }
  52.279 +    ultimately have ?thesis using aa' bb'
  52.280 +      by (simp add: x y Nadd_def INum_def normNum_def Let_def) }
  52.281 +  ultimately show ?thesis by blast
  52.282 +qed
  52.283 +
  52.284 +lemma Nmul[simp]: "INum (x *\<^sub>N y) = INum x * (INum y:: 'a :: {field_char_0, field_inverse_zero})"
  52.285 +proof -
  52.286 +  let ?z = "0::'a"
  52.287 +  obtain a b where x: "x = (a, b)" by (cases x)
  52.288 +  obtain a' b' where y: "y = (a', b')" by (cases y)
  52.289 +  { assume "a=0 \<or> a'= 0 \<or> b = 0 \<or> b' = 0"
  52.290 +    hence ?thesis
  52.291 +      apply (cases "a=0", simp_all add: x y Nmul_def INum_def Let_def)
  52.292 +      apply (cases "b=0", simp_all)
  52.293 +      apply (cases "a'=0", simp_all)
  52.294 +      done }
  52.295 +  moreover
  52.296 +  { assume z: "a \<noteq> 0" "a' \<noteq> 0" "b \<noteq> 0" "b' \<noteq> 0"
  52.297 +    let ?g="gcd (a*a') (b*b')"
  52.298 +    have gz: "?g \<noteq> 0" using z by simp
  52.299 +    from z of_int_div[where ?'a = 'a, OF gz gcd_dvd1_int[where x="a*a'" and y="b*b'"]]
  52.300 +      of_int_div[where ?'a = 'a , OF gz gcd_dvd2_int[where x="a*a'" and y="b*b'"]]
  52.301 +    have ?thesis by (simp add: Nmul_def x y Let_def INum_def) }
  52.302 +  ultimately show ?thesis by blast
  52.303 +qed
  52.304 +
  52.305 +lemma Nneg[simp]: "INum (~\<^sub>N x) = - (INum x ::'a:: field)"
  52.306 +  by (simp add: Nneg_def split_def INum_def)
  52.307 +
  52.308 +lemma Nsub[simp]: "INum (x -\<^sub>N y) = INum x - (INum y:: 'a :: {field_char_0, field_inverse_zero})"
  52.309 +  by (simp add: Nsub_def split_def)
  52.310 +
  52.311 +lemma Ninv[simp]: "INum (Ninv x) = (1::'a :: field_inverse_zero) / (INum x)"
  52.312 +  by (simp add: Ninv_def INum_def split_def)
  52.313 +
  52.314 +lemma Ndiv[simp]: "INum (x \<div>\<^sub>N y) = INum x / (INum y ::'a :: {field_char_0, field_inverse_zero})"
  52.315 +  by (simp add: Ndiv_def)
  52.316 +
  52.317 +lemma Nlt0_iff[simp]:
  52.318 +  assumes nx: "isnormNum x"
  52.319 +  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero})< 0) = 0>\<^sub>N x"
  52.320 +proof -
  52.321 +  obtain a b where x: "x = (a, b)" by (cases x)
  52.322 +  { assume "a = 0" hence ?thesis by (simp add: x Nlt0_def INum_def) }
  52.323 +  moreover
  52.324 +  { assume a: "a \<noteq> 0" hence b: "(of_int b::'a) > 0"
  52.325 +      using nx by (simp add: x isnormNum_def)
  52.326 +    from pos_divide_less_eq[OF b, where b="of_int a" and a="0::'a"]
  52.327 +    have ?thesis by (simp add: x Nlt0_def INum_def) }
  52.328 +  ultimately show ?thesis by blast
  52.329 +qed
  52.330 +
  52.331 +lemma Nle0_iff[simp]:
  52.332 +  assumes nx: "isnormNum x"
  52.333 +  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero}) \<le> 0) = 0\<ge>\<^sub>N x"
  52.334 +proof -
  52.335 +  obtain a b where x: "x = (a, b)" by (cases x)
  52.336 +  { assume "a = 0" hence ?thesis by (simp add: x Nle0_def INum_def) }
  52.337 +  moreover
  52.338 +  { assume a: "a \<noteq> 0" hence b: "(of_int b :: 'a) > 0"
  52.339 +      using nx by (simp add: x isnormNum_def)
  52.340 +    from pos_divide_le_eq[OF b, where b="of_int a" and a="0::'a"]
  52.341 +    have ?thesis by (simp add: x Nle0_def INum_def) }
  52.342 +  ultimately show ?thesis by blast
  52.343 +qed
  52.344 +
  52.345 +lemma Ngt0_iff[simp]:
  52.346 +  assumes nx: "isnormNum x"
  52.347 +  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero})> 0) = 0<\<^sub>N x"
  52.348 +proof -
  52.349 +  obtain a b where x: "x = (a, b)" by (cases x)
  52.350 +  { assume "a = 0" hence ?thesis by (simp add: x Ngt0_def INum_def) }
  52.351 +  moreover
  52.352 +  { assume a: "a \<noteq> 0" hence b: "(of_int b::'a) > 0" using nx
  52.353 +      by (simp add: x isnormNum_def)
  52.354 +    from pos_less_divide_eq[OF b, where b="of_int a" and a="0::'a"]
  52.355 +    have ?thesis by (simp add: x Ngt0_def INum_def) }
  52.356 +  ultimately show ?thesis by blast
  52.357 +qed
  52.358 +
  52.359 +lemma Nge0_iff[simp]:
  52.360 +  assumes nx: "isnormNum x"
  52.361 +  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero}) \<ge> 0) = 0\<le>\<^sub>N x"
  52.362 +proof -
  52.363 +  obtain a b where x: "x = (a, b)" by (cases x)
  52.364 +  { assume "a = 0" hence ?thesis by (simp add: x Nge0_def INum_def) }
  52.365 +  moreover
  52.366 +  { assume "a \<noteq> 0" hence b: "(of_int b::'a) > 0" using nx
  52.367 +      by (simp add: x isnormNum_def)
  52.368 +    from pos_le_divide_eq[OF b, where b="of_int a" and a="0::'a"]
  52.369 +    have ?thesis by (simp add: x Nge0_def INum_def) }
  52.370 +  ultimately show ?thesis by blast
  52.371 +qed
  52.372 +
  52.373 +lemma Nlt_iff[simp]:
  52.374 +  assumes nx: "isnormNum x" and ny: "isnormNum y"
  52.375 +  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero}) < INum y) = (x <\<^sub>N y)"
  52.376 +proof -
  52.377 +  let ?z = "0::'a"
  52.378 +  have "((INum x ::'a) < INum y) = (INum (x -\<^sub>N y) < ?z)"
  52.379 +    using nx ny by simp
  52.380 +  also have "\<dots> = (0>\<^sub>N (x -\<^sub>N y))"
  52.381 +    using Nlt0_iff[OF Nsub_normN[OF ny]] by simp
  52.382 +  finally show ?thesis by (simp add: Nlt_def)
  52.383 +qed
  52.384 +
  52.385 +lemma Nle_iff[simp]:
  52.386 +  assumes nx: "isnormNum x" and ny: "isnormNum y"
  52.387 +  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero})\<le> INum y) = (x \<le>\<^sub>N y)"
  52.388 +proof -
  52.389 +  have "((INum x ::'a) \<le> INum y) = (INum (x -\<^sub>N y) \<le> (0::'a))"
  52.390 +    using nx ny by simp
  52.391 +  also have "\<dots> = (0\<ge>\<^sub>N (x -\<^sub>N y))"
  52.392 +    using Nle0_iff[OF Nsub_normN[OF ny]] by simp
  52.393 +  finally show ?thesis by (simp add: Nle_def)
  52.394 +qed
  52.395 +
  52.396 +lemma Nadd_commute:
  52.397 +  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  52.398 +  shows "x +\<^sub>N y = y +\<^sub>N x"
  52.399 +proof -
  52.400 +  have n: "isnormNum (x +\<^sub>N y)" "isnormNum (y +\<^sub>N x)" by simp_all
  52.401 +  have "(INum (x +\<^sub>N y)::'a) = INum (y +\<^sub>N x)" by simp
  52.402 +  with isnormNum_unique[OF n] show ?thesis by simp
  52.403 +qed
  52.404 +
  52.405 +lemma [simp]:
  52.406 +  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  52.407 +  shows "(0, b) +\<^sub>N y = normNum y"
  52.408 +    and "(a, 0) +\<^sub>N y = normNum y"
  52.409 +    and "x +\<^sub>N (0, b) = normNum x"
  52.410 +    and "x +\<^sub>N (a, 0) = normNum x"
  52.411 +  apply (simp add: Nadd_def split_def)
  52.412 +  apply (simp add: Nadd_def split_def)
  52.413 +  apply (subst Nadd_commute, simp add: Nadd_def split_def)
  52.414 +  apply (subst Nadd_commute, simp add: Nadd_def split_def)
  52.415 +  done
  52.416 +
  52.417 +lemma normNum_nilpotent_aux[simp]:
  52.418 +  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  52.419 +  assumes nx: "isnormNum x"
  52.420 +  shows "normNum x = x"
  52.421 +proof -
  52.422 +  let ?a = "normNum x"
  52.423 +  have n: "isnormNum ?a" by simp
  52.424 +  have th: "INum ?a = (INum x ::'a)" by simp
  52.425 +  with isnormNum_unique[OF n nx] show ?thesis by simp
  52.426 +qed
  52.427 +
  52.428 +lemma normNum_nilpotent[simp]:
  52.429 +  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  52.430 +  shows "normNum (normNum x) = normNum x"
  52.431 +  by simp
  52.432 +
  52.433 +lemma normNum0[simp]: "normNum (0,b) = 0\<^sub>N" "normNum (a,0) = 0\<^sub>N"
  52.434 +  by (simp_all add: normNum_def)
  52.435 +
  52.436 +lemma normNum_Nadd:
  52.437 +  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  52.438 +  shows "normNum (x +\<^sub>N y) = x +\<^sub>N y" by simp
  52.439 +
  52.440 +lemma Nadd_normNum1[simp]:
  52.441 +  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  52.442 +  shows "normNum x +\<^sub>N y = x +\<^sub>N y"
  52.443 +proof -
  52.444 +  have n: "isnormNum (normNum x +\<^sub>N y)" "isnormNum (x +\<^sub>N y)" by simp_all
  52.445 +  have "INum (normNum x +\<^sub>N y) = INum x + (INum y :: 'a)" by simp
  52.446 +  also have "\<dots> = INum (x +\<^sub>N y)" by simp
  52.447 +  finally show ?thesis using isnormNum_unique[OF n] by simp
  52.448 +qed
  52.449 +
  52.450 +lemma Nadd_normNum2[simp]:
  52.451 +  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  52.452 +  shows "x +\<^sub>N normNum y = x +\<^sub>N y"
  52.453 +proof -
  52.454 +  have n: "isnormNum (x +\<^sub>N normNum y)" "isnormNum (x +\<^sub>N y)" by simp_all
  52.455 +  have "INum (x +\<^sub>N normNum y) = INum x + (INum y :: 'a)" by simp
  52.456 +  also have "\<dots> = INum (x +\<^sub>N y)" by simp
  52.457 +  finally show ?thesis using isnormNum_unique[OF n] by simp
  52.458 +qed
  52.459 +
  52.460 +lemma Nadd_assoc:
  52.461 +  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  52.462 +  shows "x +\<^sub>N y +\<^sub>N z = x +\<^sub>N (y +\<^sub>N z)"
  52.463 +proof -
  52.464 +  have n: "isnormNum (x +\<^sub>N y +\<^sub>N z)" "isnormNum (x +\<^sub>N (y +\<^sub>N z))" by simp_all
  52.465 +  have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" by simp
  52.466 +  with isnormNum_unique[OF n] show ?thesis by simp
  52.467 +qed
  52.468 +
  52.469 +lemma Nmul_commute: "isnormNum x \<Longrightarrow> isnormNum y \<Longrightarrow> x *\<^sub>N y = y *\<^sub>N x"
  52.470 +  by (simp add: Nmul_def split_def Let_def gcd_commute_int mult_commute)
  52.471 +
  52.472 +lemma Nmul_assoc:
  52.473 +  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  52.474 +  assumes nx: "isnormNum x" and ny: "isnormNum y" and nz: "isnormNum z"
  52.475 +  shows "x *\<^sub>N y *\<^sub>N z = x *\<^sub>N (y *\<^sub>N z)"
  52.476 +proof -
  52.477 +  from nx ny nz have n: "isnormNum (x *\<^sub>N y *\<^sub>N z)" "isnormNum (x *\<^sub>N (y *\<^sub>N z))"
  52.478 +    by simp_all
  52.479 +  have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" by simp
  52.480 +  with isnormNum_unique[OF n] show ?thesis by simp
  52.481 +qed
  52.482 +
  52.483 +lemma Nsub0:
  52.484 +  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  52.485 +  assumes x: "isnormNum x" and y: "isnormNum y"
  52.486 +  shows "x -\<^sub>N y = 0\<^sub>N \<longleftrightarrow> x = y"
  52.487 +proof -
  52.488 +  fix h :: 'a
  52.489 +  from isnormNum_unique[where 'a = 'a, OF Nsub_normN[OF y], where y="0\<^sub>N"]
  52.490 +  have "(x -\<^sub>N y = 0\<^sub>N) = (INum (x -\<^sub>N y) = (INum 0\<^sub>N :: 'a)) " by simp
  52.491 +  also have "\<dots> = (INum x = (INum y :: 'a))" by simp
  52.492 +  also have "\<dots> = (x = y)" using x y by simp
  52.493 +  finally show ?thesis .
  52.494 +qed
  52.495 +
  52.496 +lemma Nmul0[simp]: "c *\<^sub>N 0\<^sub>N = 0\<^sub>N" " 0\<^sub>N *\<^sub>N c = 0\<^sub>N"
  52.497 +  by (simp_all add: Nmul_def Let_def split_def)
  52.498 +
  52.499 +lemma Nmul_eq0[simp]:
  52.500 +  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  52.501 +  assumes nx: "isnormNum x" and ny: "isnormNum y"
  52.502 +  shows "x*\<^sub>N y = 0\<^sub>N \<longleftrightarrow> x = 0\<^sub>N \<or> y = 0\<^sub>N"
  52.503 +proof -
  52.504 +  fix h :: 'a
  52.505 +  obtain a b where x: "x = (a, b)" by (cases x)
  52.506 +  obtain a' b' where y: "y = (a', b')" by (cases y)
  52.507 +  have n0: "isnormNum 0\<^sub>N" by simp
  52.508 +  show ?thesis using nx ny
  52.509 +    apply (simp only: isnormNum_unique[where ?'a = 'a, OF  Nmul_normN[OF nx ny] n0, symmetric]
  52.510 +      Nmul[where ?'a = 'a])
  52.511 +    apply (simp add: x y INum_def split_def isnormNum_def split: split_if_asm)
  52.512 +    done
  52.513 +qed
  52.514 +
  52.515 +lemma Nneg_Nneg[simp]: "~\<^sub>N (~\<^sub>N c) = c"
  52.516 +  by (simp add: Nneg_def split_def)
  52.517 +
  52.518 +lemma Nmul1[simp]:
  52.519 +    "isnormNum c \<Longrightarrow> (1)\<^sub>N *\<^sub>N c = c"
  52.520 +    "isnormNum c \<Longrightarrow> c *\<^sub>N (1)\<^sub>N = c"
  52.521 +  apply (simp_all add: Nmul_def Let_def split_def isnormNum_def)
  52.522 +  apply (cases "fst c = 0", simp_all, cases c, simp_all)+
  52.523 +  done
  52.524 +
  52.525 +end
    53.1 --- a/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy	Mon Nov 11 17:34:44 2013 +0100
    53.2 +++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy	Mon Nov 11 17:44:21 2013 +0100
    53.3 @@ -5,7 +5,7 @@
    53.4  header {* Implementation and verification of multivariate polynomials *}
    53.5  
    53.6  theory Reflected_Multivariate_Polynomial
    53.7 -imports Complex_Main "~~/src/HOL/Library/Abstract_Rat" Polynomial_List
    53.8 +imports Complex_Main Rat_Pair Polynomial_List
    53.9  begin
   53.10  
   53.11  subsection{* Datatype of polynomial expressions *}
    54.1 --- a/src/HOL/Decision_Procs/mir_tac.ML	Mon Nov 11 17:34:44 2013 +0100
    54.2 +++ b/src/HOL/Decision_Procs/mir_tac.ML	Mon Nov 11 17:44:21 2013 +0100
    54.3 @@ -34,7 +34,7 @@
    54.4               @{thm "divide_zero"}, 
    54.5               @{thm "divide_divide_eq_left"}, @{thm "times_divide_eq_right"}, 
    54.6               @{thm "times_divide_eq_left"}, @{thm "divide_divide_eq_right"},
    54.7 -             @{thm "diff_minus"}, @{thm "minus_divide_left"}]
    54.8 +             @{thm uminus_add_conv_diff [symmetric]}, @{thm "minus_divide_left"}]
    54.9  val comp_ths = ths @ comp_arith @ @{thms simp_thms};
   54.10  
   54.11  
    55.1 --- a/src/HOL/Deriv.thy	Mon Nov 11 17:34:44 2013 +0100
    55.2 +++ b/src/HOL/Deriv.thy	Mon Nov 11 17:44:21 2013 +0100
    55.3 @@ -98,7 +98,7 @@
    55.4  
    55.5  lemma FDERIV_diff[simp, FDERIV_intros]:
    55.6    "(f has_derivative f') F \<Longrightarrow> (g has_derivative g') F \<Longrightarrow> ((\<lambda>x. f x - g x) has_derivative (\<lambda>x. f' x - g' x)) F"
    55.7 -  by (simp only: diff_minus FDERIV_add FDERIV_minus)
    55.8 +  by (simp only: diff_conv_add_uminus FDERIV_add FDERIV_minus)
    55.9  
   55.10  abbreviation
   55.11    -- {* Frechet derivative: D is derivative of function f at x within s *}
   55.12 @@ -718,13 +718,13 @@
   55.13        ((%x. (f(x)-f(a)) / (x-a)) -- a --> D)"
   55.14  apply (rule iffI)
   55.15  apply (drule_tac k="- a" in LIM_offset)
   55.16 -apply (simp add: diff_minus)
   55.17 +apply simp
   55.18  apply (drule_tac k="a" in LIM_offset)
   55.19  apply (simp add: add_commute)
   55.20  done
   55.21  
   55.22  lemma DERIV_iff2: "(DERIV f x :> D) \<longleftrightarrow> (\<lambda>z. (f z - f x) / (z - x)) --x --> D"
   55.23 -  by (simp add: deriv_def diff_minus [symmetric] DERIV_LIM_iff)
   55.24 +  by (simp add: deriv_def DERIV_LIM_iff)
   55.25  
   55.26  lemma DERIV_cong_ev: "x = y \<Longrightarrow> eventually (\<lambda>x. f x = g x) (nhds x) \<Longrightarrow> u = v \<Longrightarrow>
   55.27      DERIV f x :> u \<longleftrightarrow> DERIV g y :> v"
   55.28 @@ -758,8 +758,7 @@
   55.29      let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))"
   55.30      show "\<forall>z. f z - f x = ?g z * (z-x)" by simp
   55.31      show "isCont ?g x" using der
   55.32 -      by (simp add: isCont_iff DERIV_iff diff_minus
   55.33 -               cong: LIM_equal [rule_format])
   55.34 +      by (simp add: isCont_iff DERIV_iff cong: LIM_equal [rule_format])
   55.35      show "?g x = l" by simp
   55.36    qed
   55.37  next
   55.38 @@ -787,7 +786,7 @@
   55.39  proof -
   55.40    from l der [THEN DERIV_D, THEN LIM_D [where r = "l"]]
   55.41    have "\<exists>s > 0. (\<forall>z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < l)"
   55.42 -    by (simp add: diff_minus)
   55.43 +    by simp
   55.44    then obtain s
   55.45          where s:   "0 < s"
   55.46            and all: "!!z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < l"
   55.47 @@ -798,8 +797,7 @@
   55.48      fix h::real
   55.49      assume "0 < h" "h < s"
   55.50      with all [of h] show "f x < f (x+h)"
   55.51 -    proof (simp add: abs_if pos_less_divide_eq diff_minus [symmetric]
   55.52 -    split add: split_if_asm)
   55.53 +    proof (simp add: abs_if pos_less_divide_eq split add: split_if_asm)
   55.54        assume "~ (f (x+h) - f x) / h < l" and h: "0 < h"
   55.55        with l
   55.56        have "0 < (f (x+h) - f x) / h" by arith
   55.57 @@ -817,7 +815,7 @@
   55.58  proof -
   55.59    from l der [THEN DERIV_D, THEN LIM_D [where r = "-l"]]
   55.60    have "\<exists>s > 0. (\<forall>z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < -l)"
   55.61 -    by (simp add: diff_minus)
   55.62 +    by simp
   55.63    then obtain s
   55.64          where s:   "0 < s"
   55.65            and all: "!!z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < -l"
   55.66 @@ -828,8 +826,7 @@
   55.67      fix h::real
   55.68      assume "0 < h" "h < s"
   55.69      with all [of "-h"] show "f x < f (x-h)"
   55.70 -    proof (simp add: abs_if pos_less_divide_eq diff_minus [symmetric]
   55.71 -    split add: split_if_asm)
   55.72 +    proof (simp add: abs_if pos_less_divide_eq split add: split_if_asm)
   55.73        assume " - ((f (x-h) - f x) / h) < l" and h: "0 < h"
   55.74        with l
   55.75        have "0 < (f (x-h) - f x) / h" by arith
   55.76 @@ -1131,7 +1128,7 @@
   55.77  apply (rule linorder_cases [of a b], auto)
   55.78  apply (drule_tac [!] f = f in MVT)
   55.79  apply (auto dest: DERIV_isCont DERIV_unique simp add: differentiable_def)
   55.80 -apply (auto dest: DERIV_unique simp add: ring_distribs diff_minus)
   55.81 +apply (auto dest: DERIV_unique simp add: ring_distribs)
   55.82  done
   55.83  
   55.84  lemma DERIV_const_ratio_const2:
    56.1 --- a/src/HOL/Divides.thy	Mon Nov 11 17:34:44 2013 +0100
    56.2 +++ b/src/HOL/Divides.thy	Mon Nov 11 17:44:21 2013 +0100
    56.3 @@ -53,6 +53,16 @@
    56.4    shows "(a + b * c) div b = c + a div b"
    56.5    using assms div_mult_self1 [of b a c] by (simp add: mult_commute)
    56.6  
    56.7 +lemma div_mult_self3 [simp]:
    56.8 +  assumes "b \<noteq> 0"
    56.9 +  shows "(c * b + a) div b = c + a div b"
   56.10 +  using assms by (simp add: add.commute)
   56.11 +
   56.12 +lemma div_mult_self4 [simp]:
   56.13 +  assumes "b \<noteq> 0"
   56.14 +  shows "(b * c + a) div b = c + a div b"
   56.15 +  using assms by (simp add: add.commute)
   56.16 +
   56.17  lemma mod_mult_self1 [simp]: "(a + c * b) mod b = a mod b"
   56.18  proof (cases "b = 0")
   56.19    case True then show ?thesis by simp
   56.20 @@ -70,9 +80,18 @@
   56.21    then show ?thesis by simp
   56.22  qed
   56.23  
   56.24 -lemma mod_mult_self2 [simp]: "(a + b * c) mod b = a mod b"
   56.25 +lemma mod_mult_self2 [simp]: 
   56.26 +  "(a + b * c) mod b = a mod b"
   56.27    by (simp add: mult_commute [of b])
   56.28  
   56.29 +lemma mod_mult_self3 [simp]:
   56.30 +  "(c * b + a) mod b = a mod b"
   56.31 +  by (simp add: add.commute)
   56.32 +
   56.33 +lemma mod_mult_self4 [simp]:
   56.34 +  "(b * c + a) mod b = a mod b"
   56.35 +  by (simp add: add.commute)
   56.36 +
   56.37  lemma div_mult_self1_is_id [simp]: "b \<noteq> 0 \<Longrightarrow> b * a div b = a"
   56.38    using div_mult_self2 [of b 0 a] by simp
   56.39  
   56.40 @@ -420,24 +439,23 @@
   56.41  
   56.42  text {* Subtraction respects modular equivalence. *}
   56.43  
   56.44 -lemma mod_diff_left_eq: "(a - b) mod c = (a mod c - b) mod c"
   56.45 -  unfolding diff_minus
   56.46 -  by (intro mod_add_cong mod_minus_cong) simp_all
   56.47 -
   56.48 -lemma mod_diff_right_eq: "(a - b) mod c = (a - b mod c) mod c"
   56.49 -  unfolding diff_minus
   56.50 -  by (intro mod_add_cong mod_minus_cong) simp_all
   56.51 -
   56.52 -lemma mod_diff_eq: "(a - b) mod c = (a mod c - b mod c) mod c"
   56.53 -  unfolding diff_minus
   56.54 -  by (intro mod_add_cong mod_minus_cong) simp_all
   56.55 +lemma mod_diff_left_eq:
   56.56 +  "(a - b) mod c = (a mod c - b) mod c"
   56.57 +  using mod_add_cong [of a c "a mod c" "- b" "- b"] by simp
   56.58 +
   56.59 +lemma mod_diff_right_eq:
   56.60 +  "(a - b) mod c = (a - b mod c) mod c"
   56.61 +  using mod_add_cong [of a c a "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp
   56.62 +
   56.63 +lemma mod_diff_eq:
   56.64 +  "(a - b) mod c = (a mod c - b mod c) mod c"
   56.65 +  using mod_add_cong [of a c "a mod c" "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp
   56.66  
   56.67  lemma mod_diff_cong:
   56.68    assumes "a mod c = a' mod c"
   56.69    assumes "b mod c = b' mod c"
   56.70    shows "(a - b) mod c = (a' - b') mod c"
   56.71 -  unfolding diff_minus using assms
   56.72 -  by (intro mod_add_cong mod_minus_cong)
   56.73 +  using assms mod_add_cong [of a c a' "- b" "- b'"] mod_minus_cong [of b c "b'"] by simp
   56.74  
   56.75  lemma dvd_neg_div: "y dvd x \<Longrightarrow> -x div y = - (x div y)"
   56.76  apply (case_tac "y = 0") apply simp
   56.77 @@ -477,6 +495,34 @@
   56.78  lemma mod_minus1_right [simp]: "a mod (-1) = 0"
   56.79    using mod_minus_right [of a 1] by simp
   56.80  
   56.81 +lemma minus_mod_self2 [simp]: 
   56.82 +  "(a - b) mod b = a mod b"
   56.83 +  by (simp add: mod_diff_right_eq)
   56.84 +
   56.85 +lemma minus_mod_self1 [simp]: 
   56.86 +  "(b - a) mod b = - a mod b"
   56.87 +  using mod_add_self2 [of "- a" b] by simp
   56.88 +
   56.89 +end
   56.90 +
   56.91 +class semiring_div_parity = semiring_div + semiring_numeral +
   56.92 +  assumes parity: "a mod 2 = 0 \<or> a mod 2 = 1"
   56.93 +begin
   56.94 +
   56.95 +lemma parity_cases [case_names even odd]:
   56.96 +  assumes "a mod 2 = 0 \<Longrightarrow> P"
   56.97 +  assumes "a mod 2 = 1 \<Longrightarrow> P"
   56.98 +  shows P
   56.99 +  using assms parity by blast
  56.100 +
  56.101 +lemma not_mod_2_eq_0_eq_1 [simp]:
  56.102 +  "a mod 2 \<noteq> 0 \<longleftrightarrow> a mod 2 = 1"
  56.103 +  by (cases a rule: parity_cases) simp_all
  56.104 +
  56.105 +lemma not_mod_2_eq_1_eq_0 [simp]:
  56.106 +  "a mod 2 \<noteq> 1 \<longleftrightarrow> a mod 2 = 0"
  56.107 +  by (cases a rule: parity_cases) simp_all
  56.108 +
  56.109  end
  56.110  
  56.111  
  56.112 @@ -490,7 +536,6 @@
  56.113    and less technical class hierarchy.
  56.114  *}
  56.115  
  56.116 -
  56.117  class semiring_numeral_div = linordered_semidom + minus + semiring_div +
  56.118    assumes diff_invert_add1: "a + b = c \<Longrightarrow> a = c - b"
  56.119      and le_add_diff_inverse2: "b \<le> a \<Longrightarrow> a - b + b = a"
  56.120 @@ -510,18 +555,21 @@
  56.121    "a - 0 = a"
  56.122    by (rule diff_invert_add1 [symmetric]) simp
  56.123  
  56.124 -lemma parity:
  56.125 -  "a mod 2 = 0 \<or> a mod 2 = 1"
  56.126 -proof (rule ccontr)
  56.127 -  assume "\<not> (a mod 2 = 0 \<or> a mod 2 = 1)"
  56.128 -  then have "a mod 2 \<noteq> 0" and "a mod 2 \<noteq> 1" by simp_all
  56.129 -  have "0 < 2" by simp
  56.130 -  with pos_mod_bound pos_mod_sign have "0 \<le> a mod 2" "a mod 2 < 2" by simp_all
  56.131 -  with `a mod 2 \<noteq> 0` have "0 < a mod 2" by simp
  56.132 -  with discrete have "1 \<le> a mod 2" by simp
  56.133 -  with `a mod 2 \<noteq> 1` have "1 < a mod 2" by simp
  56.134 -  with discrete have "2 \<le> a mod 2" by simp
  56.135 -  with `a mod 2 < 2` show False by simp
  56.136 +subclass semiring_div_parity
  56.137 +proof
  56.138 +  fix a
  56.139 +  show "a mod 2 = 0 \<or> a mod 2 = 1"
  56.140 +  proof (rule ccontr)
  56.141 +    assume "\<not> (a mod 2 = 0 \<or> a mod 2 = 1)"
  56.142 +    then have "a mod 2 \<noteq> 0" and "a mod 2 \<noteq> 1" by simp_all
  56.143 +    have "0 < 2" by simp
  56.144 +    with pos_mod_bound pos_mod_sign have "0 \<le> a mod 2" "a mod 2 < 2" by simp_all
  56.145 +    with `a mod 2 \<noteq> 0` have "0 < a mod 2" by simp
  56.146 +    with discrete have "1 \<le> a mod 2" by simp
  56.147 +    with `a mod 2 \<noteq> 1` have "1 < a mod 2" by simp
  56.148 +    with discrete have "2 \<le> a mod 2" by simp
  56.149 +    with `a mod 2 < 2` show False by simp
  56.150 +  qed
  56.151  qed
  56.152  
  56.153  lemma divmod_digit_1:
  56.154 @@ -1697,7 +1745,7 @@
  56.155    val div_mod_eqs = map mk_meta_eq [@{thm div_mod_equality}, @{thm div_mod_equality2}];
  56.156  
  56.157    val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac 
  56.158 -    (@{thm diff_minus} :: @{thms add_0s} @ @{thms add_ac}))
  56.159 +    (@{thm diff_conv_add_uminus} :: @{thms add_0s} @ @{thms add_ac}))
  56.160  )
  56.161  *}
  56.162  
  56.163 @@ -2568,11 +2616,6 @@
  56.164    "Suc 0 mod numeral v' = nat (1 mod numeral v')"
  56.165    by (subst nat_mod_distrib) simp_all
  56.166  
  56.167 -lemma mod_2_not_eq_zero_eq_one_int:
  56.168 -  fixes k :: int
  56.169 -  shows "k mod 2 \<noteq> 0 \<longleftrightarrow> k mod 2 = 1"
  56.170 -  by auto
  56.171 -
  56.172  instance int :: semiring_numeral_div
  56.173    by intro_classes (auto intro: zmod_le_nonneg_dividend
  56.174      simp add: zmult_div_cancel
    57.1 --- a/src/HOL/Enum.thy	Mon Nov 11 17:34:44 2013 +0100
    57.2 +++ b/src/HOL/Enum.thy	Mon Nov 11 17:44:21 2013 +0100
    57.3 @@ -156,11 +156,11 @@
    57.4    "Id = image (\<lambda>x. (x, x)) (set Enum.enum)"
    57.5    by (auto intro: imageI in_enum)
    57.6  
    57.7 -lemma tranclp_unfold [code, no_atp]:
    57.8 +lemma tranclp_unfold [code]:
    57.9    "tranclp r a b \<longleftrightarrow> (a, b) \<in> trancl {(x, y). r x y}"
   57.10    by (simp add: trancl_def)
   57.11  
   57.12 -lemma rtranclp_rtrancl_eq [code, no_atp]:
   57.13 +lemma rtranclp_rtrancl_eq [code]:
   57.14    "rtranclp r x y \<longleftrightarrow> (x, y) \<in> rtrancl {(x, y). r x y}"
   57.15    by (simp add: rtrancl_def)
   57.16  
   57.17 @@ -178,13 +178,9 @@
   57.18  
   57.19  lemma [code]:
   57.20    fixes xs :: "('a::finite \<times> 'a) list"
   57.21 -  shows "acc (set xs) = bacc (set xs) (card_UNIV TYPE('a))"
   57.22 +  shows "Wellfounded.acc (set xs) = bacc (set xs) (card_UNIV TYPE('a))"
   57.23    by (simp add: card_UNIV_def acc_bacc_eq)
   57.24  
   57.25 -lemma [code]:
   57.26 -  "accp r = (\<lambda>x. x \<in> acc {(x, y). r x y})"
   57.27 -  by (simp add: acc_def)
   57.28 -
   57.29  
   57.30  subsection {* Default instances for @{class enum} *}
   57.31  
    58.1 --- a/src/HOL/Fields.thy	Mon Nov 11 17:34:44 2013 +0100
    58.2 +++ b/src/HOL/Fields.thy	Mon Nov 11 17:44:21 2013 +0100
    58.3 @@ -152,11 +152,11 @@
    58.4  lemma nonzero_minus_divide_divide: "b \<noteq> 0 ==> (-a) / (-b) = a / b"
    58.5    by (simp add: divide_inverse nonzero_inverse_minus_eq)
    58.6  
    58.7 -lemma divide_minus_left [simp, no_atp]: "(-a) / b = - (a / b)"
    58.8 +lemma divide_minus_left [simp]: "(-a) / b = - (a / b)"
    58.9    by (simp add: divide_inverse)
   58.10  
   58.11  lemma diff_divide_distrib: "(a - b) / c = a / c - b / c"
   58.12 -  by (simp add: diff_minus add_divide_distrib)
   58.13 +  using add_divide_distrib [of a "- b" c] by simp
   58.14  
   58.15  lemma nonzero_eq_divide_eq [field_simps]: "c \<noteq> 0 \<Longrightarrow> a = b / c \<longleftrightarrow> a * c = b"
   58.16  proof -
   58.17 @@ -252,7 +252,7 @@
   58.18     ==> inverse a + inverse b = (a + b) * inverse a * inverse b"
   58.19  by (simp add: division_ring_inverse_add mult_ac)
   58.20  
   58.21 -lemma nonzero_mult_divide_mult_cancel_left [simp, no_atp]:
   58.22 +lemma nonzero_mult_divide_mult_cancel_left [simp]:
   58.23  assumes [simp]: "b\<noteq>0" and [simp]: "c\<noteq>0" shows "(c*a)/(c*b) = a/b"
   58.24  proof -
   58.25    have "(c*a)/(c*b) = c * a * (inverse b * inverse c)"
   58.26 @@ -263,7 +263,7 @@
   58.27      finally show ?thesis by (simp add: divide_inverse)
   58.28  qed
   58.29  
   58.30 -lemma nonzero_mult_divide_mult_cancel_right [simp, no_atp]:
   58.31 +lemma nonzero_mult_divide_mult_cancel_right [simp]:
   58.32    "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (a * c) / (b * c) = a / b"
   58.33  by (simp add: mult_commute [of _ c])
   58.34  
   58.35 @@ -275,7 +275,7 @@
   58.36    fraction, like a*b*c / x*y*z. The rationale for that is unclear, but
   58.37    many proofs seem to need them.*}
   58.38  
   58.39 -lemmas times_divide_eq [no_atp] = times_divide_eq_right times_divide_eq_left
   58.40 +lemmas times_divide_eq = times_divide_eq_right times_divide_eq_left
   58.41  
   58.42  lemma add_frac_eq:
   58.43    assumes "y \<noteq> 0" and "z \<noteq> 0"
   58.44 @@ -291,27 +291,27 @@
   58.45  
   58.46  text{*Special Cancellation Simprules for Division*}
   58.47  
   58.48 -lemma nonzero_mult_divide_cancel_right [simp, no_atp]:
   58.49 +lemma nonzero_mult_divide_cancel_right [simp]:
   58.50    "b \<noteq> 0 \<Longrightarrow> a * b / b = a"
   58.51    using nonzero_mult_divide_mult_cancel_right [of 1 b a] by simp
   58.52  
   58.53 -lemma nonzero_mult_divide_cancel_left [simp, no_atp]:
   58.54 +lemma nonzero_mult_divide_cancel_left [simp]:
   58.55    "a \<noteq> 0 \<Longrightarrow> a * b / a = b"
   58.56  using nonzero_mult_divide_mult_cancel_left [of 1 a b] by simp
   58.57  
   58.58 -lemma nonzero_divide_mult_cancel_right [simp, no_atp]:
   58.59 +lemma nonzero_divide_mult_cancel_right [simp]:
   58.60    "\<lbrakk>a \<noteq> 0; b \<noteq> 0\<rbrakk> \<Longrightarrow> b / (a * b) = 1 / a"
   58.61  using nonzero_mult_divide_mult_cancel_right [of a b 1] by simp
   58.62  
   58.63 -lemma nonzero_divide_mult_cancel_left [simp, no_atp]:
   58.64 +lemma nonzero_divide_mult_cancel_left [simp]:
   58.65    "\<lbrakk>a \<noteq> 0; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / (a * b) = 1 / b"
   58.66  using nonzero_mult_divide_mult_cancel_left [of b a 1] by simp
   58.67  
   58.68 -lemma nonzero_mult_divide_mult_cancel_left2 [simp, no_atp]:
   58.69 +lemma nonzero_mult_divide_mult_cancel_left2 [simp]:
   58.70    "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (c * a) / (b * c) = a / b"
   58.71  using nonzero_mult_divide_mult_cancel_left [of b c a] by (simp add: mult_ac)
   58.72  
   58.73 -lemma nonzero_mult_divide_mult_cancel_right2 [simp, no_atp]:
   58.74 +lemma nonzero_mult_divide_mult_cancel_right2 [simp]:
   58.75    "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (a * c) / (c * b) = a / b"
   58.76  using nonzero_mult_divide_mult_cancel_right [of b c a] by (simp add: mult_ac)
   58.77  
   58.78 @@ -383,18 +383,18 @@
   58.79  apply simp_all
   58.80  done
   58.81  
   58.82 -lemma divide_divide_eq_right [simp, no_atp]:
   58.83 +lemma divide_divide_eq_right [simp]:
   58.84    "a / (b / c) = (a * c) / b"
   58.85    by (simp add: divide_inverse mult_ac)
   58.86  
   58.87 -lemma divide_divide_eq_left [simp, no_atp]:
   58.88 +lemma divide_divide_eq_left [simp]:
   58.89    "(a / b) / c = a / (b * c)"
   58.90    by (simp add: divide_inverse mult_assoc)
   58.91  
   58.92  
   58.93  text {*Special Cancellation Simprules for Division*}
   58.94  
   58.95 -lemma mult_divide_mult_cancel_left_if [simp,no_atp]:
   58.96 +lemma mult_divide_mult_cancel_left_if [simp]:
   58.97    shows "(c * a) / (c * b) = (if c = 0 then 0 else a / b)"
   58.98    by (simp add: mult_divide_mult_cancel_left)
   58.99  
  58.100 @@ -405,7 +405,7 @@
  58.101    "- (a / b) = a / - b"
  58.102    by (simp add: divide_inverse)
  58.103  
  58.104 -lemma divide_minus_right [simp, no_atp]:
  58.105 +lemma divide_minus_right [simp]:
  58.106    "a / - b = - (a / b)"
  58.107    by (simp add: divide_inverse)
  58.108  
  58.109 @@ -427,29 +427,29 @@
  58.110    "inverse x = 1 \<longleftrightarrow> x = 1"
  58.111    by (insert inverse_eq_iff_eq [of x 1], simp) 
  58.112  
  58.113 -lemma divide_eq_0_iff [simp, no_atp]:
  58.114 +lemma divide_eq_0_iff [simp]:
  58.115    "a / b = 0 \<longleftrightarrow> a = 0 \<or> b = 0"
  58.116    by (simp add: divide_inverse)
  58.117  
  58.118 -lemma divide_cancel_right [simp, no_atp]:
  58.119 +lemma divide_cancel_right [simp]:
  58.120    "a / c = b / c \<longleftrightarrow> c = 0 \<or> a = b"
  58.121    apply (cases "c=0", simp)
  58.122    apply (simp add: divide_inverse)
  58.123    done
  58.124  
  58.125 -lemma divide_cancel_left [simp, no_atp]:
  58.126 +lemma divide_cancel_left [simp]:
  58.127    "c / a = c / b \<longleftrightarrow> c = 0 \<or> a = b" 
  58.128    apply (cases "c=0", simp)
  58.129    apply (simp add: divide_inverse)
  58.130    done
  58.131  
  58.132 -lemma divide_eq_1_iff [simp, no_atp]:
  58.133 +lemma divide_eq_1_iff [simp]:
  58.134    "a / b = 1 \<longleftrightarrow> b \<noteq> 0 \<and> a = b"
  58.135    apply (cases "b=0", simp)
  58.136    apply (simp add: right_inverse_eq)
  58.137    done
  58.138  
  58.139 -lemma one_eq_divide_iff [simp, no_atp]:
  58.140 +lemma one_eq_divide_iff [simp]:
  58.141    "1 = a / b \<longleftrightarrow> b \<noteq> 0 \<and> a = b"
  58.142    by (simp add: eq_commute [of 1])
  58.143  
  58.144 @@ -559,7 +559,7 @@
  58.145  done
  58.146  
  58.147  text{*Both premises are essential. Consider -1 and 1.*}
  58.148 -lemma inverse_less_iff_less [simp,no_atp]:
  58.149 +lemma inverse_less_iff_less [simp]:
  58.150    "0 < a \<Longrightarrow> 0 < b \<Longrightarrow> inverse a < inverse b \<longleftrightarrow> b < a"
  58.151    by (blast intro: less_imp_inverse_less dest: inverse_less_imp_less) 
  58.152  
  58.153 @@ -567,7 +567,7 @@
  58.154    "a \<le> b \<Longrightarrow> 0 < a \<Longrightarrow> inverse b \<le> inverse a"
  58.155    by (force simp add: le_less less_imp_inverse_less)
  58.156  
  58.157 -lemma inverse_le_iff_le [simp,no_atp]:
  58.158 +lemma inverse_le_iff_le [simp]:
  58.159    "0 < a \<Longrightarrow> 0 < b \<Longrightarrow> inverse a \<le> inverse b \<longleftrightarrow> b \<le> a"
  58.160    by (blast intro: le_imp_inverse_le dest: inverse_le_imp_le) 
  58.161  
  58.162 @@ -601,7 +601,7 @@
  58.163  apply (simp add: nonzero_inverse_minus_eq) 
  58.164  done
  58.165  
  58.166 -lemma inverse_less_iff_less_neg [simp,no_atp]:
  58.167 +lemma inverse_less_iff_less_neg [simp]:
  58.168    "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> inverse a < inverse b \<longleftrightarrow> b < a"
  58.169  apply (insert inverse_less_iff_less [of "-b" "-a"])
  58.170  apply (simp del: inverse_less_iff_less 
  58.171 @@ -612,7 +612,7 @@
  58.172    "a \<le> b \<Longrightarrow> b < 0 ==> inverse b \<le> inverse a"
  58.173    by (force simp add: le_less less_imp_inverse_less_neg)
  58.174  
  58.175 -lemma inverse_le_iff_le_neg [simp,no_atp]:
  58.176 +lemma inverse_le_iff_le_neg [simp]:
  58.177    "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> inverse a \<le> inverse b \<longleftrightarrow> b \<le> a"
  58.178    by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg) 
  58.179  
  58.180 @@ -713,11 +713,9 @@
  58.181  sign_simps} to @{text field_simps} because the former can lead to case
  58.182  explosions. *}
  58.183  
  58.184 -lemmas sign_simps [no_atp] = algebra_simps
  58.185 -  zero_less_mult_iff mult_less_0_iff
  58.186 +lemmas sign_simps = algebra_simps zero_less_mult_iff mult_less_0_iff
  58.187  
  58.188 -lemmas (in -) sign_simps [no_atp] = algebra_simps
  58.189 -  zero_less_mult_iff mult_less_0_iff
  58.190 +lemmas (in -) sign_simps = algebra_simps zero_less_mult_iff mult_less_0_iff
  58.191  
  58.192  (* Only works once linear arithmetic is installed:
  58.193  text{*An example:*}
  58.194 @@ -847,7 +845,7 @@
  58.195    fix x y :: 'a
  58.196    from less_add_one show "\<exists>y. x < y" .. 
  58.197    from less_add_one have "x + (- 1) < (x + 1) + (- 1)" by (rule add_strict_right_mono)
  58.198 -  then have "x - 1 < x + 1 - 1" by (simp only: diff_minus [symmetric])
  58.199 +  then have "x - 1 < x + 1 - 1" by simp
  58.200    then have "x - 1 < x" by (simp add: algebra_simps)
  58.201    then show "\<exists>y. y < x" ..
  58.202    show "x < y \<Longrightarrow> \<exists>z>x. z < y" by (blast intro!: less_half_sum gt_half_sum)
  58.203 @@ -998,13 +996,13 @@
  58.204  
  58.205  text{*Simplify expressions equated with 1*}
  58.206  
  58.207 -lemma zero_eq_1_divide_iff [simp,no_atp]:
  58.208 +lemma zero_eq_1_divide_iff [simp]:
  58.209       "(0 = 1/a) = (a = 0)"
  58.210  apply (cases "a=0", simp)
  58.211  apply (auto simp add: nonzero_eq_divide_eq)
  58.212  done
  58.213  
  58.214 -lemma one_divide_eq_0_iff [simp,no_atp]:
  58.215 +lemma one_divide_eq_0_iff [simp]:
  58.216       "(1/a = 0) = (a = 0)"
  58.217  apply (cases "a=0", simp)
  58.218  apply (insert zero_neq_one [THEN not_sym])
  58.219 @@ -1013,19 +1011,19 @@
  58.220  
  58.221  text{*Simplify expressions such as @{text "0 < 1/x"} to @{text "0 < x"}*}
  58.222  
  58.223 -lemma zero_le_divide_1_iff [simp, no_atp]:
  58.224 +lemma zero_le_divide_1_iff [simp]:
  58.225    "0 \<le> 1 / a \<longleftrightarrow> 0 \<le> a"
  58.226    by (simp add: zero_le_divide_iff)
  58.227  
  58.228 -lemma zero_less_divide_1_iff [simp, no_atp]:
  58.229 +lemma zero_less_divide_1_iff [simp]:
  58.230    "0 < 1 / a \<longleftrightarrow> 0 < a"
  58.231    by (simp add: zero_less_divide_iff)
  58.232  
  58.233 -lemma divide_le_0_1_iff [simp, no_atp]:
  58.234 +lemma divide_le_0_1_iff [simp]:
  58.235    "1 / a \<le> 0 \<longleftrightarrow> a \<le> 0"
  58.236    by (simp add: divide_le_0_iff)
  58.237  
  58.238 -lemma divide_less_0_1_iff [simp, no_atp]:
  58.239 +lemma divide_less_0_1_iff [simp]:
  58.240    "1 / a < 0 \<longleftrightarrow> a < 0"
  58.241    by (simp add: divide_less_0_iff)
  58.242  
  58.243 @@ -1080,62 +1078,62 @@
  58.244  
  58.245  text{*Simplify quotients that are compared with the value 1.*}
  58.246  
  58.247 -lemma le_divide_eq_1 [no_atp]:
  58.248 +lemma le_divide_eq_1:
  58.249    "(1 \<le> b / a) = ((0 < a & a \<le> b) | (a < 0 & b \<le> a))"
  58.250  by (auto simp add: le_divide_eq)
  58.251  
  58.252 -lemma divide_le_eq_1 [no_atp]:
  58.253 +lemma divide_le_eq_1:
  58.254    "(b / a \<le> 1) = ((0 < a & b \<le> a) | (a < 0 & a \<le> b) | a=0)"
  58.255  by (auto simp add: divide_le_eq)
  58.256  
  58.257 -lemma less_divide_eq_1 [no_atp]:
  58.258 +lemma less_divide_eq_1:
  58.259    "(1 < b / a) = ((0 < a & a < b) | (a < 0 & b < a))"
  58.260  by (auto simp add: less_divide_eq)
  58.261  
  58.262 -lemma divide_less_eq_1 [no_atp]:
  58.263 +lemma divide_less_eq_1:
  58.264    "(b / a < 1) = ((0 < a & b < a) | (a < 0 & a < b) | a=0)"
  58.265  by (auto simp add: divide_less_eq)
  58.266  
  58.267  
  58.268  text {*Conditional Simplification Rules: No Case Splits*}
  58.269  
  58.270 -lemma le_divide_eq_1_pos [simp,no_atp]:
  58.271 +lemma le_divide_eq_1_pos [simp]:
  58.272    "0 < a \<Longrightarrow> (1 \<le> b/a) = (a \<le> b)"
  58.273  by (auto simp add: le_divide_eq)
  58.274  
  58.275 -lemma le_divide_eq_1_neg [simp,no_atp]:
  58.276 +lemma le_divide_eq_1_neg [simp]:
  58.277    "a < 0 \<Longrightarrow> (1 \<le> b/a) = (b \<le> a)"
  58.278  by (auto simp add: le_divide_eq)
  58.279  
  58.280 -lemma divide_le_eq_1_pos [simp,no_atp]:
  58.281 +lemma divide_le_eq_1_pos [simp]:
  58.282    "0 < a \<Longrightarrow> (b/a \<le> 1) = (b \<le> a)"
  58.283  by (auto simp add: divide_le_eq)
  58.284  
  58.285 -lemma divide_le_eq_1_neg [simp,no_atp]:
  58.286 +lemma divide_le_eq_1_neg [simp]:
  58.287    "a < 0 \<Longrightarrow> (b/a \<le> 1) = (a \<le> b)"
  58.288  by (auto simp add: divide_le_eq)
  58.289  
  58.290 -lemma less_divide_eq_1_pos [simp,no_atp]:
  58.291 +lemma less_divide_eq_1_pos [simp]:
  58.292    "0 < a \<Longrightarrow> (1 < b/a) = (a < b)"
  58.293  by (auto simp add: less_divide_eq)
  58.294  
  58.295 -lemma less_divide_eq_1_neg [simp,no_atp]:
  58.296 +lemma less_divide_eq_1_neg [simp]:
  58.297    "a < 0 \<Longrightarrow> (1 < b/a) = (b < a)"
  58.298  by (auto simp add: less_divide_eq)
  58.299  
  58.300 -lemma divide_less_eq_1_pos [simp,no_atp]:
  58.301 +lemma divide_less_eq_1_pos [simp]:
  58.302    "0 < a \<Longrightarrow> (b/a < 1) = (b < a)"
  58.303  by (auto simp add: divide_less_eq)
  58.304  
  58.305 -lemma divide_less_eq_1_neg [simp,no_atp]:
  58.306 +lemma divide_less_eq_1_neg [simp]:
  58.307    "a < 0 \<Longrightarrow> b/a < 1 <-> a < b"
  58.308  by (auto simp add: divide_less_eq)
  58.309  
  58.310 -lemma eq_divide_eq_1 [simp,no_atp]:
  58.311 +lemma eq_divide_eq_1 [simp]:
  58.312    "(1 = b/a) = ((a \<noteq> 0 & a = b))"
  58.313  by (auto simp add: eq_divide_eq)
  58.314  
  58.315 -lemma divide_eq_eq_1 [simp,no_atp]:
  58.316 +lemma divide_eq_eq_1 [simp]:
  58.317    "(b/a = 1) = ((a \<noteq> 0 & a = b))"
  58.318  by (auto simp add: divide_eq_eq)
  58.319  
    59.1 --- a/src/HOL/Finite_Set.thy	Mon Nov 11 17:34:44 2013 +0100
    59.2 +++ b/src/HOL/Finite_Set.thy	Mon Nov 11 17:44:21 2013 +0100
    59.3 @@ -1188,7 +1188,7 @@
    59.4    "card A > 0 \<Longrightarrow> finite A"
    59.5    by (rule ccontr) simp
    59.6  
    59.7 -lemma card_0_eq [simp, no_atp]:
    59.8 +lemma card_0_eq [simp]:
    59.9    "finite A \<Longrightarrow> card A = 0 \<longleftrightarrow> A = {}"
   59.10    by (auto dest: mk_disjoint_insert)
   59.11  
   59.12 @@ -1620,8 +1620,7 @@
   59.13    show False by simp (blast dest: Suc_neq_Zero surjD)
   59.14  qed
   59.15  
   59.16 -(* Often leads to bogus ATP proofs because of reduced type information, hence no_atp *)
   59.17 -lemma infinite_UNIV_char_0 [no_atp]:
   59.18 +lemma infinite_UNIV_char_0:
   59.19    "\<not> finite (UNIV :: 'a::semiring_char_0 set)"
   59.20  proof
   59.21    assume "finite (UNIV :: 'a set)"
    60.1 --- a/src/HOL/Fun.thy	Mon Nov 11 17:34:44 2013 +0100
    60.2 +++ b/src/HOL/Fun.thy	Mon Nov 11 17:44:21 2013 +0100
    60.3 @@ -775,7 +775,7 @@
    60.4  
    60.5  subsection {* Cantor's Paradox *}
    60.6  
    60.7 -lemma Cantors_paradox [no_atp]:
    60.8 +lemma Cantors_paradox:
    60.9    "\<not>(\<exists>f. f ` A = Pow A)"
   60.10  proof clarify
   60.11    fix f assume "f ` A = Pow A" hence *: "Pow A \<le> f ` A" by blast
    61.1 --- a/src/HOL/GCD.thy	Mon Nov 11 17:34:44 2013 +0100
    61.2 +++ b/src/HOL/GCD.thy	Mon Nov 11 17:44:21 2013 +0100
    61.3 @@ -1555,8 +1555,8 @@
    61.4  interpretation gcd_lcm_complete_lattice_nat:
    61.5    complete_lattice Gcd Lcm gcd Rings.dvd "\<lambda>m n. m dvd n \<and> \<not> n dvd m" lcm 1 "0::nat"
    61.6  where
    61.7 -  "complete_lattice.INFI Gcd A f = Gcd (f ` A :: nat set)"
    61.8 -  and "complete_lattice.SUPR Lcm A f = Lcm (f ` A)"
    61.9 +  "Inf.INFI Gcd A f = Gcd (f ` A :: nat set)"
   61.10 +  and "Sup.SUPR Lcm A f = Lcm (f ` A)"
   61.11  proof -
   61.12    show "class.complete_lattice Gcd Lcm gcd Rings.dvd (\<lambda>m n. m dvd n \<and> \<not> n dvd m) lcm 1 (0::nat)"
   61.13    proof
   61.14 @@ -1574,8 +1574,8 @@
   61.15    qed
   61.16    then interpret gcd_lcm_complete_lattice_nat:
   61.17      complete_lattice Gcd Lcm gcd Rings.dvd "\<lambda>m n. m dvd n \<and> \<not> n dvd m" lcm 1 "0::nat" .
   61.18 -  from gcd_lcm_complete_lattice_nat.INF_def show "complete_lattice.INFI Gcd A f = Gcd (f ` A)" .
   61.19 -  from gcd_lcm_complete_lattice_nat.SUP_def show "complete_lattice.SUPR Lcm A f = Lcm (f ` A)" .
   61.20 +  from gcd_lcm_complete_lattice_nat.INF_def show "Inf.INFI Gcd A f = Gcd (f ` A)" .
   61.21 +  from gcd_lcm_complete_lattice_nat.SUP_def show "Sup.SUPR Lcm A f = Lcm (f ` A)" .
   61.22  qed
   61.23  
   61.24  lemma Lcm_empty_nat: "Lcm {} = (1::nat)"
    62.1 --- a/src/HOL/Groebner_Basis.thy	Mon Nov 11 17:34:44 2013 +0100
    62.2 +++ b/src/HOL/Groebner_Basis.thy	Mon Nov 11 17:44:21 2013 +0100
    62.3 @@ -10,20 +10,23 @@
    62.4  
    62.5  subsection {* Groebner Bases *}
    62.6  
    62.7 -lemmas bool_simps = simp_thms(1-34)
    62.8 +lemmas bool_simps = simp_thms(1-34) -- {* FIXME move to @{theory HOL} *}
    62.9 +
   62.10 +lemma nnf_simps: -- {* FIXME shadows fact binding in @{theory HOL} *}
   62.11 +  "(\<not>(P \<and> Q)) = (\<not>P \<or> \<not>Q)" "(\<not>(P \<or> Q)) = (\<not>P \<and> \<not>Q)"
   62.12 +  "(P \<longrightarrow> Q) = (\<not>P \<or> Q)"
   62.13 +  "(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P"
   62.14 +  by blast+
   62.15  
   62.16  lemma dnf:
   62.17 -    "(P & (Q | R)) = ((P&Q) | (P&R))" "((Q | R) & P) = ((Q&P) | (R&P))"
   62.18 -    "(P \<and> Q) = (Q \<and> P)" "(P \<or> Q) = (Q \<or> P)"
   62.19 +  "(P & (Q | R)) = ((P&Q) | (P&R))"
   62.20 +  "((Q | R) & P) = ((Q&P) | (R&P))"
   62.21 +  "(P \<and> Q) = (Q \<and> P)"
   62.22 +  "(P \<or> Q) = (Q \<or> P)"
   62.23    by blast+
   62.24  
   62.25  lemmas weak_dnf_simps = dnf bool_simps
   62.26  
   62.27 -lemma nnf_simps:
   62.28 -    "(\<not>(P \<and> Q)) = (\<not>P \<or> \<not>Q)" "(\<not>(P \<or> Q)) = (\<not>P \<and> \<not>Q)" "(P \<longrightarrow> Q) = (\<not>P \<or> Q)"
   62.29 -    "(P = Q) = ((P \<and> Q) \<or> (\<not>P \<and> \<not> Q))" "(\<not> \<not>(P)) = P"
   62.30 -  by blast+
   62.31 -
   62.32  lemma PFalse:
   62.33      "P \<equiv> False \<Longrightarrow> \<not> P"
   62.34      "\<not> P \<Longrightarrow> (P \<equiv> False)"
    63.1 --- a/src/HOL/Groups.thy	Mon Nov 11 17:34:44 2013 +0100
    63.2 +++ b/src/HOL/Groups.thy	Mon Nov 11 17:44:21 2013 +0100
    63.3 @@ -321,9 +321,13 @@
    63.4  
    63.5  class group_add = minus + uminus + monoid_add +
    63.6    assumes left_minus [simp]: "- a + a = 0"
    63.7 -  assumes diff_minus: "a - b = a + (- b)"
    63.8 +  assumes add_uminus_conv_diff [simp]: "a + (- b) = a - b"
    63.9  begin
   63.10  
   63.11 +lemma diff_conv_add_uminus:
   63.12 +  "a - b = a + (- b)"
   63.13 +  by simp
   63.14 +
   63.15  lemma minus_unique:
   63.16    assumes "a + b = 0" shows "- a = b"
   63.17  proof -
   63.18 @@ -332,8 +336,6 @@
   63.19    finally show ?thesis .
   63.20  qed
   63.21  
   63.22 -lemmas equals_zero_I = minus_unique (* legacy name *)
   63.23 -
   63.24  lemma minus_zero [simp]: "- 0 = 0"
   63.25  proof -
   63.26    have "0 + 0 = 0" by (rule add_0_right)
   63.27 @@ -346,13 +348,17 @@
   63.28    thus "- (- a) = a" by (rule minus_unique)
   63.29  qed
   63.30  
   63.31 -lemma right_minus [simp]: "a + - a = 0"
   63.32 +lemma right_minus: "a + - a = 0"
   63.33  proof -
   63.34    have "a + - a = - (- a) + - a" by simp
   63.35    also have "\<dots> = 0" by (rule left_minus)
   63.36    finally show ?thesis .
   63.37  qed
   63.38  
   63.39 +lemma diff_self [simp]:
   63.40 +  "a - a = 0"
   63.41 +  using right_minus [of a] by simp
   63.42 +
   63.43  subclass cancel_semigroup_add
   63.44  proof
   63.45    fix a b c :: 'a
   63.46 @@ -367,41 +373,57 @@
   63.47    then show "b = c" unfolding add_assoc by simp
   63.48  qed
   63.49  
   63.50 -lemma minus_add_cancel: "- a + (a + b) = b"
   63.51 -by (simp add: add_assoc [symmetric])
   63.52 +lemma minus_add_cancel [simp]:
   63.53 +  "- a + (a + b) = b"
   63.54 +  by (simp add: add_assoc [symmetric])
   63.55  
   63.56 -lemma add_minus_cancel: "a + (- a + b) = b"
   63.57 -by (simp add: add_assoc [symmetric])
   63.58 +lemma add_minus_cancel [simp]:
   63.59 +  "a + (- a + b) = b"
   63.60 +  by (simp add: add_assoc [symmetric])
   63.61  
   63.62 -lemma minus_add: "- (a + b) = - b + - a"
   63.63 +lemma diff_add_cancel [simp]:
   63.64 +  "a - b + b = a"
   63.65 +  by (simp only: diff_conv_add_uminus add_assoc) simp
   63.66 +
   63.67 +lemma add_diff_cancel [simp]:
   63.68 +  "a + b - b = a"
   63.69 +  by (simp only: diff_conv_add_uminus add_assoc) simp
   63.70 +
   63.71 +lemma minus_add:
   63.72 +  "- (a + b) = - b + - a"
   63.73  proof -
   63.74    have "(a + b) + (- b + - a) = 0"
   63.75 -    by (simp add: add_assoc add_minus_cancel)
   63.76 -  thus "- (a + b) = - b + - a"
   63.77 +    by (simp only: add_assoc add_minus_cancel) simp
   63.78 +  then show "- (a + b) = - b + - a"
   63.79      by (rule minus_unique)
   63.80  qed
   63.81  
   63.82 -lemma right_minus_eq: "a - b = 0 \<longleftrightarrow> a = b"
   63.83 +lemma right_minus_eq [simp]:
   63.84 +  "a - b = 0 \<longleftrightarrow> a = b"
   63.85  proof
   63.86    assume "a - b = 0"
   63.87 -  have "a = (a - b) + b" by (simp add:diff_minus add_assoc)
   63.88 +  have "a = (a - b) + b" by (simp add: add_assoc)
   63.89    also have "\<dots> = b" using `a - b = 0` by simp
   63.90    finally show "a = b" .
   63.91  next
   63.92 -  assume "a = b" thus "a - b = 0" by (simp add: diff_minus)
   63.93 +  assume "a = b" thus "a - b = 0" by simp
   63.94  qed
   63.95  
   63.96 -lemma diff_self [simp]: "a - a = 0"
   63.97 -by (simp add: diff_minus)
   63.98 +lemma eq_iff_diff_eq_0:
   63.99 +  "a = b \<longleftrightarrow> a - b = 0"
  63.100 +  by (fact right_minus_eq [symmetric])
  63.101  
  63.102 -lemma diff_0 [simp]: "0 - a = - a"
  63.103 -by (simp add: diff_minus)
  63.104 +lemma diff_0 [simp]:
  63.105 +  "0 - a = - a"
  63.106 +  by (simp only: diff_conv_add_uminus add_0_left)
  63.107  
  63.108 -lemma diff_0_right [simp]: "a - 0 = a" 
  63.109 -by (simp add: diff_minus)
  63.110 +lemma diff_0_right [simp]:
  63.111 +  "a - 0 = a" 
  63.112 +  by (simp only: diff_conv_add_uminus minus_zero add_0_right)
  63.113  
  63.114 -lemma diff_minus_eq_add [simp]: "a - - b = a + b"
  63.115 -by (simp add: diff_minus)
  63.116 +lemma diff_minus_eq_add [simp]:
  63.117 +  "a - - b = a + b"
  63.118 +  by (simp only: diff_conv_add_uminus minus_minus)
  63.119  
  63.120  lemma neg_equal_iff_equal [simp]:
  63.121    "- a = - b \<longleftrightarrow> a = b" 
  63.122 @@ -416,11 +438,11 @@
  63.123  
  63.124  lemma neg_equal_0_iff_equal [simp]:
  63.125    "- a = 0 \<longleftrightarrow> a = 0"
  63.126 -by (subst neg_equal_iff_equal [symmetric], simp)
  63.127 +  by (subst neg_equal_iff_equal [symmetric]) simp
  63.128  
  63.129  lemma neg_0_equal_iff_equal [simp]:
  63.130    "0 = - a \<longleftrightarrow> 0 = a"
  63.131 -by (subst neg_equal_iff_equal [symmetric], simp)
  63.132 +  by (subst neg_equal_iff_equal [symmetric]) simp
  63.133  
  63.134  text{*The next two equations can make the simplifier loop!*}
  63.135  
  63.136 @@ -438,15 +460,8 @@
  63.137    thus ?thesis by (simp add: eq_commute)
  63.138  qed
  63.139  
  63.140 -lemma diff_add_cancel: "a - b + b = a"
  63.141 -by (simp add: diff_minus add_assoc)
  63.142 -
  63.143 -lemma add_diff_cancel: "a + b - b = a"
  63.144 -by (simp add: diff_minus add_assoc)
  63.145 -
  63.146 -declare diff_minus[symmetric, algebra_simps, field_simps]
  63.147 -
  63.148 -lemma eq_neg_iff_add_eq_0: "a = - b \<longleftrightarrow> a + b = 0"
  63.149 +lemma eq_neg_iff_add_eq_0:
  63.150 +  "a = - b \<longleftrightarrow> a + b = 0"
  63.151  proof
  63.152    assume "a = - b" then show "a + b = 0" by simp
  63.153  next
  63.154 @@ -456,72 +471,88 @@
  63.155    ultimately show "a = - b" by simp
  63.156  qed
  63.157  
  63.158 -lemma add_eq_0_iff: "x + y = 0 \<longleftrightarrow> y = - x"
  63.159 -  unfolding eq_neg_iff_add_eq_0 [symmetric]
  63.160 -  by (rule equation_minus_iff)
  63.161 +lemma add_eq_0_iff2:
  63.162 +  "a + b = 0 \<longleftrightarrow> a = - b"
  63.163 +  by (fact eq_neg_iff_add_eq_0 [symmetric])
  63.164  
  63.165 -lemma minus_diff_eq [simp]: "- (a - b) = b - a"
  63.166 -  by (simp add: diff_minus minus_add)
  63.167 +lemma neg_eq_iff_add_eq_0:
  63.168 +  "- a = b \<longleftrightarrow> a + b = 0"
  63.169 +  by (auto simp add: add_eq_0_iff2)
  63.170  
  63.171 -lemma add_diff_eq[algebra_simps, field_simps]: "a + (b - c) = (a + b) - c"
  63.172 -  by (simp add: diff_minus add_assoc)
  63.173 +lemma add_eq_0_iff:
  63.174 +  "a + b = 0 \<longleftrightarrow> b = - a"
  63.175 +  by (auto simp add: neg_eq_iff_add_eq_0 [symmetric])
  63.176  
  63.177 -lemma diff_eq_eq[algebra_simps, field_simps]: "a - b = c \<longleftrightarrow> a = c + b"
  63.178 -  by (auto simp add: diff_minus add_assoc)
  63.179 +lemma minus_diff_eq [simp]:
  63.180 +  "- (a - b) = b - a"
  63.181 +  by (simp only: neg_eq_iff_add_eq_0 diff_conv_add_uminus add_assoc minus_add_cancel) simp
  63.182  
  63.183 -lemma eq_diff_eq[algebra_simps, field_simps]: "a = c - b \<longleftrightarrow> a + b = c"
  63.184 -  by (auto simp add: diff_minus add_assoc)
  63.185 +lemma add_diff_eq [algebra_simps, field_simps]:
  63.186 +  "a + (b - c) = (a + b) - c"
  63.187 +  by (simp only: diff_conv_add_uminus add_assoc)
  63.188  
  63.189 -lemma diff_diff_eq2[algebra_simps, field_simps]: "a - (b - c) = (a + c) - b"
  63.190 -  by (simp add: diff_minus minus_add add_assoc)
  63.191 +lemma diff_add_eq_diff_diff_swap:
  63.192 +  "a - (b + c) = a - c - b"
  63.193 +  by (simp only: diff_conv_add_uminus add_assoc minus_add)
  63.194  
  63.195 -lemma eq_iff_diff_eq_0: "a = b \<longleftrightarrow> a - b = 0"
  63.196 -  by (fact right_minus_eq [symmetric])
  63.197 +lemma diff_eq_eq [algebra_simps, field_simps]:
  63.198 +  "a - b = c \<longleftrightarrow> a = c + b"
  63.199 +  by auto
  63.200 +
  63.201 +lemma eq_diff_eq [algebra_simps, field_simps]:
  63.202 +  "a = c - b \<longleftrightarrow> a + b = c"
  63.203 +  by auto
  63.204 +
  63.205 +lemma diff_diff_eq2 [algebra_simps, field_simps]:
  63.206 +  "a - (b - c) = (a + c) - b"
  63.207 +  by (simp only: diff_conv_add_uminus add_assoc) simp
  63.208  
  63.209  lemma diff_eq_diff_eq:
  63.210    "a - b = c - d \<Longrightarrow> a = b \<longleftrightarrow> c = d"
  63.211 -  by (simp add: eq_iff_diff_eq_0 [of a b] eq_iff_diff_eq_0 [of c d])
  63.212 +  by (simp only: eq_iff_diff_eq_0 [of a b] eq_iff_diff_eq_0 [of c d])
  63.213  
  63.214  end
  63.215  
  63.216  class ab_group_add = minus + uminus + comm_monoid_add +
  63.217    assumes ab_left_minus: "- a + a = 0"
  63.218 -  assumes ab_diff_minus: "a - b = a + (- b)"
  63.219 +  assumes ab_add_uminus_conv_diff: "a - b = a + (- b)"
  63.220  begin
  63.221  
  63.222  subclass group_add
  63.223 -  proof qed (simp_all add: ab_left_minus ab_diff_minus)
  63.224 +  proof qed (simp_all add: ab_left_minus ab_add_uminus_conv_diff)
  63.225  
  63.226  subclass cancel_comm_monoid_add
  63.227  proof
  63.228    fix a b c :: 'a
  63.229    assume "a + b = a + c"
  63.230    then have "- a + a + b = - a + a + c"
  63.231 -    unfolding add_assoc by simp
  63.232 +    by (simp only: add_assoc)
  63.233    then show "b = c" by simp
  63.234  qed
  63.235  
  63.236 -lemma uminus_add_conv_diff[algebra_simps, field_simps]:
  63.237 +lemma uminus_add_conv_diff [simp]:
  63.238    "- a + b = b - a"
  63.239 -by (simp add:diff_minus add_commute)
  63.240 +  by (simp add: add_commute)
  63.241  
  63.242  lemma minus_add_distrib [simp]:
  63.243    "- (a + b) = - a + - b"
  63.244 -by (rule minus_unique) (simp add: add_ac)
  63.245 +  by (simp add: algebra_simps)
  63.246  
  63.247 -lemma diff_add_eq[algebra_simps, field_simps]: "(a - b) + c = (a + c) - b"
  63.248 -by (simp add: diff_minus add_ac)
  63.249 +lemma diff_add_eq [algebra_simps, field_simps]:
  63.250 +  "(a - b) + c = (a + c) - b"
  63.251 +  by (simp add: algebra_simps)
  63.252  
  63.253 -lemma diff_diff_eq[algebra_simps, field_simps]: "(a - b) - c = a - (b + c)"
  63.254 -by (simp add: diff_minus add_ac)
  63.255 +lemma diff_diff_eq [algebra_simps, field_simps]:
  63.256 +  "(a - b) - c = a - (b + c)"
  63.257 +  by (simp add: algebra_simps)
  63.258  
  63.259 -(* FIXME: duplicates right_minus_eq from class group_add *)
  63.260 -(* but only this one is declared as a simp rule. *)
  63.261 -lemma diff_eq_0_iff_eq [simp, no_atp]: "a - b = 0 \<longleftrightarrow> a = b"
  63.262 -  by (rule right_minus_eq)
  63.263 +lemma diff_add_eq_diff_diff:
  63.264 +  "a - (b + c) = a - b - c"
  63.265 +  using diff_add_eq_diff_diff_swap [of a c b] by (simp add: add.commute)
  63.266  
  63.267 -lemma add_diff_cancel_left: "(c + a) - (c + b) = a - b"
  63.268 -  by (simp add: diff_minus add_ac)
  63.269 +lemma add_diff_cancel_left [simp]:
  63.270 +  "(c + a) - (c + b) = a - b"
  63.271 +  by (simp add: algebra_simps)
  63.272  
  63.273  end
  63.274  
  63.275 @@ -622,19 +653,19 @@
  63.276  
  63.277  lemma add_less_cancel_left [simp]:
  63.278    "c + a < c + b \<longleftrightarrow> a < b"
  63.279 -by (blast intro: add_less_imp_less_left add_strict_left_mono) 
  63.280 +  by (blast intro: add_less_imp_less_left add_strict_left_mono) 
  63.281  
  63.282  lemma add_less_cancel_right [simp]:
  63.283    "a + c < b + c \<longleftrightarrow> a < b"
  63.284 -by (blast intro: add_less_imp_less_right add_strict_right_mono)
  63.285 +  by (blast intro: add_less_imp_less_right add_strict_right_mono)
  63.286  
  63.287  lemma add_le_cancel_left [simp]:
  63.288    "c + a \<le> c + b \<longleftrightarrow> a \<le> b"
  63.289 -by (auto, drule add_le_imp_le_left, simp_all add: add_left_mono) 
  63.290 +  by (auto, drule add_le_imp_le_left, simp_all add: add_left_mono) 
  63.291  
  63.292  lemma add_le_cancel_right [simp]:
  63.293    "a + c \<le> b + c \<longleftrightarrow> a \<le> b"
  63.294 -by (simp add: add_commute [of a c] add_commute [of b c])
  63.295 +  by (simp add: add_commute [of a c] add_commute [of b c])
  63.296  
  63.297  lemma add_le_imp_le_right:
  63.298    "a + c \<le> b + c \<Longrightarrow> a \<le> b"
  63.299 @@ -806,6 +837,22 @@
  63.300    then show "x + y = 0" by simp
  63.301  qed
  63.302  
  63.303 +lemma add_increasing:
  63.304 +  "0 \<le> a \<Longrightarrow> b \<le> c \<Longrightarrow> b \<le> a + c"
  63.305 +  by (insert add_mono [of 0 a b c], simp)
  63.306 +
  63.307 +lemma add_increasing2:
  63.308 +  "0 \<le> c \<Longrightarrow> b \<le> a \<Longrightarrow> b \<le> a + c"
  63.309 +  by (simp add: add_increasing add_commute [of a])
  63.310 +
  63.311 +lemma add_strict_increasing:
  63.312 +  "0 < a \<Longrightarrow> b \<le> c \<Longrightarrow> b < a + c"
  63.313 +  by (insert add_less_le_mono [of 0 a b c], simp)
  63.314 +
  63.315 +lemma add_strict_increasing2:
  63.316 +  "0 \<le> a \<Longrightarrow> b < c \<Longrightarrow> b < a + c"
  63.317 +  by (insert add_le_less_mono [of 0 a b c], simp)
  63.318 +
  63.319  end
  63.320  
  63.321  class ordered_ab_group_add =
  63.322 @@ -825,21 +872,53 @@
  63.323  
  63.324  subclass ordered_comm_monoid_add ..
  63.325  
  63.326 +lemma add_less_same_cancel1 [simp]:
  63.327 +  "b + a < b \<longleftrightarrow> a < 0"
  63.328 +  using add_less_cancel_left [of _ _ 0] by simp
  63.329 +
  63.330 +lemma add_less_same_cancel2 [simp]:
  63.331 +  "a + b < b \<longleftrightarrow> a < 0"
  63.332 +  using add_less_cancel_right [of _ _ 0] by simp
  63.333 +
  63.334 +lemma less_add_same_cancel1 [simp]:
  63.335 +  "a < a + b \<longleftrightarrow> 0 < b"
  63.336 +  using add_less_cancel_left [of _ 0] by simp
  63.337 +
  63.338 +lemma less_add_same_cancel2 [simp]:
  63.339 +  "a < b + a \<longleftrightarrow> 0 < b"
  63.340 +  using add_less_cancel_right [of 0] by simp
  63.341 +
  63.342 +lemma add_le_same_cancel1 [simp]:
  63.343 +  "b + a \<le> b \<longleftrightarrow> a \<le> 0"
  63.344 +  using add_le_cancel_left [of _ _ 0] by simp
  63.345 +
  63.346 +lemma add_le_same_cancel2 [simp]:
  63.347 +  "a + b \<le> b \<longleftrightarrow> a \<le> 0"
  63.348 +  using add_le_cancel_right [of _ _ 0] by simp
  63.349 +
  63.350 +lemma le_add_same_cancel1 [simp]:
  63.351 +  "a \<le> a + b \<longleftrightarrow> 0 \<le> b"
  63.352 +  using add_le_cancel_left [of _ 0] by simp
  63.353 +
  63.354 +lemma le_add_same_cancel2 [simp]:
  63.355 +  "a \<le> b + a \<longleftrightarrow> 0 \<le> b"
  63.356 +  using add_le_cancel_right [of 0] by simp
  63.357 +
  63.358  lemma max_diff_distrib_left:
  63.359    shows "max x y - z = max (x - z) (y - z)"
  63.360 -by (simp add: diff_minus, rule max_add_distrib_left) 
  63.361 +  using max_add_distrib_left [of x y "- z"] by simp
  63.362  
  63.363  lemma min_diff_distrib_left:
  63.364    shows "min x y - z = min (x - z) (y - z)"
  63.365 -by (simp add: diff_minus, rule min_add_distrib_left) 
  63.366 +  using min_add_distrib_left [of x y "- z"] by simp
  63.367  
  63.368  lemma le_imp_neg_le:
  63.369    assumes "a \<le> b" shows "-b \<le> -a"
  63.370  proof -
  63.371    have "-a+a \<le> -a+b" using `a \<le> b` by (rule add_left_mono) 
  63.372 -  hence "0 \<le> -a+b" by simp
  63.373 -  hence "0 + (-b) \<le> (-a + b) + (-b)" by (rule add_right_mono) 
  63.374 -  thus ?thesis by (simp add: add_assoc)
  63.375 +  then have "0 \<le> -a+b" by simp
  63.376 +  then have "0 + (-b) \<le> (-a + b) + (-b)" by (rule add_right_mono) 
  63.377 +  then show ?thesis by (simp add: algebra_simps)
  63.378  qed
  63.379  
  63.380  lemma neg_le_iff_le [simp]: "- b \<le> - a \<longleftrightarrow> a \<le> b"
  63.381 @@ -896,35 +975,37 @@
  63.382  lemma minus_le_iff: "- a \<le> b \<longleftrightarrow> - b \<le> a"
  63.383  by (auto simp add: le_less minus_less_iff)
  63.384  
  63.385 -lemma diff_less_0_iff_less [simp, no_atp]:
  63.386 +lemma diff_less_0_iff_less [simp]:
  63.387    "a - b < 0 \<longleftrightarrow> a < b"
  63.388  proof -
  63.389 -  have "a - b < 0 \<longleftrightarrow> a + (- b) < b + (- b)" by (simp add: diff_minus)
  63.390 +  have "a - b < 0 \<longleftrightarrow> a + (- b) < b + (- b)" by simp
  63.391    also have "... \<longleftrightarrow> a < b" by (simp only: add_less_cancel_right)
  63.392    finally show ?thesis .
  63.393  qed
  63.394  
  63.395  lemmas less_iff_diff_less_0 = diff_less_0_iff_less [symmetric]
  63.396  
  63.397 -lemma diff_less_eq[algebra_simps, field_simps]: "a - b < c \<longleftrightarrow> a < c + b"
  63.398 +lemma diff_less_eq [algebra_simps, field_simps]:
  63.399 +  "a - b < c \<longleftrightarrow> a < c + b"
  63.400  apply (subst less_iff_diff_less_0 [of a])
  63.401  apply (rule less_iff_diff_less_0 [of _ c, THEN ssubst])
  63.402 -apply (simp add: diff_minus add_ac)
  63.403 +apply (simp add: algebra_simps)
  63.404  done
  63.405  
  63.406 -lemma less_diff_eq[algebra_simps, field_simps]: "a < c - b \<longleftrightarrow> a + b < c"
  63.407 +lemma less_diff_eq[algebra_simps, field_simps]:
  63.408 +  "a < c - b \<longleftrightarrow> a + b < c"
  63.409  apply (subst less_iff_diff_less_0 [of "a + b"])
  63.410  apply (subst less_iff_diff_less_0 [of a])
  63.411 -apply (simp add: diff_minus add_ac)
  63.412 +apply (simp add: algebra_simps)
  63.413  done
  63.414  
  63.415  lemma diff_le_eq[algebra_simps, field_simps]: "a - b \<le> c \<longleftrightarrow> a \<le> c + b"
  63.416 -by (auto simp add: le_less diff_less_eq diff_add_cancel add_diff_cancel)
  63.417 +by (auto simp add: le_less diff_less_eq )
  63.418  
  63.419  lemma le_diff_eq[algebra_simps, field_simps]: "a \<le> c - b \<longleftrightarrow> a + b \<le> c"
  63.420 -by (auto simp add: le_less less_diff_eq diff_add_cancel add_diff_cancel)
  63.421 +by (auto simp add: le_less less_diff_eq)
  63.422  
  63.423 -lemma diff_le_0_iff_le [simp, no_atp]:
  63.424 +lemma diff_le_0_iff_le [simp]:
  63.425    "a - b \<le> 0 \<longleftrightarrow> a \<le> b"
  63.426    by (simp add: algebra_simps)
  63.427  
  63.428 @@ -992,63 +1073,6 @@
  63.429  
  63.430  subclass linordered_cancel_ab_semigroup_add ..
  63.431  
  63.432 -lemma neg_less_eq_nonneg [simp]:
  63.433 -  "- a \<le> a \<longleftrightarrow> 0 \<le> a"
  63.434 -proof
  63.435 -  assume A: "- a \<le> a" show "0 \<le> a"
  63.436 -  proof (rule classical)
  63.437 -    assume "\<not> 0 \<le> a"
  63.438 -    then have "a < 0" by auto
  63.439 -    with A have "- a < 0" by (rule le_less_trans)
  63.440 -    then show ?thesis by auto
  63.441 -  qed
  63.442 -next
  63.443 -  assume A: "0 \<le> a" show "- a \<le> a"
  63.444 -  proof (rule order_trans)
  63.445 -    show "- a \<le> 0" using A by (simp add: minus_le_iff)
  63.446 -  next
  63.447 -    show "0 \<le> a" using A .
  63.448 -  qed
  63.449 -qed
  63.450 -
  63.451 -lemma neg_less_nonneg [simp]:
  63.452 -  "- a < a \<longleftrightarrow> 0 < a"
  63.453 -proof
  63.454 -  assume A: "- a < a" show "0 < a"
  63.455 -  proof (rule classical)
  63.456 -    assume "\<not> 0 < a"
  63.457 -    then have "a \<le> 0" by auto
  63.458 -    with A have "- a < 0" by (rule less_le_trans)
  63.459 -    then show ?thesis by auto
  63.460 -  qed
  63.461 -next
  63.462 -  assume A: "0 < a" show "- a < a"
  63.463 -  proof (rule less_trans)
  63.464 -    show "- a < 0" using A by (simp add: minus_le_iff)
  63.465 -  next
  63.466 -    show "0 < a" using A .
  63.467 -  qed
  63.468 -qed
  63.469 -
  63.470 -lemma less_eq_neg_nonpos [simp]:
  63.471 -  "a \<le> - a \<longleftrightarrow> a \<le> 0"
  63.472 -proof
  63.473 -  assume A: "a \<le> - a" show "a \<le> 0"
  63.474 -  proof (rule classical)
  63.475 -    assume "\<not> a \<le> 0"
  63.476 -    then have "0 < a" by auto
  63.477 -    then have "0 < - a" using A by (rule less_le_trans)
  63.478 -    then show ?thesis by auto
  63.479 -  qed
  63.480 -next
  63.481 -  assume A: "a \<le> 0" show "a \<le> - a"
  63.482 -  proof (rule order_trans)
  63.483 -    show "0 \<le> - a" using A by (simp add: minus_le_iff)
  63.484 -  next
  63.485 -    show "a \<le> 0" using A .
  63.486 -  qed
  63.487 -qed
  63.488 -
  63.489  lemma equal_neg_zero [simp]:
  63.490    "a = - a \<longleftrightarrow> a = 0"
  63.491  proof
  63.492 @@ -1070,6 +1094,37 @@
  63.493    "- a = a \<longleftrightarrow> a = 0"
  63.494    by (auto dest: sym)
  63.495  
  63.496 +lemma neg_less_eq_nonneg [simp]:
  63.497 +  "- a \<le> a \<longleftrightarrow> 0 \<le> a"
  63.498 +proof
  63.499 +  assume A: "- a \<le> a" show "0 \<le> a"
  63.500 +  proof (rule classical)
  63.501 +    assume "\<not> 0 \<le> a"
  63.502 +    then have "a < 0" by auto
  63.503 +    with A have "- a < 0" by (rule le_less_trans)
  63.504 +    then show ?thesis by auto
  63.505 +  qed
  63.506 +next
  63.507 +  assume A: "0 \<le> a" show "- a \<le> a"
  63.508 +  proof (rule order_trans)
  63.509 +    show "- a \<le> 0" using A by (simp add: minus_le_iff)
  63.510 +  next
  63.511 +    show "0 \<le> a" using A .
  63.512 +  qed
  63.513 +qed
  63.514 +
  63.515 +lemma neg_less_pos [simp]:
  63.516 +  "- a < a \<longleftrightarrow> 0 < a"
  63.517 +  by (auto simp add: less_le)
  63.518 +
  63.519 +lemma less_eq_neg_nonpos [simp]:
  63.520 +  "a \<le> - a \<longleftrightarrow> a \<le> 0"
  63.521 +  using neg_less_eq_nonneg [of "- a"] by simp
  63.522 +
  63.523 +lemma less_neg_neg [simp]:
  63.524 +  "a < - a \<longleftrightarrow> a < 0"
  63.525 +  using neg_less_pos [of "- a"] by simp
  63.526 +
  63.527  lemma double_zero [simp]:
  63.528    "a + a = 0 \<longleftrightarrow> a = 0"
  63.529  proof
  63.530 @@ -1088,7 +1143,7 @@
  63.531    assume "0 < a + a"
  63.532    then have "0 - a < a" by (simp only: diff_less_eq)
  63.533    then have "- a < a" by simp
  63.534 -  then show "0 < a" by (simp only: neg_less_nonneg)
  63.535 +  then show "0 < a" by simp
  63.536  next
  63.537    assume "0 < a"
  63.538    with this have "0 + 0 < a + a"
  63.539 @@ -1116,24 +1171,6 @@
  63.540    then show ?thesis by simp
  63.541  qed
  63.542  
  63.543 -lemma le_minus_self_iff:
  63.544 -  "a \<le> - a \<longleftrightarrow> a \<le> 0"
  63.545 -proof -
  63.546 -  from add_le_cancel_left [of "- a" "a + a" 0]
  63.547 -  have "a \<le> - a \<longleftrightarrow> a + a \<le> 0" 
  63.548 -    by (simp add: add_assoc [symmetric])
  63.549 -  thus ?thesis by simp
  63.550 -qed
  63.551 -
  63.552 -lemma minus_le_self_iff:
  63.553 -  "- a \<le> a \<longleftrightarrow> 0 \<le> a"
  63.554 -proof -
  63.555 -  from add_le_cancel_left [of "- a" 0 "a + a"]
  63.556 -  have "- a \<le> a \<longleftrightarrow> 0 \<le> a + a" 
  63.557 -    by (simp add: add_assoc [symmetric])
  63.558 -  thus ?thesis by simp
  63.559 -qed
  63.560 -
  63.561  lemma minus_max_eq_min:
  63.562    "- max x y = min (-x) (-y)"
  63.563    by (auto simp add: max_def min_def)
  63.564 @@ -1144,27 +1181,6 @@
  63.565  
  63.566  end
  63.567  
  63.568 -context ordered_comm_monoid_add
  63.569 -begin
  63.570 -
  63.571 -lemma add_increasing:
  63.572 -  "0 \<le> a \<Longrightarrow> b \<le> c \<Longrightarrow> b \<le> a + c"
  63.573 -  by (insert add_mono [of 0 a b c], simp)
  63.574 -
  63.575 -lemma add_increasing2:
  63.576 -  "0 \<le> c \<Longrightarrow> b \<le> a \<Longrightarrow> b \<le> a + c"
  63.577 -  by (simp add: add_increasing add_commute [of a])
  63.578 -
  63.579 -lemma add_strict_increasing:
  63.580 -  "0 < a \<Longrightarrow> b \<le> c \<Longrightarrow> b < a + c"
  63.581 -  by (insert add_less_le_mono [of 0 a b c], simp)
  63.582 -
  63.583 -lemma add_strict_increasing2:
  63.584 -  "0 \<le> a \<Longrightarrow> b < c \<Longrightarrow> b < a + c"
  63.585 -  by (insert add_le_less_mono [of 0 a b c], simp)
  63.586 -
  63.587 -end
  63.588 -
  63.589  class abs =
  63.590    fixes abs :: "'a \<Rightarrow> 'a"
  63.591  begin
  63.592 @@ -1231,7 +1247,7 @@
  63.593  lemma abs_zero [simp]: "\<bar>0\<bar> = 0"
  63.594  by simp
  63.595  
  63.596 -lemma abs_0_eq [simp, no_atp]: "0 = \<bar>a\<bar> \<longleftrightarrow> a = 0"
  63.597 +lemma abs_0_eq [simp]: "0 = \<bar>a\<bar> \<longleftrightarrow> a = 0"
  63.598  proof -
  63.599    have "0 = \<bar>a\<bar> \<longleftrightarrow> \<bar>a\<bar> = 0" by (simp only: eq_ac)
  63.600    thus ?thesis by simp
  63.601 @@ -1299,7 +1315,7 @@
  63.602  lemma abs_triangle_ineq2: "\<bar>a\<bar> - \<bar>b\<bar> \<le> \<bar>a - b\<bar>"
  63.603  proof -
  63.604    have "\<bar>a\<bar> = \<bar>b + (a - b)\<bar>"
  63.605 -    by (simp add: algebra_simps add_diff_cancel)
  63.606 +    by (simp add: algebra_simps)
  63.607    then have "\<bar>a\<bar> \<le> \<bar>b\<bar> + \<bar>a - b\<bar>"
  63.608      by (simp add: abs_triangle_ineq)
  63.609    then show ?thesis
  63.610 @@ -1314,14 +1330,14 @@
  63.611  
  63.612  lemma abs_triangle_ineq4: "\<bar>a - b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
  63.613  proof -
  63.614 -  have "\<bar>a - b\<bar> = \<bar>a + - b\<bar>" by (subst diff_minus, rule refl)
  63.615 +  have "\<bar>a - b\<bar> = \<bar>a + - b\<bar>" by (simp add: algebra_simps)
  63.616    also have "... \<le> \<bar>a\<bar> + \<bar>- b\<bar>" by (rule abs_triangle_ineq)
  63.617    finally show ?thesis by simp
  63.618  qed
  63.619  
  63.620  lemma abs_diff_triangle_ineq: "\<bar>a + b - (c + d)\<bar> \<le> \<bar>a - c\<bar> + \<bar>b - d\<bar>"
  63.621  proof -
  63.622 -  have "\<bar>a + b - (c+d)\<bar> = \<bar>(a-c) + (b-d)\<bar>" by (simp add: diff_minus add_ac)
  63.623 +  have "\<bar>a + b - (c+d)\<bar> = \<bar>(a-c) + (b-d)\<bar>" by (simp add: algebra_simps)
  63.624    also have "... \<le> \<bar>a-c\<bar> + \<bar>b-d\<bar>" by (rule abs_triangle_ineq)
  63.625    finally show ?thesis .
  63.626  qed
  63.627 @@ -1341,7 +1357,7 @@
  63.628  
  63.629  subsection {* Tools setup *}
  63.630  
  63.631 -lemma add_mono_thms_linordered_semiring [no_atp]:
  63.632 +lemma add_mono_thms_linordered_semiring:
  63.633    fixes i j k :: "'a\<Colon>ordered_ab_semigroup_add"
  63.634    shows "i \<le> j \<and> k \<le> l \<Longrightarrow> i + k \<le> j + l"
  63.635      and "i = j \<and> k \<le> l \<Longrightarrow> i + k \<le> j + l"
  63.636 @@ -1349,7 +1365,7 @@
  63.637      and "i = j \<and> k = l \<Longrightarrow> i + k = j + l"
  63.638  by (rule add_mono, clarify+)+
  63.639  
  63.640 -lemma add_mono_thms_linordered_field [no_atp]:
  63.641 +lemma add_mono_thms_linordered_field:
  63.642    fixes i j k :: "'a\<Colon>ordered_cancel_ab_semigroup_add"
  63.643    shows "i < j \<and> k = l \<Longrightarrow> i + k < j + l"
  63.644      and "i = j \<and> k < l \<Longrightarrow> i + k < j + l"
  63.645 @@ -1362,10 +1378,5 @@
  63.646  code_identifier
  63.647    code_module Groups \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
  63.648  
  63.649 -
  63.650 -text {* Legacy *}
  63.651 -
  63.652 -lemmas diff_def = diff_minus
  63.653 -
  63.654  end
  63.655  
    64.1 --- a/src/HOL/Hahn_Banach/Bounds.thy	Mon Nov 11 17:34:44 2013 +0100
    64.2 +++ b/src/HOL/Hahn_Banach/Bounds.thy	Mon Nov 11 17:44:21 2013 +0100
    64.3 @@ -57,25 +57,7 @@
    64.4    finally show ?thesis .
    64.5  qed
    64.6  
    64.7 -lemma lub_compat: "lub A x = isLub UNIV A x"
    64.8 -proof -
    64.9 -  have "isUb UNIV A = (\<lambda>x. A *<= x \<and> x \<in> UNIV)"
   64.10 -    by (rule ext) (simp only: isUb_def)
   64.11 -  then show ?thesis
   64.12 -    by (simp only: lub_def isLub_def leastP_def setge_def setle_def) blast
   64.13 -qed
   64.14 -
   64.15 -lemma real_complete:
   64.16 -  fixes A :: "real set"
   64.17 -  assumes nonempty: "\<exists>a. a \<in> A"
   64.18 -    and ex_upper: "\<exists>y. \<forall>a \<in> A. a \<le> y"
   64.19 -  shows "\<exists>x. lub A x"
   64.20 -proof -
   64.21 -  from ex_upper have "\<exists>y. isUb UNIV A y"
   64.22 -    unfolding isUb_def setle_def by blast
   64.23 -  with nonempty have "\<exists>x. isLub UNIV A x"
   64.24 -    by (rule reals_complete)
   64.25 -  then show ?thesis by (simp only: lub_compat)
   64.26 -qed
   64.27 +lemma real_complete: "\<exists>a::real. a \<in> A \<Longrightarrow> \<exists>y. \<forall>a \<in> A. a \<le> y \<Longrightarrow> \<exists>x. lub A x"
   64.28 +  by (intro exI[of _ "Sup A"]) (auto intro!: cSup_upper cSup_least simp: lub_def)
   64.29  
   64.30  end
    65.1 --- a/src/HOL/Hahn_Banach/Vector_Space.thy	Mon Nov 11 17:34:44 2013 +0100
    65.2 +++ b/src/HOL/Hahn_Banach/Vector_Space.thy	Mon Nov 11 17:44:21 2013 +0100
    65.3 @@ -112,7 +112,7 @@
    65.4  proof -
    65.5    assume x: "x \<in> V"
    65.6    have " (a - b) \<cdot> x = (a + - b) \<cdot> x"
    65.7 -    by (simp add: diff_minus)
    65.8 +    by simp
    65.9    also from x have "\<dots> = a \<cdot> x + (- b) \<cdot> x"
   65.10      by (rule add_mult_distrib2)
   65.11    also from x have "\<dots> = a \<cdot> x + - (b \<cdot> x)"
    66.1 --- a/src/HOL/Hilbert_Choice.thy	Mon Nov 11 17:34:44 2013 +0100
    66.2 +++ b/src/HOL/Hilbert_Choice.thy	Mon Nov 11 17:44:21 2013 +0100
    66.3 @@ -741,7 +741,7 @@
    66.4  | "bacc r (Suc n) = (bacc r n \<union> {x. \<forall>y. (y, x) \<in> r \<longrightarrow> y \<in> bacc r n})"
    66.5  
    66.6  lemma bacc_subseteq_acc:
    66.7 -  "bacc r n \<subseteq> acc r"
    66.8 +  "bacc r n \<subseteq> Wellfounded.acc r"
    66.9    by (induct n) (auto intro: acc.intros)
   66.10  
   66.11  lemma bacc_mono:
   66.12 @@ -761,10 +761,10 @@
   66.13  
   66.14  lemma acc_subseteq_bacc:
   66.15    assumes "finite r"
   66.16 -  shows "acc r \<subseteq> (\<Union>n. bacc r n)"
   66.17 +  shows "Wellfounded.acc r \<subseteq> (\<Union>n. bacc r n)"
   66.18  proof
   66.19    fix x
   66.20 -  assume "x : acc r"
   66.21 +  assume "x : Wellfounded.acc r"
   66.22    then have "\<exists> n. x : bacc r n"
   66.23    proof (induct x arbitrary: rule: acc.induct)
   66.24      case (accI x)
   66.25 @@ -788,7 +788,7 @@
   66.26  lemma acc_bacc_eq:
   66.27    fixes A :: "('a :: finite \<times> 'a) set"
   66.28    assumes "finite A"
   66.29 -  shows "acc A = bacc A (card (UNIV :: 'a set))"
   66.30 +  shows "Wellfounded.acc A = bacc A (card (UNIV :: 'a set))"
   66.31    using assms by (metis acc_subseteq_bacc bacc_subseteq_acc bacc_upper_bound order_eq_iff)
   66.32  
   66.33  
    67.1 --- a/src/HOL/IMP/AExp.thy	Mon Nov 11 17:34:44 2013 +0100
    67.2 +++ b/src/HOL/IMP/AExp.thy	Mon Nov 11 17:44:21 2013 +0100
    67.3 @@ -33,11 +33,12 @@
    67.4    "_State" :: "updbinds => 'a" ("<_>")
    67.5  translations
    67.6    "_State ms" == "_Update <> ms"
    67.7 +  "_State (_updbinds b bs)" <= "_Update (_State b) bs"
    67.8  
    67.9  text {* \noindent
   67.10    We can now write a series of updates to the function @{text "\<lambda>x. 0"} compactly:
   67.11  *}
   67.12 -lemma "<a := Suc 0, b := 2> = (<> (a := Suc 0)) (b := 2)"
   67.13 +lemma "<a := 1, b := 2> = (<> (a := 1)) (b := (2::int))"
   67.14    by (rule refl)
   67.15  
   67.16  value "aval (Plus (V ''x'') (N 5)) <''x'' := 7>"
    68.1 --- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_ITP.thy	Mon Nov 11 17:34:44 2013 +0100
    68.2 +++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_ITP.thy	Mon Nov 11 17:44:21 2013 +0100
    68.3 @@ -135,8 +135,6 @@
    68.4  
    68.5  subsubsection "Ascending Chain Condition"
    68.6  
    68.7 -hide_const (open) acc
    68.8 -
    68.9  abbreviation "strict r == r \<inter> -(r^-1)"
   68.10  abbreviation "acc r == wf((strict r)^-1)"
   68.11  
    69.1 --- a/src/HOL/IMP/Big_Step.thy	Mon Nov 11 17:34:44 2013 +0100
    69.2 +++ b/src/HOL/IMP/Big_Step.thy	Mon Nov 11 17:44:21 2013 +0100
    69.3 @@ -268,11 +268,9 @@
    69.4  subsection "Execution is deterministic"
    69.5  
    69.6  text {* This proof is automatic. *}
    69.7 -text_raw{*\snip{BigStepDeterministic}{0}{1}{% *}
    69.8 +
    69.9  theorem big_step_determ: "\<lbrakk> (c,s) \<Rightarrow> t; (c,s) \<Rightarrow> u \<rbrakk> \<Longrightarrow> u = t"
   69.10    by (induction arbitrary: u rule: big_step.induct) blast+
   69.11 -text_raw{*}%endsnip*}
   69.12 -
   69.13  
   69.14  text {*
   69.15    This is the proof as you might present it in a lecture. The remaining
    70.1 --- a/src/HOL/IMP/Fold.thy	Mon Nov 11 17:34:44 2013 +0100
    70.2 +++ b/src/HOL/IMP/Fold.thy	Mon Nov 11 17:44:21 2013 +0100
    70.3 @@ -1,5 +1,3 @@
    70.4 -header "Constant Folding"
    70.5 -
    70.6  theory Fold imports Sem_Equiv Vars begin
    70.7  
    70.8  subsection "Simple folding of arithmetic expressions"
    71.1 --- a/src/HOL/IMP/Sem_Equiv.thy	Mon Nov 11 17:34:44 2013 +0100
    71.2 +++ b/src/HOL/IMP/Sem_Equiv.thy	Mon Nov 11 17:44:21 2013 +0100
    71.3 @@ -1,9 +1,11 @@
    71.4 -header "Semantic Equivalence up to a Condition"
    71.5 +header "Constant Folding"
    71.6  
    71.7  theory Sem_Equiv
    71.8  imports Big_Step
    71.9  begin
   71.10  
   71.11 +subsection "Semantic Equivalence up to a Condition"
   71.12 +
   71.13  type_synonym assn = "state \<Rightarrow> bool"
   71.14  
   71.15  definition
    72.1 --- a/src/HOL/IMP/Small_Step.thy	Mon Nov 11 17:34:44 2013 +0100
    72.2 +++ b/src/HOL/IMP/Small_Step.thy	Mon Nov 11 17:44:21 2013 +0100
    72.3 @@ -4,7 +4,6 @@
    72.4  
    72.5  subsection "The transition relation"
    72.6  
    72.7 -text_raw{*\snip{SmallStepDef}{0}{2}{% *}
    72.8  inductive
    72.9    small_step :: "com * state \<Rightarrow> com * state \<Rightarrow> bool" (infix "\<rightarrow>" 55)
   72.10  where
   72.11 @@ -18,7 +17,6 @@
   72.12  
   72.13  While:   "(WHILE b DO c,s) \<rightarrow>
   72.14              (IF b THEN c;; WHILE b DO c ELSE SKIP,s)"
   72.15 -text_raw{*}%endsnip*}
   72.16  
   72.17  
   72.18  abbreviation
    73.1 --- a/src/HOL/Int.thy	Mon Nov 11 17:34:44 2013 +0100
    73.2 +++ b/src/HOL/Int.thy	Mon Nov 11 17:44:21 2013 +0100
    73.3 @@ -220,7 +220,7 @@
    73.4    by (transfer fixing: uminus) clarsimp
    73.5  
    73.6  lemma of_int_diff [simp]: "of_int (w - z) = of_int w - of_int z"
    73.7 -by (simp add: diff_minus Groups.diff_minus)
    73.8 +  using of_int_add [of w "- z"] by simp
    73.9  
   73.10  lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z"
   73.11    by (transfer fixing: times) (clarsimp simp add: algebra_simps of_nat_mult)
   73.12 @@ -349,12 +349,33 @@
   73.13    shows P
   73.14    using assms by (blast dest: nat_0_le sym)
   73.15  
   73.16 -lemma nat_eq_iff: "(nat w = m) = (if 0 \<le> w then w = int m else m=0)"
   73.17 +lemma nat_eq_iff:
   73.18 +  "nat w = m \<longleftrightarrow> (if 0 \<le> w then w = int m else m = 0)"
   73.19    by transfer (clarsimp simp add: le_imp_diff_is_add)
   73.20 + 
   73.21 +corollary nat_eq_iff2:
   73.22 +  "m = nat w \<longleftrightarrow> (if 0 \<le> w then w = int m else m = 0)"
   73.23 +  using nat_eq_iff [of w m] by auto
   73.24  
   73.25 -corollary nat_eq_iff2: "(m = nat w) = (if 0 \<le> w then w = int m else m=0)"
   73.26 -by (simp only: eq_commute [of m] nat_eq_iff)
   73.27 +lemma nat_0 [simp]:
   73.28 +  "nat 0 = 0"
   73.29 +  by (simp add: nat_eq_iff)
   73.30  
   73.31 +lemma nat_1 [simp]:
   73.32 +  "nat 1 = Suc 0"
   73.33 +  by (simp add: nat_eq_iff)
   73.34 +
   73.35 +lemma nat_numeral [simp]:
   73.36 +  "nat (numeral k) = numeral k"
   73.37 +  by (simp add: nat_eq_iff)
   73.38 +
   73.39 +lemma nat_neg_numeral [simp]:
   73.40 +  "nat (neg_numeral k) = 0"
   73.41 +  by simp
   73.42 +
   73.43 +lemma nat_2: "nat 2 = Suc (Suc 0)"
   73.44 +  by simp
   73.45 + 
   73.46  lemma nat_less_iff: "0 \<le> w ==> (nat w < m) = (w < of_nat m)"
   73.47    by transfer (clarsimp, arith)
   73.48  
   73.49 @@ -374,12 +395,16 @@
   73.50  by (insert zless_nat_conj [of 0], auto)
   73.51  
   73.52  lemma nat_add_distrib:
   73.53 -     "[| (0::int) \<le> z;  0 \<le> z' |] ==> nat (z+z') = nat z + nat z'"
   73.54 +  "0 \<le> z \<Longrightarrow> 0 \<le> z' \<Longrightarrow> nat (z + z') = nat z + nat z'"
   73.55    by transfer clarsimp
   73.56  
   73.57 +lemma nat_diff_distrib':
   73.58 +  "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> nat (x - y) = nat x - nat y"
   73.59 +  by transfer clarsimp
   73.60 + 
   73.61  lemma nat_diff_distrib:
   73.62 -     "[| (0::int) \<le> z';  z' \<le> z |] ==> nat (z-z') = nat z - nat z'"
   73.63 -  by transfer clarsimp
   73.64 +  "0 \<le> z' \<Longrightarrow> z' \<le> z \<Longrightarrow> nat (z - z') = nat z - nat z'"
   73.65 +  by (rule nat_diff_distrib') auto
   73.66  
   73.67  lemma nat_zminus_int [simp]: "nat (- int n) = 0"
   73.68    by transfer simp
   73.69 @@ -399,6 +424,11 @@
   73.70  
   73.71  end
   73.72  
   73.73 +lemma diff_nat_numeral [simp]: 
   73.74 +  "(numeral v :: nat) - numeral v' = nat (numeral v - numeral v')"
   73.75 +  by (simp only: nat_diff_distrib' zero_le_numeral nat_numeral)
   73.76 +
   73.77 +
   73.78  text {* For termination proofs: *}
   73.79  lemma measure_function_int[measure_function]: "is_measure (nat o abs)" ..
   73.80  
   73.81 @@ -450,7 +480,7 @@
   73.82        It is proved here because attribute @{text arith_split} is not available
   73.83        in theory @{text Rings}.
   73.84        But is it really better than just rewriting with @{text abs_if}?*}
   73.85 -lemma abs_split [arith_split,no_atp]:
   73.86 +lemma abs_split [arith_split, no_atp]:
   73.87       "P(abs(a::'a::linordered_idom)) = ((0 \<le> a --> P a) & (a < 0 --> P(-a)))"
   73.88  by (force dest: order_less_le_trans simp add: abs_if linorder_not_less)
   73.89  
   73.90 @@ -722,14 +752,11 @@
   73.91  
   73.92  subsection {* Setting up simplification procedures *}
   73.93  
   73.94 +lemmas of_int_simps =
   73.95 +  of_int_0 of_int_1 of_int_add of_int_mult
   73.96 +
   73.97  lemmas int_arith_rules =
   73.98 -  neg_le_iff_le numeral_One
   73.99 -  minus_zero diff_minus left_minus right_minus
  73.100 -  mult_zero_left mult_zero_right mult_1_left mult_1_right
  73.101 -  mult_minus_left mult_minus_right
  73.102 -  minus_add_distrib minus_minus mult_assoc
  73.103 -  of_nat_0 of_nat_1 of_nat_Suc of_nat_add of_nat_mult
  73.104 -  of_int_0 of_int_1 of_int_add of_int_mult
  73.105 +  numeral_One more_arith_simps of_nat_simps of_int_simps
  73.106  
  73.107  ML_file "Tools/int_arith.ML"
  73.108  declaration {* K Int_Arith.setup *}
  73.109 @@ -768,16 +795,6 @@
  73.110  subsection{*The functions @{term nat} and @{term int}*}
  73.111  
  73.112  text{*Simplify the term @{term "w + - z"}*}
  73.113 -lemmas diff_int_def_symmetric = diff_def [where 'a=int, symmetric, simp]
  73.114 -
  73.115 -lemma nat_0 [simp]: "nat 0 = 0"
  73.116 -by (simp add: nat_eq_iff)
  73.117 -
  73.118 -lemma nat_1 [simp]: "nat 1 = Suc 0"
  73.119 -by (subst nat_eq_iff, simp)
  73.120 -
  73.121 -lemma nat_2: "nat 2 = Suc (Suc 0)"
  73.122 -by (subst nat_eq_iff, simp)
  73.123  
  73.124  lemma one_less_nat_eq [simp]: "(Suc 0 < nat z) = (1 < z)"
  73.125  apply (insert zless_nat_conj [of 1 z])
  73.126 @@ -860,31 +877,10 @@
  73.127                if d < 0 then 0 else nat d)"
  73.128  by (simp add: Let_def nat_diff_distrib [symmetric])
  73.129  
  73.130 -(* nat_diff_distrib has too-strong premises *)
  73.131 -lemma nat_diff_distrib': "\<lbrakk>0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> nat (x - y) = nat x - nat y"
  73.132 -apply (rule int_int_eq [THEN iffD1], clarsimp)
  73.133 -apply (subst of_nat_diff)
  73.134 -apply (rule nat_mono, simp_all)
  73.135 -done
  73.136 -
  73.137 -lemma nat_numeral [simp]:
  73.138 -  "nat (numeral k) = numeral k"
  73.139 -  by (simp add: nat_eq_iff)
  73.140 -
  73.141 -lemma nat_neg_numeral [simp]:
  73.142 -  "nat (neg_numeral k) = 0"
  73.143 -  by simp
  73.144 -
  73.145 -lemma diff_nat_numeral [simp]: 
  73.146 -  "(numeral v :: nat) - numeral v' = nat (numeral v - numeral v')"
  73.147 -  by (simp only: nat_diff_distrib' zero_le_numeral nat_numeral)
  73.148 -
  73.149  lemma nat_numeral_diff_1 [simp]:
  73.150    "numeral v - (1::nat) = nat (numeral v - 1)"
  73.151    using diff_nat_numeral [of v Num.One] by simp
  73.152  
  73.153 -lemmas nat_arith = diff_nat_numeral
  73.154 -
  73.155  
  73.156  subsection "Induction principles for int"
  73.157  
  73.158 @@ -1158,32 +1154,32 @@
  73.159  
  73.160  text{*To Simplify Inequalities Where One Side is the Constant 1*}
  73.161  
  73.162 -lemma less_minus_iff_1 [simp,no_atp]:
  73.163 +lemma less_minus_iff_1 [simp]:
  73.164    fixes b::"'b::linordered_idom"
  73.165    shows "(1 < - b) = (b < -1)"
  73.166  by auto
  73.167  
  73.168 -lemma le_minus_iff_1 [simp,no_atp]:
  73.169 +lemma le_minus_iff_1 [simp]:
  73.170    fixes b::"'b::linordered_idom"
  73.171    shows "(1 \<le> - b) = (b \<le> -1)"
  73.172  by auto
  73.173  
  73.174 -lemma equation_minus_iff_1 [simp,no_atp]:
  73.175 +lemma equation_minus_iff_1 [simp]:
  73.176    fixes b::"'b::ring_1"
  73.177    shows "(1 = - b) = (b = -1)"
  73.178  by (subst equation_minus_iff, auto)
  73.179  
  73.180 -lemma minus_less_iff_1 [simp,no_atp]:
  73.181 +lemma minus_less_iff_1 [simp]:
  73.182    fixes a::"'b::linordered_idom"
  73.183    shows "(- a < 1) = (-1 < a)"
  73.184  by auto
  73.185  
  73.186 -lemma minus_le_iff_1 [simp,no_atp]:
  73.187 +lemma minus_le_iff_1 [simp]:
  73.188    fixes a::"'b::linordered_idom"
  73.189    shows "(- a \<le> 1) = (-1 \<le> a)"
  73.190  by auto
  73.191  
  73.192 -lemma minus_equation_iff_1 [simp,no_atp]:
  73.193 +lemma minus_equation_iff_1 [simp]:
  73.194    fixes a::"'b::ring_1"
  73.195    shows "(- a = 1) = (a = -1)"
  73.196  by (subst minus_equation_iff, auto)
  73.197 @@ -1509,10 +1505,13 @@
  73.198    "sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)"
  73.199    "sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1"
  73.200    "sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1"
  73.201 -  unfolding sub_def dup_def numeral.simps Pos_def Neg_def
  73.202 -    neg_numeral_def numeral_BitM
  73.203 -  by (simp_all only: algebra_simps)
  73.204 -
  73.205 +  apply (simp_all only: sub_def dup_def numeral.simps Pos_def Neg_def
  73.206 +    neg_numeral_def numeral_BitM)
  73.207 +  apply (simp_all only: algebra_simps minus_diff_eq)
  73.208 +  apply (simp_all only: add.commute [of _ "- (numeral n + numeral n)"])
  73.209 +  apply (simp_all only: minus_add add.assoc left_minus)
  73.210 +  apply (simp_all only: algebra_simps right_minus)
  73.211 +  done
  73.212  
  73.213  text {* Implementations *}
  73.214  
    74.1 --- a/src/HOL/Library/Abstract_Rat.thy	Mon Nov 11 17:34:44 2013 +0100
    74.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    74.3 @@ -1,521 +0,0 @@
    74.4 -(*  Title:      HOL/Library/Abstract_Rat.thy
    74.5 -    Author:     Amine Chaieb
    74.6 -*)
    74.7 -
    74.8 -header {* Abstract rational numbers *}
    74.9 -
   74.10 -theory Abstract_Rat
   74.11 -imports Complex_Main
   74.12 -begin
   74.13 -
   74.14 -type_synonym Num = "int \<times> int"
   74.15 -
   74.16 -abbreviation Num0_syn :: Num  ("0\<^sub>N")
   74.17 -  where "0\<^sub>N \<equiv> (0, 0)"
   74.18 -
   74.19 -abbreviation Numi_syn :: "int \<Rightarrow> Num"  ("'((_)')\<^sub>N")
   74.20 -  where "(i)\<^sub>N \<equiv> (i, 1)"
   74.21 -
   74.22 -definition isnormNum :: "Num \<Rightarrow> bool" where
   74.23 -  "isnormNum = (\<lambda>(a,b). (if a = 0 then b = 0 else b > 0 \<and> gcd a b = 1))"
   74.24 -
   74.25 -definition normNum :: "Num \<Rightarrow> Num" where
   74.26 -  "normNum = (\<lambda>(a,b).
   74.27 -    (if a=0 \<or> b = 0 then (0,0) else
   74.28 -      (let g = gcd a b
   74.29 -       in if b > 0 then (a div g, b div g) else (- (a div g), - (b div g)))))"
   74.30 -
   74.31 -declare gcd_dvd1_int[presburger] gcd_dvd2_int[presburger]
   74.32 -
   74.33 -lemma normNum_isnormNum [simp]: "isnormNum (normNum x)"
   74.34 -proof -
   74.35 -  obtain a b where x: "x = (a, b)" by (cases x)
   74.36 -  { assume "a=0 \<or> b = 0" hence ?thesis by (simp add: x normNum_def isnormNum_def) }
   74.37 -  moreover
   74.38 -  { assume anz: "a \<noteq> 0" and bnz: "b \<noteq> 0"
   74.39 -    let ?g = "gcd a b"
   74.40 -    let ?a' = "a div ?g"
   74.41 -    let ?b' = "b div ?g"
   74.42 -    let ?g' = "gcd ?a' ?b'"
   74.43 -    from anz bnz have "?g \<noteq> 0" by simp  with gcd_ge_0_int[of a b]
   74.44 -    have gpos: "?g > 0" by arith
   74.45 -    have gdvd: "?g dvd a" "?g dvd b" by arith+
   74.46 -    from dvd_mult_div_cancel[OF gdvd(1)] dvd_mult_div_cancel[OF gdvd(2)] anz bnz
   74.47 -    have nz': "?a' \<noteq> 0" "?b' \<noteq> 0" by - (rule notI, simp)+
   74.48 -    from anz bnz have stupid: "a \<noteq> 0 \<or> b \<noteq> 0" by arith
   74.49 -    from div_gcd_coprime_int[OF stupid] have gp1: "?g' = 1" .
   74.50 -    from bnz have "b < 0 \<or> b > 0" by arith
   74.51 -    moreover
   74.52 -    { assume b: "b > 0"
   74.53 -      from b have "?b' \<ge> 0"
   74.54 -        by (presburger add: pos_imp_zdiv_nonneg_iff[OF gpos])
   74.55 -      with nz' have b': "?b' > 0" by arith
   74.56 -      from b b' anz bnz nz' gp1 have ?thesis
   74.57 -        by (simp add: x isnormNum_def normNum_def Let_def split_def) }
   74.58 -    moreover {
   74.59 -      assume b: "b < 0"
   74.60 -      { assume b': "?b' \<ge> 0"
   74.61 -        from gpos have th: "?g \<ge> 0" by arith
   74.62 -        from mult_nonneg_nonneg[OF th b'] dvd_mult_div_cancel[OF gdvd(2)]
   74.63 -        have False using b by arith }
   74.64 -      hence b': "?b' < 0" by (presburger add: linorder_not_le[symmetric])
   74.65 -      from anz bnz nz' b b' gp1 have ?thesis
   74.66 -        by (simp add: x isnormNum_def normNum_def Let_def split_def) }
   74.67 -    ultimately have ?thesis by blast
   74.68 -  }
   74.69 -  ultimately show ?thesis by blast
   74.70 -qed
   74.71 -
   74.72 -text {* Arithmetic over Num *}
   74.73 -
   74.74 -definition Nadd :: "Num \<Rightarrow> Num \<Rightarrow> Num"  (infixl "+\<^sub>N" 60) where
   74.75 -  "Nadd = (\<lambda>(a,b) (a',b'). if a = 0 \<or> b = 0 then normNum(a',b')
   74.76 -    else if a'=0 \<or> b' = 0 then normNum(a,b)
   74.77 -    else normNum(a*b' + b*a', b*b'))"
   74.78 -
   74.79 -definition Nmul :: "Num \<Rightarrow> Num \<Rightarrow> Num"  (infixl "*\<^sub>N" 60) where
   74.80 -  "Nmul = (\<lambda>(a,b) (a',b'). let g = gcd (a*a') (b*b')
   74.81 -    in (a*a' div g, b*b' div g))"
   74.82 -
   74.83 -definition Nneg :: "Num \<Rightarrow> Num" ("~\<^sub>N")
   74.84 -  where "Nneg \<equiv> (\<lambda>(a,b). (-a,b))"
   74.85 -
   74.86 -definition Nsub :: "Num \<Rightarrow> Num \<Rightarrow> Num"  (infixl "-\<^sub>N" 60)
   74.87 -  where "Nsub = (\<lambda>a b. a +\<^sub>N ~\<^sub>N b)"
   74.88 -
   74.89 -definition Ninv :: "Num \<Rightarrow> Num"
   74.90 -  where "Ninv = (\<lambda>(a,b). if a < 0 then (-b, \<bar>a\<bar>) else (b,a))"
   74.91 -
   74.92 -definition Ndiv :: "Num \<Rightarrow> Num \<Rightarrow> Num"  (infixl "\<div>\<^sub>N" 60)
   74.93 -  where "Ndiv = (\<lambda>a b. a *\<^sub>N Ninv b)"
   74.94 -
   74.95 -lemma Nneg_normN[simp]: "isnormNum x \<Longrightarrow> isnormNum (~\<^sub>N x)"
   74.96 -  by (simp add: isnormNum_def Nneg_def split_def)
   74.97 -
   74.98 -lemma Nadd_normN[simp]: "isnormNum (x +\<^sub>N y)"
   74.99 -  by (simp add: Nadd_def split_def)
  74.100 -
  74.101 -lemma Nsub_normN[simp]: "\<lbrakk> isnormNum y\<rbrakk> \<Longrightarrow> isnormNum (x -\<^sub>N y)"
  74.102 -  by (simp add: Nsub_def split_def)
  74.103 -
  74.104 -lemma Nmul_normN[simp]:
  74.105 -  assumes xn: "isnormNum x" and yn: "isnormNum y"
  74.106 -  shows "isnormNum (x *\<^sub>N y)"
  74.107 -proof -
  74.108 -  obtain a b where x: "x = (a, b)" by (cases x)
  74.109 -  obtain a' b' where y: "y = (a', b')" by (cases y)
  74.110 -  { assume "a = 0"
  74.111 -    hence ?thesis using xn x y
  74.112 -      by (simp add: isnormNum_def Let_def Nmul_def split_def) }
  74.113 -  moreover
  74.114 -  { assume "a' = 0"
  74.115 -    hence ?thesis using yn x y
  74.116 -      by (simp add: isnormNum_def Let_def Nmul_def split_def) }
  74.117 -  moreover
  74.118 -  { assume a: "a \<noteq>0" and a': "a'\<noteq>0"
  74.119 -    hence bp: "b > 0" "b' > 0" using xn yn x y by (simp_all add: isnormNum_def)
  74.120 -    from mult_pos_pos[OF bp] have "x *\<^sub>N y = normNum (a * a', b * b')"
  74.121 -      using x y a a' bp by (simp add: Nmul_def Let_def split_def normNum_def)
  74.122 -    hence ?thesis by simp }
  74.123 -  ultimately show ?thesis by blast
  74.124 -qed
  74.125 -
  74.126 -lemma Ninv_normN[simp]: "isnormNum x \<Longrightarrow> isnormNum (Ninv x)"
  74.127 -  by (simp add: Ninv_def isnormNum_def split_def)
  74.128 -    (cases "fst x = 0", auto simp add: gcd_commute_int)
  74.129 -
  74.130 -lemma isnormNum_int[simp]:
  74.131 -  "isnormNum 0\<^sub>N" "isnormNum ((1::int)\<^sub>N)" "i \<noteq> 0 \<Longrightarrow> isnormNum (i)\<^sub>N"
  74.132 -  by (simp_all add: isnormNum_def)
  74.133 -
  74.134 -
  74.135 -text {* Relations over Num *}
  74.136 -
  74.137 -definition Nlt0:: "Num \<Rightarrow> bool"  ("0>\<^sub>N")
  74.138 -  where "Nlt0 = (\<lambda>(a,b). a < 0)"
  74.139 -
  74.140 -definition Nle0:: "Num \<Rightarrow> bool"  ("0\<ge>\<^sub>N")
  74.141 -  where "Nle0 = (\<lambda>(a,b). a \<le> 0)"
  74.142 -
  74.143 -definition Ngt0:: "Num \<Rightarrow> bool"  ("0<\<^sub>N")
  74.144 -  where "Ngt0 = (\<lambda>(a,b). a > 0)"
  74.145 -
  74.146 -definition Nge0:: "Num \<Rightarrow> bool"  ("0\<le>\<^sub>N")
  74.147 -  where "Nge0 = (\<lambda>(a,b). a \<ge> 0)"
  74.148 -
  74.149 -definition Nlt :: "Num \<Rightarrow> Num \<Rightarrow> bool"  (infix "<\<^sub>N" 55)
  74.150 -  where "Nlt = (\<lambda>a b. 0>\<^sub>N (a -\<^sub>N b))"
  74.151 -
  74.152 -definition Nle :: "Num \<Rightarrow> Num \<Rightarrow> bool"  (infix "\<le>\<^sub>N" 55)
  74.153 -  where "Nle = (\<lambda>a b. 0\<ge>\<^sub>N (a -\<^sub>N b))"
  74.154 -
  74.155 -definition "INum = (\<lambda>(a,b). of_int a / of_int b)"
  74.156 -
  74.157 -lemma INum_int [simp]: "INum (i)\<^sub>N = ((of_int i) ::'a::field)" "INum 0\<^sub>N = (0::'a::field)"
  74.158 -  by (simp_all add: INum_def)
  74.159 -
  74.160 -lemma isnormNum_unique[simp]:
  74.161 -  assumes na: "isnormNum x" and nb: "isnormNum y"
  74.162 -  shows "((INum x ::'a::{field_char_0, field_inverse_zero}) = INum y) = (x = y)" (is "?lhs = ?rhs")
  74.163 -proof
  74.164 -  obtain a b where x: "x = (a, b)" by (cases x)
  74.165 -  obtain a' b' where y: "y = (a', b')" by (cases y)
  74.166 -  assume H: ?lhs
  74.167 -  { assume "a = 0 \<or> b = 0 \<or> a' = 0 \<or> b' = 0"
  74.168 -    hence ?rhs using na nb H
  74.169 -      by (simp add: x y INum_def split_def isnormNum_def split: split_if_asm) }
  74.170 -  moreover
  74.171 -  { assume az: "a \<noteq> 0" and bz: "b \<noteq> 0" and a'z: "a'\<noteq>0" and b'z: "b'\<noteq>0"
  74.172 -    from az bz a'z b'z na nb have pos: "b > 0" "b' > 0" by (simp_all add: x y isnormNum_def)
  74.173 -    from H bz b'z have eq: "a * b' = a'*b"
  74.174 -      by (simp add: x y INum_def eq_divide_eq divide_eq_eq of_int_mult[symmetric] del: of_int_mult)
  74.175 -    from az a'z na nb have gcd1: "gcd a b = 1" "gcd b a = 1" "gcd a' b' = 1" "gcd b' a' = 1"
  74.176 -      by (simp_all add: x y isnormNum_def add: gcd_commute_int)
  74.177 -    from eq have raw_dvd: "a dvd a' * b" "b dvd b' * a" "a' dvd a * b'" "b' dvd b * a'"
  74.178 -      apply -
  74.179 -      apply algebra
  74.180 -      apply algebra
  74.181 -      apply simp
  74.182 -      apply algebra
  74.183 -      done
  74.184 -    from zdvd_antisym_abs[OF coprime_dvd_mult_int[OF gcd1(2) raw_dvd(2)]
  74.185 -        coprime_dvd_mult_int[OF gcd1(4) raw_dvd(4)]]
  74.186 -      have eq1: "b = b'" using pos by arith
  74.187 -      with eq have "a = a'" using pos by simp
  74.188 -      with eq1 have ?rhs by (simp add: x y) }
  74.189 -  ultimately show ?rhs by blast
  74.190 -next
  74.191 -  assume ?rhs thus ?lhs by simp
  74.192 -qed
  74.193 -
  74.194 -
  74.195 -lemma isnormNum0[simp]:
  74.196 -    "isnormNum x \<Longrightarrow> (INum x = (0::'a::{field_char_0, field_inverse_zero})) = (x = 0\<^sub>N)"
  74.197 -  unfolding INum_int(2)[symmetric]
  74.198 -  by (rule isnormNum_unique) simp_all
  74.199 -
  74.200 -lemma of_int_div_aux: "d ~= 0 ==> ((of_int x)::'a::field_char_0) / (of_int d) =
  74.201 -    of_int (x div d) + (of_int (x mod d)) / ((of_int d)::'a)"
  74.202 -proof -
  74.203 -  assume "d ~= 0"
  74.204 -  let ?t = "of_int (x div d) * ((of_int d)::'a) + of_int(x mod d)"
  74.205 -  let ?f = "\<lambda>x. x / of_int d"
  74.206 -  have "x = (x div d) * d + x mod d"
  74.207 -    by auto
  74.208 -  then have eq: "of_int x = ?t"
  74.209 -    by (simp only: of_int_mult[symmetric] of_int_add [symmetric])
  74.210 -  then have "of_int x / of_int d = ?t / of_int d"
  74.211 -    using cong[OF refl[of ?f] eq] by simp
  74.212 -  then show ?thesis by (simp add: add_divide_distrib algebra_simps `d ~= 0`)
  74.213 -qed
  74.214 -
  74.215 -lemma of_int_div: "(d::int) ~= 0 ==> d dvd n ==>
  74.216 -    (of_int(n div d)::'a::field_char_0) = of_int n / of_int d"
  74.217 -  apply (frule of_int_div_aux [of d n, where ?'a = 'a])
  74.218 -  apply simp
  74.219 -  apply (simp add: dvd_eq_mod_eq_0)
  74.220 -  done
  74.221 -
  74.222 -
  74.223 -lemma normNum[simp]: "INum (normNum x) = (INum x :: 'a::{field_char_0, field_inverse_zero})"
  74.224 -proof -
  74.225 -  obtain a b where x: "x = (a, b)" by (cases x)
  74.226 -  { assume "a = 0 \<or> b = 0"
  74.227 -    hence ?thesis by (simp add: x INum_def normNum_def split_def Let_def) }
  74.228 -  moreover
  74.229 -  { assume a: "a \<noteq> 0" and b: "b \<noteq> 0"
  74.230 -    let ?g = "gcd a b"
  74.231 -    from a b have g: "?g \<noteq> 0"by simp
  74.232 -    from of_int_div[OF g, where ?'a = 'a]
  74.233 -    have ?thesis by (auto simp add: x INum_def normNum_def split_def Let_def) }
  74.234 -  ultimately show ?thesis by blast
  74.235 -qed
  74.236 -
  74.237 -lemma INum_normNum_iff:
  74.238 -  "(INum x ::'a::{field_char_0, field_inverse_zero}) = INum y \<longleftrightarrow> normNum x = normNum y"
  74.239 -  (is "?lhs = ?rhs")
  74.240 -proof -
  74.241 -  have "normNum x = normNum y \<longleftrightarrow> (INum (normNum x) :: 'a) = INum (normNum y)"
  74.242 -    by (simp del: normNum)
  74.243 -  also have "\<dots> = ?lhs" by simp
  74.244 -  finally show ?thesis by simp
  74.245 -qed
  74.246 -
  74.247 -lemma Nadd[simp]: "INum (x +\<^sub>N y) = INum x + (INum y :: 'a :: {field_char_0, field_inverse_zero})"
  74.248 -proof -
  74.249 -  let ?z = "0:: 'a"
  74.250 -  obtain a b where x: "x = (a, b)" by (cases x)
  74.251 -  obtain a' b' where y: "y = (a', b')" by (cases y)
  74.252 -  { assume "a=0 \<or> a'= 0 \<or> b =0 \<or> b' = 0"
  74.253 -    hence ?thesis
  74.254 -      apply (cases "a=0", simp_all add: x y Nadd_def)
  74.255 -      apply (cases "b= 0", simp_all add: INum_def)
  74.256 -       apply (cases "a'= 0", simp_all)
  74.257 -       apply (cases "b'= 0", simp_all)
  74.258 -       done }
  74.259 -  moreover
  74.260 -  { assume aa': "a \<noteq> 0" "a'\<noteq> 0" and bb': "b \<noteq> 0" "b' \<noteq> 0"
  74.261 -    { assume z: "a * b' + b * a' = 0"
  74.262 -      hence "of_int (a*b' + b*a') / (of_int b* of_int b') = ?z" by simp
  74.263 -      hence "of_int b' * of_int a / (of_int b * of_int b') +
  74.264 -          of_int b * of_int a' / (of_int b * of_int b') = ?z"
  74.265 -        by (simp add:add_divide_distrib)
  74.266 -      hence th: "of_int a / of_int b + of_int a' / of_int b' = ?z" using bb' aa'
  74.267 -        by simp
  74.268 -      from z aa' bb' have ?thesis
  74.269 -        by (simp add: x y th Nadd_def normNum_def INum_def split_def) }
  74.270 -    moreover {
  74.271 -      assume z: "a * b' + b * a' \<noteq> 0"
  74.272 -      let ?g = "gcd (a * b' + b * a') (b*b')"
  74.273 -      have gz: "?g \<noteq> 0" using z by simp
  74.274 -      have ?thesis using aa' bb' z gz
  74.275 -        of_int_div[where ?'a = 'a, OF gz gcd_dvd1_int[where x="a * b' + b * a'" and y="b*b'"]]
  74.276 -        of_int_div[where ?'a = 'a, OF gz gcd_dvd2_int[where x="a * b' + b * a'" and y="b*b'"]]
  74.277 -        by (simp add: x y Nadd_def INum_def normNum_def Let_def add_divide_distrib) }
  74.278 -    ultimately have ?thesis using aa' bb'
  74.279 -      by (simp add: x y Nadd_def INum_def normNum_def Let_def) }
  74.280 -  ultimately show ?thesis by blast
  74.281 -qed
  74.282 -
  74.283 -lemma Nmul[simp]: "INum (x *\<^sub>N y) = INum x * (INum y:: 'a :: {field_char_0, field_inverse_zero})"
  74.284 -proof -
  74.285 -  let ?z = "0::'a"
  74.286 -  obtain a b where x: "x = (a, b)" by (cases x)
  74.287 -  obtain a' b' where y: "y = (a', b')" by (cases y)
  74.288 -  { assume "a=0 \<or> a'= 0 \<or> b = 0 \<or> b' = 0"
  74.289 -    hence ?thesis
  74.290 -      apply (cases "a=0", simp_all add: x y Nmul_def INum_def Let_def)
  74.291 -      apply (cases "b=0", simp_all)
  74.292 -      apply (cases "a'=0", simp_all)
  74.293 -      done }
  74.294 -  moreover
  74.295 -  { assume z: "a \<noteq> 0" "a' \<noteq> 0" "b \<noteq> 0" "b' \<noteq> 0"
  74.296 -    let ?g="gcd (a*a') (b*b')"
  74.297 -    have gz: "?g \<noteq> 0" using z by simp
  74.298 -    from z of_int_div[where ?'a = 'a, OF gz gcd_dvd1_int[where x="a*a'" and y="b*b'"]]
  74.299 -      of_int_div[where ?'a = 'a , OF gz gcd_dvd2_int[where x="a*a'" and y="b*b'"]]
  74.300 -    have ?thesis by (simp add: Nmul_def x y Let_def INum_def) }
  74.301 -  ultimately show ?thesis by blast
  74.302 -qed
  74.303 -
  74.304 -lemma Nneg[simp]: "INum (~\<^sub>N x) = - (INum x ::'a:: field)"
  74.305 -  by (simp add: Nneg_def split_def INum_def)
  74.306 -
  74.307 -lemma Nsub[simp]: "INum (x -\<^sub>N y) = INum x - (INum y:: 'a :: {field_char_0, field_inverse_zero})"
  74.308 -  by (simp add: Nsub_def split_def)
  74.309 -
  74.310 -lemma Ninv[simp]: "INum (Ninv x) = (1::'a :: field_inverse_zero) / (INum x)"
  74.311 -  by (simp add: Ninv_def INum_def split_def)
  74.312 -
  74.313 -lemma Ndiv[simp]: "INum (x \<div>\<^sub>N y) = INum x / (INum y ::'a :: {field_char_0, field_inverse_zero})"
  74.314 -  by (simp add: Ndiv_def)
  74.315 -
  74.316 -lemma Nlt0_iff[simp]:
  74.317 -  assumes nx: "isnormNum x"
  74.318 -  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero})< 0) = 0>\<^sub>N x"
  74.319 -proof -
  74.320 -  obtain a b where x: "x = (a, b)" by (cases x)
  74.321 -  { assume "a = 0" hence ?thesis by (simp add: x Nlt0_def INum_def) }
  74.322 -  moreover
  74.323 -  { assume a: "a \<noteq> 0" hence b: "(of_int b::'a) > 0"
  74.324 -      using nx by (simp add: x isnormNum_def)
  74.325 -    from pos_divide_less_eq[OF b, where b="of_int a" and a="0::'a"]
  74.326 -    have ?thesis by (simp add: x Nlt0_def INum_def) }
  74.327 -  ultimately show ?thesis by blast
  74.328 -qed
  74.329 -
  74.330 -lemma Nle0_iff[simp]:
  74.331 -  assumes nx: "isnormNum x"
  74.332 -  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero}) \<le> 0) = 0\<ge>\<^sub>N x"
  74.333 -proof -
  74.334 -  obtain a b where x: "x = (a, b)" by (cases x)
  74.335 -  { assume "a = 0" hence ?thesis by (simp add: x Nle0_def INum_def) }
  74.336 -  moreover
  74.337 -  { assume a: "a \<noteq> 0" hence b: "(of_int b :: 'a) > 0"
  74.338 -      using nx by (simp add: x isnormNum_def)
  74.339 -    from pos_divide_le_eq[OF b, where b="of_int a" and a="0::'a"]
  74.340 -    have ?thesis by (simp add: x Nle0_def INum_def) }
  74.341 -  ultimately show ?thesis by blast
  74.342 -qed
  74.343 -
  74.344 -lemma Ngt0_iff[simp]:
  74.345 -  assumes nx: "isnormNum x"
  74.346 -  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero})> 0) = 0<\<^sub>N x"
  74.347 -proof -
  74.348 -  obtain a b where x: "x = (a, b)" by (cases x)
  74.349 -  { assume "a = 0" hence ?thesis by (simp add: x Ngt0_def INum_def) }
  74.350 -  moreover
  74.351 -  { assume a: "a \<noteq> 0" hence b: "(of_int b::'a) > 0" using nx
  74.352 -      by (simp add: x isnormNum_def)
  74.353 -    from pos_less_divide_eq[OF b, where b="of_int a" and a="0::'a"]
  74.354 -    have ?thesis by (simp add: x Ngt0_def INum_def) }
  74.355 -  ultimately show ?thesis by blast
  74.356 -qed
  74.357 -
  74.358 -lemma Nge0_iff[simp]:
  74.359 -  assumes nx: "isnormNum x"
  74.360 -  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero}) \<ge> 0) = 0\<le>\<^sub>N x"
  74.361 -proof -
  74.362 -  obtain a b where x: "x = (a, b)" by (cases x)
  74.363 -  { assume "a = 0" hence ?thesis by (simp add: x Nge0_def INum_def) }
  74.364 -  moreover
  74.365 -  { assume "a \<noteq> 0" hence b: "(of_int b::'a) > 0" using nx
  74.366 -      by (simp add: x isnormNum_def)
  74.367 -    from pos_le_divide_eq[OF b, where b="of_int a" and a="0::'a"]
  74.368 -    have ?thesis by (simp add: x Nge0_def INum_def) }
  74.369 -  ultimately show ?thesis by blast
  74.370 -qed
  74.371 -
  74.372 -lemma Nlt_iff[simp]:
  74.373 -  assumes nx: "isnormNum x" and ny: "isnormNum y"
  74.374 -  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero}) < INum y) = (x <\<^sub>N y)"
  74.375 -proof -
  74.376 -  let ?z = "0::'a"
  74.377 -  have "((INum x ::'a) < INum y) = (INum (x -\<^sub>N y) < ?z)"
  74.378 -    using nx ny by simp
  74.379 -  also have "\<dots> = (0>\<^sub>N (x -\<^sub>N y))"
  74.380 -    using Nlt0_iff[OF Nsub_normN[OF ny]] by simp
  74.381 -  finally show ?thesis by (simp add: Nlt_def)
  74.382 -qed
  74.383 -
  74.384 -lemma Nle_iff[simp]:
  74.385 -  assumes nx: "isnormNum x" and ny: "isnormNum y"
  74.386 -  shows "((INum x :: 'a :: {field_char_0, linordered_field_inverse_zero})\<le> INum y) = (x \<le>\<^sub>N y)"
  74.387 -proof -
  74.388 -  have "((INum x ::'a) \<le> INum y) = (INum (x -\<^sub>N y) \<le> (0::'a))"
  74.389 -    using nx ny by simp
  74.390 -  also have "\<dots> = (0\<ge>\<^sub>N (x -\<^sub>N y))"
  74.391 -    using Nle0_iff[OF Nsub_normN[OF ny]] by simp
  74.392 -  finally show ?thesis by (simp add: Nle_def)
  74.393 -qed
  74.394 -
  74.395 -lemma Nadd_commute:
  74.396 -  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  74.397 -  shows "x +\<^sub>N y = y +\<^sub>N x"
  74.398 -proof -
  74.399 -  have n: "isnormNum (x +\<^sub>N y)" "isnormNum (y +\<^sub>N x)" by simp_all
  74.400 -  have "(INum (x +\<^sub>N y)::'a) = INum (y +\<^sub>N x)" by simp
  74.401 -  with isnormNum_unique[OF n] show ?thesis by simp
  74.402 -qed
  74.403 -
  74.404 -lemma [simp]:
  74.405 -  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  74.406 -  shows "(0, b) +\<^sub>N y = normNum y"
  74.407 -    and "(a, 0) +\<^sub>N y = normNum y"
  74.408 -    and "x +\<^sub>N (0, b) = normNum x"
  74.409 -    and "x +\<^sub>N (a, 0) = normNum x"
  74.410 -  apply (simp add: Nadd_def split_def)
  74.411 -  apply (simp add: Nadd_def split_def)
  74.412 -  apply (subst Nadd_commute, simp add: Nadd_def split_def)
  74.413 -  apply (subst Nadd_commute, simp add: Nadd_def split_def)
  74.414 -  done
  74.415 -
  74.416 -lemma normNum_nilpotent_aux[simp]:
  74.417 -  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  74.418 -  assumes nx: "isnormNum x"
  74.419 -  shows "normNum x = x"
  74.420 -proof -
  74.421 -  let ?a = "normNum x"
  74.422 -  have n: "isnormNum ?a" by simp
  74.423 -  have th: "INum ?a = (INum x ::'a)" by simp
  74.424 -  with isnormNum_unique[OF n nx] show ?thesis by simp
  74.425 -qed
  74.426 -
  74.427 -lemma normNum_nilpotent[simp]:
  74.428 -  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  74.429 -  shows "normNum (normNum x) = normNum x"
  74.430 -  by simp
  74.431 -
  74.432 -lemma normNum0[simp]: "normNum (0,b) = 0\<^sub>N" "normNum (a,0) = 0\<^sub>N"
  74.433 -  by (simp_all add: normNum_def)
  74.434 -
  74.435 -lemma normNum_Nadd:
  74.436 -  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  74.437 -  shows "normNum (x +\<^sub>N y) = x +\<^sub>N y" by simp
  74.438 -
  74.439 -lemma Nadd_normNum1[simp]:
  74.440 -  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  74.441 -  shows "normNum x +\<^sub>N y = x +\<^sub>N y"
  74.442 -proof -
  74.443 -  have n: "isnormNum (normNum x +\<^sub>N y)" "isnormNum (x +\<^sub>N y)" by simp_all
  74.444 -  have "INum (normNum x +\<^sub>N y) = INum x + (INum y :: 'a)" by simp
  74.445 -  also have "\<dots> = INum (x +\<^sub>N y)" by simp
  74.446 -  finally show ?thesis using isnormNum_unique[OF n] by simp
  74.447 -qed
  74.448 -
  74.449 -lemma Nadd_normNum2[simp]:
  74.450 -  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  74.451 -  shows "x +\<^sub>N normNum y = x +\<^sub>N y"
  74.452 -proof -
  74.453 -  have n: "isnormNum (x +\<^sub>N normNum y)" "isnormNum (x +\<^sub>N y)" by simp_all
  74.454 -  have "INum (x +\<^sub>N normNum y) = INum x + (INum y :: 'a)" by simp
  74.455 -  also have "\<dots> = INum (x +\<^sub>N y)" by simp
  74.456 -  finally show ?thesis using isnormNum_unique[OF n] by simp
  74.457 -qed
  74.458 -
  74.459 -lemma Nadd_assoc:
  74.460 -  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  74.461 -  shows "x +\<^sub>N y +\<^sub>N z = x +\<^sub>N (y +\<^sub>N z)"
  74.462 -proof -
  74.463 -  have n: "isnormNum (x +\<^sub>N y +\<^sub>N z)" "isnormNum (x +\<^sub>N (y +\<^sub>N z))" by simp_all
  74.464 -  have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" by simp
  74.465 -  with isnormNum_unique[OF n] show ?thesis by simp
  74.466 -qed
  74.467 -
  74.468 -lemma Nmul_commute: "isnormNum x \<Longrightarrow> isnormNum y \<Longrightarrow> x *\<^sub>N y = y *\<^sub>N x"
  74.469 -  by (simp add: Nmul_def split_def Let_def gcd_commute_int mult_commute)
  74.470 -
  74.471 -lemma Nmul_assoc:
  74.472 -  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  74.473 -  assumes nx: "isnormNum x" and ny: "isnormNum y" and nz: "isnormNum z"
  74.474 -  shows "x *\<^sub>N y *\<^sub>N z = x *\<^sub>N (y *\<^sub>N z)"
  74.475 -proof -
  74.476 -  from nx ny nz have n: "isnormNum (x *\<^sub>N y *\<^sub>N z)" "isnormNum (x *\<^sub>N (y *\<^sub>N z))"
  74.477 -    by simp_all
  74.478 -  have "INum (x +\<^sub>N y +\<^sub>N z) = (INum (x +\<^sub>N (y +\<^sub>N z)) :: 'a)" by simp
  74.479 -  with isnormNum_unique[OF n] show ?thesis by simp
  74.480 -qed
  74.481 -
  74.482 -lemma Nsub0:
  74.483 -  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  74.484 -  assumes x: "isnormNum x" and y: "isnormNum y"
  74.485 -  shows "x -\<^sub>N y = 0\<^sub>N \<longleftrightarrow> x = y"
  74.486 -proof -
  74.487 -  fix h :: 'a
  74.488 -  from isnormNum_unique[where 'a = 'a, OF Nsub_normN[OF y], where y="0\<^sub>N"]
  74.489 -  have "(x -\<^sub>N y = 0\<^sub>N) = (INum (x -\<^sub>N y) = (INum 0\<^sub>N :: 'a)) " by simp
  74.490 -  also have "\<dots> = (INum x = (INum y :: 'a))" by simp
  74.491 -  also have "\<dots> = (x = y)" using x y by simp
  74.492 -  finally show ?thesis .
  74.493 -qed
  74.494 -
  74.495 -lemma Nmul0[simp]: "c *\<^sub>N 0\<^sub>N = 0\<^sub>N" " 0\<^sub>N *\<^sub>N c = 0\<^sub>N"
  74.496 -  by (simp_all add: Nmul_def Let_def split_def)
  74.497 -
  74.498 -lemma Nmul_eq0[simp]:
  74.499 -  assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})"
  74.500 -  assumes nx: "isnormNum x" and ny: "isnormNum y"
  74.501 -  shows "x*\<^sub>N y = 0\<^sub>N \<longleftrightarrow> x = 0\<^sub>N \<or> y = 0\<^sub>N"
  74.502 -proof -
  74.503 -  fix h :: 'a
  74.504 -  obtain a b where x: "x = (a, b)" by (cases x)
  74.505 -  obtain a' b' where y: "y = (a', b')" by (cases y)
  74.506 -  have n0: "isnormNum 0\<^sub>N" by simp
  74.507 -  show ?thesis using nx ny
  74.508 -    apply (simp only: isnormNum_unique[where ?'a = 'a, OF  Nmul_normN[OF nx ny] n0, symmetric]
  74.509 -      Nmul[where ?'a = 'a])
  74.510 -    apply (simp add: x y INum_def split_def isnormNum_def split: split_if_asm)
  74.511 -    done
  74.512 -qed
  74.513 -
  74.514 -lemma Nneg_Nneg[simp]: "~\<^sub>N (~\<^sub>N c) = c"
  74.515 -  by (simp add: Nneg_def split_def)
  74.516 -
  74.517 -lemma Nmul1[simp]:
  74.518 -    "isnormNum c \<Longrightarrow> (1)\<^sub>N *\<^sub>N c = c"
  74.519 -    "isnormNum c \<Longrightarrow> c *\<^sub>N (1)\<^sub>N = c"
  74.520 -  apply (simp_all add: Nmul_def Let_def split_def isnormNum_def)
  74.521 -  apply (cases "fst c = 0", simp_all, cases c, simp_all)+
  74.522 -  done
  74.523 -
  74.524 -end
    75.1 --- a/src/HOL/Library/BigO.thy	Mon Nov 11 17:34:44 2013 +0100
    75.2 +++ b/src/HOL/Library/BigO.thy	Mon Nov 11 17:44:21 2013 +0100
    75.3 @@ -215,7 +215,7 @@
    75.4      f : lb +o O(g)"
    75.5    apply (rule set_minus_imp_plus)
    75.6    apply (rule bigo_bounded)
    75.7 -  apply (auto simp add: diff_minus fun_Compl_def func_plus)
    75.8 +  apply (auto simp add: fun_Compl_def func_plus)
    75.9    apply (drule_tac x = x in spec)+
   75.10    apply force
   75.11    apply (drule_tac x = x in spec)+
   75.12 @@ -390,7 +390,7 @@
   75.13    apply (rule set_minus_imp_plus)
   75.14    apply (drule set_plus_imp_minus)
   75.15    apply (drule bigo_minus)
   75.16 -  apply (simp add: diff_minus)
   75.17 +  apply simp
   75.18    done
   75.19  
   75.20  lemma bigo_minus3: "O(-f) = O(f)"
   75.21 @@ -446,7 +446,7 @@
   75.22    apply (rule bigo_minus)
   75.23    apply (subst set_minus_plus)
   75.24    apply assumption
   75.25 -  apply  (simp add: diff_minus add_ac)
   75.26 +  apply (simp add: add_ac)
   75.27    done
   75.28  
   75.29  lemma bigo_add_commute: "(f : g +o O(h)) = (g : f +o O(h))"
   75.30 @@ -545,10 +545,9 @@
   75.31  
   75.32  lemma bigo_compose2: "f =o g +o O(h) ==> (%x. f(k x)) =o (%x. g(k x)) +o 
   75.33      O(%x. h(k x))"
   75.34 -  apply (simp only: set_minus_plus [symmetric] diff_minus fun_Compl_def
   75.35 -      func_plus)
   75.36 -  apply (erule bigo_compose1)
   75.37 -done
   75.38 +  apply (simp only: set_minus_plus [symmetric] fun_Compl_def func_plus)
   75.39 +  apply (drule bigo_compose1) apply (simp add: fun_diff_def)
   75.40 +  done
   75.41  
   75.42  
   75.43  subsection {* Setsum *}
   75.44 @@ -779,7 +778,7 @@
   75.45    apply (subst abs_of_nonneg)
   75.46    apply (drule_tac x = x in spec) back
   75.47    apply (simp add: algebra_simps)
   75.48 -  apply (subst diff_minus)+
   75.49 +  apply (subst diff_conv_add_uminus)+
   75.50    apply (rule add_right_mono)
   75.51    apply (erule spec)
   75.52    apply (rule order_trans) 
   75.53 @@ -803,7 +802,7 @@
   75.54    apply (subst abs_of_nonneg)
   75.55    apply (drule_tac x = x in spec) back
   75.56    apply (simp add: algebra_simps)
   75.57 -  apply (subst diff_minus)+
   75.58 +  apply (subst diff_conv_add_uminus)+
   75.59    apply (rule add_left_mono)
   75.60    apply (rule le_imp_neg_le)
   75.61    apply (erule spec)
    76.1 --- a/src/HOL/Library/Code_Target_Int.thy	Mon Nov 11 17:34:44 2013 +0100
    76.2 +++ b/src/HOL/Library/Code_Target_Int.thy	Mon Nov 11 17:44:21 2013 +0100
    76.3 @@ -99,7 +99,7 @@
    76.4  proof -
    76.5    from mod_div_equality have *: "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp
    76.6    show ?thesis
    76.7 -    by (simp add: Let_def divmod_int_mod_div mod_2_not_eq_zero_eq_one_int
    76.8 +    by (simp add: Let_def divmod_int_mod_div not_mod_2_eq_0_eq_1
    76.9        of_int_add [symmetric]) (simp add: * mult_commute)
   76.10  qed
   76.11  
    77.1 --- a/src/HOL/Library/ContNotDenum.thy	Mon Nov 11 17:34:44 2013 +0100
    77.2 +++ b/src/HOL/Library/ContNotDenum.thy	Mon Nov 11 17:44:21 2013 +0100
    77.3 @@ -131,17 +131,15 @@
    77.4  
    77.5    -- "A denotes the set of all left-most points of all the intervals ..."
    77.6    moreover obtain A where Adef: "A = ?g ` \<nat>" by simp
    77.7 -  ultimately have "\<exists>x. x\<in>A"
    77.8 +  ultimately have "A \<noteq> {}"
    77.9    proof -
   77.10      have "(0::nat) \<in> \<nat>" by simp
   77.11 -    moreover have "?g 0 = ?g 0" by simp
   77.12 -    ultimately have "?g 0 \<in> ?g ` \<nat>" by (rule  rev_image_eqI)
   77.13 -    with Adef have "?g 0 \<in> A" by simp
   77.14 -    thus ?thesis ..
   77.15 +    with Adef show ?thesis
   77.16 +      by blast
   77.17    qed
   77.18  
   77.19    -- "Now show that A is bounded above ..."
   77.20 -  moreover have "\<exists>y. isUb (UNIV::real set) A y"
   77.21 +  moreover have "bdd_above A"
   77.22    proof -
   77.23      {
   77.24        fix n
   77.25 @@ -177,18 +175,11 @@
   77.26        obtain a and b where "f 0 = closed_int a b" and alb: "a \<le> b" by blast
   77.27      ultimately have "\<forall>n. ?g n \<in> closed_int a b" by auto
   77.28      with alb have "\<forall>n. ?g n \<le> b" using closed_int_most by blast
   77.29 -    with Adef have "\<forall>y\<in>A. y\<le>b" by auto
   77.30 -    hence "A *<= b" by (unfold setle_def)
   77.31 -    moreover have "b \<in> (UNIV::real set)" by simp
   77.32 -    ultimately have "A *<= b \<and> b \<in> (UNIV::real set)" by simp
   77.33 -    hence "isUb (UNIV::real set) A b" by (unfold isUb_def)
   77.34 -    thus ?thesis by auto
   77.35 +    with Adef show "bdd_above A" by auto
   77.36    qed
   77.37 -  -- "by the Axiom Of Completeness, A has a least upper bound ..."
   77.38 -  ultimately have "\<exists>t. isLub UNIV A t" by (rule reals_complete)
   77.39  
   77.40    -- "denote this least upper bound as t ..."
   77.41 -  then obtain t where tdef: "isLub UNIV A t" ..
   77.42 +  def tdef: t == "Sup A"
   77.43  
   77.44    -- "and finally show that this least upper bound is in all the intervals..."
   77.45    have "\<forall>n. t \<in> f n"
   77.46 @@ -229,82 +220,76 @@
   77.47          with Adef have "(?g n) \<in> A" by auto
   77.48          ultimately show ?thesis by simp
   77.49        qed 
   77.50 -      with tdef show "a \<le> t" by (rule isLubD2)
   77.51 +      with `bdd_above A` show "a \<le> t"
   77.52 +        unfolding tdef by (intro cSup_upper)
   77.53      qed
   77.54      moreover have "t \<le> b"
   77.55 -    proof -
   77.56 -      have "isUb UNIV A b"
   77.57 -      proof -
   77.58 +      unfolding tdef
   77.59 +    proof (rule cSup_least)
   77.60 +      {
   77.61 +        from alb int have
   77.62 +          ain: "b\<in>f n \<and> (\<forall>x\<in>f n. x \<le> b)" using closed_int_most by blast
   77.63 +        
   77.64 +        have subsetd: "\<forall>m. \<forall>n. f (n + m) \<subseteq> f n"
   77.65 +        proof (rule allI, induct_tac m)
   77.66 +          show "\<forall>n. f (n + 0) \<subseteq> f n" by simp
   77.67 +        next
   77.68 +          fix m n
   77.69 +          assume pp: "\<forall>p. f (p + n) \<subseteq> f p"
   77.70 +          {
   77.71 +            fix p
   77.72 +            from pp have "f (p + n) \<subseteq> f p" by simp
   77.73 +            moreover from subset have "f (Suc (p + n)) \<subseteq> f (p + n)" by auto
   77.74 +            hence "f (p + (Suc n)) \<subseteq> f (p + n)" by simp
   77.75 +            ultimately have "f (p + (Suc n)) \<subseteq> f p" by simp
   77.76 +          }
   77.77 +          thus "\<forall>p. f (p + Suc n) \<subseteq> f p" ..
   77.78 +        qed 
   77.79 +        have subsetm: "\<forall>\<alpha> \<beta>. \<alpha> \<ge> \<beta> \<longrightarrow> (f \<alpha>) \<subseteq> (f \<beta>)"
   77.80 +        proof ((rule allI)+, rule impI)
   77.81 +          fix \<alpha>::nat and \<beta>::nat
   77.82 +          assume "\<beta> \<le> \<alpha>"
   77.83 +          hence "\<exists>k. \<alpha> = \<beta> + k" by (simp only: le_iff_add)
   77.84 +          then obtain k where "\<alpha> = \<beta> + k" ..
   77.85 +          moreover
   77.86 +          from subsetd have "f (\<beta> + k) \<subseteq> f \<beta>" by simp
   77.87 +          ultimately show "f \<alpha> \<subseteq> f \<beta>" by auto
   77.88 +        qed 
   77.89 +        
   77.90 +        fix m   
   77.91          {
   77.92 -          from alb int have
   77.93 -            ain: "b\<in>f n \<and> (\<forall>x\<in>f n. x \<le> b)" using closed_int_most by blast
   77.94 -          
   77.95 -          have subsetd: "\<forall>m. \<forall>n. f (n + m) \<subseteq> f n"
   77.96 -          proof (rule allI, induct_tac m)
   77.97 -            show "\<forall>n. f (n + 0) \<subseteq> f n" by simp
   77.98 -          next
   77.99 -            fix m n
  77.100 -            assume pp: "\<forall>p. f (p + n) \<subseteq> f p"
  77.101 -            {
  77.102 -              fix p
  77.103 -              from pp have "f (p + n) \<subseteq> f p" by simp
  77.104 -              moreover from subset have "f (Suc (p + n)) \<subseteq> f (p + n)" by auto
  77.105 -              hence "f (p + (Suc n)) \<subseteq> f (p + n)" by simp
  77.106 -              ultimately have "f (p + (Suc n)) \<subseteq> f p" by simp
  77.107 -            }
  77.108 -            thus "\<forall>p. f (p + Suc n) \<subseteq> f p" ..
  77.109 -          qed 
  77.110 -          have subsetm: "\<forall>\<alpha> \<beta>. \<alpha> \<ge> \<beta> \<longrightarrow> (f \<alpha>) \<subseteq> (f \<beta>)"
  77.111 -          proof ((rule allI)+, rule impI)
  77.112 -            fix \<alpha>::nat and \<beta>::nat
  77.113 -            assume "\<beta> \<le> \<alpha>"
  77.114 -            hence "\<exists>k. \<alpha> = \<beta> + k" by (simp only: le_iff_add)
  77.115 -            then obtain k where "\<alpha> = \<beta> + k" ..
  77.116 -            moreover
  77.117 -            from subsetd have "f (\<beta> + k) \<subseteq> f \<beta>" by simp
  77.118 -            ultimately show "f \<alpha> \<subseteq> f \<beta>" by auto
  77.119 -          qed 
  77.120 -          
  77.121 -          fix m   
  77.122 -          {
  77.123 -            assume "m \<ge> n"
  77.124 -            with subsetm have "f m \<subseteq> f n" by simp
  77.125 -            with ain have "\<forall>x\<in>f m. x \<le> b" by auto
  77.126 -            moreover
  77.127 -            from gdef have "?g m \<in> f m \<and> (\<forall>x\<in>f m. ?g m \<le> x)" by simp
  77.128 -            ultimately have "?g m \<le> b" by auto
  77.129 -          }
  77.130 +          assume "m \<ge> n"
  77.131 +          with subsetm have "f m \<subseteq> f n" by simp
  77.132 +          with ain have "\<forall>x\<in>f m. x \<le> b" by auto
  77.133            moreover
  77.134 -          {
  77.135 -            assume "\<not>(m \<ge> n)"
  77.136 -            hence "m < n" by simp
  77.137 -            with subsetm have sub: "(f n) \<subseteq> (f m)" by simp
  77.138 -            from closed obtain ma and mb where
  77.139 -              "f m = closed_int ma mb \<and> ma \<le> mb" by blast
  77.140 -            hence one: "ma \<le> mb" and fm: "f m = closed_int ma mb" by auto 
  77.141 -            from one alb sub fm int have "ma \<le> b" using closed_subset by blast
  77.142 -            moreover have "(?g m) = ma"
  77.143 -            proof -
  77.144 -              from gdef have "?g m \<in> f m \<and> (\<forall>x\<in>f m. ?g m \<le> x)" ..
  77.145 -              moreover from one have
  77.146 -                "ma \<in> closed_int ma mb \<and> (\<forall>x\<in>closed_int ma mb. ma \<le> x)"
  77.147 -                by (rule closed_int_least)
  77.148 -              with fm have "ma\<in>f m \<and> (\<forall>x\<in>f m. ma \<le> x)" by simp
  77.149 -              ultimately have "ma \<le> ?g m \<and> ?g m \<le> ma" by auto
  77.150 -              thus "?g m = ma" by auto
  77.151 -            qed
  77.152 -            ultimately have "?g m \<le> b" by simp
  77.153 -          } 
  77.154 -          ultimately have "?g m \<le> b" by (rule case_split)
  77.155 +          from gdef have "?g m \<in> f m \<and> (\<forall>x\<in>f m. ?g m \<le> x)" by simp
  77.156 +          ultimately have "?g m \<le> b" by auto
  77.157          }
  77.158 -        with Adef have "\<forall>y\<in>A. y\<le>b" by auto
  77.159 -        hence "A *<= b" by (unfold setle_def)
  77.160 -        moreover have "b \<in> (UNIV::real set)" by simp
  77.161 -        ultimately have "A *<= b \<and> b \<in> (UNIV::real set)" by simp
  77.162 -        thus "isUb (UNIV::real set) A b" by (unfold isUb_def)
  77.163 -      qed
  77.164 -      with tdef show "t \<le> b" by (rule isLub_le_isUb)
  77.165 -    qed
  77.166 +        moreover
  77.167 +        {
  77.168 +          assume "\<not>(m \<ge> n)"
  77.169 +          hence "m < n" by simp
  77.170 +          with subsetm have sub: "(f n) \<subseteq> (f m)" by simp
  77.171 +          from closed obtain ma and mb where
  77.172 +            "f m = closed_int ma mb \<and> ma \<le> mb" by blast
  77.173 +          hence one: "ma \<le> mb" and fm: "f m = closed_int ma mb" by auto 
  77.174 +          from one alb sub fm int have "ma \<le> b" using closed_subset by blast
  77.175 +          moreover have "(?g m) = ma"
  77.176 +          proof -
  77.177 +            from gdef have "?g m \<in> f m \<and> (\<forall>x\<in>f m. ?g m \<le> x)" ..
  77.178 +            moreover from one have
  77.179 +              "ma \<in> closed_int ma mb \<and> (\<forall>x\<in>closed_int ma mb. ma \<le> x)"
  77.180 +              by (rule closed_int_least)
  77.181 +            with fm have "ma\<in>f m \<and> (\<forall>x\<in>f m. ma \<le> x)" by simp
  77.182 +            ultimately have "ma \<le> ?g m \<and> ?g m \<le> ma" by auto
  77.183 +            thus "?g m = ma" by auto
  77.184 +          qed
  77.185 +          ultimately have "?g m \<le> b" by simp
  77.186 +        } 
  77.187 +        ultimately have "?g m \<le> b" by (rule case_split)
  77.188 +      }
  77.189 +      with Adef show "\<And>y. y \<in> A \<Longrightarrow> y \<le> b" by auto
  77.190 +    qed fact
  77.191      ultimately have "t \<in> closed_int a b" by (rule closed_mem)
  77.192      with int show "t \<in> f n" by simp
  77.193    qed
    78.1 --- a/src/HOL/Library/Continuity.thy	Mon Nov 11 17:34:44 2013 +0100
    78.2 +++ b/src/HOL/Library/Continuity.thy	Mon Nov 11 17:44:21 2013 +0100
    78.3 @@ -19,7 +19,8 @@
    78.4    "continuous F \<longleftrightarrow> (\<forall>M. chain M \<longrightarrow> F (SUP i. M i) = (SUP i. F (M i)))"
    78.5  
    78.6  lemma SUP_nat_conv:
    78.7 -  "(SUP n. M n) = sup (M 0) (SUP n. M(Suc n))"
    78.8 +  fixes M :: "nat \<Rightarrow> 'a::complete_lattice"
    78.9 +  shows "(SUP n. M n) = sup (M 0) (SUP n. M(Suc n))"
   78.10  apply(rule order_antisym)
   78.11   apply(rule SUP_least)
   78.12   apply(case_tac n)
    79.1 --- a/src/HOL/Library/Convex.thy	Mon Nov 11 17:34:44 2013 +0100
    79.2 +++ b/src/HOL/Library/Convex.thy	Mon Nov 11 17:44:21 2013 +0100
    79.3 @@ -362,7 +362,7 @@
    79.4    shows "convex {x - y| x y. x \<in> s \<and> y \<in> t}"
    79.5  proof -
    79.6    have "{x - y| x y. x \<in> s \<and> y \<in> t} = {x + y |x y. x \<in> s \<and> y \<in> uminus ` t}"
    79.7 -    unfolding diff_def by auto
    79.8 +    by (auto simp add: diff_conv_add_uminus simp del: add_uminus_conv_diff)
    79.9    then show ?thesis
   79.10      using convex_sums[OF assms(1) convex_negations[OF assms(2)]] by auto
   79.11  qed
    80.1 --- a/src/HOL/Library/FSet.thy	Mon Nov 11 17:34:44 2013 +0100
    80.2 +++ b/src/HOL/Library/FSet.thy	Mon Nov 11 17:44:21 2013 +0100
    80.3 @@ -101,19 +101,25 @@
    80.4  lemma finite_Sup: "\<exists>z. finite z \<and> (\<forall>a. a \<in> X \<longrightarrow> a \<le> z) \<Longrightarrow> finite (Sup X)"
    80.5  by (auto intro: finite_subset)
    80.6  
    80.7 +lemma transfer_bdd_below[transfer_rule]: "(set_rel (pcr_fset op =) ===> op =) bdd_below bdd_below"
    80.8 +  by auto
    80.9 +
   80.10  instance
   80.11  proof 
   80.12    fix x z :: "'a fset"
   80.13    fix X :: "'a fset set"
   80.14    {
   80.15 -    assume "x \<in> X" "(\<And>a. a \<in> X \<Longrightarrow> z |\<subseteq>| a)" 
   80.16 +    assume "x \<in> X" "bdd_below X" 
   80.17      then show "Inf X |\<subseteq>| x"  by transfer auto
   80.18    next
   80.19      assume "X \<noteq> {}" "(\<And>x. x \<in> X \<Longrightarrow> z |\<subseteq>| x)"
   80.20      then show "z |\<subseteq>| Inf X" by transfer (clarsimp, blast)
   80.21    next
   80.22 -    assume "x \<in> X" "(\<And>a. a \<in> X \<Longrightarrow> a |\<subseteq>| z)"
   80.23 -    then show "x |\<subseteq>| Sup X" by transfer (auto intro!: finite_Sup)
   80.24 +    assume "x \<in> X" "bdd_above X"
   80.25 +    then obtain z where "x \<in> X" "(\<And>x. x \<in> X \<Longrightarrow> x |\<subseteq>| z)"
   80.26 +      by (auto simp: bdd_above_def)
   80.27 +    then show "x |\<subseteq>| Sup X"
   80.28 +      by transfer (auto intro!: finite_Sup)
   80.29    next
   80.30      assume "X \<noteq> {}" "(\<And>x. x \<in> X \<Longrightarrow> x |\<subseteq>| z)"
   80.31      then show "Sup X |\<subseteq>| z" by transfer (clarsimp, blast)
    81.1 --- a/src/HOL/Library/Float.thy	Mon Nov 11 17:34:44 2013 +0100
    81.2 +++ b/src/HOL/Library/Float.thy	Mon Nov 11 17:44:21 2013 +0100
    81.3 @@ -88,7 +88,7 @@
    81.4    done
    81.5  
    81.6  lemma minus_float[simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> x - y \<in> float"
    81.7 -  unfolding ab_diff_minus by (intro uminus_float plus_float)
    81.8 +  using plus_float [of x "- y"] by simp
    81.9  
   81.10  lemma abs_float[simp]: "x \<in> float \<Longrightarrow> abs x \<in> float"
   81.11    by (cases x rule: linorder_cases[of 0]) auto
   81.12 @@ -960,7 +960,7 @@
   81.13    also have "... < (1 / 2) * 2 powr real (rat_precision n (int x) (int y))"
   81.14      apply (rule mult_strict_right_mono) by (insert assms) auto
   81.15    also have "\<dots> = 2 powr real (rat_precision n (int x) (int y) - 1)"
   81.16 -    by (simp add: powr_add diff_def powr_neg_numeral)
   81.17 +    using powr_add [of 2 _ "- 1", simplified add_uminus_conv_diff] by (simp add: powr_neg_numeral)
   81.18    also have "\<dots> = 2 ^ nat (rat_precision n (int x) (int y) - 1)"
   81.19      using rat_precision_pos[of x y n] assms by (simp add: powr_realpow[symmetric])
   81.20    also have "\<dots> \<le> 2 ^ nat (rat_precision n (int x) (int y)) - 1"
    82.1 --- a/src/HOL/Library/Formal_Power_Series.thy	Mon Nov 11 17:34:44 2013 +0100
    82.2 +++ b/src/HOL/Library/Formal_Power_Series.thy	Mon Nov 11 17:44:21 2013 +0100
    82.3 @@ -231,7 +231,7 @@
    82.4  proof
    82.5    fix a b :: "'a fps"
    82.6    show "- a + a = 0" by (simp add: fps_ext)
    82.7 -  show "a - b = a + - b" by (simp add: fps_ext diff_minus)
    82.8 +  show "a + - b = a - b" by (simp add: fps_ext)
    82.9  qed
   82.10  
   82.11  instance fps :: (ab_group_add) ab_group_add
   82.12 @@ -348,10 +348,10 @@
   82.13  instance fps :: (ring) ring ..
   82.14  
   82.15  instance fps :: (ring_1) ring_1
   82.16 -  by (intro_classes, auto simp add: diff_minus distrib_right)
   82.17 +  by (intro_classes, auto simp add: distrib_right)
   82.18  
   82.19  instance fps :: (comm_ring_1) comm_ring_1
   82.20 -  by (intro_classes, auto simp add: diff_minus distrib_right)
   82.21 +  by (intro_classes, auto simp add: distrib_right)
   82.22  
   82.23  instance fps :: (ring_no_zero_divisors) ring_no_zero_divisors
   82.24  proof
   82.25 @@ -451,7 +451,7 @@
   82.26  
   82.27  definition
   82.28    dist_fps_def: "dist (a::'a fps) b =
   82.29 -    (if (\<exists>n. a$n \<noteq> b$n) then inverse (2 ^ The (leastP (\<lambda>n. a$n \<noteq> b$n))) else 0)"
   82.30 +    (if (\<exists>n. a$n \<noteq> b$n) then inverse (2 ^ (LEAST n. a$n \<noteq> b$n)) else 0)"
   82.31  
   82.32  lemma dist_fps_ge0: "dist (a::'a fps) b \<ge> 0"
   82.33    by (simp add: dist_fps_def)
   82.34 @@ -467,34 +467,6 @@
   82.35  
   82.36  end
   82.37  
   82.38 -lemma fps_nonzero_least_unique:
   82.39 -  assumes a0: "a \<noteq> 0"
   82.40 -  shows "\<exists>!n. leastP (\<lambda>n. a$n \<noteq> 0) n"
   82.41 -proof -
   82.42 -  from fps_nonzero_nth_minimal [of a] a0
   82.43 -  obtain n where "a$n \<noteq> 0" "\<forall>m < n. a$m = 0" by blast
   82.44 -  then have ln: "leastP (\<lambda>n. a$n \<noteq> 0) n"
   82.45 -    by (auto simp add: leastP_def setge_def not_le [symmetric])
   82.46 -  moreover
   82.47 -  {
   82.48 -    fix m
   82.49 -    assume "leastP (\<lambda>n. a $ n \<noteq> 0) m"
   82.50 -    then have "m = n" using ln
   82.51 -      apply (auto simp add: leastP_def setge_def)
   82.52 -      apply (erule allE[where x=n])
   82.53 -      apply (erule allE[where x=m])
   82.54 -      apply simp
   82.55 -      done
   82.56 -  }
   82.57 -  ultimately show ?thesis by blast
   82.58 -qed
   82.59 -
   82.60 -lemma fps_eq_least_unique:
   82.61 -  assumes "(a::('a::ab_group_add) fps) \<noteq> b"
   82.62 -  shows "\<exists>! n. leastP (\<lambda>n. a$n \<noteq> b$n) n"
   82.63 -  using fps_nonzero_least_unique[of "a - b"] assms
   82.64 -  by auto
   82.65 -
   82.66  instantiation fps :: (comm_ring_1) metric_space
   82.67  begin
   82.68  
   82.69 @@ -540,31 +512,15 @@
   82.70    moreover
   82.71    {
   82.72      assume ab: "a \<noteq> b" and ac: "a \<noteq> c" and bc: "b \<noteq> c"
   82.73 -    let ?P = "\<lambda>a b n. a$n \<noteq> b$n"
   82.74 -    from fps_eq_least_unique[OF ab] fps_eq_least_unique[OF ac]
   82.75 -      fps_eq_least_unique[OF bc]
   82.76 -    obtain nab nac nbc where nab: "leastP (?P a b) nab"
   82.77 -      and nac: "leastP (?P a c) nac"
   82.78 -      and nbc: "leastP (?P b c) nbc" by blast
   82.79 -    from nab have nab': "\<And>m. m < nab \<Longrightarrow> a$m = b$m" "a$nab \<noteq> b$nab"
   82.80 -      by (auto simp add: leastP_def setge_def)
   82.81 -    from nac have nac': "\<And>m. m < nac \<Longrightarrow> a$m = c$m" "a$nac \<noteq> c$nac"
   82.82 -      by (auto simp add: leastP_def setge_def)
   82.83 -    from nbc have nbc': "\<And>m. m < nbc \<Longrightarrow> b$m = c$m" "b$nbc \<noteq> c$nbc"
   82.84 -      by (auto simp add: leastP_def setge_def)
   82.85 -
   82.86 -    have th0: "\<And>(a::'a fps) b. a \<noteq> b \<longleftrightarrow> (\<exists>n. a$n \<noteq> b$n)"
   82.87 -      by (simp add: fps_eq_iff)
   82.88 -    from ab ac bc nab nac nbc
   82.89 -    have dab: "dist a b = inverse (2 ^ nab)"
   82.90 -      and dac: "dist a c = inverse (2 ^ nac)"
   82.91 -      and dbc: "dist b c = inverse (2 ^ nbc)"
   82.92 -      unfolding th0
   82.93 -      apply (simp_all add: dist_fps_def)
   82.94 -      apply (erule the1_equality[OF fps_eq_least_unique[OF ab]])
   82.95 -      apply (erule the1_equality[OF fps_eq_least_unique[OF ac]])
   82.96 -      apply (erule the1_equality[OF fps_eq_least_unique[OF bc]])
   82.97 -      done
   82.98 +    def n \<equiv> "\<lambda>a b::'a fps. LEAST n. a$n \<noteq> b$n"
   82.99 +    then have n': "\<And>m a b. m < n a b \<Longrightarrow> a$m = b$m"
  82.100 +      by (auto dest: not_less_Least)
  82.101 +
  82.102 +    from ab ac bc
  82.103 +    have dab: "dist a b = inverse (2 ^ n a b)"
  82.104 +      and dac: "dist a c = inverse (2 ^ n a c)"
  82.105 +      and dbc: "dist b c = inverse (2 ^ n b c)"
  82.106 +      by (simp_all add: dist_fps_def n_def fps_eq_iff)
  82.107      from ab ac bc have nz: "dist a b \<noteq> 0" "dist a c \<noteq> 0" "dist b c \<noteq> 0"
  82.108        unfolding th by simp_all
  82.109      from nz have pos: "dist a b > 0" "dist a c > 0" "dist b c > 0"
  82.110 @@ -575,11 +531,13 @@
  82.111        assume h: "dist a b > dist a c + dist b c"
  82.112        then have gt: "dist a b > dist a c" "dist a b > dist b c"
  82.113          using pos by auto
  82.114 -      from gt have gtn: "nab < nbc" "nab < nac"
  82.115 +      from gt have gtn: "n a b < n b c" "n a b < n a c"
  82.116          unfolding dab dbc dac by (auto simp add: th1)
  82.117 -      from nac'(1)[OF gtn(2)] nbc'(1)[OF gtn(1)]
  82.118 -      have "a $ nab = b $ nab" by simp
  82.119 -      with nab'(2) have False  by simp
  82.120 +      from n'[OF gtn(2)] n'(1)[OF gtn(1)]
  82.121 +      have "a $ n a b = b $ n a b" by simp
  82.122 +      moreover have "a $ n a b \<noteq> b $ n a b"
  82.123 +         unfolding n_def by (rule LeastI_ex) (insert ab, simp add: fps_eq_iff)
  82.124 +      ultimately have False by contradiction
  82.125      }
  82.126      then have "dist a b \<le> dist a c + dist b c"
  82.127        by (auto simp add: not_le[symmetric])
  82.128 @@ -649,17 +607,12 @@
  82.129        moreover
  82.130        {
  82.131          assume neq: "?s n \<noteq> a"
  82.132 -        from fps_eq_least_unique[OF neq]
  82.133 -        obtain k where k: "leastP (\<lambda>i. ?s n $ i \<noteq> a$i) k" by blast
  82.134 -        have th0: "\<And>(a::'a fps) b. a \<noteq> b \<longleftrightarrow> (\<exists>n. a$n \<noteq> b$n)"
  82.135 -          by (simp add: fps_eq_iff)
  82.136 +        def k \<equiv> "LEAST i. ?s n $ i \<noteq> a $ i"
  82.137          from neq have dth: "dist (?s n) a = (1/2)^k"
  82.138 -          unfolding th0 dist_fps_def
  82.139 -          unfolding the1_equality[OF fps_eq_least_unique[OF neq], OF k]
  82.140 -          by (auto simp add: inverse_eq_divide power_divide)
  82.141 -
  82.142 -        from k have kn: "k > n"
  82.143 -          by (simp add: leastP_def setge_def fps_sum_rep_nth split:split_if_asm)
  82.144 +          by (auto simp add: dist_fps_def inverse_eq_divide power_divide k_def fps_eq_iff)
  82.145 +
  82.146 +        from neq have kn: "k > n"
  82.147 +          by (auto simp: fps_sum_rep_nth not_le k_def fps_eq_iff split: split_if_asm intro: LeastI2_ex)
  82.148          then have "dist (?s n) a < (1/2)^n" unfolding dth
  82.149            by (auto intro: power_strict_decreasing)
  82.150          also have "\<dots> <= (1/2)^n0" using nn0
  82.151 @@ -888,7 +841,7 @@
  82.152    using fps_deriv_linear[of 1 f 1 g] by simp
  82.153  
  82.154  lemma fps_deriv_sub[simp]: "fps_deriv ((f:: ('a::comm_ring_1) fps) - g) = fps_deriv f - fps_deriv g"
  82.155 -  unfolding diff_minus by simp
  82.156 +  using fps_deriv_add [of f "- g"] by simp
  82.157  
  82.158  lemma fps_deriv_const[simp]: "fps_deriv (fps_const c) = 0"
  82.159    by (simp add: fps_ext fps_deriv_def fps_const_def)
  82.160 @@ -978,7 +931,7 @@
  82.161  
  82.162  lemma fps_nth_deriv_sub[simp]:
  82.163    "fps_nth_deriv n ((f:: ('a::comm_ring_1) fps) - g) = fps_nth_deriv n f - fps_nth_deriv n g"
  82.164 -  unfolding diff_minus fps_nth_deriv_add by simp
  82.165 +  using fps_nth_deriv_add [of n f "- g"] by simp
  82.166  
  82.167  lemma fps_nth_deriv_0[simp]: "fps_nth_deriv n 0 = 0"
  82.168    by (induct n) simp_all
  82.169 @@ -2634,7 +2587,7 @@
  82.170    by (simp add: fps_eq_iff fps_compose_nth field_simps setsum_negf[symmetric])
  82.171  
  82.172  lemma fps_compose_sub_distrib: "(a - b) oo (c::'a::ring_1 fps) = (a oo c) - (b oo c)"
  82.173 -  unfolding diff_minus fps_compose_uminus fps_compose_add_distrib ..
  82.174 +  using fps_compose_add_distrib [of a "- b" c] by (simp add: fps_compose_uminus)
  82.175  
  82.176  lemma X_fps_compose: "X oo a = Abs_fps (\<lambda>n. if n = 0 then (0::'a::comm_ring_1) else a$n)"
  82.177    by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum_delta)
  82.178 @@ -3797,20 +3750,14 @@
  82.179    assumes "dist f g < inverse (2 ^ i)"
  82.180      and"j \<le> i"
  82.181    shows "f $ j = g $ j"
  82.182 -proof (cases "f = g")
  82.183 -  case False
  82.184 -  hence "\<exists>n. f $ n \<noteq> g $ n" by (simp add: fps_eq_iff)
  82.185 -  with assms have "i < The (leastP (\<lambda>n. f $ n \<noteq> g $ n))"
  82.186 +proof (rule ccontr)
  82.187 +  assume "f $ j \<noteq> g $ j"
  82.188 +  then have "\<exists>n. f $ n \<noteq> g $ n" by auto
  82.189 +  with assms have "i < (LEAST n. f $ n \<noteq> g $ n)"
  82.190      by (simp add: split_if_asm dist_fps_def)
  82.191 -  moreover
  82.192 -  from fps_eq_least_unique[OF `f \<noteq> g`]
  82.193 -  obtain n where n: "leastP (\<lambda>n. f$n \<noteq> g$n) n" "The (leastP (\<lambda>n. f $ n \<noteq> g $ n)) = n" by blast
  82.194 -  moreover from n have "\<And>m. m < n \<Longrightarrow> f$m = g$m" "f$n \<noteq> g$n"
  82.195 -    by (auto simp add: leastP_def setge_def)
  82.196 -  ultimately show ?thesis using `j \<le> i` by simp
  82.197 -next
  82.198 -  case True
  82.199 -  then show ?thesis by simp
  82.200 +  also have "\<dots> \<le> j"
  82.201 +    using `f $ j \<noteq> g $ j` by (auto intro: Least_le)
  82.202 +  finally show False using `j \<le> i` by simp
  82.203  qed
  82.204  
  82.205  lemma nth_equal_imp_dist_less:
  82.206 @@ -3819,18 +3766,13 @@
  82.207  proof (cases "f = g")
  82.208    case False
  82.209    hence "\<exists>n. f $ n \<noteq> g $ n" by (simp add: fps_eq_iff)
  82.210 -  with assms have "dist f g = inverse (2 ^ (The (leastP (\<lambda>n. f $ n \<noteq> g $ n))))"
  82.211 +  with assms have "dist f g = inverse (2 ^ (LEAST n. f $ n \<noteq> g $ n))"
  82.212      by (simp add: split_if_asm dist_fps_def)
  82.213    moreover
  82.214 -  from fps_eq_least_unique[OF `f \<noteq> g`]
  82.215 -  obtain n where "leastP (\<lambda>n. f$n \<noteq> g$n) n" "The (leastP (\<lambda>n. f $ n \<noteq> g $ n)) = n" by blast
  82.216 -  with assms have "i < The (leastP (\<lambda>n. f $ n \<noteq> g $ n))"
  82.217 -    by (metis (full_types) leastPD1 not_le)
  82.218 +  from assms `\<exists>n. f $ n \<noteq> g $ n` have "i < (LEAST n. f $ n \<noteq> g $ n)"
  82.219 +    by (metis (mono_tags) LeastI not_less)
  82.220    ultimately show ?thesis by simp
  82.221 -next
  82.222 -  case True
  82.223 -  then show ?thesis by simp
  82.224 -qed
  82.225 +qed simp
  82.226  
  82.227  lemma dist_less_eq_nth_equal: "dist f g < inverse (2 ^ i) \<longleftrightarrow> (\<forall>j \<le> i. f $ j = g $ j)"
  82.228    using dist_less_imp_nth_equal nth_equal_imp_dist_less by blast
    83.1 --- a/src/HOL/Library/Fraction_Field.thy	Mon Nov 11 17:34:44 2013 +0100
    83.2 +++ b/src/HOL/Library/Fraction_Field.thy	Mon Nov 11 17:44:21 2013 +0100
    83.3 @@ -132,7 +132,7 @@
    83.4  lemma diff_fract [simp]:
    83.5    assumes "b \<noteq> 0" and "d \<noteq> 0"
    83.6    shows "Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
    83.7 -  using assms by (simp add: diff_fract_def diff_minus)
    83.8 +  using assms by (simp add: diff_fract_def)
    83.9  
   83.10  definition mult_fract_def:
   83.11    "q * r = Abs_fract (\<Union>x \<in> Rep_fract q. \<Union>y \<in> Rep_fract r.
    84.1 --- a/src/HOL/Library/Function_Algebras.thy	Mon Nov 11 17:34:44 2013 +0100
    84.2 +++ b/src/HOL/Library/Function_Algebras.thy	Mon Nov 11 17:44:21 2013 +0100
    84.3 @@ -83,10 +83,10 @@
    84.4  
    84.5  instance "fun" :: (type, group_add) group_add
    84.6    by default
    84.7 -    (simp_all add: fun_eq_iff diff_minus)
    84.8 +    (simp_all add: fun_eq_iff)
    84.9  
   84.10  instance "fun" :: (type, ab_group_add) ab_group_add
   84.11 -  by default (simp_all add: diff_minus)
   84.12 +  by default simp_all
   84.13  
   84.14  
   84.15  text {* Multiplicative structures *}
    85.1 --- a/src/HOL/Library/Fundamental_Theorem_Algebra.thy	Mon Nov 11 17:34:44 2013 +0100
    85.2 +++ b/src/HOL/Library/Fundamental_Theorem_Algebra.thy	Mon Nov 11 17:44:21 2013 +0100
    85.3 @@ -156,28 +156,11 @@
    85.4  text{* An alternative useful formulation of completeness of the reals *}
    85.5  lemma real_sup_exists: assumes ex: "\<exists>x. P x" and bz: "\<exists>z. \<forall>x. P x \<longrightarrow> x < z"
    85.6    shows "\<exists>(s::real). \<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < s"
    85.7 -proof-
    85.8 -  from ex bz obtain x Y where x: "P x" and Y: "\<And>x. P x \<Longrightarrow> x < Y"  by blast
    85.9 -  from ex have thx:"\<exists>x. x \<in> Collect P" by blast
   85.10 -  from bz have thY: "\<exists>Y. isUb UNIV (Collect P) Y"
   85.11 -    by(auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def order_le_less)
   85.12 -  from reals_complete[OF thx thY] obtain L where L: "isLub UNIV (Collect P) L"
   85.13 -    by blast
   85.14 -  from Y[OF x] have xY: "x < Y" .
   85.15 -  from L have L': "\<forall>x. P x \<longrightarrow> x \<le> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)
   85.16 -  from Y have Y': "\<forall>x. P x \<longrightarrow> x \<le> Y"
   85.17 -    apply (clarsimp, atomize (full)) by auto
   85.18 -  from L Y' have "L \<le> Y" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)
   85.19 -  {fix y
   85.20 -    {fix z assume z: "P z" "y < z"
   85.21 -      from L' z have "y < L" by auto }
   85.22 -    moreover
   85.23 -    {assume yL: "y < L" "\<forall>z. P z \<longrightarrow> \<not> y < z"
   85.24 -      hence nox: "\<forall>z. P z \<longrightarrow> y \<ge> z" by auto
   85.25 -      from nox L have "y \<ge> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)
   85.26 -      with yL(1) have False  by arith}
   85.27 -    ultimately have "(\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < L" by blast}
   85.28 -  thus ?thesis by blast
   85.29 +proof
   85.30 +  from bz have "bdd_above (Collect P)"
   85.31 +    by (force intro: less_imp_le)
   85.32 +  then show "\<forall>y. (\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < Sup (Collect P)"
   85.33 +    using ex bz by (subst less_cSup_iff) auto
   85.34  qed
   85.35  
   85.36  subsection {* Fundamental theorem of algebra *}
   85.37 @@ -224,12 +207,12 @@
   85.38      from unimodular_reduce_norm[OF th0] o
   85.39      have "\<exists>v. cmod (complex_of_real (cmod b) / b + v^n) < 1"
   85.40        apply (cases "cmod (complex_of_real (cmod b) / b + 1) < 1", rule_tac x="1" in exI, simp)
   85.41 -      apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1", rule_tac x="-1" in exI, simp add: diff_minus)
   85.42 +      apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1", rule_tac x="-1" in exI, simp del: minus_one add: minus_one [symmetric])
   85.43        apply (cases "cmod (complex_of_real (cmod b) / b + ii) < 1")
   85.44        apply (cases "even m", rule_tac x="ii" in exI, simp add: m power_mult)
   85.45        apply (rule_tac x="- ii" in exI, simp add: m power_mult)
   85.46 -      apply (cases "even m", rule_tac x="- ii" in exI, simp add: m power_mult diff_minus)
   85.47 -      apply (rule_tac x="ii" in exI, simp add: m power_mult diff_minus)
   85.48 +      apply (cases "even m", rule_tac x="- ii" in exI, simp add: m power_mult)
   85.49 +      apply (rule_tac x="ii" in exI, simp add: m power_mult)
   85.50        done
   85.51      then obtain v where v: "cmod (complex_of_real (cmod b) / b + v^n) < 1" by blast
   85.52      let ?w = "v / complex_of_real (root n (cmod b))"
   85.53 @@ -954,7 +937,7 @@
   85.54  
   85.55  lemma mpoly_sub_conv:
   85.56    "poly p (x::complex) - poly q x \<equiv> poly p x + -1 * poly q x"
   85.57 -  by (simp add: diff_minus)
   85.58 +  by simp
   85.59  
   85.60  lemma poly_pad_rule: "poly p x = 0 ==> poly (pCons 0 p) x = (0::complex)" by simp
   85.61  
    86.1 --- a/src/HOL/Library/Glbs.thy	Mon Nov 11 17:34:44 2013 +0100
    86.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    86.3 @@ -1,79 +0,0 @@
    86.4 -(* Author: Amine Chaieb, University of Cambridge *)
    86.5 -
    86.6 -header {* Definitions of Lower Bounds and Greatest Lower Bounds, analogous to Lubs *}
    86.7 -
    86.8 -theory Glbs
    86.9 -imports Lubs
   86.10 -begin
   86.11 -
   86.12 -definition greatestP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a::ord \<Rightarrow> bool"
   86.13 -  where "greatestP P x = (P x \<and> Collect P *<=  x)"
   86.14 -
   86.15 -definition isLb :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"
   86.16 -  where "isLb R S x = (x <=* S \<and> x: R)"
   86.17 -
   86.18 -definition isGlb :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"
   86.19 -  where "isGlb R S x = greatestP (isLb R S) x"
   86.20 -
   86.21 -definition lbs :: "'a set \<Rightarrow> 'a::ord set \<Rightarrow> 'a set"
   86.22 -  where "lbs R S = Collect (isLb R S)"
   86.23 -
   86.24 -
   86.25 -subsection {* Rules about the Operators @{term greatestP}, @{term isLb}
   86.26 -  and @{term isGlb} *}
   86.27 -
   86.28 -lemma leastPD1: "greatestP P x \<Longrightarrow> P x"
   86.29 -  by (simp add: greatestP_def)
   86.30 -
   86.31 -lemma greatestPD2: "greatestP P x \<Longrightarrow> Collect P *<= x"
   86.32 -  by (simp add: greatestP_def)
   86.33 -
   86.34 -lemma greatestPD3: "greatestP P x \<Longrightarrow> y: Collect P \<Longrightarrow> x \<ge> y"
   86.35 -  by (blast dest!: greatestPD2 setleD)
   86.36 -
   86.37 -lemma isGlbD1: "isGlb R S x \<Longrightarrow> x <=* S"
   86.38 -  by (simp add: isGlb_def isLb_def greatestP_def)
   86.39 -
   86.40 -lemma isGlbD1a: "isGlb R S x \<Longrightarrow> x: R"
   86.41 -  by (simp add: isGlb_def isLb_def greatestP_def)
   86.42 -
   86.43 -lemma isGlb_isLb: "isGlb R S x \<Longrightarrow> isLb R S x"
   86.44 -  unfolding isLb_def by (blast dest: isGlbD1 isGlbD1a)
   86.45 -
   86.46 -lemma isGlbD2: "isGlb R S x \<Longrightarrow> y : S \<Longrightarrow> y \<ge> x"
   86.47 -  by (blast dest!: isGlbD1 setgeD)
   86.48 -
   86.49 -lemma isGlbD3: "isGlb R S x \<Longrightarrow> greatestP (isLb R S) x"
   86.50 -  by (simp add: isGlb_def)
   86.51 -
   86.52 -lemma isGlbI1: "greatestP (isLb R S) x \<Longrightarrow> isGlb R S x"
   86.53 -  by (simp add: isGlb_def)
   86.54 -
   86.55 -lemma isGlbI2: "isLb R S x \<Longrightarrow> Collect (isLb R S) *<= x \<Longrightarrow> isGlb R S x"
   86.56 -  by (simp add: isGlb_def greatestP_def)
   86.57 -
   86.58 -lemma isLbD: "isLb R S x \<Longrightarrow> y : S \<Longrightarrow> y \<ge> x"
   86.59 -  by (simp add: isLb_def setge_def)
   86.60 -
   86.61 -lemma isLbD2: "isLb R S x \<Longrightarrow> x <=* S "
   86.62 -  by (simp add: isLb_def)
   86.63 -
   86.64 -lemma isLbD2a: "isLb R S x \<Longrightarrow> x: R"
   86.65 -  by (simp add: isLb_def)
   86.66 -
   86.67 -lemma isLbI: "x <=* S \<Longrightarrow> x: R \<Longrightarrow> isLb R S x"
   86.68 -  by (simp add: isLb_def)
   86.69 -
   86.70 -lemma isGlb_le_isLb: "isGlb R S x \<Longrightarrow> isLb R S y \<Longrightarrow> x \<ge> y"
   86.71 -  unfolding isGlb_def by (blast intro!: greatestPD3)
   86.72 -
   86.73 -lemma isGlb_ubs: "isGlb R S x \<Longrightarrow> lbs R S *<= x"
   86.74 -  unfolding lbs_def isGlb_def by (rule greatestPD2)
   86.75 -
   86.76 -lemma isGlb_unique: "[| isGlb R S x; isGlb R S y |] ==> x = (y::'a::linorder)"
   86.77 -  apply (frule isGlb_isLb)
   86.78 -  apply (frule_tac x = y in isGlb_isLb)
   86.79 -  apply (blast intro!: order_antisym dest!: isGlb_le_isLb)
   86.80 -  done
   86.81 -
   86.82 -end
    87.1 --- a/src/HOL/Library/Inner_Product.thy	Mon Nov 11 17:34:44 2013 +0100
    87.2 +++ b/src/HOL/Library/Inner_Product.thy	Mon Nov 11 17:44:21 2013 +0100
    87.3 @@ -41,7 +41,7 @@
    87.4    using inner_add_left [of x "- x" y] by simp
    87.5  
    87.6  lemma inner_diff_left: "inner (x - y) z = inner x z - inner y z"
    87.7 -  by (simp add: diff_minus inner_add_left)
    87.8 +  using inner_add_left [of x "- y" z] by simp
    87.9  
   87.10  lemma inner_setsum_left: "inner (\<Sum>x\<in>A. f x) y = (\<Sum>x\<in>A. inner (f x) y)"
   87.11    by (cases "finite A", induct set: finite, simp_all add: inner_add_left)
    88.1 --- a/src/HOL/Library/Lattice_Algebras.thy	Mon Nov 11 17:34:44 2013 +0100
    88.2 +++ b/src/HOL/Library/Lattice_Algebras.thy	Mon Nov 11 17:44:21 2013 +0100
    88.3 @@ -13,9 +13,7 @@
    88.4    apply (rule antisym)
    88.5    apply (simp_all add: le_infI)
    88.6    apply (rule add_le_imp_le_left [of "uminus a"])
    88.7 -  apply (simp only: add_assoc [symmetric], simp)
    88.8 -  apply rule
    88.9 -  apply (rule add_le_imp_le_left[of "a"], simp only: add_assoc[symmetric], simp)+
   88.10 +  apply (simp only: add_assoc [symmetric], simp add: diff_le_eq add.commute)
   88.11    done
   88.12  
   88.13  lemma add_inf_distrib_right: "inf a b + c = inf (a + c) (b + c)"
   88.14 @@ -33,11 +31,10 @@
   88.15  lemma add_sup_distrib_left: "a + sup b c = sup (a + b) (a + c)"
   88.16    apply (rule antisym)
   88.17    apply (rule add_le_imp_le_left [of "uminus a"])
   88.18 -  apply (simp only: add_assoc[symmetric], simp)
   88.19 -  apply rule
   88.20 +  apply (simp only: add_assoc [symmetric], simp)
   88.21 +  apply (simp add: le_diff_eq add.commute)
   88.22 +  apply (rule le_supI)
   88.23    apply (rule add_le_imp_le_left [of "a"], simp only: add_assoc[symmetric], simp)+
   88.24 -  apply (rule le_supI)
   88.25 -  apply (simp_all)
   88.26    done
   88.27  
   88.28  lemma add_sup_distrib_right: "sup a b + c = sup (a+c) (b+c)"
   88.29 @@ -87,9 +84,15 @@
   88.30  lemma neg_inf_eq_sup: "- inf a b = sup (-a) (-b)"
   88.31    by (simp add: inf_eq_neg_sup)
   88.32  
   88.33 +lemma diff_inf_eq_sup: "a - inf b c = a + sup (- b) (- c)"
   88.34 +  using neg_inf_eq_sup [of b c, symmetric] by simp
   88.35 +
   88.36  lemma neg_sup_eq_inf: "- sup a b = inf (-a) (-b)"
   88.37    by (simp add: sup_eq_neg_inf)
   88.38  
   88.39 +lemma diff_sup_eq_inf: "a - sup b c = a + inf (- b) (- c)"
   88.40 +  using neg_sup_eq_inf [of b c, symmetric] by simp
   88.41 +
   88.42  lemma add_eq_inf_sup: "a + b = sup a b + inf a b"
   88.43  proof -
   88.44    have "0 = - inf 0 (a-b) + inf (a-b) 0"
   88.45 @@ -97,8 +100,8 @@
   88.46    hence "0 = sup 0 (b-a) + inf (a-b) 0"
   88.47      by (simp add: inf_eq_neg_sup)
   88.48    hence "0 = (-a + sup a b) + (inf a b + (-b))"
   88.49 -    by (simp add: add_sup_distrib_left add_inf_distrib_right) (simp add: algebra_simps)
   88.50 -  thus ?thesis by (simp add: algebra_simps)
   88.51 +    by (simp only: add_sup_distrib_left add_inf_distrib_right) simp
   88.52 +  then show ?thesis by (simp add: algebra_simps)
   88.53  qed
   88.54  
   88.55  
   88.56 @@ -251,7 +254,7 @@
   88.57      apply assumption
   88.58      apply (rule notI)
   88.59      unfolding double_zero [symmetric, of a]
   88.60 -    apply simp
   88.61 +    apply blast
   88.62      done
   88.63  qed
   88.64  
   88.65 @@ -259,7 +262,8 @@
   88.66    "a + a \<le> 0 \<longleftrightarrow> a \<le> 0"
   88.67  proof -
   88.68    have "a + a \<le> 0 \<longleftrightarrow> 0 \<le> - (a + a)" by (subst le_minus_iff, simp)
   88.69 -  moreover have "\<dots> \<longleftrightarrow> a \<le> 0" by simp
   88.70 +  moreover have "\<dots> \<longleftrightarrow> a \<le> 0"
   88.71 +    by (simp only: minus_add_distrib zero_le_double_add_iff_zero_le_single_add) simp
   88.72    ultimately show ?thesis by blast
   88.73  qed
   88.74  
   88.75 @@ -267,11 +271,12 @@
   88.76    "a + a < 0 \<longleftrightarrow> a < 0"
   88.77  proof -
   88.78    have "a + a < 0 \<longleftrightarrow> 0 < - (a + a)" by (subst less_minus_iff, simp)
   88.79 -  moreover have "\<dots> \<longleftrightarrow> a < 0" by simp
   88.80 +  moreover have "\<dots> \<longleftrightarrow> a < 0"
   88.81 +    by (simp only: minus_add_distrib zero_less_double_add_iff_zero_less_single_add) simp
   88.82    ultimately show ?thesis by blast
   88.83  qed
   88.84  
   88.85 -declare neg_inf_eq_sup [simp] neg_sup_eq_inf [simp]
   88.86 +declare neg_inf_eq_sup [simp] neg_sup_eq_inf [simp] diff_inf_eq_sup [simp] diff_sup_eq_inf [simp]
   88.87  
   88.88  lemma le_minus_self_iff: "a \<le> - a \<longleftrightarrow> a \<le> 0"
   88.89  proof -
   88.90 @@ -326,7 +331,7 @@
   88.91    then have "0 \<le> sup a (- a)" unfolding abs_lattice .
   88.92    then have "sup (sup a (- a)) 0 = sup a (- a)" by (rule sup_absorb1)
   88.93    then show ?thesis
   88.94 -    by (simp add: add_sup_inf_distribs sup_aci pprt_def nprt_def diff_minus abs_lattice)
   88.95 +    by (simp add: add_sup_inf_distribs ac_simps pprt_def nprt_def abs_lattice)
   88.96  qed
   88.97  
   88.98  subclass ordered_ab_group_add_abs
   88.99 @@ -355,16 +360,17 @@
  88.100    show "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
  88.101    proof -
  88.102      have g:"abs a + abs b = sup (a+b) (sup (-a-b) (sup (-a+b) (a + (-b))))" (is "_=sup ?m ?n")
  88.103 -      by (simp add: abs_lattice add_sup_inf_distribs sup_aci diff_minus)
  88.104 +      by (simp add: abs_lattice add_sup_inf_distribs sup_aci ac_simps)
  88.105      have a: "a + b <= sup ?m ?n" by simp
  88.106      have b: "- a - b <= ?n" by simp
  88.107      have c: "?n <= sup ?m ?n" by simp
  88.108      from b c have d: "-a-b <= sup ?m ?n" by (rule order_trans)
  88.109 -    have e:"-a-b = -(a+b)" by (simp add: diff_minus)
  88.110 +    have e:"-a-b = -(a+b)" by simp
  88.111      from a d e have "abs(a+b) <= sup ?m ?n"
  88.112        apply -
  88.113        apply (drule abs_leI)
  88.114 -      apply auto
  88.115 +      apply (simp_all only: algebra_simps ac_simps minus_add)
  88.116 +      apply (metis add_uminus_conv_diff d sup_commute uminus_add_conv_diff)
  88.117        done
  88.118      with g[symmetric] show ?thesis by simp
  88.119    qed
  88.120 @@ -421,14 +427,12 @@
  88.121    }
  88.122    note b = this[OF refl[of a] refl[of b]]
  88.123    have xy: "- ?x <= ?y"
  88.124 -    apply (simp)
  88.125 -    apply (rule order_trans [OF add_nonpos_nonpos add_nonneg_nonneg])
  88.126 -    apply (simp_all add: mult_nonneg_nonneg mult_nonpos_nonpos)
  88.127 +    apply simp
  88.128 +    apply (metis (full_types) add_increasing add_uminus_conv_diff lattice_ab_group_add_class.minus_le_self_iff minus_add_distrib mult_nonneg_nonneg mult_nonpos_nonpos nprt_le_zero zero_le_pprt)
  88.129      done
  88.130    have yx: "?y <= ?x"
  88.131 -    apply (simp add:diff_minus)
  88.132 -    apply (rule order_trans [OF add_nonpos_nonpos add_nonneg_nonneg])
  88.133 -    apply (simp_all add: mult_nonneg_nonpos mult_nonpos_nonneg)
  88.134 +    apply simp
  88.135 +    apply (metis (full_types) add_nonpos_nonpos add_uminus_conv_diff lattice_ab_group_add_class.le_minus_self_iff minus_add_distrib mult_nonneg_nonpos mult_nonpos_nonneg nprt_le_zero zero_le_pprt)
  88.136      done
  88.137    have i1: "a*b <= abs a * abs b" by (simp only: a b yx)
  88.138    have i2: "- (abs a * abs b) <= a*b" by (simp only: a b xy)
  88.139 @@ -549,7 +553,7 @@
  88.140      by simp
  88.141    then have "-(- nprt a1 * pprt b2 + - nprt a2 * nprt b2 + - pprt a1 * pprt b1 + - pprt a2 * nprt b1) <= a * b"
  88.142      by (simp only: minus_le_iff)
  88.143 -  then show ?thesis by simp
  88.144 +  then show ?thesis by (simp add: algebra_simps)
  88.145  qed
  88.146  
  88.147  instance int :: lattice_ring
  88.148 @@ -567,3 +571,4 @@
  88.149  qed
  88.150  
  88.151  end
  88.152 +
    89.1 --- a/src/HOL/Library/Library.thy	Mon Nov 11 17:34:44 2013 +0100
    89.2 +++ b/src/HOL/Library/Library.thy	Mon Nov 11 17:44:21 2013 +0100
    89.3 @@ -1,7 +1,6 @@
    89.4  (*<*)
    89.5  theory Library
    89.6  imports
    89.7 -  Abstract_Rat
    89.8    AList
    89.9    BigO
   89.10    Binomial
   89.11 @@ -65,7 +64,6 @@
   89.12    Sublist
   89.13    Sum_of_Squares
   89.14    Transitive_Closure_Table
   89.15 -  Univ_Poly
   89.16    Wfrec
   89.17    While_Combinator
   89.18    Zorn
    90.1 --- a/src/HOL/Library/Liminf_Limsup.thy	Mon Nov 11 17:34:44 2013 +0100
    90.2 +++ b/src/HOL/Library/Liminf_Limsup.thy	Mon Nov 11 17:44:21 2013 +0100
    90.3 @@ -21,19 +21,21 @@
    90.4    by (blast intro: less_imp_le less_trans le_less_trans dest: dense)
    90.5  
    90.6  lemma SUPR_pair:
    90.7 -  "(SUP i : A. SUP j : B. f i j) = (SUP p : A \<times> B. f (fst p) (snd p))"
    90.8 +  fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _ :: complete_lattice"
    90.9 +  shows "(SUP i : A. SUP j : B. f i j) = (SUP p : A \<times> B. f (fst p) (snd p))"
   90.10    by (rule antisym) (auto intro!: SUP_least SUP_upper2)
   90.11  
   90.12  lemma INFI_pair:
   90.13 -  "(INF i : A. INF j : B. f i j) = (INF p : A \<times> B. f (fst p) (snd p))"
   90.14 +  fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _ :: complete_lattice"
   90.15 +  shows "(INF i : A. INF j : B. f i j) = (INF p : A \<times> B. f (fst p) (snd p))"
   90.16    by (rule antisym) (auto intro!: INF_greatest INF_lower2)
   90.17  
   90.18  subsubsection {* @{text Liminf} and @{text Limsup} *}
   90.19  
   90.20 -definition
   90.21 +definition Liminf :: "'a filter \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b :: complete_lattice" where
   90.22    "Liminf F f = (SUP P:{P. eventually P F}. INF x:{x. P x}. f x)"
   90.23  
   90.24 -definition
   90.25 +definition Limsup :: "'a filter \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b :: complete_lattice" where
   90.26    "Limsup F f = (INF P:{P. eventually P F}. SUP x:{x. P x}. f x)"
   90.27  
   90.28  abbreviation "liminf \<equiv> Liminf sequentially"
   90.29 @@ -50,21 +52,17 @@
   90.30      (\<And>y. (\<And>P. eventually P F \<Longrightarrow> y \<le> SUPR (Collect P) f) \<Longrightarrow> y \<le> x) \<Longrightarrow> Limsup F f = x"
   90.31    unfolding Limsup_def by (auto intro!: INF_eqI)
   90.32  
   90.33 -lemma liminf_SUPR_INFI:
   90.34 -  fixes f :: "nat \<Rightarrow> 'a :: complete_lattice"
   90.35 -  shows "liminf f = (SUP n. INF m:{n..}. f m)"
   90.36 +lemma liminf_SUPR_INFI: "liminf f = (SUP n. INF m:{n..}. f m)"
   90.37    unfolding Liminf_def eventually_sequentially
   90.38    by (rule SUPR_eq) (auto simp: atLeast_def intro!: INF_mono)
   90.39  
   90.40 -lemma limsup_INFI_SUPR:
   90.41 -  fixes f :: "nat \<Rightarrow> 'a :: complete_lattice"
   90.42 -  shows "limsup f = (INF n. SUP m:{n..}. f m)"
   90.43 +lemma limsup_INFI_SUPR: "limsup f = (INF n. SUP m:{n..}. f m)"
   90.44    unfolding Limsup_def eventually_sequentially
   90.45    by (rule INFI_eq) (auto simp: atLeast_def intro!: SUP_mono)
   90.46  
   90.47  lemma Limsup_const: 
   90.48    assumes ntriv: "\<not> trivial_limit F"
   90.49 -  shows "Limsup F (\<lambda>x. c) = (c::'a::complete_lattice)"
   90.50 +  shows "Limsup F (\<lambda>x. c) = c"
   90.51  proof -
   90.52    have *: "\<And>P. Ex P \<longleftrightarrow> P \<noteq> (\<lambda>x. False)" by auto
   90.53    have "\<And>P. eventually P F \<Longrightarrow> (SUP x : {x. P x}. c) = c"
   90.54 @@ -77,7 +75,7 @@
   90.55  
   90.56  lemma Liminf_const:
   90.57    assumes ntriv: "\<not> trivial_limit F"
   90.58 -  shows "Liminf F (\<lambda>x. c) = (c::'a::complete_lattice)"
   90.59 +  shows "Liminf F (\<lambda>x. c) = c"
   90.60  proof -
   90.61    have *: "\<And>P. Ex P \<longleftrightarrow> P \<noteq> (\<lambda>x. False)" by auto
   90.62    have "\<And>P. eventually P F \<Longrightarrow> (INF x : {x. P x}. c) = c"
   90.63 @@ -89,7 +87,6 @@
   90.64  qed
   90.65  
   90.66  lemma Liminf_mono:
   90.67 -  fixes f g :: "'a => 'b :: complete_lattice"
   90.68    assumes ev: "eventually (\<lambda>x. f x \<le> g x) F"
   90.69    shows "Liminf F f \<le> Liminf F g"
   90.70    unfolding Liminf_def
   90.71 @@ -101,13 +98,11 @@
   90.72  qed
   90.73  
   90.74  lemma Liminf_eq:
   90.75 -  fixes f g :: "'a \<Rightarrow> 'b :: complete_lattice"
   90.76    assumes "eventually (\<lambda>x. f x = g x) F"
   90.77    shows "Liminf F f = Liminf F g"
   90.78    by (intro antisym Liminf_mono eventually_mono[OF _ assms]) auto
   90.79  
   90.80  lemma Limsup_mono:
   90.81 -  fixes f g :: "'a \<Rightarrow> 'b :: complete_lattice"
   90.82    assumes ev: "eventually (\<lambda>x. f x \<le> g x) F"
   90.83    shows "Limsup F f \<le> Limsup F g"
   90.84    unfolding Limsup_def
   90.85 @@ -119,18 +114,16 @@
   90.86  qed
   90.87  
   90.88  lemma Limsup_eq:
   90.89 -  fixes f g :: "'a \<Rightarrow> 'b :: complete_lattice"
   90.90    assumes "eventually (\<lambda>x. f x = g x) net"
   90.91    shows "Limsup net f = Limsup net g"
   90.92    by (intro antisym Limsup_mono eventually_mono[OF _ assms]) auto
   90.93  
   90.94  lemma Liminf_le_Limsup:
   90.95 -  fixes f :: "'a \<Rightarrow> 'b::complete_lattice"
   90.96    assumes ntriv: "\<not> trivial_limit F"
   90.97    shows "Liminf F f \<le> Limsup F f"
   90.98    unfolding Limsup_def Liminf_def
   90.99 -  apply (rule complete_lattice_class.SUP_least)
  90.100 -  apply (rule complete_lattice_class.INF_greatest)
  90.101 +  apply (rule SUP_least)
  90.102 +  apply (rule INF_greatest)
  90.103  proof safe
  90.104    fix P Q assume "eventually P F" "eventually Q F"
  90.105    then have "eventually (\<lambda>x. P x \<and> Q x) F" (is "eventually ?C F") by (rule eventually_conj)
  90.106 @@ -146,14 +139,12 @@
  90.107  qed
  90.108  
  90.109  lemma Liminf_bounded:
  90.110 -  fixes X Y :: "'a \<Rightarrow> 'b::complete_lattice"
  90.111    assumes ntriv: "\<not> trivial_limit F"
  90.112    assumes le: "eventually (\<lambda>n. C \<le> X n) F"
  90.113    shows "C \<le> Liminf F X"
  90.114    using Liminf_mono[OF le] Liminf_const[OF ntriv, of C] by simp
  90.115  
  90.116  lemma Limsup_bounded:
  90.117 -  fixes X Y :: "'a \<Rightarrow> 'b::complete_lattice"
  90.118    assumes ntriv: "\<not> trivial_limit F"
  90.119    assumes le: "eventually (\<lambda>n. X n \<le> C) F"
  90.120    shows "Limsup F X \<le> C"
    91.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
    91.2 +++ b/src/HOL/Library/Lubs_Glbs.thy	Mon Nov 11 17:44:21 2013 +0100
    91.3 @@ -0,0 +1,245 @@
    91.4 +(*  Title:      HOL/Library/Lubs_Glbs.thy
    91.5 +    Author:     Jacques D. Fleuriot, University of Cambridge
    91.6 +    Author:     Amine Chaieb, University of Cambridge *)
    91.7 +
    91.8 +header {* Definitions of Least Upper Bounds and Greatest Lower Bounds *}
    91.9 +
   91.10 +theory Lubs_Glbs
   91.11 +imports Complex_Main
   91.12 +begin
   91.13 +
   91.14 +text {* Thanks to suggestions by James Margetson *}
   91.15 +
   91.16 +definition setle :: "'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"  (infixl "*<=" 70)
   91.17 +  where "S *<= x = (ALL y: S. y \<le> x)"
   91.18 +
   91.19 +definition setge :: "'a::ord \<Rightarrow> 'a set \<Rightarrow> bool"  (infixl "<=*" 70)
   91.20 +  where "x <=* S = (ALL y: S. x \<le> y)"
   91.21 +
   91.22 +
   91.23 +subsection {* Rules for the Relations @{text "*<="} and @{text "<=*"} *}
   91.24 +
   91.25 +lemma setleI: "ALL y: S. y \<le> x \<Longrightarrow> S *<= x"
   91.26 +  by (simp add: setle_def)
   91.27 +
   91.28 +lemma setleD: "S *<= x \<Longrightarrow> y: S \<Longrightarrow> y \<le> x"
   91.29 +  by (simp add: setle_def)
   91.30 +
   91.31 +lemma setgeI: "ALL y: S. x \<le> y \<Longrightarrow> x <=* S"
   91.32 +  by (simp add: setge_def)
   91.33 +
   91.34 +lemma setgeD: "x <=* S \<Longrightarrow> y: S \<Longrightarrow> x \<le> y"
   91.35 +  by (simp add: setge_def)
   91.36 +
   91.37 +
   91.38 +definition leastP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a::ord \<Rightarrow> bool"
   91.39 +  where "leastP P x = (P x \<and> x <=* Collect P)"
   91.40 +
   91.41 +definition isUb :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"
   91.42 +  where "isUb R S x = (S *<= x \<and> x: R)"
   91.43 +
   91.44 +definition isLub :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"
   91.45 +  where "isLub R S x = leastP (isUb R S) x"
   91.46 +
   91.47 +definition ubs :: "'a set \<Rightarrow> 'a::ord set \<Rightarrow> 'a set"
   91.48 +  where "ubs R S = Collect (isUb R S)"
   91.49 +
   91.50 +
   91.51 +subsection {* Rules about the Operators @{term leastP}, @{term ub} and @{term lub} *}
   91.52 +
   91.53 +lemma leastPD1: "leastP P x \<Longrightarrow> P x"
   91.54 +  by (simp add: leastP_def)
   91.55 +
   91.56 +lemma leastPD2: "leastP P x \<Longrightarrow> x <=* Collect P"
   91.57 +  by (simp add: leastP_def)
   91.58 +
   91.59 +lemma leastPD3: "leastP P x \<Longrightarrow> y: Collect P \<Longrightarrow> x \<le> y"
   91.60 +  by (blast dest!: leastPD2 setgeD)
   91.61 +
   91.62 +lemma isLubD1: "isLub R S x \<Longrightarrow> S *<= x"
   91.63 +  by (simp add: isLub_def isUb_def leastP_def)
   91.64 +
   91.65 +lemma isLubD1a: "isLub R S x \<Longrightarrow> x: R"
   91.66 +  by (simp add: isLub_def isUb_def leastP_def)
   91.67 +
   91.68 +lemma isLub_isUb: "isLub R S x \<Longrightarrow> isUb R S x"
   91.69 +  unfolding isUb_def by (blast dest: isLubD1 isLubD1a)
   91.70 +
   91.71 +lemma isLubD2: "isLub R S x \<Longrightarrow> y : S \<Longrightarrow> y \<le> x"
   91.72 +  by (blast dest!: isLubD1 setleD)
   91.73 +
   91.74 +lemma isLubD3: "isLub R S x \<Longrightarrow> leastP (isUb R S) x"
   91.75 +  by (simp add: isLub_def)
   91.76 +
   91.77 +lemma isLubI1: "leastP(isUb R S) x \<Longrightarrow> isLub R S x"
   91.78 +  by (simp add: isLub_def)
   91.79 +
   91.80 +lemma isLubI2: "isUb R S x \<Longrightarrow> x <=* Collect (isUb R S) \<Longrightarrow> isLub R S x"
   91.81 +  by (simp add: isLub_def leastP_def)
   91.82 +
   91.83 +lemma isUbD: "isUb R S x \<Longrightarrow> y : S \<Longrightarrow> y \<le> x"
   91.84 +  by (simp add: isUb_def setle_def)
   91.85 +
   91.86 +lemma isUbD2: "isUb R S x \<Longrightarrow> S *<= x"
   91.87 +  by (simp add: isUb_def)
   91.88 +
   91.89 +lemma isUbD2a: "isUb R S x \<Longrightarrow> x: R"
   91.90 +  by (simp add: isUb_def)
   91.91 +
   91.92 +lemma isUbI: "S *<= x \<Longrightarrow> x: R \<Longrightarrow> isUb R S x"
   91.93 +  by (simp add: isUb_def)
   91.94 +
   91.95 +lemma isLub_le_isUb: "isLub R S x \<Longrightarrow> isUb R S y \<Longrightarrow> x \<le> y"
   91.96 +  unfolding isLub_def by (blast intro!: leastPD3)
   91.97 +
   91.98 +lemma isLub_ubs: "isLub R S x \<Longrightarrow> x <=* ubs R S"
   91.99 +  unfolding ubs_def isLub_def by (rule leastPD2)
  91.100 +
  91.101 +lemma isLub_unique: "[| isLub R S x; isLub R S y |] ==> x = (y::'a::linorder)"
  91.102 +  apply (frule isLub_isUb)
  91.103 +  apply (frule_tac x = y in isLub_isUb)
  91.104 +  apply (blast intro!: order_antisym dest!: isLub_le_isUb)
  91.105 +  done
  91.106 +
  91.107 +lemma isUb_UNIV_I: "(\<And>y. y \<in> S \<Longrightarrow> y \<le> u) \<Longrightarrow> isUb UNIV S u"
  91.108 +  by (simp add: isUbI setleI)
  91.109 +
  91.110 +
  91.111 +definition greatestP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a::ord \<Rightarrow> bool"
  91.112 +  where "greatestP P x = (P x \<and> Collect P *<=  x)"
  91.113 +
  91.114 +definition isLb :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"
  91.115 +  where "isLb R S x = (x <=* S \<and> x: R)"
  91.116 +
  91.117 +definition isGlb :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"
  91.118 +  where "isGlb R S x = greatestP (isLb R S) x"
  91.119 +
  91.120 +definition lbs :: "'a set \<Rightarrow> 'a::ord set \<Rightarrow> 'a set"
  91.121 +  where "lbs R S = Collect (isLb R S)"
  91.122 +
  91.123 +
  91.124 +subsection {* Rules about the Operators @{term greatestP}, @{term isLb} and @{term isGlb} *}
  91.125 +
  91.126 +lemma greatestPD1: "greatestP P x \<Longrightarrow> P x"
  91.127 +  by (simp add: greatestP_def)
  91.128 +
  91.129 +lemma greatestPD2: "greatestP P x \<Longrightarrow> Collect P *<= x"
  91.130 +  by (simp add: greatestP_def)
  91.131 +
  91.132 +lemma greatestPD3: "greatestP P x \<Longrightarrow> y: Collect P \<Longrightarrow> x \<ge> y"
  91.133 +  by (blast dest!: greatestPD2 setleD)
  91.134 +
  91.135 +lemma isGlbD1: "isGlb R S x \<Longrightarrow> x <=* S"
  91.136 +  by (simp add: isGlb_def isLb_def greatestP_def)
  91.137 +
  91.138 +lemma isGlbD1a: "isGlb R S x \<Longrightarrow> x: R"
  91.139 +  by (simp add: isGlb_def isLb_def greatestP_def)
  91.140 +
  91.141 +lemma isGlb_isLb: "isGlb R S x \<Longrightarrow> isLb R S x"
  91.142 +  unfolding isLb_def by (blast dest: isGlbD1 isGlbD1a)
  91.143 +
  91.144 +lemma isGlbD2: "isGlb R S x \<Longrightarrow> y : S \<Longrightarrow> y \<ge> x"
  91.145 +  by (blast dest!: isGlbD1 setgeD)
  91.146 +
  91.147 +lemma isGlbD3: "isGlb R S x \<Longrightarrow> greatestP (isLb R S) x"
  91.148 +  by (simp add: isGlb_def)
  91.149 +
  91.150 +lemma isGlbI1: "greatestP (isLb R S) x \<Longrightarrow> isGlb R S x"
  91.151 +  by (simp add: isGlb_def)
  91.152 +
  91.153 +lemma isGlbI2: "isLb R S x \<Longrightarrow> Collect (isLb R S) *<= x \<Longrightarrow> isGlb R S x"
  91.154 +  by (simp add: isGlb_def greatestP_def)
  91.155 +
  91.156 +lemma isLbD: "isLb R S x \<Longrightarrow> y : S \<Longrightarrow> y \<ge> x"
  91.157 +  by (simp add: isLb_def setge_def)
  91.158 +
  91.159 +lemma isLbD2: "isLb R S x \<Longrightarrow> x <=* S "
  91.160 +  by (simp add: isLb_def)
  91.161 +
  91.162 +lemma isLbD2a: "isLb R S x \<Longrightarrow> x: R"
  91.163 +  by (simp add: isLb_def)
  91.164 +
  91.165 +lemma isLbI: "x <=* S \<Longrightarrow> x: R \<Longrightarrow> isLb R S x"
  91.166 +  by (simp add: isLb_def)
  91.167 +
  91.168 +lemma isGlb_le_isLb: "isGlb R S x \<Longrightarrow> isLb R S y \<Longrightarrow> x \<ge> y"
  91.169 +  unfolding isGlb_def by (blast intro!: greatestPD3)
  91.170 +
  91.171 +lemma isGlb_ubs: "isGlb R S x \<Longrightarrow> lbs R S *<= x"
  91.172 +  unfolding lbs_def isGlb_def by (rule greatestPD2)
  91.173 +
  91.174 +lemma isGlb_unique: "[| isGlb R S x; isGlb R S y |] ==> x = (y::'a::linorder)"
  91.175 +  apply (frule isGlb_isLb)
  91.176 +  apply (frule_tac x = y in isGlb_isLb)
  91.177 +  apply (blast intro!: order_antisym dest!: isGlb_le_isLb)
  91.178 +  done
  91.179 +
  91.180 +lemma bdd_above_setle: "bdd_above A \<longleftrightarrow> (\<exists>a. A *<= a)"
  91.181 +  by (auto simp: bdd_above_def setle_def)
  91.182 +
  91.183 +lemma bdd_below_setge: "bdd_below A \<longleftrightarrow> (\<exists>a. a <=* A)"
  91.184 +  by (auto simp: bdd_below_def setge_def)
  91.185 +
  91.186 +lemma isLub_cSup: 
  91.187 +  "(S::'a :: conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> (\<exists>b. S *<= b) \<Longrightarrow> isLub UNIV S (Sup S)"
  91.188 +  by  (auto simp add: isLub_def setle_def leastP_def isUb_def
  91.189 +            intro!: setgeI cSup_upper cSup_least)
  91.190 +
  91.191 +lemma isGlb_cInf: 
  91.192 +  "(S::'a :: conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> (\<exists>b. b <=* S) \<Longrightarrow> isGlb UNIV S (Inf S)"
  91.193 +  by  (auto simp add: isGlb_def setge_def greatestP_def isLb_def
  91.194 +            intro!: setleI cInf_lower cInf_greatest)
  91.195 +
  91.196 +lemma cSup_le: "(S::'a::conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> S *<= b \<Longrightarrow> Sup S \<le> b"
  91.197 +  by (metis cSup_least setle_def)
  91.198 +
  91.199 +lemma cInf_ge: "(S::'a :: conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> b <=* S \<Longrightarrow> Inf S \<ge> b"
  91.200 +  by (metis cInf_greatest setge_def)
  91.201 +
  91.202 +lemma cSup_bounds:
  91.203 +  fixes S :: "'a :: conditionally_complete_lattice set"
  91.204 +  shows "S \<noteq> {} \<Longrightarrow> a <=* S \<Longrightarrow> S *<= b \<Longrightarrow> a \<le> Sup S \<and> Sup S \<le> b"
  91.205 +  using cSup_least[of S b] cSup_upper2[of _ S a]
  91.206 +  by (auto simp: bdd_above_setle setge_def setle_def)
  91.207 +
  91.208 +lemma cSup_unique: "(S::'a :: {conditionally_complete_linorder, no_bot} set) *<= b \<Longrightarrow> (\<forall>b'<b. \<exists>x\<in>S. b' < x) \<Longrightarrow> Sup S = b"
  91.209 +  by (rule cSup_eq) (auto simp: not_le[symmetric] setle_def)
  91.210 +
  91.211 +lemma cInf_unique: "b <=* (S::'a :: {conditionally_complete_linorder, no_top} set) \<Longrightarrow> (\<forall>b'>b. \<exists>x\<in>S. b' > x) \<Longrightarrow> Inf S = b"
  91.212 +  by (rule cInf_eq) (auto simp: not_le[symmetric] setge_def)
  91.213 +
  91.214 +text{* Use completeness of reals (supremum property) to show that any bounded sequence has a least upper bound*}
  91.215 +
  91.216 +lemma reals_complete: "\<exists>X. X \<in> S \<Longrightarrow> \<exists>Y. isUb (UNIV::real set) S Y \<Longrightarrow> \<exists>t. isLub (UNIV :: real set) S t"
  91.217 +  by (intro exI[of _ "Sup S"] isLub_cSup) (auto simp: setle_def isUb_def intro!: cSup_upper)
  91.218 +
  91.219 +lemma Bseq_isUb: "\<And>X :: nat \<Rightarrow> real. Bseq X \<Longrightarrow> \<exists>U. isUb (UNIV::real set) {x. \<exists>n. X n = x} U"
  91.220 +  by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff)
  91.221 +
  91.222 +lemma Bseq_isLub: "\<And>X :: nat \<Rightarrow> real. Bseq X \<Longrightarrow> \<exists>U. isLub (UNIV::real set) {x. \<exists>n. X n = x} U"
  91.223 +  by (blast intro: reals_complete Bseq_isUb)
  91.224 +
  91.225 +lemma isLub_mono_imp_LIMSEQ:
  91.226 +  fixes X :: "nat \<Rightarrow> real"
  91.227 +  assumes u: "isLub UNIV {x. \<exists>n. X n = x} u" (* FIXME: use 'range X' *)
  91.228 +  assumes X: "\<forall>m n. m \<le> n \<longrightarrow> X m \<le> X n"
  91.229 +  shows "X ----> u"
  91.230 +proof -
  91.231 +  have "X ----> (SUP i. X i)"
  91.232 +    using u[THEN isLubD1] X
  91.233 +    by (intro LIMSEQ_incseq_SUP) (auto simp: incseq_def image_def eq_commute bdd_above_setle)
  91.234 +  also have "(SUP i. X i) = u"
  91.235 +    using isLub_cSup[of "range X"] u[THEN isLubD1]
  91.236 +    by (intro isLub_unique[OF _ u]) (auto simp add: SUP_def image_def eq_commute)
  91.237 +  finally show ?thesis .
  91.238 +qed
  91.239 +
  91.240 +lemmas real_isGlb_unique = isGlb_unique[where 'a=real]
  91.241 +
  91.242 +lemma real_le_inf_subset: "t \<noteq> {} \<Longrightarrow> t \<subseteq> s \<Longrightarrow> \<exists>b. b <=* s \<Longrightarrow> Inf s \<le> Inf (t::real set)"
  91.243 +  by (rule cInf_superset_mono) (auto simp: bdd_below_setge)
  91.244 +
  91.245 +lemma real_ge_sup_subset: "t \<noteq> {} \<Longrightarrow> t \<subseteq> s \<Longrightarrow> \<exists>b. s *<= b \<Longrightarrow> Sup s \<ge> Sup (t::real set)"
  91.246 +  by (rule cSup_subset_mono) (auto simp: bdd_above_setle)
  91.247 +
  91.248 +end
    92.1 --- a/src/HOL/Library/Multiset.thy	Mon Nov 11 17:34:44 2013 +0100
    92.2 +++ b/src/HOL/Library/Multiset.thy	Mon Nov 11 17:44:21 2013 +0100
    92.3 @@ -1533,10 +1533,10 @@
    92.4    qed
    92.5  qed
    92.6  
    92.7 -lemma all_accessible: "wf r ==> \<forall>M. M \<in> acc (mult1 r)"
    92.8 +lemma all_accessible: "wf r ==> \<forall>M. M \<in> Wellfounded.acc (mult1 r)"
    92.9  proof
   92.10    let ?R = "mult1 r"
   92.11 -  let ?W = "acc ?R"
   92.12 +  let ?W = "Wellfounded.acc ?R"
   92.13    {
   92.14      fix M M0 a
   92.15      assume M0: "M0 \<in> ?W"
    93.1 --- a/src/HOL/Library/Polynomial.thy	Mon Nov 11 17:34:44 2013 +0100
    93.2 +++ b/src/HOL/Library/Polynomial.thy	Mon Nov 11 17:44:21 2013 +0100
    93.3 @@ -667,7 +667,7 @@
    93.4    show "- p + p = 0"
    93.5      by (simp add: poly_eq_iff)
    93.6    show "p - q = p + - q"
    93.7 -    by (simp add: poly_eq_iff diff_minus)
    93.8 +    by (simp add: poly_eq_iff)
    93.9  qed
   93.10  
   93.11  end
   93.12 @@ -714,15 +714,15 @@
   93.13  
   93.14  lemma degree_diff_le_max: "degree (p - q) \<le> max (degree p) (degree q)"
   93.15    using degree_add_le [where p=p and q="-q"]
   93.16 -  by (simp add: diff_minus)
   93.17 +  by simp
   93.18  
   93.19  lemma degree_diff_le:
   93.20    "\<lbrakk>degree p \<le> n; degree q \<le> n\<rbrakk> \<Longrightarrow> degree (p - q) \<le> n"
   93.21 -  by (simp add: diff_minus degree_add_le)
   93.22 +  using degree_add_le [of p n "- q"] by simp
   93.23  
   93.24  lemma degree_diff_less:
   93.25    "\<lbrakk>degree p < n; degree q < n\<rbrakk> \<Longrightarrow> degree (p - q) < n"
   93.26 -  by (simp add: diff_minus degree_add_less)
   93.27 +  using degree_add_less [of p n "- q"] by simp
   93.28  
   93.29  lemma add_monom: "monom a n + monom b n = monom (a + b) n"
   93.30    by (rule poly_eqI) simp
   93.31 @@ -793,7 +793,7 @@
   93.32  lemma poly_diff [simp]:
   93.33    fixes x :: "'a::comm_ring"
   93.34    shows "poly (p - q) x = poly p x - poly q x"
   93.35 -  by (simp add: diff_minus)
   93.36 +  using poly_add [of p "- q" x] by simp
   93.37  
   93.38  lemma poly_setsum: "poly (\<Sum>k\<in>A. p k) x = (\<Sum>k\<in>A. poly (p k) x)"
   93.39    by (induct A rule: infinite_finite_induct) simp_all
    94.1 --- a/src/HOL/Library/Product_plus.thy	Mon Nov 11 17:34:44 2013 +0100
    94.2 +++ b/src/HOL/Library/Product_plus.thy	Mon Nov 11 17:44:21 2013 +0100
    94.3 @@ -104,7 +104,7 @@
    94.4    (cancel_comm_monoid_add, cancel_comm_monoid_add) cancel_comm_monoid_add ..
    94.5  
    94.6  instance prod :: (group_add, group_add) group_add
    94.7 -  by default (simp_all add: prod_eq_iff diff_minus)
    94.8 +  by default (simp_all add: prod_eq_iff)
    94.9  
   94.10  instance prod :: (ab_group_add, ab_group_add) ab_group_add
   94.11    by default (simp_all add: prod_eq_iff)
    95.1 --- a/src/HOL/Library/RBT_Set.thy	Mon Nov 11 17:34:44 2013 +0100
    95.2 +++ b/src/HOL/Library/RBT_Set.thy	Mon Nov 11 17:44:21 2013 +0100
    95.3 @@ -756,7 +756,8 @@
    95.4  declare Inf_Set_fold[folded Inf'_def, code]
    95.5  
    95.6  lemma INFI_Set_fold [code]:
    95.7 -  "INFI (Set t) f = fold_keys (inf \<circ> f) t top"
    95.8 +  fixes f :: "_ \<Rightarrow> 'a::complete_lattice"
    95.9 +  shows "INFI (Set t) f = fold_keys (inf \<circ> f) t top"
   95.10  proof -
   95.11    have "comp_fun_commute ((inf :: 'a \<Rightarrow> 'a \<Rightarrow> 'a) \<circ> f)" 
   95.12      by default (auto simp add: fun_eq_iff ac_simps)
   95.13 @@ -796,7 +797,8 @@
   95.14  declare Sup_Set_fold[folded Sup'_def, code]
   95.15  
   95.16  lemma SUPR_Set_fold [code]:
   95.17 -  "SUPR (Set t) f = fold_keys (sup \<circ> f) t bot"
   95.18 +  fixes f :: "_ \<Rightarrow> 'a::complete_lattice"
   95.19 +  shows "SUPR (Set t) f = fold_keys (sup \<circ> f) t bot"
   95.20  proof -
   95.21    have "comp_fun_commute ((sup :: 'a \<Rightarrow> 'a \<Rightarrow> 'a) \<circ> f)" 
   95.22      by default (auto simp add: fun_eq_iff ac_simps)
    96.1 --- a/src/HOL/Library/Set_Algebras.thy	Mon Nov 11 17:34:44 2013 +0100
    96.2 +++ b/src/HOL/Library/Set_Algebras.thy	Mon Nov 11 17:44:21 2013 +0100
    96.3 @@ -190,12 +190,12 @@
    96.4    done
    96.5  
    96.6  lemma set_plus_imp_minus: "(a::'a::ab_group_add) : b +o C ==> (a - b) : C"
    96.7 -  by (auto simp add: elt_set_plus_def add_ac diff_minus)
    96.8 +  by (auto simp add: elt_set_plus_def add_ac)
    96.9  
   96.10  lemma set_minus_imp_plus: "(a::'a::ab_group_add) - b : C ==> a : b +o C"
   96.11 -  apply (auto simp add: elt_set_plus_def add_ac diff_minus)
   96.12 +  apply (auto simp add: elt_set_plus_def add_ac)
   96.13    apply (subgoal_tac "a = (a + - b) + b")
   96.14 -   apply (rule bexI, assumption, assumption)
   96.15 +   apply (rule bexI, assumption)
   96.16    apply (auto simp add: add_ac)
   96.17    done
   96.18  
    97.1 --- a/src/HOL/Library/Univ_Poly.thy	Mon Nov 11 17:34:44 2013 +0100
    97.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
    97.3 @@ -1,1053 +0,0 @@
    97.4 -(*  Title:      HOL/Library/Univ_Poly.thy
    97.5 -    Author:     Amine Chaieb
    97.6 -*)
    97.7 -
    97.8 -header {* Univariate Polynomials *}
    97.9 -
   97.10 -theory Univ_Poly
   97.11 -imports Main
   97.12 -begin
   97.13 -
   97.14 -text{* Application of polynomial as a function. *}
   97.15 -
   97.16 -primrec (in semiring_0) poly :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a"
   97.17 -where
   97.18 -  poly_Nil:  "poly [] x = 0"
   97.19 -| poly_Cons: "poly (h#t) x = h + x * poly t x"
   97.20 -
   97.21 -
   97.22 -subsection{*Arithmetic Operations on Polynomials*}
   97.23 -
   97.24 -text{*addition*}
   97.25 -
   97.26 -primrec (in semiring_0) padd :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "+++" 65)
   97.27 -where
   97.28 -  padd_Nil:  "[] +++ l2 = l2"
   97.29 -| padd_Cons: "(h#t) +++ l2 = (if l2 = [] then h#t else (h + hd l2)#(t +++ tl l2))"
   97.30 -
   97.31 -text{*Multiplication by a constant*}
   97.32 -primrec (in semiring_0) cmult :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "%*" 70) where
   97.33 -  cmult_Nil:  "c %* [] = []"
   97.34 -| cmult_Cons: "c %* (h#t) = (c * h)#(c %* t)"
   97.35 -
   97.36 -text{*Multiplication by a polynomial*}
   97.37 -primrec (in semiring_0) pmult :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "***" 70)
   97.38 -where
   97.39 -  pmult_Nil:  "[] *** l2 = []"
   97.40 -| pmult_Cons: "(h#t) *** l2 = (if t = [] then h %* l2
   97.41 -                              else (h %* l2) +++ ((0) # (t *** l2)))"
   97.42 -
   97.43 -text{*Repeated multiplication by a polynomial*}
   97.44 -primrec (in semiring_0) mulexp :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a  list \<Rightarrow> 'a list" where
   97.45 -  mulexp_zero:  "mulexp 0 p q = q"
   97.46 -| mulexp_Suc:   "mulexp (Suc n) p q = p *** mulexp n p q"
   97.47 -
   97.48 -text{*Exponential*}
   97.49 -primrec (in semiring_1) pexp :: "'a list \<Rightarrow> nat \<Rightarrow> 'a list"  (infixl "%^" 80) where
   97.50 -  pexp_0:   "p %^ 0 = [1]"
   97.51 -| pexp_Suc: "p %^ (Suc n) = p *** (p %^ n)"
   97.52 -
   97.53 -text{*Quotient related value of dividing a polynomial by x + a*}
   97.54 -(* Useful for divisor properties in inductive proofs *)
   97.55 -primrec (in field) "pquot" :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list"
   97.56 -where
   97.57 -  pquot_Nil:  "pquot [] a= []"
   97.58 -| pquot_Cons: "pquot (h#t) a =
   97.59 -    (if t = [] then [h] else (inverse(a) * (h - hd( pquot t a)))#(pquot t a))"
   97.60 -
   97.61 -text{*normalization of polynomials (remove extra 0 coeff)*}
   97.62 -primrec (in semiring_0) pnormalize :: "'a list \<Rightarrow> 'a list" where
   97.63 -  pnormalize_Nil:  "pnormalize [] = []"
   97.64 -| pnormalize_Cons: "pnormalize (h#p) =
   97.65 -    (if pnormalize p = [] then (if h = 0 then [] else [h]) else h # pnormalize p)"
   97.66 -
   97.67 -definition (in semiring_0) "pnormal p = ((pnormalize p = p) \<and> p \<noteq> [])"
   97.68 -definition (in semiring_0) "nonconstant p = (pnormal p \<and> (\<forall>x. p \<noteq> [x]))"
   97.69 -text{*Other definitions*}
   97.70 -
   97.71 -definition (in ring_1) poly_minus :: "'a list \<Rightarrow> 'a list" ("-- _" [80] 80)
   97.72 -  where "-- p = (- 1) %* p"
   97.73 -
   97.74 -definition (in semiring_0) divides :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"  (infixl "divides" 70)
   97.75 -  where "p1 divides p2 = (\<exists>q. poly p2 = poly(p1 *** q))"
   97.76 -
   97.77 -lemma (in semiring_0) dividesI:
   97.78 -  "poly p2 = poly (p1 *** q) \<Longrightarrow> p1 divides p2"
   97.79 -  by (auto simp add: divides_def)
   97.80 -
   97.81 -lemma (in semiring_0) dividesE:
   97.82 -  assumes "p1 divides p2"
   97.83 -  obtains q where "poly p2 = poly (p1 *** q)"
   97.84 -  using assms by (auto simp add: divides_def)
   97.85 -
   97.86 -    --{*order of a polynomial*}
   97.87 -definition (in ring_1) order :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" where
   97.88 -  "order a p = (SOME n. ([-a, 1] %^ n) divides p \<and> ~ (([-a, 1] %^ (Suc n)) divides p))"
   97.89 -
   97.90 -     --{*degree of a polynomial*}
   97.91 -definition (in semiring_0) degree :: "'a list \<Rightarrow> nat"
   97.92 -  where "degree p = length (pnormalize p) - 1"
   97.93 -
   97.94 -     --{*squarefree polynomials --- NB with respect to real roots only.*}
   97.95 -definition (in ring_1) rsquarefree :: "'a list \<Rightarrow> bool"
   97.96 -  where "rsquarefree p \<longleftrightarrow> poly p \<noteq> poly [] \<and> (\<forall>a. order a p = 0 \<or> order a p = 1)"
   97.97 -
   97.98 -context semiring_0
   97.99 -begin
  97.100 -
  97.101 -lemma padd_Nil2[simp]: "p +++ [] = p"
  97.102 -  by (induct p) auto
  97.103 -
  97.104 -lemma padd_Cons_Cons: "(h1 # p1) +++ (h2 # p2) = (h1 + h2) # (p1 +++ p2)"
  97.105 -  by auto
  97.106 -
  97.107 -lemma pminus_Nil: "-- [] = []"
  97.108 -  by (simp add: poly_minus_def)
  97.109 -
  97.110 -lemma pmult_singleton: "[h1] *** p1 = h1 %* p1" by simp
  97.111 -
  97.112 -end
  97.113 -
  97.114 -lemma (in semiring_1) poly_ident_mult[simp]: "1 %* t = t" by (induct t) auto
  97.115 -
  97.116 -lemma (in semiring_0) poly_simple_add_Cons[simp]: "[a] +++ ((0)#t) = (a#t)"
  97.117 -  by simp
  97.118 -
  97.119 -text{*Handy general properties*}
  97.120 -
  97.121 -lemma (in comm_semiring_0) padd_commut: "b +++ a = a +++ b"
  97.122 -proof (induct b arbitrary: a)
  97.123 -  case Nil
  97.124 -  thus ?case by auto
  97.125 -next
  97.126 -  case (Cons b bs a)
  97.127 -  thus ?case by (cases a) (simp_all add: add_commute)
  97.128 -qed
  97.129 -
  97.130 -lemma (in comm_semiring_0) padd_assoc: "\<forall>b c. (a +++ b) +++ c = a +++ (b +++ c)"
  97.131 -  apply (induct a)
  97.132 -  apply (simp, clarify)
  97.133 -  apply (case_tac b, simp_all add: add_ac)
  97.134 -  done
  97.135 -
  97.136 -lemma (in semiring_0) poly_cmult_distr: "a %* ( p +++ q) = (a %* p +++ a %* q)"
  97.137 -  apply (induct p arbitrary: q)
  97.138 -  apply simp
  97.139 -  apply (case_tac q, simp_all add: distrib_left)
  97.140 -  done
  97.141 -
  97.142 -lemma (in ring_1) pmult_by_x[simp]: "[0, 1] *** t = ((0)#t)"
  97.143 -  apply (induct t)
  97.144 -  apply simp
  97.145 -  apply (auto simp add: padd_commut)
  97.146 -  apply (case_tac t, auto)
  97.147 -  done
  97.148 -
  97.149 -text{*properties of evaluation of polynomials.*}
  97.150 -
  97.151 -lemma (in semiring_0) poly_add: "poly (p1 +++ p2) x = poly p1 x + poly p2 x"
  97.152 -proof(induct p1 arbitrary: p2)
  97.153 -  case Nil
  97.154 -  thus ?case by simp
  97.155 -next
  97.156 -  case (Cons a as p2)
  97.157 -  thus ?case
  97.158 -    by (cases p2) (simp_all  add: add_ac distrib_left)
  97.159 -qed
  97.160 -
  97.161 -lemma (in comm_semiring_0) poly_cmult: "poly (c %* p) x = c * poly p x"
  97.162 -  apply (induct p)
  97.163 -  apply (case_tac [2] "x = zero")
  97.164 -  apply (auto simp add: distrib_left mult_ac)
  97.165 -  done
  97.166 -
  97.167 -lemma (in comm_semiring_0) poly_cmult_map: "poly (map (op * c) p) x = c*poly p x"
  97.168 -  by (induct p) (auto simp add: distrib_left mult_ac)
  97.169 -
  97.170 -lemma (in comm_ring_1) poly_minus: "poly (-- p) x = - (poly p x)"
  97.171 -  apply (simp add: poly_minus_def)
  97.172 -  apply (auto simp add: poly_cmult)
  97.173 -  done
  97.174 -
  97.175 -lemma (in comm_semiring_0) poly_mult: "poly (p1 *** p2) x = poly p1 x * poly p2 x"
  97.176 -proof (induct p1 arbitrary: p2)
  97.177 -  case Nil
  97.178 -  thus ?case by simp
  97.179 -next
  97.180 -  case (Cons a as p2)
  97.181 -  thus ?case by (cases as)
  97.182 -    (simp_all add: poly_cmult poly_add distrib_right distrib_left mult_ac)
  97.183 -qed
  97.184 -
  97.185 -class idom_char_0 = idom + ring_char_0
  97.186 -
  97.187 -lemma (in comm_ring_1) poly_exp: "poly (p %^ n) x = (poly p x) ^ n"
  97.188 -  by (induct n) (auto simp add: poly_cmult poly_mult)
  97.189 -
  97.190 -text{*More Polynomial Evaluation Lemmas*}
  97.191 -
  97.192 -lemma (in semiring_0) poly_add_rzero[simp]: "poly (a +++ []) x = poly a x"
  97.193 -  by simp
  97.194 -
  97.195 -lemma (in comm_semiring_0) poly_mult_assoc: "poly ((a *** b) *** c) x = poly (a *** (b *** c)) x"
  97.196 -  by (simp add: poly_mult mult_assoc)
  97.197 -
  97.198 -lemma (in semiring_0) poly_mult_Nil2[simp]: "poly (p *** []) x = 0"
  97.199 -  by (induct p) auto
  97.200 -
  97.201 -lemma (in comm_semiring_1) poly_exp_add: "poly (p %^ (n + d)) x = poly( p %^ n *** p %^ d) x"
  97.202 -  by (induct n) (auto simp add: poly_mult mult_assoc)
  97.203 -
  97.204 -subsection{*Key Property: if @{term "f(a) = 0"} then @{term "(x - a)"} divides
  97.205 - @{term "p(x)"} *}
  97.206 -
  97.207 -lemma (in comm_ring_1) lemma_poly_linear_rem: "\<forall>h. \<exists>q r. h#t = [r] +++ [-a, 1] *** q"
  97.208 -proof(induct t)
  97.209 -  case Nil
  97.210 -  { fix h have "[h] = [h] +++ [- a, 1] *** []" by simp }
  97.211 -  thus ?case by blast
  97.212 -next
  97.213 -  case (Cons  x xs)
  97.214 -  { fix h
  97.215 -    from Cons.hyps[rule_format, of x]
  97.216 -    obtain q r where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast
  97.217 -    have "h#x#xs = [a*r + h] +++ [-a, 1] *** (r#q)"
  97.218 -      using qr by (cases q) (simp_all add: algebra_simps)
  97.219 -    hence "\<exists>q r. h#x#xs = [r] +++ [-a, 1] *** q" by blast}
  97.220 -  thus ?case by blast
  97.221 -qed
  97.222 -
  97.223 -lemma (in comm_ring_1) poly_linear_rem: "\<exists>q r. h#t = [r] +++ [-a, 1] *** q"
  97.224 -  using lemma_poly_linear_rem [where t = t and a = a] by auto
  97.225 -
  97.226 -
  97.227 -lemma (in comm_ring_1) poly_linear_divides: "(poly p a = 0) = ((p = []) | (\<exists>q. p = [-a, 1] *** q))"
  97.228 -proof -
  97.229 -  { assume p: "p = []" hence ?thesis by simp }
  97.230 -  moreover
  97.231 -  {
  97.232 -    fix x xs assume p: "p = x#xs"
  97.233 -    {
  97.234 -      fix q assume "p = [-a, 1] *** q"
  97.235 -      hence "poly p a = 0" by (simp add: poly_add poly_cmult)
  97.236 -    }
  97.237 -    moreover
  97.238 -    { assume p0: "poly p a = 0"
  97.239 -      from poly_linear_rem[of x xs a] obtain q r
  97.240 -      where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast
  97.241 -      have "r = 0" using p0 by (simp only: p qr poly_mult poly_add) simp
  97.242 -      hence "\<exists>q. p = [- a, 1] *** q"
  97.243 -        using p qr
  97.244 -        apply -
  97.245 -        apply (rule exI[where x=q])
  97.246 -        apply auto
  97.247 -        apply (cases q)
  97.248 -        apply auto
  97.249 -        done
  97.250 -    }
  97.251 -    ultimately have ?thesis using p by blast
  97.252 -  }
  97.253 -  ultimately show ?thesis by (cases p) auto
  97.254 -qed
  97.255 -
  97.256 -lemma (in semiring_0) lemma_poly_length_mult[simp]: "\<forall>h k a. length (k %* p +++  (h # (a %* p))) = Suc (length p)"
  97.257 -  by (induct p) auto
  97.258 -
  97.259 -lemma (in semiring_0) lemma_poly_length_mult2[simp]: "\<forall>h k. length (k %* p +++  (h # p)) = Suc (length p)"
  97.260 -  by (induct p) auto
  97.261 -
  97.262 -lemma (in ring_1) poly_length_mult[simp]: "length([-a,1] *** q) = Suc (length q)"
  97.263 -  by auto
  97.264 -
  97.265 -subsection{*Polynomial length*}
  97.266 -
  97.267 -lemma (in semiring_0) poly_cmult_length[simp]: "length (a %* p) = length p"
  97.268 -  by (induct p) auto
  97.269 -
  97.270 -lemma (in semiring_0) poly_add_length: "length (p1 +++ p2) = max (length p1) (length p2)"
  97.271 -  by (induct p1 arbitrary: p2) (simp_all, arith)
  97.272 -
  97.273 -lemma (in semiring_0) poly_root_mult_length[simp]: "length([a,b] *** p) = Suc (length p)"
  97.274 -  by (simp add: poly_add_length)
  97.275 -
  97.276 -lemma (in idom) poly_mult_not_eq_poly_Nil[simp]:
  97.277 -  "poly (p *** q) x \<noteq> poly [] x \<longleftrightarrow> poly p x \<noteq> poly [] x \<and> poly q x \<noteq> poly [] x"
  97.278 -  by (auto simp add: poly_mult)
  97.279 -
  97.280 -lemma (in idom) poly_mult_eq_zero_disj: "poly (p *** q) x = 0 \<longleftrightarrow> poly p x = 0 \<or> poly q x = 0"
  97.281 -  by (auto simp add: poly_mult)
  97.282 -
  97.283 -text{*Normalisation Properties*}
  97.284 -
  97.285 -lemma (in semiring_0) poly_normalized_nil: "(pnormalize p = []) --> (poly p x = 0)"
  97.286 -  by (induct p) auto
  97.287 -
  97.288 -text{*A nontrivial polynomial of degree n has no more than n roots*}
  97.289 -lemma (in idom) poly_roots_index_lemma:
  97.290 -   assumes p: "poly p x \<noteq> poly [] x" and n: "length p = n"
  97.291 -  shows "\<exists>i. \<forall>x. poly p x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)"
  97.292 -  using p n
  97.293 -proof (induct n arbitrary: p x)
  97.294 -  case 0
  97.295 -  thus ?case by simp
  97.296 -next
  97.297 -  case (Suc n p x)
  97.298 -  {
  97.299 -    assume C: "\<And>i. \<exists>x. poly p x = 0 \<and> (\<forall>m\<le>Suc n. x \<noteq> i m)"
  97.300 -    from Suc.prems have p0: "poly p x \<noteq> 0" "p\<noteq> []" by auto
  97.301 -    from p0(1)[unfolded poly_linear_divides[of p x]]
  97.302 -    have "\<forall>q. p \<noteq> [- x, 1] *** q" by blast
  97.303 -    from C obtain a where a: "poly p a = 0" by blast
  97.304 -    from a[unfolded poly_linear_divides[of p a]] p0(2)
  97.305 -    obtain q where q: "p = [-a, 1] *** q" by blast
  97.306 -    have lg: "length q = n" using q Suc.prems(2) by simp
  97.307 -    from q p0 have qx: "poly q x \<noteq> poly [] x"
  97.308 -      by (auto simp add: poly_mult poly_add poly_cmult)
  97.309 -    from Suc.hyps[OF qx lg] obtain i where
  97.310 -      i: "\<forall>x. poly q x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)" by blast
  97.311 -    let ?i = "\<lambda>m. if m = Suc n then a else i m"
  97.312 -    from C[of ?i] obtain y where y: "poly p y = 0" "\<forall>m\<le> Suc n. y \<noteq> ?i m"
  97.313 -      by blast
  97.314 -    from y have "y = a \<or> poly q y = 0"
  97.315 -      by (simp only: q poly_mult_eq_zero_disj poly_add) (simp add: algebra_simps)
  97.316 -    with i[rule_format, of y] y(1) y(2) have False
  97.317 -      apply auto
  97.318 -      apply (erule_tac x = "m" in allE)
  97.319 -      apply auto
  97.320 -      done
  97.321 -  }
  97.322 -  thus ?case by blast
  97.323 -qed
  97.324 -
  97.325 -
  97.326 -lemma (in idom) poly_roots_index_length:
  97.327 -  "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>i. \<forall>x. (poly p x = 0) \<longrightarrow> (\<exists>n. n \<le> length p \<and> x = i n)"
  97.328 -  by (blast intro: poly_roots_index_lemma)
  97.329 -
  97.330 -lemma (in idom) poly_roots_finite_lemma1:
  97.331 -  "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>N i. \<forall>x. (poly p x = 0) \<longrightarrow> (\<exists>n. (n::nat) < N \<and> x = i n)"
  97.332 -  apply (drule poly_roots_index_length, safe)
  97.333 -  apply (rule_tac x = "Suc (length p)" in exI)
  97.334 -  apply (rule_tac x = i in exI)
  97.335 -  apply (simp add: less_Suc_eq_le)
  97.336 -  done
  97.337 -
  97.338 -lemma (in idom) idom_finite_lemma:
  97.339 -  assumes P: "\<forall>x. P x --> (\<exists>n. n < length j \<and> x = j!n)"
  97.340 -  shows "finite {x. P x}"
  97.341 -proof -
  97.342 -  let ?M = "{x. P x}"
  97.343 -  let ?N = "set j"
  97.344 -  have "?M \<subseteq> ?N" using P by auto
  97.345 -  thus ?thesis using finite_subset by auto
  97.346 -qed
  97.347 -
  97.348 -lemma (in idom) poly_roots_finite_lemma2:
  97.349 -  "poly p x \<noteq> poly [] x \<Longrightarrow> \<exists>i. \<forall>x. poly p x = 0 \<longrightarrow> x \<in> set i"
  97.350 -  apply (drule poly_roots_index_length, safe)
  97.351 -  apply (rule_tac x="map (\<lambda>n. i n) [0 ..< Suc (length p)]" in exI)
  97.352 -  apply (auto simp add: image_iff)
  97.353 -  apply (erule_tac x="x" in allE, clarsimp)
  97.354 -  apply (case_tac "n = length p")
  97.355 -  apply (auto simp add: order_le_less)
  97.356 -  done
  97.357 -
  97.358 -lemma (in ring_char_0) UNIV_ring_char_0_infinte: "\<not> (finite (UNIV:: 'a set))"
  97.359 -proof
  97.360 -  assume F: "finite (UNIV :: 'a set)"
  97.361 -  have "finite (UNIV :: nat set)"
  97.362 -  proof (rule finite_imageD)
  97.363 -    have "of_nat ` UNIV \<subseteq> UNIV" by simp
  97.364 -    then show "finite (of_nat ` UNIV :: 'a set)" using F by (rule finite_subset)
  97.365 -    show "inj (of_nat :: nat \<Rightarrow> 'a)" by (simp add: inj_on_def)
  97.366 -  qed
  97.367 -  with infinite_UNIV_nat show False ..
  97.368 -qed
  97.369 -
  97.370 -lemma (in idom_char_0) poly_roots_finite: "poly p \<noteq> poly [] \<longleftrightarrow> finite {x. poly p x = 0}"
  97.371 -proof
  97.372 -  assume H: "poly p \<noteq> poly []"
  97.373 -  show "finite {x. poly p x = (0::'a)}"
  97.374 -    using H
  97.375 -    apply -
  97.376 -    apply (erule contrapos_np, rule ext)
  97.377 -    apply (rule ccontr)
  97.378 -    apply (clarify dest!: poly_roots_finite_lemma2)
  97.379 -    using finite_subset
  97.380 -  proof -
  97.381 -    fix x i
  97.382 -    assume F: "\<not> finite {x. poly p x = (0\<Colon>'a)}"
  97.383 -      and P: "\<forall>x. poly p x = (0\<Colon>'a) \<longrightarrow> x \<in> set i"
  97.384 -    let ?M= "{x. poly p x = (0\<Colon>'a)}"
  97.385 -    from P have "?M \<subseteq> set i" by auto
  97.386 -    with finite_subset F show False by auto
  97.387 -  qed
  97.388 -next
  97.389 -  assume F: "finite {x. poly p x = (0\<Colon>'a)}"
  97.390 -  show "poly p \<noteq> poly []" using F UNIV_ring_char_0_infinte by auto
  97.391 -qed
  97.392 -
  97.393 -text{*Entirety and Cancellation for polynomials*}
  97.394 -
  97.395 -lemma (in idom_char_0) poly_entire_lemma2:
  97.396 -  assumes p0: "poly p \<noteq> poly []"
  97.397 -    and q0: "poly q \<noteq> poly []"
  97.398 -  shows "poly (p***q) \<noteq> poly []"
  97.399 -proof -
  97.400 -  let ?S = "\<lambda>p. {x. poly p x = 0}"
  97.401 -  have "?S (p *** q) = ?S p \<union> ?S q" by (auto simp add: poly_mult)
  97.402 -  with p0 q0 show ?thesis  unfolding poly_roots_finite by auto
  97.403 -qed
  97.404 -
  97.405 -lemma (in idom_char_0) poly_entire:
  97.406 -  "poly (p *** q) = poly [] \<longleftrightarrow> poly p = poly [] \<or> poly q = poly []"
  97.407 -  using poly_entire_lemma2[of p q]
  97.408 -  by (auto simp add: fun_eq_iff poly_mult)
  97.409 -
  97.410 -lemma (in idom_char_0) poly_entire_neg:
  97.411 -  "poly (p *** q) \<noteq> poly [] \<longleftrightarrow> poly p \<noteq> poly [] \<and> poly q \<noteq> poly []"
  97.412 -  by (simp add: poly_entire)
  97.413 -
  97.414 -lemma fun_eq: "f = g \<longleftrightarrow> (\<forall>x. f x = g x)"
  97.415 -  by auto
  97.416 -
  97.417 -lemma (in comm_ring_1) poly_add_minus_zero_iff:
  97.418 -  "poly (p +++ -- q) = poly [] \<longleftrightarrow> poly p = poly q"
  97.419 -  by (auto simp add: algebra_simps poly_add poly_minus_def fun_eq poly_cmult)
  97.420 -
  97.421 -lemma (in comm_ring_1) poly_add_minus_mult_eq:
  97.422 -  "poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))"
  97.423 -  by (auto simp add: poly_add poly_minus_def fun_eq poly_mult poly_cmult distrib_left)
  97.424 -
  97.425 -subclass (in idom_char_0) comm_ring_1 ..
  97.426 -
  97.427 -lemma (in idom_char_0) poly_mult_left_cancel:
  97.428 -  "poly (p *** q) = poly (p *** r) \<longleftrightarrow> poly p = poly [] \<or> poly q = poly r"
  97.429 -proof -
  97.430 -  have "poly (p *** q) = poly (p *** r) \<longleftrightarrow> poly (p *** q +++ -- (p *** r)) = poly []"
  97.431 -    by (simp only: poly_add_minus_zero_iff)
  97.432 -  also have "\<dots> \<longleftrightarrow> poly p = poly [] \<or> poly q = poly r"
  97.433 -    by (auto intro: simp add: poly_add_minus_mult_eq poly_entire poly_add_minus_zero_iff)
  97.434 -  finally show ?thesis .
  97.435 -qed
  97.436 -
  97.437 -lemma (in idom) poly_exp_eq_zero[simp]:
  97.438 -  "poly (p %^ n) = poly [] \<longleftrightarrow> poly p = poly [] \<and> n \<noteq> 0"
  97.439 -  apply (simp only: fun_eq add: HOL.all_simps [symmetric])
  97.440 -  apply (rule arg_cong [where f = All])
  97.441 -  apply (rule ext)
  97.442 -  apply (induct n)
  97.443 -  apply (auto simp add: poly_exp poly_mult)
  97.444 -  done
  97.445 -
  97.446 -lemma (in comm_ring_1) poly_prime_eq_zero[simp]: "poly [a,1] \<noteq> poly []"
  97.447 -  apply (simp add: fun_eq)
  97.448 -  apply (rule_tac x = "minus one a" in exI)
  97.449 -  apply (unfold diff_minus)
  97.450 -  apply (subst add_commute)
  97.451 -  apply (subst add_assoc)
  97.452 -  apply simp
  97.453 -  done
  97.454 -
  97.455 -lemma (in idom) poly_exp_prime_eq_zero: "poly ([a, 1] %^ n) \<noteq> poly []"
  97.456 -  by auto
  97.457 -
  97.458 -text{*A more constructive notion of polynomials being trivial*}
  97.459 -
  97.460 -lemma (in idom_char_0) poly_zero_lemma': "poly (h # t) = poly [] \<Longrightarrow> h = 0 \<and> poly t = poly []"
  97.461 -  apply (simp add: fun_eq)
  97.462 -  apply (case_tac "h = zero")
  97.463 -  apply (drule_tac [2] x = zero in spec, auto)
  97.464 -  apply (cases "poly t = poly []", simp)
  97.465 -proof -
  97.466 -  fix x
  97.467 -  assume H: "\<forall>x. x = (0\<Colon>'a) \<or> poly t x = (0\<Colon>'a)"
  97.468 -    and pnz: "poly t \<noteq> poly []"
  97.469 -  let ?S = "{x. poly t x = 0}"
  97.470 -  from H have "\<forall>x. x \<noteq>0 \<longrightarrow> poly t x = 0" by blast
  97.471 -  hence th: "?S \<supseteq> UNIV - {0}" by auto
  97.472 -  from poly_roots_finite pnz have th': "finite ?S" by blast
  97.473 -  from finite_subset[OF th th'] UNIV_ring_char_0_infinte show "poly t x = (0\<Colon>'a)"
  97.474 -    by simp
  97.475 -qed
  97.476 -
  97.477 -lemma (in idom_char_0) poly_zero: "(poly p = poly []) = list_all (%c. c = 0) p"
  97.478 -  apply (induct p)
  97.479 -  apply simp
  97.480 -  apply (rule iffI)
  97.481 -  apply (drule poly_zero_lemma', auto)
  97.482 -  done
  97.483 -
  97.484 -lemma (in idom_char_0) poly_0: "list_all (\<lambda>c. c = 0) p \<Longrightarrow> poly p x = 0"
  97.485 -  unfolding poly_zero[symmetric] by simp
  97.486 -
  97.487 -
  97.488 -
  97.489 -text{*Basics of divisibility.*}
  97.490 -
  97.491 -lemma (in idom) poly_primes:
  97.492 -  "[a, 1] divides (p *** q) \<longleftrightarrow> [a, 1] divides p \<or> [a, 1] divides q"
  97.493 -  apply (auto simp add: divides_def fun_eq poly_mult poly_add poly_cmult distrib_right [symmetric])
  97.494 -  apply (drule_tac x = "uminus a" in spec)
  97.495 -  apply (simp add: poly_linear_divides poly_add poly_cmult distrib_right [symmetric])
  97.496 -  apply (cases "p = []")
  97.497 -  apply (rule exI[where x="[]"])
  97.498 -  apply simp
  97.499 -  apply (cases "q = []")
  97.500 -  apply (erule allE[where x="[]"], simp)
  97.501 -
  97.502 -  apply clarsimp
  97.503 -  apply (cases "\<exists>q\<Colon>'a list. p = a %* q +++ ((0\<Colon>'a) # q)")
  97.504 -  apply (clarsimp simp add: poly_add poly_cmult)
  97.505 -  apply (rule_tac x="qa" in exI)
  97.506 -  apply (simp add: distrib_right [symmetric])
  97.507 -  apply clarsimp
  97.508 -
  97.509 -  apply (auto simp add: poly_linear_divides poly_add poly_cmult distrib_right [symmetric])
  97.510 -  apply (rule_tac x = "pmult qa q" in exI)
  97.511 -  apply (rule_tac [2] x = "pmult p qa" in exI)
  97.512 -  apply (auto simp add: poly_add poly_mult poly_cmult mult_ac)
  97.513 -  done
  97.514 -
  97.515 -lemma (in comm_semiring_1) poly_divides_refl[simp]: "p divides p"
  97.516 -  apply (simp add: divides_def)
  97.517 -  apply (rule_tac x = "[one]" in exI)
  97.518 -  apply (auto simp add: poly_mult fun_eq)
  97.519 -  done
  97.520 -
  97.521 -lemma (in comm_semiring_1) poly_divides_trans: "p divides q \<Longrightarrow> q divides r \<Longrightarrow> p divides r"
  97.522 -  apply (simp add: divides_def, safe)
  97.523 -  apply (rule_tac x = "pmult qa qaa" in exI)
  97.524 -  apply (auto simp add: poly_mult fun_eq mult_assoc)
  97.525 -  done
  97.526 -
  97.527 -lemma (in comm_semiring_1) poly_divides_exp: "m \<le> n \<Longrightarrow> (p %^ m) divides (p %^ n)"
  97.528 -  apply (auto simp add: le_iff_add)
  97.529 -  apply (induct_tac k)
  97.530 -  apply (rule_tac [2] poly_divides_trans)
  97.531 -  apply (auto simp add: divides_def)
  97.532 -  apply (rule_tac x = p in exI)
  97.533 -  apply (auto simp add: poly_mult fun_eq mult_ac)
  97.534 -  done
  97.535 -
  97.536 -lemma (in comm_semiring_1) poly_exp_divides:
  97.537 -  "(p %^ n) divides q \<Longrightarrow> m \<le> n \<Longrightarrow> (p %^ m) divides q"
  97.538 -  by (blast intro: poly_divides_exp poly_divides_trans)
  97.539 -
  97.540 -lemma (in comm_semiring_0) poly_divides_add:
  97.541 -  "p divides q \<Longrightarrow> p divides r \<Longrightarrow> p divides (q +++ r)"
  97.542 -  apply (simp add: divides_def, auto)
  97.543 -  apply (rule_tac x = "padd qa qaa" in exI)
  97.544 -  apply (auto simp add: poly_add fun_eq poly_mult distrib_left)
  97.545 -  done
  97.546 -
  97.547 -lemma (in comm_ring_1) poly_divides_diff:
  97.548 -  "p divides q \<Longrightarrow> p divides (q +++ r) \<Longrightarrow> p divides r"
  97.549 -  apply (simp add: divides_def, auto)
  97.550 -  apply (rule_tac x = "padd qaa (poly_minus qa)" in exI)
  97.551 -  apply (auto simp add: poly_add fun_eq poly_mult poly_minus algebra_simps)
  97.552 -  done
  97.553 -
  97.554 -lemma (in comm_ring_1) poly_divides_diff2:
  97.555 -  "p divides r \<Longrightarrow> p divides (q +++ r) \<Longrightarrow> p divides q"
  97.556 -  apply (erule poly_divides_diff)
  97.557 -  apply (auto simp add: poly_add fun_eq poly_mult divides_def add_ac)
  97.558 -  done
  97.559 -
  97.560 -lemma (in semiring_0) poly_divides_zero: "poly p = poly [] \<Longrightarrow> q divides p"
  97.561 -  apply (simp add: divides_def)
  97.562 -  apply (rule exI[where x="[]"])
  97.563 -  apply (auto simp add: fun_eq poly_mult)
  97.564 -  done
  97.565 -
  97.566 -lemma (in semiring_0) poly_divides_zero2 [simp]: "q divides []"
  97.567 -  apply (simp add: divides_def)
  97.568 -  apply (rule_tac x = "[]" in exI)
  97.569 -  apply (auto simp add: fun_eq)
  97.570 -  done
  97.571 -
  97.572 -text{*At last, we can consider the order of a root.*}
  97.573 -
  97.574 -lemma (in idom_char_0) poly_order_exists_lemma:
  97.575 -  assumes lp: "length p = d"
  97.576 -    and p: "poly p \<noteq> poly []"
  97.577 -  shows "\<exists>n q. p = mulexp n [-a, 1] q \<and> poly q a \<noteq> 0"
  97.578 -  using lp p
  97.579 -proof (induct d arbitrary: p)
  97.580 -  case 0
  97.581 -  thus ?case by simp
  97.582 -next
  97.583 -  case (Suc n p)
  97.584 -  show ?case
  97.585 -  proof (cases "poly p a = 0")
  97.586 -    case True
  97.587 -    from Suc.prems have h: "length p = Suc n" "poly p \<noteq> poly []" by auto
  97.588 -    hence pN: "p \<noteq> []" by auto
  97.589 -    from True[unfolded poly_linear_divides] pN obtain q where q: "p = [-a, 1] *** q"
  97.590 -      by blast
  97.591 -    from q h True have qh: "length q = n" "poly q \<noteq> poly []"
  97.592 -      apply -
  97.593 -      apply simp
  97.594 -      apply (simp only: fun_eq)
  97.595 -      apply (rule ccontr)
  97.596 -      apply (simp add: fun_eq poly_add poly_cmult)
  97.597 -      done
  97.598 -    from Suc.hyps[OF qh] obtain m r where mr: "q = mulexp m [-a,1] r" "poly r a \<noteq> 0"
  97.599 -      by blast
  97.600 -    from mr q have "p = mulexp (Suc m) [-a,1] r \<and> poly r a \<noteq> 0" by simp
  97.601 -    then show ?thesis by blast
  97.602 -  next
  97.603 -    case False
  97.604 -    then show ?thesis
  97.605 -      using Suc.prems
  97.606 -      apply simp
  97.607 -      apply (rule exI[where x="0::nat"])
  97.608 -      apply simp
  97.609 -      done
  97.610 -  qed
  97.611 -qed
  97.612 -
  97.613 -
  97.614 -lemma (in comm_semiring_1) poly_mulexp: "poly (mulexp n p q) x = (poly p x) ^ n * poly q x"
  97.615 -  by (induct n) (auto simp add: poly_mult mult_ac)
  97.616 -
  97.617 -lemma (in comm_semiring_1) divides_left_mult:
  97.618 -  assumes d:"(p***q) divides r" shows "p divides r \<and> q divides r"
  97.619 -proof-
  97.620 -  from d obtain t where r:"poly r = poly (p***q *** t)"
  97.621 -    unfolding divides_def by blast
  97.622 -  hence "poly r = poly (p *** (q *** t))"
  97.623 -    "poly r = poly (q *** (p***t))" by(auto simp add: fun_eq poly_mult mult_ac)
  97.624 -  thus ?thesis unfolding divides_def by blast
  97.625 -qed
  97.626 -
  97.627 -
  97.628 -(* FIXME: Tidy up *)
  97.629 -
  97.630 -lemma (in semiring_1) zero_power_iff: "0 ^ n = (if n = 0 then 1 else 0)"
  97.631 -  by (induct n) simp_all
  97.632 -
  97.633 -lemma (in idom_char_0) poly_order_exists:
  97.634 -  assumes "length p = d" and "poly p \<noteq> poly []"
  97.635 -  shows "\<exists>n. [- a, 1] %^ n divides p \<and> \<not> [- a, 1] %^ Suc n divides p"
  97.636 -proof -
  97.637 -  from assms have "\<exists>n q. p = mulexp n [- a, 1] q \<and> poly q a \<noteq> 0"
  97.638 -    by (rule poly_order_exists_lemma)
  97.639 -  then obtain n q where p: "p = mulexp n [- a, 1] q" and "poly q a \<noteq> 0" by blast
  97.640 -  have "[- a, 1] %^ n divides mulexp n [- a, 1] q"
  97.641 -  proof (rule dividesI)
  97.642 -    show "poly (mulexp n [- a, 1] q) = poly ([- a, 1] %^ n *** q)"
  97.643 -      by (induct n) (simp_all add: poly_add poly_cmult poly_mult distrib_left mult_ac)
  97.644 -  qed
  97.645 -  moreover have "\<not> [- a, 1] %^ Suc n divides mulexp n [- a, 1] q"
  97.646 -  proof
  97.647 -    assume "[- a, 1] %^ Suc n divides mulexp n [- a, 1] q"
  97.648 -    then obtain m where "poly (mulexp n [- a, 1] q) = poly ([- a, 1] %^ Suc n *** m)"
  97.649 -      by (rule dividesE)
  97.650 -    moreover have "poly (mulexp n [- a, 1] q) \<noteq> poly ([- a, 1] %^ Suc n *** m)"
  97.651 -    proof (induct n)
  97.652 -      case 0 show ?case
  97.653 -      proof (rule ccontr)
  97.654 -        assume "\<not> poly (mulexp 0 [- a, 1] q) \<noteq> poly ([- a, 1] %^ Suc 0 *** m)"
  97.655 -        then have "poly q a = 0"
  97.656 -          by (simp add: poly_add poly_cmult)
  97.657 -        with `poly q a \<noteq> 0` show False by simp
  97.658 -      qed
  97.659 -    next
  97.660 -      case (Suc n) show ?case
  97.661 -        by (rule pexp_Suc [THEN ssubst], rule ccontr)
  97.662 -          (simp add: poly_mult_left_cancel poly_mult_assoc Suc del: pmult_Cons pexp_Suc)
  97.663 -    qed
  97.664 -    ultimately show False by simp
  97.665 -  qed
  97.666 -  ultimately show ?thesis by (auto simp add: p)
  97.667 -qed
  97.668 -
  97.669 -lemma (in semiring_1) poly_one_divides[simp]: "[1] divides p"
  97.670 -  by (auto simp add: divides_def)
  97.671 -
  97.672 -lemma (in idom_char_0) poly_order:
  97.673 -  "poly p \<noteq> poly [] \<Longrightarrow> \<exists>!n. ([-a, 1] %^ n) divides p \<and> \<not> (([-a, 1] %^ Suc n) divides p)"
  97.674 -  apply (auto intro: poly_order_exists simp add: less_linear simp del: pmult_Cons pexp_Suc)
  97.675 -  apply (cut_tac x = y and y = n in less_linear)
  97.676 -  apply (drule_tac m = n in poly_exp_divides)
  97.677 -  apply (auto dest: Suc_le_eq [THEN iffD2, THEN [2] poly_exp_divides]
  97.678 -              simp del: pmult_Cons pexp_Suc)
  97.679 -  done
  97.680 -
  97.681 -text{*Order*}
  97.682 -
  97.683 -lemma some1_equalityD: "n = (SOME n. P n) \<Longrightarrow> \<exists>!n. P n \<Longrightarrow> P n"
  97.684 -  by (blast intro: someI2)
  97.685 -
  97.686 -lemma (in idom_char_0) order:
  97.687 -      "(([-a, 1] %^ n) divides p \<and>
  97.688 -        ~(([-a, 1] %^ (Suc n)) divides p)) =
  97.689 -        ((n = order a p) \<and> ~(poly p = poly []))"
  97.690 -  apply (unfold order_def)
  97.691 -  apply (rule iffI)
  97.692 -  apply (blast dest: poly_divides_zero intro!: some1_equality [symmetric] poly_order)
  97.693 -  apply (blast intro!: poly_order [THEN [2] some1_equalityD])
  97.694 -  done
  97.695 -
  97.696 -lemma (in idom_char_0) order2:
  97.697 -  "poly p \<noteq> poly [] \<Longrightarrow>
  97.698 -    ([-a, 1] %^ (order a p)) divides p \<and> \<not> (([-a, 1] %^ (Suc (order a p))) divides p)"
  97.699 -  by (simp add: order del: pexp_Suc)
  97.700 -
  97.701 -lemma (in idom_char_0) order_unique:
  97.702 -  "poly p \<noteq> poly [] \<Longrightarrow> ([-a, 1] %^ n) divides p \<Longrightarrow> ~(([-a, 1] %^ (Suc n)) divides p) \<Longrightarrow>
  97.703 -    n = order a p"
  97.704 -  using order [of a n p] by auto
  97.705 -
  97.706 -lemma (in idom_char_0) order_unique_lemma:
  97.707 -  "poly p \<noteq> poly [] \<and> ([-a, 1] %^ n) divides p \<and> ~(([-a, 1] %^ (Suc n)) divides p) \<Longrightarrow>
  97.708 -    n = order a p"
  97.709 -  by (blast intro: order_unique)
  97.710 -
  97.711 -lemma (in ring_1) order_poly: "poly p = poly q \<Longrightarrow> order a p = order a q"
  97.712 -  by (auto simp add: fun_eq divides_def poly_mult order_def)
  97.713 -
  97.714 -lemma (in semiring_1) pexp_one[simp]: "p %^ (Suc 0) = p"
  97.715 -  by (induct "p") auto
  97.716 -
  97.717 -lemma (in comm_ring_1) lemma_order_root:
  97.718 -  "0 < n \<and> [- a, 1] %^ n divides p \<and> ~ [- a, 1] %^ (Suc n) divides p \<Longrightarrow> poly p a = 0"
  97.719 -  by (induct n arbitrary: a p) (auto simp add: divides_def poly_mult simp del: pmult_Cons)
  97.720 -
  97.721 -lemma (in idom_char_0) order_root:
  97.722 -  "poly p a = 0 \<longleftrightarrow> poly p = poly [] \<or> order a p \<noteq> 0"
  97.723 -  apply (cases "poly p = poly []")
  97.724 -  apply auto
  97.725 -  apply (simp add: poly_linear_divides del: pmult_Cons, safe)
  97.726 -  apply (drule_tac [!] a = a in order2)
  97.727 -  apply (rule ccontr)
  97.728 -  apply (simp add: divides_def poly_mult fun_eq del: pmult_Cons, blast)
  97.729 -  using neq0_conv
  97.730 -  apply (blast intro: lemma_order_root)
  97.731 -  done
  97.732 -
  97.733 -lemma (in idom_char_0) order_divides:
  97.734 -  "([-a, 1] %^ n) divides p \<longleftrightarrow> poly p = poly [] \<or> n \<le> order a p"
  97.735 -  apply (cases "poly p = poly []")
  97.736 -  apply auto
  97.737 -  apply (simp add: divides_def fun_eq poly_mult)
  97.738 -  apply (rule_tac x = "[]" in exI)
  97.739 -  apply (auto dest!: order2 [where a=a] intro: poly_exp_divides simp del: pexp_Suc)
  97.740 -  done
  97.741 -
  97.742 -lemma (in idom_char_0) order_decomp:
  97.743 -  "poly p \<noteq> poly [] \<Longrightarrow> \<exists>q. poly p = poly (([-a, 1] %^ (order a p)) *** q) \<and> ~([-a, 1] divides q)"
  97.744 -  apply (unfold divides_def)
  97.745 -  apply (drule order2 [where a = a])
  97.746 -  apply (simp add: divides_def del: pexp_Suc pmult_Cons, safe)
  97.747 -  apply (rule_tac x = q in exI, safe)
  97.748 -  apply (drule_tac x = qa in spec)
  97.749 -  apply (auto simp add: poly_mult fun_eq poly_exp mult_ac simp del: pmult_Cons)
  97.750 -  done
  97.751 -
  97.752 -text{*Important composition properties of orders.*}
  97.753 -lemma order_mult:
  97.754 -  "poly (p *** q) \<noteq> poly [] \<Longrightarrow>
  97.755 -    order a (p *** q) = order a p + order (a::'a::{idom_char_0}) q"
  97.756 -  apply (cut_tac a = a and p = "p *** q" and n = "order a p + order a q" in order)
  97.757 -  apply (auto simp add: poly_entire simp del: pmult_Cons)
  97.758 -  apply (drule_tac a = a in order2)+
  97.759 -  apply safe
  97.760 -  apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe)
  97.761 -  apply (rule_tac x = "qa *** qaa" in exI)
  97.762 -  apply (simp add: poly_mult mult_ac del: pmult_Cons)
  97.763 -  apply (drule_tac a = a in order_decomp)+
  97.764 -  apply safe
  97.765 -  apply (subgoal_tac "[-a,1] divides (qa *** qaa) ")
  97.766 -  apply (simp add: poly_primes del: pmult_Cons)
  97.767 -  apply (auto simp add: divides_def simp del: pmult_Cons)
  97.768 -  apply (rule_tac x = qb in exI)
  97.769 -  apply (subgoal_tac "poly ([-a, 1] %^ (order a p) *** (qa *** qaa)) = poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))")
  97.770 -  apply (drule poly_mult_left_cancel [THEN iffD1], force)
  97.771 -  apply (subgoal_tac "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (qa *** qaa))) = poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))) ")
  97.772 -  apply (drule poly_mult_left_cancel [THEN iffD1], force)
  97.773 -  apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons)
  97.774 -  done
  97.775 -
  97.776 -lemma (in idom_char_0) order_mult:
  97.777 -  assumes "poly (p *** q) \<noteq> poly []"
  97.778 -  shows "order a (p *** q) = order a p + order a q"
  97.779 -  using assms
  97.780 -  apply (cut_tac a = a and p = "pmult p q" and n = "order a p + order a q" in order)
  97.781 -  apply (auto simp add: poly_entire simp del: pmult_Cons)
  97.782 -  apply (drule_tac a = a in order2)+
  97.783 -  apply safe
  97.784 -  apply (simp add: divides_def fun_eq poly_exp_add poly_mult del: pmult_Cons, safe)
  97.785 -  apply (rule_tac x = "pmult qa qaa" in exI)
  97.786 -  apply (simp add: poly_mult mult_ac del: pmult_Cons)
  97.787 -  apply (drule_tac a = a in order_decomp)+
  97.788 -  apply safe
  97.789 -  apply (subgoal_tac "[uminus a, one] divides pmult qa qaa")
  97.790 -  apply (simp add: poly_primes del: pmult_Cons)
  97.791 -  apply (auto simp add: divides_def simp del: pmult_Cons)
  97.792 -  apply (rule_tac x = qb in exI)
  97.793 -  apply (subgoal_tac "poly (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa)) =
  97.794 -    poly (pmult (pexp [uminus a, one] (?order a p)) (pmult [uminus a, one] qb))")
  97.795 -  apply (drule poly_mult_left_cancel [THEN iffD1], force)
  97.796 -  apply (subgoal_tac "poly (pmult (pexp [uminus a, one] (order a q))
  97.797 -      (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa))) =
  97.798 -    poly (pmult (pexp [uminus a, one] (order a q))
  97.799 -      (pmult (pexp [uminus a, one] (order a p)) (pmult [uminus a, one] qb)))")
  97.800 -  apply (drule poly_mult_left_cancel [THEN iffD1], force)
  97.801 -  apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons)
  97.802 -  done
  97.803 -
  97.804 -lemma (in idom_char_0) order_root2: "poly p \<noteq> poly [] \<Longrightarrow> poly p a = 0 \<longleftrightarrow> order a p \<noteq> 0"
  97.805 -  by (rule order_root [THEN ssubst]) auto
  97.806 -
  97.807 -lemma (in semiring_1) pmult_one[simp]: "[1] *** p = p" by auto
  97.808 -
  97.809 -lemma (in semiring_0) poly_Nil_zero: "poly [] = poly [0]"
  97.810 -  by (simp add: fun_eq)
  97.811 -
  97.812 -lemma (in idom_char_0) rsquarefree_decomp:
  97.813 -  "rsquarefree p \<Longrightarrow> poly p a = 0 \<Longrightarrow>
  97.814 -    \<exists>q. poly p = poly ([-a, 1] *** q) \<and> poly q a \<noteq> 0"
  97.815 -  apply (simp add: rsquarefree_def, safe)
  97.816 -  apply (frule_tac a = a in order_decomp)
  97.817 -  apply (drule_tac x = a in spec)
  97.818 -  apply (drule_tac a = a in order_root2 [symmetric])
  97.819 -  apply (auto simp del: pmult_Cons)
  97.820 -  apply (rule_tac x = q in exI, safe)
  97.821 -  apply (simp add: poly_mult fun_eq)
  97.822 -  apply (drule_tac p1 = q in poly_linear_divides [THEN iffD1])
  97.823 -  apply (simp add: divides_def del: pmult_Cons, safe)
  97.824 -  apply (drule_tac x = "[]" in spec)
  97.825 -  apply (auto simp add: fun_eq)
  97.826 -  done
  97.827 -
  97.828 -
  97.829 -text{*Normalization of a polynomial.*}
  97.830 -
  97.831 -lemma (in semiring_0) poly_normalize[simp]: "poly (pnormalize p) = poly p"
  97.832 -  by (induct p) (auto simp add: fun_eq)
  97.833 -
  97.834 -text{*The degree of a polynomial.*}
  97.835 -
  97.836 -lemma (in semiring_0) lemma_degree_zero: "list_all (%c. c = 0) p \<longleftrightarrow> pnormalize p = []"
  97.837 -  by (induct p) auto
  97.838 -
  97.839 -lemma (in idom_char_0) degree_zero:
  97.840 -  assumes "poly p = poly []"
  97.841 -  shows "degree p = 0"
  97.842 -  using assms
  97.843 -  by (cases "pnormalize p = []") (auto simp add: degree_def poly_zero lemma_degree_zero)
  97.844 -
  97.845 -lemma (in semiring_0) pnormalize_sing: "(pnormalize [x] = [x]) \<longleftrightarrow> x \<noteq> 0"
  97.846 -  by simp
  97.847 -
  97.848 -lemma (in semiring_0) pnormalize_pair: "y \<noteq> 0 \<longleftrightarrow> (pnormalize [x, y] = [x, y])"
  97.849 -  by simp
  97.850 -
  97.851 -lemma (in semiring_0) pnormal_cons: "pnormal p \<Longrightarrow> pnormal (c#p)"
  97.852 -  unfolding pnormal_def by simp
  97.853 -
  97.854 -lemma (in semiring_0) pnormal_tail: "p\<noteq>[] \<Longrightarrow> pnormal (c#p) \<Longrightarrow> pnormal p"
  97.855 -  unfolding pnormal_def by(auto split: split_if_asm)
  97.856 -
  97.857 -
  97.858 -lemma (in semiring_0) pnormal_last_nonzero: "pnormal p \<Longrightarrow> last p \<noteq> 0"
  97.859 -  by (induct p) (simp_all add: pnormal_def split: split_if_asm)
  97.860 -
  97.861 -lemma (in semiring_0) pnormal_length: "pnormal p \<Longrightarrow> 0 < length p"
  97.862 -  unfolding pnormal_def length_greater_0_conv by blast
  97.863 -
  97.864 -lemma (in semiring_0) pnormal_last_length: "0 < length p \<Longrightarrow> last p \<noteq> 0 \<Longrightarrow> pnormal p"
  97.865 -  by (induct p) (auto simp: pnormal_def  split: split_if_asm)
  97.866 -
  97.867 -
  97.868 -lemma (in semiring_0) pnormal_id: "pnormal p \<longleftrightarrow> 0 < length p \<and> last p \<noteq> 0"
  97.869 -  using pnormal_last_length pnormal_length pnormal_last_nonzero by blast
  97.870 -
  97.871 -lemma (in idom_char_0) poly_Cons_eq:
  97.872 -  "poly (c # cs) = poly (d # ds) \<longleftrightarrow> c = d \<and> poly cs = poly ds"
  97.873 -  (is "?lhs \<longleftrightarrow> ?rhs")
  97.874 -proof
  97.875 -  assume eq: ?lhs
  97.876 -  hence "\<And>x. poly ((c#cs) +++ -- (d#ds)) x = 0"
  97.877 -    by (simp only: poly_minus poly_add algebra_simps) simp
  97.878 -  hence "poly ((c#cs) +++ -- (d#ds)) = poly []" by(simp add: fun_eq_iff)
  97.879 -  hence "c = d \<and> list_all (\<lambda>x. x=0) ((cs +++ -- ds))"
  97.880 -    unfolding poly_zero by (simp add: poly_minus_def algebra_simps)
  97.881 -  hence "c = d \<and> (\<forall>x. poly (cs +++ -- ds) x = 0)"
  97.882 -    unfolding poly_zero[symmetric] by simp
  97.883 -  then show ?rhs by (simp add: poly_minus poly_add algebra_simps fun_eq_iff)
  97.884 -next
  97.885 -  assume ?rhs
  97.886 -  then show ?lhs by(simp add:fun_eq_iff)
  97.887 -qed
  97.888 -
  97.889 -lemma (in idom_char_0) pnormalize_unique: "poly p = poly q \<Longrightarrow> pnormalize p = pnormalize q"
  97.890 -proof (induct q arbitrary: p)
  97.891 -  case Nil
  97.892 -  thus ?case by (simp only: poly_zero lemma_degree_zero) simp
  97.893 -next
  97.894 -  case (Cons c cs p)
  97.895 -  thus ?case
  97.896 -  proof (induct p)
  97.897 -    case Nil
  97.898 -    hence "poly [] = poly (c#cs)" by blast
  97.899 -    then have "poly (c#cs) = poly [] " by simp
  97.900 -    thus ?case by (simp only: poly_zero lemma_degree_zero) simp
  97.901 -  next
  97.902 -    case (Cons d ds)
  97.903 -    hence eq: "poly (d # ds) = poly (c # cs)" by blast
  97.904 -    hence eq': "\<And>x. poly (d # ds) x = poly (c # cs) x" by simp
  97.905 -    hence "poly (d # ds) 0 = poly (c # cs) 0" by blast
  97.906 -    hence dc: "d = c" by auto
  97.907 -    with eq have "poly ds = poly cs"
  97.908 -      unfolding  poly_Cons_eq by simp
  97.909 -    with Cons.prems have "pnormalize ds = pnormalize cs" by blast
  97.910 -    with dc show ?case by simp
  97.911 -  qed
  97.912 -qed
  97.913 -
  97.914 -lemma (in idom_char_0) degree_unique:
  97.915 -  assumes pq: "poly p = poly q"
  97.916 -  shows "degree p = degree q"
  97.917 -  using pnormalize_unique[OF pq] unfolding degree_def by simp
  97.918 -
  97.919 -lemma (in semiring_0) pnormalize_length:
  97.920 -  "length (pnormalize p) \<le> length p" by (induct p) auto
  97.921 -
  97.922 -lemma (in semiring_0) last_linear_mul_lemma:
  97.923 -  "last ((a %* p) +++ (x#(b %* p))) = (if p = [] then x else b * last p)"
  97.924 -  apply (induct p arbitrary: a x b)
  97.925 -  apply auto
  97.926 -  apply (subgoal_tac "padd (cmult aa p) (times b a # cmult b p) \<noteq> []")
  97.927 -  apply simp
  97.928 -  apply (induct_tac p)
  97.929 -  apply auto
  97.930 -  done
  97.931 -
  97.932 -lemma (in semiring_1) last_linear_mul:
  97.933 -  assumes p: "p \<noteq> []"
  97.934 -  shows "last ([a,1] *** p) = last p"
  97.935 -proof -
  97.936 -  from p obtain c cs where cs: "p = c#cs" by (cases p) auto
  97.937 -  from cs have eq: "[a,1] *** p = (a %* (c#cs)) +++ (0#(1 %* (c#cs)))"
  97.938 -    by (simp add: poly_cmult_distr)
  97.939 -  show ?thesis using cs
  97.940 -    unfolding eq last_linear_mul_lemma by simp
  97.941 -qed
  97.942 -
  97.943 -lemma (in semiring_0) pnormalize_eq: "last p \<noteq> 0 \<Longrightarrow> pnormalize p = p"
  97.944 -  by (induct p) (auto split: split_if_asm)
  97.945 -
  97.946 -lemma (in semiring_0) last_pnormalize: "pnormalize p \<noteq> [] \<Longrightarrow> last (pnormalize p) \<noteq> 0"
  97.947 -  by (induct p) auto
  97.948 -
  97.949 -lemma (in semiring_0) pnormal_degree: "last p \<noteq> 0 \<Longrightarrow> degree p = length p - 1"
  97.950 -  using pnormalize_eq[of p] unfolding degree_def by simp
  97.951 -
  97.952 -lemma (in semiring_0) poly_Nil_ext: "poly [] = (\<lambda>x. 0)"
  97.953 -  by (rule ext) simp
  97.954 -
  97.955 -lemma (in idom_char_0) linear_mul_degree:
  97.956 -  assumes p: "poly p \<noteq> poly []"
  97.957 -  shows "degree ([a,1] *** p) = degree p + 1"
  97.958 -proof -
  97.959 -  from p have pnz: "pnormalize p \<noteq> []"
  97.960 -    unfolding poly_zero lemma_degree_zero .
  97.961 -
  97.962 -  from last_linear_mul[OF pnz, of a] last_pnormalize[OF pnz]
  97.963 -  have l0: "last ([a, 1] *** pnormalize p) \<noteq> 0" by simp
  97.964 -  from last_pnormalize[OF pnz] last_linear_mul[OF pnz, of a]
  97.965 -    pnormal_degree[OF l0] pnormal_degree[OF last_pnormalize[OF pnz]] pnz
  97.966 -
  97.967 -  have th: "degree ([a,1] *** pnormalize p) = degree (pnormalize p) + 1"
  97.968 -    by simp
  97.969 -
  97.970 -  have eqs: "poly ([a,1] *** pnormalize p) = poly ([a,1] *** p)"
  97.971 -    by (rule ext) (simp add: poly_mult poly_add poly_cmult)
  97.972 -  from degree_unique[OF eqs] th
  97.973 -  show ?thesis by (simp add: degree_unique[OF poly_normalize])
  97.974 -qed
  97.975 -
  97.976 -lemma (in idom_char_0) linear_pow_mul_degree:
  97.977 -  "degree([a,1] %^n *** p) = (if poly p = poly [] then 0 else degree p + n)"
  97.978 -proof (induct n arbitrary: a p)
  97.979 -  case (0 a p)
  97.980 -  show ?case
  97.981 -  proof (cases "poly p = poly []")
  97.982 -    case True
  97.983 -    then show ?thesis
  97.984 -      using degree_unique[OF True] by (simp add: degree_def)
  97.985 -  next
  97.986 -    case False
  97.987 -    then show ?thesis by (auto simp add: poly_Nil_ext)
  97.988 -  qed
  97.989 -next
  97.990 -  case (Suc n a p)
  97.991 -  have eq: "poly ([a,1] %^(Suc n) *** p) = poly ([a,1] %^ n *** ([a,1] *** p))"
  97.992 -    apply (rule ext)
  97.993 -    apply (simp add: poly_mult poly_add poly_cmult)
  97.994 -    apply (simp add: mult_ac add_ac distrib_left)
  97.995 -    done
  97.996 -  note deq = degree_unique[OF eq]
  97.997 -  show ?case
  97.998 -  proof (cases "poly p = poly []")
  97.999 -    case True
 97.1000 -    with eq have eq': "poly ([a,1] %^(Suc n) *** p) = poly []"
 97.1001 -      apply -
 97.1002 -      apply (rule ext)
 97.1003 -      apply (simp add: poly_mult poly_cmult poly_add)
 97.1004 -      done
 97.1005 -    from degree_unique[OF eq'] True show ?thesis
 97.1006 -      by (simp add: degree_def)
 97.1007 -  next
 97.1008 -    case False
 97.1009 -    then have ap: "poly ([a,1] *** p) \<noteq> poly []"
 97.1010 -      using poly_mult_not_eq_poly_Nil unfolding poly_entire by auto
 97.1011 -    have eq: "poly ([a,1] %^(Suc n) *** p) = poly ([a,1]%^n *** ([a,1] *** p))"
 97.1012 -      by (rule ext, simp add: poly_mult poly_add poly_exp poly_cmult algebra_simps)
 97.1013 -    from ap have ap': "(poly ([a,1] *** p) = poly []) = False"
 97.1014 -      by blast
 97.1015 -    have th0: "degree ([a,1]%^n *** ([a,1] *** p)) = degree ([a,1] *** p) + n"
 97.1016 -      apply (simp only: Suc.hyps[of a "pmult [a,one] p"] ap')
 97.1017 -      apply simp
 97.1018 -      done
 97.1019 -    from degree_unique[OF eq] ap False th0 linear_mul_degree[OF False, of a]
 97.1020 -    show ?thesis by (auto simp del: poly.simps)
 97.1021 -  qed
 97.1022 -qed
 97.1023 -
 97.1024 -lemma (in idom_char_0) order_degree:
 97.1025 -  assumes p0: "poly p \<noteq> poly []"
 97.1026 -  shows "order a p \<le> degree p"
 97.1027 -proof -
 97.1028 -  from order2[OF p0, unfolded divides_def]
 97.1029 -  obtain q where q: "poly p = poly ([- a, 1]%^ (order a p) *** q)" by blast
 97.1030 -  {
 97.1031 -    assume "poly q = poly []"
 97.1032 -    with q p0 have False by (simp add: poly_mult poly_entire)
 97.1033 -  }
 97.1034 -  with degree_unique[OF q, unfolded linear_pow_mul_degree] show ?thesis
 97.1035 -    by auto
 97.1036 -qed
 97.1037 -
 97.1038 -text{*Tidier versions of finiteness of roots.*}
 97.1039 -
 97.1040 -lemma (in idom_char_0) poly_roots_finite_set:
 97.1041 -  "poly p \<noteq> poly [] \<Longrightarrow> finite {x. poly p x = 0}"
 97.1042 -  unfolding poly_roots_finite .
 97.1043 -
 97.1044 -text{*bound for polynomial.*}
 97.1045 -
 97.1046 -lemma poly_mono: "abs(x) \<le> k \<Longrightarrow> abs(poly p (x::'a::{linordered_idom})) \<le> poly (map abs p) k"
 97.1047 -  apply (induct p)
 97.1048 -  apply auto
 97.1049 -  apply (rule_tac y = "abs a + abs (x * poly p x)" in order_trans)
 97.1050 -  apply (rule abs_triangle_ineq)
 97.1051 -  apply (auto intro!: mult_mono simp add: abs_mult)
 97.1052 -  done
 97.1053 -
 97.1054 -lemma (in semiring_0) poly_Sing: "poly [c] x = c" by simp
 97.1055 -
 97.1056 -end
    98.1 --- a/src/HOL/Library/While_Combinator.thy	Mon Nov 11 17:34:44 2013 +0100
    98.2 +++ b/src/HOL/Library/While_Combinator.thy	Mon Nov 11 17:44:21 2013 +0100
    98.3 @@ -307,37 +307,44 @@
    98.4  by Nipkow (the theories are in the AFP entry Flyspeck by Nipkow)
    98.5  and the AFP article Executable Transitive Closures by René Thiemann. *}
    98.6  
    98.7 -definition rtrancl_while :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a list) \<Rightarrow> 'a
    98.8 -  \<Rightarrow> ('a list * 'a set) option"
    98.9 -where "rtrancl_while p f x =
   98.10 -  while_option (%(ws,_). ws \<noteq> [] \<and> p(hd ws))
   98.11 -    ((%(ws,Z).
   98.12 -     let x = hd ws; new = filter (\<lambda>y. y \<notin> Z) (f x)
   98.13 -     in (new @ tl ws, set new \<union> Z)))
   98.14 -    ([x],{x})"
   98.15 +context
   98.16 +fixes p :: "'a \<Rightarrow> bool"
   98.17 +and f :: "'a \<Rightarrow> 'a list"
   98.18 +and x :: 'a
   98.19 +begin
   98.20  
   98.21 -lemma rtrancl_while_Some: assumes "rtrancl_while p f x = Some(ws,Z)"
   98.22 +fun rtrancl_while_test :: "'a list \<times> 'a set \<Rightarrow> bool"
   98.23 +where "rtrancl_while_test (ws,_) = (ws \<noteq> [] \<and> p(hd ws))"
   98.24 +
   98.25 +fun rtrancl_while_step :: "'a list \<times> 'a set \<Rightarrow> 'a list \<times> 'a set"
   98.26 +where "rtrancl_while_step (ws, Z) =
   98.27 +  (let x = hd ws; new = remdups (filter (\<lambda>y. y \<notin> Z) (f x))
   98.28 +  in (new @ tl ws, set new \<union> Z))"
   98.29 +
   98.30 +definition rtrancl_while :: "('a list * 'a set) option"
   98.31 +where "rtrancl_while = while_option rtrancl_while_test rtrancl_while_step ([x],{x})"
   98.32 +
   98.33 +fun rtrancl_while_invariant :: "'a list \<times> 'a set \<Rightarrow> bool"
   98.34 +where "rtrancl_while_invariant (ws, Z) =
   98.35 +   (x \<in> Z \<and> set ws \<subseteq> Z \<and> distinct ws \<and> {(x,y). y \<in> set(f x)} `` (Z - set ws) \<subseteq> Z \<and>
   98.36 +    Z \<subseteq> {(x,y). y \<in> set(f x)}^* `` {x} \<and> (\<forall>z\<in>Z - set ws. p z))"
   98.37 +
   98.38 +lemma rtrancl_while_invariant: 
   98.39 +  assumes inv: "rtrancl_while_invariant st" and test: "rtrancl_while_test st"
   98.40 +  shows   "rtrancl_while_invariant (rtrancl_while_step st)"
   98.41 +proof (cases st)
   98.42 +  fix ws Z assume st: "st = (ws, Z)"
   98.43 +  with test obtain h t where "ws = h # t" "p h" by (cases ws) auto
   98.44 +  with inv st show ?thesis by (auto intro: rtrancl.rtrancl_into_rtrancl)
   98.45 +qed
   98.46 +
   98.47 +lemma rtrancl_while_Some: assumes "rtrancl_while = Some(ws,Z)"
   98.48  shows "if ws = []
   98.49         then Z = {(x,y). y \<in> set(f x)}^* `` {x} \<and> (\<forall>z\<in>Z. p z)
   98.50         else \<not>p(hd ws) \<and> hd ws \<in> {(x,y). y \<in> set(f x)}^* `` {x}"
   98.51 -proof-
   98.52 -  let ?test = "(%(ws,_). ws \<noteq> [] \<and> p(hd ws))"
   98.53 -  let ?step = "(%(ws,Z).
   98.54 -     let x = hd ws; new = filter (\<lambda>y. y \<notin> Z) (f x)
   98.55 -     in (new @ tl ws, set new \<union> Z))"
   98.56 -  let ?R = "{(x,y). y \<in> set(f x)}"
   98.57 -  let ?Inv = "%(ws,Z). x \<in> Z \<and> set ws \<subseteq> Z \<and> ?R `` (Z - set ws) \<subseteq> Z \<and>
   98.58 -                       Z \<subseteq> ?R^* `` {x} \<and> (\<forall>z\<in>Z - set ws. p z)"
   98.59 -  { fix ws Z assume 1: "?Inv(ws,Z)" and 2: "?test(ws,Z)"
   98.60 -    from 2 obtain v vs where [simp]: "ws = v#vs" by (auto simp: neq_Nil_conv)
   98.61 -    have "?Inv(?step (ws,Z))" using 1 2
   98.62 -      by (auto intro: rtrancl.rtrancl_into_rtrancl)
   98.63 -  } note inv = this
   98.64 -  hence "!!p. ?Inv p \<Longrightarrow> ?test p \<Longrightarrow> ?Inv(?step p)"
   98.65 -    apply(tactic {* split_all_tac @{context} 1 *})
   98.66 -    using inv by iprover
   98.67 -  moreover have "?Inv ([x],{x})" by (simp)
   98.68 -  ultimately have I: "?Inv (ws,Z)"
   98.69 +proof -
   98.70 +  have "rtrancl_while_invariant ([x],{x})" by simp
   98.71 +  with rtrancl_while_invariant have I: "rtrancl_while_invariant (ws,Z)"
   98.72      by (rule while_option_rule[OF _ assms[unfolded rtrancl_while_def]])
   98.73    { assume "ws = []"
   98.74      hence ?thesis using I
   98.75 @@ -350,4 +357,41 @@
   98.76    ultimately show ?thesis by simp
   98.77  qed
   98.78  
   98.79 +lemma rtrancl_while_finite_Some:
   98.80 +  assumes "finite ({(x, y). y \<in> set (f x)}\<^sup>* `` {x})" (is "finite ?Cl")
   98.81 +  shows "\<exists>y. rtrancl_while = Some y"
   98.82 +proof -
   98.83 +  let ?R = "(\<lambda>(_, Z). card (?Cl - Z)) <*mlex*> (\<lambda>(ws, _). length ws) <*mlex*> {}"
   98.84 +  have "wf ?R" by (blast intro: wf_mlex)
   98.85 +  then show ?thesis unfolding rtrancl_while_def
   98.86 +  proof (rule wf_rel_while_option_Some[of ?R rtrancl_while_invariant])
   98.87 +    fix st assume *: "rtrancl_while_invariant st \<and> rtrancl_while_test st"
   98.88 +    hence I: "rtrancl_while_invariant (rtrancl_while_step st)"
   98.89 +      by (blast intro: rtrancl_while_invariant)
   98.90 +    show "(rtrancl_while_step st, st) \<in> ?R"
   98.91 +    proof (cases st)
   98.92 +      fix ws Z let ?ws = "fst (rtrancl_while_step st)" and ?Z = "snd (rtrancl_while_step st)"
   98.93 +      assume st: "st = (ws, Z)"
   98.94 +      with * obtain h t where ws: "ws = h # t" "p h" by (cases ws) auto
   98.95 +      { assume "remdups (filter (\<lambda>y. y \<notin> Z) (f h)) \<noteq> []"
   98.96 +        then obtain z where "z \<in> set (remdups (filter (\<lambda>y. y \<notin> Z) (f h)))" by fastforce
   98.97 +        with st ws I have "Z \<subset> ?Z" "Z \<subseteq> ?Cl" "?Z \<subseteq> ?Cl" by auto
   98.98 +        with assms have "card (?Cl - ?Z) < card (?Cl - Z)" by (blast intro: psubset_card_mono)
   98.99 +        with st ws have ?thesis unfolding mlex_prod_def by simp
  98.100 +      }
  98.101 +      moreover
  98.102 +      { assume "remdups (filter (\<lambda>y. y \<notin> Z) (f h)) = []"
  98.103 +        with st ws have "?Z = Z" "?ws = t"  by (auto simp: filter_empty_conv)
  98.104 +        with st ws have ?thesis unfolding mlex_prod_def by simp
  98.105 +      }
  98.106 +      ultimately show ?thesis by blast
  98.107 +    qed
  98.108 +  qed (simp_all add: rtrancl_while_invariant)
  98.109 +qed
  98.110 +
  98.111  end
  98.112 +
  98.113 +hide_const (open) rtrancl_while_test rtrancl_while_step rtrancl_while_invariant
  98.114 +hide_fact (open) rtrancl_while_invariant
  98.115 +
  98.116 +end
    99.1 --- a/src/HOL/Lifting_Set.thy	Mon Nov 11 17:34:44 2013 +0100
    99.2 +++ b/src/HOL/Lifting_Set.thy	Mon Nov 11 17:44:21 2013 +0100
    99.3 @@ -153,7 +153,7 @@
    99.4    unfolding fun_rel_def set_rel_def by fast
    99.5  
    99.6  lemma SUPR_parametric [transfer_rule]:
    99.7 -  "(set_rel R ===> (R ===> op =) ===> op =) SUPR SUPR"
    99.8 +  "(set_rel R ===> (R ===> op =) ===> op =) SUPR (SUPR :: _ \<Rightarrow> _ \<Rightarrow> _::complete_lattice)"
    99.9  proof(rule fun_relI)+
   99.10    fix A B f and g :: "'b \<Rightarrow> 'c"
   99.11    assume AB: "set_rel R A B"
   100.1 --- a/src/HOL/Limits.thy	Mon Nov 11 17:34:44 2013 +0100
   100.2 +++ b/src/HOL/Limits.thy	Mon Nov 11 17:44:21 2013 +0100
   100.3 @@ -138,6 +138,18 @@
   100.4  lemma BseqI: "[| 0 < K; \<forall>n. norm (X n) \<le> K |] ==> Bseq X"
   100.5  by (auto simp add: Bseq_def)
   100.6  
   100.7 +lemma Bseq_bdd_above: "Bseq (X::nat \<Rightarrow> real) \<Longrightarrow> bdd_above (range X)"
   100.8 +proof (elim BseqE, intro bdd_aboveI2)
   100.9 +  fix K n assume "0 < K" "\<forall>n. norm (X n) \<le> K" then show "X n \<le> K"
  100.10 +    by (auto elim!: allE[of _ n])
  100.11 +qed
  100.12 +
  100.13 +lemma Bseq_bdd_below: "Bseq (X::nat \<Rightarrow> real) \<Longrightarrow> bdd_below (range X)"
  100.14 +proof (elim BseqE, intro bdd_belowI2)
  100.15 +  fix K n assume "0 < K" "\<forall>n. norm (X n) \<le> K" then show "- K \<le> X n"
  100.16 +    by (auto elim!: allE[of _ n])
  100.17 +qed
  100.18 +
  100.19  lemma lemma_NBseq_def:
  100.20    "(\<exists>K > 0. \<forall>n. norm (X n) \<le> K) = (\<exists>N. \<forall>n. norm (X n) \<le> real(Suc N))"
  100.21  proof safe
  100.22 @@ -179,7 +191,7 @@
  100.23  apply (rule_tac x = K in exI, simp)
  100.24  apply (rule exI [where x = 0], auto)
  100.25  apply (erule order_less_le_trans, simp)
  100.26 -apply (drule_tac x=n in spec, fold diff_minus)
  100.27 +apply (drule_tac x=n in spec)
  100.28  apply (drule order_trans [OF norm_triangle_ineq2])
  100.29  apply simp
  100.30  done
  100.31 @@ -192,9 +204,11 @@
  100.32    then obtain K
  100.33      where *: "0 < K" and **: "\<And>n. norm (X n) \<le> K" by (auto simp add: Bseq_def)
  100.34    from * have "0 < K + norm (X 0)" by (rule order_less_le_trans) simp
  100.35 -  moreover from ** have "\<forall>n. norm (X n + - X 0) \<le> K + norm (X 0)"
  100.36 -    by (auto intro: order_trans norm_triangle_ineq)
  100.37 -  ultimately show ?Q by blast
  100.38 +  from ** have "\<forall>n. norm (X n - X 0) \<le> K + norm (X 0)"
  100.39 +    by (auto intro: order_trans norm_triangle_ineq4)
  100.40 +  then have "\<forall>n. norm (X n + - X 0) \<le> K + norm (X 0)"
  100.41 +    by simp
  100.42 +  with `0 < K + norm (X 0)` show ?Q by blast
  100.43  next
  100.44    assume ?Q then show ?P by (auto simp add: Bseq_iff2)
  100.45  qed
  100.46 @@ -205,20 +219,9 @@
  100.47  apply (drule_tac x = n in spec, arith)
  100.48  done
  100.49  
  100.50 +
  100.51  subsubsection{*Upper Bounds and Lubs of Bounded Sequences*}
  100.52  
  100.53 -lemma Bseq_isUb:
  100.54 -  "!!(X::nat=>real). Bseq X ==> \<exists>U. isUb (UNIV::real set) {x. \<exists>n. X n = x} U"
  100.55 -by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff)
  100.56 -
  100.57 -text{* Use completeness of reals (supremum property)
  100.58 -   to show that any bounded sequence has a least upper bound*}
  100.59 -
  100.60 -lemma Bseq_isLub:
  100.61 -  "!!(X::nat=>real). Bseq X ==>
  100.62 -   \<exists>U. isLub (UNIV::real set) {x. \<exists>n. X n = x} U"
  100.63 -by (blast intro: reals_complete Bseq_isUb)
  100.64 -
  100.65  lemma Bseq_minus_iff: "Bseq (%n. -(X n) :: 'a :: real_normed_vector) = Bseq X"
  100.66    by (simp add: Bseq_def)
  100.67  
  100.68 @@ -342,7 +345,7 @@
  100.69    unfolding Zfun_def by simp
  100.70  
  100.71  lemma Zfun_diff: "\<lbrakk>Zfun f F; Zfun g F\<rbrakk> \<Longrightarrow> Zfun (\<lambda>x. f x - g x) F"
  100.72 -  by (simp only: diff_minus Zfun_add Zfun_minus)
  100.73 +  using Zfun_add [of f F "\<lambda>x. - g x"] by (simp add: Zfun_minus)
  100.74  
  100.75  lemma (in bounded_linear) Zfun:
  100.76    assumes g: "Zfun g F"
  100.77 @@ -532,7 +535,7 @@
  100.78  lemma tendsto_diff [tendsto_intros]:
  100.79    fixes a b :: "'a::real_normed_vector"
  100.80    shows "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x - g x) ---> a - b) F"
  100.81 -  by (simp add: diff_minus tendsto_add tendsto_minus)
  100.82 +  using tendsto_add [of f a F "\<lambda>x. - g x" "- b"] by (simp add: tendsto_minus)
  100.83  
  100.84  lemma continuous_diff [continuous_intros]:
  100.85    fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
  100.86 @@ -1355,40 +1358,29 @@
  100.87  
  100.88  text {* A monotone sequence converges to its least upper bound. *}
  100.89  
  100.90 -lemma isLub_mono_imp_LIMSEQ:
  100.91 -  fixes X :: "nat \<Rightarrow> real"
  100.92 -  assumes u: "isLub UNIV {x. \<exists>n. X n = x} u" (* FIXME: use 'range X' *)
  100.93 -  assumes X: "\<forall>m n. m \<le> n \<longrightarrow> X m \<le> X n"
  100.94 -  shows "X ----> u"
  100.95 -proof (rule LIMSEQ_I)
  100.96 -  have 1: "\<forall>n. X n \<le> u"
  100.97 -    using isLubD2 [OF u] by auto
  100.98 -  have "\<forall>y. (\<forall>n. X n \<le> y) \<longrightarrow> u \<le> y"
  100.99 -    using isLub_le_isUb [OF u] by (auto simp add: isUb_def setle_def)
 100.100 -  hence 2: "\<forall>y<u. \<exists>n. y < X n"
 100.101 -    by (metis not_le)
 100.102 -  fix r :: real assume "0 < r"
 100.103 -  hence "u - r < u" by simp
 100.104 -  hence "\<exists>m. u - r < X m" using 2 by simp
 100.105 -  then obtain m where "u - r < X m" ..
 100.106 -  with X have "\<forall>n\<ge>m. u - r < X n"
 100.107 -    by (fast intro: less_le_trans)
 100.108 -  hence "\<exists>m. \<forall>n\<ge>m. u - r < X n" ..
 100.109 -  thus "\<exists>m. \<forall>n\<ge>m. norm (X n - u) < r"
 100.110 -    using 1 by (simp add: diff_less_eq add_commute)
 100.111 -qed
 100.112 +lemma LIMSEQ_incseq_SUP:
 100.113 +  fixes X :: "nat \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology}"
 100.114 +  assumes u: "bdd_above (range X)"
 100.115 +  assumes X: "incseq X"
 100.116 +  shows "X ----> (SUP i. X i)"
 100.117 +  by (rule order_tendstoI)
 100.118 +     (auto simp: eventually_sequentially u less_cSUP_iff intro: X[THEN incseqD] less_le_trans cSUP_lessD[OF u])
 100.119  
 100.120 -text{*A standard proof of the theorem for monotone increasing sequence*}
 100.121 -
 100.122 -lemma Bseq_mono_convergent:
 100.123 -   "Bseq X \<Longrightarrow> \<forall>m. \<forall>n \<ge> m. X m \<le> X n \<Longrightarrow> convergent (X::nat=>real)"
 100.124 -  by (metis Bseq_isLub isLub_mono_imp_LIMSEQ convergentI)
 100.125 +lemma LIMSEQ_decseq_INF:
 100.126 +  fixes X :: "nat \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology}"
 100.127 +  assumes u: "bdd_below (range X)"
 100.128 +  assumes X: "decseq X"
 100.129 +  shows "X ----> (INF i. X i)"
 100.130 +  by (rule order_tendstoI)
 100.131 +     (auto simp: eventually_sequentially u cINF_less_iff intro: X[THEN decseqD] le_less_trans less_cINF_D[OF u])
 100.132  
 100.133  text{*Main monotonicity theorem*}
 100.134  
 100.135  lemma Bseq_monoseq_convergent: "Bseq X \<Longrightarrow> monoseq X \<Longrightarrow> convergent (X::nat\<Rightarrow>real)"
 100.136 -  by (metis monoseq_iff incseq_def decseq_eq_incseq convergent_minus_iff Bseq_minus_iff
 100.137 -            Bseq_mono_convergent)
 100.138 +  by (auto simp: monoseq_iff convergent_def intro: LIMSEQ_decseq_INF LIMSEQ_incseq_SUP dest: Bseq_bdd_above Bseq_bdd_below)
 100.139 +
 100.140 +lemma Bseq_mono_convergent: "Bseq X \<Longrightarrow> (\<forall>m n. m \<le> n \<longrightarrow> X m \<le> X n) \<Longrightarrow> convergent (X::nat\<Rightarrow>real)"
 100.141 +  by (auto intro!: Bseq_monoseq_convergent incseq_imp_monoseq simp: incseq_def)
 100.142  
 100.143  lemma Cauchy_iff:
 100.144    fixes X :: "nat \<Rightarrow> 'a::real_normed_vector"
   101.1 --- a/src/HOL/List.thy	Mon Nov 11 17:34:44 2013 +0100
   101.2 +++ b/src/HOL/List.thy	Mon Nov 11 17:44:21 2013 +0100
   101.3 @@ -902,7 +902,7 @@
   101.4  lemma self_append_conv [iff]: "(xs = xs @ ys) = (ys = [])"
   101.5  by (induct xs) auto
   101.6  
   101.7 -lemma append_eq_append_conv [simp, no_atp]:
   101.8 +lemma append_eq_append_conv [simp]:
   101.9   "length xs = length ys \<or> length us = length vs
  101.10   ==> (xs@us = ys@vs) = (xs=ys \<and> us=vs)"
  101.11  apply (induct xs arbitrary: ys)
  101.12 @@ -934,7 +934,7 @@
  101.13  lemma self_append_conv2 [iff]: "(ys = xs @ ys) = (xs = [])"
  101.14  using append_same_eq [of "[]"] by auto
  101.15  
  101.16 -lemma hd_Cons_tl [simp,no_atp]: "xs \<noteq> [] ==> hd xs # tl xs = xs"
  101.17 +lemma hd_Cons_tl [simp]: "xs \<noteq> [] ==> hd xs # tl xs = xs"
  101.18  by (induct xs) auto
  101.19  
  101.20  lemma hd_append: "hd (xs @ ys) = (if xs = [] then hd ys else hd xs)"
  101.21 @@ -1178,7 +1178,7 @@
  101.22  lemma singleton_rev_conv [simp]: "([x] = rev xs) = (xs = [x])"
  101.23  by (cases xs) auto
  101.24  
  101.25 -lemma rev_is_rev_conv [iff, no_atp]: "(rev xs = rev ys) = (xs = ys)"
  101.26 +lemma rev_is_rev_conv [iff]: "(rev xs = rev ys) = (xs = ys)"
  101.27  apply (induct xs arbitrary: ys, force)
  101.28  apply (case_tac ys, simp, force)
  101.29  done
  101.30 @@ -5084,10 +5084,10 @@
  101.31    for A :: "'a set"
  101.32  where
  101.33      Nil [intro!, simp]: "[]: lists A"
  101.34 -  | Cons [intro!, simp, no_atp]: "[| a: A; l: lists A|] ==> a#l : lists A"
  101.35 -
  101.36 -inductive_cases listsE [elim!,no_atp]: "x#l : lists A"
  101.37 -inductive_cases listspE [elim!,no_atp]: "listsp A (x # l)"
  101.38 +  | Cons [intro!, simp]: "[| a: A; l: lists A|] ==> a#l : lists A"
  101.39 +
  101.40 +inductive_cases listsE [elim!]: "x#l : lists A"
  101.41 +inductive_cases listspE [elim!]: "listsp A (x # l)"
  101.42  
  101.43  inductive_simps listsp_simps[code]:
  101.44    "listsp A []"
  101.45 @@ -5129,15 +5129,15 @@
  101.46  
  101.47  lemmas in_lists_conv_set [code_unfold] = in_listsp_conv_set [to_set]
  101.48  
  101.49 -lemma in_listspD [dest!,no_atp]: "listsp A xs ==> \<forall>x\<in>set xs. A x"
  101.50 +lemma in_listspD [dest!]: "listsp A xs ==> \<forall>x\<in>set xs. A x"
  101.51  by (rule in_listsp_conv_set [THEN iffD1])
  101.52  
  101.53 -lemmas in_listsD [dest!,no_atp] = in_listspD [to_set]
  101.54 -
  101.55 -lemma in_listspI [intro!,no_atp]: "\<forall>x\<in>set xs. A x ==> listsp A xs"
  101.56 +lemmas in_listsD [dest!] = in_listspD [to_set]
  101.57 +
  101.58 +lemma in_listspI [intro!]: "\<forall>x\<in>set xs. A x ==> listsp A xs"
  101.59  by (rule in_listsp_conv_set [THEN iffD2])
  101.60  
  101.61 -lemmas in_listsI [intro!,no_atp] = in_listspI [to_set]
  101.62 +lemmas in_listsI [intro!] = in_listspI [to_set]
  101.63  
  101.64  lemma lists_eq_set: "lists A = {xs. set xs <= A}"
  101.65  by auto
  101.66 @@ -5514,15 +5514,15 @@
  101.67  text{* Accessible part and wellfoundedness: *}
  101.68  
  101.69  lemma Cons_acc_listrel1I [intro!]:
  101.70 -  "x \<in> acc r \<Longrightarrow> xs \<in> acc (listrel1 r) \<Longrightarrow> (x # xs) \<in> acc (listrel1 r)"
  101.71 -apply (induct arbitrary: xs set: acc)
  101.72 +  "x \<in> Wellfounded.acc r \<Longrightarrow> xs \<in> Wellfounded.acc (listrel1 r) \<Longrightarrow> (x # xs) \<in> Wellfounded.acc (listrel1 r)"
  101.73 +apply (induct arbitrary: xs set: Wellfounded.acc)
  101.74  apply (erule thin_rl)
  101.75  apply (erule acc_induct)
  101.76  apply (rule accI)
  101.77  apply (blast)
  101.78  done
  101.79  
  101.80 -lemma lists_accD: "xs \<in> lists (acc r) \<Longrightarrow> xs \<in> acc (listrel1 r)"
  101.81 +lemma lists_accD: "xs \<in> lists (Wellfounded.acc r) \<Longrightarrow> xs \<in> Wellfounded.acc (listrel1 r)"
  101.82  apply (induct set: lists)
  101.83   apply (rule accI)
  101.84   apply simp
  101.85 @@ -5530,8 +5530,8 @@
  101.86  apply (fast dest: acc_downward)
  101.87  done
  101.88  
  101.89 -lemma lists_accI: "xs \<in> acc (listrel1 r) \<Longrightarrow> xs \<in> lists (acc r)"
  101.90 -apply (induct set: acc)
  101.91 +lemma lists_accI: "xs \<in> Wellfounded.acc (listrel1 r) \<Longrightarrow> xs \<in> lists (Wellfounded.acc r)"
  101.92 +apply (induct set: Wellfounded.acc)
  101.93  apply clarify
  101.94  apply (rule accI)
  101.95  apply (fastforce dest!: in_set_conv_decomp[THEN iffD1] simp: listrel1_def)
   102.1 --- a/src/HOL/Lubs.thy	Mon Nov 11 17:34:44 2013 +0100
   102.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
   102.3 @@ -1,103 +0,0 @@
   102.4 -(*  Title:      HOL/Lubs.thy
   102.5 -    Author:     Jacques D. Fleuriot, University of Cambridge
   102.6 -*)
   102.7 -
   102.8 -header {* Definitions of Upper Bounds and Least Upper Bounds *}
   102.9 -
  102.10 -theory Lubs
  102.11 -imports Main
  102.12 -begin
  102.13 -
  102.14 -text {* Thanks to suggestions by James Margetson *}
  102.15 -
  102.16 -definition setle :: "'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"  (infixl "*<=" 70)
  102.17 -  where "S *<= x = (ALL y: S. y \<le> x)"
  102.18 -
  102.19 -definition setge :: "'a::ord \<Rightarrow> 'a set \<Rightarrow> bool"  (infixl "<=*" 70)
  102.20 -  where "x <=* S = (ALL y: S. x \<le> y)"
  102.21 -
  102.22 -definition leastP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a::ord \<Rightarrow> bool"
  102.23 -  where "leastP P x = (P x \<and> x <=* Collect P)"
  102.24 -
  102.25 -definition isUb :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"
  102.26 -  where "isUb R S x = (S *<= x \<and> x: R)"
  102.27 -
  102.28 -definition isLub :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool"
  102.29 -  where "isLub R S x = leastP (isUb R S) x"
  102.30 -
  102.31 -definition ubs :: "'a set \<Rightarrow> 'a::ord set \<Rightarrow> 'a set"
  102.32 -  where "ubs R S = Collect (isUb R S)"
  102.33 -
  102.34 -
  102.35 -subsection {* Rules for the Relations @{text "*<="} and @{text "<=*"} *}
  102.36 -
  102.37 -lemma setleI: "ALL y: S. y \<le> x \<Longrightarrow> S *<= x"
  102.38 -  by (simp add: setle_def)
  102.39 -
  102.40 -lemma setleD: "S *<= x \<Longrightarrow> y: S \<Longrightarrow> y \<le> x"
  102.41 -  by (simp add: setle_def)
  102.42 -
  102.43 -lemma setgeI: "ALL y: S. x \<le> y \<Longrightarrow> x <=* S"
  102.44 -  by (simp add: setge_def)
  102.45 -
  102.46 -lemma setgeD: "x <=* S \<Longrightarrow> y: S \<Longrightarrow> x \<le> y"
  102.47 -  by (simp add: setge_def)
  102.48 -
  102.49 -
  102.50 -subsection {* Rules about the Operators @{term leastP}, @{term ub} and @{term lub} *}
  102.51 -
  102.52 -lemma leastPD1: "leastP P x \<Longrightarrow> P x"
  102.53 -  by (simp add: leastP_def)
  102.54 -
  102.55 -lemma leastPD2: "leastP P x \<Longrightarrow> x <=* Collect P"
  102.56 -  by (simp add: leastP_def)
  102.57 -
  102.58 -lemma leastPD3: "leastP P x \<Longrightarrow> y: Collect P \<Longrightarrow> x \<le> y"
  102.59 -  by (blast dest!: leastPD2 setgeD)
  102.60 -
  102.61 -lemma isLubD1: "isLub R S x \<Longrightarrow> S *<= x"
  102.62 -  by (simp add: isLub_def isUb_def leastP_def)
  102.63 -
  102.64 -lemma isLubD1a: "isLub R S x \<Longrightarrow> x: R"
  102.65 -  by (simp add: isLub_def isUb_def leastP_def)
  102.66 -
  102.67 -lemma isLub_isUb: "isLub R S x \<Longrightarrow> isUb R S x"
  102.68 -  unfolding isUb_def by (blast dest: isLubD1 isLubD1a)
  102.69 -
  102.70 -lemma isLubD2: "isLub R S x \<Longrightarrow> y : S \<Longrightarrow> y \<le> x"
  102.71 -  by (blast dest!: isLubD1 setleD)
  102.72 -
  102.73 -lemma isLubD3: "isLub R S x \<Longrightarrow> leastP (isUb R S) x"
  102.74 -  by (simp add: isLub_def)
  102.75 -
  102.76 -lemma isLubI1: "leastP(isUb R S) x \<Longrightarrow> isLub R S x"
  102.77 -  by (simp add: isLub_def)
  102.78 -
  102.79 -lemma isLubI2: "isUb R S x \<Longrightarrow> x <=* Collect (isUb R S) \<Longrightarrow> isLub R S x"
  102.80 -  by (simp add: isLub_def leastP_def)
  102.81 -
  102.82 -lemma isUbD: "isUb R S x \<Longrightarrow> y : S \<Longrightarrow> y \<le> x"
  102.83 -  by (simp add: isUb_def setle_def)
  102.84 -
  102.85 -lemma isUbD2: "isUb R S x \<Longrightarrow> S *<= x"
  102.86 -  by (simp add: isUb_def)
  102.87 -
  102.88 -lemma isUbD2a: "isUb R S x \<Longrightarrow> x: R"
  102.89 -  by (simp add: isUb_def)
  102.90 -
  102.91 -lemma isUbI: "S *<= x \<Longrightarrow> x: R \<Longrightarrow> isUb R S x"
  102.92 -  by (simp add: isUb_def)
  102.93 -
  102.94 -lemma isLub_le_isUb: "isLub R S x \<Longrightarrow> isUb R S y \<Longrightarrow> x \<le> y"
  102.95 -  unfolding isLub_def by (blast intro!: leastPD3)
  102.96 -
  102.97 -lemma isLub_ubs: "isLub R S x \<Longrightarrow> x <=* ubs R S"
  102.98 -  unfolding ubs_def isLub_def by (rule leastPD2)
  102.99 -
 102.100 -lemma isLub_unique: "[| isLub R S x; isLub R S y |] ==> x = (y::'a::linorder)"
 102.101 -  apply (frule isLub_isUb)
 102.102 -  apply (frule_tac x = y in isLub_isUb)
 102.103 -  apply (blast intro!: order_antisym dest!: isLub_le_isUb)
 102.104 -  done
 102.105 -
 102.106 -end
   103.1 --- a/src/HOL/Matrix_LP/LP.thy	Mon Nov 11 17:34:44 2013 +0100
   103.2 +++ b/src/HOL/Matrix_LP/LP.thy	Mon Nov 11 17:44:21 2013 +0100
   103.3 @@ -72,8 +72,7 @@
   103.4  proof -
   103.5    have "0 <= A - A1"    
   103.6    proof -
   103.7 -    have 1: "A - A1 = A + (- A1)" by simp
   103.8 -    show ?thesis by (simp only: 1 add_right_mono[of A1 A "-A1", simplified, simplified assms])
   103.9 +    from assms add_right_mono [of A1 A "- A1"] show ?thesis by simp
  103.10    qed
  103.11    then have "abs (A-A1) = A-A1" by (rule abs_of_nonneg)
  103.12    with assms show "abs (A-A1) <= (A2-A1)" by simp
  103.13 @@ -147,9 +146,9 @@
  103.14    then have "c * x <= y * b - (y * A - c) * x" by (simp add: le_diff_eq)
  103.15    then have cx: "c * x <= y * b + (c - y * A) * x" by (simp add: algebra_simps)
  103.16    have s2: "c - y * A <= c2 - y * A1"
  103.17 -    by (simp add: diff_minus assms add_mono mult_left_mono)
  103.18 +    by (simp add: assms add_mono mult_left_mono algebra_simps)
  103.19    have s1: "c1 - y * A2 <= c - y * A"
  103.20 -    by (simp add: diff_minus assms add_mono mult_left_mono)
  103.21 +    by (simp add: assms add_mono mult_left_mono algebra_simps)
  103.22    have prts: "(c - y * A) * x <= ?C"
  103.23      apply (simp add: Let_def)
  103.24      apply (rule mult_le_prts)
   104.1 --- a/src/HOL/Matrix_LP/Matrix.thy	Mon Nov 11 17:34:44 2013 +0100
   104.2 +++ b/src/HOL/Matrix_LP/Matrix.thy	Mon Nov 11 17:44:21 2013 +0100
   104.3 @@ -1542,8 +1542,8 @@
   104.4    fix A B :: "'a matrix"
   104.5    show "- A + A = 0" 
   104.6      by (simp add: plus_matrix_def minus_matrix_def Rep_matrix_inject[symmetric] ext)
   104.7 -  show "A - B = A + - B"
   104.8 -    by (simp add: plus_matrix_def diff_matrix_def minus_matrix_def Rep_matrix_inject [symmetric] ext diff_minus)
   104.9 +  show "A + - B = A - B"
  104.10 +    by (simp add: plus_matrix_def diff_matrix_def minus_matrix_def Rep_matrix_inject [symmetric] ext)
  104.11  qed
  104.12  
  104.13  instance matrix :: (ab_group_add) ab_group_add
   105.1 --- a/src/HOL/Meson.thy	Mon Nov 11 17:34:44 2013 +0100
   105.2 +++ b/src/HOL/Meson.thy	Mon Nov 11 17:44:21 2013 +0100
   105.3 @@ -132,45 +132,45 @@
   105.4  text{* Combinator translation helpers *}
   105.5  
   105.6  definition COMBI :: "'a \<Rightarrow> 'a" where
   105.7 -[no_atp]: "COMBI P = P"
   105.8 +"COMBI P = P"
   105.9  
  105.10  definition COMBK :: "'a \<Rightarrow> 'b \<Rightarrow> 'a" where
  105.11 -[no_atp]: "COMBK P Q = P"
  105.12 +"COMBK P Q = P"
  105.13  
  105.14 -definition COMBB :: "('b => 'c) \<Rightarrow> ('a => 'b) \<Rightarrow> 'a \<Rightarrow> 'c" where [no_atp]:
  105.15 +definition COMBB :: "('b => 'c) \<Rightarrow> ('a => 'b) \<Rightarrow> 'a \<Rightarrow> 'c" where
  105.16  "COMBB P Q R = P (Q R)"
  105.17  
  105.18  definition COMBC :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'c" where
  105.19 -[no_atp]: "COMBC P Q R = P R Q"
  105.20 +"COMBC P Q R = P R Q"
  105.21  
  105.22  definition COMBS :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c" where
  105.23 -[no_atp]: "COMBS P Q R = P R (Q R)"
  105.24 +"COMBS P Q R = P R (Q R)"
  105.25  
  105.26 -lemma abs_S [no_atp]: "\<lambda>x. (f x) (g x) \<equiv> COMBS f g"
  105.27 +lemma abs_S: "\<lambda>x. (f x) (g x) \<equiv> COMBS f g"
  105.28  apply (rule eq_reflection)
  105.29  apply (rule ext) 
  105.30  apply (simp add: COMBS_def) 
  105.31  done
  105.32  
  105.33 -lemma abs_I [no_atp]: "\<lambda>x. x \<equiv> COMBI"
  105.34 +lemma abs_I: "\<lambda>x. x \<equiv> COMBI"
  105.35  apply (rule eq_reflection)
  105.36  apply (rule ext) 
  105.37  apply (simp add: COMBI_def) 
  105.38  done
  105.39  
  105.40 -lemma abs_K [no_atp]: "\<lambda>x. y \<equiv> COMBK y"
  105.41 +lemma abs_K: "\<lambda>x. y \<equiv> COMBK y"
  105.42  apply (rule eq_reflection)
  105.43  apply (rule ext) 
  105.44  apply (simp add: COMBK_def) 
  105.45  done
  105.46  
  105.47 -lemma abs_B [no_atp]: "\<lambda>x. a (g x) \<equiv> COMBB a g"
  105.48 +lemma abs_B: "\<lambda>x. a (g x) \<equiv> COMBB a g"
  105.49  apply (rule eq_reflection)
  105.50  apply (rule ext) 
  105.51  apply (simp add: COMBB_def) 
  105.52  done
  105.53  
  105.54 -lemma abs_C [no_atp]: "\<lambda>x. (f x) b \<equiv> COMBC f b"
  105.55 +lemma abs_C: "\<lambda>x. (f x) b \<equiv> COMBC f b"
  105.56  apply (rule eq_reflection)
  105.57  apply (rule ext) 
  105.58  apply (simp add: COMBC_def) 
  105.59 @@ -180,7 +180,7 @@
  105.60  subsection {* Skolemization helpers *}
  105.61  
  105.62  definition skolem :: "'a \<Rightarrow> 'a" where
  105.63 -[no_atp]: "skolem = (\<lambda>x. x)"
  105.64 +"skolem = (\<lambda>x. x)"
  105.65  
  105.66  lemma skolem_COMBK_iff: "P \<longleftrightarrow> skolem (COMBK P (i\<Colon>nat))"
  105.67  unfolding skolem_def COMBK_def by (rule refl)
   106.1 --- a/src/HOL/Metis.thy	Mon Nov 11 17:34:44 2013 +0100
   106.2 +++ b/src/HOL/Metis.thy	Mon Nov 11 17:44:21 2013 +0100
   106.3 @@ -16,7 +16,7 @@
   106.4  subsection {* Literal selection and lambda-lifting helpers *}
   106.5  
   106.6  definition select :: "'a \<Rightarrow> 'a" where
   106.7 -[no_atp]: "select = (\<lambda>x. x)"
   106.8 +"select = (\<lambda>x. x)"
   106.9  
  106.10  lemma not_atomize: "(\<not> A \<Longrightarrow> False) \<equiv> Trueprop A"
  106.11  by (cut_tac atomize_not [of "\<not> A"]) simp
  106.12 @@ -30,7 +30,7 @@
  106.13  lemma select_FalseI: "False \<Longrightarrow> select False" by simp
  106.14  
  106.15  definition lambda :: "'a \<Rightarrow> 'a" where
  106.16 -[no_atp]: "lambda = (\<lambda>x. x)"
  106.17 +"lambda = (\<lambda>x. x)"
  106.18  
  106.19  lemma eq_lambdaI: "x \<equiv> y \<Longrightarrow> x \<equiv> lambda y"
  106.20  unfolding lambda_def by assumption
   107.1 --- a/src/HOL/Metis_Examples/Big_O.thy	Mon Nov 11 17:34:44 2013 +0100
   107.2 +++ b/src/HOL/Metis_Examples/Big_O.thy	Mon Nov 11 17:44:21 2013 +0100
   107.3 @@ -493,8 +493,10 @@
   107.4  
   107.5  lemma bigo_compose2:
   107.6  "f =o g +o O(h) \<Longrightarrow> (\<lambda>x. f(k x)) =o (\<lambda>x. g(k x)) +o O(\<lambda>x. h(k x))"
   107.7 -apply (simp only: set_minus_plus [symmetric] diff_minus fun_Compl_def func_plus)
   107.8 -by (erule bigo_compose1)
   107.9 +apply (simp only: set_minus_plus [symmetric] fun_Compl_def func_plus)
  107.10 +apply (drule bigo_compose1 [of "f - g" h k])
  107.11 +apply (simp add: fun_diff_def)
  107.12 +done
  107.13  
  107.14  subsection {* Setsum *}
  107.15  
   108.1 --- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Mon Nov 11 17:34:44 2013 +0100
   108.2 +++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Mon Nov 11 17:44:21 2013 +0100
   108.3 @@ -439,7 +439,7 @@
   108.4                    term_order |> the_default I)
   108.5              #> (Option.map (Config.put ATP_Systems.force_sos)
   108.6                    force_sos |> the_default I))
   108.7 -    val params as {max_facts, slice, ...} =
   108.8 +    val params as {max_facts, ...} =
   108.9        Sledgehammer_Isar.default_params ctxt
  108.10           ([("verbose", "true"),
  108.11             ("fact_filter", fact_filter),
  108.12 @@ -454,8 +454,7 @@
  108.13            |> sh_minimizeLST (*don't confuse the two minimization flags*)
  108.14            |> max_new_mono_instancesLST
  108.15            |> max_mono_itersLST)
  108.16 -    val default_max_facts =
  108.17 -      Sledgehammer_Provers.default_max_facts_of_prover ctxt slice prover_name
  108.18 +    val default_max_facts = Sledgehammer_Provers.default_max_facts_of_prover ctxt prover_name
  108.19      val is_appropriate_prop =
  108.20        Sledgehammer_Provers.is_appropriate_prop_of_prover ctxt prover_name
  108.21      val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal goal i ctxt
  108.22 @@ -494,7 +493,7 @@
  108.23                       |> Output.urgent_message)
  108.24          val prover = get_prover ctxt prover_name params goal facts
  108.25          val problem =
  108.26 -          {state = st', goal = goal, subgoal = i,
  108.27 +          {comment = "", state = st', goal = goal, subgoal = i,
  108.28             subgoal_count = Sledgehammer_Util.subgoal_count st, factss = factss}
  108.29        in prover params (K (K (K ""))) problem end)) ()
  108.30        handle TimeLimit.TimeOut => failed ATP_Proof.TimedOut
  108.31 @@ -584,6 +583,7 @@
  108.32          ({pre=st, log, ...}: Mirabelle.run_args) =
  108.33    let
  108.34      val ctxt = Proof.context_of st
  108.35 +    val {goal, ...} = Proof.goal st
  108.36      val n0 = length (these (!named_thms))
  108.37      val prover_name = get_prover_name ctxt args
  108.38      val type_enc = AList.lookup (op =) args type_encK |> the_default type_enc_default
  108.39 @@ -613,7 +613,7 @@
  108.40            true 1 (Sledgehammer_Util.subgoal_count st)
  108.41      val _ = log separator
  108.42      val (used_facts, (preplay, message, message_tail)) =
  108.43 -      minimize st NONE (these (!named_thms))
  108.44 +      minimize st goal NONE (these (!named_thms))
  108.45      val msg = message (Lazy.force preplay) ^ message_tail
  108.46    in
  108.47      case used_facts of
   109.1 --- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML	Mon Nov 11 17:34:44 2013 +0100
   109.2 +++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML	Mon Nov 11 17:44:21 2013 +0100
   109.3 @@ -110,18 +110,14 @@
   109.4         SOME proofs =>
   109.5         let
   109.6           val {context = ctxt, facts = chained_ths, goal} = Proof.goal pre
   109.7 -         val prover = AList.lookup (op =) args "prover"
   109.8 -                      |> the_default default_prover
   109.9 -         val params as {max_facts, slice, ...} =
  109.10 +         val prover = AList.lookup (op =) args "prover" |> the_default default_prover
  109.11 +         val params as {max_facts, ...} =
  109.12             Sledgehammer_Isar.default_params ctxt args
  109.13 -         val default_max_facts =
  109.14 -           Sledgehammer_Provers.default_max_facts_of_prover ctxt slice prover
  109.15 +         val default_max_facts = Sledgehammer_Provers.default_max_facts_of_prover ctxt prover
  109.16           val is_appropriate_prop =
  109.17 -           Sledgehammer_Provers.is_appropriate_prop_of_prover ctxt
  109.18 -               default_prover
  109.19 +           Sledgehammer_Provers.is_appropriate_prop_of_prover ctxt default_prover
  109.20           val relevance_fudge =
  109.21 -           extract_relevance_fudge args
  109.22 -               (Sledgehammer_Provers.relevance_fudge_of_prover ctxt prover)
  109.23 +           extract_relevance_fudge args Sledgehammer_MePo.default_relevance_fudge
  109.24           val subgoal = 1
  109.25           val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal goal subgoal ctxt
  109.26           val ho_atp = Sledgehammer_Provers.is_ho_atp ctxt prover
  109.27 @@ -132,9 +128,9 @@
  109.28                 Sledgehammer_Fact.no_fact_override reserved css_table chained_ths
  109.29                 hyp_ts concl_t
  109.30             |> filter (is_appropriate_prop o prop_of o snd)
  109.31 +           |> Sledgehammer_Fact.drop_duplicate_facts
  109.32             |> Sledgehammer_MePo.mepo_suggested_facts ctxt params
  109.33 -                  default_prover (the_default default_max_facts max_facts)
  109.34 -                  (SOME relevance_fudge) hyp_ts concl_t
  109.35 +                  (the_default default_max_facts max_facts) (SOME relevance_fudge) hyp_ts concl_t
  109.36              |> map (fst o fst)
  109.37           val (found_facts, lost_facts) =
  109.38             flat proofs |> sort_distinct string_ord
   110.1 --- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy	Mon Nov 11 17:34:44 2013 +0100
   110.2 +++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy	Mon Nov 11 17:44:21 2013 +0100
   110.3 @@ -1099,8 +1099,8 @@
   110.4    shows "(\<lambda>x. m *s x + c) o (\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) = id"
   110.5    "(\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) o (\<lambda>x. m *s x + c) = id"
   110.6    using m0
   110.7 -  apply (auto simp add: fun_eq_iff vector_add_ldistrib)
   110.8 -  apply (simp_all add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1[symmetric])
   110.9 +  apply (auto simp add: fun_eq_iff vector_add_ldistrib diff_conv_add_uminus simp del: add_uminus_conv_diff)
  110.10 +  apply (simp_all add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1 [symmetric])
  110.11    done
  110.12  
  110.13  lemma vector_affinity_eq:
  110.14 @@ -1114,7 +1114,7 @@
  110.15      using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
  110.16  next
  110.17    assume h: "x = inverse m *s y + - (inverse m *s c)"
  110.18 -  show "m *s x + c = y" unfolding h diff_minus[symmetric]
  110.19 +  show "m *s x + c = y" unfolding h
  110.20      using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
  110.21  qed
  110.22  
   111.1 --- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Mon Nov 11 17:34:44 2013 +0100
   111.2 +++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Mon Nov 11 17:44:21 2013 +0100
   111.3 @@ -858,9 +858,10 @@
   111.4    assumes "affine_parallel A B"
   111.5    shows "affine_parallel B A"
   111.6  proof -
   111.7 -  from assms obtain a where "B = (\<lambda>x. a + x) ` A"
   111.8 +  from assms obtain a where B: "B = (\<lambda>x. a + x) ` A"
   111.9      unfolding affine_parallel_def by auto
  111.10 -  then show ?thesis
  111.11 +  have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff)
  111.12 +  from B show ?thesis
  111.13      using translation_galois [of B a A]
  111.14      unfolding affine_parallel_def by auto
  111.15  qed
  111.16 @@ -980,6 +981,7 @@
  111.17    assumes "affine S" "a \<in> S"
  111.18    shows "subspace ((\<lambda>x. (-a)+x) ` S)"
  111.19  proof -
  111.20 +  have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff)
  111.21    have "affine ((\<lambda>x. (-a)+x) ` S)"
  111.22      using  affine_translation assms by auto
  111.23    moreover have "0 : ((\<lambda>x. (-a)+x) ` S)"
  111.24 @@ -992,15 +994,12 @@
  111.25    assumes "L \<equiv> {y. \<exists>x \<in> S. (-a)+x=y}"
  111.26    shows "subspace L & affine_parallel S L"
  111.27  proof -
  111.28 -  have par: "affine_parallel S L"
  111.29 -    unfolding affine_parallel_def using assms by auto
  111.30 +  from assms have "L = plus (- a) ` S" by auto
  111.31 +  then have par: "affine_parallel S L"
  111.32 +    unfolding affine_parallel_def .. 
  111.33    then have "affine L" using assms parallel_is_affine by auto
  111.34    moreover have "0 \<in> L"
  111.35 -    using assms
  111.36 -    apply auto
  111.37 -    using exI[of "(\<lambda>x. x:S \<and> -a+x=0)" a]
  111.38 -    apply auto
  111.39 -    done
  111.40 +    using assms by auto
  111.41    ultimately show ?thesis
  111.42      using subspace_affine par by auto
  111.43  qed
  111.44 @@ -2390,7 +2389,7 @@
  111.45    ultimately have h1: "affine hull ((\<lambda>x. a + x) `  S) \<subseteq> (\<lambda>x. a + x) ` (affine hull S)"
  111.46      by (metis hull_minimal)
  111.47    have "affine((\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) `  S)))"
  111.48 -    using affine_translation affine_affine_hull by auto
  111.49 +    using affine_translation affine_affine_hull by (auto simp del: uminus_add_conv_diff)
  111.50    moreover have "(\<lambda>x. -a + x) ` (\<lambda>x. a + x) `  S \<subseteq> (\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) `  S))"
  111.51      using hull_subset[of "(\<lambda>x. a + x) `  S"] by auto
  111.52    moreover have "S = (\<lambda>x. -a + x) ` (\<lambda>x. a + x) `  S"
  111.53 @@ -2478,7 +2477,7 @@
  111.54      using affine_dependent_translation_eq[of "(insert a S)" "-a"]
  111.55        affine_dependent_imp_dependent2 assms
  111.56        dependent_imp_affine_dependent[of a S]
  111.57 -    by auto
  111.58 +    by (auto simp del: uminus_add_conv_diff)
  111.59  qed
  111.60  
  111.61  lemma affine_dependent_iff_dependent2:
  111.62 @@ -2512,7 +2511,7 @@
  111.63      then have "insert 0 ((\<lambda>x. -a+x) ` (s - {a})) = (\<lambda>x. -a+x) ` s"
  111.64        by auto
  111.65      then have "span ((\<lambda>x. -a+x) ` (s - {a}))=span ((\<lambda>x. -a+x) ` s)"
  111.66 -      using span_insert_0[of "op + (- a) ` (s - {a})"] by auto
  111.67 +      using span_insert_0[of "op + (- a) ` (s - {a})"] by (auto simp del: uminus_add_conv_diff)
  111.68      moreover have "{x - a |x. x \<in> (s - {a})} = ((\<lambda>x. -a+x) ` (s - {a}))"
  111.69        by auto
  111.70      moreover have "insert a (s - {a}) = insert a s"
  111.71 @@ -2652,7 +2651,7 @@
  111.72      moreover have h1: "card ((\<lambda>x. -a + x) ` (B-{a})) = card (B-{a})"
  111.73         apply (rule card_image)
  111.74         using translate_inj_on
  111.75 -       apply auto
  111.76 +       apply (auto simp del: uminus_add_conv_diff)
  111.77         done
  111.78      ultimately have "card (B-{a}) > 0" by auto
  111.79      then have *: "finite (B - {a})"
  111.80 @@ -4507,38 +4506,30 @@
  111.81      apply (erule_tac x="x - y" in ballE)
  111.82      apply (auto simp add: inner_diff)
  111.83      done
  111.84 -  def k \<equiv> "Sup ((\<lambda>x. inner a x) ` t)"
  111.85 +  def k \<equiv> "SUP x:t. a \<bullet> x"
  111.86    show ?thesis
  111.87      apply (rule_tac x="-a" in exI)
  111.88      apply (rule_tac x="-(k + b / 2)" in exI)
  111.89 -    apply rule
  111.90 -    apply rule
  111.91 -    defer
  111.92 -    apply rule
  111.93 +    apply (intro conjI ballI)
  111.94      unfolding inner_minus_left and neg_less_iff_less
  111.95    proof -
  111.96 -    from ab have "((\<lambda>x. inner a x) ` t) *<= (inner a y - b)"
  111.97 -      apply (erule_tac x=y in ballE)
  111.98 -      apply (rule setleI)
  111.99 -      using `y \<in> s`
 111.100 -      apply auto
 111.101 -      done
 111.102 -    then have k: "isLub UNIV ((\<lambda>x. inner a x) ` t) k"
 111.103 +    fix x assume "x \<in> t"
 111.104 +    then have "inner a x - b / 2 < k"
 111.105        unfolding k_def
 111.106 -      apply (rule_tac isLub_cSup)
 111.107 -      using assms(5)
 111.108 -      apply auto
 111.109 -      done
 111.110 -    fix x
 111.111 -    assume "x \<in> t"
 111.112 -    then show "inner a x < (k + b / 2)"
 111.113 -      using `0<b` and isLubD2[OF k, of "inner a x"] by auto
 111.114 +    proof (subst less_cSUP_iff)
 111.115 +      show "t \<noteq> {}" by fact
 111.116 +      show "bdd_above (op \<bullet> a ` t)"
 111.117 +        using ab[rule_format, of y] `y \<in> s`
 111.118 +        by (intro bdd_aboveI2[where M="inner a y - b"]) (auto simp: field_simps intro: less_imp_le)
 111.119 +    qed (auto intro!: bexI[of _ x] `0<b`)
 111.120 +    then show "inner a x < k + b / 2"
 111.121 +      by auto
 111.122    next
 111.123      fix x
 111.124      assume "x \<in> s"
 111.125      then have "k \<le> inner a x - b"
 111.126        unfolding k_def
 111.127 -      apply (rule_tac cSup_least)
 111.128 +      apply (rule_tac cSUP_least)
 111.129        using assms(5)
 111.130        using ab[THEN bspec[where x=x]]
 111.131        apply auto
 111.132 @@ -4627,20 +4618,14 @@
 111.133    from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]]
 111.134    obtain a where "a \<noteq> 0" "\<forall>x\<in>{x - y |x y. x \<in> t \<and> y \<in> s}. 0 \<le> inner a x"
 111.135      using assms(3-5) by auto
 111.136 -  then have "\<forall>x\<in>t. \<forall>y\<in>s. inner a y \<le> inner a x"
 111.137 +  then have *: "\<And>x y. x \<in> t \<Longrightarrow> y \<in> s \<Longrightarrow> inner a y \<le> inner a x"
 111.138      by (force simp add: inner_diff)
 111.139 -  then show ?thesis
 111.140 -    apply (rule_tac x=a in exI)
 111.141 -    apply (rule_tac x="Sup ((\<lambda>x. inner a x) ` s)" in exI)
 111.142 +  then have bdd: "bdd_above ((op \<bullet> a)`s)"
 111.143 +    using `t \<noteq> {}` by (auto intro: bdd_aboveI2[OF *])
 111.144 +  show ?thesis
 111.145      using `a\<noteq>0`
 111.146 -    apply auto
 111.147 -    apply (rule isLub_cSup[THEN isLubD2])
 111.148 -    prefer 4
 111.149 -    apply (rule cSup_least)
 111.150 -    using assms(3-5)
 111.151 -    apply (auto simp add: setle_def)
 111.152 -    apply metis
 111.153 -    done
 111.154 +    by (intro exI[of _ a] exI[of _ "SUP x:s. a \<bullet> x"])
 111.155 +       (auto intro!: cSUP_upper bdd cSUP_least `a \<noteq> 0` `s \<noteq> {}` *)
 111.156  qed
 111.157  
 111.158  
 111.159 @@ -8725,7 +8710,7 @@
 111.160      using interior_subset[of I] `x \<in> interior I` by auto
 111.161  
 111.162    have "Inf (?F x) \<le> (f x - f y) / (x - y)"
 111.163 -  proof (rule cInf_lower2)
 111.164 +  proof (intro bdd_belowI cInf_lower2)
 111.165      show "(f x - f t) / (x - t) \<in> ?F x"
 111.166        using `t \<in> I` `x < t` by auto
 111.167      show "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)"
   112.1 --- a/src/HOL/Multivariate_Analysis/Derivative.thy	Mon Nov 11 17:34:44 2013 +0100
   112.2 +++ b/src/HOL/Multivariate_Analysis/Derivative.thy	Mon Nov 11 17:44:21 2013 +0100
   112.3 @@ -516,7 +516,7 @@
   112.4        unfolding e_def
   112.5        using c[THEN conjunct1]
   112.6        using norm_minus_cancel[of "f' i - f'' i"]
   112.7 -      by (auto simp add: add.commute ab_diff_minus)
   112.8 +      by auto
   112.9      finally show False
  112.10        using c
  112.11        using d[THEN conjunct2,rule_format,of "x + c *\<^sub>R i"]
   113.1 --- a/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Mon Nov 11 17:34:44 2013 +0100
   113.2 +++ b/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Mon Nov 11 17:44:21 2013 +0100
   113.3 @@ -660,7 +660,7 @@
   113.4      assume "S \<noteq> {}"
   113.5      { assume ex: "\<exists>B. \<forall>x\<in>S. B \<le> x"
   113.6        then have *: "\<forall>x\<in>S. Inf S \<le> x"
   113.7 -        using cInf_lower_EX[of _ S] ex by metis
   113.8 +        using cInf_lower[of _ S] ex by (metis bdd_below_def)
   113.9        then have "Inf S \<in> S"
  113.10          apply (subst closed_contains_Inf)
  113.11          using ex `S \<noteq> {}` `closed S`
  113.12 @@ -1193,12 +1193,12 @@
  113.13  qed
  113.14  
  113.15  lemma Liminf_at:
  113.16 -  fixes f :: "'a::metric_space \<Rightarrow> _"
  113.17 +  fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_lattice"
  113.18    shows "Liminf (at x) f = (SUP e:{0<..}. INF y:(ball x e - {x}). f y)"
  113.19    using Liminf_within[of x UNIV f] by simp
  113.20  
  113.21  lemma Limsup_at:
  113.22 -  fixes f :: "'a::metric_space \<Rightarrow> _"
  113.23 +  fixes f :: "'a::metric_space \<Rightarrow> 'b::complete_lattice"
  113.24    shows "Limsup (at x) f = (INF e:{0<..}. SUP y:(ball x e - {x}). f y)"
  113.25    using Limsup_within[of x UNIV f] by simp
  113.26  
  113.27 @@ -1209,7 +1209,7 @@
  113.28    apply (subst inf_commute)
  113.29    apply (subst SUP_inf)
  113.30    apply (intro SUP_cong[OF refl])
  113.31 -  apply (cut_tac A="ball x b - {x}" and B="{x}" and M=f in INF_union)
  113.32 +  apply (cut_tac A="ball x xa - {x}" and B="{x}" and M=f in INF_union)
  113.33    apply (simp add: INF_def del: inf_ereal_def)
  113.34    done
  113.35  
   114.1 --- a/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy	Mon Nov 11 17:34:44 2013 +0100
   114.2 +++ b/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy	Mon Nov 11 17:44:21 2013 +0100
   114.3 @@ -115,7 +115,7 @@
   114.4  instance vec :: (cancel_comm_monoid_add, finite) cancel_comm_monoid_add ..
   114.5  
   114.6  instance vec :: (group_add, finite) group_add
   114.7 -  by default (simp_all add: vec_eq_iff diff_minus)
   114.8 +  by default (simp_all add: vec_eq_iff)
   114.9  
  114.10  instance vec :: (ab_group_add, finite) ab_group_add
  114.11    by default (simp_all add: vec_eq_iff)
   115.1 --- a/src/HOL/Multivariate_Analysis/Integration.thy	Mon Nov 11 17:34:44 2013 +0100
   115.2 +++ b/src/HOL/Multivariate_Analysis/Integration.thy	Mon Nov 11 17:44:21 2013 +0100
   115.3 @@ -13,7 +13,7 @@
   115.4  lemma cSup_abs_le: (* TODO: is this really needed? *)
   115.5    fixes S :: "real set"
   115.6    shows "S \<noteq> {} \<Longrightarrow> (\<forall>x\<in>S. \<bar>x\<bar> \<le> a) \<Longrightarrow> \<bar>Sup S\<bar> \<le> a"
   115.7 -  by (auto simp add: abs_le_interval_iff intro: cSup_least) (metis cSup_upper2)
   115.8 +  by (auto simp add: abs_le_interval_iff intro: cSup_least) (metis cSup_upper2 bdd_aboveI)
   115.9  
  115.10  lemma cInf_abs_ge: (* TODO: is this really needed? *)
  115.11    fixes S :: "real set"
  115.12 @@ -28,10 +28,10 @@
  115.13  proof -
  115.14    have th: "\<And>(x::real) l e. \<bar>x - l\<bar> \<le> e \<longleftrightarrow> l - e \<le> x \<and> x \<le> l + e"
  115.15      by arith
  115.16 -  then show ?thesis
  115.17 -    using S b cSup_bounds[of S "l - e" "l+e"]
  115.18 -    unfolding th
  115.19 -    by (auto simp add: setge_def setle_def)
  115.20 +  have "bdd_above S"
  115.21 +    using b by (auto intro!: bdd_aboveI[of _ "l + e"])
  115.22 +  with S b show ?thesis
  115.23 +    unfolding th by (auto intro!: cSup_upper2 cSup_least)
  115.24  qed
  115.25  
  115.26  lemma cInf_asclose: (* TODO: is this really needed? *)
  115.27 @@ -44,9 +44,7 @@
  115.28      by auto
  115.29    also have "\<dots> \<le> e"
  115.30      apply (rule cSup_asclose)
  115.31 -    apply (auto simp add: S)
  115.32 -    apply (metis abs_minus_add_cancel b add_commute diff_minus)
  115.33 -    done
  115.34 +    using abs_minus_add_cancel b by (auto simp add: S)
  115.35    finally have "\<bar>- Sup (uminus ` S) - l\<bar> \<le> e" .
  115.36    then show ?thesis
  115.37      by (simp add: Inf_real_def)
  115.38 @@ -72,39 +70,6 @@
  115.39    shows "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> a \<ge> Inf S \<longleftrightarrow> (\<exists>x\<in>S. a \<ge> x)"
  115.40    by (metis cInf_eq_Min Min_le_iff)
  115.41  
  115.42 -lemma Inf: (* rename *)
  115.43 -  fixes S :: "real set"
  115.44 -  shows "S \<noteq> {} \<Longrightarrow> (\<exists>b. b <=* S) \<Longrightarrow> isGlb UNIV S (Inf S)"
  115.45 -  by (auto simp add: isLb_def setle_def setge_def isGlb_def greatestP_def
  115.46 -    intro: cInf_lower cInf_greatest)
  115.47 -
  115.48 -lemma real_le_inf_subset:
  115.49 -  assumes "t \<noteq> {}"
  115.50 -    and "t \<subseteq> s"
  115.51 -    and "\<exists>b. b <=* s"
  115.52 -  shows "Inf s \<le> Inf (t::real set)"
  115.53 -  apply (rule isGlb_le_isLb)
  115.54 -  apply (rule Inf[OF assms(1)])
  115.55 -  apply (insert assms)
  115.56 -  apply (erule exE)
  115.57 -  apply (rule_tac x = b in exI)
  115.58 -  apply (auto simp: isLb_def setge_def intro: cInf_lower cInf_greatest)
  115.59 -  done
  115.60 -
  115.61 -lemma real_ge_sup_subset:
  115.62 -  fixes t :: "real set"
  115.63 -  assumes "t \<noteq> {}"
  115.64 -    and "t \<subseteq> s"
  115.65 -    and "\<exists>b. s *<= b"
  115.66 -  shows "Sup s \<ge> Sup t"
  115.67 -  apply (rule isLub_le_isUb)
  115.68 -  apply (rule isLub_cSup[OF assms(1)])
  115.69 -  apply (insert assms)
  115.70 -  apply (erule exE)
  115.71 -  apply (rule_tac x = b in exI)
  115.72 -  apply (auto simp: isUb_def setle_def intro: cSup_upper cSup_least)
  115.73 -  done
  115.74 -
  115.75  (*declare not_less[simp] not_le[simp]*)
  115.76  
  115.77  lemmas scaleR_simps = scaleR_zero_left scaleR_minus_left scaleR_left_diff_distrib
  115.78 @@ -380,7 +345,7 @@
  115.79                  using Basis_le_norm[OF k, of "?z - y"] unfolding dist_norm by auto
  115.80                then have "y\<bullet>k < a\<bullet>k"
  115.81                  using e[THEN conjunct1] k
  115.82 -                by (auto simp add: field_simps as inner_Basis inner_simps)
  115.83 +                by (auto simp add: field_simps abs_less_iff as inner_Basis inner_simps)
  115.84                then have "y \<notin> i"
  115.85                  unfolding ab mem_interval by (auto intro!: bexI[OF _ k])
  115.86                then show False using yi by auto
  115.87 @@ -488,24 +453,24 @@
  115.88  subsection {* Bounds on intervals where they exist. *}
  115.89  
  115.90  definition interval_upperbound :: "('a::ordered_euclidean_space) set \<Rightarrow> 'a"
  115.91 -  where "interval_upperbound s = (\<Sum>i\<in>Basis. Sup {a. \<exists>x\<in>s. x\<bullet>i = a} *\<^sub>R i)"
  115.92 +  where "interval_upperbound s = (\<Sum>i\<in>Basis. (SUP x:s. x\<bullet>i) *\<^sub>R i)"
  115.93  
  115.94  definition interval_lowerbound :: "('a::ordered_euclidean_space) set \<Rightarrow> 'a"
  115.95 -  where "interval_lowerbound s = (\<Sum>i\<in>Basis. Inf {a. \<exists>x\<in>s. x\<bullet>i = a} *\<^sub>R i)"
  115.96 +  where "interval_lowerbound s = (\<Sum>i\<in>Basis. (INF x:s. x\<bullet>i) *\<^sub>R i)"
  115.97  
  115.98  lemma interval_upperbound[simp]:
  115.99    "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
 115.100      interval_upperbound {a..b} = (b::'a::ordered_euclidean_space)"
 115.101    unfolding interval_upperbound_def euclidean_representation_setsum
 115.102 -  by (auto simp del: ex_simps simp add: Bex_def ex_simps[symmetric] eucl_le[where 'a='a] setle_def
 115.103 -      intro!: cSup_unique)
 115.104 +  by (auto simp del: ex_simps simp add: Bex_def ex_simps[symmetric] eucl_le[where 'a='a] SUP_def
 115.105 +           intro!: cSup_eq)
 115.106  
 115.107  lemma interval_lowerbound[simp]:
 115.108    "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
 115.109      interval_lowerbound {a..b} = (a::'a::ordered_euclidean_space)"
 115.110    unfolding interval_lowerbound_def euclidean_representation_setsum
 115.111 -  by (auto simp del: ex_simps simp add: Bex_def ex_simps[symmetric] eucl_le[where 'a='a] setge_def
 115.112 -      intro!: cInf_unique)
 115.113 +  by (auto simp del: ex_simps simp add: Bex_def ex_simps[symmetric] eucl_le[where 'a='a] INF_def
 115.114 +           intro!: cInf_eq)
 115.115  
 115.116  lemmas interval_bounds = interval_upperbound interval_lowerbound
 115.117  
 115.118 @@ -6629,7 +6594,7 @@
 115.119  lemma interval_bound_sing[simp]:
 115.120    "interval_upperbound {a} = a"
 115.121    "interval_lowerbound {a} = a"
 115.122 -  unfolding interval_upperbound_def interval_lowerbound_def
 115.123 +  unfolding interval_upperbound_def interval_lowerbound_def SUP_def INF_def
 115.124    by (auto simp: euclidean_representation)
 115.125  
 115.126  lemma additive_tagged_division_1:
 115.127 @@ -11238,37 +11203,26 @@
 115.128  lemma bounded_variation_absolutely_integrable_interval:
 115.129    fixes f :: "'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
 115.130    assumes "f integrable_on {a..b}"
 115.131 -    and "\<forall>d. d division_of {a..b} \<longrightarrow> setsum (\<lambda>k. norm(integral k f)) d \<le> B"
 115.132 +    and *: "\<forall>d. d division_of {a..b} \<longrightarrow> setsum (\<lambda>k. norm(integral k f)) d \<le> B"
 115.133    shows "f absolutely_integrable_on {a..b}"
 115.134  proof -
 115.135 -  let ?S = "(\<lambda>d. setsum (\<lambda>k. norm(integral k f)) d) ` {d. d division_of {a..b} }"
 115.136 -  def i \<equiv> "Sup ?S"
 115.137 -  have i: "isLub UNIV ?S i"
 115.138 -    unfolding i_def
 115.139 -    apply (rule isLub_cSup)
 115.140 -    apply (rule elementary_interval)
 115.141 -    defer
 115.142 -    apply (rule_tac x=B in exI)
 115.143 -    apply (rule setleI)
 115.144 -    using assms(2)
 115.145 -    apply auto
 115.146 -    done
 115.147 +  let ?f = "\<lambda>d. \<Sum>k\<in>d. norm (integral k f)" and ?D = "{d. d division_of {a..b}}"
 115.148 +  have D_1: "?D \<noteq> {}"
 115.149 +    by (rule elementary_interval[of a b]) auto
 115.150 +  have D_2: "bdd_above (?f`?D)"
 115.151 +    by (metis * mem_Collect_eq bdd_aboveI2)
 115.152 +  note D = D_1 D_2
 115.153 +  let ?S = "SUP x:?D. ?f x"
 115.154    show ?thesis
 115.155      apply rule
 115.156      apply (rule assms)
 115.157      apply rule
 115.158 -    apply (subst has_integral[of _ i])
 115.159 +    apply (subst has_integral[of _ ?S])
 115.160    proof safe
 115.161      case goal1
 115.162 -    then have "i - e / 2 \<notin> Collect (isUb UNIV (setsum (\<lambda>k. norm (integral k f)) `
 115.163 -      {d. d division_of {a..b}}))"
 115.164 -      using isLub_ubs[OF i,rule_format]
 115.165 -      unfolding setge_def ubs_def
 115.166 -      by auto
 115.167 -    then have "\<exists>y. y division_of {a..b} \<and> i - e / 2 < (\<Sum>k\<in>y. norm (integral k f))"
 115.168 -      unfolding mem_Collect_eq isUb_def setle_def
 115.169 -      by (simp add: not_le)
 115.170 -    then guess d .. note d=conjunctD2[OF this]
 115.171 +    then have "?S - e / 2 < ?S" by simp
 115.172 +    then obtain d where d: "d division_of {a..b}" "?S - e / 2 < (\<Sum>k\<in>d. norm (integral k f))"
 115.173 +      unfolding less_cSUP_iff[OF D] by auto
 115.174      note d' = division_ofD[OF this(1)]
 115.175  
 115.176      have "\<forall>x. \<exists>e>0. \<forall>i\<in>d. x \<notin> i \<longrightarrow> ball x e \<inter> i = {}"
 115.177 @@ -11453,21 +11407,17 @@
 115.178          done
 115.179        note snd_p = division_ofD[OF division_of_tagged_division[OF p(1)]]
 115.180  
 115.181 -      have *: "\<And>sni sni' sf sf'. abs (sf' - sni') < e / 2 \<longrightarrow> i - e / 2 < sni \<and> sni' \<le> i \<and>
 115.182 -        sni \<le> sni' \<and> sf' = sf \<longrightarrow> abs (sf - i) < e"
 115.183 +      have *: "\<And>sni sni' sf sf'. abs (sf' - sni') < e / 2 \<longrightarrow> ?S - e / 2 < sni \<and> sni' \<le> ?S \<and>
 115.184 +        sni \<le> sni' \<and> sf' = sf \<longrightarrow> abs (sf - ?S) < e"
 115.185          by arith
 115.186 -      show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x)) - i) < e"
 115.187 +      show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x)) - ?S) < e"
 115.188          unfolding real_norm_def
 115.189          apply (rule *[rule_format,OF **])
 115.190          apply safe
 115.191          apply(rule d(2))
 115.192        proof -
 115.193 -        case goal1
 115.194 -        show ?case unfolding sum_p'
 115.195 -          apply (rule isLubD2[OF i])
 115.196 -          using division_of_tagged_division[OF p'']
 115.197 -          apply auto
 115.198 -          done
 115.199 +        case goal1 show ?case
 115.200 +          by (auto simp: sum_p' division_of_tagged_division[OF p''] D intro!: cSUP_upper)
 115.201        next
 115.202          case goal2
 115.203          have *: "{k \<inter> l | k l. k \<in> d \<and> l \<in> snd ` p} =
 115.204 @@ -11758,18 +11708,13 @@
 115.205      and "\<forall>d. d division_of (\<Union>d) \<longrightarrow> setsum (\<lambda>k. norm (integral k f)) d \<le> B"
 115.206    shows "f absolutely_integrable_on UNIV"
 115.207  proof (rule absolutely_integrable_onI, fact, rule)
 115.208 -  let ?S = "(\<lambda>d. setsum (\<lambda>k. norm(integral k f)) d) ` {d. d division_of  (\<Union>d)}"
 115.209 -  def i \<equiv> "Sup ?S"
 115.210 -  have i: "isLub UNIV ?S i"
 115.211 -    unfolding i_def
 115.212 -    apply (rule isLub_cSup)
 115.213 -    apply (rule elementary_interval)
 115.214 -    defer
 115.215 -    apply (rule_tac x=B in exI)
 115.216 -    apply (rule setleI)
 115.217 -    using assms(2)
 115.218 -    apply auto
 115.219 -    done
 115.220 +  let ?f = "\<lambda>d. \<Sum>k\<in>d. norm (integral k f)" and ?D = "{d. d division_of  (\<Union>d)}"
 115.221 +  have D_1: "?D \<noteq> {}"
 115.222 +    by (rule elementary_interval) auto
 115.223 +  have D_2: "bdd_above (?f`?D)"
 115.224 +    by (intro bdd_aboveI2[where M=B] assms(2)[rule_format]) simp
 115.225 +  note D = D_1 D_2
 115.226 +  let ?S = "SUP d:?D. ?f d"
 115.227    have f_int: "\<And>a b. f absolutely_integrable_on {a..b}"
 115.228      apply (rule bounded_variation_absolutely_integrable_interval[where B=B])
 115.229      apply (rule integrable_on_subinterval[OF assms(1)])
 115.230 @@ -11778,7 +11723,7 @@
 115.231      apply (rule assms(2)[rule_format])
 115.232      apply auto
 115.233      done
 115.234 -  show "((\<lambda>x. norm (f x)) has_integral i) UNIV"
 115.235 +  show "((\<lambda>x. norm (f x)) has_integral ?S) UNIV"
 115.236      apply (subst has_integral_alt')
 115.237      apply safe
 115.238    proof -
 115.239 @@ -11787,16 +11732,11 @@
 115.240        using f_int[of a b] by auto
 115.241    next
 115.242      case goal2
 115.243 -    have "\<exists>y\<in>setsum (\<lambda>k. norm (integral k f)) ` {d. d division_of \<Union>d}. \<not> y \<le> i - e"
 115.244 +    have "\<exists>y\<in>setsum (\<lambda>k. norm (integral k f)) ` {d. d division_of \<Union>d}. \<not> y \<le> ?S - e"
 115.245      proof (rule ccontr)
 115.246        assume "\<not> ?thesis"
 115.247 -      then have "i \<le> i - e"
 115.248 -        apply -
 115.249 -        apply (rule isLub_le_isUb[OF i])
 115.250 -        apply (rule isUbI)
 115.251 -        unfolding setle_def
 115.252 -        apply auto
 115.253 -        done
 115.254 +      then have "?S \<le> ?S - e"
 115.255 +        by (intro cSUP_least[OF D(1)]) auto
 115.256        then show False
 115.257          using goal2 by auto
 115.258      qed
 115.259 @@ -11813,9 +11753,9 @@
 115.260      proof -
 115.261        fix a b :: 'n
 115.262        assume ab: "ball 0 (K + 1) \<subseteq> {a..b}"
 115.263 -      have *: "\<forall>s s1. i - e < s1 \<and> s1 \<le> s \<and> s < i + e \<longrightarrow> abs (s - i) < e"
 115.264 +      have *: "\<forall>s s1. ?S - e < s1 \<and> s1 \<le> s \<and> s < ?S + e \<longrightarrow> abs (s - ?S) < e"
 115.265          by arith
 115.266 -      show "norm (integral {a..b} (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) - i) < e"
 115.267 +      show "norm (integral {a..b} (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) - ?S) < e"
 115.268          unfolding real_norm_def
 115.269          apply (rule *[rule_format])
 115.270          apply safe
 115.271 @@ -11867,10 +11807,10 @@
 115.272          from henstock_lemma[OF f(1) `e/2>0`] guess d2 . note d2=this
 115.273          from fine_division_exists[OF gauge_inter[OF d1(1) d2(1)], of a b] guess p .
 115.274          note p=this(1) conjunctD2[OF this(2)[unfolded fine_inter]]
 115.275 -        have *: "\<And>sf sf' si di. sf' = sf \<longrightarrow> si \<le> i \<longrightarrow> abs (sf - si) < e / 2 \<longrightarrow>
 115.276 -          abs (sf' - di) < e / 2 \<longrightarrow> di < i + e"
 115.277 +        have *: "\<And>sf sf' si di. sf' = sf \<longrightarrow> si \<le> ?S \<longrightarrow> abs (sf - si) < e / 2 \<longrightarrow>
 115.278 +          abs (sf' - di) < e / 2 \<longrightarrow> di < ?S + e"
 115.279            by arith
 115.280 -        show "integral {a..b} (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) < i + e"
 115.281 +        show "integral {a..b} (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) < ?S + e"
 115.282            apply (subst if_P)
 115.283            apply rule
 115.284          proof (rule *[rule_format])
 115.285 @@ -11893,18 +11833,12 @@
 115.286              apply (subst abs_of_nonneg)
 115.287              apply auto
 115.288              done
 115.289 -          show "(\<Sum>(x, k)\<in>p. norm (integral k f)) \<le> i"
 115.290 +          show "(\<Sum>(x, k)\<in>p. norm (integral k f)) \<le> ?S"
 115.291 +            using partial_division_of_tagged_division[of p "{a..b}"] p(1)
 115.292              apply (subst setsum_over_tagged_division_lemma[OF p(1)])
 115.293 -            defer
 115.294 -            apply (rule isLubD2[OF i])
 115.295 -            unfolding image_iff
 115.296 -            apply (rule_tac x="snd ` p" in bexI)
 115.297 -            unfolding mem_Collect_eq
 115.298 -            defer
 115.299 -            apply (rule partial_division_of_tagged_division[of _ "{a..b}"])
 115.300 -            using p(1)
 115.301 -            unfolding tagged_division_of_def
 115.302 -            apply auto
 115.303 +            apply (simp add: integral_null)
 115.304 +            apply (intro cSUP_upper2[OF D(2), of "snd ` p"])
 115.305 +            apply (auto simp: tagged_partial_division_of_def)
 115.306              done
 115.307          qed
 115.308        qed
 115.309 @@ -11975,7 +11909,7 @@
 115.310      and "g absolutely_integrable_on s"
 115.311    shows "(\<lambda>x. f x - g x) absolutely_integrable_on s"
 115.312    using absolutely_integrable_add[OF assms(1) absolutely_integrable_neg[OF assms(2)]]
 115.313 -  by (simp only: algebra_simps)
 115.314 +  by (simp add: algebra_simps)
 115.315  
 115.316  lemma absolutely_integrable_linear:
 115.317    fixes f :: "'m::ordered_euclidean_space \<Rightarrow> 'n::ordered_euclidean_space"
 115.318 @@ -12380,11 +12314,22 @@
 115.319  lemma dominated_convergence:
 115.320    fixes f :: "nat \<Rightarrow> 'n::ordered_euclidean_space \<Rightarrow> real"
 115.321    assumes "\<And>k. (f k) integrable_on s" "h integrable_on s"
 115.322 -    and "\<And>k. \<forall>x \<in> s. norm(f k x) \<le> (h x)"
 115.323 +    and "\<And>k. \<forall>x \<in> s. norm (f k x) \<le> h x"
 115.324      and "\<forall>x \<in> s. ((\<lambda>k. f k x) ---> g x) sequentially"
 115.325    shows "g integrable_on s"
 115.326      and "((\<lambda>k. integral s (f k)) ---> integral s g) sequentially"
 115.327  proof -
 115.328 +  have bdd_below[simp]: "\<And>x P. x \<in> s \<Longrightarrow> bdd_below {f n x |n. P n}"
 115.329 +  proof (safe intro!: bdd_belowI)
 115.330 +    fix n x show "x \<in> s \<Longrightarrow> - h x \<le> f n x"
 115.331 +      using assms(3)[rule_format, of x n] by simp
 115.332 +  qed
 115.333 +  have bdd_above[simp]: "\<And>x P. x \<in> s \<Longrightarrow> bdd_above {f n x |n. P n}"
 115.334 +  proof (safe intro!: bdd_aboveI)
 115.335 +    fix n x show "x \<in> s \<Longrightarrow> f n x \<le> h x"
 115.336 +      using assms(3)[rule_format, of x n] by simp
 115.337 +  qed
 115.338 +
 115.339    have "\<And>m. (\<lambda>x. Inf {f j x |j. m \<le> j}) integrable_on s \<and>
 115.340      ((\<lambda>k. integral s (\<lambda>x. Inf {f j x |j. j \<in> {m..m + k}})) --->
 115.341      integral s (\<lambda>x. Inf {f j x |j. m \<le> j}))sequentially"
 115.342 @@ -12424,66 +12369,32 @@
 115.343      fix x
 115.344      assume x: "x \<in> s"
 115.345      show "Inf {f j x |j. j \<in> {m..m + Suc k}} \<le> Inf {f j x |j. j \<in> {m..m + k}}"
 115.346 -      apply (rule cInf_ge)
 115.347 -      unfolding setge_def
 115.348 -      defer
 115.349 -      apply rule
 115.350 -      apply (subst cInf_finite_le_iff)
 115.351 -      prefer 3
 115.352 -      apply (rule_tac x=xa in bexI)
 115.353 -      apply auto
 115.354 -      done
 115.355 -    let ?S = "{f j x| j.  m \<le> j}"
 115.356 -    def i \<equiv> "Inf ?S"
 115.357 -    show "((\<lambda>k. Inf {f j x |j. j \<in> {m..m + k}}) ---> i) sequentially"
 115.358 +      by (rule cInf_superset_mono) auto
 115.359 +    let ?S = "{f j x| j. m \<le> j}"
 115.360 +    show "((\<lambda>k. Inf {f j x |j. j \<in> {m..m + k}}) ---> Inf ?S) sequentially"
 115.361      proof (rule LIMSEQ_I)
 115.362        case goal1
 115.363        note r = this
 115.364 -      have i: "isGlb UNIV ?S i"
 115.365 -        unfolding i_def
 115.366 -        apply (rule Inf)
 115.367 -        defer
 115.368 -        apply (rule_tac x="- h x - 1" in exI)
 115.369 -        unfolding setge_def
 115.370 -      proof safe
 115.371 -        case goal1
 115.372 -        then show ?case using assms(3)[rule_format,OF x, of j] by auto
 115.373 -      qed auto
 115.374 -
 115.375 -      have "\<exists>y\<in>?S. \<not> y \<ge> i + r"
 115.376 -      proof(rule ccontr)
 115.377 -        assume "\<not> ?thesis"
 115.378 -        then have "i \<ge> i + r"
 115.379 -          apply -
 115.380 -          apply (rule isGlb_le_isLb[OF i])
 115.381 -          apply (rule isLbI)
 115.382 -          unfolding setge_def
 115.383 -          apply fastforce+
 115.384 -          done
 115.385 -        then show False using r by auto
 115.386 -      qed
 115.387 -      then guess y .. note y=this[unfolded not_le]
 115.388 -      from this(1)[unfolded mem_Collect_eq] guess N .. note N=conjunctD2[OF this]
 115.389 +
 115.390 +      have "\<exists>y\<in>?S. y < Inf ?S + r"
 115.391 +        by (subst cInf_less_iff[symmetric]) (auto simp: `x\<in>s` r)
 115.392 +      then obtain N where N: "f N x < Inf ?S + r" "m \<le> N"
 115.393 +        by blast
 115.394  
 115.395        show ?case
 115.396          apply (rule_tac x=N in exI)
 115.397        proof safe
 115.398          case goal1
 115.399 -        have *: "\<And>y ix. y < i + r \<longrightarrow> i \<le> ix \<longrightarrow> ix \<le> y \<longrightarrow> abs(ix - i) < r"
 115.400 +        have *: "\<And>y ix. y < Inf ?S + r \<longrightarrow> Inf ?S \<le> ix \<longrightarrow> ix \<le> y \<longrightarrow> abs(ix - Inf ?S) < r"
 115.401            by arith
 115.402          show ?case
 115.403            unfolding real_norm_def
 115.404 -            apply (rule *[rule_format,OF y(2)])
 115.405 -            unfolding i_def
 115.406 -            apply (rule real_le_inf_subset)
 115.407 -            prefer 3
 115.408 -            apply (rule,rule isGlbD1[OF i])
 115.409 -            prefer 3
 115.410 -            apply (subst cInf_finite_le_iff)
 115.411 -            prefer 3
 115.412 -            apply (rule_tac x=y in bexI)
 115.413 +            apply (rule *[rule_format, OF N(1)])
 115.414 +            apply (rule cInf_superset_mono, auto simp: `x\<in>s`) []
 115.415 +            apply (rule cInf_lower)
 115.416              using N goal1
 115.417 -            apply auto
 115.418 +            apply auto []
 115.419 +            apply simp
 115.420              done
 115.421        qed
 115.422      qed
 115.423 @@ -12527,65 +12438,27 @@
 115.424      fix x
 115.425      assume x: "x\<in>s"
 115.426      show "Sup {f j x |j. j \<in> {m..m + Suc k}} \<ge> Sup {f j x |j. j \<in> {m..m + k}}"
 115.427 -      apply (rule cSup_le)
 115.428 -      unfolding setle_def
 115.429 -      defer
 115.430 -      apply rule
 115.431 -      apply (subst cSup_finite_ge_iff)
 115.432 -      prefer 3
 115.433 -      apply (rule_tac x=y in bexI)
 115.434 -      apply auto
 115.435 -      done
 115.436 -    let ?S = "{f j x| j.  m \<le> j}"
 115.437 -    def i \<equiv> "Sup ?S"
 115.438 -    show "((\<lambda>k. Sup {f j x |j. j \<in> {m..m + k}}) ---> i) sequentially"
 115.439 +      by (rule cSup_subset_mono) auto
 115.440 +    let ?S = "{f j x| j. m \<le> j}"
 115.441 +    show "((\<lambda>k. Sup {f j x |j. j \<in> {m..m + k}}) ---> Sup ?S) sequentially"
 115.442      proof (rule LIMSEQ_I)
 115.443        case goal1 note r=this
 115.444 -      have i: "isLub UNIV ?S i"
 115.445 -        unfolding i_def
 115.446 -        apply (rule isLub_cSup)
 115.447 -        defer
 115.448 -        apply (rule_tac x="h x" in exI)
 115.449 -        unfolding setle_def
 115.450 -      proof safe
 115.451 -        case goal1
 115.452 -        then show ?case
 115.453 -          using assms(3)[rule_format,OF x, of j] by auto
 115.454 -      qed auto
 115.455 -
 115.456 -      have "\<exists>y\<in>?S. \<not> y \<le> i - r"
 115.457 -      proof (rule ccontr)
 115.458 -        assume "\<not> ?thesis"
 115.459 -        then have "i \<le> i - r"
 115.460 -          apply -
 115.461 -          apply (rule isLub_le_isUb[OF i])
 115.462 -          apply (rule isUbI)
 115.463 -          unfolding setle_def
 115.464 -          apply fastforce+
 115.465 -          done
 115.466 -        then show False
 115.467 -          using r by auto
 115.468 -      qed
 115.469 -      then guess y .. note y=this[unfolded not_le]
 115.470 -      from this(1)[unfolded mem_Collect_eq] guess N .. note N=conjunctD2[OF this]
 115.471 +      have "\<exists>y\<in>?S. Sup ?S - r < y"
 115.472 +        by (subst less_cSup_iff[symmetric]) (auto simp: r `x\<in>s`)
 115.473 +      then obtain N where N: "Sup ?S - r < f N x" "m \<le> N"
 115.474 +        by blast
 115.475  
 115.476        show ?case
 115.477          apply (rule_tac x=N in exI)
 115.478        proof safe
 115.479          case goal1
 115.480 -        have *: "\<And>y ix. i - r < y \<longrightarrow> ix \<le> i \<longrightarrow> y \<le> ix \<longrightarrow> abs(ix - i) < r"
 115.481 +        have *: "\<And>y ix. Sup ?S - r < y \<longrightarrow> ix \<le> Sup ?S \<longrightarrow> y \<le> ix \<longrightarrow> abs(ix - Sup ?S) < r"
 115.482            by arith
 115.483          show ?case
 115.484 -          unfolding real_norm_def
 115.485 -          apply (rule *[rule_format,OF y(2)])
 115.486 -          unfolding i_def
 115.487 -          apply (rule real_ge_sup_subset)
 115.488 -          prefer 3
 115.489 -          apply (rule, rule isLubD1[OF i])
 115.490 -          prefer 3
 115.491 -          apply (subst cSup_finite_ge_iff)
 115.492 -          prefer 3
 115.493 -          apply (rule_tac x = y in bexI)
 115.494 +          apply simp
 115.495 +          apply (rule *[rule_format, OF N(1)])
 115.496 +          apply (rule cSup_subset_mono, auto simp: `x\<in>s`) []
 115.497 +          apply (subst cSup_upper)
 115.498            using N goal1
 115.499            apply auto
 115.500            done
 115.501 @@ -12618,17 +12491,7 @@
 115.502  
 115.503      have *: "\<And>x y::real. x \<ge> - y \<Longrightarrow> - x \<le> y" by auto
 115.504      show "Inf {f j x |j. k \<le> j} \<le> Inf {f j x |j. Suc k \<le> j}"
 115.505 -      apply -
 115.506 -      apply (rule real_le_inf_subset)
 115.507 -      prefer 3
 115.508 -      unfolding setge_def
 115.509 -      apply (rule_tac x="- h x" in exI)
 115.510 -      apply safe
 115.511 -      apply (rule *)
 115.512 -      using assms(3)[rule_format,OF x]
 115.513 -      unfolding real_norm_def abs_le_iff
 115.514 -      apply auto
 115.515 -      done
 115.516 +      by (intro cInf_superset_mono) (auto simp: `x\<in>s`)
 115.517  
 115.518      show "(\<lambda>k::nat. Inf {f j x |j. k \<le> j}) ----> g x"
 115.519      proof (rule LIMSEQ_I)
 115.520 @@ -12676,16 +12539,7 @@
 115.521      assume x: "x \<in> s"
 115.522  
 115.523      show "Sup {f j x |j. k \<le> j} \<ge> Sup {f j x |j. Suc k \<le> j}"
 115.524 -      apply -
 115.525 -      apply (rule real_ge_sup_subset)
 115.526 -      prefer 3
 115.527 -      unfolding setle_def
 115.528 -      apply (rule_tac x="h x" in exI)
 115.529 -      apply safe
 115.530 -      using assms(3)[rule_format,OF x]
 115.531 -      unfolding real_norm_def abs_le_iff
 115.532 -      apply auto
 115.533 -      done
 115.534 +      by (rule cSup_subset_mono) (auto simp: `x\<in>s`)
 115.535      show "((\<lambda>k. Sup {f j x |j. k \<le> j}) ---> g x) sequentially"
 115.536      proof (rule LIMSEQ_I)
 115.537        case goal1
 115.538 @@ -12714,42 +12568,18 @@
 115.539      from LIMSEQ_D [OF inc2(2) goal1] guess N1 .. note N1=this[unfolded real_norm_def]
 115.540      from LIMSEQ_D [OF dec2(2) goal1] guess N2 .. note N2=this[unfolded real_norm_def]
 115.541      show ?case
 115.542 -      apply (rule_tac x="N1+N2" in exI, safe)
 115.543 -    proof -
 115.544 +    proof (rule_tac x="N1+N2" in exI, safe)
 115.545        fix n
 115.546        assume n: "n \<ge> N1 + N2"
 115.547        have *: "\<And>i0 i i1 g. \<bar>i0 - g\<bar> < r \<longrightarrow> \<bar>i1 - g\<bar> < r \<longrightarrow> i0 \<le> i \<longrightarrow> i \<le> i1 \<longrightarrow> \<bar>i - g\<bar> < r"
 115.548          by arith
 115.549        show "norm (integral s (f n) - integral s g) < r"
 115.550          unfolding real_norm_def
 115.551 -        apply (rule *[rule_format,OF N1[rule_format] N2[rule_format], of n n])
 115.552 -      proof -
 115.553 +      proof (rule *[rule_format,OF N1[rule_format] N2[rule_format], of n n])
 115.554          show "integral s (\<lambda>x. Inf {f j x |j. n \<le> j}) \<le> integral s (f n)"
 115.555 -        proof (rule integral_le[OF dec1(1) assms(1)], safe)
 115.556 -          fix x
 115.557 -          assume x: "x \<in> s"
 115.558 -          have *: "\<And>x y::real. x \<ge> - y \<Longrightarrow> - x \<le> y" by auto
 115.559 -          show "Inf {f j x |j. n \<le> j} \<le> f n x"
 115.560 -            apply (rule cInf_lower[where z="- h x"])
 115.561 -            defer
 115.562 -            apply (rule *)
 115.563 -            using assms(3)[rule_format,OF x]
 115.564 -            unfolding real_norm_def abs_le_iff
 115.565 -            apply auto
 115.566 -            done
 115.567 -        qed
 115.568 +          by (rule integral_le[OF dec1(1) assms(1)]) (auto intro!: cInf_lower)
 115.569          show "integral s (f n) \<le> integral s (\<lambda>x. Sup {f j x |j. n \<le> j})"
 115.570 -        proof (rule integral_le[OF assms(1) inc1(1)], safe)
 115.571 -          fix x
 115.572 -          assume x: "x \<in> s"
 115.573 -          show "f n x \<le> Sup {f j x |j. n \<le> j}"
 115.574 -            apply (rule cSup_upper[where z="h x"])
 115.575 -            defer
 115.576 -            using assms(3)[rule_format,OF x]
 115.577 -            unfolding real_norm_def abs_le_iff
 115.578 -            apply auto
 115.579 -            done
 115.580 -        qed
 115.581 +          by (rule integral_le[OF assms(1) inc1(1)]) (auto intro!: cSup_upper)
 115.582        qed (insert n, auto)
 115.583      qed
 115.584    qed
   116.1 --- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy	Mon Nov 11 17:34:44 2013 +0100
   116.2 +++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy	Mon Nov 11 17:44:21 2013 +0100
   116.3 @@ -303,7 +303,7 @@
   116.4    by (metis linear_iff)
   116.5  
   116.6  lemma linear_sub: "linear f \<Longrightarrow> f (x - y) = f x - f y"
   116.7 -  by (simp add: diff_minus linear_add linear_neg)
   116.8 +  using linear_add [of f x "- y"] by (simp add: linear_neg)
   116.9  
  116.10  lemma linear_setsum:
  116.11    assumes lin: "linear f"
  116.12 @@ -384,10 +384,10 @@
  116.13    using bilinear_radd [OF assms, of x 0 0 ] by (simp add: eq_add_iff field_simps)
  116.14  
  116.15  lemma bilinear_lsub: "bilinear h \<Longrightarrow> h (x - y) z = h x z - h y z"
  116.16 -  by (simp  add: diff_minus bilinear_ladd bilinear_lneg)
  116.17 +  using bilinear_ladd [of h x "- y"] by (simp add: bilinear_lneg)
  116.18  
  116.19  lemma bilinear_rsub: "bilinear h \<Longrightarrow> h z (x - y) = h z x - h z y"
  116.20 -  by (simp  add: diff_minus bilinear_radd bilinear_rneg)
  116.21 +  using bilinear_radd [of h _ x "- y"] by (simp add: bilinear_rneg)
  116.22  
  116.23  lemma bilinear_setsum:
  116.24    assumes bh: "bilinear h"
  116.25 @@ -730,7 +730,7 @@
  116.26    by (metis scaleR_minus1_left subspace_mul)
  116.27  
  116.28  lemma subspace_sub: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x - y \<in> S"
  116.29 -  by (metis diff_minus subspace_add subspace_neg)
  116.30 +  using subspace_add [of S x "- y"] by (simp add: subspace_neg)
  116.31  
  116.32  lemma (in real_vector) subspace_setsum:
  116.33    assumes sA: "subspace A"
  116.34 @@ -1021,8 +1021,7 @@
  116.35      apply safe
  116.36      apply (rule_tac x=k in exI, simp)
  116.37      apply (erule rev_image_eqI [OF SigmaI [OF rangeI]])
  116.38 -    apply simp
  116.39 -    apply (rule right_minus)
  116.40 +    apply auto
  116.41      done
  116.42    then show ?thesis by simp
  116.43  qed
  116.44 @@ -2064,7 +2063,7 @@
  116.45        using C by simp
  116.46      have "orthogonal ?a y"
  116.47        unfolding orthogonal_def
  116.48 -      unfolding inner_diff inner_setsum_left diff_eq_0_iff_eq
  116.49 +      unfolding inner_diff inner_setsum_left right_minus_eq
  116.50        unfolding setsum_diff1' [OF `finite C` `y \<in> C`]
  116.51        apply (clarsimp simp add: inner_commute[of y a])
  116.52        apply (rule setsum_0')
  116.53 @@ -3026,7 +3025,7 @@
  116.54          norm x * (norm y * (y \<bullet> x) - norm x * norm y * norm y) =  0)"
  116.55        using x y
  116.56        unfolding inner_simps
  116.57 -      unfolding power2_norm_eq_inner[symmetric] power2_eq_square diff_eq_0_iff_eq
  116.58 +      unfolding power2_norm_eq_inner[symmetric] power2_eq_square right_minus_eq
  116.59        apply (simp add: inner_commute)
  116.60        apply (simp add: field_simps)
  116.61        apply metis
   117.1 --- a/src/HOL/Multivariate_Analysis/Operator_Norm.thy	Mon Nov 11 17:34:44 2013 +0100
   117.2 +++ b/src/HOL/Multivariate_Analysis/Operator_Norm.thy	Mon Nov 11 17:44:21 2013 +0100
   117.3 @@ -8,7 +8,7 @@
   117.4  imports Linear_Algebra
   117.5  begin
   117.6  
   117.7 -definition "onorm f = Sup {norm (f x)| x. norm x = 1}"
   117.8 +definition "onorm f = (SUP x:{x. norm x = 1}. norm (f x))"
   117.9  
  117.10  lemma norm_bound_generalize:
  117.11    fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
  117.12 @@ -67,25 +67,22 @@
  117.13    shows "norm (f x) \<le> onorm f * norm x"
  117.14      and "\<forall>x. norm (f x) \<le> b * norm x \<Longrightarrow> onorm f \<le> b"
  117.15  proof -
  117.16 -  let ?S = "{norm (f x) |x. norm x = 1}"
  117.17 +  let ?S = "(\<lambda>x. norm (f x))`{x. norm x = 1}"
  117.18    have "norm (f (SOME i. i \<in> Basis)) \<in> ?S"
  117.19      by (auto intro!: exI[of _ "SOME i. i \<in> Basis"] norm_Basis SOME_Basis)
  117.20    then have Se: "?S \<noteq> {}"
  117.21      by auto
  117.22 -  from linear_bounded[OF lf] have b: "\<exists> b. ?S *<= b"
  117.23 -    unfolding norm_bound_generalize[OF lf, symmetric]
  117.24 -    by (auto simp add: setle_def)
  117.25 -  from isLub_cSup[OF Se b, unfolded onorm_def[symmetric]]
  117.26 -  show "norm (f x) \<le> onorm f * norm x"
  117.27 +  from linear_bounded[OF lf] have b: "bdd_above ?S"
  117.28 +    unfolding norm_bound_generalize[OF lf, symmetric] by auto
  117.29 +  then show "norm (f x) \<le> onorm f * norm x"
  117.30      apply -
  117.31      apply (rule spec[where x = x])
  117.32      unfolding norm_bound_generalize[OF lf, symmetric]
  117.33 -    apply (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)
  117.34 +    apply (auto simp: onorm_def intro!: cSUP_upper)
  117.35      done
  117.36    show "\<forall>x. norm (f x) \<le> b * norm x \<Longrightarrow> onorm f \<le> b"
  117.37 -    using isLub_cSup[OF Se b, unfolded onorm_def[symmetric]]
  117.38      unfolding norm_bound_generalize[OF lf, symmetric]
  117.39 -    by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)
  117.40 +    using Se by (auto simp: onorm_def intro!: cSUP_least b)
  117.41  qed
  117.42  
  117.43  lemma onorm_pos_le:
  117.44 @@ -107,18 +104,8 @@
  117.45    apply arith
  117.46    done
  117.47  
  117.48 -lemma onorm_const:
  117.49 -  "onorm (\<lambda>x::'a::euclidean_space. y::'b::euclidean_space) = norm y"
  117.50 -proof -
  117.51 -  let ?f = "\<lambda>x::'a. y::'b"
  117.52 -  have th: "{norm (?f x)| x. norm x = 1} = {norm y}"
  117.53 -    by (auto simp: SOME_Basis intro!: exI[of _ "SOME i. i \<in> Basis"])
  117.54 -  show ?thesis
  117.55 -    unfolding onorm_def th
  117.56 -    apply (rule cSup_unique)
  117.57 -    apply (simp_all  add: setle_def)
  117.58 -    done
  117.59 -qed
  117.60 +lemma onorm_const: "onorm (\<lambda>x::'a::euclidean_space. y::'b::euclidean_space) = norm y"
  117.61 +  using SOME_Basis by (auto simp add: onorm_def intro!: cSUP_const norm_Basis)
  117.62  
  117.63  lemma onorm_pos_lt:
  117.64    fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   118.1 --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Mon Nov 11 17:34:44 2013 +0100
   118.2 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Mon Nov 11 17:44:21 2013 +0100
   118.3 @@ -10,7 +10,6 @@
   118.4  imports
   118.5    Complex_Main
   118.6    "~~/src/HOL/Library/Countable_Set"
   118.7 -  "~~/src/HOL/Library/Glbs"
   118.8    "~~/src/HOL/Library/FuncSet"
   118.9    Linear_Algebra
  118.10    Norm_Arith
  118.11 @@ -28,8 +27,6 @@
  118.12  lemma lim_subseq: "subseq r \<Longrightarrow> s ----> l \<Longrightarrow> (s \<circ> r) ----> l"
  118.13    by (rule LIMSEQ_subseq_LIMSEQ)
  118.14  
  118.15 -lemmas real_isGlb_unique = isGlb_unique[where 'a=real]
  118.16 -
  118.17  lemma countable_PiE:
  118.18    "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> countable (F i)) \<Longrightarrow> countable (PiE I F)"
  118.19    by (induct I arbitrary: F rule: finite_induct) (auto simp: PiE_insert_eq)
  118.20 @@ -789,18 +786,20 @@
  118.21    "a - b \<ge> c \<longleftrightarrow> a \<ge> c + b"
  118.22    by arith+
  118.23  
  118.24 -lemma open_ball[intro, simp]: "open (ball x e)"
  118.25 -  unfolding open_dist ball_def mem_Collect_eq Ball_def
  118.26 -  unfolding dist_commute
  118.27 -  apply clarify
  118.28 -  apply (rule_tac x="e - dist xa x" in exI)
  118.29 -  using dist_triangle_alt[where z=x]
  118.30 -  apply (clarsimp simp add: diff_less_iff)
  118.31 -  apply atomize
  118.32 -  apply (erule_tac x="y" in allE)
  118.33 -  apply (erule_tac x="xa" in allE)
  118.34 -  apply arith
  118.35 -  done
  118.36 +lemma open_vimage: (* TODO: move to Topological_Spaces.thy *)
  118.37 +  assumes "open s" and "continuous_on UNIV f"
  118.38 +  shows "open (vimage f s)"
  118.39 +  using assms unfolding continuous_on_open_vimage [OF open_UNIV]
  118.40 +  by simp
  118.41 +
  118.42 +lemma open_ball [intro, simp]: "open (ball x e)"
  118.43 +proof -
  118.44 +  have "open (dist x -` {..<e})"
  118.45 +    by (intro open_vimage open_lessThan continuous_on_intros)
  118.46 +  also have "dist x -` {..<e} = ball x e"
  118.47 +    by auto
  118.48 +  finally show ?thesis .
  118.49 +qed
  118.50  
  118.51  lemma open_contains_ball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. ball x e \<subseteq> S)"
  118.52    unfolding open_dist subset_eq mem_ball Ball_def dist_commute ..
  118.53 @@ -1553,7 +1552,7 @@
  118.54        fix y
  118.55        assume "y \<in> {x<..} \<inter> I"
  118.56        with False bnd have "Inf (f ` ({x<..} \<inter> I)) \<le> f y"
  118.57 -        by (auto intro: cInf_lower)
  118.58 +        by (auto intro!: cInf_lower bdd_belowI2)
  118.59        with a have "a < f y"
  118.60          by (blast intro: less_le_trans)
  118.61      }
  118.62 @@ -1889,7 +1888,6 @@
  118.63  lemma closed_sequential_limits:
  118.64    fixes S :: "'a::first_countable_topology set"
  118.65    shows "closed S \<longleftrightarrow> (\<forall>x l. (\<forall>n. x n \<in> S) \<and> (x ---> l) sequentially \<longrightarrow> l \<in> S)"
  118.66 -  unfolding closed_limpt
  118.67    using closure_sequential [where 'a='a] closure_closed [where 'a='a]
  118.68      closed_limpt [where 'a='a] islimpt_sequential [where 'a='a] mem_delete [where 'a='a]
  118.69    by metis
  118.70 @@ -1908,17 +1906,17 @@
  118.71  
  118.72  lemma closure_contains_Inf:
  118.73    fixes S :: "real set"
  118.74 -  assumes "S \<noteq> {}" "\<forall>x\<in>S. B \<le> x"
  118.75 +  assumes "S \<noteq> {}" "bdd_below S"
  118.76    shows "Inf S \<in> closure S"
  118.77  proof -
  118.78    have *: "\<forall>x\<in>S. Inf S \<le> x"
  118.79 -    using cInf_lower_EX[of _ S] assms by metis
  118.80 +    using cInf_lower[of _ S] assms by metis
  118.81    {
  118.82      fix e :: real
  118.83      assume "e > 0"
  118.84      then have "Inf S < Inf S + e" by simp
  118.85      with assms obtain x where "x \<in> S" "x < Inf S + e"
  118.86 -      by (subst (asm) cInf_less_iff[of _ B]) auto
  118.87 +      by (subst (asm) cInf_less_iff) auto
  118.88      with * have "\<exists>x\<in>S. dist x (Inf S) < e"
  118.89        by (intro bexI[of _ x]) (auto simp add: dist_real_def)
  118.90    }
  118.91 @@ -1927,12 +1925,9 @@
  118.92  
  118.93  lemma closed_contains_Inf:
  118.94    fixes S :: "real set"
  118.95 -  assumes "S \<noteq> {}" "\<forall>x\<in>S. B \<le> x"
  118.96 -    and "closed S"
  118.97 -  shows "Inf S \<in> S"
  118.98 +  shows "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> closed S \<Longrightarrow> Inf S \<in> S"
  118.99    by (metis closure_contains_Inf closure_closed assms)
 118.100  
 118.101 -
 118.102  lemma not_trivial_limit_within_ball:
 118.103    "\<not> trivial_limit (at x within S) \<longleftrightarrow> (\<forall>e>0. S \<inter> ball x e - {x} \<noteq> {})"
 118.104    (is "?lhs = ?rhs")
 118.105 @@ -1974,29 +1969,25 @@
 118.106  
 118.107  subsection {* Infimum Distance *}
 118.108  
 118.109 -definition "infdist x A = (if A = {} then 0 else Inf {dist x a|a. a \<in> A})"
 118.110 -
 118.111 -lemma infdist_notempty: "A \<noteq> {} \<Longrightarrow> infdist x A = Inf {dist x a|a. a \<in> A}"
 118.112 +definition "infdist x A = (if A = {} then 0 else INF a:A. dist x a)"
 118.113 +
 118.114 +lemma bdd_below_infdist[intro, simp]: "bdd_below (dist x`A)"
 118.115 +  by (auto intro!: zero_le_dist)
 118.116 +
 118.117 +lemma infdist_notempty: "A \<noteq> {} \<Longrightarrow> infdist x A = (INF a:A. dist x a)"
 118.118    by (simp add: infdist_def)
 118.119  
 118.120  lemma infdist_nonneg: "0 \<le> infdist x A"
 118.121 -  by (auto simp add: infdist_def intro: cInf_greatest)
 118.122 -
 118.123 -lemma infdist_le:
 118.124 -  assumes "a \<in> A"
 118.125 -    and "d = dist x a"
 118.126 -  shows "infdist x A \<le> d"
 118.127 -  using assms by (auto intro!: cInf_lower[where z=0] simp add: infdist_def)
 118.128 -
 118.129 -lemma infdist_zero[simp]:
 118.130 -  assumes "a \<in> A"
 118.131 -  shows "infdist a A = 0"
 118.132 -proof -
 118.133 -  from infdist_le[OF assms, of "dist a a"] have "infdist a A \<le> 0"
 118.134 -    by auto
 118.135 -  with infdist_nonneg[of a A] assms show "infdist a A = 0"
 118.136 -    by auto
 118.137 -qed
 118.138 +  by (auto simp add: infdist_def intro: cINF_greatest)
 118.139 +
 118.140 +lemma infdist_le: "a \<in> A \<Longrightarrow> infdist x A \<le> dist x a"
 118.141 +  by (auto intro: cINF_lower simp add: infdist_def)
 118.142 +
 118.143 +lemma infdist_le2: "a \<in> A \<Longrightarrow> dist x a \<le> d \<Longrightarrow> infdist x A \<le> d"
 118.144 +  by (auto intro!: cINF_lower2 simp add: infdist_def)
 118.145 +
 118.146 +lemma infdist_zero[simp]: "a \<in> A \<Longrightarrow> infdist a A = 0"
 118.147 +  by (auto intro!: antisym infdist_nonneg infdist_le2)
 118.148  
 118.149  lemma infdist_triangle: "infdist x A \<le> infdist y A + dist x y"
 118.150  proof (cases "A = {}")
 118.151 @@ -2015,18 +2006,11 @@
 118.152        by auto
 118.153      show "infdist x A \<le> d"
 118.154        unfolding infdist_notempty[OF `A \<noteq> {}`]
 118.155 -    proof (rule cInf_lower2)
 118.156 -      show "dist x a \<in> {dist x a |a. a \<in> A}"
 118.157 -        using `a \<in> A` by auto
 118.158 +    proof (rule cINF_lower2)
 118.159 +      show "a \<in> A" by fact
 118.160        show "dist x a \<le> d"
 118.161          unfolding d by (rule dist_triangle)
 118.162 -      fix d
 118.163 -      assume "d \<in> {dist x a |a. a \<in> A}"
 118.164 -      then obtain a where "a \<in> A" "d = dist x a"
 118.165 -        by auto
 118.166 -      then show "infdist x A \<le> d"
 118.167 -        by (rule infdist_le)
 118.168 -    qed
 118.169 +    qed simp
 118.170    qed
 118.171    also have "\<dots> = dist x y + infdist y A"
 118.172    proof (rule cInf_eq, safe)
 118.173 @@ -2039,7 +2023,7 @@
 118.174      assume inf: "\<And>d. d \<in> {dist x y + dist y a |a. a \<in> A} \<Longrightarrow> i \<le> d"
 118.175      then have "i - dist x y \<le> infdist y A"
 118.176        unfolding infdist_notempty[OF `A \<noteq> {}`] using `a \<in> A`
 118.177 -      by (intro cInf_greatest) (auto simp: field_simps)
 118.178 +      by (intro cINF_greatest) (auto simp: field_simps)
 118.179      then show "i \<le> dist x y + infdist y A"
 118.180        by simp
 118.181    qed
 118.182 @@ -2078,7 +2062,7 @@
 118.183      assume "\<not> (\<exists>y\<in>A. dist y x < e)"
 118.184      then have "infdist x A \<ge> e" using `a \<in> A`
 118.185        unfolding infdist_def
 118.186 -      by (force simp: dist_commute intro: cInf_greatest)
 118.187 +      by (force simp: dist_commute intro: cINF_greatest)
 118.188      with x `e > 0` show False by auto
 118.189    qed
 118.190  qed
 118.191 @@ -2129,15 +2113,20 @@
 118.192  
 118.193  subsection {* More properties of closed balls *}
 118.194  
 118.195 +lemma closed_vimage: (* TODO: move to Topological_Spaces.thy *)
 118.196 +  assumes "closed s" and "continuous_on UNIV f"
 118.197 +  shows "closed (vimage f s)"
 118.198 +  using assms unfolding continuous_on_closed_vimage [OF closed_UNIV]
 118.199 +  by simp
 118.200 +
 118.201  lemma closed_cball: "closed (cball x e)"
 118.202 -  unfolding cball_def closed_def
 118.203 -  unfolding Collect_neg_eq [symmetric] not_le
 118.204 -  apply (clarsimp simp add: open_dist, rename_tac y)
 118.205 -  apply (rule_tac x="dist x y - e" in exI, clarsimp)
 118.206 -  apply (rename_tac x')
 118.207 -  apply (cut_tac x=x and y=x' and z=y in dist_triangle)
 118.208 -  apply simp
 118.209 -  done
 118.210 +proof -
 118.211 +  have "closed (dist x -` {..e})"
 118.212 +    by (intro closed_vimage closed_atMost continuous_on_intros)
 118.213 +  also have "dist x -` {..e} = cball x e"
 118.214 +    by auto
 118.215 +  finally show ?thesis .
 118.216 +qed
 118.217  
 118.218  lemma open_contains_cball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0.  cball x e \<subseteq> S)"
 118.219  proof -
 118.220 @@ -2645,11 +2634,19 @@
 118.221  
 118.222  text{* Some theorems on sups and infs using the notion "bounded". *}
 118.223  
 118.224 -lemma bounded_real:
 118.225 -  fixes S :: "real set"
 118.226 -  shows "bounded S \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. abs x \<le> a)"
 118.227 +lemma bounded_real: "bounded (S::real set) \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. \<bar>x\<bar> \<le> a)"
 118.228    by (simp add: bounded_iff)
 118.229  
 118.230 +lemma bounded_imp_bdd_above: "bounded S \<Longrightarrow> bdd_above (S :: real set)"
 118.231 +  by (auto simp: bounded_def bdd_above_def dist_real_def)
 118.232 +     (metis abs_le_D1 abs_minus_commute diff_le_eq)
 118.233 +
 118.234 +lemma bounded_imp_bdd_below: "bounded S \<Longrightarrow> bdd_below (S :: real set)"
 118.235 +  by (auto simp: bounded_def bdd_below_def dist_real_def)
 118.236 +     (metis abs_le_D1 add_commute diff_le_eq)
 118.237 +
 118.238 +(* TODO: remove the following lemmas about Inf and Sup, is now in conditionally complete lattice *)
 118.239 +
 118.240  lemma bounded_has_Sup:
 118.241    fixes S :: "real set"
 118.242    assumes "bounded S"
 118.243 @@ -2657,22 +2654,14 @@
 118.244    shows "\<forall>x\<in>S. x \<le> Sup S"
 118.245      and "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b"
 118.246  proof
 118.247 -  fix x
 118.248 -  assume "x\<in>S"
 118.249 -  then show "x \<le> Sup S"
 118.250 -    by (metis cSup_upper abs_le_D1 assms(1) bounded_real)
 118.251 -next
 118.252    show "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b"
 118.253      using assms by (metis cSup_least)
 118.254 -qed
 118.255 +qed (metis cSup_upper assms(1) bounded_imp_bdd_above)
 118.256  
 118.257  lemma Sup_insert:
 118.258    fixes S :: "real set"
 118.259    shows "bounded S \<Longrightarrow> Sup (insert x S) = (if S = {} then x else max x (Sup S))"
 118.260 -  apply (subst cSup_insert_If)
 118.261 -  apply (rule bounded_has_Sup(1)[of S, rule_format])
 118.262 -  apply (auto simp: sup_max)
 118.263 -  done
 118.264 +  by (auto simp: bounded_imp_bdd_above sup_max cSup_insert_If)
 118.265  
 118.266  lemma Sup_insert_finite:
 118.267    fixes S :: "real set"
 118.268 @@ -2689,24 +2678,14 @@
 118.269    shows "\<forall>x\<in>S. x \<ge> Inf S"
 118.270      and "\<forall>b. (\<forall>x\<in>S. x \<ge> b) \<longrightarrow> Inf S \<ge> b"
 118.271  proof
 118.272 -  fix x
 118.273 -  assume "x \<in> S"
 118.274 -  from assms(1) obtain a where a: "\<forall>x\<in>S. \<bar>x\<bar> \<le> a"
 118.275 -    unfolding bounded_real by auto
 118.276 -  then show "x \<ge> Inf S" using `x \<in> S`
 118.277 -    by (metis cInf_lower_EX abs_le_D2 minus_le_iff)
 118.278 -next
 118.279    show "\<forall>b. (\<forall>x\<in>S. x \<ge> b) \<longrightarrow> Inf S \<ge> b"
 118.280      using assms by (metis cInf_greatest)
 118.281 -qed
 118.282 +qed (metis cInf_lower assms(1) bounded_imp_bdd_below)
 118.283  
 118.284  lemma Inf_insert:
 118.285    fixes S :: "real set"
 118.286    shows "bounded S \<Longrightarrow> Inf (insert x S) = (if S = {} then x else min x (Inf S))"
 118.287 -  apply (subst cInf_insert_if)
 118.288 -  apply (rule bounded_has_Inf(1)[of S, rule_format])
 118.289 -  apply (auto simp: inf_min)
 118.290 -  done
 118.291 +  by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_If)
 118.292  
 118.293  lemma Inf_insert_finite:
 118.294    fixes S :: "real set"
 118.295 @@ -3298,6 +3277,50 @@
 118.296    where "seq_compact S \<longleftrightarrow>
 118.297      (\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow> (\<exists>l\<in>S. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially))"
 118.298  
 118.299 +lemma seq_compactI:
 118.300 +  assumes "\<And>f. \<forall>n. f n \<in> S \<Longrightarrow> \<exists>l\<in>S. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
 118.301 +  shows "seq_compact S"
 118.302 +  unfolding seq_compact_def using assms by fast
 118.303 +
 118.304 +lemma seq_compactE:
 118.305 +  assumes "seq_compact S" "\<forall>n. f n \<in> S"
 118.306 +  obtains l r where "l \<in> S" "subseq r" "((f \<circ> r) ---> l) sequentially"
 118.307 +  using assms unfolding seq_compact_def by fast
 118.308 +
 118.309 +lemma closed_sequentially: (* TODO: move upwards *)
 118.310 +  assumes "closed s" and "\<forall>n. f n \<in> s" and "f ----> l"
 118.311 +  shows "l \<in> s"
 118.312 +proof (rule ccontr)
 118.313 +  assume "l \<notin> s"
 118.314 +  with `closed s` and `f ----> l` have "eventually (\<lambda>n. f n \<in> - s) sequentially"
 118.315 +    by (fast intro: topological_tendstoD)
 118.316 +  with `\<forall>n. f n \<in> s` show "False"
 118.317 +    by simp
 118.318 +qed
 118.319 +
 118.320 +lemma seq_compact_inter_closed:
 118.321 +  assumes "seq_compact s" and "closed t"
 118.322 +  shows "seq_compact (s \<inter> t)"
 118.323 +proof (rule seq_compactI)
 118.324 +  fix f assume "\<forall>n::nat. f n \<in> s \<inter> t"
 118.325 +  hence "\<forall>n. f n \<in> s" and "\<forall>n. f n \<in> t"
 118.326 +    by simp_all
 118.327 +  from `seq_compact s` and `\<forall>n. f n \<in> s`
 118.328 +  obtain l r where "l \<in> s" and r: "subseq r" and l: "(f \<circ> r) ----> l"
 118.329 +    by (rule seq_compactE)
 118.330 +  from `\<forall>n. f n \<in> t` have "\<forall>n. (f \<circ> r) n \<in> t"
 118.331 +    by simp
 118.332 +  from `closed t` and this and l have "l \<in> t"
 118.333 +    by (rule closed_sequentially)
 118.334 +  with `l \<in> s` and r and l show "\<exists>l\<in>s \<inter> t. \<exists>r. subseq r \<and> (f \<circ> r) ----> l"
 118.335 +    by fast
 118.336 +qed
 118.337 +
 118.338 +lemma seq_compact_closed_subset:
 118.339 +  assumes "closed s" and "s \<subseteq> t" and "seq_compact t"
 118.340 +  shows "seq_compact s"
 118.341 +  using assms seq_compact_inter_closed [of t s] by (simp add: Int_absorb1)
 118.342 +
 118.343  lemma seq_compact_imp_countably_compact:
 118.344    fixes U :: "'a :: first_countable_topology set"
 118.345    assumes "seq_compact U"
 118.346 @@ -3410,16 +3433,6 @@
 118.347      using `x \<in> U` by (auto simp: convergent_def comp_def)
 118.348  qed
 118.349  
 118.350 -lemma seq_compactI:
 118.351 -  assumes "\<And>f. \<forall>n. f n \<in> S \<Longrightarrow> \<exists>l\<in>S. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
 118.352 -  shows "seq_compact S"
 118.353 -  unfolding seq_compact_def using assms by fast
 118.354 -
 118.355 -lemma seq_compactE:
 118.356 -  assumes "seq_compact S" "\<forall>n. f n \<in> S"
 118.357 -  obtains l r where "l \<in> S" "subseq r" "((f \<circ> r) ---> l) sequentially"
 118.358 -  using assms unfolding seq_compact_def by fast
 118.359 -
 118.360  lemma countably_compact_imp_acc_point:
 118.361    assumes "countably_compact s"
 118.362      and "countable t"
 118.363 @@ -3654,6 +3667,8 @@
 118.364    "\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x \<in> s. x islimpt t) \<Longrightarrow> bounded s"
 118.365    using compact_imp_bounded unfolding compact_eq_bolzano_weierstrass .
 118.366  
 118.367 +subsection {* Metric spaces with the Heine-Borel property *}
 118.368 +
 118.369  text {*
 118.370    A metric space (or topological vector space) is said to have the
 118.371    Heine-Borel property if every closed and bounded subset is compact.
 118.372 @@ -3678,7 +3693,7 @@
 118.373    from f have fr: "\<forall>n. (f \<circ> r) n \<in> s"
 118.374      by simp
 118.375    have "l \<in> s" using `closed s` fr l
 118.376 -    unfolding closed_sequential_limits by blast
 118.377 +    by (rule closed_sequentially)
 118.378    show "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"
 118.379      using `l \<in> s` r l by blast
 118.380  qed
 118.381 @@ -3859,11 +3874,21 @@
 118.382      using l r by fast
 118.383  qed
 118.384  
 118.385 -subsubsection{* Completeness *}
 118.386 +subsubsection {* Completeness *}
 118.387  
 118.388  definition complete :: "'a::metric_space set \<Rightarrow> bool"
 118.389    where "complete s \<longleftrightarrow> (\<forall>f. (\<forall>n. f n \<in> s) \<and> Cauchy f \<longrightarrow> (\<exists>l\<in>s. f ----> l))"
 118.390  
 118.391 +lemma completeI:
 118.392 +  assumes "\<And>f. \<forall>n. f n \<in> s \<Longrightarrow> Cauchy f \<Longrightarrow> \<exists>l\<in>s. f ----> l"
 118.393 +  shows "complete s"
 118.394 +  using assms unfolding complete_def by fast
 118.395 +
 118.396 +lemma completeE:
 118.397 +  assumes "complete s" and "\<forall>n. f n \<in> s" and "Cauchy f"
 118.398 +  obtains l where "l \<in> s" and "f ----> l"
 118.399 +  using assms unfolding complete_def by fast
 118.400 +
 118.401  lemma compact_imp_complete:
 118.402    assumes "compact s"
 118.403    shows "complete s"
 118.404 @@ -4085,49 +4110,57 @@
 118.405  
 118.406  instance euclidean_space \<subseteq> banach ..
 118.407  
 118.408 -lemma complete_univ: "complete (UNIV :: 'a::complete_space set)"
 118.409 -proof (simp add: complete_def, rule, rule)
 118.410 -  fix f :: "nat \<Rightarrow> 'a"
 118.411 -  assume "Cauchy f"
 118.412 +lemma complete_UNIV: "complete (UNIV :: ('a::complete_space) set)"
 118.413 +proof (rule completeI)
 118.414 +  fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"
 118.415    then have "convergent f" by (rule Cauchy_convergent)
 118.416 -  then show "\<exists>l. f ----> l" unfolding convergent_def .
 118.417 +  then show "\<exists>l\<in>UNIV. f ----> l" unfolding convergent_def by simp
 118.418  qed
 118.419  
 118.420  lemma complete_imp_closed:
 118.421    assumes "complete s"
 118.422    shows "closed s"
 118.423 -proof -
 118.424 -  {
 118.425 -    fix x
 118.426 -    assume "x islimpt s"
 118.427 -    then obtain f where f: "\<forall>n. f n \<in> s - {x}" "(f ---> x) sequentially"
 118.428 -      unfolding islimpt_sequential by auto
 118.429 -    then obtain l where l: "l\<in>s" "(f ---> l) sequentially"
 118.430 -      using `complete s`[unfolded complete_def] using LIMSEQ_imp_Cauchy[of f x] by auto
 118.431 -    then have "x \<in> s"
 118.432 -      using tendsto_unique[of sequentially f l x] trivial_limit_sequentially f(2) by auto
 118.433 -  }
 118.434 -  then show "closed s" unfolding closed_limpt by auto
 118.435 -qed
 118.436 +proof (unfold closed_sequential_limits, clarify)
 118.437 +  fix f x assume "\<forall>n. f n \<in> s" and "f ----> x"
 118.438 +  from `f ----> x` have "Cauchy f"
 118.439 +    by (rule LIMSEQ_imp_Cauchy)
 118.440 +  with `complete s` and `\<forall>n. f n \<in> s` obtain l where "l \<in> s" and "f ----> l"
 118.441 +    by (rule completeE)
 118.442 +  from `f ----> x` and `f ----> l` have "x = l"
 118.443 +    by (rule LIMSEQ_unique)
 118.444 +  with `l \<in> s` show "x \<in> s"
 118.445 +    by simp
 118.446 +qed
 118.447 +
 118.448 +lemma complete_inter_closed:
 118.449 +  assumes "complete s" and "closed t"
 118.450 +  shows "complete (s \<inter> t)"
 118.451 +proof (rule completeI)
 118.452 +  fix f assume "\<forall>n. f n \<in> s \<inter> t" and "Cauchy f"
 118.453 +  then have "\<forall>n. f n \<in> s" and "\<forall>n. f n \<in> t"
 118.454 +    by simp_all
 118.455 +  from `complete s` obtain l where "l \<in> s" and "f ----> l"
 118.456 +    using `\<forall>n. f n \<in> s` and `Cauchy f` by (rule completeE)
 118.457 +  from `closed t` and `\<forall>n. f n \<in> t` and `f ----> l` have "l \<in> t"
 118.458 +    by (rule closed_sequentially)
 118.459 +  with `l \<in> s` and `f ----> l` show "\<exists>l\<in>s \<inter> t. f ----> l"
 118.460 +    by fast
 118.461 +qed
 118.462 +
 118.463 +lemma complete_closed_subset:
 118.464 +  assumes "closed s" and "s \<subseteq> t" and "complete t"
 118.465 +  shows "complete s"
 118.466 +  using assms complete_inter_closed [of t s] by (simp add: Int_absorb1)
 118.467  
 118.468  lemma complete_eq_closed:
 118.469 -  fixes s :: "'a::complete_space set"
 118.470 -  shows "complete s \<longleftrightarrow> closed s" (is "?lhs = ?rhs")
 118.471 +  fixes s :: "('a::complete_space) set"
 118.472 +  shows "complete s \<longleftrightarrow> closed s"
 118.473  proof
 118.474 -  assume ?lhs
 118.475 -  then show ?rhs by (rule complete_imp_closed)
 118.476 +  assume "closed s" then show "complete s"
 118.477 +    using subset_UNIV complete_UNIV by (rule complete_closed_subset)
 118.478  next
 118.479 -  assume ?rhs
 118.480 -  {
 118.481 -    fix f
 118.482 -    assume as:"\<forall>n::nat. f n \<in> s" "Cauchy f"
 118.483 -    then obtain l where "(f ---> l) sequentially"
 118.484 -      using complete_univ[unfolded complete_def, THEN spec[where x=f]] by auto
 118.485 -    then have "\<exists>l\<in>s. (f ---> l) sequentially"
 118.486 -      using `?rhs`[unfolded closed_sequential_limits, THEN spec[where x=f], THEN spec[where x=l]]
 118.487 -      using as(1) by auto
 118.488 -  }
 118.489 -  then show ?lhs unfolding complete_def by auto
 118.490 +  assume "complete s" then show "closed s"
 118.491 +    by (rule complete_imp_closed)
 118.492  qed
 118.493  
 118.494  lemma convergent_eq_cauchy:
 118.495 @@ -4142,13 +4175,13 @@
 118.496  
 118.497  lemma compact_cball[simp]:
 118.498    fixes x :: "'a::heine_borel"
 118.499 -  shows "compact(cball x e)"
 118.500 +  shows "compact (cball x e)"
 118.501    using compact_eq_bounded_closed bounded_cball closed_cball
 118.502    by blast
 118.503  
 118.504  lemma compact_frontier_bounded[intro]:
 118.505    fixes s :: "'a::heine_borel set"
 118.506 -  shows "bounded s \<Longrightarrow> compact(frontier s)"
 118.507 +  shows "bounded s \<Longrightarrow> compact (frontier s)"
 118.508    unfolding frontier_def
 118.509    using compact_eq_bounded_closed
 118.510    by blast
 118.511 @@ -4168,68 +4201,51 @@
 118.512  subsection {* Bounded closed nest property (proof does not use Heine-Borel) *}
 118.513  
 118.514  lemma bounded_closed_nest:
 118.515 -  assumes "\<forall>n. closed(s n)"
 118.516 -    and "\<forall>n. (s n \<noteq> {})"
 118.517 -    and "(\<forall>m n. m \<le> n --> s n \<subseteq> s m)"
 118.518 -    and "bounded(s 0)"
 118.519 -  shows "\<exists>a::'a::heine_borel. \<forall>n::nat. a \<in> s(n)"
 118.520 +  fixes s :: "nat \<Rightarrow> ('a::heine_borel) set"
 118.521 +  assumes "\<forall>n. closed (s n)"
 118.522 +    and "\<forall>n. s n \<noteq> {}"
 118.523 +    and "\<forall>m n. m \<le> n \<longrightarrow> s n \<subseteq> s m"
 118.524 +    and "bounded (s 0)"
 118.525 +  shows "\<exists>a. \<forall>n. a \<in> s n"
 118.526  proof -
 118.527 -  from assms(2) obtain x where x:"\<forall>n::nat. x n \<in> s n"
 118.528 -    using choice[of "\<lambda>n x. x\<in> s n"] by auto
 118.529 -  from assms(4,1) have *:"seq_compact (s 0)"
 118.530 -    using bounded_closed_imp_seq_compact[of "s 0"] by auto
 118.531 -
 118.532 -  then obtain l r where lr:"l\<in>s 0" "subseq r" "((x \<circ> r) ---> l) sequentially"
 118.533 -    unfolding seq_compact_def
 118.534 -    apply (erule_tac x=x in allE)
 118.535 -    using x using assms(3)
 118.536 -    apply blast
 118.537 -    done
 118.538 -
 118.539 -  {
 118.540 +  from assms(2) obtain x where x: "\<forall>n. x n \<in> s n"
 118.541 +    using choice[of "\<lambda>n x. x \<in> s n"] by auto
 118.542 +  from assms(4,1) have "seq_compact (s 0)"
 118.543 +    by (simp add: bounded_closed_imp_seq_compact)
 118.544 +  then obtain l r where lr: "l \<in> s 0" "subseq r" "(x \<circ> r) ----> l"
 118.545 +    using x and assms(3) unfolding seq_compact_def by blast
 118.546 +  have "\<forall>n. l \<in> s n"
 118.547 +  proof
 118.548      fix n :: nat
 118.549 -    {
 118.550 -      fix e :: real
 118.551 -      assume "e>0"
 118.552 -      with lr(3) obtain N where N:"\<forall>m\<ge>N. dist ((x \<circ> r) m) l < e"
 118.553 -        unfolding LIMSEQ_def by auto
 118.554 -      then have "dist ((x \<circ> r) (max N n)) l < e" by auto
 118.555 -      moreover
 118.556 -      have "r (max N n) \<ge> n" using lr(2) using seq_suble[of r "max N n"]
 118.557 -        by auto
 118.558 -      then have "(x \<circ> r) (max N n) \<in> s n"
 118.559 -        using x
 118.560 -        apply (erule_tac x=n in allE)
 118.561 -        using x
 118.562 -        apply (erule_tac x="r (max N n)" in allE)
 118.563 -        using assms(3)
 118.564 -        apply (erule_tac x=n in allE)
 118.565 -        apply (erule_tac x="r (max N n)" in allE)
 118.566 -        apply auto
 118.567 -        done
 118.568 -      ultimately have "\<exists>y\<in>s n. dist y l < e"
 118.569 -        by auto
 118.570 -    }
 118.571 -    then have "l \<in> s n"
 118.572 -      using closed_approachable[of "s n" l] assms(1) by blast
 118.573 -  }
 118.574 -  then show ?thesis by auto
 118.575 +    have "closed (s n)"
 118.576 +      using assms(1) by simp
 118.577 +    moreover have "\<forall>i. (x \<circ> r) i \<in> s i"
 118.578 +      using x and assms(3) and lr(2) [THEN seq_suble] by auto
 118.579 +    then have "\<forall>i. (x \<circ> r) (i + n) \<in> s n"
 118.580 +      using assms(3) by (fast intro!: le_add2)
 118.581 +    moreover have "(\<lambda>i. (x \<circ> r) (i + n)) ----> l"
 118.582 +      using lr(3) by (rule LIMSEQ_ignore_initial_segment)
 118.583 +    ultimately show "l \<in> s n"
 118.584 +      by (rule closed_sequentially)
 118.585 +  qed
 118.586 +  then show ?thesis ..
 118.587  qed
 118.588  
 118.589  text {* Decreasing case does not even need compactness, just completeness. *}
 118.590  
 118.591  lemma decreasing_closed_nest:
 118.592 +  fixes s :: "nat \<Rightarrow> ('a::complete_space) set"
 118.593    assumes
 118.594 -    "\<forall>n. closed(s n)"
 118.595 -    "\<forall>n. (s n \<noteq> {})"
 118.596 -    "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
 118.597 -    "\<forall>e>0. \<exists>n. \<forall>x \<in> (s n). \<forall> y \<in> (s n). dist x y < e"
 118.598 -  shows "\<exists>a::'a::complete_space. \<forall>n::nat. a \<in> s n"
 118.599 -proof-
 118.600 -  have "\<forall>n. \<exists> x. x\<in>s n"
 118.601 +    "\<forall>n. closed (s n)"
 118.602 +    "\<forall>n. s n \<noteq> {}"
 118.603 +    "\<forall>m n. m \<le> n \<longrightarrow> s n \<subseteq> s m"
 118.604 +    "\<forall>e>0. \<exists>n. \<forall>x\<in>s n. \<forall>y\<in>s n. dist x y < e"
 118.605 +  shows "\<exists>a. \<forall>n. a \<in> s n"
 118.606 +proof -
 118.607 +  have "\<forall>n. \<exists>x. x \<in> s n"
 118.608      using assms(2) by auto
 118.609    then have "\<exists>t. \<forall>n. t n \<in> s n"
 118.610 -    using choice[of "\<lambda> n x. x \<in> s n"] by auto
 118.611 +    using choice[of "\<lambda>n x. x \<in> s n"] by auto
 118.612    then obtain t where t: "\<forall>n. t n \<in> s n" by auto
 118.613    {
 118.614      fix e :: real
 118.615 @@ -4250,7 +4266,7 @@
 118.616    then have "Cauchy t"
 118.617      unfolding cauchy_def by auto
 118.618    then obtain l where l:"(t ---> l) sequentially"
 118.619 -    using complete_univ unfolding complete_def by auto
 118.620 +    using complete_UNIV unfolding complete_def by auto
 118.621    {
 118.622      fix n :: nat
 118.623      {
 118.624 @@ -4285,7 +4301,7 @@
 118.625    assumes
 118.626      "\<forall>n. closed(s n)"
 118.627      "\<forall>n. s n \<noteq> {}"
 118.628 -    "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
 118.629 +    "\<forall>m n. m \<le> n \<longrightarrow> s n \<subseteq> s m"
 118.630      "\<forall>e>0. \<exists>n. \<forall>x \<in> (s n). \<forall> y\<in>(s n). dist x y < e"
 118.631    shows "\<exists>a. \<Inter>(range s) = {a}"
 118.632  proof -
 118.633 @@ -4815,8 +4831,8 @@
 118.634    assumes "uniformly_continuous_on s f"
 118.635      and "uniformly_continuous_on s g"
 118.636    shows "uniformly_continuous_on s (\<lambda>x. f x - g x)"
 118.637 -  unfolding ab_diff_minus using assms
 118.638 -  by (intro uniformly_continuous_on_add uniformly_continuous_on_minus)
 118.639 +  using assms uniformly_continuous_on_add [of s f "- g"]
 118.640 +    by (simp add: fun_Compl_def uniformly_continuous_on_minus)
 118.641  
 118.642  text{* Continuity of all kinds is preserved under composition. *}
 118.643  
 118.644 @@ -5637,8 +5653,6 @@
 118.645      apply auto
 118.646      apply (rule_tac x= xa in exI)
 118.647      apply auto
 118.648 -    apply (rule_tac x=xa in exI)
 118.649 -    apply auto
 118.650      done
 118.651    then show ?thesis
 118.652      using compact_sums[OF assms(1) compact_negations[OF assms(2)]] by auto
 118.653 @@ -5686,24 +5700,27 @@
 118.654  
 118.655  text {* We can state this in terms of diameter of a set. *}
 118.656  
 118.657 -definition "diameter s = (if s = {} then 0::real else Sup {dist x y | x y. x \<in> s \<and> y \<in> s})"
 118.658 +definition diameter :: "'a::metric_space set \<Rightarrow> real" where
 118.659 +  "diameter S = (if S = {} then 0 else SUP (x,y):S\<times>S. dist x y)"
 118.660  
 118.661  lemma diameter_bounded_bound:
 118.662    fixes s :: "'a :: metric_space set"
 118.663    assumes s: "bounded s" "x \<in> s" "y \<in> s"
 118.664    shows "dist x y \<le> diameter s"
 118.665  proof -
 118.666 -  let ?D = "{dist x y |x y. x \<in> s \<and> y \<in> s}"
 118.667    from s obtain z d where z: "\<And>x. x \<in> s \<Longrightarrow> dist z x \<le> d"
 118.668      unfolding bounded_def by auto
 118.669 -  have "dist x y \<le> Sup ?D"
 118.670 -  proof (rule cSup_upper, safe)
 118.671 +  have "bdd_above (split dist ` (s\<times>s))"
 118.672 +  proof (intro bdd_aboveI, safe)
 118.673      fix a b
 118.674      assume "a \<in> s" "b \<in> s"
 118.675      with z[of a] z[of b] dist_triangle[of a b z]
 118.676      show "dist a b \<le> 2 * d"
 118.677        by (simp add: dist_commute)
 118.678 -  qed (insert s, auto)
 118.679 +  qed
 118.680 +  moreover have "(x,y) \<in> s\<times>s" using s by auto
 118.681 +  ultimately have "dist x y \<le> (SUP (x,y):s\<times>s. dist x y)"
 118.682 +    by (rule cSUP_upper2) simp
 118.683    with `x \<in> s` show ?thesis
 118.684      by (auto simp add: diameter_def)
 118.685  qed
 118.686 @@ -5714,16 +5731,12 @@
 118.687      and d: "0 < d" "d < diameter s"
 118.688    shows "\<exists>x\<in>s. \<exists>y\<in>s. d < dist x y"
 118.689  proof (rule ccontr)
 118.690 -  let ?D = "{dist x y |x y. x \<in> s \<and> y \<in> s}"
 118.691    assume contr: "\<not> ?thesis"
 118.692 -  moreover
 118.693 -  from d have "s \<noteq> {}"
 118.694 -    by (auto simp: diameter_def)
 118.695 -  then have "?D \<noteq> {}" by auto
 118.696 -  ultimately have "Sup ?D \<le> d"
 118.697 -    by (intro cSup_least) (auto simp: not_less)
 118.698 -  with `d < diameter s` `s \<noteq> {}` show False
 118.699 -    by (auto simp: diameter_def)
 118.700 +  moreover have "s \<noteq> {}"
 118.701 +    using d by (auto simp add: diameter_def)
 118.702 +  ultimately have "diameter s \<le> d"
 118.703 +    by (auto simp: not_less diameter_def intro!: cSUP_least)
 118.704 +  with `d < diameter s` show False by auto
 118.705  qed
 118.706  
 118.707  lemma diameter_bounded:
 118.708 @@ -5746,7 +5759,7 @@
 118.709    then have "diameter s \<le> dist x y"
 118.710      unfolding diameter_def
 118.711      apply clarsimp
 118.712 -    apply (rule cSup_least)
 118.713 +    apply (rule cSUP_least)
 118.714      apply fast+
 118.715      done
 118.716    then show ?thesis
 118.717 @@ -6989,7 +7002,8 @@
 118.718    unfolding homeomorphic_minimal
 118.719    apply (rule_tac x="\<lambda>x. a + x" in exI)
 118.720    apply (rule_tac x="\<lambda>x. -a + x" in exI)
 118.721 -  using continuous_on_add[OF continuous_on_const continuous_on_id]
 118.722 +  using continuous_on_add [OF continuous_on_const continuous_on_id, of s a]
 118.723 +    continuous_on_add [OF continuous_on_const continuous_on_id, of "plus a ` s" "- a"]
 118.724    apply auto
 118.725    done
 118.726  
 118.727 @@ -7350,14 +7364,14 @@
 118.728      fix y
 118.729      assume "a \<le> y" "y \<le> b" "m > 0"
 118.730      then have "m *\<^sub>R a + c \<le> m *\<^sub>R y + c" and "m *\<^sub>R y + c \<le> m *\<^sub>R b + c"
 118.731 -      unfolding eucl_le[where 'a='a] by (auto simp: inner_simps)
 118.732 +      unfolding eucl_le[where 'a='a] by (auto simp: inner_distrib)
 118.733    }
 118.734    moreover
 118.735    {
 118.736      fix y
 118.737      assume "a \<le> y" "y \<le> b" "m < 0"
 118.738      then have "m *\<^sub>R b + c \<le> m *\<^sub>R y + c" and "m *\<^sub>R y + c \<le> m *\<^sub>R a + c"
 118.739 -      unfolding eucl_le[where 'a='a] by (auto simp add: mult_left_mono_neg inner_simps)
 118.740 +      unfolding eucl_le[where 'a='a] by (auto simp add: mult_left_mono_neg inner_distrib)
 118.741    }
 118.742    moreover
 118.743    {
 118.744 @@ -7366,7 +7380,7 @@
 118.745      then have "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
 118.746        unfolding image_iff Bex_def mem_interval eucl_le[where 'a='a]
 118.747        apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"])
 118.748 -      apply (auto simp add: pos_le_divide_eq pos_divide_le_eq mult_commute diff_le_iff inner_simps)
 118.749 +      apply (auto simp add: pos_le_divide_eq pos_divide_le_eq mult_commute diff_le_iff inner_distrib inner_diff_left)
 118.750        done
 118.751    }
 118.752    moreover
 118.753 @@ -7376,7 +7390,7 @@
 118.754      then have "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
 118.755        unfolding image_iff Bex_def mem_interval eucl_le[where 'a='a]
 118.756        apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"])
 118.757 -      apply (auto simp add: neg_le_divide_eq neg_divide_le_eq mult_commute diff_le_iff inner_simps)
 118.758 +      apply (auto simp add: neg_le_divide_eq neg_divide_le_eq mult_commute diff_le_iff inner_distrib inner_diff_left)
 118.759        done
 118.760    }
 118.761    ultimately show ?thesis using False by auto
   119.1 --- a/src/HOL/NSA/CLim.thy	Mon Nov 11 17:34:44 2013 +0100
   119.2 +++ b/src/HOL/NSA/CLim.thy	Mon Nov 11 17:44:21 2013 +0100
   119.3 @@ -22,11 +22,11 @@
   119.4  lemma all_shift: "(\<forall>x::'a::comm_ring_1. P x) = (\<forall>x. P (x-a))";
   119.5  apply auto 
   119.6  apply (drule_tac x="x+a" in spec) 
   119.7 -apply (simp add: diff_minus add_assoc) 
   119.8 +apply (simp add: add_assoc) 
   119.9  done
  119.10  
  119.11  lemma complex_add_minus_iff [simp]: "(x + - a = (0::complex)) = (x=a)"
  119.12 -by (simp add: diff_eq_eq diff_minus [symmetric])
  119.13 +by (simp add: diff_eq_eq)
  119.14  
  119.15  lemma complex_add_eq_0_iff [iff]: "(x+y = (0::complex)) = (y = -x)"
  119.16  apply auto
   120.1 --- a/src/HOL/NSA/HDeriv.thy	Mon Nov 11 17:34:44 2013 +0100
   120.2 +++ b/src/HOL/NSA/HDeriv.thy	Mon Nov 11 17:44:21 2013 +0100
   120.3 @@ -81,8 +81,7 @@
   120.4  text{*second equivalence *}
   120.5  lemma NSDERIV_NSLIM_iff2:
   120.6       "(NSDERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) -- x --NS> D)"
   120.7 -by (simp add: NSDERIV_NSLIM_iff DERIV_LIM_iff  diff_minus [symmetric]
   120.8 -              LIM_NSLIM_iff [symmetric])
   120.9 +  by (simp add: NSDERIV_NSLIM_iff DERIV_LIM_iff LIM_NSLIM_iff [symmetric])
  120.10  
  120.11  (* while we're at it! *)
  120.12  
  120.13 @@ -120,11 +119,10 @@
  120.14                   hypreal_of_real (f x))\<approx> (hypreal_of_real D) * h)"
  120.15  apply (auto simp add: nsderiv_def)
  120.16  apply (case_tac "h = (0::hypreal) ")
  120.17 -apply (auto simp add: diff_minus)
  120.18 +apply auto
  120.19  apply (drule_tac x = h in bspec)
  120.20  apply (drule_tac [2] c = h in approx_mult1)
  120.21 -apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
  120.22 -            simp add: diff_minus)
  120.23 +apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD])
  120.24  done
  120.25  
  120.26  lemma NSDERIVD3:
  120.27 @@ -135,8 +133,7 @@
  120.28  apply (auto simp add: nsderiv_def)
  120.29  apply (rule ccontr, drule_tac x = h in bspec)
  120.30  apply (drule_tac [2] c = h in approx_mult1)
  120.31 -apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
  120.32 -            simp add: mult_assoc diff_minus)
  120.33 +apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] simp add: mult_assoc)
  120.34  done
  120.35  
  120.36  text{*Differentiability implies continuity
  120.37 @@ -174,7 +171,7 @@
  120.38  apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def)
  120.39  apply (auto simp add: add_divide_distrib diff_divide_distrib dest!: spec)
  120.40  apply (drule_tac b = "star_of Da" and d = "star_of Db" in approx_add)
  120.41 -apply (auto simp add: diff_minus add_ac)
  120.42 +apply (auto simp add: add_ac algebra_simps)
  120.43  done
  120.44  
  120.45  text{*Product of functions - Proof is trivial but tedious
  120.46 @@ -234,9 +231,11 @@
  120.47    hence deriv: "(\<lambda>h. - ((f(x+h) - f x) / h)) -- 0 --NS> - D"
  120.48      by (rule NSLIM_minus)
  120.49    have "\<forall>h. - ((f (x + h) - f x) / h) = (- f (x + h) + f x) / h"
  120.50 -    by (simp add: minus_divide_left diff_minus)
  120.51 +    by (simp add: minus_divide_left)
  120.52    with deriv
  120.53 -  show "(\<lambda>h. (- f (x + h) + f x) / h) -- 0 --NS> - D" by simp
  120.54 +  have "(\<lambda>h. (- f (x + h) + f x) / h) -- 0 --NS> - D" by simp
  120.55 +  then show "(\<lambda>h. (f (x + h) - f x) / h) -- 0 --NS> D \<Longrightarrow>
  120.56 +    (\<lambda>h. (f x - f (x + h)) / h) -- 0 --NS> - D" by simp
  120.57  qed
  120.58  
  120.59  text{*Subtraction*}
  120.60 @@ -244,11 +243,8 @@
  120.61  by (blast dest: NSDERIV_add NSDERIV_minus)
  120.62  
  120.63  lemma NSDERIV_diff:
  120.64 -     "[| NSDERIV f x :> Da; NSDERIV g x :> Db |]
  120.65 -      ==> NSDERIV (%x. f x - g x) x :> Da-Db"
  120.66 -apply (simp add: diff_minus)
  120.67 -apply (blast intro: NSDERIV_add_minus)
  120.68 -done
  120.69 +  "NSDERIV f x :> Da \<Longrightarrow> NSDERIV g x :> Db \<Longrightarrow> NSDERIV (\<lambda>x. f x - g x) x :> Da-Db"
  120.70 +  using NSDERIV_add_minus [of f x Da g Db] by simp
  120.71  
  120.72  text{*  Similarly to the above, the chain rule admits an entirely
  120.73     straightforward derivation. Compare this with Harrison's
  120.74 @@ -294,7 +290,7 @@
  120.75                     - star_of (f (g x)))
  120.76                / (( *f* g) (star_of(x) + xa) - star_of (g x))
  120.77               \<approx> star_of(Da)"
  120.78 -by (auto simp add: NSDERIV_NSLIM_iff2 NSLIM_def diff_minus [symmetric])
  120.79 +by (auto simp add: NSDERIV_NSLIM_iff2 NSLIM_def)
  120.80  
  120.81  (*--------------------------------------------------------------
  120.82     from other version of differentiability
  120.83 @@ -354,13 +350,23 @@
  120.84      from h_Inf have "h * star_of x \<in> Infinitesimal" by (rule Infinitesimal_HFinite_mult) simp
  120.85      with assms have "inverse (- (h * star_of x) + - (star_of x * star_of x)) \<approx>
  120.86        inverse (- (star_of x * star_of x))"
  120.87 -      by (auto intro: inverse_add_Infinitesimal_approx2
  120.88 +      apply - apply (rule inverse_add_Infinitesimal_approx2)
  120.89 +      apply (auto
  120.90          dest!: hypreal_of_real_HFinite_diff_Infinitesimal
  120.91          simp add: inverse_minus_eq [symmetric] HFinite_minus_iff)
  120.92 -    with not_0 `h \<noteq> 0` assms have "(inverse (star_of x + h) - inverse (star_of x)) / h \<approx>
  120.93 +      done
  120.94 +    moreover from not_0 `h \<noteq> 0` assms
  120.95 +      have "inverse (- (h * star_of x) + - (star_of x * star_of x)) =
  120.96 +        (inverse (star_of x + h) - inverse (star_of x)) / h"
  120.97 +      apply (simp add: division_ring_inverse_diff nonzero_inverse_mult_distrib [symmetric]
  120.98 +        nonzero_inverse_minus_eq [symmetric] ac_simps ring_distribs)
  120.99 +      apply (subst nonzero_inverse_minus_eq [symmetric])
 120.100 +      using distrib_right [symmetric, of h "star_of x" "star_of x"] apply simp
 120.101 +      apply (simp add: field_simps) 
 120.102 +      done
 120.103 +    ultimately have "(inverse (star_of x + h) - inverse (star_of x)) / h \<approx>
 120.104        - (inverse (star_of x) * inverse (star_of x))"
 120.105 -      by (simp add: inverse_add nonzero_inverse_mult_distrib [symmetric]
 120.106 -        nonzero_inverse_minus_eq [symmetric] add_ac mult_ac diff_minus ring_distribs)
 120.107 +      using assms by (simp add: nonzero_inverse_mult_distrib [symmetric] nonzero_inverse_minus_eq [symmetric])
 120.108    } then show ?thesis by (simp add: nsderiv_def)
 120.109  qed
 120.110  
   121.1 --- a/src/HOL/NSA/HLim.thy	Mon Nov 11 17:34:44 2013 +0100
   121.2 +++ b/src/HOL/NSA/HLim.thy	Mon Nov 11 17:44:21 2013 +0100
   121.3 @@ -71,7 +71,7 @@
   121.4  
   121.5  lemma NSLIM_diff:
   121.6    "\<lbrakk>f -- x --NS> l; g -- x --NS> m\<rbrakk> \<Longrightarrow> (\<lambda>x. f x - g x) -- x --NS> (l - m)"
   121.7 -by (simp only: diff_minus NSLIM_add NSLIM_minus)
   121.8 +  by (simp only: NSLIM_add NSLIM_minus diff_conv_add_uminus)
   121.9  
  121.10  lemma NSLIM_add_minus: "[| f -- x --NS> l; g -- x --NS> m |] ==> (%x. f(x) + -g(x)) -- x --NS> (l + -m)"
  121.11  by (simp only: NSLIM_add NSLIM_minus)
  121.12 @@ -95,7 +95,7 @@
  121.13  
  121.14  lemma NSLIM_zero_cancel: "(%x. f(x) - l) -- x --NS> 0 ==> f -- x --NS> l"
  121.15  apply (drule_tac g = "%x. l" and m = l in NSLIM_add)
  121.16 -apply (auto simp add: diff_minus add_assoc)
  121.17 +apply (auto simp add: add_assoc)
  121.18  done
  121.19  
  121.20  lemma NSLIM_const_not_eq:
  121.21 @@ -243,14 +243,14 @@
  121.22  apply (drule_tac [2] x = "- star_of a + x" in spec, safe, simp)
  121.23  apply (erule mem_infmal_iff [THEN iffD2, THEN Infinitesimal_add_approx_self [THEN approx_sym]])
  121.24  apply (erule_tac [3] approx_minus_iff2 [THEN iffD1])
  121.25 - prefer 2 apply (simp add: add_commute diff_minus [symmetric])
  121.26 + prefer 2 apply (simp add: add_commute)
  121.27  apply (rule_tac x = x in star_cases)
  121.28  apply (rule_tac [2] x = x in star_cases)
  121.29 -apply (auto simp add: starfun star_of_def star_n_minus star_n_add add_assoc approx_refl star_n_zero_num)
  121.30 +apply (auto simp add: starfun star_of_def star_n_minus star_n_add add_assoc star_n_zero_num)
  121.31  done
  121.32  
  121.33  lemma NSLIM_isCont_iff: "(f -- a --NS> f a) = ((%h. f(a + h)) -- 0 --NS> f a)"
  121.34 -by (rule NSLIM_h_iff)
  121.35 +  by (fact NSLIM_h_iff)
  121.36  
  121.37  lemma isNSCont_minus: "isNSCont f a ==> isNSCont (%x. - f x) a"
  121.38  by (simp add: isNSCont_def)
   122.1 --- a/src/HOL/NSA/HSEQ.thy	Mon Nov 11 17:34:44 2013 +0100
   122.2 +++ b/src/HOL/NSA/HSEQ.thy	Mon Nov 11 17:44:21 2013 +0100
   122.3 @@ -73,14 +73,14 @@
   122.4  lemma NSLIMSEQ_minus_cancel: "(%n. -(X n)) ----NS> -a ==> X ----NS> a"
   122.5  by (drule NSLIMSEQ_minus, simp)
   122.6  
   122.7 +lemma NSLIMSEQ_diff:
   122.8 +     "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n - Y n) ----NS> a - b"
   122.9 +  using NSLIMSEQ_add [of X a "- Y" "- b"] by (simp add: NSLIMSEQ_minus fun_Compl_def)
  122.10 +
  122.11  (* FIXME: delete *)
  122.12  lemma NSLIMSEQ_add_minus:
  122.13       "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n + -Y n) ----NS> a + -b"
  122.14 -by (simp add: NSLIMSEQ_add NSLIMSEQ_minus)
  122.15 -
  122.16 -lemma NSLIMSEQ_diff:
  122.17 -     "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n - Y n) ----NS> a - b"
  122.18 -by (simp add: diff_minus NSLIMSEQ_add NSLIMSEQ_minus)
  122.19 +  by (simp add: NSLIMSEQ_diff)
  122.20  
  122.21  lemma NSLIMSEQ_diff_const: "f ----NS> a ==> (%n.(f n - b)) ----NS> a - b"
  122.22  by (simp add: NSLIMSEQ_diff NSLIMSEQ_const)
  122.23 @@ -233,11 +233,11 @@
  122.24  
  122.25  lemma NSLIMSEQ_inverse_real_of_nat_add_minus:
  122.26       "(%n. r + -inverse(real(Suc n))) ----NS> r"
  122.27 -by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat_add_minus)
  122.28 +  using LIMSEQ_inverse_real_of_nat_add_minus by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric])
  122.29  
  122.30  lemma NSLIMSEQ_inverse_real_of_nat_add_minus_mult:
  122.31       "(%n. r*( 1 + -inverse(real(Suc n)))) ----NS> r"
  122.32 -by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat_add_minus_mult)
  122.33 +  using LIMSEQ_inverse_real_of_nat_add_minus_mult by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric])
  122.34  
  122.35  
  122.36  subsection {* Convergence *}
   123.1 --- a/src/HOL/NSA/HSeries.thy	Mon Nov 11 17:34:44 2013 +0100
   123.2 +++ b/src/HOL/NSA/HSeries.thy	Mon Nov 11 17:44:21 2013 +0100
   123.3 @@ -131,7 +131,7 @@
   123.4  apply (auto simp add: approx_refl)
   123.5  apply (drule approx_sym [THEN approx_minus_iff [THEN iffD1]])
   123.6  apply (auto dest: approx_hrabs 
   123.7 -            simp add: sumhr_split_diff diff_minus [symmetric])
   123.8 +            simp add: sumhr_split_diff)
   123.9  done
  123.10  
  123.11  (*----------------------------------------------------------------
  123.12 @@ -172,7 +172,7 @@
  123.13  apply (rule approx_minus_iff [THEN iffD2, THEN approx_sym])
  123.14  apply (rule_tac [2] approx_minus_iff [THEN iffD2])
  123.15  apply (auto dest: approx_hrabs_zero_cancel 
  123.16 -            simp add: sumhr_split_diff diff_minus [symmetric])
  123.17 +            simp add: sumhr_split_diff)
  123.18  done
  123.19  
  123.20  
   124.1 --- a/src/HOL/NSA/HTranscendental.thy	Mon Nov 11 17:34:44 2013 +0100
   124.2 +++ b/src/HOL/NSA/HTranscendental.thy	Mon Nov 11 17:44:21 2013 +0100
   124.3 @@ -258,7 +258,7 @@
   124.4              simp add: mult_assoc)
   124.5  apply (rule approx_add_right_cancel [where d="-1"])
   124.6  apply (rule approx_sym [THEN [2] approx_trans2])
   124.7 -apply (auto simp add: diff_minus mem_infmal_iff)
   124.8 +apply (auto simp add: mem_infmal_iff minus_one [symmetric] simp del: minus_one)
   124.9  done
  124.10  
  124.11  lemma STAR_exp_epsilon [simp]: "( *f* exp) epsilon @= 1"
  124.12 @@ -450,7 +450,7 @@
  124.13  apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
  124.14              simp add: mult_assoc)
  124.15  apply (rule approx_add_right_cancel [where d = "-1"])
  124.16 -apply (simp add: diff_minus)
  124.17 +apply (simp add: minus_one [symmetric] del: minus_one)
  124.18  done
  124.19  
  124.20  lemma STAR_tan_zero [simp]: "( *f* tan) 0 = 0"
  124.21 @@ -587,7 +587,7 @@
  124.22       "x \<in> Infinitesimal ==> ( *f* cos) x @= 1 - x\<^sup>2"
  124.23  apply (rule STAR_cos_Infinitesimal [THEN approx_trans])
  124.24  apply (auto simp add: Infinitesimal_approx_minus [symmetric] 
  124.25 -            diff_minus add_assoc [symmetric] numeral_2_eq_2)
  124.26 +            add_assoc [symmetric] numeral_2_eq_2)
  124.27  done
  124.28  
  124.29  lemma STAR_cos_Infinitesimal_approx2:
   125.1 --- a/src/HOL/NSA/NSA.thy	Mon Nov 11 17:34:44 2013 +0100
   125.2 +++ b/src/HOL/NSA/NSA.thy	Mon Nov 11 17:44:21 2013 +0100
   125.3 @@ -6,7 +6,7 @@
   125.4  header{*Infinite Numbers, Infinitesimals, Infinitely Close Relation*}
   125.5  
   125.6  theory NSA
   125.7 -imports HyperDef
   125.8 +imports HyperDef "~~/src/HOL/Library/Lubs_Glbs"
   125.9  begin
  125.10  
  125.11  definition
  125.12 @@ -368,7 +368,7 @@
  125.13  
  125.14  lemma Infinitesimal_diff:
  125.15       "[| x \<in> Infinitesimal;  y \<in> Infinitesimal |] ==> x-y \<in> Infinitesimal"
  125.16 -by (simp add: diff_minus Infinitesimal_add)
  125.17 +  using Infinitesimal_add [of x "- y"] by simp
  125.18  
  125.19  lemma Infinitesimal_mult:
  125.20    fixes x y :: "'a::real_normed_algebra star"
  125.21 @@ -491,7 +491,9 @@
  125.22       "[|(x::hypreal) \<in> HInfinite; y \<le> 0; x \<le> 0|] ==> (x + y): HInfinite"
  125.23  apply (drule HInfinite_minus_iff [THEN iffD2])
  125.24  apply (rule HInfinite_minus_iff [THEN iffD1])
  125.25 -apply (auto intro: HInfinite_add_ge_zero)
  125.26 +apply (simp only: minus_add add.commute)
  125.27 +apply (rule HInfinite_add_ge_zero)
  125.28 +apply simp_all
  125.29  done
  125.30  
  125.31  lemma HInfinite_add_lt_zero:
  125.32 @@ -620,7 +622,7 @@
  125.33  by (simp add: approx_def)
  125.34  
  125.35  lemma approx_minus_iff2: " (x @= y) = (-y + x @= 0)"
  125.36 -by (simp add: approx_def diff_minus add_commute)
  125.37 +by (simp add: approx_def add_commute)
  125.38  
  125.39  lemma approx_refl [iff]: "x @= x"
  125.40  by (simp add: approx_def Infinitesimal_def)
  125.41 @@ -637,7 +639,7 @@
  125.42  lemma approx_trans: "[| x @= y; y @= z |] ==> x @= z"
  125.43  apply (simp add: approx_def)
  125.44  apply (drule (1) Infinitesimal_add)
  125.45 -apply (simp add: diff_minus)
  125.46 +apply simp
  125.47  done
  125.48  
  125.49  lemma approx_trans2: "[| r @= x; s @= x |] ==> r @= s"
  125.50 @@ -687,7 +689,7 @@
  125.51  lemma approx_minus: "a @= b ==> -a @= -b"
  125.52  apply (rule approx_minus_iff [THEN iffD2, THEN approx_sym])
  125.53  apply (drule approx_minus_iff [THEN iffD1])
  125.54 -apply (simp add: add_commute diff_minus)
  125.55 +apply (simp add: add_commute)
  125.56  done
  125.57  
  125.58  lemma approx_minus2: "-a @= -b ==> a @= b"
  125.59 @@ -700,7 +702,7 @@
  125.60  by (blast intro!: approx_add approx_minus)
  125.61  
  125.62  lemma approx_diff: "[| a @= b; c @= d |] ==> a - c @= b - d"
  125.63 -by (simp only: diff_minus approx_add approx_minus)
  125.64 +  using approx_add [of a b "- c" "- d"] by simp
  125.65  
  125.66  lemma approx_mult1:
  125.67    fixes a b c :: "'a::real_normed_algebra star"
  125.68 @@ -1213,7 +1215,9 @@
  125.69           r \<in> Reals;  0 < r |]
  125.70        ==> -(x + -t) \<le> r"
  125.71  apply (subgoal_tac "(t + -r \<le> x)") 
  125.72 -apply (auto intro: lemma_st_part_le2)
  125.73 +apply simp
  125.74 +apply (rule lemma_st_part_le2)
  125.75 +apply auto
  125.76  done
  125.77  
  125.78  lemma lemma_SReal_ub:
  125.79 @@ -1238,7 +1242,7 @@
  125.80        ==> x + -t \<noteq> r"
  125.81  apply auto
  125.82  apply (frule isLubD1a [THEN Reals_minus])
  125.83 -apply (drule Reals_add_cancel, assumption)
  125.84 +using Reals_add_cancel [of x "- t"] apply simp
  125.85  apply (drule_tac x = x in lemma_SReal_lub)
  125.86  apply (drule hypreal_isLub_unique, assumption, auto)
  125.87  done
  125.88 @@ -1250,8 +1254,7 @@
  125.89        ==> -(x + -t) \<noteq> r"
  125.90  apply (auto)
  125.91  apply (frule isLubD1a)
  125.92 -apply (drule Reals_add_cancel, assumption)
  125.93 -apply (drule_tac a = "-x" in Reals_minus, simp)
  125.94 +using Reals_add_cancel [of "- x" t] apply simp
  125.95  apply (drule_tac x = x in lemma_SReal_lub)
  125.96  apply (drule hypreal_isLub_unique, assumption, auto)
  125.97  done
   126.1 --- a/src/HOL/NSA/NSCA.thy	Mon Nov 11 17:34:44 2013 +0100
   126.2 +++ b/src/HOL/NSA/NSCA.thy	Mon Nov 11 17:44:21 2013 +0100
   126.3 @@ -165,7 +165,7 @@
   126.4  
   126.5  lemma approx_hcmod_approx_zero: "(x @= y) = (hcmod (y - x) @= 0)"
   126.6  apply (subst hnorm_minus_commute)
   126.7 -apply (simp add: approx_def Infinitesimal_hcmod_iff diff_minus)
   126.8 +apply (simp add: approx_def Infinitesimal_hcmod_iff)
   126.9  done
  126.10  
  126.11  lemma approx_approx_zero_iff: "(x @= 0) = (hcmod x @= 0)"
  126.12 @@ -178,14 +178,14 @@
  126.13       "u @= 0 ==> hcmod(x + u) - hcmod x \<in> Infinitesimal"
  126.14  apply (drule approx_approx_zero_iff [THEN iffD1])
  126.15  apply (rule_tac e = "hcmod u" and e' = "- hcmod u" in Infinitesimal_interval2)
  126.16 -apply (auto simp add: mem_infmal_iff [symmetric] diff_minus)
  126.17 +apply (auto simp add: mem_infmal_iff [symmetric])
  126.18  apply (rule_tac c1 = "hcmod x" in add_le_cancel_left [THEN iffD1])
  126.19 -apply (auto simp add: diff_minus [symmetric])
  126.20 +apply auto
  126.21  done
  126.22  
  126.23  lemma approx_hcmod_add_hcmod: "u @= 0 ==> hcmod(x + u) @= hcmod x"
  126.24  apply (rule approx_minus_iff [THEN iffD2])
  126.25 -apply (auto intro: Infinitesimal_hcmod_add_diff simp add: mem_infmal_iff [symmetric] diff_minus [symmetric])
  126.26 +apply (auto intro: Infinitesimal_hcmod_add_diff simp add: mem_infmal_iff [symmetric])
  126.27  done
  126.28  
  126.29  
   127.1 --- a/src/HOL/NSA/StarDef.thy	Mon Nov 11 17:34:44 2013 +0100
   127.2 +++ b/src/HOL/NSA/StarDef.thy	Mon Nov 11 17:44:21 2013 +0100
   127.3 @@ -803,7 +803,7 @@
   127.4  instance star :: (ab_group_add) ab_group_add
   127.5  apply (intro_classes)
   127.6  apply (transfer, rule left_minus)
   127.7 -apply (transfer, rule diff_minus)
   127.8 +apply (transfer, rule diff_conv_add_uminus)
   127.9  done
  127.10  
  127.11  instance star :: (ordered_ab_semigroup_add) ordered_ab_semigroup_add
   128.1 --- a/src/HOL/Nat.thy	Mon Nov 11 17:34:44 2013 +0100
   128.2 +++ b/src/HOL/Nat.thy	Mon Nov 11 17:44:21 2013 +0100
   128.3 @@ -327,7 +327,7 @@
   128.4     apply auto
   128.5    done
   128.6  
   128.7 -lemma one_eq_mult_iff [simp,no_atp]: "(Suc 0 = m * n) = (m = Suc 0 & n = Suc 0)"
   128.8 +lemma one_eq_mult_iff [simp]: "(Suc 0 = m * n) = (m = Suc 0 & n = Suc 0)"
   128.9    apply (rule trans)
  128.10    apply (rule_tac [2] mult_eq_1_iff, fastforce)
  128.11    done
  128.12 @@ -369,8 +369,8 @@
  128.13  | "Suc m \<le> n \<longleftrightarrow> (case n of 0 \<Rightarrow> False | Suc n \<Rightarrow> m \<le> n)"
  128.14  
  128.15  declare less_eq_nat.simps [simp del]
  128.16 -lemma [code]: "(0\<Colon>nat) \<le> n \<longleftrightarrow> True" by (simp add: less_eq_nat.simps)
  128.17  lemma le0 [iff]: "0 \<le> (n\<Colon>nat)" by (simp add: less_eq_nat.simps)
  128.18 +lemma [code]: "(0\<Colon>nat) \<le> n \<longleftrightarrow> True" by simp
  128.19  
  128.20  definition less_nat where
  128.21    less_eq_Suc_le: "n < m \<longleftrightarrow> Suc n \<le> m"
  128.22 @@ -491,7 +491,7 @@
  128.23  lemma less_Suc0 [iff]: "(n < Suc 0) = (n = 0)"
  128.24    by (simp add: less_Suc_eq)
  128.25  
  128.26 -lemma less_one [iff, no_atp]: "(n < (1::nat)) = (n = 0)"
  128.27 +lemma less_one [iff]: "(n < (1::nat)) = (n = 0)"
  128.28    unfolding One_nat_def by (rule less_Suc0)
  128.29  
  128.30  lemma Suc_mono: "m < n ==> Suc m < Suc n"
  128.31 @@ -659,7 +659,7 @@
  128.32  lemma gr0_conv_Suc: "(0 < n) = (\<exists>m. n = Suc m)"
  128.33  by (fast intro: not0_implies_Suc)
  128.34  
  128.35 -lemma not_gr0 [iff,no_atp]: "!!n::nat. (~ (0 < n)) = (n = 0)"
  128.36 +lemma not_gr0 [iff]: "!!n::nat. (~ (0 < n)) = (n = 0)"
  128.37  using neq0_conv by blast
  128.38  
  128.39  lemma Suc_le_D: "(Suc n \<le> m') ==> (? m. m' = Suc m)"
  128.40 @@ -1396,10 +1396,10 @@
  128.41  
  128.42  text{*Special cases where either operand is zero*}
  128.43  
  128.44 -lemma of_nat_0_eq_iff [simp, no_atp]: "0 = of_nat n \<longleftrightarrow> 0 = n"
  128.45 +lemma of_nat_0_eq_iff [simp]: "0 = of_nat n \<longleftrightarrow> 0 = n"
  128.46    by (fact of_nat_eq_iff [of 0 n, unfolded of_nat_0])
  128.47  
  128.48 -lemma of_nat_eq_0_iff [simp, no_atp]: "of_nat m = 0 \<longleftrightarrow> m = 0"
  128.49 +lemma of_nat_eq_0_iff [simp]: "of_nat m = 0 \<longleftrightarrow> m = 0"
  128.50    by (fact of_nat_eq_iff [of m 0, unfolded of_nat_0])
  128.51  
  128.52  end
  128.53 @@ -1432,7 +1432,7 @@
  128.54  
  128.55  text{*Special cases where either operand is zero*}
  128.56  
  128.57 -lemma of_nat_le_0_iff [simp, no_atp]: "of_nat m \<le> 0 \<longleftrightarrow> m = 0"
  128.58 +lemma of_nat_le_0_iff [simp]: "of_nat m \<le> 0 \<longleftrightarrow> m = 0"
  128.59    by (rule of_nat_le_iff [of _ 0, simplified])
  128.60  
  128.61  lemma of_nat_0_less_iff [simp]: "0 < of_nat n \<longleftrightarrow> 0 < n"
  128.62 @@ -1872,12 +1872,12 @@
  128.63    shows "m dvd n + q \<longleftrightarrow> m dvd n"
  128.64    using assms by (simp add: dvd_plus_eq_right add_commute [of n])
  128.65  
  128.66 -lemma less_dvd_minus:
  128.67 +lemma less_eq_dvd_minus:
  128.68    fixes m n :: nat
  128.69 -  assumes "m < n"
  128.70 -  shows "m dvd n \<longleftrightarrow> m dvd (n - m)"
  128.71 +  assumes "m \<le> n"
  128.72 +  shows "m dvd n \<longleftrightarrow> m dvd n - m"
  128.73  proof -
  128.74 -  from assms have "n = m + (n - m)" by arith
  128.75 +  from assms have "n = m + (n - m)" by simp
  128.76    then obtain q where "n = m + q" ..
  128.77    then show ?thesis by (simp add: dvd_reduce add_commute [of m])
  128.78  qed
   129.1 --- a/src/HOL/Nitpick.thy	Mon Nov 11 17:34:44 2013 +0100
   129.2 +++ b/src/HOL/Nitpick.thy	Mon Nov 11 17:44:21 2013 +0100
   129.3 @@ -33,7 +33,7 @@
   129.4  Alternative definitions.
   129.5  *}
   129.6  
   129.7 -lemma Ex1_unfold [nitpick_unfold, no_atp]:
   129.8 +lemma Ex1_unfold [nitpick_unfold]:
   129.9  "Ex1 P \<equiv> \<exists>x. {x. P x} = {x}"
  129.10  apply (rule eq_reflection)
  129.11  apply (simp add: Ex1_def set_eq_iff)
  129.12 @@ -46,18 +46,18 @@
  129.13   apply (erule_tac x = y in allE)
  129.14  by auto
  129.15  
  129.16 -lemma rtrancl_unfold [nitpick_unfold, no_atp]: "r\<^sup>* \<equiv> (r\<^sup>+)\<^sup>="
  129.17 +lemma rtrancl_unfold [nitpick_unfold]: "r\<^sup>* \<equiv> (r\<^sup>+)\<^sup>="
  129.18    by (simp only: rtrancl_trancl_reflcl)
  129.19  
  129.20 -lemma rtranclp_unfold [nitpick_unfold, no_atp]:
  129.21 +lemma rtranclp_unfold [nitpick_unfold]:
  129.22  "rtranclp r a b \<equiv> (a = b \<or> tranclp r a b)"
  129.23  by (rule eq_reflection) (auto dest: rtranclpD)
  129.24  
  129.25 -lemma tranclp_unfold [nitpick_unfold, no_atp]:
  129.26 +lemma tranclp_unfold [nitpick_unfold]:
  129.27  "tranclp r a b \<equiv> (a, b) \<in> trancl {(x, y). r x y}"
  129.28  by (simp add: trancl_def)
  129.29  
  129.30 -lemma [nitpick_simp, no_atp]:
  129.31 +lemma [nitpick_simp]:
  129.32  "of_nat n = (if n = 0 then 0 else 1 + of_nat (n - 1))"
  129.33  by (cases n) auto
  129.34  
  129.35 @@ -85,18 +85,18 @@
  129.36  \textit{specialize} optimization.
  129.37  *}
  129.38  
  129.39 -lemma The_psimp [nitpick_psimp, no_atp]:
  129.40 +lemma The_psimp [nitpick_psimp]:
  129.41    "P = (op =) x \<Longrightarrow> The P = x"
  129.42    by auto
  129.43  
  129.44 -lemma Eps_psimp [nitpick_psimp, no_atp]:
  129.45 +lemma Eps_psimp [nitpick_psimp]:
  129.46  "\<lbrakk>P x; \<not> P y; Eps P = y\<rbrakk> \<Longrightarrow> Eps P = x"
  129.47  apply (cases "P (Eps P)")
  129.48   apply auto
  129.49  apply (erule contrapos_np)
  129.50  by (rule someI)
  129.51  
  129.52 -lemma unit_case_unfold [nitpick_unfold, no_atp]:
  129.53 +lemma unit_case_unfold [nitpick_unfold]:
  129.54  "unit_case x u \<equiv> x"
  129.55  apply (subgoal_tac "u = ()")
  129.56   apply (simp only: unit.cases)
  129.57 @@ -104,14 +104,14 @@
  129.58  
  129.59  declare unit.cases [nitpick_simp del]
  129.60  
  129.61 -lemma nat_case_unfold [nitpick_unfold, no_atp]:
  129.62 +lemma nat_case_unfold [nitpick_unfold]:
  129.63  "nat_case x f n \<equiv> if n = 0 then x else f (n - 1)"
  129.64  apply (rule eq_reflection)
  129.65  by (cases n) auto
  129.66  
  129.67  declare nat.cases [nitpick_simp del]
  129.68  
  129.69 -lemma list_size_simp [nitpick_simp, no_atp]:
  129.70 +lemma list_size_simp [nitpick_simp]:
  129.71  "list_size f xs = (if xs = [] then 0
  129.72                     else Suc (f (hd xs) + list_size f (tl xs)))"
  129.73  "size xs = (if xs = [] then 0 else Suc (size (tl xs)))"
   130.1 --- a/src/HOL/Num.thy	Mon Nov 11 17:34:44 2013 +0100
   130.2 +++ b/src/HOL/Num.thy	Mon Nov 11 17:44:21 2013 +0100
   130.3 @@ -407,7 +407,7 @@
   130.4    apply (simp add: add_assoc [symmetric], simp add: add_assoc)
   130.5    apply (rule_tac a=x in add_left_imp_eq)
   130.6    apply (rule_tac a=x in add_right_imp_eq)
   130.7 -  apply (simp add: add_assoc minus_add_cancel add_minus_cancel)
   130.8 +  apply (simp add: add_assoc)
   130.9    apply (simp add: add_assoc, simp add: add_assoc [symmetric])
  130.10    done
  130.11  
  130.12 @@ -418,7 +418,7 @@
  130.13  lemmas is_num_normalize =
  130.14    add_assoc is_num_add_commute is_num_add_left_commute
  130.15    is_num.intros is_num_numeral
  130.16 -  diff_minus minus_add add_minus_cancel minus_add_cancel
  130.17 +  minus_add
  130.18  
  130.19  definition dbl :: "'a \<Rightarrow> 'a" where "dbl x = x + x"
  130.20  definition dbl_inc :: "'a \<Rightarrow> 'a" where "dbl_inc x = x + x + 1"
  130.21 @@ -435,24 +435,21 @@
  130.22    "dbl 0 = 0"
  130.23    "dbl 1 = 2"
  130.24    "dbl (numeral k) = numeral (Bit0 k)"
  130.25 -  unfolding dbl_def neg_numeral_def numeral.simps
  130.26 -  by (simp_all add: minus_add)
  130.27 +  by (simp_all add: dbl_def neg_numeral_def numeral.simps minus_add)
  130.28  
  130.29  lemma dbl_inc_simps [simp]:
  130.30    "dbl_inc (neg_numeral k) = neg_numeral (BitM k)"
  130.31    "dbl_inc 0 = 1"
  130.32    "dbl_inc 1 = 3"
  130.33    "dbl_inc (numeral k) = numeral (Bit1 k)"
  130.34 -  unfolding dbl_inc_def neg_numeral_def numeral.simps numeral_BitM
  130.35 -  by (simp_all add: is_num_normalize)
  130.36 +  by (simp_all add: dbl_inc_def neg_numeral_def numeral.simps numeral_BitM is_num_normalize algebra_simps del: add_uminus_conv_diff)
  130.37  
  130.38  lemma dbl_dec_simps [simp]:
  130.39    "dbl_dec (neg_numeral k) = neg_numeral (Bit1 k)"
  130.40    "dbl_dec 0 = -1"
  130.41    "dbl_dec 1 = 1"
  130.42    "dbl_dec (numeral k) = numeral (BitM k)"
  130.43 -  unfolding dbl_dec_def neg_numeral_def numeral.simps numeral_BitM
  130.44 -  by (simp_all add: is_num_normalize)
  130.45 +  by (simp_all add: dbl_dec_def neg_numeral_def numeral.simps numeral_BitM is_num_normalize)
  130.46  
  130.47  lemma sub_num_simps [simp]:
  130.48    "sub One One = 0"
  130.49 @@ -464,38 +461,35 @@
  130.50    "sub (Bit0 k) (Bit1 l) = dbl_dec (sub k l)"
  130.51    "sub (Bit1 k) (Bit0 l) = dbl_inc (sub k l)"
  130.52    "sub (Bit1 k) (Bit1 l) = dbl (sub k l)"
  130.53 -  unfolding dbl_def dbl_dec_def dbl_inc_def sub_def
  130.54 -  unfolding neg_numeral_def numeral.simps numeral_BitM
  130.55 -  by (simp_all add: is_num_normalize)
  130.56 +  by (simp_all add: dbl_def dbl_dec_def dbl_inc_def sub_def neg_numeral_def numeral.simps
  130.57 +    numeral_BitM is_num_normalize del: add_uminus_conv_diff add: diff_conv_add_uminus)
  130.58  
  130.59  lemma add_neg_numeral_simps:
  130.60    "numeral m + neg_numeral n = sub m n"
  130.61    "neg_numeral m + numeral n = sub n m"
  130.62    "neg_numeral m + neg_numeral n = neg_numeral (m + n)"
  130.63 -  unfolding sub_def diff_minus neg_numeral_def numeral_add numeral.simps
  130.64 -  by (simp_all add: is_num_normalize)
  130.65 +  by (simp_all add: sub_def neg_numeral_def numeral_add numeral.simps is_num_normalize
  130.66 +    del: add_uminus_conv_diff add: diff_conv_add_uminus)
  130.67  
  130.68  lemma add_neg_numeral_special:
  130.69    "1 + neg_numeral m = sub One m"
  130.70    "neg_numeral m + 1 = sub One m"
  130.71 -  unfolding sub_def diff_minus neg_numeral_def numeral_add numeral.simps
  130.72 -  by (simp_all add: is_num_normalize)
  130.73 +  by (simp_all add: sub_def neg_numeral_def numeral_add numeral.simps is_num_normalize)
  130.74  
  130.75  lemma diff_numeral_simps:
  130.76    "numeral m - numeral n = sub m n"
  130.77    "numeral m - neg_numeral n = numeral (m + n)"
  130.78    "neg_numeral m - numeral n = neg_numeral (m + n)"
  130.79    "neg_numeral m - neg_numeral n = sub n m"
  130.80 -  unfolding neg_numeral_def sub_def diff_minus numeral_add numeral.simps
  130.81 -  by (simp_all add: is_num_normalize)
  130.82 +  by (simp_all add: neg_numeral_def sub_def numeral_add numeral.simps is_num_normalize
  130.83 +    del: add_uminus_conv_diff add: diff_conv_add_uminus)
  130.84  
  130.85  lemma diff_numeral_special:
  130.86    "1 - numeral n = sub One n"
  130.87    "1 - neg_numeral n = numeral (One + n)"
  130.88    "numeral m - 1 = sub m One"
  130.89    "neg_numeral m - 1 = neg_numeral (m + One)"
  130.90 -  unfolding neg_numeral_def sub_def diff_minus numeral_add numeral.simps
  130.91 -  by (simp_all add: is_num_normalize)
  130.92 +  by (simp_all add: neg_numeral_def sub_def numeral_add numeral.simps add: is_num_normalize)
  130.93  
  130.94  lemma minus_one: "- 1 = -1"
  130.95    unfolding neg_numeral_def numeral.simps ..
  130.96 @@ -1078,6 +1072,16 @@
  130.97    BitM.simps dbl_simps dbl_inc_simps dbl_dec_simps
  130.98    abs_zero abs_one arith_extra_simps
  130.99  
 130.100 +lemmas more_arith_simps =
 130.101 +  neg_le_iff_le
 130.102 +  minus_zero left_minus right_minus
 130.103 +  mult_1_left mult_1_right
 130.104 +  mult_minus_left mult_minus_right
 130.105 +  minus_add_distrib minus_minus mult_assoc
 130.106 +
 130.107 +lemmas of_nat_simps =
 130.108 +  of_nat_0 of_nat_1 of_nat_Suc of_nat_add of_nat_mult
 130.109 +
 130.110  text {* Simplification of relational operations *}
 130.111  
 130.112  lemmas eq_numeral_extra =
 130.113 @@ -1089,6 +1093,38 @@
 130.114    less_numeral_simps less_neg_numeral_simps less_numeral_extra
 130.115    eq_numeral_simps eq_neg_numeral_simps eq_numeral_extra
 130.116  
 130.117 +lemma Let_numeral [simp]: "Let (numeral v) f = f (numeral v)"
 130.118 +  -- {* Unfold all @{text let}s involving constants *}
 130.119 +  unfolding Let_def ..
 130.120 +
 130.121 +lemma Let_neg_numeral [simp]: "Let (neg_numeral v) f = f (neg_numeral v)"
 130.122 +  -- {* Unfold all @{text let}s involving constants *}
 130.123 +  unfolding Let_def ..
 130.124 +
 130.125 +declaration {*
 130.126 +let 
 130.127 +  fun number_of thy T n =
 130.128 +    if not (Sign.of_sort thy (T, @{sort numeral}))
 130.129 +    then raise CTERM ("number_of", [])
 130.130 +    else Numeral.mk_cnumber (Thm.ctyp_of thy T) n;
 130.131 +in
 130.132 +  K (
 130.133 +    Lin_Arith.add_simps (@{thms arith_simps} @ @{thms more_arith_simps}
 130.134 +      @ @{thms rel_simps}
 130.135 +      @ @{thms pred_numeral_simps}
 130.136 +      @ @{thms arith_special numeral_One}
 130.137 +      @ @{thms of_nat_simps})
 130.138 +    #> Lin_Arith.add_simps [@{thm Suc_numeral},
 130.139 +      @{thm Let_numeral}, @{thm Let_neg_numeral}, @{thm Let_0}, @{thm Let_1},
 130.140 +      @{thm le_Suc_numeral}, @{thm le_numeral_Suc},
 130.141 +      @{thm less_Suc_numeral}, @{thm less_numeral_Suc},
 130.142 +      @{thm Suc_eq_numeral}, @{thm eq_numeral_Suc},
 130.143 +      @{thm mult_Suc}, @{thm mult_Suc_right},
 130.144 +      @{thm of_nat_numeral}]
 130.145 +    #> Lin_Arith.set_number_of number_of)
 130.146 +end
 130.147 +*}
 130.148 +
 130.149  
 130.150  subsubsection {* Simplification of arithmetic when nested to the right. *}
 130.151  
 130.152 @@ -1119,4 +1155,3 @@
 130.153  
 130.154  end
 130.155  
 130.156 -
   131.1 --- a/src/HOL/Number_Theory/Cong.thy	Mon Nov 11 17:34:44 2013 +0100
   131.2 +++ b/src/HOL/Number_Theory/Cong.thy	Mon Nov 11 17:44:21 2013 +0100
   131.3 @@ -543,7 +543,8 @@
   131.4    apply (subgoal_tac "a * b = (-a * -b)")
   131.5    apply (erule ssubst)
   131.6    apply (subst zmod_zmult2_eq)
   131.7 -  apply (auto simp add: mod_add_left_eq)
   131.8 +  apply (auto simp add: mod_add_left_eq mod_minus_right div_minus_right)
   131.9 +  apply (metis mod_diff_left_eq mod_diff_right_eq mod_mult_self1_is_0 semiring_numeral_div_class.diff_zero)+
  131.10    done
  131.11  
  131.12  lemma cong_to_1_nat: "([(a::nat) = 1] (mod n)) \<Longrightarrow> (n dvd (a - 1))"
   132.1 --- a/src/HOL/Number_Theory/Eratosthenes.thy	Mon Nov 11 17:34:44 2013 +0100
   132.2 +++ b/src/HOL/Number_Theory/Eratosthenes.thy	Mon Nov 11 17:44:21 2013 +0100
   132.3 @@ -75,7 +75,7 @@
   132.4  lemma numbers_of_marks_mark_out:
   132.5    "numbers_of_marks n (mark_out m bs) = {q \<in> numbers_of_marks n bs. \<not> Suc m dvd Suc q - n}"
   132.6    by (auto simp add: numbers_of_marks_def mark_out_def in_set_enumerate_eq image_iff
   132.7 -    nth_enumerate_eq less_dvd_minus)
   132.8 +    nth_enumerate_eq less_eq_dvd_minus)
   132.9  
  132.10  
  132.11  text {* Auxiliary operation for efficient implementation  *}
  132.12 @@ -128,7 +128,7 @@
  132.13      by (simp add: mark_out_aux_def)
  132.14    show ?thesis2
  132.15      by (auto simp add: mark_out_code [symmetric] mark_out_aux_def mark_out_def
  132.16 -      enumerate_Suc_eq in_set_enumerate_eq less_dvd_minus)
  132.17 +      enumerate_Suc_eq in_set_enumerate_eq less_eq_dvd_minus)
  132.18    { def v \<equiv> "Suc m" and w \<equiv> "Suc n"
  132.19      fix q
  132.20      assume "m + n \<le> q"
   133.1 --- a/src/HOL/Number_Theory/Primes.thy	Mon Nov 11 17:34:44 2013 +0100
   133.2 +++ b/src/HOL/Number_Theory/Primes.thy	Mon Nov 11 17:44:21 2013 +0100
   133.3 @@ -74,8 +74,9 @@
   133.4  subsection {* Primes *}
   133.5  
   133.6  lemma prime_odd_nat: "prime (p::nat) \<Longrightarrow> p > 2 \<Longrightarrow> odd p"
   133.7 -  unfolding prime_nat_def
   133.8 -  by (metis gcd_lcm_complete_lattice_nat.bot_least nat_even_iff_2_dvd nat_neq_iff odd_1_nat)
   133.9 +  apply (auto simp add: prime_nat_def even_def dvd_eq_mod_eq_0)
  133.10 +  apply (metis dvd_eq_mod_eq_0 even_Suc even_def mod_by_1 nat_dvd_not_less not_mod_2_eq_0_eq_1 zero_less_numeral)
  133.11 +  done
  133.12  
  133.13  lemma prime_odd_int: "prime (p::int) \<Longrightarrow> p > 2 \<Longrightarrow> odd p"
  133.14    unfolding prime_int_def
   134.1 --- a/src/HOL/Numeral_Simprocs.thy	Mon Nov 11 17:34:44 2013 +0100
   134.2 +++ b/src/HOL/Numeral_Simprocs.thy	Mon Nov 11 17:44:21 2013 +0100
   134.3 @@ -13,7 +13,7 @@
   134.4  ML_file "~~/src/Provers/Arith/extract_common_term.ML"
   134.5  
   134.6  lemmas semiring_norm =
   134.7 -  Let_def arith_simps nat_arith rel_simps
   134.8 +  Let_def arith_simps diff_nat_numeral rel_simps
   134.9    if_False if_True
  134.10    add_0 add_Suc add_numeral_left
  134.11    add_neg_numeral_left mult_numeral_left
  134.12 @@ -278,27 +278,8 @@
  134.13    ("((l::nat) * m) dvd n" | "(l::nat) dvd (m * n)") =
  134.14    {* fn phi => Nat_Numeral_Simprocs.dvd_cancel_factor *}
  134.15  
  134.16 -(* FIXME: duplicate rule warnings for:
  134.17 -  ring_distribs
  134.18 -  numeral_plus_numeral numeral_times_numeral
  134.19 -  numeral_eq_iff numeral_less_iff numeral_le_iff
  134.20 -  numeral_neq_zero zero_neq_numeral zero_less_numeral
  134.21 -  if_True if_False *)
  134.22  declaration {* 
  134.23 -  K (Lin_Arith.add_simps ([@{thm Suc_numeral}, @{thm int_numeral}])
  134.24 -  #> Lin_Arith.add_simps (@{thms ring_distribs} @ [@{thm Let_numeral}, @{thm Let_neg_numeral}, @{thm Let_0}, @{thm Let_1},
  134.25 -     @{thm nat_0}, @{thm nat_1},
  134.26 -     @{thm numeral_plus_numeral}, @{thm diff_nat_numeral}, @{thm numeral_times_numeral},
  134.27 -     @{thm numeral_eq_iff}, @{thm numeral_less_iff}, @{thm numeral_le_iff},
  134.28 -     @{thm le_Suc_numeral}, @{thm le_numeral_Suc},
  134.29 -     @{thm less_Suc_numeral}, @{thm less_numeral_Suc},
  134.30 -     @{thm Suc_eq_numeral}, @{thm eq_numeral_Suc},
  134.31 -     @{thm mult_Suc}, @{thm mult_Suc_right},
  134.32 -     @{thm add_Suc}, @{thm add_Suc_right},
  134.33 -     @{thm numeral_neq_zero}, @{thm zero_neq_numeral}, @{thm zero_less_numeral},
  134.34 -     @{thm of_int_numeral}, @{thm of_nat_numeral}, @{thm nat_numeral},
  134.35 -     @{thm if_True}, @{thm if_False}])
  134.36 -  #> Lin_Arith.add_simprocs
  134.37 +  K (Lin_Arith.add_simprocs
  134.38        [@{simproc semiring_assoc_fold},
  134.39         @{simproc int_combine_numerals},
  134.40         @{simproc inteq_cancel_numerals},
   135.1 --- a/src/HOL/Old_Number_Theory/Pocklington.thy	Mon Nov 11 17:34:44 2013 +0100
   135.2 +++ b/src/HOL/Old_Number_Theory/Pocklington.thy	Mon Nov 11 17:44:21 2013 +0100
   135.3 @@ -348,15 +348,11 @@
   135.4    let ?w = "x mod (a*b)"
   135.5    have wab: "?w < a*b" by (simp add: mod_less_divisor[OF abpos])
   135.6    from xq12(1) have "?w mod a = ((m + q1 * a) mod (a*b)) mod a" by simp
   135.7 -  also have "\<dots> = m mod a" apply (simp add: mod_mult2_eq)
   135.8 -    apply (subst mod_add_left_eq)
   135.9 -    by simp
  135.10 +  also have "\<dots> = m mod a" by (simp add: mod_mult2_eq)
  135.11    finally have th1: "[?w = m] (mod a)" by (simp add: modeq_def)
  135.12    from xq12(2) have "?w mod b = ((n + q2 * b) mod (a*b)) mod b" by simp
  135.13    also have "\<dots> = ((n + q2 * b) mod (b*a)) mod b" by (simp add: mult_commute)
  135.14 -  also have "\<dots> = n mod b" apply (simp add: mod_mult2_eq)
  135.15 -    apply (subst mod_add_left_eq)
  135.16 -    by simp
  135.17 +  also have "\<dots> = n mod b" by (simp add: mod_mult2_eq)
  135.18    finally have th2: "[?w = n] (mod b)" by (simp add: modeq_def)
  135.19    {fix y assume H: "y < a*b" "[y = m] (mod a)" "[y = n] (mod b)"
  135.20      with th1 th2 have H': "[y = ?w] (mod a)" "[y = ?w] (mod b)"
   136.1 --- a/src/HOL/Orderings.thy	Mon Nov 11 17:34:44 2013 +0100
   136.2 +++ b/src/HOL/Orderings.thy	Mon Nov 11 17:44:21 2013 +0100
   136.3 @@ -998,7 +998,7 @@
   136.4      (!!x y. x > y ==> f x > f y) ==> f a > c"
   136.5  by (subgoal_tac "f a > f b", force, force)
   136.6  
   136.7 -lemmas xtrans = xt1 xt2 xt3 xt4 xt5 xt6 xt7 xt8 xt9 [no_atp]
   136.8 +lemmas xtrans = xt1 xt2 xt3 xt4 xt5 xt6 xt7 xt8 xt9
   136.9  
  136.10  (* 
  136.11    Since "a >= b" abbreviates "b <= a", the abbreviation "..." stands
   137.1 --- a/src/HOL/Parity.thy	Mon Nov 11 17:34:44 2013 +0100
   137.2 +++ b/src/HOL/Parity.thy	Mon Nov 11 17:44:21 2013 +0100
   137.3 @@ -9,26 +9,52 @@
   137.4  imports Main
   137.5  begin
   137.6  
   137.7 -class even_odd = 
   137.8 -  fixes even :: "'a \<Rightarrow> bool"
   137.9 -
  137.10 -abbreviation
  137.11 -  odd :: "'a\<Colon>even_odd \<Rightarrow> bool" where
  137.12 -  "odd x \<equiv> \<not> even x"
  137.13 -
  137.14 -instantiation nat and int  :: even_odd
  137.15 +class even_odd = semiring_div_parity
  137.16  begin
  137.17  
  137.18 -definition
  137.19 -  even_def [presburger]: "even x \<longleftrightarrow> (x\<Colon>int) mod 2 = 0"
  137.20 +definition even :: "'a \<Rightarrow> bool"
  137.21 +where
  137.22 +  even_def [presburger]: "even a \<longleftrightarrow> a mod 2 = 0"
  137.23  
  137.24 -definition
  137.25 -  even_nat_def [presburger]: "even x \<longleftrightarrow> even (int x)"
  137.26 +lemma even_iff_2_dvd [algebra]:
  137.27 +  "even a \<longleftrightarrow> 2 dvd a"
  137.28 +  by (simp add: even_def dvd_eq_mod_eq_0)
  137.29  
  137.30 -instance ..
  137.31 +lemma even_zero [simp]:
  137.32 +  "even 0"
  137.33 +  by (simp add: even_def)
  137.34 +
  137.35 +lemma even_times_anything:
  137.36 +  "even a \<Longrightarrow> even (a * b)"
  137.37 +  by (simp add: even_iff_2_dvd)
  137.38 +
  137.39 +lemma anything_times_even:
  137.40 +  "even a \<Longrightarrow> even (b * a)"
  137.41 +  by (simp add: even_iff_2_dvd)
  137.42 +
  137.43 +abbreviation odd :: "'a \<Rightarrow> bool"
  137.44 +where
  137.45 +  "odd a \<equiv> \<not> even a"
  137.46 +
  137.47 +lemma odd_times_odd:
  137.48 +  "odd a \<Longrightarrow> odd b \<Longrightarrow> odd (a * b)" 
  137.49 +  by (auto simp add: even_def mod_mult_left_eq)
  137.50 +
  137.51 +lemma even_product [simp, presburger]:
  137.52 +  "even (a * b) \<longleftrightarrow> even a \<or> even b"
  137.53 +  apply (auto simp add: even_times_anything anything_times_even)
  137.54 +  apply (rule ccontr)
  137.55 +  apply (auto simp add: odd_times_odd)
  137.56 +  done
  137.57  
  137.58  end
  137.59  
  137.60 +instance nat and int  :: even_odd ..
  137.61 +
  137.62 +lemma even_nat_def [presburger]:
  137.63 +  "even x \<longleftrightarrow> even (int x)"
  137.64 +  by (auto simp add: even_def int_eq_iff int_mult nat_mult_distrib)
  137.65 +  
  137.66  lemma transfer_int_nat_relations:
  137.67    "even (int x) \<longleftrightarrow> even x"
  137.68    by (simp add: even_nat_def)
  137.69 @@ -37,13 +63,13 @@
  137.70    transfer_int_nat_relations
  137.71  ]
  137.72  
  137.73 -lemma even_zero_int[simp]: "even (0::int)" by presburger
  137.74 +lemma odd_one_int [simp]:
  137.75 +  "odd (1::int)"
  137.76 +  by presburger
  137.77  
  137.78 -lemma odd_one_int[simp]: "odd (1::int)" by presburger
  137.79 -
  137.80 -lemma even_zero_nat[simp]: "even (0::nat)" by presburger
  137.81 -
  137.82 -lemma odd_1_nat [simp]: "odd (1::nat)" by presburger
  137.83 +lemma odd_1_nat [simp]:
  137.84 +  "odd (1::nat)"
  137.85 +  by presburger
  137.86  
  137.87  lemma even_numeral_int [simp]: "even (numeral (Num.Bit0 k) :: int)"
  137.88    unfolding even_def by simp
  137.89 @@ -52,7 +78,7 @@
  137.90    unfolding even_def by simp
  137.91  
  137.92  (* TODO: proper simp rules for Num.Bit0, Num.Bit1 *)
  137.93 -declare even_def[of "neg_numeral v", simp] for v
  137.94 +declare even_def [of "neg_numeral v", simp] for v
  137.95  
  137.96  lemma even_numeral_nat [simp]: "even (numeral (Num.Bit0 k) :: nat)"
  137.97    unfolding even_nat_def by simp
  137.98 @@ -62,34 +88,8 @@
  137.99  
 137.100  subsection {* Even and odd are mutually exclusive *}
 137.101  
 137.102 -lemma int_pos_lt_two_imp_zero_or_one:
 137.103 -    "0 <= x ==> (x::int) < 2 ==> x = 0 | x = 1"
 137.104 -  by presburger
 137.105 -
 137.106 -lemma neq_one_mod_two [simp, presburger]: 
 137.107 -  "((x::int) mod 2 ~= 0) = (x mod 2 = 1)" by presburger
 137.108 -
 137.109  
 137.110  subsection {* Behavior under integer arithmetic operations *}
 137.111 -declare dvd_def[algebra]
 137.112 -lemma nat_even_iff_2_dvd[algebra]: "even (x::nat) \<longleftrightarrow> 2 dvd x"
 137.113 -  by presburger
 137.114 -lemma int_even_iff_2_dvd[algebra]: "even (x::int) \<longleftrightarrow> 2 dvd x"
 137.115 -  by presburger
 137.116 -
 137.117 -lemma even_times_anything: "even (x::int) ==> even (x * y)"
 137.118 -  by algebra
 137.119 -
 137.120 -lemma anything_times_even: "even (y::int) ==> even (x * y)" by algebra
 137.121 -
 137.122 -lemma odd_times_odd: "odd (x::int) ==> odd y ==> odd (x * y)" 
 137.123 -  by (simp add: even_def mod_mult_right_eq)
 137.124 -
 137.125 -lemma even_product[simp,presburger]: "even((x::int) * y) = (even x | even y)"
 137.126 -  apply (auto simp add: even_times_anything anything_times_even)
 137.127 -  apply (rule ccontr)
 137.128 -  apply (auto simp add: odd_times_odd)
 137.129 -  done
 137.130  
 137.131  lemma even_plus_even: "even (x::int) ==> even y ==> even (x + y)"
 137.132  by presburger
 137.133 @@ -158,10 +158,6 @@
 137.134  
 137.135  subsection {* Equivalent definitions *}
 137.136  
 137.137 -lemma nat_lt_two_imp_zero_or_one:
 137.138 -  "(x::nat) < Suc (Suc 0) ==> x = 0 | x = Suc 0"
 137.139 -by presburger
 137.140 -
 137.141  lemma even_nat_mod_two_eq_zero: "even (x::nat) ==> x mod (Suc (Suc 0)) = 0"
 137.142  by presburger
 137.143  
 137.144 @@ -189,45 +185,19 @@
 137.145  
 137.146  subsection {* Parity and powers *}
 137.147  
 137.148 -lemma  minus_one_even_odd_power:
 137.149 -     "(even x --> (- 1::'a::{comm_ring_1})^x = 1) &
 137.150 -      (odd x --> (- 1::'a)^x = - 1)"
 137.151 -  apply (induct x)
 137.152 -  apply (rule conjI)
 137.153 -  apply simp
 137.154 -  apply (insert even_zero_nat, blast)
 137.155 -  apply simp
 137.156 -  done
 137.157 +lemma (in comm_ring_1) neg_power_if:
 137.158 +  "(- a) ^ n = (if even n then (a ^ n) else - (a ^ n))"
 137.159 +  by (induct n) simp_all
 137.160  
 137.161 -lemma minus_one_even_power [simp]:
 137.162 -    "even x ==> (- 1::'a::{comm_ring_1})^x = 1"
 137.163 -  using minus_one_even_odd_power by blast
 137.164 +lemma (in comm_ring_1)
 137.165 +  shows minus_one_even_power [simp]: "even n \<Longrightarrow> (- 1) ^ n = 1"
 137.166 +  and minus_one_odd_power [simp]: "odd n \<Longrightarrow> (- 1) ^ n = - 1"
 137.167 +  by (simp_all add: neg_power_if del: minus_one)
 137.168  
 137.169 -lemma minus_one_odd_power [simp]:
 137.170 -    "odd x ==> (- 1::'a::{comm_ring_1})^x = - 1"
 137.171 -  using minus_one_even_odd_power by blast
 137.172 -
 137.173 -lemma neg_one_even_odd_power:
 137.174 -     "(even x --> (-1::'a::{comm_ring_1})^x = 1) &
 137.175 -      (odd x --> (-1::'a)^x = -1)"
 137.176 -  apply (induct x)
 137.177 -  apply (simp, simp)
 137.178 -  done
 137.179 -
 137.180 -lemma neg_one_even_power [simp]:
 137.181 -    "even x ==> (-1::'a::{comm_ring_1})^x = 1"
 137.182 -  using neg_one_even_odd_power by blast
 137.183 -
 137.184 -lemma neg_one_odd_power [simp]:
 137.185 -    "odd x ==> (-1::'a::{comm_ring_1})^x = -1"
 137.186 -  using neg_one_even_odd_power by blast
 137.187 -
 137.188 -lemma neg_power_if:
 137.189 -     "(-x::'a::{comm_ring_1}) ^ n =
 137.190 -      (if even n then (x ^ n) else -(x ^ n))"
 137.191 -  apply (induct n)
 137.192 -  apply simp_all
 137.193 -  done
 137.194 +lemma (in comm_ring_1)
 137.195 +  shows neg_one_even_power [simp]: "even n \<Longrightarrow> (-1) ^ n = 1"
 137.196 +  and neg_one_odd_power [simp]: "odd n \<Longrightarrow> (-1) ^ n = - 1"
 137.197 +  by (simp_all add: minus_one [symmetric] del: minus_one)
 137.198  
 137.199  lemma zero_le_even_power: "even n ==>
 137.200      0 <= (x::'a::{linordered_ring,monoid_mult}) ^ n"
 137.201 @@ -244,7 +214,7 @@
 137.202  apply (metis field_power_not_zero divisors_zero order_antisym_conv zero_le_square)
 137.203  done
 137.204  
 137.205 -lemma zero_le_power_eq[presburger]: "(0 <= (x::'a::{linordered_idom}) ^ n) =
 137.206 +lemma zero_le_power_eq [presburger]: "(0 <= (x::'a::{linordered_idom}) ^ n) =
 137.207      (even n | (odd n & 0 <= x))"
 137.208    apply auto
 137.209    apply (subst zero_le_odd_power [symmetric])
 137.210 @@ -277,9 +247,6 @@
 137.211    apply (simp add: zero_le_even_power)
 137.212    done
 137.213  
 137.214 -lemma zero_less_power_nat_eq[presburger]: "(0 < (x::nat) ^ n) = (n = 0 | 0 < x)"
 137.215 -  by (induct n) auto
 137.216 -
 137.217  lemma power_minus_even [simp]: "even n ==>
 137.218      (- x)^n = (x^n::'a::{comm_ring_1})"
 137.219    apply (subst power_minus)
 137.220 @@ -336,13 +303,6 @@
 137.221  
 137.222  lemma odd_add [simp]: "odd(m + n::nat) = (odd m \<noteq> odd n)" by presburger
 137.223  
 137.224 -lemma div_Suc: "Suc a div c = a div c + Suc 0 div c +
 137.225 -    (a mod c + Suc 0 mod c) div c" 
 137.226 -  apply (subgoal_tac "Suc a = a + Suc 0")
 137.227 -  apply (erule ssubst)
 137.228 -  apply (rule div_add1_eq, simp)
 137.229 -  done
 137.230 -
 137.231  lemma lemma_even_div2 [simp]: "even (n::nat) ==> (n + 1) div 2 = n div 2" by presburger
 137.232  
 137.233  lemma lemma_not_even_div2 [simp]: "~even n ==> (n + 1) div 2 = Suc (n div 2)"
 137.234 @@ -359,31 +319,29 @@
 137.235  text {* Simplify, when the exponent is a numeral *}
 137.236  
 137.237  lemmas zero_le_power_eq_numeral [simp] =
 137.238 -    zero_le_power_eq [of _ "numeral w"] for w
 137.239 +  zero_le_power_eq [of _ "numeral w"] for w
 137.240  
 137.241  lemmas zero_less_power_eq_numeral [simp] =
 137.242 -    zero_less_power_eq [of _ "numeral w"] for w
 137.243 +  zero_less_power_eq [of _ "numeral w"] for w
 137.244  
 137.245  lemmas power_le_zero_eq_numeral [simp] =
 137.246 -    power_le_zero_eq [of _ "numeral w"] for w
 137.247 +  power_le_zero_eq [of _ "numeral w"] for w
 137.248  
 137.249  lemmas power_less_zero_eq_numeral [simp] =
 137.250 -    power_less_zero_eq [of _ "numeral w"] for w
 137.251 +  power_less_zero_eq [of _ "numeral w"] for w
 137.252  
 137.253  lemmas zero_less_power_nat_eq_numeral [simp] =
 137.254 -    zero_less_power_nat_eq [of _ "numeral w"] for w
 137.255 +  nat_zero_less_power_iff [of _ "numeral w"] for w
 137.256  
 137.257 -lemmas power_eq_0_iff_numeral [simp] = power_eq_0_iff [of _ "numeral w"] for w
 137.258 +lemmas power_eq_0_iff_numeral [simp] =
 137.259 +  power_eq_0_iff [of _ "numeral w"] for w
 137.260  
 137.261 -lemmas power_even_abs_numeral [simp] = power_even_abs [of "numeral w" _] for w
 137.262 +lemmas power_even_abs_numeral [simp] =
 137.263 +  power_even_abs [of "numeral w" _] for w
 137.264  
 137.265  
 137.266  subsection {* An Equivalence for @{term [source] "0 \<le> a^n"} *}
 137.267  
 137.268 -lemma even_power_le_0_imp_0:
 137.269 -    "a ^ (2*k) \<le> (0::'a::{linordered_idom}) ==> a=0"
 137.270 -  by (induct k) (auto simp add: zero_le_mult_iff mult_le_0_iff)
 137.271 -
 137.272  lemma zero_le_power_iff[presburger]:
 137.273    "(0 \<le> a^n) = (0 \<le> (a::'a::{linordered_idom}) | even n)"
 137.274  proof cases
 137.275 @@ -395,9 +353,10 @@
 137.276    assume odd: "odd n"
 137.277    then obtain k where "n = Suc(2*k)"
 137.278      by (auto simp add: odd_nat_equiv_def2 numeral_2_eq_2)
 137.279 -  thus ?thesis
 137.280 -    by (auto simp add: zero_le_mult_iff zero_le_even_power
 137.281 -             dest!: even_power_le_0_imp_0)
 137.282 +  moreover have "a ^ (2 * k) \<le> 0 \<Longrightarrow> a = 0"
 137.283 +    by (induct k) (auto simp add: zero_le_mult_iff mult_le_0_iff)
 137.284 +  ultimately show ?thesis
 137.285 +    by (auto simp add: zero_le_mult_iff zero_le_even_power)
 137.286  qed
 137.287  
 137.288  
 137.289 @@ -409,7 +368,6 @@
 137.290  lemma odd_plus_one_div_two: "odd (x::int) ==> (x + 1) div 2 = x div 2 + 1" by presburger
 137.291  
 137.292  lemma [presburger]: "(Suc x) div Suc (Suc 0) = x div Suc (Suc 0) \<longleftrightarrow> even x" by presburger
 137.293 -lemma [presburger]: "(Suc x) div Suc (Suc 0) = x div Suc (Suc 0) \<longleftrightarrow> even x" by presburger
 137.294  lemma even_nat_plus_one_div_two: "even (x::nat) ==>
 137.295      (Suc x) div Suc (Suc 0) = x div Suc (Suc 0)" by presburger
 137.296  
 137.297 @@ -417,3 +375,4 @@
 137.298      (Suc x) div Suc (Suc 0) = Suc (x div Suc (Suc 0))" by presburger
 137.299  
 137.300  end
 137.301 +
   138.1 --- a/src/HOL/Power.thy	Mon Nov 11 17:34:44 2013 +0100
   138.2 +++ b/src/HOL/Power.thy	Mon Nov 11 17:44:21 2013 +0100
   138.3 @@ -530,7 +530,7 @@
   138.4    "abs ((-a) ^ n) = abs (a ^ n)"
   138.5    by (simp add: power_abs)
   138.6  
   138.7 -lemma zero_less_power_abs_iff [simp, no_atp]:
   138.8 +lemma zero_less_power_abs_iff [simp]:
   138.9    "0 < abs a ^ n \<longleftrightarrow> a \<noteq> 0 \<or> n = 0"
  138.10  proof (induct n)
  138.11    case 0 show ?case by simp
  138.12 @@ -730,8 +730,18 @@
  138.13    fixes m n :: nat
  138.14    assumes "m\<^sup>2 \<le> n"
  138.15    shows "m \<le> n"
  138.16 -  using assms by (cases m) (simp_all add: power2_eq_square)
  138.17 -
  138.18 +proof (cases m)
  138.19 +  case 0 then show ?thesis by simp
  138.20 +next
  138.21 +  case (Suc k)
  138.22 +  show ?thesis
  138.23 +  proof (rule ccontr)
  138.24 +    assume "\<not> m \<le> n"
  138.25 +    then have "n < m" by simp
  138.26 +    with assms Suc show False
  138.27 +      by (auto simp add: algebra_simps) (simp add: power2_eq_square)
  138.28 +  qed
  138.29 +qed
  138.30  
  138.31  
  138.32  subsection {* Code generator tweak *}
   139.1 --- a/src/HOL/Presburger.thy	Mon Nov 11 17:34:44 2013 +0100
   139.2 +++ b/src/HOL/Presburger.thy	Mon Nov 11 17:44:21 2013 +0100
   139.3 @@ -360,11 +360,15 @@
   139.4    apply simp
   139.5    done
   139.6  
   139.7 -lemma zdvd_mono: assumes not0: "(k::int) \<noteq> 0"
   139.8 -shows "((m::int) dvd t) \<equiv> (k*m dvd k*t)" 
   139.9 -  using not0 by (simp add: dvd_def)
  139.10 +lemma zdvd_mono:
  139.11 +  fixes k m t :: int
  139.12 +  assumes "k \<noteq> 0"
  139.13 +  shows "m dvd t \<equiv> k * m dvd k * t" 
  139.14 +  using assms by simp
  139.15  
  139.16 -lemma uminus_dvd_conv: "(d dvd (t::int)) \<equiv> (-d dvd t)" "(d dvd (t::int)) \<equiv> (d dvd -t)"
  139.17 +lemma uminus_dvd_conv:
  139.18 +  fixes d t :: int
  139.19 +  shows "d dvd t \<equiv> - d dvd t" and "d dvd t \<equiv> d dvd - t"
  139.20    by simp_all
  139.21  
  139.22  text {* \bigskip Theorems for transforming predicates on nat to predicates on @{text int}*}
  139.23 @@ -406,24 +410,23 @@
  139.24    end
  139.25  *} "Cooper's algorithm for Presburger arithmetic"
  139.26  
  139.27 -declare dvd_eq_mod_eq_0[symmetric, presburger]
  139.28 -declare mod_1[presburger] 
  139.29 -declare mod_0[presburger]
  139.30 -declare mod_by_1[presburger]
  139.31 -declare mod_self[presburger]
  139.32 -declare mod_by_0[presburger]
  139.33 -declare mod_div_trivial[presburger]
  139.34 -declare div_mod_equality2[presburger]
  139.35 -declare div_mod_equality[presburger]
  139.36 -declare mod_div_equality2[presburger]
  139.37 -declare mod_div_equality[presburger]
  139.38 -declare mod_mult_self1[presburger]
  139.39 -declare mod_mult_self2[presburger]
  139.40 -declare div_mod_equality[presburger]
  139.41 -declare div_mod_equality2[presburger]
  139.42 +declare dvd_eq_mod_eq_0 [symmetric, presburger]
  139.43 +declare mod_1 [presburger] 
  139.44 +declare mod_0 [presburger]
  139.45 +declare mod_by_1 [presburger]
  139.46 +declare mod_self [presburger]
  139.47 +declare div_by_0 [presburger]
  139.48 +declare mod_by_0 [presburger]
  139.49 +declare mod_div_trivial [presburger]
  139.50 +declare div_mod_equality2 [presburger]
  139.51 +declare div_mod_equality [presburger]
  139.52 +declare mod_div_equality2 [presburger]
  139.53 +declare mod_div_equality [presburger]
  139.54 +declare mod_mult_self1 [presburger]
  139.55 +declare mod_mult_self2 [presburger]
  139.56  declare mod2_Suc_Suc[presburger]
  139.57 -lemma [presburger]: "(a::int) div 0 = 0" and [presburger]: "a mod 0 = a"
  139.58 -by simp_all
  139.59 +declare not_mod_2_eq_0_eq_1 [presburger] 
  139.60 +declare nat_zero_less_power_iff [presburger]
  139.61  
  139.62  lemma [presburger, algebra]: "m mod 2 = (1::nat) \<longleftrightarrow> \<not> 2 dvd m " by presburger
  139.63  lemma [presburger, algebra]: "m mod 2 = Suc 0 \<longleftrightarrow> \<not> 2 dvd m " by presburger
  139.64 @@ -432,3 +435,4 @@
  139.65  lemma [presburger, algebra]: "m mod 2 = (1::int) \<longleftrightarrow> \<not> 2 dvd m " by presburger
  139.66  
  139.67  end
  139.68 +
   140.1 --- a/src/HOL/Probability/Borel_Space.thy	Mon Nov 11 17:34:44 2013 +0100
   140.2 +++ b/src/HOL/Probability/Borel_Space.thy	Mon Nov 11 17:44:21 2013 +0100
   140.3 @@ -677,7 +677,7 @@
   140.4    assumes f: "f \<in> borel_measurable M"
   140.5    assumes g: "g \<in> borel_measurable M"
   140.6    shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
   140.7 -  unfolding diff_minus using assms by simp
   140.8 +  using borel_measurable_add [of f M "- g"] assms by (simp add: fun_Compl_def)
   140.9  
  140.10  lemma borel_measurable_times[measurable (raw)]:
  140.11    fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_algebra}"
  140.12 @@ -719,7 +719,8 @@
  140.13    proof cases
  140.14      assume "b \<noteq> 0"
  140.15      with `open S` have "open ((\<lambda>x. (- a + x) /\<^sub>R b) ` S)" (is "open ?S")
  140.16 -      by (auto intro!: open_affinity simp: scaleR_add_right)
  140.17 +      using open_affinity [of S "inverse b" "- a /\<^sub>R b"]
  140.18 +      by (auto simp: algebra_simps)
  140.19      hence "?S \<in> sets borel" by auto
  140.20      moreover
  140.21      from `b \<noteq> 0` have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S"
   141.1 --- a/src/HOL/Probability/Lebesgue_Integration.thy	Mon Nov 11 17:34:44 2013 +0100
   141.2 +++ b/src/HOL/Probability/Lebesgue_Integration.thy	Mon Nov 11 17:44:21 2013 +0100
   141.3 @@ -1528,7 +1528,7 @@
   141.4      using mono by auto
   141.5    ultimately show ?thesis using fg
   141.6      by (auto intro!: add_mono positive_integral_mono_AE real_of_ereal_positive_mono
   141.7 -             simp: positive_integral_positive lebesgue_integral_def diff_minus)
   141.8 +             simp: positive_integral_positive lebesgue_integral_def algebra_simps)
   141.9  qed
  141.10  
  141.11  lemma integral_mono:
  141.12 @@ -1732,7 +1732,7 @@
  141.13    shows "integrable M (\<lambda>t. f t - g t)"
  141.14    and "(\<integral> t. f t - g t \<partial>M) = integral\<^sup>L M f - integral\<^sup>L M g"
  141.15    using integral_add[OF f integral_minus(1)[OF g]]
  141.16 -  unfolding diff_minus integral_minus(2)[OF g]
  141.17 +  unfolding integral_minus(2)[OF g]
  141.18    by auto
  141.19  
  141.20  lemma integral_indicator[simp, intro]:
   142.1 --- a/src/HOL/Probability/Lebesgue_Measure.thy	Mon Nov 11 17:34:44 2013 +0100
   142.2 +++ b/src/HOL/Probability/Lebesgue_Measure.thy	Mon Nov 11 17:44:21 2013 +0100
   142.3 @@ -1138,7 +1138,7 @@
   142.4      show "\<And>i x. 0 \<le> ?f i x"
   142.5        using nonneg by (auto split: split_indicator)
   142.6    qed
   142.7 -  also have "\<dots> = (SUP i::nat. F (a + real i) - F a)"
   142.8 +  also have "\<dots> = (SUP i::nat. ereal (F (a + real i) - F a))"
   142.9      by (subst positive_integral_FTC_atLeastAtMost[OF f_borel f nonneg]) auto
  142.10    also have "\<dots> = T - F a"
  142.11    proof (rule SUP_Lim_ereal)
   143.1 --- a/src/HOL/Product_Type.thy	Mon Nov 11 17:34:44 2013 +0100
   143.2 +++ b/src/HOL/Product_Type.thy	Mon Nov 11 17:44:21 2013 +0100
   143.3 @@ -75,10 +75,10 @@
   143.4    f} rather than by @{term [source] "%u. f ()"}.
   143.5  *}
   143.6  
   143.7 -lemma unit_abs_eta_conv [simp, no_atp]: "(%u::unit. f ()) = f"
   143.8 +lemma unit_abs_eta_conv [simp]: "(%u::unit. f ()) = f"
   143.9    by (rule ext) simp
  143.10  
  143.11 -lemma UNIV_unit [no_atp]:
  143.12 +lemma UNIV_unit:
  143.13    "UNIV = {()}" by auto
  143.14  
  143.15  instantiation unit :: default
  143.16 @@ -586,10 +586,10 @@
  143.17     to quite time-consuming computations (in particular for nested tuples) *)
  143.18  setup {* map_theory_claset (fn ctxt => ctxt addSbefore ("split_conv_tac", split_conv_tac)) *}
  143.19  
  143.20 -lemma split_eta_SetCompr [simp,no_atp]: "(%u. EX x y. u = (x, y) & P (x, y)) = P"
  143.21 +lemma split_eta_SetCompr [simp, no_atp]: "(%u. EX x y. u = (x, y) & P (x, y)) = P"
  143.22    by (rule ext) fast
  143.23  
  143.24 -lemma split_eta_SetCompr2 [simp,no_atp]: "(%u. EX x y. u = (x, y) & P x y) = split P"
  143.25 +lemma split_eta_SetCompr2 [simp, no_atp]: "(%u. EX x y. u = (x, y) & P x y) = split P"
  143.26    by (rule ext) fast
  143.27  
  143.28  lemma split_part [simp]: "(%(a,b). P & Q a b) = (%ab. P & split Q ab)"
   144.1 --- a/src/HOL/Proofs/Lambda/ListOrder.thy	Mon Nov 11 17:34:44 2013 +0100
   144.2 +++ b/src/HOL/Proofs/Lambda/ListOrder.thy	Mon Nov 11 17:44:21 2013 +0100
   144.3 @@ -89,15 +89,15 @@
   144.4    done
   144.5  
   144.6  lemma Cons_acc_step1I [intro!]:
   144.7 -    "accp r x ==> accp (step1 r) xs \<Longrightarrow> accp (step1 r) (x # xs)"
   144.8 -  apply (induct arbitrary: xs set: accp)
   144.9 +    "Wellfounded.accp r x ==> Wellfounded.accp (step1 r) xs \<Longrightarrow> Wellfounded.accp (step1 r) (x # xs)"
  144.10 +  apply (induct arbitrary: xs set: Wellfounded.accp)
  144.11    apply (erule thin_rl)
  144.12    apply (erule accp_induct)
  144.13    apply (rule accp.accI)
  144.14    apply blast
  144.15    done
  144.16  
  144.17 -lemma lists_accD: "listsp (accp r) xs ==> accp (step1 r) xs"
  144.18 +lemma lists_accD: "listsp (Wellfounded.accp r) xs ==> Wellfounded.accp (step1 r) xs"
  144.19    apply (induct set: listsp)
  144.20     apply (rule accp.accI)
  144.21     apply simp
  144.22 @@ -113,8 +113,8 @@
  144.23    apply force
  144.24    done
  144.25  
  144.26 -lemma lists_accI: "accp (step1 r) xs ==> listsp (accp r) xs"
  144.27 -  apply (induct set: accp)
  144.28 +lemma lists_accI: "Wellfounded.accp (step1 r) xs ==> listsp (Wellfounded.accp r) xs"
  144.29 +  apply (induct set: Wellfounded.accp)
  144.30    apply clarify
  144.31    apply (rule accp.accI)
  144.32    apply (drule_tac r=r in ex_step1I, assumption)
   145.1 --- a/src/HOL/ROOT	Mon Nov 11 17:34:44 2013 +0100
   145.2 +++ b/src/HOL/ROOT	Mon Nov 11 17:44:21 2013 +0100
   145.3 @@ -733,6 +733,7 @@
   145.4    theories [condition = ISABELLE_FULL_TEST]
   145.5      Misc_Codatatype
   145.6      Misc_Datatype
   145.7 +    Misc_Primcorec
   145.8      Misc_Primrec
   145.9  
  145.10  session "HOL-Word" (main) in Word = HOL +
   146.1 --- a/src/HOL/Rat.thy	Mon Nov 11 17:34:44 2013 +0100
   146.2 +++ b/src/HOL/Rat.thy	Mon Nov 11 17:44:21 2013 +0100
   146.3 @@ -468,7 +468,7 @@
   146.4      unfolding less_eq_rat_def less_rat_def
   146.5      by (auto, drule (1) positive_add, simp add: positive_zero)
   146.6    show "a \<le> b \<Longrightarrow> c + a \<le> c + b"
   146.7 -    unfolding less_eq_rat_def less_rat_def by (auto simp: diff_minus)
   146.8 +    unfolding less_eq_rat_def less_rat_def by auto
   146.9    show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
  146.10      by (rule sgn_rat_def)
  146.11    show "a \<le> b \<or> b \<le> a"
  146.12 @@ -665,7 +665,7 @@
  146.13    by transfer simp
  146.14  
  146.15  lemma of_rat_diff: "of_rat (a - b) = of_rat a - of_rat b"
  146.16 -by (simp only: diff_minus of_rat_add of_rat_minus)
  146.17 +  using of_rat_add [of a "- b"] by (simp add: of_rat_minus)
  146.18  
  146.19  lemma of_rat_mult: "of_rat (a * b) = of_rat a * of_rat b"
  146.20  apply transfer
   147.1 --- a/src/HOL/Real.thy	Mon Nov 11 17:34:44 2013 +0100
   147.2 +++ b/src/HOL/Real.thy	Mon Nov 11 17:44:21 2013 +0100
   147.3 @@ -98,7 +98,7 @@
   147.4  lemma vanishes_diff:
   147.5    assumes X: "vanishes X" and Y: "vanishes Y"
   147.6    shows "vanishes (\<lambda>n. X n - Y n)"
   147.7 -unfolding diff_minus by (intro vanishes_add vanishes_minus X Y)
   147.8 +  unfolding diff_conv_add_uminus by (intro vanishes_add vanishes_minus X Y)
   147.9  
  147.10  lemma vanishes_mult_bounded:
  147.11    assumes X: "\<exists>a>0. \<forall>n. \<bar>X n\<bar> < a"
  147.12 @@ -170,7 +170,7 @@
  147.13  lemma cauchy_diff [simp]:
  147.14    assumes X: "cauchy X" and Y: "cauchy Y"
  147.15    shows "cauchy (\<lambda>n. X n - Y n)"
  147.16 -using assms unfolding diff_minus by simp
  147.17 +  using assms unfolding diff_conv_add_uminus by (simp del: add_uminus_conv_diff)
  147.18  
  147.19  lemma cauchy_imp_bounded:
  147.20    assumes "cauchy X" shows "\<exists>b>0. \<forall>n. \<bar>X n\<bar> < b"
  147.21 @@ -456,7 +456,7 @@
  147.22  lemma diff_Real:
  147.23    assumes X: "cauchy X" and Y: "cauchy Y"
  147.24    shows "Real X - Real Y = Real (\<lambda>n. X n - Y n)"
  147.25 -  unfolding minus_real_def diff_minus
  147.26 +  unfolding minus_real_def
  147.27    by (simp add: minus_Real add_Real X Y)
  147.28  
  147.29  lemma mult_Real:
  147.30 @@ -607,7 +607,7 @@
  147.31      unfolding less_eq_real_def less_real_def
  147.32      by (auto, drule (1) positive_add, simp add: positive_zero)
  147.33    show "a \<le> b \<Longrightarrow> c + a \<le> c + b"
  147.34 -    unfolding less_eq_real_def less_real_def by (auto simp: diff_minus) (* by auto *)
  147.35 +    unfolding less_eq_real_def less_real_def by auto
  147.36      (* FIXME: Procedure int_combine_numerals: c + b - (c + a) \<equiv> b + - a *)
  147.37      (* Should produce c + b - (c + a) \<equiv> b - a *)
  147.38    show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
  147.39 @@ -919,24 +919,20 @@
  147.40      using 1 2 3 by (rule_tac x="Real B" in exI, simp)
  147.41  qed
  147.42  
  147.43 -
  147.44  instantiation real :: linear_continuum
  147.45  begin
  147.46  
  147.47  subsection{*Supremum of a set of reals*}
  147.48  
  147.49 -definition
  147.50 -  Sup_real_def: "Sup X \<equiv> LEAST z::real. \<forall>x\<in>X. x\<le>z"
  147.51 -
  147.52 -definition
  147.53 -  Inf_real_def: "Inf (X::real set) \<equiv> - Sup (uminus ` X)"
  147.54 +definition "Sup X = (LEAST z::real. \<forall>x\<in>X. x \<le> z)"
  147.55 +definition "Inf (X::real set) = - Sup (uminus ` X)"
  147.56  
  147.57  instance
  147.58  proof
  147.59 -  { fix z x :: real and X :: "real set"
  147.60 -    assume x: "x \<in> X" and z: "\<And>x. x \<in> X \<Longrightarrow> x \<le> z"
  147.61 +  { fix x :: real and X :: "real set"
  147.62 +    assume x: "x \<in> X" "bdd_above X"
  147.63      then obtain s where s: "\<forall>y\<in>X. y \<le> s" "\<And>z. \<forall>y\<in>X. y \<le> z \<Longrightarrow> s \<le> z"
  147.64 -      using complete_real[of X] by blast
  147.65 +      using complete_real[of X] unfolding bdd_above_def by blast
  147.66      then show "x \<le> Sup X"
  147.67        unfolding Sup_real_def by (rule LeastI2_order) (auto simp: x) }
  147.68    note Sup_upper = this
  147.69 @@ -952,32 +948,15 @@
  147.70      finally show "Sup X \<le> z" . }
  147.71    note Sup_least = this
  147.72  
  147.73 -  { fix x z :: real and X :: "real set"
  147.74 -    assume x: "x \<in> X" and z: "\<And>x. x \<in> X \<Longrightarrow> z \<le> x"
  147.75 -    have "-x \<le> Sup (uminus ` X)"
  147.76 -      by (rule Sup_upper[of _ _ "- z"]) (auto simp add: image_iff x z)
  147.77 -    then show "Inf X \<le> x" 
  147.78 -      by (auto simp add: Inf_real_def) }
  147.79 -
  147.80 -  { fix z :: real and X :: "real set"
  147.81 -    assume x: "X \<noteq> {}" and z: "\<And>x. x \<in> X \<Longrightarrow> z \<le> x"
  147.82 -    have "Sup (uminus ` X) \<le> -z"
  147.83 -      using x z by (force intro: Sup_least)
  147.84 -    then show "z \<le> Inf X" 
  147.85 -        by (auto simp add: Inf_real_def) }
  147.86 -
  147.87 +  { fix x :: real and X :: "real set" assume x: "x \<in> X" "bdd_below X" then show "Inf X \<le> x"
  147.88 +      using Sup_upper[of "-x" "uminus ` X"] by (auto simp: Inf_real_def) }
  147.89 +  { fix z :: real and X :: "real set" assume "X \<noteq> {}" "\<And>x. x \<in> X \<Longrightarrow> z \<le> x" then show "z \<le> Inf X"
  147.90 +      using Sup_least[of "uminus ` X" "- z"] by (force simp: Inf_real_def) }
  147.91    show "\<exists>a b::real. a \<noteq> b"
  147.92      using zero_neq_one by blast
  147.93  qed
  147.94  end
  147.95  
  147.96 -text {*
  147.97 -  \medskip Completeness properties using @{text "isUb"}, @{text "isLub"}:
  147.98 -*}
  147.99 -
 147.100 -lemma reals_complete: "\<exists>X. X \<in> S \<Longrightarrow> \<exists>Y. isUb (UNIV::real set) S Y \<Longrightarrow> \<exists>t. isLub (UNIV :: real set) S t"
 147.101 -  by (intro exI[of _ "Sup S"] isLub_cSup) (auto simp: setle_def isUb_def intro: cSup_upper)
 147.102 -
 147.103  
 147.104  subsection {* Hiding implementation details *}
 147.105  
   148.1 --- a/src/HOL/Real_Vector_Spaces.thy	Mon Nov 11 17:34:44 2013 +0100
   148.2 +++ b/src/HOL/Real_Vector_Spaces.thy	Mon Nov 11 17:44:21 2013 +0100
   148.3 @@ -31,7 +31,7 @@
   148.4  qed
   148.5  
   148.6  lemma diff: "f (x - y) = f x - f y"
   148.7 -by (simp add: add minus diff_minus)
   148.8 +  using add [of x "- y"] by (simp add: minus)
   148.9  
  148.10  lemma setsum: "f (setsum g A) = (\<Sum>x\<in>A. f (g x))"
  148.11  apply (cases "finite A")
  148.12 @@ -553,8 +553,7 @@
  148.13  proof -
  148.14    have "norm (a + - b) \<le> norm a + norm (- b)"
  148.15      by (rule norm_triangle_ineq)
  148.16 -  thus ?thesis
  148.17 -    by (simp only: diff_minus norm_minus_cancel)
  148.18 +  then show ?thesis by simp
  148.19  qed
  148.20  
  148.21  lemma norm_diff_ineq:
  148.22 @@ -571,7 +570,7 @@
  148.23    shows "norm ((a + b) - (c + d)) \<le> norm (a - c) + norm (b - d)"
  148.24  proof -
  148.25    have "norm ((a + b) - (c + d)) = norm ((a - c) + (b - d))"
  148.26 -    by (simp add: diff_minus add_ac)
  148.27 +    by (simp add: algebra_simps)
  148.28    also have "\<dots> \<le> norm (a - c) + norm (b - d)"
  148.29      by (rule norm_triangle_ineq)
  148.30    finally show ?thesis .
  148.31 @@ -1426,9 +1425,6 @@
  148.32    @{term "{r::real. \<exists>N. \<forall>n\<ge>N. r < X n}"}
  148.33  *}
  148.34  
  148.35 -lemma isUb_UNIV_I: "(\<And>y. y \<in> S \<Longrightarrow> y \<le> u) \<Longrightarrow> isUb UNIV S u"
  148.36 -by (simp add: isUbI setleI)
  148.37 -
  148.38  lemma increasing_LIMSEQ:
  148.39    fixes f :: "nat \<Rightarrow> real"
  148.40    assumes inc: "\<And>n. f n \<le> f (Suc n)"
  148.41 @@ -1455,40 +1451,33 @@
  148.42    then have mem_S: "\<And>N x. \<forall>n\<ge>N. x < X n \<Longrightarrow> x \<in> S" by auto
  148.43  
  148.44    { fix N x assume N: "\<forall>n\<ge>N. X n < x"
  148.45 -  have "isUb UNIV S x"
  148.46 -  proof (rule isUb_UNIV_I)
  148.47    fix y::real assume "y \<in> S"
  148.48    hence "\<exists>M. \<forall>n\<ge>M. y < X n"
  148.49      by (simp add: S_def)
  148.50    then obtain M where "\<forall>n\<ge>M. y < X n" ..
  148.51    hence "y < X (max M N)" by simp
  148.52    also have "\<dots> < x" using N by simp
  148.53 -  finally show "y \<le> x"
  148.54 -    by (rule order_less_imp_le)
  148.55 -  qed }
  148.56 +  finally have "y \<le> x"
  148.57 +    by (rule order_less_imp_le) }
  148.58    note bound_isUb = this 
  148.59  
  148.60 -  have "\<exists>u. isLub UNIV S u"
  148.61 -  proof (rule reals_complete)
  148.62    obtain N where "\<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m) (X n) < 1"
  148.63      using X[THEN metric_CauchyD, OF zero_less_one] by auto
  148.64    hence N: "\<forall>n\<ge>N. dist (X n) (X N) < 1" by simp
  148.65 -  show "\<exists>x. x \<in> S"
  148.66 -  proof
  148.67 +  have [simp]: "S \<noteq> {}"
  148.68 +  proof (intro exI ex_in_conv[THEN iffD1])
  148.69      from N have "\<forall>n\<ge>N. X N - 1 < X n"
  148.70        by (simp add: abs_diff_less_iff dist_real_def)
  148.71      thus "X N - 1 \<in> S" by (rule mem_S)
  148.72    qed
  148.73 -  show "\<exists>u. isUb UNIV S u"
  148.74 +  have [simp]: "bdd_above S"
  148.75    proof
  148.76      from N have "\<forall>n\<ge>N. X n < X N + 1"
  148.77        by (simp add: abs_diff_less_iff dist_real_def)
  148.78 -    thus "isUb UNIV S (X N + 1)"
  148.79 +    thus "\<And>s. s \<in> S \<Longrightarrow>  s \<le> X N + 1"
  148.80        by (rule bound_isUb)
  148.81    qed
  148.82 -  qed
  148.83 -  then obtain x where x: "isLub UNIV S x" ..
  148.84 -  have "X ----> x"
  148.85 +  have "X ----> Sup S"
  148.86    proof (rule metric_LIMSEQ_I)
  148.87    fix r::real assume "0 < r"
  148.88    hence r: "0 < r/2" by simp
  148.89 @@ -1500,17 +1489,18 @@
  148.90  
  148.91    from N have "\<forall>n\<ge>N. X N - r/2 < X n" by fast
  148.92    hence "X N - r/2 \<in> S" by (rule mem_S)
  148.93 -  hence 1: "X N - r/2 \<le> x" using x isLub_isUb isUbD by fast
  148.94 +  hence 1: "X N - r/2 \<le> Sup S" by (simp add: cSup_upper)
  148.95  
  148.96    from N have "\<forall>n\<ge>N. X n < X N + r/2" by fast
  148.97 -  hence "isUb UNIV S (X N + r/2)" by (rule bound_isUb)
  148.98 -  hence 2: "x \<le> X N + r/2" using x isLub_le_isUb by fast
  148.99 +  from bound_isUb[OF this]
 148.100 +  have 2: "Sup S \<le> X N + r/2"
 148.101 +    by (intro cSup_least) simp_all
 148.102  
 148.103 -  show "\<exists>N. \<forall>n\<ge>N. dist (X n) x < r"
 148.104 +  show "\<exists>N. \<forall>n\<ge>N. dist (X n) (Sup S) < r"
 148.105    proof (intro exI allI impI)
 148.106      fix n assume n: "N \<le> n"
 148.107      from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp+
 148.108 -    thus "dist (X n) x < r" using 1 2
 148.109 +    thus "dist (X n) (Sup S) < r" using 1 2
 148.110        by (simp add: abs_diff_less_iff dist_real_def)
 148.111    qed
 148.112    qed
   149.1 --- a/src/HOL/Record.thy	Mon Nov 11 17:34:44 2013 +0100
   149.2 +++ b/src/HOL/Record.thy	Mon Nov 11 17:44:21 2013 +0100
   149.3 @@ -399,7 +399,7 @@
   149.4  lemma refl_conj_eq: "Q = R \<Longrightarrow> P \<and> Q \<longleftrightarrow> P \<and> R"
   149.5    by simp
   149.6  
   149.7 -lemma iso_tuple_UNIV_I [no_atp]: "x \<in> UNIV \<equiv> True"
   149.8 +lemma iso_tuple_UNIV_I: "x \<in> UNIV \<equiv> True"
   149.9    by simp
  149.10  
  149.11  lemma iso_tuple_True_simp: "(True \<Longrightarrow> PROP P) \<equiv> PROP P"
   150.1 --- a/src/HOL/Relation.thy	Mon Nov 11 17:34:44 2013 +0100
   150.2 +++ b/src/HOL/Relation.thy	Mon Nov 11 17:44:21 2013 +0100
   150.3 @@ -478,7 +478,7 @@
   150.4  lemma Id_on_eqI: "a = b ==> a : A ==> (a, b) : Id_on A"
   150.5    by (simp add: Id_on_def)
   150.6  
   150.7 -lemma Id_onI [intro!,no_atp]: "a : A ==> (a, a) : Id_on A"
   150.8 +lemma Id_onI [intro!]: "a : A ==> (a, a) : Id_on A"
   150.9    by (rule Id_on_eqI) (rule refl)
  150.10  
  150.11  lemma Id_onE [elim!]:
  150.12 @@ -939,8 +939,6 @@
  150.13  where
  150.14    "r `` s = {y. \<exists>x\<in>s. (x, y) \<in> r}"
  150.15  
  150.16 -declare Image_def [no_atp]
  150.17 -
  150.18  lemma Image_iff: "(b : r``A) = (EX x:A. (x, b) : r)"
  150.19    by (simp add: Image_def)
  150.20  
  150.21 @@ -950,7 +948,7 @@
  150.22  lemma Image_singleton_iff [iff]: "(b : r``{a}) = ((a, b) : r)"
  150.23    by (rule Image_iff [THEN trans]) simp
  150.24  
  150.25 -lemma ImageI [intro,no_atp]: "(a, b) : r ==> a : A ==> b : r``A"
  150.26 +lemma ImageI [intro]: "(a, b) : r ==> a : A ==> b : r``A"
  150.27    by (unfold Image_def) blast
  150.28  
  150.29  lemma ImageE [elim!]:
   151.1 --- a/src/HOL/Rings.thy	Mon Nov 11 17:34:44 2013 +0100
   151.2 +++ b/src/HOL/Rings.thy	Mon Nov 11 17:44:21 2013 +0100
   151.3 @@ -86,7 +86,20 @@
   151.4  lemma one_neq_zero [simp]: "1 \<noteq> 0"
   151.5  by (rule not_sym) (rule zero_neq_one)
   151.6  
   151.7 -end
   151.8 +definition of_bool :: "bool \<Rightarrow> 'a"
   151.9 +where
  151.10 +  "of_bool p = (if p then 1 else 0)" 
  151.11 +
  151.12 +lemma of_bool_eq [simp, code]:
  151.13 +  "of_bool False = 0"
  151.14 +  "of_bool True = 1"
  151.15 +  by (simp_all add: of_bool_def)
  151.16 +
  151.17 +lemma of_bool_eq_iff:
  151.18 +  "of_bool p = of_bool q \<longleftrightarrow> p = q"
  151.19 +  by (simp add: of_bool_def)
  151.20 +
  151.21 +end  
  151.22  
  151.23  class semiring_1 = zero_neq_one + semiring_0 + monoid_mult
  151.24  
  151.25 @@ -127,7 +140,7 @@
  151.26    then show ?thesis ..
  151.27  qed
  151.28  
  151.29 -lemma dvd_0_left_iff [no_atp, simp]: "0 dvd a \<longleftrightarrow> a = 0"
  151.30 +lemma dvd_0_left_iff [simp]: "0 dvd a \<longleftrightarrow> a = 0"
  151.31  by (auto intro: dvd_refl elim!: dvdE)
  151.32  
  151.33  lemma dvd_0_right [iff]: "a dvd 0"
  151.34 @@ -233,8 +246,8 @@
  151.35  by (rule minus_unique) (simp add: distrib_left [symmetric]) 
  151.36  
  151.37  text{*Extract signs from products*}
  151.38 -lemmas mult_minus_left [simp, no_atp] = minus_mult_left [symmetric]
  151.39 -lemmas mult_minus_right [simp,no_atp] = minus_mult_right [symmetric]
  151.40 +lemmas mult_minus_left [simp] = minus_mult_left [symmetric]
  151.41 +lemmas mult_minus_right [simp] = minus_mult_right [symmetric]
  151.42  
  151.43  lemma minus_mult_minus [simp]: "- a * - b = a * b"
  151.44  by simp
  151.45 @@ -242,13 +255,15 @@
  151.46  lemma minus_mult_commute: "- a * b = a * - b"
  151.47  by simp
  151.48  
  151.49 -lemma right_diff_distrib[algebra_simps, field_simps]: "a * (b - c) = a * b - a * c"
  151.50 -by (simp add: distrib_left diff_minus)
  151.51 +lemma right_diff_distrib [algebra_simps, field_simps]:
  151.52 +  "a * (b - c) = a * b - a * c"
  151.53 +  using distrib_left [of a b "-c "] by simp
  151.54  
  151.55 -lemma left_diff_distrib[algebra_simps, field_simps]: "(a - b) * c = a * c - b * c"
  151.56 -by (simp add: distrib_right diff_minus)
  151.57 +lemma left_diff_distrib [algebra_simps, field_simps]:
  151.58 +  "(a - b) * c = a * c - b * c"
  151.59 +  using distrib_right [of a "- b" c] by simp
  151.60  
  151.61 -lemmas ring_distribs[no_atp] =
  151.62 +lemmas ring_distribs =
  151.63    distrib_left distrib_right left_diff_distrib right_diff_distrib
  151.64  
  151.65  lemma eq_add_iff1:
  151.66 @@ -261,7 +276,7 @@
  151.67  
  151.68  end
  151.69  
  151.70 -lemmas ring_distribs[no_atp] =
  151.71 +lemmas ring_distribs =
  151.72    distrib_left distrib_right left_diff_distrib right_diff_distrib
  151.73  
  151.74  class comm_ring = comm_semiring + ab_group_add
  151.75 @@ -318,8 +333,9 @@
  151.76    then show "- x dvd y" ..
  151.77  qed
  151.78  
  151.79 -lemma dvd_diff[simp]: "x dvd y \<Longrightarrow> x dvd z \<Longrightarrow> x dvd (y - z)"
  151.80 -by (simp only: diff_minus dvd_add dvd_minus_iff)
  151.81 +lemma dvd_diff [simp]:
  151.82 +  "x dvd y \<Longrightarrow> x dvd z \<Longrightarrow> x dvd (y - z)"
  151.83 +  using dvd_add [of x y "- z"] by simp
  151.84  
  151.85  end
  151.86  
  151.87 @@ -336,7 +352,7 @@
  151.88  qed
  151.89  
  151.90  text{*Cancellation of equalities with a common factor*}
  151.91 -lemma mult_cancel_right [simp, no_atp]:
  151.92 +lemma mult_cancel_right [simp]:
  151.93    "a * c = b * c \<longleftrightarrow> c = 0 \<or> a = b"
  151.94  proof -
  151.95    have "(a * c = b * c) = ((a - b) * c = 0)"
  151.96 @@ -344,7 +360,7 @@
  151.97    thus ?thesis by (simp add: disj_commute)
  151.98  qed
  151.99  
 151.100 -lemma mult_cancel_left [simp, no_atp]:
 151.101 +lemma mult_cancel_left [simp]:
 151.102    "c * a = c * b \<longleftrightarrow> c = 0 \<or> a = b"
 151.103  proof -
 151.104    have "(c * a = c * b) = (c * (a - b) = 0)"
 151.105 @@ -742,9 +758,7 @@
 151.106  proof
 151.107    fix a b
 151.108    show "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
 151.109 -    by (auto simp add: abs_if not_less)
 151.110 -    (auto simp del: minus_add_distrib simp add: minus_add_distrib [symmetric],
 151.111 -     auto intro!: less_imp_le add_neg_neg)
 151.112 +    by (auto simp add: abs_if not_le not_less algebra_simps simp del: add.commute dest: add_neg_neg add_nonneg_nonneg)
 151.113  qed (auto simp add: abs_if)
 151.114  
 151.115  lemma zero_le_square [simp]: "0 \<le> a * a"
 151.116 @@ -1048,7 +1062,7 @@
 151.117  
 151.118  text {* Simprules for comparisons where common factors can be cancelled. *}
 151.119  
 151.120 -lemmas mult_compare_simps[no_atp] =
 151.121 +lemmas mult_compare_simps =
 151.122      mult_le_cancel_right mult_le_cancel_left
 151.123      mult_le_cancel_right1 mult_le_cancel_right2
 151.124      mult_le_cancel_left1 mult_le_cancel_left2
 151.125 @@ -1131,10 +1145,6 @@
 151.126    thus ?thesis by (simp add: ac cpos mult_strict_mono) 
 151.127  qed
 151.128  
 151.129 -lemma less_minus_self_iff:
 151.130 -  "a < - a \<longleftrightarrow> a < 0"
 151.131 -  by (simp only: less_le less_eq_neg_nonpos equal_neg_zero)
 151.132 -
 151.133  lemma abs_less_iff:
 151.134    "\<bar>a\<bar> < b \<longleftrightarrow> a < b \<and> - a < b" 
 151.135    by (simp add: less_le abs_le_iff) (auto simp add: abs_if)
   152.1 --- a/src/HOL/Semiring_Normalization.thy	Mon Nov 11 17:34:44 2013 +0100
   152.2 +++ b/src/HOL/Semiring_Normalization.thy	Mon Nov 11 17:44:21 2013 +0100
   152.3 @@ -137,7 +137,7 @@
   152.4  lemma normalizing_ring_rules:
   152.5    "- x = (- 1) * x"
   152.6    "x - y = x + (- y)"
   152.7 -  by (simp_all add: diff_minus)
   152.8 +  by simp_all
   152.9  
  152.10  lemmas normalizing_comm_ring_1_axioms =
  152.11    comm_ring_1_axioms [normalizer
   153.1 --- a/src/HOL/Series.thy	Mon Nov 11 17:34:44 2013 +0100
   153.2 +++ b/src/HOL/Series.thy	Mon Nov 11 17:44:21 2013 +0100
   153.3 @@ -34,7 +34,7 @@
   153.4  
   153.5  lemma sumr_diff_mult_const:
   153.6   "setsum f {0..<n} - (real n*r) = setsum (%i. f i - r) {0..<n::nat}"
   153.7 -by (simp add: diff_minus setsum_addf real_of_nat_def)
   153.8 +  by (simp add: setsum_subtractf real_of_nat_def)
   153.9  
  153.10  lemma real_setsum_nat_ivl_bounded:
  153.11       "(!!p. p < n \<Longrightarrow> f(p) \<le> K)
   154.1 --- a/src/HOL/Set.thy	Mon Nov 11 17:34:44 2013 +0100
   154.2 +++ b/src/HOL/Set.thy	Mon Nov 11 17:44:21 2013 +0100
   154.3 @@ -87,7 +87,7 @@
   154.4    then show ?thesis by simp
   154.5  qed
   154.6  
   154.7 -lemma set_eq_iff [no_atp]:
   154.8 +lemma set_eq_iff:
   154.9    "A = B \<longleftrightarrow> (\<forall>x. x \<in> A \<longleftrightarrow> x \<in> B)"
  154.10    by (auto intro:set_eqI)
  154.11  
  154.12 @@ -495,7 +495,7 @@
  154.13    by (simp add: less_eq_set_def le_fun_def)
  154.14    -- {* Rule in Modus Ponens style. *}
  154.15  
  154.16 -lemma rev_subsetD [no_atp,intro?]: "c \<in> A ==> A \<subseteq> B ==> c \<in> B"
  154.17 +lemma rev_subsetD [intro?]: "c \<in> A ==> A \<subseteq> B ==> c \<in> B"
  154.18    -- {* The same, with reversed premises for use with @{text erule} --
  154.19        cf @{text rev_mp}. *}
  154.20    by (rule subsetD)
  154.21 @@ -504,13 +504,13 @@
  154.22    \medskip Converts @{prop "A \<subseteq> B"} to @{prop "x \<in> A ==> x \<in> B"}.
  154.23  *}
  154.24  
  154.25 -lemma subsetCE [no_atp,elim]: "A \<subseteq> B ==> (c \<notin> A ==> P) ==> (c \<in> B ==> P) ==> P"
  154.26 +lemma subsetCE [elim]: "A \<subseteq> B ==> (c \<notin> A ==> P) ==> (c \<in> B ==> P) ==> P"
  154.27    -- {* Classical elimination rule. *}
  154.28    by (auto simp add: less_eq_set_def le_fun_def)
  154.29  
  154.30 -lemma subset_eq [no_atp]: "A \<le> B = (\<forall>x\<in>A. x \<in> B)" by blast
  154.31 -
  154.32 -lemma contra_subsetD [no_atp]: "A \<subseteq> B ==> c \<notin> B ==> c \<notin> A"
  154.33 +lemma subset_eq: "A \<le> B = (\<forall>x\<in>A. x \<in> B)" by blast
  154.34 +
  154.35 +lemma contra_subsetD: "A \<subseteq> B ==> c \<notin> B ==> c \<notin> A"
  154.36    by blast
  154.37  
  154.38  lemma subset_refl: "A \<subseteq> A"
  154.39 @@ -845,11 +845,11 @@
  154.40  
  154.41  subsubsection {* Singletons, using insert *}
  154.42  
  154.43 -lemma singletonI [intro!,no_atp]: "a : {a}"
  154.44 +lemma singletonI [intro!]: "a : {a}"
  154.45      -- {* Redundant? But unlike @{text insertCI}, it proves the subgoal immediately! *}
  154.46    by (rule insertI1)
  154.47  
  154.48 -lemma singletonD [dest!,no_atp]: "b : {a} ==> b = a"
  154.49 +lemma singletonD [dest!]: "b : {a} ==> b = a"
  154.50    by blast
  154.51  
  154.52  lemmas singletonE = singletonD [elim_format]
  154.53 @@ -860,11 +860,11 @@
  154.54  lemma singleton_inject [dest!]: "{a} = {b} ==> a = b"
  154.55    by blast
  154.56  
  154.57 -lemma singleton_insert_inj_eq [iff,no_atp]:
  154.58 +lemma singleton_insert_inj_eq [iff]:
  154.59       "({b} = insert a A) = (a = b & A \<subseteq> {b})"
  154.60    by blast
  154.61  
  154.62 -lemma singleton_insert_inj_eq' [iff,no_atp]:
  154.63 +lemma singleton_insert_inj_eq' [iff]:
  154.64       "(insert a A = {b}) = (a = b & A \<subseteq> {b})"
  154.65    by blast
  154.66  
  154.67 @@ -898,7 +898,7 @@
  154.68  *}
  154.69  
  154.70  definition image :: "('a => 'b) => 'a set => 'b set" (infixr "`" 90) where
  154.71 -  image_def [no_atp]: "f ` A = {y. EX x:A. y = f(x)}"
  154.72 +  image_def: "f ` A = {y. EX x:A. y = f(x)}"
  154.73  
  154.74  abbreviation
  154.75    range :: "('a => 'b) => 'b set" where -- "of function"
  154.76 @@ -930,7 +930,7 @@
  154.77  lemma image_iff: "(z : f`A) = (EX x:A. z = f x)"
  154.78    by blast
  154.79  
  154.80 -lemma image_subset_iff [no_atp]: "(f`A \<subseteq> B) = (\<forall>x\<in>A. f x \<in> B)"
  154.81 +lemma image_subset_iff: "(f`A \<subseteq> B) = (\<forall>x\<in>A. f x \<in> B)"
  154.82    -- {* This rewrite rule would confuse users if made default. *}
  154.83    by blast
  154.84  
  154.85 @@ -1009,10 +1009,10 @@
  154.86  
  154.87  subsubsection {* The ``proper subset'' relation *}
  154.88  
  154.89 -lemma psubsetI [intro!,no_atp]: "A \<subseteq> B ==> A \<noteq> B ==> A \<subset> B"
  154.90 +lemma psubsetI [intro!]: "A \<subseteq> B ==> A \<noteq> B ==> A \<subset> B"
  154.91    by (unfold less_le) blast
  154.92  
  154.93 -lemma psubsetE [elim!,no_atp]:
  154.94 +lemma psubsetE [elim!]:
  154.95      "[|A \<subset> B;  [|A \<subseteq> B; ~ (B\<subseteq>A)|] ==> R|] ==> R"
  154.96    by (unfold less_le) blast
  154.97  
  154.98 @@ -1184,12 +1184,12 @@
  154.99  lemma insert_inter_insert[simp]: "insert a A \<inter> insert a B = insert a (A \<inter> B)"
 154.100    by blast
 154.101  
 154.102 -lemma insert_disjoint [simp,no_atp]:
 154.103 +lemma insert_disjoint [simp]:
 154.104   "(insert a A \<inter> B = {}) = (a \<notin> B \<and> A \<inter> B = {})"
 154.105   "({} = insert a A \<inter> B) = (a \<notin> B \<and> {} = A \<inter> B)"
 154.106    by auto
 154.107  
 154.108 -lemma disjoint_insert [simp,no_atp]:
 154.109 +lemma disjoint_insert [simp]:
 154.110   "(B \<inter> insert a A = {}) = (a \<notin> B \<and> B \<inter> A = {})"
 154.111   "({} = A \<inter> insert b B) = (b \<notin> A \<and> {} = A \<inter> B)"
 154.112    by auto
 154.113 @@ -1221,7 +1221,7 @@
 154.114  by blast
 154.115  
 154.116  
 154.117 -lemma image_Collect [no_atp]: "f ` {x. P x} = {f x | x. P x}"
 154.118 +lemma image_Collect: "f ` {x. P x} = {f x | x. P x}"
 154.119    -- {* NOT suitable as a default simprule: the RHS isn't simpler than the LHS,
 154.120        with its implicit quantifier and conjunction.  Also image enjoys better
 154.121        equational properties than does the RHS. *}
 154.122 @@ -1244,7 +1244,7 @@
 154.123  
 154.124  text {* \medskip @{text range}. *}
 154.125  
 154.126 -lemma full_SetCompr_eq [no_atp]: "{u. \<exists>x. u = f x} = range f"
 154.127 +lemma full_SetCompr_eq: "{u. \<exists>x. u = f x} = range f"
 154.128    by auto
 154.129  
 154.130  lemma range_composition: "range (\<lambda>x. f (g x)) = f`range g"
 154.131 @@ -1301,10 +1301,10 @@
 154.132  lemma Int_Un_distrib2: "(B \<union> C) \<inter> A = (B \<inter> A) \<union> (C \<inter> A)"
 154.133    by (fact inf_sup_distrib2)
 154.134  
 154.135 -lemma Int_UNIV [simp,no_atp]: "(A \<inter> B = UNIV) = (A = UNIV & B = UNIV)"
 154.136 +lemma Int_UNIV [simp]: "(A \<inter> B = UNIV) = (A = UNIV & B = UNIV)"
 154.137    by (fact inf_eq_top_iff) (* already simp *)
 154.138  
 154.139 -lemma Int_subset_iff [no_atp, simp]: "(C \<subseteq> A \<inter> B) = (C \<subseteq> A & C \<subseteq> B)"
 154.140 +lemma Int_subset_iff [simp]: "(C \<subseteq> A \<inter> B) = (C \<subseteq> A & C \<subseteq> B)"
 154.141    by (fact le_inf_iff)
 154.142  
 154.143  lemma Int_Collect: "(x \<in> A \<inter> {x. P x}) = (x \<in> A & P x)"
 154.144 @@ -1395,7 +1395,7 @@
 154.145  lemma Un_empty [iff]: "(A \<union> B = {}) = (A = {} & B = {})"
 154.146    by (fact sup_eq_bot_iff) (* FIXME: already simp *)
 154.147  
 154.148 -lemma Un_subset_iff [no_atp, simp]: "(A \<union> B \<subseteq> C) = (A \<subseteq> C & B \<subseteq> C)"
 154.149 +lemma Un_subset_iff [simp]: "(A \<union> B \<subseteq> C) = (A \<subseteq> C & B \<subseteq> C)"
 154.150    by (fact le_sup_iff)
 154.151  
 154.152  lemma Un_Diff_Int: "(A - B) \<union> (A \<inter> B) = A"
 154.153 @@ -1467,7 +1467,7 @@
 154.154  lemma Diff_eq: "A - B = A \<inter> (-B)"
 154.155    by blast
 154.156  
 154.157 -lemma Diff_eq_empty_iff [simp,no_atp]: "(A - B = {}) = (A \<subseteq> B)"
 154.158 +lemma Diff_eq_empty_iff [simp]: "(A - B = {}) = (A \<subseteq> B)"
 154.159    by blast
 154.160  
 154.161  lemma Diff_cancel [simp]: "A - A = {}"
 154.162 @@ -1488,7 +1488,7 @@
 154.163  lemma Diff_UNIV [simp]: "A - UNIV = {}"
 154.164    by blast
 154.165  
 154.166 -lemma Diff_insert0 [simp,no_atp]: "x \<notin> A ==> A - insert x B = A - B"
 154.167 +lemma Diff_insert0 [simp]: "x \<notin> A ==> A - insert x B = A - B"
 154.168    by blast
 154.169  
 154.170  lemma Diff_insert: "A - insert a B = A - B - {a}"
 154.171 @@ -1568,7 +1568,7 @@
 154.172  lemma ex_bool_eq: "(\<exists>b. P b) \<longleftrightarrow> P True \<or> P False"
 154.173    by (auto intro: bool_contrapos)
 154.174  
 154.175 -lemma UNIV_bool [no_atp]: "UNIV = {False, True}"
 154.176 +lemma UNIV_bool: "UNIV = {False, True}"
 154.177    by (auto intro: bool_induct)
 154.178  
 154.179  text {* \medskip @{text Pow} *}
 154.180 @@ -1597,7 +1597,7 @@
 154.181  lemma set_eq_subset: "(A = B) = (A \<subseteq> B & B \<subseteq> A)"
 154.182    by blast
 154.183  
 154.184 -lemma subset_iff [no_atp]: "(A \<subseteq> B) = (\<forall>t. t \<in> A --> t \<in> B)"
 154.185 +lemma subset_iff: "(A \<subseteq> B) = (\<forall>t. t \<in> A --> t \<in> B)"
 154.186    by blast
 154.187  
 154.188  lemma subset_iff_psubset_eq: "(A \<subseteq> B) = ((A \<subset> B) | (A = B))"
 154.189 @@ -1754,7 +1754,7 @@
 154.190    -- {* monotonicity *}
 154.191    by blast
 154.192  
 154.193 -lemma vimage_image_eq [no_atp]: "f -` (f ` A) = {y. EX x:A. f x = f y}"
 154.194 +lemma vimage_image_eq: "f -` (f ` A) = {y. EX x:A. f x = f y}"
 154.195  by (blast intro: sym)
 154.196  
 154.197  lemma image_vimage_subset: "f ` (f -` A) <= A"
   155.1 --- a/src/HOL/Set_Interval.thy	Mon Nov 11 17:34:44 2013 +0100
   155.2 +++ b/src/HOL/Set_Interval.thy	Mon Nov 11 17:44:21 2013 +0100
   155.3 @@ -180,19 +180,19 @@
   155.4  context ord
   155.5  begin
   155.6  
   155.7 -lemma greaterThanLessThan_iff [simp,no_atp]:
   155.8 +lemma greaterThanLessThan_iff [simp]:
   155.9    "(i : {l<..<u}) = (l < i & i < u)"
  155.10  by (simp add: greaterThanLessThan_def)
  155.11  
  155.12 -lemma atLeastLessThan_iff [simp,no_atp]:
  155.13 +lemma atLeastLessThan_iff [simp]:
  155.14    "(i : {l..<u}) = (l <= i & i < u)"
  155.15  by (simp add: atLeastLessThan_def)
  155.16  
  155.17 -lemma greaterThanAtMost_iff [simp,no_atp]:
  155.18 +lemma greaterThanAtMost_iff [simp]:
  155.19    "(i : {l<..u}) = (l < i & i <= u)"
  155.20  by (simp add: greaterThanAtMost_def)
  155.21  
  155.22 -lemma atLeastAtMost_iff [simp,no_atp]:
  155.23 +lemma atLeastAtMost_iff [simp]:
  155.24    "(i : {l..u}) = (l <= i & i <= u)"
  155.25  by (simp add: atLeastAtMost_def)
  155.26  
  155.27 @@ -1196,7 +1196,7 @@
  155.28  
  155.29  subsubsection {* Some Subset Conditions *}
  155.30  
  155.31 -lemma ivl_subset [simp,no_atp]:
  155.32 +lemma ivl_subset [simp]:
  155.33   "({i..<j} \<subseteq> {m..<n}) = (j \<le> i | m \<le> i & j \<le> (n::'a::linorder))"
  155.34  apply(auto simp:linorder_not_le)
  155.35  apply(rule ccontr)
   156.1 --- a/src/HOL/TPTP/MaSh_Export.thy	Mon Nov 11 17:34:44 2013 +0100
   156.2 +++ b/src/HOL/TPTP/MaSh_Export.thy	Mon Nov 11 17:44:21 2013 +0100
   156.3 @@ -63,7 +63,7 @@
   156.4  
   156.5  ML {*
   156.6  if do_it then
   156.7 -  generate_isar_dependencies @{context} thys linearize
   156.8 +  generate_isar_dependencies @{context} range thys linearize
   156.9        (prefix ^ "mash_dependencies")
  156.10  else
  156.11    ()
   157.1 --- a/src/HOL/TPTP/atp_theory_export.ML	Mon Nov 11 17:34:44 2013 +0100
   157.2 +++ b/src/HOL/TPTP/atp_theory_export.ML	Mon Nov 11 17:44:21 2013 +0100
   157.3 @@ -34,8 +34,8 @@
   157.4  val prefix = Library.prefix
   157.5  val fact_name_of = prefix fact_prefix o ascii_of
   157.6  
   157.7 -fun atp_of_format (THF (Polymorphic, _, _)) = dummy_thfN
   157.8 -  | atp_of_format (THF (Monomorphic, _, _)) = satallaxN
   157.9 +fun atp_of_format (THF (Polymorphic, _)) = dummy_thfN
  157.10 +  | atp_of_format (THF (Monomorphic, _)) = satallaxN
  157.11    | atp_of_format (DFG Polymorphic) = spass_polyN
  157.12    | atp_of_format (DFG Monomorphic) = spassN
  157.13    | atp_of_format (TFF Polymorphic) = alt_ergoN
   158.1 --- a/src/HOL/TPTP/mash_eval.ML	Mon Nov 11 17:34:44 2013 +0100
   158.2 +++ b/src/HOL/TPTP/mash_eval.ML	Mon Nov 11 17:44:21 2013 +0100
   158.3 @@ -147,7 +147,7 @@
   158.4                    |> map fact_of_raw_fact
   158.5                  val ctxt = ctxt |> set_file_name method prob_dir_name
   158.6                  val res as {outcome, ...} =
   158.7 -                  run_prover_for_mash ctxt params prover facts goal
   158.8 +                  run_prover_for_mash ctxt params prover name facts goal
   158.9                  val ok = if is_none outcome then 1 else 0
  158.10                in (str_of_result method facts res, ok) end
  158.11            val ress =
   159.1 --- a/src/HOL/TPTP/mash_export.ML	Mon Nov 11 17:34:44 2013 +0100
   159.2 +++ b/src/HOL/TPTP/mash_export.ML	Mon Nov 11 17:44:21 2013 +0100
   159.3 @@ -14,7 +14,7 @@
   159.4    val generate_features :
   159.5      Proof.context -> string -> theory list -> string -> unit
   159.6    val generate_isar_dependencies :
   159.7 -    Proof.context -> theory list -> bool -> string -> unit
   159.8 +    Proof.context -> int * int option -> theory list -> bool -> string -> unit
   159.9    val generate_prover_dependencies :
  159.10      Proof.context -> params -> int * int option -> theory list -> bool -> string
  159.11      -> unit
  159.12 @@ -79,9 +79,7 @@
  159.13        let
  159.14          val name = nickname_of_thm th
  159.15          val feats =
  159.16 -          features_of ctxt prover (theory_of_thm th) 0 Symtab.empty stature
  159.17 -                      [prop_of th]
  159.18 -          |> map fst
  159.19 +          features_of ctxt (theory_of_thm th) 0 Symtab.empty stature [prop_of th] |> map fst
  159.20          val s =
  159.21            encode_str name ^ ": " ^ encode_strs (sort string_ord feats) ^ "\n"
  159.22        in File.append path s end
  159.23 @@ -136,7 +134,7 @@
  159.24    in File.write_list path lines end
  159.25  
  159.26  fun generate_isar_dependencies ctxt =
  159.27 -  generate_isar_or_prover_dependencies ctxt NONE (1, NONE)
  159.28 +  generate_isar_or_prover_dependencies ctxt NONE
  159.29  
  159.30  fun generate_prover_dependencies ctxt params =
  159.31    generate_isar_or_prover_dependencies ctxt (SOME params)
  159.32 @@ -164,8 +162,7 @@
  159.33            val isar_deps = isar_dependencies_of name_tabs th
  159.34            val do_query = not (is_bad_query ctxt ho_atp step j th isar_deps)
  159.35            val goal_feats =
  159.36 -            features_of ctxt prover (theory_of_thm th) (num_old_facts + j)
  159.37 -                        const_tab stature [prop_of th]
  159.38 +            features_of ctxt (theory_of_thm th) (num_old_facts + j) const_tab stature [prop_of th]
  159.39              |> sort_wrt fst
  159.40            val access_facts =
  159.41              (if linearize then take (j - 1) new_facts
  159.42 @@ -176,8 +173,7 @@
  159.43            val parents = if linearize then prevs else parents
  159.44            fun extra_features_of (((_, stature), th), weight) =
  159.45              [prop_of th]
  159.46 -            |> features_of ctxt prover (theory_of_thm th) (num_old_facts + j)
  159.47 -                           const_tab stature
  159.48 +            |> features_of ctxt (theory_of_thm th) (num_old_facts + j) const_tab stature
  159.49              |> map (apsnd (fn r => weight * extra_feature_factor * r))
  159.50            val query =
  159.51              if do_query then
  159.52 @@ -249,8 +245,8 @@
  159.53                val suggs =
  159.54                  old_facts
  159.55                  |> linearize ? filter_accessible_from th
  159.56 -                |> Sledgehammer_MePo.mepo_suggested_facts ctxt params prover
  159.57 -                       max_suggs NONE hyp_ts concl_t
  159.58 +                |> Sledgehammer_Fact.drop_duplicate_facts
  159.59 +                |> Sledgehammer_MePo.mepo_suggested_facts ctxt params max_suggs NONE hyp_ts concl_t
  159.60                  |> map (nickname_of_thm o snd)
  159.61              in encode_str name ^ ": " ^ encode_strs suggs ^ "\n" end
  159.62          end
   160.1 --- a/src/HOL/TPTP/sledgehammer_tactics.ML	Mon Nov 11 17:34:44 2013 +0100
   160.2 +++ b/src/HOL/TPTP/sledgehammer_tactics.ML	Mon Nov 11 17:44:21 2013 +0100
   160.3 @@ -27,24 +27,22 @@
   160.4  fun run_prover override_params fact_override i n ctxt goal =
   160.5    let
   160.6      val mode = Normal
   160.7 -    val params as {provers, max_facts, slice, ...} =
   160.8 -      default_params ctxt override_params
   160.9 +    val params as {provers, max_facts, ...} = default_params ctxt override_params
  160.10      val name = hd provers
  160.11      val prover = get_prover ctxt mode name
  160.12 -    val default_max_facts = default_max_facts_of_prover ctxt slice name
  160.13 +    val default_max_facts = default_max_facts_of_prover ctxt name
  160.14      val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal goal i ctxt
  160.15      val ho_atp = exists (is_ho_atp ctxt) provers
  160.16      val reserved = reserved_isar_keyword_table ()
  160.17      val css_table = clasimpset_rule_table_of ctxt
  160.18      val facts =
  160.19 -      nearly_all_facts ctxt ho_atp fact_override reserved css_table [] hyp_ts
  160.20 -                       concl_t
  160.21 +      nearly_all_facts ctxt ho_atp fact_override reserved css_table [] hyp_ts concl_t
  160.22        |> relevant_facts ctxt params name
  160.23               (the_default default_max_facts max_facts) fact_override hyp_ts
  160.24               concl_t
  160.25        |> hd |> snd
  160.26      val problem =
  160.27 -      {state = Proof.init ctxt, goal = goal, subgoal = i, subgoal_count = n,
  160.28 +      {comment = "", state = Proof.init ctxt, goal = goal, subgoal = i, subgoal_count = n,
  160.29         factss = [("", facts)]}
  160.30    in
  160.31      (case prover params (K (K (K ""))) problem of
   161.1 --- a/src/HOL/Tools/ATP/atp_problem.ML	Mon Nov 11 17:34:44 2013 +0100
   161.2 +++ b/src/HOL/Tools/ATP/atp_problem.ML	Mon Nov 11 17:44:21 2013 +0100
   161.3 @@ -33,14 +33,13 @@
   161.4  
   161.5    datatype polymorphism = Monomorphic | Polymorphic
   161.6    datatype thf_choice = THF_Without_Choice | THF_With_Choice
   161.7 -  datatype thf_defs = THF_Without_Defs | THF_With_Defs
   161.8  
   161.9    datatype atp_format =
  161.10      CNF |
  161.11      CNF_UEQ |
  161.12      FOF |
  161.13      TFF of polymorphism |
  161.14 -    THF of polymorphism * thf_choice * thf_defs |
  161.15 +    THF of polymorphism * thf_choice |
  161.16      DFG of polymorphism
  161.17  
  161.18    datatype atp_formula_role =
  161.19 @@ -179,14 +178,13 @@
  161.20  
  161.21  datatype polymorphism = Monomorphic | Polymorphic
  161.22  datatype thf_choice = THF_Without_Choice | THF_With_Choice
  161.23 -datatype thf_defs = THF_Without_Defs | THF_With_Defs
  161.24  
  161.25  datatype atp_format =
  161.26    CNF |
  161.27    CNF_UEQ |
  161.28    FOF |
  161.29    TFF of polymorphism |
  161.30 -  THF of polymorphism * thf_choice * thf_defs |
  161.31 +  THF of polymorphism * thf_choice |
  161.32    DFG of polymorphism
  161.33  
  161.34  datatype atp_formula_role =
   162.1 --- a/src/HOL/Tools/ATP/atp_problem_generate.ML	Mon Nov 11 17:34:44 2013 +0100
   162.2 +++ b/src/HOL/Tools/ATP/atp_problem_generate.ML	Mon Nov 11 17:44:21 2013 +0100
   162.3 @@ -91,6 +91,8 @@
   162.4    val atp_logical_consts : string list
   162.5    val atp_irrelevant_consts : string list
   162.6    val atp_widely_irrelevant_consts : string list
   162.7 +  val is_irrelevant_const : string -> bool
   162.8 +  val is_widely_irrelevant_const : string -> bool
   162.9    val atp_schematic_consts_of : term -> typ list Symtab.table
  162.10    val is_type_enc_higher_order : type_enc -> bool
  162.11    val is_type_enc_polymorphic : type_enc -> bool
  162.12 @@ -405,19 +407,25 @@
  162.13  (* These are ignored anyway by the relevance filter (unless they appear in
  162.14     higher-order places) but not by the monomorphizer. *)
  162.15  val atp_logical_consts =
  162.16 -  [@{const_name all}, @{const_name "==>"}, @{const_name "=="},
  162.17 -   @{const_name Trueprop}, @{const_name All}, @{const_name Ex},
  162.18 +  [@{const_name prop}, @{const_name Pure.conjunction}, @{const_name all}, @{const_name "==>"},
  162.19 +   @{const_name "=="}, @{const_name Trueprop}, @{const_name All}, @{const_name Ex},
  162.20     @{const_name Ex1}, @{const_name Ball}, @{const_name Bex}]
  162.21  
  162.22  (* These are either simplified away by "Meson.presimplify" (most of the time) or
  162.23     handled specially via "fFalse", "fTrue", ..., "fequal". *)
  162.24  val atp_irrelevant_consts =
  162.25 -  [@{const_name False}, @{const_name True}, @{const_name Not},
  162.26 -   @{const_name conj}, @{const_name disj}, @{const_name implies},
  162.27 -   @{const_name HOL.eq}, @{const_name If}, @{const_name Let}]
  162.28 +  [@{const_name False}, @{const_name True}, @{const_name Not}, @{const_name conj},
  162.29 +   @{const_name disj}, @{const_name implies}, @{const_name HOL.eq}, @{const_name If},
  162.30 +   @{const_name Let}]
  162.31  
  162.32  val atp_widely_irrelevant_consts = atp_logical_consts @ atp_irrelevant_consts
  162.33  
  162.34 +val atp_irrelevant_const_tab = Symtab.make (map (rpair ()) atp_irrelevant_consts)
  162.35 +val atp_widely_irrelevant_const_tab = Symtab.make (map (rpair ()) atp_widely_irrelevant_consts)
  162.36 +
  162.37 +val is_irrelevant_const = Symtab.defined atp_irrelevant_const_tab
  162.38 +val is_widely_irrelevant_const = Symtab.defined atp_widely_irrelevant_const_tab
  162.39 +
  162.40  fun add_schematic_const (x as (_, T)) =
  162.41    Monomorph.typ_has_tvars T ? Symtab.insert_list (op =) x
  162.42  val add_schematic_consts_of =
  162.43 @@ -689,11 +697,9 @@
  162.44      Raw_Polymorphic With_Phantom_Type_Vars
  162.45    | no_type_classes poly = poly
  162.46  
  162.47 -fun adjust_type_enc (THF (Polymorphic, choice, _))
  162.48 -                    (Native (order, poly, level)) =
  162.49 +fun adjust_type_enc (THF (Polymorphic, choice)) (Native (order, poly, level)) =
  162.50      Native (adjust_order choice order, no_type_classes poly, level)
  162.51 -  | adjust_type_enc (THF (Monomorphic, choice, _))
  162.52 -                         (Native (order, _, level)) =
  162.53 +  | adjust_type_enc (THF (Monomorphic, choice)) (Native (order, _, level)) =
  162.54      Native (adjust_order choice order, Mangled_Monomorphic, level)
  162.55    | adjust_type_enc (TFF Monomorphic) (Native (_, _, level)) =
  162.56      Native (First_Order, Mangled_Monomorphic, level)
  162.57 @@ -1309,7 +1315,7 @@
  162.58       atomic_types = atomic_Ts}
  162.59    end
  162.60  
  162.61 -fun is_format_with_defs (THF (_, _, THF_With_Defs)) = true
  162.62 +fun is_format_with_defs (THF _) = true
  162.63    | is_format_with_defs _ = false
  162.64  
  162.65  fun make_fact ctxt format type_enc iff_for_eq
  162.66 @@ -1557,8 +1563,7 @@
  162.67                    end
  162.68                  | NONE =>
  162.69                    let
  162.70 -                    val pred_sym = top_level andalso not bool_vars
  162.71 -                    val ary =
  162.72 +                    val max_ary =
  162.73                        case unprefix_and_unascii const_prefix s of
  162.74                          SOME s =>
  162.75                          (if String.isSubstring uncurried_alias_sep s then
  162.76 @@ -1568,15 +1573,16 @@
  162.77                             SOME ary0 => Int.min (ary0, ary)
  162.78                           | NONE => ary)
  162.79                        | NONE => ary
  162.80 +                    val pred_sym = top_level andalso max_ary = ary andalso not bool_vars
  162.81                      val min_ary =
  162.82                        case app_op_level of
  162.83 -                        Min_App_Op => ary
  162.84 +                        Min_App_Op => max_ary
  162.85                        | Full_App_Op_And_Predicator => 0
  162.86 -                      | _ => fold (consider_var_ary T) fun_var_Ts ary
  162.87 +                      | _ => fold (consider_var_ary T) fun_var_Ts max_ary
  162.88                    in
  162.89                      Symtab.update_new (s,
  162.90                          {pred_sym = pred_sym, min_ary = min_ary,
  162.91 -                         max_ary = ary, types = [T], in_conj = conj_fact})
  162.92 +                         max_ary = max_ary, types = [T], in_conj = conj_fact})
  162.93                          sym_tab
  162.94                    end)
  162.95               end
  162.96 @@ -2742,12 +2748,10 @@
  162.97    else
  162.98      NONE
  162.99  
 162.100 -fun ctrss_of_descr descr =
 162.101 -  map_filter (ctrs_of_datatype descr) descr
 162.102 +fun ctrss_of_descr descr = map_filter (ctrs_of_datatype descr) descr
 162.103  
 162.104  fun all_ctrss_of_datatypes thy =
 162.105 -  Symtab.fold (snd #> #descr #> ctrss_of_descr #> append) (Datatype.get_all thy)
 162.106 -              []
 162.107 +  Symtab.fold (snd #> #descr #> ctrss_of_descr #> append) (Datatype.get_all thy) []
 162.108  
 162.109  val app_op_and_predicator_threshold = 45
 162.110  
 162.111 @@ -2757,35 +2761,26 @@
 162.112    let
 162.113      val thy = Proof_Context.theory_of ctxt
 162.114      val type_enc = type_enc |> adjust_type_enc format
 162.115 +    val exporter = (mode = Exporter)
 162.116 +    val completish = (mode = Sledgehammer_Completish)
 162.117      (* Forcing explicit applications is expensive for polymorphic encodings,
 162.118         because it takes only one existential variable ranging over "'a => 'b" to
 162.119         ruin everything. Hence we do it only if there are few facts (which is
 162.120         normally the case for "metis" and the minimizer). *)
 162.121      val app_op_level =
 162.122 -      if mode = Sledgehammer_Completish then
 162.123 +      if completish then
 162.124          Full_App_Op_And_Predicator
 162.125 -      else if length facts + length hyp_ts
 162.126 -              > app_op_and_predicator_threshold then
 162.127 -        if is_type_enc_polymorphic type_enc then Min_App_Op
 162.128 -        else Sufficient_App_Op
 162.129 +      else if length facts + length hyp_ts >= app_op_and_predicator_threshold then
 162.130 +        if is_type_enc_polymorphic type_enc then Min_App_Op else Sufficient_App_Op
 162.131        else
 162.132          Sufficient_App_Op_And_Predicator
 162.133 -    val exporter = (mode = Exporter)
 162.134 -    val completish = (mode = Sledgehammer_Completish)
 162.135      val lam_trans =
 162.136 -      if lam_trans = keep_lamsN andalso
 162.137 -         not (is_type_enc_higher_order type_enc) then
 162.138 -        liftingN
 162.139 -      else
 162.140 -        lam_trans
 162.141 -    val (fact_names, classes, conjs, facts, subclass_pairs, tcon_clauses,
 162.142 -         lifted) =
 162.143 -      translate_formulas ctxt prem_role format type_enc lam_trans preproc hyp_ts
 162.144 -                         concl_t facts
 162.145 -    val (_, sym_tab0) =
 162.146 -      sym_table_of_facts ctxt type_enc app_op_level conjs facts
 162.147 -    val mono =
 162.148 -      conjs @ facts |> mononotonicity_info_of_facts ctxt type_enc completish
 162.149 +      if lam_trans = keep_lamsN andalso not (is_type_enc_higher_order type_enc) then liftingN
 162.150 +      else lam_trans
 162.151 +    val (fact_names, classes, conjs, facts, subclass_pairs, tcon_clauses, lifted) =
 162.152 +      translate_formulas ctxt prem_role format type_enc lam_trans preproc hyp_ts concl_t facts
 162.153 +    val (_, sym_tab0) = sym_table_of_facts ctxt type_enc app_op_level conjs facts
 162.154 +    val mono = conjs @ facts |> mononotonicity_info_of_facts ctxt type_enc completish
 162.155      val ctrss = all_ctrss_of_datatypes thy
 162.156      fun firstorderize in_helper =
 162.157        firstorderize_fact thy ctrss type_enc
   163.1 --- a/src/HOL/Tools/ATP/atp_systems.ML	Mon Nov 11 17:34:44 2013 +0100
   163.2 +++ b/src/HOL/Tools/ATP/atp_systems.ML	Mon Nov 11 17:44:21 2013 +0100
   163.3 @@ -211,7 +211,7 @@
   163.4  
   163.5  (* agsyHOL *)
   163.6  
   163.7 -val agsyhol_thf0 = THF (Monomorphic, THF_Without_Choice, THF_With_Defs)
   163.8 +val agsyhol_thf0 = THF (Monomorphic, THF_Without_Choice)
   163.9  
  163.10  val agsyhol_config : atp_config =
  163.11    {exec = K (["AGSYHOL_HOME"], ["agsyHOL"]),
  163.12 @@ -461,7 +461,7 @@
  163.13  
  163.14  (* LEO-II supports definitions, but it performs significantly better on our
  163.15     benchmarks when they are not used. *)
  163.16 -val leo2_thf0 = THF (Monomorphic, THF_Without_Choice, THF_Without_Defs)
  163.17 +val leo2_thf0 = THF (Monomorphic, THF_Without_Choice)
  163.18  
  163.19  val leo2_config : atp_config =
  163.20    {exec = K (["LEO2_HOME"], ["leo.opt", "leo"]),
  163.21 @@ -488,7 +488,7 @@
  163.22  (* Satallax *)
  163.23  
  163.24  (* Choice is disabled until there is proper reconstruction for it. *)
  163.25 -val satallax_thf0 = THF (Monomorphic, THF_Without_Choice, THF_With_Defs)
  163.26 +val satallax_thf0 = THF (Monomorphic, THF_Without_Choice)
  163.27  
  163.28  val satallax_config : atp_config =
  163.29    {exec = K (["SATALLAX_HOME"], ["satallax.opt", "satallax"]),
  163.30 @@ -653,7 +653,7 @@
  163.31     best_max_mono_iters = default_max_mono_iters,
  163.32     best_max_new_mono_instances = default_max_new_mono_instances}
  163.33  
  163.34 -val dummy_thf_format = THF (Polymorphic, THF_With_Choice, THF_With_Defs)
  163.35 +val dummy_thf_format = THF (Polymorphic, THF_With_Choice)
  163.36  val dummy_thf_config =
  163.37    dummy_config Hypothesis dummy_thf_format "poly_native_higher" false
  163.38  val dummy_thf = (dummy_thfN, fn () => dummy_thf_config)
   164.1 --- a/src/HOL/Tools/Function/function_common.ML	Mon Nov 11 17:34:44 2013 +0100
   164.2 +++ b/src/HOL/Tools/Function/function_common.ML	Mon Nov 11 17:44:21 2013 +0100
   164.3 @@ -119,7 +119,7 @@
   164.4  
   164.5  fun PROFILE msg = if !profile then timeap_msg msg else I
   164.6  
   164.7 -val acc_const_name = @{const_name accp}
   164.8 +val acc_const_name = @{const_name Wellfounded.accp}
   164.9  fun mk_acc domT R =
  164.10    Const (acc_const_name, (domT --> domT --> HOLogic.boolT) --> domT --> HOLogic.boolT) $ R 
  164.11  
   165.1 --- a/src/HOL/Tools/Nitpick/nitpick_hol.ML	Mon Nov 11 17:34:44 2013 +0100
   165.2 +++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML	Mon Nov 11 17:44:21 2013 +0100
   165.3 @@ -1648,8 +1648,7 @@
   165.4      fun do_numeral depth Ts mult T t0 t1 =
   165.5        (if is_number_type ctxt T then
   165.6           let
   165.7 -           val j = mult * (HOLogic.dest_num t1)
   165.8 -                   |> T = nat_T ? Integer.max 0
   165.9 +           val j = mult * HOLogic.dest_num t1
  165.10           in
  165.11             if j = 1 then
  165.12               raise SAME ()
   166.1 --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML	Mon Nov 11 17:34:44 2013 +0100
   166.2 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML	Mon Nov 11 17:44:21 2013 +0100
   166.3 @@ -280,13 +280,15 @@
   166.4              |> map (fn (resultt, (_, prems)) =>
   166.5                Logic.list_implies (prems, HOLogic.mk_Trueprop (list_comb (pred, args @ [resultt]))))
   166.6            end
   166.7 -      val intr_ts = maps mk_intros ((funs ~~ preds) ~~ (argss' ~~ rhss))
   166.8 +      val intr_tss = map mk_intros ((funs ~~ preds) ~~ (argss' ~~ rhss))
   166.9        val (intrs, thy') = thy
  166.10          |> Sign.add_consts_i
  166.11            (map (fn Const (name, T) => (Binding.name (Long_Name.base_name name), T, NoSyn))
  166.12             dst_preds)
  166.13          |> fold_map Specification.axiom
  166.14 -            (map (fn t => ((Binding.name ("unnamed_axiom_" ^ serial_string ()), []), t)) intr_ts)
  166.15 +            (map_index (fn (j, (predname, t)) =>
  166.16 +                ((Binding.name (Long_Name.base_name predname ^ "_intro_" ^ string_of_int (j + 1)), []), t))
  166.17 +              (maps (uncurry (map o pair)) (prednames ~~ intr_tss)))
  166.18        val specs = map (fn predname => (predname,
  166.19            map Drule.export_without_context (filter (Predicate_Compile_Aux.is_intro predname) intrs)))
  166.20          dst_prednames
   167.1 --- a/src/HOL/Tools/Qelim/qelim.ML	Mon Nov 11 17:34:44 2013 +0100
   167.2 +++ b/src/HOL/Tools/Qelim/qelim.ML	Mon Nov 11 17:44:21 2013 +0100
   167.3 @@ -37,7 +37,7 @@
   167.4       val th = Thm.abstract_rule s x ((conv env' then_conv ncv env') p')
   167.5                     |> Drule.arg_cong_rule e
   167.6       val th' = simpex_conv (Thm.rhs_of th)
   167.7 -     val (l,r) = Thm.dest_equals (cprop_of th')
   167.8 +     val (_, r) = Thm.dest_equals (cprop_of th')
   167.9      in if Thm.is_reflexive th' then Thm.transitive th (qcv env (Thm.rhs_of th))
  167.10         else Thm.transitive (Thm.transitive th th') (conv env r) end
  167.11    | Const(@{const_name Ex},_)$ _ => (Thm.eta_long_conversion then_conv conv env) p
   168.1 --- a/src/HOL/Tools/SMT/smt_normalize.ML	Mon Nov 11 17:34:44 2013 +0100
   168.2 +++ b/src/HOL/Tools/SMT/smt_normalize.ML	Mon Nov 11 17:44:21 2013 +0100
   168.3 @@ -532,9 +532,7 @@
   168.4    *)
   168.5  
   168.6    fun is_strange_number ctxt (t as Const (@{const_name neg_numeral}, _) $ _) =
   168.7 -        (case try HOLogic.dest_number t of
   168.8 -          SOME (_, i) => SMT_Builtin.is_builtin_num ctxt t andalso i < 2
   168.9 -        | NONE => false)
  168.10 +        SMT_Builtin.is_builtin_num ctxt t
  168.11      | is_strange_number _ _ = false
  168.12  
  168.13    val pos_num_ss =
   169.1 --- a/src/HOL/Tools/Sledgehammer/MaSh/etc/settings	Mon Nov 11 17:34:44 2013 +0100
   169.2 +++ b/src/HOL/Tools/Sledgehammer/MaSh/etc/settings	Mon Nov 11 17:44:21 2013 +0100
   169.3 @@ -3,4 +3,6 @@
   169.4  ISABELLE_SLEDGEHAMMER_MASH="$COMPONENT"
   169.5  
   169.6  # MASH=yes
   169.7 -MASH_PORT=9255
   169.8 +if [ -z "$MASH_PORT" ]; then
   169.9 +  MASH_PORT=9255
  169.10 +fi
   170.1 --- a/src/HOL/Tools/Sledgehammer/MaSh/src/dictionaries.py	Mon Nov 11 17:34:44 2013 +0100
   170.2 +++ b/src/HOL/Tools/Sledgehammer/MaSh/src/dictionaries.py	Mon Nov 11 17:44:21 2013 +0100
   170.3 @@ -4,11 +4,12 @@
   170.4  #
   170.5  # Persistent dictionaries: accessibility, dependencies, and features.
   170.6  
   170.7 -import logging,sys
   170.8 +import sys
   170.9  from os.path import join
  170.10  from Queue import Queue
  170.11  from readData import create_accessible_dict,create_dependencies_dict
  170.12  from cPickle import load,dump
  170.13 +from exceptions import LookupError
  170.14  
  170.15  class Dictionaries(object):
  170.16      '''
  170.17 @@ -56,7 +57,6 @@
  170.18          self.changed = True
  170.19  
  170.20      def create_feature_dict(self,inputFile):
  170.21 -        logger = logging.getLogger('create_feature_dict')
  170.22          self.featureDict = {}
  170.23          IS = open(inputFile,'r')
  170.24          for line in IS:
  170.25 @@ -64,7 +64,7 @@
  170.26              name = line[0]
  170.27              # Name Id
  170.28              if self.nameIdDict.has_key(name):
  170.29 -                logger.warning('%s appears twice in the feature file. Aborting.',name)
  170.30 +                raise LookupError('%s appears twice in the feature file. Aborting.'% name)
  170.31                  sys.exit(-1)
  170.32              else:
  170.33                  self.nameIdDict[name] = self.maxNameId
  170.34 @@ -134,6 +134,13 @@
  170.35                          unexpandedQueue.put(a)
  170.36          return list(accessibles)
  170.37  
  170.38 +    def parse_unExpAcc(self,line):
  170.39 +        try:
  170.40 +            unExpAcc = [self.nameIdDict[a.strip()] for a in line.split()]            
  170.41 +        except:
  170.42 +            raise LookupError('Cannot find the accessibles:%s. Accessibles need to be introduced before referring to them.' % line)
  170.43 +        return unExpAcc
  170.44 +
  170.45      def parse_fact(self,line):
  170.46          """
  170.47          Parses a single line, extracting accessibles, features, and dependencies.
  170.48 @@ -147,8 +154,9 @@
  170.49          nameId = self.get_name_id(name)
  170.50          line = line[1].split(';')
  170.51          # Accessible Ids
  170.52 -        unExpAcc = [self.nameIdDict[a.strip()] for a in line[0].split()]
  170.53 -        self.accessibleDict[nameId] = unExpAcc
  170.54 +        #unExpAcc = [self.nameIdDict[a.strip()] for a in line[0].split()]
  170.55 +        #self.accessibleDict[nameId] = unExpAcc
  170.56 +        self.accessibleDict[nameId] = self.parse_unExpAcc(line[0])
  170.57          features = self.get_features(line)
  170.58          self.featureDict[nameId] = features
  170.59          self.dependenciesDict[nameId] = [self.nameIdDict[d.strip()] for d in line[2].split()]        
  170.60 @@ -180,7 +188,7 @@
  170.61          name = None
  170.62          numberOfPredictions = None
  170.63  
  170.64 -        # Check whether there is a problem name:
  170.65 +        # How many predictions should be returned:
  170.66          tmp = line.split('#')
  170.67          if len(tmp) == 2:
  170.68              numberOfPredictions = int(tmp[0].strip())
  170.69 @@ -194,8 +202,11 @@
  170.70  
  170.71          # line = accessibles;features
  170.72          line = line.split(';')
  170.73 +        features = self.get_features(line)
  170.74 +        
  170.75          # Accessible Ids, expand and store the accessibles.
  170.76 -        unExpAcc = [self.nameIdDict[a.strip()] for a in line[0].split()]
  170.77 +        #unExpAcc = [self.nameIdDict[a.strip()] for a in line[0].split()]
  170.78 +        unExpAcc = self.parse_unExpAcc(line[0])        
  170.79          if len(self.expandedAccessibles.keys())>=100:
  170.80              self.expandedAccessibles = {}
  170.81              self.changed = True
  170.82 @@ -205,7 +216,7 @@
  170.83                  self.expandedAccessibles[accId] = self.expand_accessibles(accIdAcc)
  170.84                  self.changed = True
  170.85          accessibles = self.expand_accessibles(unExpAcc)
  170.86 -        features = self.get_features(line)
  170.87 +        
  170.88          # Get hints:
  170.89          if len(line) == 3:
  170.90              hints = [self.nameIdDict[d.strip()] for d in line[2].split()]
   171.1 --- a/src/HOL/Tools/Sledgehammer/MaSh/src/mash.py	Mon Nov 11 17:34:44 2013 +0100
   171.2 +++ b/src/HOL/Tools/Sledgehammer/MaSh/src/mash.py	Mon Nov 11 17:44:21 2013 +0100
   171.3 @@ -22,13 +22,25 @@
   171.4  from parameters import init_parser
   171.5  
   171.6  def communicate(data,host,port):
   171.7 -    sock = socket.socket(socket.AF_INET, socket.SOCK_STREAM)
   171.8 +    logger = logging.getLogger('communicate')
   171.9 +    sock = socket.socket(socket.AF_INET, socket.SOCK_STREAM)    
  171.10      try:
  171.11          sock.connect((host,port))
  171.12 -        sock.sendall(data)
  171.13 -        received = sock.recv(4194304)
  171.14 -    except:
  171.15 -        logger = logging.getLogger('communicate')
  171.16 +        sock.sendall(data+'\n')        
  171.17 +        received = ''
  171.18 +        cont = True
  171.19 +        counter = 0
  171.20 +        while cont and counter < 100000:
  171.21 +            rec = sock.recv(4096)
  171.22 +            if rec.endswith('stop'):
  171.23 +                cont = False
  171.24 +                received += rec[:-4]
  171.25 +            else:
  171.26 +                received += rec
  171.27 +            counter += 1
  171.28 +        if rec == '':
  171.29 +            logger.warning('No response from server. Check server log for details.')
  171.30 +    except:        
  171.31          logger.warning('Communication with server failed.')
  171.32          received = -1
  171.33      finally:
   172.1 --- a/src/HOL/Tools/Sledgehammer/MaSh/src/server.py	Mon Nov 11 17:34:44 2013 +0100
   172.2 +++ b/src/HOL/Tools/Sledgehammer/MaSh/src/server.py	Mon Nov 11 17:44:21 2013 +0100
   172.3 @@ -19,7 +19,8 @@
   172.4  from stats import Statistics
   172.5  
   172.6  
   172.7 -class ThreadingTCPServer(SocketServer.ThreadingTCPServer): 
   172.8 +class ThreadingTCPServer(SocketServer.ThreadingTCPServer):
   172.9 +    
  172.10      def __init__(self, *args, **kwargs):
  172.11          SocketServer.ThreadingTCPServer.__init__(self,*args, **kwargs)
  172.12          self.manager = Manager()
  172.13 @@ -27,8 +28,17 @@
  172.14          self.idle_timeout = 28800.0 # 8 hours in seconds
  172.15          self.idle_timer = Timer(self.idle_timeout, self.shutdown)
  172.16          self.idle_timer.start()        
  172.17 +        self.model = None
  172.18 +        self.dicts = None
  172.19 +        self.callCounter = 0
  172.20          
  172.21      def save(self):
  172.22 +        if self.model == None or self.dicts == None:
  172.23 +            try:
  172.24 +                self.logger.warning('Cannot save nonexisting models.')
  172.25 +            except:
  172.26 +                pass
  172.27 +            return
  172.28          # Save Models
  172.29          self.model.save(self.args.modelFile)
  172.30          self.dicts.save(self.args.dictsFile)
  172.31 @@ -40,7 +50,7 @@
  172.32          self.save()          
  172.33          self.shutdown()
  172.34  
  172.35 -class MaShHandler(SocketServer.BaseRequestHandler):
  172.36 +class MaShHandler(SocketServer.StreamRequestHandler):
  172.37  
  172.38      def init(self,argv):
  172.39          if argv == '':
  172.40 @@ -140,7 +150,7 @@
  172.41          #predictionValues = [str(x) for x in predictionValues[:numberOfPredictions]]
  172.42          #predictionsStringList = ['%s=%s' % (predictionNames[i],predictionValues[i]) for i in range(len(predictionNames))]
  172.43          #predictionsString = string.join(predictionsStringList,' ')
  172.44 -        predictionsString = string.join(predictionNames,' ')
  172.45 +        predictionsString = string.join(predictionNames,' ')        
  172.46          outString = '%s: %s' % (name,predictionsString)
  172.47          self.request.sendall(outString)
  172.48      
  172.49 @@ -154,15 +164,15 @@
  172.50  
  172.51      def handle(self):
  172.52          # self.request is the TCP socket connected to the client
  172.53 -        self.data = self.request.recv(4194304).strip()
  172.54          self.server.lock.acquire()
  172.55 +        self.data = self.rfile.readline().strip()
  172.56          try:
  172.57              # Update idle shutdown timer
  172.58              self.server.idle_timer.cancel()
  172.59              self.server.idle_timer = Timer(self.server.idle_timeout, self.server.save_and_shutdown)
  172.60              self.server.idle_timer.start()        
  172.61  
  172.62 -            self.startTime = time()  
  172.63 +            self.startTime = time()
  172.64              if self.data == 'shutdown':
  172.65                  self.shutdown()         
  172.66              elif self.data == 'save':
  172.67 @@ -189,12 +199,19 @@
  172.68              else:
  172.69                  self.request.sendall('Unspecified input format: \n%s',self.data)
  172.70              self.server.callCounter += 1
  172.71 +            self.request.sendall('stop')                       
  172.72 +        except: # catch exceptions
  172.73 +            #print 'Caught an error. Check %s for more details' % (self.server.args.log+'server')
  172.74 +            logging.exception('')
  172.75          finally:
  172.76              self.server.lock.release()
  172.77  
  172.78  if __name__ == "__main__":
  172.79 -    HOST, PORT = sys.argv[1:]    
  172.80 -    #HOST, PORT = "localhost", 9255
  172.81 +    if not len(sys.argv[1:]) == 2:
  172.82 +        print 'No Arguments for HOST and PORT found. Using localhost and 9255'
  172.83 +        HOST, PORT = "localhost", 9255
  172.84 +    else:
  172.85 +        HOST, PORT = sys.argv[1:]
  172.86      SocketServer.TCPServer.allow_reuse_address = True
  172.87      server = ThreadingTCPServer((HOST, int(PORT)), MaShHandler)
  172.88      server.serve_forever()        
   173.1 --- a/src/HOL/Tools/Sledgehammer/MaSh/src/sparseNaiveBayes.py	Mon Nov 11 17:34:44 2013 +0100
   173.2 +++ b/src/HOL/Tools/Sledgehammer/MaSh/src/sparseNaiveBayes.py	Mon Nov 11 17:44:21 2013 +0100
   173.3 @@ -78,6 +78,8 @@
   173.4              self.counts[dep][0] -= 1
   173.5              for f,_w in features.items():
   173.6                  self.counts[dep][1][f] -= 1
   173.7 +                if self.counts[dep][1][f] == 0:
   173.8 +                    del self.counts[dep][1][f]
   173.9  
  173.10  
  173.11      def overwrite(self,problemId,newDependencies,dicts):
   174.1 --- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact.ML	Mon Nov 11 17:34:44 2013 +0100
   174.2 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact.ML	Mon Nov 11 17:44:21 2013 +0100
   174.3 @@ -39,6 +39,7 @@
   174.4    val nearly_all_facts :
   174.5      Proof.context -> bool -> fact_override -> unit Symtab.table
   174.6      -> status Termtab.table -> thm list -> term list -> term -> raw_fact list
   174.7 +  val drop_duplicate_facts : raw_fact list -> raw_fact list
   174.8  end;
   174.9  
  174.10  structure Sledgehammer_Fact : SLEDGEHAMMER_FACT =
  174.11 @@ -57,6 +58,11 @@
  174.12     del : (Facts.ref * Attrib.src list) list,
  174.13     only : bool}
  174.14  
  174.15 +(* gracefully handle huge background theories *)
  174.16 +val max_facts_for_duplicates = 50000
  174.17 +val max_facts_for_complex_check = 25000
  174.18 +val max_simps_for_clasimpset = 10000
  174.19 +
  174.20  (* experimental feature *)
  174.21  val instantiate_inducts =
  174.22    Attrib.setup_config_bool @{binding sledgehammer_instantiate_inducts} (K false)
  174.23 @@ -102,6 +108,7 @@
  174.24                       body_type T = @{typ bool}
  174.25                     | _ => false)
  174.26  
  174.27 +(* TODO: get rid of *)
  174.28  fun normalize_vars t =
  174.29    let
  174.30      fun normT (Type (s, Ts)) = fold_map normT Ts #>> curry Type s
  174.31 @@ -126,25 +133,31 @@
  174.32    in fst (norm t (([], 0), ([], 0))) end
  174.33  
  174.34  fun status_of_thm css name th =
  174.35 -  let val t = normalize_vars (prop_of th) in
  174.36 -    (* FIXME: use structured name *)
  174.37 -    if String.isSubstring ".induct" name andalso may_be_induction t then
  174.38 -      Induction
  174.39 -    else case Termtab.lookup css t of
  174.40 -      SOME status => status
  174.41 -    | NONE =>
  174.42 -      let val concl = Logic.strip_imp_concl t in
  174.43 -        case try (HOLogic.dest_eq o HOLogic.dest_Trueprop) concl of
  174.44 -          SOME lrhss =>
  174.45 -          let
  174.46 -            val prems = Logic.strip_imp_prems t
  174.47 -            val t' = Logic.list_implies (prems, Logic.mk_equals lrhss)
  174.48 -          in
  174.49 -            Termtab.lookup css t' |> the_default General
  174.50 -          end
  174.51 -        | NONE => General
  174.52 -      end
  174.53 -  end
  174.54 +  if Termtab.is_empty css then
  174.55 +    General
  174.56 +  else
  174.57 +    let val t = prop_of th in
  174.58 +      (* FIXME: use structured name *)
  174.59 +      if String.isSubstring ".induct" name andalso may_be_induction t then
  174.60 +        Induction
  174.61 +      else
  174.62 +        let val t = normalize_vars t in
  174.63 +          case Termtab.lookup css t of
  174.64 +            SOME status => status
  174.65 +          | NONE =>
  174.66 +            let val concl = Logic.strip_imp_concl t in
  174.67 +              case try (HOLogic.dest_eq o HOLogic.dest_Trueprop) concl of
  174.68 +                SOME lrhss =>
  174.69 +                let
  174.70 +                  val prems = Logic.strip_imp_prems t
  174.71 +                  val t' = Logic.list_implies (prems, Logic.mk_equals lrhss)
  174.72 +                in
  174.73 +                  Termtab.lookup css t' |> the_default General
  174.74 +                end
  174.75 +              | NONE => General
  174.76 +            end
  174.77 +        end
  174.78 +    end
  174.79  
  174.80  fun stature_of_thm global assms chained css name th =
  174.81    (scope_of_thm global assms chained th, status_of_thm css name th)
  174.82 @@ -219,10 +232,11 @@
  174.83  
  174.84  val sep_that = Long_Name.separator ^ Obtain.thatN
  174.85  
  174.86 +val skolem_thesis = Name.skolem Auto_Bind.thesisN
  174.87 +
  174.88  fun is_that_fact th =
  174.89 -  String.isSuffix sep_that (Thm.get_name_hint th)
  174.90 -  andalso exists_subterm (fn Free (s, _) => s = Name.skolem Auto_Bind.thesisN
  174.91 -                           | _ => false) (prop_of th)
  174.92 +  exists_subterm (fn Free (s, _) => s = skolem_thesis | _ => false) (prop_of th)
  174.93 +  andalso String.isSuffix sep_that (Thm.get_name_hint th)
  174.94  
  174.95  datatype interest = Deal_Breaker | Interesting | Boring
  174.96  
  174.97 @@ -295,9 +309,7 @@
  174.98  fun backquote_term ctxt = close_form #> hackish_string_of_term ctxt #> backquote
  174.99  fun backquote_thm ctxt = backquote_term ctxt o prop_of
 174.100  
 174.101 -(* gracefully handle huge background theories *)
 174.102 -val max_simps_for_clasimpset = 10000
 174.103 -
 174.104 +(* TODO: rewrite to use nets and/or to reuse existing data structures *)
 174.105  fun clasimpset_rule_table_of ctxt =
 174.106    let val simps = ctxt |> simpset_of |> dest_ss |> #simps in
 174.107      if length simps >= max_simps_for_clasimpset then
 174.108 @@ -336,18 +348,16 @@
 174.109        end
 174.110    end
 174.111  
 174.112 -fun normalize_eq (t as @{const Trueprop}
 174.113 -        $ ((t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2)) =
 174.114 -    if Term_Ord.fast_term_ord (t1, t2) <> GREATER then t
 174.115 -    else HOLogic.mk_Trueprop (t0 $ t2 $ t1)
 174.116 -  | normalize_eq (t as @{const Trueprop} $ (@{const Not}
 174.117 +fun normalize_eq (@{const Trueprop} $ (t as (t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2)) =
 174.118 +    if Term_Ord.fast_term_ord (t1, t2) <> GREATER then t else t0 $ t2 $ t1
 174.119 +  | normalize_eq (@{const Trueprop} $ (t as @{const Not}
 174.120          $ ((t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2))) =
 174.121 -    if Term_Ord.fast_term_ord (t1, t2) <> GREATER then t
 174.122 -    else HOLogic.mk_Trueprop (HOLogic.mk_not (t0 $ t2 $ t1))
 174.123 +    if Term_Ord.fast_term_ord (t1, t2) <> GREATER then t else HOLogic.mk_not (t0 $ t2 $ t1)
 174.124 +  | normalize_eq (Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2) =
 174.125 +    (if Term_Ord.fast_term_ord (t1, t2) <> GREATER then (t1, t2) else (t2, t1))
 174.126 +    |> (fn (t1, t2) => HOLogic.eq_const T $ t1 $ t2)
 174.127    | normalize_eq t = t
 174.128  
 174.129 -val normalize_eq_vars = normalize_eq #> normalize_vars
 174.130 -
 174.131  fun if_thm_before th th' =
 174.132    if Theory.subthy (pairself Thm.theory_of_thm (th, th')) then th else th'
 174.133  
 174.134 @@ -361,30 +371,22 @@
 174.135  
 174.136  fun build_name_tables name_of facts =
 174.137    let
 174.138 -    fun cons_thm (_, th) =
 174.139 -      Termtab.cons_list (normalize_eq_vars (prop_of th), th)
 174.140 +    fun cons_thm (_, th) = Termtab.cons_list (normalize_vars (normalize_eq (prop_of th)), th)
 174.141      fun add_plain canon alias =
 174.142 -      Symtab.update (Thm.get_name_hint alias,
 174.143 -                     name_of (if_thm_before canon alias))
 174.144 +      Symtab.update (Thm.get_name_hint alias, name_of (if_thm_before canon alias))
 174.145      fun add_plains (_, aliases as canon :: _) = fold (add_plain canon) aliases
 174.146 -    fun add_inclass (name, target) =
 174.147 -      fold (fn s => Symtab.update (s, target)) (un_class_ify name)
 174.148 +    fun add_inclass (name, target) = fold (fn s => Symtab.update (s, target)) (un_class_ify name)
 174.149      val prop_tab = fold cons_thm facts Termtab.empty
 174.150      val plain_name_tab = Termtab.fold add_plains prop_tab Symtab.empty
 174.151      val inclass_name_tab = Symtab.fold add_inclass plain_name_tab Symtab.empty
 174.152    in (plain_name_tab, inclass_name_tab) end
 174.153  
 174.154 -fun keyed_distinct key_of xs =
 174.155 -  let val tab = fold (Termtab.default o `key_of) xs Termtab.empty in
 174.156 -    Termtab.fold (cons o snd) tab []
 174.157 -  end
 174.158 -
 174.159 -fun hashed_keyed_distinct hash_of key_of xs =
 174.160 -  let
 174.161 -    val ys = map (`hash_of) xs
 174.162 -    val sorted_ys = sort (int_ord o pairself fst) ys
 174.163 -    val grouped_ys = AList.coalesce (op =) sorted_ys
 174.164 -  in maps (keyed_distinct key_of o snd) grouped_ys end
 174.165 +fun fact_distinct eq facts =
 174.166 +  fold (fn fact as (_, th) =>
 174.167 +      Net.insert_term_safe (eq o pairself (normalize_eq o prop_of o snd))
 174.168 +        (normalize_eq (prop_of th), fact))
 174.169 +    facts Net.empty
 174.170 +  |> Net.entries
 174.171  
 174.172  fun struct_induct_rule_on th =
 174.173    case Logic.strip_horn (prop_of th) of
 174.174 @@ -447,9 +449,6 @@
 174.175  
 174.176  fun fact_count facts = Facts.fold_static (K (Integer.add 1)) facts 0
 174.177  
 174.178 -(* gracefully handle huge background theories *)
 174.179 -val max_facts_for_complex_check = 25000
 174.180 -
 174.181  fun all_facts ctxt generous ho_atp reserved add_ths chained css =
 174.182    let
 174.183      val thy = Proof_Context.theory_of ctxt
 174.184 @@ -473,12 +472,12 @@
 174.185        Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts)
 174.186      val is_blacklisted_or_something = is_blacklisted_or_something ctxt ho_atp
 174.187      fun add_facts global foldx facts =
 174.188 -      foldx (fn (name0, ths) =>
 174.189 +      foldx (fn (name0, ths) => fn accum =>
 174.190          if name0 <> "" andalso
 174.191             forall (not o member Thm.eq_thm_prop add_ths) ths andalso
 174.192             (Facts.is_concealed facts name0 orelse
 174.193              (not generous andalso is_blacklisted_or_something name0)) then
 174.194 -          I
 174.195 +          accum
 174.196          else
 174.197            let
 174.198              val n = length ths
 174.199 @@ -488,32 +487,30 @@
 174.200                  NONE => false
 174.201                | SOME ths' => eq_list Thm.eq_thm_prop (ths, ths')
 174.202            in
 174.203 -            pair n
 174.204 -            #> fold_rev (fn th => fn (j, accum) =>
 174.205 -                   (j - 1,
 174.206 -                    if not (member Thm.eq_thm_prop add_ths th) andalso
 174.207 -                       (is_likely_tautology_too_meta_or_too_technical th orelse
 174.208 -                        is_too_complex (prop_of th)) then
 174.209 -                      accum
 174.210 -                    else
 174.211 -                      let
 174.212 -                        val new =
 174.213 -                          (((fn () =>
 174.214 -                                if name0 = "" then
 174.215 -                                  backquote_thm ctxt th
 174.216 -                                else
 174.217 -                                  [Facts.extern ctxt facts name0,
 174.218 -                                   Name_Space.extern ctxt full_space name0]
 174.219 -                                  |> distinct (op =)
 174.220 -                                  |> find_first check_thms
 174.221 -                                  |> the_default name0
 174.222 -                                  |> make_name reserved multi j),
 174.223 -                             stature_of_thm global assms chained css name0 th),
 174.224 -                           th)
 174.225 -                      in
 174.226 -                        accum |> (if multi then apsnd else apfst) (cons new)
 174.227 -                      end)) ths
 174.228 -            #> snd
 174.229 +            snd (fold_rev (fn th => fn (j, accum as (uni_accum, multi_accum)) =>
 174.230 +              (j - 1,
 174.231 +               if not (member Thm.eq_thm_prop add_ths th) andalso
 174.232 +                  (is_likely_tautology_too_meta_or_too_technical th orelse
 174.233 +                   is_too_complex (prop_of th)) then
 174.234 +                 accum
 174.235 +               else
 174.236 +                 let
 174.237 +                   val new =
 174.238 +                     (((fn () =>
 174.239 +                           if name0 = "" then
 174.240 +                             backquote_thm ctxt th
 174.241 +                           else
 174.242 +                             [Facts.extern ctxt facts name0,
 174.243 +                              Name_Space.extern ctxt full_space name0]
 174.244 +                             |> find_first check_thms
 174.245 +                             |> the_default name0
 174.246 +                             |> make_name reserved multi j),
 174.247 +                        stature_of_thm global assms chained css name0 th),
 174.248 +                      th)
 174.249 +                 in
 174.250 +                   if multi then (uni_accum, new :: multi_accum)
 174.251 +                   else (new :: uni_accum, multi_accum)
 174.252 +                 end)) ths (n, accum))
 174.253            end)
 174.254    in
 174.255      (* The single-theorem names go before the multiple-theorem ones (e.g.,
 174.256 @@ -530,6 +527,7 @@
 174.257      []
 174.258    else
 174.259      let
 174.260 +      val thy = Proof_Context.theory_of ctxt
 174.261        val chained =
 174.262          chained
 174.263          |> maps (fn th => insert Thm.eq_thm_prop (zero_var_indexes th) [th])
 174.264 @@ -538,14 +536,25 @@
 174.265           maps (map (fn ((name, stature), th) => ((K name, stature), th))
 174.266                 o fact_of_ref ctxt reserved chained css) add
 174.267         else
 174.268 -         let val (add, del) = pairself (Attrib.eval_thms ctxt) (add, del) in
 174.269 -           all_facts ctxt false ho_atp reserved add chained css
 174.270 -           |> filter_out
 174.271 -                  ((member Thm.eq_thm_prop del orf No_ATPs.member ctxt) o snd)
 174.272 -           |> hashed_keyed_distinct (size_of_term o prop_of o snd)
 174.273 -                  (normalize_eq_vars o prop_of o snd)
 174.274 +        (* The "fact_distinct" call would have cleaner semantics if it called "Pattern.equiv"
 174.275 +           instead of "Pattern.matches", but it would also be slower and less precise.
 174.276 +           "Pattern.matches" throws out theorems that are strict instances of other theorems, but
 174.277 +           only if the instance is met after the general version. *)
 174.278 +         let
 174.279 +           val (add, del) = pairself (Attrib.eval_thms ctxt) (add, del)
 174.280 +           val facts =
 174.281 +             all_facts ctxt false ho_atp reserved add chained css
 174.282 +             |> filter_out ((member Thm.eq_thm_prop del orf
 174.283 +               (No_ATPs.member ctxt andf not o member Thm.eq_thm_prop add)) o snd)
 174.284 +         in
 174.285 +           facts
 174.286           end)
 174.287        |> maybe_instantiate_inducts ctxt hyp_ts concl_t
 174.288      end
 174.289  
 174.290 +fun drop_duplicate_facts facts =
 174.291 +  let val num_facts = length facts in
 174.292 +    facts |> num_facts <= max_facts_for_duplicates ? fact_distinct (op aconv)
 174.293 +  end
 174.294 +
 174.295  end;
   175.1 --- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Mon Nov 11 17:34:44 2013 +0100
   175.2 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Mon Nov 11 17:44:21 2013 +0100
   175.3 @@ -183,6 +183,7 @@
   175.4     read correctly. *)
   175.5  val implode_param = strip_spaces_except_between_idents o space_implode " "
   175.6  
   175.7 +(* FIXME: use "Generic_Data" *)
   175.8  structure Data = Theory_Data
   175.9  (
  175.10    type T = raw_param list
  175.11 @@ -196,22 +197,24 @@
  175.12  
  175.13  (* The first ATP of the list is used by Auto Sledgehammer. Because of the low
  175.14     timeout, it makes sense to put E first. *)
  175.15 -fun default_provers_param_value ctxt =
  175.16 -  [eN, spassN, vampireN, z3N, e_sineN, yicesN]
  175.17 +fun default_provers_param_value mode ctxt =
  175.18 +  [eN, spassN, z3N, vampireN, e_sineN, yicesN]
  175.19    |> map_filter (remotify_prover_if_not_installed ctxt)
  175.20 -  |> take (Multithreading.max_threads_value ())
  175.21 +  (* In "try" mode, leave at least one thread to another slow tool (e.g. Nitpick) *)
  175.22 +  |> take (Multithreading.max_threads_value () - (if mode = Try then 1 else 0))
  175.23    |> implode_param
  175.24  
  175.25  fun set_default_raw_param param thy =
  175.26    let val ctxt = Proof_Context.init_global thy in
  175.27      thy |> Data.map (AList.update (op =) (normalize_raw_param ctxt param))
  175.28    end
  175.29 -fun default_raw_params ctxt =
  175.30 +
  175.31 +fun default_raw_params mode ctxt =
  175.32    let val thy = Proof_Context.theory_of ctxt in
  175.33      Data.get thy
  175.34      |> fold (AList.default (op =))
  175.35              [("provers", [case !provers of
  175.36 -                            "" => default_provers_param_value ctxt
  175.37 +                            "" => default_provers_param_value mode ctxt
  175.38                            | s => s]),
  175.39               ("timeout", let val timeout = Options.default_int @{option sledgehammer_timeout} in
  175.40                             [if timeout <= 0 then "none"
  175.41 @@ -278,7 +281,9 @@
  175.42      val lam_trans = lookup_option lookup_string "lam_trans"
  175.43      val uncurried_aliases = lookup_option lookup_bool "uncurried_aliases"
  175.44      val learn = lookup_bool "learn"
  175.45 -    val fact_filter = lookup_option lookup_string "fact_filter"
  175.46 +    val fact_filter =
  175.47 +      lookup_option lookup_string "fact_filter"
  175.48 +      |> mode = Auto_Try ? (fn NONE => SOME mepoN | sf => sf)
  175.49      val max_facts = lookup_option lookup_int "max_facts"
  175.50      val fact_thresholds = lookup_real_pair "fact_thresholds"
  175.51      val max_mono_iters = lookup_option lookup_int "max_mono_iters"
  175.52 @@ -288,12 +293,10 @@
  175.53      val isar_compress = Real.max (1.0, lookup_real "isar_compress")
  175.54      val isar_try0 = lookup_bool "isar_try0"
  175.55      val slice = mode <> Auto_Try andalso lookup_bool "slice"
  175.56 -    val minimize =
  175.57 -      if mode = Auto_Try then NONE else lookup_option lookup_bool "minimize"
  175.58 +    val minimize = if mode = Auto_Try then NONE else lookup_option lookup_bool "minimize"
  175.59      val timeout = if mode = Auto_Try then NONE else lookup_time "timeout"
  175.60      val preplay_timeout =
  175.61 -      if mode = Auto_Try then SOME Time.zeroTime
  175.62 -      else lookup_time "preplay_timeout"
  175.63 +      if mode = Auto_Try then SOME Time.zeroTime else lookup_time "preplay_timeout"
  175.64      val expect = lookup_string "expect"
  175.65    in
  175.66      {debug = debug, verbose = verbose, overlord = overlord, spy = spy, blocking = blocking,
  175.67 @@ -305,7 +308,7 @@
  175.68       timeout = timeout, preplay_timeout = preplay_timeout, expect = expect}
  175.69    end
  175.70  
  175.71 -fun get_params mode = extract_params mode o default_raw_params
  175.72 +fun get_params mode = extract_params mode o default_raw_params mode
  175.73  fun default_params thy = get_params Normal thy o map (apsnd single)
  175.74  
  175.75  (* Sledgehammer the given subgoal *)
  175.76 @@ -385,6 +388,7 @@
  175.77         else
  175.78           ();
  175.79         mash_learn ctxt
  175.80 +           (* TODO: Use MaSh mode instead and have the special defaults hardcoded in "get_params" *)
  175.81             (get_params Normal ctxt
  175.82                  ([("timeout",
  175.83                     [string_of_real default_learn_prover_timeout]),
  175.84 @@ -415,7 +419,7 @@
  175.85         #> tap (fn thy =>
  175.86                    let val ctxt = Proof_Context.init_global thy in
  175.87                      writeln ("Default parameters for Sledgehammer:\n" ^
  175.88 -                             (case default_raw_params ctxt |> rev of
  175.89 +                             (case rev (default_raw_params Normal ctxt) of
  175.90                                  [] => "none"
  175.91                                | params =>
  175.92                                  params |> map string_of_raw_param
   176.1 --- a/src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML	Mon Nov 11 17:34:44 2013 +0100
   176.2 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML	Mon Nov 11 17:44:21 2013 +0100
   176.3 @@ -11,7 +11,6 @@
   176.4    type fact = Sledgehammer_Fact.fact
   176.5    type fact_override = Sledgehammer_Fact.fact_override
   176.6    type params = Sledgehammer_Provers.params
   176.7 -  type relevance_fudge = Sledgehammer_Provers.relevance_fudge
   176.8    type prover_result = Sledgehammer_Provers.prover_result
   176.9  
  176.10    val trace : bool Config.T
  176.11 @@ -61,10 +60,10 @@
  176.12    val thm_less : thm * thm -> bool
  176.13    val goal_of_thm : theory -> thm -> thm
  176.14    val run_prover_for_mash :
  176.15 -    Proof.context -> params -> string -> fact list -> thm -> prover_result
  176.16 +    Proof.context -> params -> string -> string -> fact list -> thm -> prover_result
  176.17    val features_of :
  176.18 -    Proof.context -> string -> theory -> int -> int Symtab.table -> stature
  176.19 -    -> term list -> (string * real) list
  176.20 +    Proof.context -> theory -> int -> int Symtab.table -> stature -> term list ->
  176.21 +    (string * real) list
  176.22    val trim_dependencies : string list -> string list option
  176.23    val isar_dependencies_of :
  176.24      string Symtab.table * string Symtab.table -> thm -> string list
  176.25 @@ -180,11 +179,10 @@
  176.26        (if background then " &" else "")
  176.27      fun run_on () =
  176.28        (Isabelle_System.bash command
  176.29 -       |> tap (fn _ => trace_msg ctxt (fn () =>
  176.30 -              case try File.read (Path.explode err_file) of
  176.31 -                NONE => "Done"
  176.32 -              | SOME "" => "Done"
  176.33 -              | SOME s => "Error: " ^ elide_string 1000 s));
  176.34 +       |> tap (fn _ =>
  176.35 +            case try File.read (Path.explode err_file) |> the_default "" of
  176.36 +              "" => trace_msg ctxt (K "Done")
  176.37 +            | s => warning ("MaSh error: " ^ elide_string 1000 s));
  176.38         read_suggs (fn () => try File.read_lines sugg_path |> these))
  176.39      fun clean_up () =
  176.40        if overlord then ()
  176.41 @@ -514,20 +512,14 @@
  176.42        val mess = mess |> map (apsnd (apfst (normalize_scores max_facts)))
  176.43        fun score_in fact (global_weight, (sels, unks)) =
  176.44          let
  176.45 -          fun score_at j =
  176.46 -            case try (nth sels) j of
  176.47 -              SOME (_, score) => SOME (global_weight * score)
  176.48 -            | NONE => NONE
  176.49 +          val score_at = try (nth sels) #> Option.map (fn (_, score) => global_weight * score)
  176.50          in
  176.51            case find_index (curry fact_eq fact o fst) sels of
  176.52 -            ~1 => (case find_index (curry fact_eq fact) unks of
  176.53 -                     ~1 => SOME 0.0
  176.54 -                   | _ => NONE)
  176.55 +            ~1 => if member fact_eq unks fact then NONE else SOME 0.0
  176.56            | rank => score_at rank
  176.57          end
  176.58        fun weight_of fact = mess |> map_filter (score_in fact) |> scaled_avg
  176.59 -      val facts =
  176.60 -        fold (union fact_eq o map fst o take max_facts o fst o snd) mess []
  176.61 +      val facts = fold (union fact_eq o map fst o take max_facts o fst o snd) mess []
  176.62      in
  176.63        facts |> map (`weight_of) |> sort (int_ord o swap o pairself fst)
  176.64              |> map snd |> take max_facts
  176.65 @@ -573,11 +565,11 @@
  176.66  
  176.67  fun goal_of_thm thy = prop_of #> freeze #> cterm_of thy #> Goal.init
  176.68  
  176.69 -fun run_prover_for_mash ctxt params prover facts goal =
  176.70 +fun run_prover_for_mash ctxt params prover goal_name facts goal =
  176.71    let
  176.72      val problem =
  176.73 -      {state = Proof.init ctxt, goal = goal, subgoal = 1, subgoal_count = 1,
  176.74 -       factss = [("", facts)]}
  176.75 +      {comment = "Goal: " ^ goal_name, state = Proof.init ctxt, goal = goal, subgoal = 1,
  176.76 +       subgoal_count = 1, factss = [("", facts)]}
  176.77    in
  176.78      get_minimizing_prover ctxt MaSh (K (K ())) prover params (K (K (K "")))
  176.79                            problem
  176.80 @@ -585,9 +577,6 @@
  176.81  
  176.82  val bad_types = [@{type_name prop}, @{type_name bool}, @{type_name fun}]
  176.83  
  176.84 -val logical_consts =
  176.85 -  [@{const_name prop}, @{const_name Pure.conjunction}] @ atp_logical_consts
  176.86 -
  176.87  val pat_tvar_prefix = "_"
  176.88  val pat_var_prefix = "_"
  176.89  
  176.90 @@ -608,15 +597,10 @@
  176.91  
  176.92  val max_pat_breadth = 10 (* FUDGE *)
  176.93  
  176.94 -fun term_features_of ctxt prover thy_name num_facts const_tab term_max_depth
  176.95 -                     type_max_depth ts =
  176.96 +fun term_features_of ctxt thy_name num_facts const_tab term_max_depth type_max_depth ts =
  176.97    let
  176.98      val thy = Proof_Context.theory_of ctxt
  176.99  
 176.100 -    fun is_built_in (x as (s, _)) args =
 176.101 -      if member (op =) logical_consts s then (true, args)
 176.102 -      else is_built_in_const_of_prover ctxt prover x args
 176.103 -
 176.104      val fixes = map snd (Variable.dest_fixes ctxt)
 176.105      val classes = Sign.classes_of thy
 176.106  
 176.107 @@ -660,11 +644,10 @@
 176.108           let val count = Symtab.lookup const_tab s |> the_default 1 in
 176.109             Real.fromInt num_facts / Real.fromInt count (* FUDGE *)
 176.110           end)
 176.111 -    fun pattify_term _ _ 0 _ = []
 176.112 -      | pattify_term _ args _ (Const (x as (s, _))) =
 176.113 -        if fst (is_built_in x args) then []
 176.114 -        else [(massage_long_name s, weight_of_const s)]
 176.115 -      | pattify_term _ _ _ (Free (s, T)) =
 176.116 +    fun pattify_term _ 0 _ = []
 176.117 +      | pattify_term _ _ (Const (s, _)) =
 176.118 +        if is_widely_irrelevant_const s then [] else [(massage_long_name s, weight_of_const s)]
 176.119 +      | pattify_term _ _ (Free (s, T)) =
 176.120          maybe_singleton_str pat_var_prefix (crude_str_of_typ T)
 176.121          |> map (rpair 0.0)
 176.122          |> (if member (op =) fixes s then
 176.123 @@ -672,36 +655,31 @@
 176.124                    (thy_name ^ Long_Name.separator ^ s)))
 176.125              else
 176.126                I)
 176.127 -      | pattify_term _ _ _ (Var (_, T)) =
 176.128 -        maybe_singleton_str pat_var_prefix (crude_str_of_typ T)
 176.129 -        |> map (rpair 0.0)
 176.130 -      | pattify_term Ts _ _ (Bound j) =
 176.131 -        maybe_singleton_str pat_var_prefix (crude_str_of_typ (nth Ts j))
 176.132 -        |> map (rpair 0.0)
 176.133 -      | pattify_term Ts args depth (t $ u) =
 176.134 +      | pattify_term _ _ (Var (_, T)) =
 176.135 +        maybe_singleton_str pat_var_prefix (crude_str_of_typ T) |> map (rpair 0.0)
 176.136 +      | pattify_term Ts _ (Bound j) =
 176.137 +        maybe_singleton_str pat_var_prefix (crude_str_of_typ (nth Ts j)) |> map (rpair 0.0)
 176.138 +      | pattify_term Ts depth (t $ u) =
 176.139          let
 176.140 -          val ps = take max_pat_breadth (pattify_term Ts (u :: args) depth t)
 176.141 -          val qs =
 176.142 -            take max_pat_breadth (("", 0.0) :: pattify_term Ts [] (depth - 1) u)
 176.143 +          val ps = take max_pat_breadth (pattify_term Ts depth t)
 176.144 +          val qs = take max_pat_breadth (("", 0.0) :: pattify_term Ts (depth - 1) u)
 176.145          in
 176.146            map_product (fn ppw as (p, pw) =>
 176.147                fn ("", _) => ppw
 176.148                 | (q, qw) => (p ^ "(" ^ q ^ ")", pw + qw)) ps qs
 176.149          end
 176.150 -      | pattify_term _ _ _ _ = []
 176.151 -    fun add_term_pat Ts = union (eq_fst (op =)) oo pattify_term Ts []
 176.152 +      | pattify_term _ _ _ = []
 176.153 +    fun add_term_pat Ts = union (eq_fst (op =)) oo pattify_term Ts
 176.154      fun add_term_pats _ 0 _ = I
 176.155        | add_term_pats Ts depth t =
 176.156          add_term_pat Ts depth t #> add_term_pats Ts (depth - 1) t
 176.157      fun add_term Ts = add_term_pats Ts term_max_depth
 176.158      fun add_subterms Ts t =
 176.159        case strip_comb t of
 176.160 -        (Const (x as (_, T)), args) =>
 176.161 -        let val (built_in, args) = is_built_in x args in
 176.162 -          (not built_in ? add_term Ts t)
 176.163 -          #> add_subtypes T
 176.164 -          #> fold (add_subterms Ts) args
 176.165 -        end
 176.166 +        (Const (s, T), args) =>
 176.167 +        (not (is_widely_irrelevant_const s) ? add_term Ts t)
 176.168 +        #> add_subtypes T
 176.169 +        #> fold (add_subterms Ts) args
 176.170        | (head, args) =>
 176.171          (case head of
 176.172             Free (_, T) => add_term Ts t #> add_subtypes T
 176.173 @@ -715,11 +693,10 @@
 176.174  val type_max_depth = 1
 176.175  
 176.176  (* TODO: Generate type classes for types? *)
 176.177 -fun features_of ctxt prover thy num_facts const_tab (scope, _) ts =
 176.178 +fun features_of ctxt thy num_facts const_tab (scope, _) ts =
 176.179    let val thy_name = Context.theory_name thy in
 176.180      thy_feature_of thy_name ::
 176.181 -    term_features_of ctxt prover thy_name num_facts const_tab term_max_depth
 176.182 -                     type_max_depth ts
 176.183 +    term_features_of ctxt thy_name num_facts const_tab term_max_depth type_max_depth ts
 176.184      |> scope <> Global ? cons local_feature
 176.185    end
 176.186  
 176.187 @@ -727,7 +704,7 @@
 176.188     isn't much to learn from such proofs. *)
 176.189  val max_dependencies = 20
 176.190  
 176.191 -val prover_default_max_facts = 50
 176.192 +val prover_default_max_facts = 25
 176.193  
 176.194  (* "type_definition_xxx" facts are characterized by their use of "CollectI". *)
 176.195  val typedef_dep = nickname_of_thm @{thm CollectI}
 176.196 @@ -785,6 +762,7 @@
 176.197      let
 176.198        val thy = Proof_Context.theory_of ctxt
 176.199        val goal = goal_of_thm thy th
 176.200 +      val name = nickname_of_thm th
 176.201        val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal goal 1 ctxt
 176.202        val facts = facts |> filter (fn (_, th') => thm_less (th', th))
 176.203        fun nickify ((_, stature), th) = ((nickname_of_thm th, stature), th)
 176.204 @@ -795,24 +773,23 @@
 176.205          else case find_first (is_dep dep) facts of
 176.206            SOME ((_, status), th) => accum @ [(("", status), th)]
 176.207          | NONE => accum (* shouldn't happen *)
 176.208 +      val mepo_facts =
 176.209 +        facts
 176.210 +        |> mepo_suggested_facts ctxt params (max_facts |> the_default prover_default_max_facts) NONE
 176.211 +             hyp_ts concl_t
 176.212        val facts =
 176.213 -        facts
 176.214 -        |> mepo_suggested_facts ctxt params prover
 176.215 -               (max_facts |> the_default prover_default_max_facts) NONE hyp_ts
 176.216 -               concl_t
 176.217 +        mepo_facts
 176.218          |> fold (add_isar_dep facts) isar_deps
 176.219          |> map nickify
 176.220 +      val num_isar_deps = length isar_deps
 176.221      in
 176.222        if verbose andalso auto_level = 0 then
 176.223 -        let val num_facts = length facts in
 176.224 -          "MaSh: " ^ quote prover ^ " on " ^ quote (nickname_of_thm th) ^
 176.225 -          " with " ^ string_of_int num_facts ^ " fact" ^ plural_s num_facts ^
 176.226 -          "."
 176.227 -          |> Output.urgent_message
 176.228 -        end
 176.229 +        "MaSh: " ^ quote prover ^ " on " ^ quote name ^ " with " ^ string_of_int num_isar_deps ^
 176.230 +        " + " ^ string_of_int (length facts - num_isar_deps) ^ " facts."
 176.231 +        |> Output.urgent_message
 176.232        else
 176.233          ();
 176.234 -      case run_prover_for_mash ctxt params prover facts goal of
 176.235 +      case run_prover_for_mash ctxt params prover name facts goal of
 176.236          {outcome = NONE, used_facts, ...} =>
 176.237          (if verbose andalso auto_level = 0 then
 176.238             let val num_facts = length used_facts in
 176.239 @@ -935,23 +912,22 @@
 176.240  
 176.241  val max_proximity_facts = 100
 176.242  
 176.243 -fun find_mash_suggestions _ _ [] _ _ raw_unknown = ([], raw_unknown)
 176.244 -  | find_mash_suggestions ctxt max_facts suggs facts chained raw_unknown =
 176.245 -    let
 176.246 -      val inter_fact = inter (eq_snd Thm.eq_thm_prop)
 176.247 -      val raw_mash = find_suggested_facts ctxt facts suggs
 176.248 -      val proximate = take max_proximity_facts facts
 176.249 -      val unknown_chained = inter_fact raw_unknown chained
 176.250 -      val unknown_proximate = inter_fact raw_unknown proximate
 176.251 -      val mess =
 176.252 -        [(0.9 (* FUDGE *), (map (rpair 1.0) unknown_chained, [])),
 176.253 -         (0.4 (* FUDGE *), (weight_facts_smoothly unknown_proximate, [])),
 176.254 -         (0.1 (* FUDGE *), (weight_facts_steeply raw_mash, raw_unknown))]
 176.255 -      val unknown =
 176.256 -        raw_unknown
 176.257 -        |> fold (subtract (eq_snd Thm.eq_thm_prop))
 176.258 -                [unknown_chained, unknown_proximate]
 176.259 -    in (mesh_facts (eq_snd Thm.eq_thm_prop) max_facts mess, unknown) end
 176.260 +fun find_mash_suggestions ctxt max_facts suggs facts chained raw_unknown =
 176.261 +  let
 176.262 +    val inter_fact = inter (eq_snd Thm.eq_thm_prop)
 176.263 +    val raw_mash = find_suggested_facts ctxt facts suggs
 176.264 +    val proximate = take max_proximity_facts facts
 176.265 +    val unknown_chained = inter_fact raw_unknown chained
 176.266 +    val unknown_proximate = inter_fact raw_unknown proximate
 176.267 +    val mess =
 176.268 +      [(0.9 (* FUDGE *), (map (rpair 1.0) unknown_chained, [])),
 176.269 +       (0.4 (* FUDGE *), (weight_facts_smoothly unknown_proximate, [])),
 176.270 +       (0.1 (* FUDGE *), (weight_facts_steeply raw_mash, raw_unknown))]
 176.271 +    val unknown =
 176.272 +      raw_unknown
 176.273 +      |> fold (subtract (eq_snd Thm.eq_thm_prop))
 176.274 +              [unknown_chained, unknown_proximate]
 176.275 +  in (mesh_facts (eq_snd Thm.eq_thm_prop) max_facts mess, unknown) end
 176.276  
 176.277  fun add_const_counts t =
 176.278    fold (fn s => Symtab.map_default (s, 0) (Integer.add 1))
 176.279 @@ -966,22 +942,23 @@
 176.280      val chained = facts |> filter (fn ((_, (scope, _)), _) => scope = Chained)
 176.281      val num_facts = length facts
 176.282      val const_tab = fold (add_const_counts o prop_of o snd) facts Symtab.empty
 176.283 +
 176.284      fun fact_has_right_theory (_, th) =
 176.285        thy_name = Context.theory_name (theory_of_thm th)
 176.286      fun chained_or_extra_features_of factor (((_, stature), th), weight) =
 176.287        [prop_of th]
 176.288 -      |> features_of ctxt prover (theory_of_thm th) num_facts const_tab stature
 176.289 +      |> features_of ctxt (theory_of_thm th) num_facts const_tab stature
 176.290        |> map (apsnd (fn r => weight * factor * r))
 176.291 +
 176.292      val (access_G, suggs) =
 176.293        peek_state ctxt overlord (fn {access_G, ...} =>
 176.294            if Graph.is_empty access_G then
 176.295 -            (access_G, [])
 176.296 +            (trace_msg ctxt (K "Nothing has been learned yet"); (access_G, []))
 176.297            else
 176.298              let
 176.299                val parents = maximal_wrt_access_graph access_G facts
 176.300                val goal_feats =
 176.301 -                features_of ctxt prover thy num_facts const_tab (Local, General)
 176.302 -                            (concl_t :: hyp_ts)
 176.303 +                features_of ctxt thy num_facts const_tab (Local, General) (concl_t :: hyp_ts)
 176.304                val chained_feats =
 176.305                  chained
 176.306                  |> map (rpair 1.0)
 176.307 @@ -1050,9 +1027,7 @@
 176.308          let
 176.309            val thy = Proof_Context.theory_of ctxt
 176.310            val name = freshish_name ()
 176.311 -          val feats =
 176.312 -            features_of ctxt prover thy 0 Symtab.empty (Local, General) [t]
 176.313 -            |> map fst
 176.314 +          val feats = features_of ctxt thy 0 Symtab.empty (Local, General) [t] |> map fst
 176.315          in
 176.316            peek_state ctxt overlord (fn {access_G, ...} =>
 176.317                let
 176.318 @@ -1151,9 +1126,7 @@
 176.319              let
 176.320                val name = nickname_of_thm th
 176.321                val feats =
 176.322 -                features_of ctxt prover (theory_of_thm th) 0 Symtab.empty
 176.323 -                            stature [prop_of th]
 176.324 -                |> map fst
 176.325 +                features_of ctxt (theory_of_thm th) 0 Symtab.empty stature [prop_of th] |> map fst
 176.326                val deps = deps_of status th |> these
 176.327                val n = n |> not (null deps) ? Integer.add 1
 176.328                val learns = (name, parents, feats, deps) :: learns
 176.329 @@ -1215,7 +1188,7 @@
 176.330                       Isar_Proof => 0
 176.331                     | Automatic_Proof => 2 * max_isar
 176.332                     | Isar_Proof_wegen_Prover_Flop => max_isar)
 176.333 -                - 500 * length (isar_dependencies_of name_tabs th)
 176.334 +                - 100 * length (isar_dependencies_of name_tabs th)
 176.335                val old_facts =
 176.336                  facts |> filter is_in_access_G
 176.337                        |> map (`priority_of)
 176.338 @@ -1228,7 +1201,7 @@
 176.339        in
 176.340          if verbose orelse auto_level < 2 then
 176.341            "Learned " ^ string_of_int n ^ " nontrivial " ^
 176.342 -          (if run_prover then "automatic" else "Isar") ^ " proof" ^ plural_s n ^
 176.343 +          (if run_prover then "automatic and " else "") ^ "Isar proof" ^ plural_s n ^
 176.344            (if verbose then " in " ^ string_of_time (Timer.checkRealTimer timer)
 176.345             else "") ^ "."
 176.346          else
 176.347 @@ -1236,15 +1209,14 @@
 176.348        end
 176.349    end
 176.350  
 176.351 -fun mash_learn ctxt (params as {provers, timeout, max_facts, ...}) fact_override
 176.352 -               chained run_prover =
 176.353 +fun mash_learn ctxt (params as {provers, timeout, ...}) fact_override chained run_prover =
 176.354    let
 176.355      val css = Sledgehammer_Fact.clasimpset_rule_table_of ctxt
 176.356      val ctxt = ctxt |> Config.put instantiate_inducts false
 176.357      val facts =
 176.358        nearly_all_facts ctxt false fact_override Symtab.empty css chained []
 176.359                         @{prop True}
 176.360 -      |> (case max_facts of NONE => I | SOME n => take n)
 176.361 +      |> sort (crude_thm_ord o pairself snd o swap)
 176.362      val num_facts = length facts
 176.363      val prover = hd provers
 176.364      fun learn auto_level run_prover =
 176.365 @@ -1345,42 +1317,36 @@
 176.366               if mash_can_suggest_facts ctxt overlord then meshN else mepoN)
 176.367            else
 176.368              (false, mepoN)
 176.369 +
 176.370 +      val unique_facts = drop_duplicate_facts facts
 176.371        val add_ths = Attrib.eval_thms ctxt add
 176.372 +
 176.373        fun in_add (_, th) = member Thm.eq_thm_prop add_ths th
 176.374        fun add_and_take accepts =
 176.375          (case add_ths of
 176.376             [] => accepts
 176.377 -         | _ => (facts |> filter in_add |> map fact_of_raw_fact) @
 176.378 +         | _ => (unique_facts |> filter in_add |> map fact_of_raw_fact) @
 176.379                  (accepts |> filter_out in_add))
 176.380          |> take max_facts
 176.381        fun mepo () =
 176.382 -        mepo_suggested_facts ctxt params prover max_facts NONE hyp_ts concl_t
 176.383 -                             facts
 176.384 -        |> weight_facts_steeply
 176.385 +        (mepo_suggested_facts ctxt params max_facts NONE hyp_ts concl_t unique_facts
 176.386 +         |> weight_facts_steeply, [])
 176.387        fun mash () =
 176.388 -        mash_suggested_facts ctxt params prover (generous_max_facts max_facts)
 176.389 -            hyp_ts concl_t facts
 176.390 +        mash_suggested_facts ctxt params prover (generous_max_facts max_facts) hyp_ts concl_t facts
 176.391          |>> weight_facts_steeply
 176.392        val mess =
 176.393          (* the order is important for the "case" expression below *)
 176.394 -        [] |> (if effective_fact_filter <> mepoN then
 176.395 -                 cons (mash_weight, (mash ()))
 176.396 -               else
 176.397 -                 I)
 176.398 -           |> (if effective_fact_filter <> mashN then
 176.399 -                 cons (mepo_weight, (mepo (), []))
 176.400 -               else
 176.401 -                 I)
 176.402 -      val mesh =
 176.403 -        mesh_facts (eq_snd Thm.eq_thm_prop) max_facts mess
 176.404 -        |> add_and_take
 176.405 +        [] |> effective_fact_filter <> mepoN ? cons (mash_weight, mash)
 176.406 +           |> effective_fact_filter <> mashN ? cons (mepo_weight, mepo)
 176.407 +           |> Par_List.map (apsnd (fn f => f ()))
 176.408 +      val mesh = mesh_facts (eq_snd Thm.eq_thm_prop) max_facts mess |> add_and_take
 176.409      in
 176.410        if save then MaSh.save ctxt overlord else ();
 176.411        case (fact_filter, mess) of
 176.412          (NONE, [(_, (mepo, _)), (_, (mash, _))]) =>
 176.413          [(meshN, mesh), (mepoN, mepo |> map fst |> add_and_take),
 176.414           (mashN, mash |> map fst |> add_and_take)]
 176.415 -      | _ => [("", mesh)]
 176.416 +      | _ => [(effective_fact_filter, mesh)]
 176.417      end
 176.418  
 176.419  fun kill_learners ctxt ({overlord, ...} : params) =
   177.1 --- a/src/HOL/Tools/Sledgehammer/sledgehammer_mepo.ML	Mon Nov 11 17:34:44 2013 +0100
   177.2 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_mepo.ML	Mon Nov 11 17:44:21 2013 +0100
   177.3 @@ -11,13 +11,33 @@
   177.4    type raw_fact = Sledgehammer_Fact.raw_fact
   177.5    type fact = Sledgehammer_Fact.fact
   177.6    type params = Sledgehammer_Provers.params
   177.7 -  type relevance_fudge = Sledgehammer_Provers.relevance_fudge
   177.8 +
   177.9 +  type relevance_fudge =
  177.10 +    {local_const_multiplier : real,
  177.11 +     worse_irrel_freq : real,
  177.12 +     higher_order_irrel_weight : real,
  177.13 +     abs_rel_weight : real,
  177.14 +     abs_irrel_weight : real,
  177.15 +     theory_const_rel_weight : real,
  177.16 +     theory_const_irrel_weight : real,
  177.17 +     chained_const_irrel_weight : real,
  177.18 +     intro_bonus : real,
  177.19 +     elim_bonus : real,
  177.20 +     simp_bonus : real,
  177.21 +     local_bonus : real,
  177.22 +     assum_bonus : real,
  177.23 +     chained_bonus : real,
  177.24 +     max_imperfect : real,
  177.25 +     max_imperfect_exp : real,
  177.26 +     threshold_divisor : real,
  177.27 +     ridiculous_threshold : real}
  177.28  
  177.29    val trace : bool Config.T
  177.30    val pseudo_abs_name : string
  177.31 +  val default_relevance_fudge : relevance_fudge
  177.32    val mepo_suggested_facts :
  177.33 -    Proof.context -> params -> string -> int -> relevance_fudge option
  177.34 -    -> term list -> term -> raw_fact list -> fact list
  177.35 +    Proof.context -> params -> int -> relevance_fudge option -> term list -> term ->
  177.36 +    raw_fact list -> fact list
  177.37  end;
  177.38  
  177.39  structure Sledgehammer_MePo : SLEDGEHAMMER_MEPO =
  177.40 @@ -36,6 +56,47 @@
  177.41  val pseudo_abs_name = sledgehammer_prefix ^ "abs"
  177.42  val theory_const_suffix = Long_Name.separator ^ " 1"
  177.43  
  177.44 +type relevance_fudge =
  177.45 +  {local_const_multiplier : real,
  177.46 +   worse_irrel_freq : real,
  177.47 +   higher_order_irrel_weight : real,
  177.48 +   abs_rel_weight : real,
  177.49 +   abs_irrel_weight : real,
  177.50 +   theory_const_rel_weight : real,
  177.51 +   theory_const_irrel_weight : real,
  177.52 +   chained_const_irrel_weight : real,
  177.53 +   intro_bonus : real,
  177.54 +   elim_bonus : real,
  177.55 +   simp_bonus : real,
  177.56 +   local_bonus : real,
  177.57 +   assum_bonus : real,
  177.58 +   chained_bonus : real,
  177.59 +   max_imperfect : real,
  177.60 +   max_imperfect_exp : real,
  177.61 +   threshold_divisor : real,
  177.62 +   ridiculous_threshold : real}
  177.63 +
  177.64 +(* FUDGE *)
  177.65 +val default_relevance_fudge =
  177.66 +  {local_const_multiplier = 1.5,
  177.67 +   worse_irrel_freq = 100.0,
  177.68 +   higher_order_irrel_weight = 1.05,
  177.69 +   abs_rel_weight = 0.5,
  177.70 +   abs_irrel_weight = 2.0,
  177.71 +   theory_const_rel_weight = 0.5,
  177.72 +   theory_const_irrel_weight = 0.25,
  177.73 +   chained_const_irrel_weight = 0.25,
  177.74 +   intro_bonus = 0.15,
  177.75 +   elim_bonus = 0.15,
  177.76 +   simp_bonus = 0.15,
  177.77 +   local_bonus = 0.55,
  177.78 +   assum_bonus = 1.05,
  177.79 +   chained_bonus = 1.5,
  177.80 +   max_imperfect = 11.5,
  177.81 +   max_imperfect_exp = 1.0,
  177.82 +   threshold_divisor = 2.0,
  177.83 +   ridiculous_threshold = 0.1}
  177.84 +
  177.85  fun order_of_type (Type (@{type_name fun}, [T1, T2])) =
  177.86      Int.max (order_of_type T1 + 1, order_of_type T2)
  177.87    | order_of_type (Type (_, Ts)) = fold (Integer.max o order_of_type) Ts 0
  177.88 @@ -43,24 +104,26 @@
  177.89  
  177.90  (* An abstraction of Isabelle types and first-order terms *)
  177.91  datatype pattern = PVar | PApp of string * pattern list
  177.92 -datatype ptype = PType of int * pattern list
  177.93 +datatype ptype = PType of int * typ list
  177.94  
  177.95 -fun string_of_pattern PVar = "_"
  177.96 -  | string_of_pattern (PApp (s, ps)) =
  177.97 -    if null ps then s else s ^ string_of_patterns ps
  177.98 -and string_of_patterns ps = "(" ^ commas (map string_of_pattern ps) ^ ")"
  177.99 -fun string_of_ptype (PType (_, ps)) = string_of_patterns ps
 177.100 +fun string_of_patternT (TVar _) = "_"
 177.101 +  | string_of_patternT (Type (s, ps)) =
 177.102 +    if null ps then s else s ^ string_of_patternsT ps
 177.103 +  | string_of_patternT (TFree (s, _)) = s
 177.104 +and string_of_patternsT ps = "(" ^ commas (map string_of_patternT ps) ^ ")"
 177.105 +fun string_of_ptype (PType (_, ps)) = string_of_patternsT ps
 177.106  
 177.107  (*Is the second type an instance of the first one?*)
 177.108 -fun match_pattern (PVar, _) = true
 177.109 -  | match_pattern (PApp _, PVar) = false
 177.110 -  | match_pattern (PApp (s, ps), PApp (t, qs)) =
 177.111 -    s = t andalso match_patterns (ps, qs)
 177.112 -and match_patterns (_, []) = true
 177.113 -  | match_patterns ([], _) = false
 177.114 -  | match_patterns (p :: ps, q :: qs) =
 177.115 -    match_pattern (p, q) andalso match_patterns (ps, qs)
 177.116 -fun match_ptype (PType (_, ps), PType (_, qs)) = match_patterns (ps, qs)
 177.117 +fun match_patternT (TVar _, _) = true
 177.118 +  | match_patternT (Type (s, ps), Type (t, qs)) =
 177.119 +    s = t andalso match_patternsT (ps, qs)
 177.120 +  | match_patternT (TFree (s, _), TFree (t, _)) = s = t
 177.121 +  | match_patternT (_, _) = false
 177.122 +and match_patternsT (_, []) = true
 177.123 +  | match_patternsT ([], _) = false
 177.124 +  | match_patternsT (p :: ps, q :: qs) =
 177.125 +    match_patternT (p, q) andalso match_patternsT (ps, qs)
 177.126 +fun match_ptype (PType (_, ps), PType (_, qs)) = match_patternsT (ps, qs)
 177.127  
 177.128  (* Is there a unifiable constant? *)
 177.129  fun pconst_mem f consts (s, ps) =
 177.130 @@ -69,14 +132,9 @@
 177.131  fun pconst_hyper_mem f const_tab (s, ps) =
 177.132    exists (curry (match_ptype o f) ps) (these (Symtab.lookup const_tab s))
 177.133  
 177.134 -fun pattern_of_type (Type (s, Ts)) = PApp (s, map pattern_of_type Ts)
 177.135 -  | pattern_of_type (TFree (s, _)) = PApp (s, [])
 177.136 -  | pattern_of_type (TVar _) = PVar
 177.137 -
 177.138  (* Pairs a constant with the list of its type instantiations. *)
 177.139  fun ptype thy const x =
 177.140 -  (if const then map pattern_of_type (these (try (Sign.const_typargs thy) x))
 177.141 -   else [])
 177.142 +  (if const then these (try (Sign.const_typargs thy) x) else [])
 177.143  fun rich_ptype thy const (s, T) =
 177.144    PType (order_of_type T, ptype thy const (s, T))
 177.145  fun rich_pconst thy const (s, T) = (s, rich_ptype thy const (s, T))
 177.146 @@ -84,9 +142,19 @@
 177.147  fun string_of_hyper_pconst (s, ps) =
 177.148    s ^ "{" ^ commas (map string_of_ptype ps) ^ "}"
 177.149  
 177.150 -(* Add a pconstant to the table, but a [] entry means a standard
 177.151 -   connective, which we ignore.*)
 177.152 -fun add_pconst_to_table (s, p) = Symtab.map_default (s, [p]) (insert (op =) p)
 177.153 +fun patternT_eq (TVar _, TVar _) = true
 177.154 +  | patternT_eq (Type (s, Ts), Type (t, Us)) = s = t andalso patternsT_eq (Ts, Us)
 177.155 +  | patternT_eq (TFree (s, _), TFree (t, _)) = (s = t)
 177.156 +  | patternT_eq _ = false
 177.157 +and patternsT_eq ([], []) = true
 177.158 +  | patternsT_eq ([], _) = false
 177.159 +  | patternsT_eq (_, []) = false
 177.160 +  | patternsT_eq (T :: Ts, U :: Us) = patternT_eq (T, U) andalso patternsT_eq (Ts, Us)
 177.161 +fun ptype_eq (PType (m, Ts), PType (n, Us)) = m = n andalso patternsT_eq (Ts, Us)
 177.162 +
 177.163 + (* Add a pconstant to the table, but a [] entry means a standard connective,
 177.164 +    which we ignore. *)
 177.165 +fun add_pconst_to_table (s, p) = Symtab.map_default (s, [p]) (insert ptype_eq p)
 177.166  
 177.167  (* Set constants tend to pull in too many irrelevant facts. We limit the damage
 177.168     by treating them more or less as if they were built-in but add their
 177.169 @@ -94,21 +162,18 @@
 177.170  val set_consts = [@{const_name Collect}, @{const_name Set.member}]
 177.171  val set_thms = @{thms Collect_mem_eq mem_Collect_eq Collect_cong}
 177.172  
 177.173 -fun add_pconsts_in_term thy is_built_in_const =
 177.174 +fun add_pconsts_in_term thy =
 177.175    let
 177.176 -    fun do_const const ext_arg (x as (s, _)) ts =
 177.177 -      let val (built_in, ts) = is_built_in_const x ts in
 177.178 -        if member (op =) set_consts s then
 177.179 -          fold (do_term ext_arg) ts
 177.180 -        else
 177.181 -          (not built_in
 177.182 -           ? add_pconst_to_table (rich_pconst thy const x))
 177.183 -          #> fold (do_term false) ts
 177.184 -      end
 177.185 +    fun do_const const (x as (s, _)) ts =
 177.186 +      if member (op =) set_consts s then
 177.187 +        fold (do_term false) ts
 177.188 +      else
 177.189 +        (not (is_irrelevant_const s) ? add_pconst_to_table (rich_pconst thy const x))
 177.190 +        #> fold (do_term false) ts
 177.191      and do_term ext_arg t =
 177.192        case strip_comb t of
 177.193 -        (Const x, ts) => do_const true ext_arg x ts
 177.194 -      | (Free x, ts) => do_const false ext_arg x ts
 177.195 +        (Const x, ts) => do_const true x ts
 177.196 +      | (Free x, ts) => do_const false x ts
 177.197        | (Abs (_, T, t'), ts) =>
 177.198          ((null ts andalso not ext_arg)
 177.199           (* Since lambdas on the right-hand side of equalities are usually
 177.200 @@ -122,7 +187,7 @@
 177.201        if T = HOLogic.boolT then do_formula else do_term ext_arg
 177.202      and do_formula t =
 177.203        case t of
 177.204 -        Const (@{const_name all}, _) $ Abs (_, T, t') => do_formula t'
 177.205 +        Const (@{const_name all}, _) $ Abs (_, _, t') => do_formula t'
 177.206        | @{const "==>"} $ t1 $ t2 => do_formula t1 #> do_formula t2
 177.207        | Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2 =>
 177.208          do_term_or_formula false T t1 #> do_term_or_formula true T t2
 177.209 @@ -130,8 +195,8 @@
 177.210        | @{const False} => I
 177.211        | @{const True} => I
 177.212        | @{const Not} $ t1 => do_formula t1
 177.213 -      | Const (@{const_name All}, _) $ Abs (_, T, t') => do_formula t'
 177.214 -      | Const (@{const_name Ex}, _) $ Abs (_, T, t') => do_formula t'
 177.215 +      | Const (@{const_name All}, _) $ Abs (_, _, t') => do_formula t'
 177.216 +      | Const (@{const_name Ex}, _) $ Abs (_, _, t') => do_formula t'
 177.217        | @{const HOL.conj} $ t1 $ t2 => do_formula t1 #> do_formula t2
 177.218        | @{const HOL.disj} $ t1 $ t2 => do_formula t1 #> do_formula t2
 177.219        | @{const HOL.implies} $ t1 $ t2 => do_formula t1 #> do_formula t2
 177.220 @@ -140,19 +205,19 @@
 177.221        | Const (@{const_name If}, Type (_, [_, Type (_, [T, _])]))
 177.222          $ t1 $ t2 $ t3 =>
 177.223          do_formula t1 #> fold (do_term_or_formula false T) [t2, t3]
 177.224 -      | Const (@{const_name Ex1}, _) $ Abs (_, T, t') => do_formula t'
 177.225 -      | Const (@{const_name Ball}, _) $ t1 $ Abs (_, T, t') =>
 177.226 +      | Const (@{const_name Ex1}, _) $ Abs (_, _, t') => do_formula t'
 177.227 +      | Const (@{const_name Ball}, _) $ t1 $ Abs (_, _, t') =>
 177.228          do_formula (t1 $ Bound ~1) #> do_formula t'
 177.229 -      | Const (@{const_name Bex}, _) $ t1 $ Abs (_, T, t') =>
 177.230 +      | Const (@{const_name Bex}, _) $ t1 $ Abs (_, _, t') =>
 177.231          do_formula (t1 $ Bound ~1) #> do_formula t'
 177.232        | (t0 as Const (_, @{typ bool})) $ t1 =>
 177.233          do_term false t0 #> do_formula t1  (* theory constant *)
 177.234        | _ => do_term false t
 177.235    in do_formula end
 177.236  
 177.237 -fun pconsts_in_fact thy is_built_in_const t =
 177.238 +fun pconsts_in_fact thy t =
 177.239    Symtab.fold (fn (s, pss) => fold (cons o pair s) pss)
 177.240 -              (Symtab.empty |> add_pconsts_in_term thy is_built_in_const t) []
 177.241 +              (Symtab.empty |> add_pconsts_in_term thy t) []
 177.242  
 177.243  (* Inserts a dummy "constant" referring to the theory name, so that relevance
 177.244     takes the given theory into account. *)
 177.245 @@ -167,9 +232,9 @@
 177.246  fun theory_const_prop_of fudge th =
 177.247    theory_constify fudge (Context.theory_name (theory_of_thm th)) (prop_of th)
 177.248  
 177.249 -fun pair_consts_fact thy is_built_in_const fudge fact =
 177.250 +fun pair_consts_fact thy fudge fact =
 177.251    case fact |> snd |> theory_const_prop_of fudge
 177.252 -            |> pconsts_in_fact thy is_built_in_const of
 177.253 +            |> pconsts_in_fact thy of
 177.254      [] => NONE
 177.255    | consts => SOME ((fact, consts), NONE)
 177.256  
 177.257 @@ -177,15 +242,22 @@
 177.258     first by constant name and second by its list of type instantiations. For the
 177.259     latter, we need a linear ordering on "pattern list". *)
 177.260  
 177.261 -fun pattern_ord p =
 177.262 +fun patternT_ord p =
 177.263    case p of
 177.264 -    (PVar, PVar) => EQUAL
 177.265 -  | (PVar, PApp _) => LESS
 177.266 -  | (PApp _, PVar) => GREATER
 177.267 -  | (PApp q1, PApp q2) =>
 177.268 -    prod_ord fast_string_ord (dict_ord pattern_ord) (q1, q2)
 177.269 -fun ptype_ord (PType p, PType q) =
 177.270 -  prod_ord (dict_ord pattern_ord) int_ord (swap p, swap q)
 177.271 +    (Type (s, ps), Type (t, qs)) =>
 177.272 +    (case fast_string_ord (s, t) of
 177.273 +      EQUAL => dict_ord patternT_ord (ps, qs)
 177.274 +    | ord => ord)
 177.275 +  | (TVar _, TVar _) => EQUAL
 177.276 +  | (TVar _, _) => LESS
 177.277 +  | (Type _, TVar _) => GREATER
 177.278 +  | (Type _, TFree _) => LESS
 177.279 +  | (TFree (s, _), TFree (t, _)) => fast_string_ord (s, t)
 177.280 +  | (TFree _, _) => GREATER
 177.281 +fun ptype_ord (PType (m, ps), PType (n, qs)) =
 177.282 +  (case dict_ord patternT_ord (ps, qs) of
 177.283 +    EQUAL => int_ord (m, n)
 177.284 +  | ord => ord)
 177.285  
 177.286  structure PType_Tab = Table(type key = ptype val ord = ptype_ord)
 177.287  
 177.288 @@ -326,23 +398,23 @@
 177.289      (accepts, more_rejects @ rejects)
 177.290    end
 177.291  
 177.292 -fun if_empty_replace_with_scope thy is_built_in_const facts sc tab =
 177.293 +fun if_empty_replace_with_scope thy facts sc tab =
 177.294    if Symtab.is_empty tab then
 177.295      Symtab.empty
 177.296 -    |> fold (add_pconsts_in_term thy is_built_in_const)
 177.297 +    |> fold (add_pconsts_in_term thy)
 177.298              (map_filter (fn ((_, (sc', _)), th) =>
 177.299                              if sc' = sc then SOME (prop_of th) else NONE) facts)
 177.300    else
 177.301      tab
 177.302  
 177.303 -fun consider_arities is_built_in_const th =
 177.304 +fun consider_arities th =
 177.305    let
 177.306      fun aux _ _ NONE = NONE
 177.307        | aux t args (SOME tab) =
 177.308          case t of
 177.309            t1 $ t2 => SOME tab |> aux t1 (t2 :: args) |> aux t2 []
 177.310 -        | Const (x as (s, _)) =>
 177.311 -          (if is_built_in_const x args |> fst then
 177.312 +        | Const (s, _) =>
 177.313 +          (if is_widely_irrelevant_const s then
 177.314               SOME tab
 177.315             else case Symtab.lookup tab s of
 177.316               NONE => SOME (Symtab.update (s, length args) tab)
 177.317 @@ -351,24 +423,24 @@
 177.318    in aux (prop_of th) [] end
 177.319  
 177.320  (* FIXME: This is currently only useful for polymorphic type encodings. *)
 177.321 -fun could_benefit_from_ext is_built_in_const facts =
 177.322 -  fold (consider_arities is_built_in_const o snd) facts (SOME Symtab.empty)
 177.323 -  |> is_none
 177.324 +fun could_benefit_from_ext facts =
 177.325 +  fold (consider_arities o snd) facts (SOME Symtab.empty) |> is_none
 177.326  
 177.327  (* High enough so that it isn't wrongly considered as very relevant (e.g., for E
 177.328     weights), but low enough so that it is unlikely to be truncated away if few
 177.329     facts are included. *)
 177.330  val special_fact_index = 45 (* FUDGE *)
 177.331  
 177.332 +fun eq_prod eqx eqy ((x1, y1), (x2, y2)) = eqx (x1, x2) andalso eqy (y1, y2)
 177.333 +
 177.334  val really_hopeless_get_kicked_out_iter = 5 (* FUDGE *)
 177.335  
 177.336 -fun relevance_filter ctxt thres0 decay max_facts is_built_in_const
 177.337 -        (fudge as {threshold_divisor, ridiculous_threshold, ...}) facts hyp_ts
 177.338 -        concl_t =
 177.339 +fun relevance_filter ctxt thres0 decay max_facts
 177.340 +        (fudge as {threshold_divisor, ridiculous_threshold, ...}) facts hyp_ts concl_t =
 177.341    let
 177.342      val thy = Proof_Context.theory_of ctxt
 177.343      val const_tab = fold (count_fact_consts thy fudge) facts Symtab.empty
 177.344 -    val add_pconsts = add_pconsts_in_term thy is_built_in_const
 177.345 +    val add_pconsts = add_pconsts_in_term thy
 177.346      val chained_ts =
 177.347        facts |> map_filter (fn ((_, (Chained, _)), th) => SOME (prop_of th)
 177.348                              | _ => NONE)
 177.349 @@ -378,8 +450,7 @@
 177.350        |> fold add_pconsts hyp_ts
 177.351        |> add_pconsts concl_t
 177.352        |> (fn tab => if Symtab.is_empty tab then chained_const_tab else tab)
 177.353 -      |> fold (if_empty_replace_with_scope thy is_built_in_const facts)
 177.354 -              [Chained, Assum, Local]
 177.355 +      |> fold (if_empty_replace_with_scope thy facts) [Chained, Assum, Local]
 177.356      fun iter j remaining_max thres rel_const_tab hopeless hopeful =
 177.357        let
 177.358          val hopeless =
 177.359 @@ -421,7 +492,8 @@
 177.360                trace_msg ctxt (fn () => "New or updated constants: " ^
 177.361                    commas (rel_const_tab'
 177.362                            |> Symtab.dest
 177.363 -                          |> subtract (op =) (rel_const_tab |> Symtab.dest)
 177.364 +                          |> subtract (eq_prod (op =) (eq_list ptype_eq))
 177.365 +                                      (rel_const_tab |> Symtab.dest)
 177.366                            |> map string_of_hyper_pconst));
 177.367                map (fst o fst) accepts @
 177.368                (if remaining_max = 0 then
 177.369 @@ -475,41 +547,33 @@
 177.370          in bef @ add @ after end
 177.371      fun insert_special_facts accepts =
 177.372        (* FIXME: get rid of "ext" here once it is treated as a helper *)
 177.373 -      [] |> could_benefit_from_ext is_built_in_const accepts ? cons @{thm ext}
 177.374 +      [] |> could_benefit_from_ext accepts ? cons @{thm ext}
 177.375           |> add_set_const_thms accepts
 177.376           |> insert_into_facts accepts
 177.377    in
 177.378 -    facts |> map_filter (pair_consts_fact thy is_built_in_const fudge)
 177.379 +    facts |> map_filter (pair_consts_fact thy fudge)
 177.380            |> iter 0 max_facts thres0 goal_const_tab []
 177.381            |> insert_special_facts
 177.382            |> tap (fn accepts => trace_msg ctxt (fn () =>
 177.383                        "Total relevant: " ^ string_of_int (length accepts)))
 177.384    end
 177.385  
 177.386 -fun mepo_suggested_facts ctxt
 177.387 -        ({fact_thresholds = (thres0, thres1), ...} : params) prover
 177.388 -        max_facts fudge hyp_ts concl_t facts =
 177.389 +fun mepo_suggested_facts ctxt ({fact_thresholds = (thres0, thres1), ...} : params) max_facts fudge
 177.390 +      hyp_ts concl_t facts =
 177.391    let
 177.392      val thy = Proof_Context.theory_of ctxt
 177.393 -    val is_built_in_const =
 177.394 -      Sledgehammer_Provers.is_built_in_const_of_prover ctxt prover
 177.395 -    val fudge =
 177.396 -      case fudge of
 177.397 -        SOME fudge => fudge
 177.398 -      | NONE => Sledgehammer_Provers.relevance_fudge_of_prover ctxt prover
 177.399 +    val fudge = fudge |> the_default default_relevance_fudge
 177.400      val decay = Math.pow ((1.0 - thres1) / (1.0 - thres0),
 177.401                            1.0 / Real.fromInt (max_facts + 1))
 177.402    in
 177.403 -    trace_msg ctxt (fn () => "Considering " ^ string_of_int (length facts) ^
 177.404 -                             " facts");
 177.405 +    trace_msg ctxt (fn () => "Considering " ^ string_of_int (length facts) ^ " facts");
 177.406      (if thres1 < 0.0 then
 177.407         facts
 177.408 -     else if thres0 > 1.0 orelse thres0 > thres1 then
 177.409 +     else if thres0 > 1.0 orelse thres0 > thres1 orelse max_facts <= 0 then
 177.410         []
 177.411       else
 177.412 -       relevance_filter ctxt thres0 decay max_facts is_built_in_const fudge
 177.413 -           facts hyp_ts
 177.414 -           (concl_t |> theory_constify fudge (Context.theory_name thy)))
 177.415 +       relevance_filter ctxt thres0 decay max_facts fudge facts hyp_ts
 177.416 +         (concl_t |> theory_constify fudge (Context.theory_name thy)))
 177.417      |> map fact_of_raw_fact
 177.418    end
 177.419  
   178.1 --- a/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML	Mon Nov 11 17:34:44 2013 +0100
   178.2 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML	Mon Nov 11 17:44:21 2013 +0100
   178.3 @@ -18,7 +18,7 @@
   178.4    val auto_minimize_max_time : real Config.T
   178.5    val minimize_facts :
   178.6      (string -> thm list -> unit) -> string -> params -> bool -> int -> int
   178.7 -    -> Proof.state -> play Lazy.lazy option
   178.8 +    -> Proof.state -> thm -> play Lazy.lazy option
   178.9      -> ((string * stature) * thm list) list
  178.10      -> ((string * stature) * thm list) list option
  178.11         * (play Lazy.lazy * (play -> string) * string)
  178.12 @@ -59,7 +59,7 @@
  178.13                   max_new_mono_instances, type_enc, strict, lam_trans,
  178.14                   uncurried_aliases, isar_proofs, isar_compress, isar_try0,
  178.15                   preplay_timeout, ...} : params)
  178.16 -               silent (prover : prover) timeout i n state facts =
  178.17 +               silent (prover : prover) timeout i n state goal facts =
  178.18    let
  178.19      val _ =
  178.20        print silent (fn () =>
  178.21 @@ -70,7 +70,6 @@
  178.22               | _ => ""
  178.23             else
  178.24               "") ^ "...")
  178.25 -    val {goal, ...} = Proof.goal state
  178.26      val facts = facts |> maps (fn (n, ths) => map (pair n) ths)
  178.27      val params =
  178.28        {debug = debug, verbose = verbose, overlord = overlord, spy = spy, blocking = true,
  178.29 @@ -82,7 +81,8 @@
  178.30         slice = false, minimize = SOME false, timeout = timeout, preplay_timeout = preplay_timeout,
  178.31         expect = ""}
  178.32      val problem =
  178.33 -      {state = state, goal = goal, subgoal = i, subgoal_count = n, factss = [("", facts)]}
  178.34 +      {comment = "", state = state, goal = goal, subgoal = i, subgoal_count = n,
  178.35 +       factss = [("", facts)]}
  178.36      val result as {outcome, used_facts, run_time, ...} =
  178.37        prover params (K (K (K ""))) problem
  178.38    in
  178.39 @@ -191,12 +191,12 @@
  178.40    end
  178.41  
  178.42  fun minimize_facts do_learn prover_name (params as {learn, timeout, ...}) silent
  178.43 -                   i n state preplay0 facts =
  178.44 +                   i n state goal preplay0 facts =
  178.45    let
  178.46      val ctxt = Proof.context_of state
  178.47      val prover =
  178.48        get_prover ctxt (if silent then Auto_Minimize else Minimize) prover_name
  178.49 -    fun test timeout = test_facts params silent prover timeout i n state
  178.50 +    fun test timeout = test_facts params silent prover timeout i n state goal
  178.51      val (chained, non_chained) = List.partition is_fact_chained facts
  178.52      (* Push chained facts to the back, so that they are less likely to be
  178.53         kicked out by the linear minimization algorithm. *)
  178.54 @@ -272,7 +272,7 @@
  178.55  
  178.56  fun maybe_minimize ctxt mode do_learn name
  178.57          (params as {verbose, isar_proofs, minimize, ...})
  178.58 -        ({state, subgoal, subgoal_count, ...} : prover_problem)
  178.59 +        ({state, goal, subgoal, subgoal_count, ...} : prover_problem)
  178.60          (result as {outcome, used_facts, used_from, run_time, preplay, message,
  178.61                      message_tail} : prover_result) =
  178.62    if is_some outcome orelse null used_facts then
  178.63 @@ -319,8 +319,8 @@
  178.64        val (used_facts, (preplay, message, _)) =
  178.65          if minimize then
  178.66            minimize_facts do_learn minimize_name params
  178.67 -              (mode <> Normal orelse not verbose) subgoal subgoal_count state
  178.68 -              (SOME preplay)
  178.69 +              (not verbose orelse (mode <> Normal andalso mode <> MaSh)) subgoal
  178.70 +              subgoal_count state goal (SOME preplay)
  178.71                (filter_used_facts true used_facts (map (apsnd single) used_from))
  178.72            |>> Option.map (map fst)
  178.73          else
  178.74 @@ -342,12 +342,10 @@
  178.75  fun run_minimize (params as {provers, ...}) do_learn i refs state =
  178.76    let
  178.77      val ctxt = Proof.context_of state
  178.78 +    val {goal, facts = chained_ths, ...} = Proof.goal state
  178.79      val reserved = reserved_isar_keyword_table ()
  178.80 -    val chained_ths = #facts (Proof.goal state)
  178.81      val css = clasimpset_rule_table_of ctxt
  178.82 -    val facts =
  178.83 -      refs |> maps (map (apsnd single)
  178.84 -                    o fact_of_ref ctxt reserved chained_ths css)
  178.85 +    val facts = refs |> maps (map (apsnd single) o fact_of_ref ctxt reserved chained_ths css)
  178.86    in
  178.87      case subgoal_count state of
  178.88        0 => Output.urgent_message "No subgoal!"
  178.89 @@ -355,7 +353,7 @@
  178.90               [] => error "No prover is set."
  178.91             | prover :: _ =>
  178.92               (kill_provers ();
  178.93 -              minimize_facts do_learn prover params false i n state NONE facts
  178.94 +              minimize_facts do_learn prover params false i n state goal NONE facts
  178.95                |> (fn (_, (preplay, message, message_tail)) =>
  178.96                       message (Lazy.force preplay) ^ message_tail
  178.97                       |> Output.urgent_message))
   179.1 --- a/src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML	Mon Nov 11 17:34:44 2013 +0100
   179.2 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML	Mon Nov 11 17:44:21 2013 +0100
   179.3 @@ -44,28 +44,9 @@
   179.4       preplay_timeout : Time.time option,
   179.5       expect : string}
   179.6  
   179.7 -  type relevance_fudge =
   179.8 -    {local_const_multiplier : real,
   179.9 -     worse_irrel_freq : real,
  179.10 -     higher_order_irrel_weight : real,
  179.11 -     abs_rel_weight : real,
  179.12 -     abs_irrel_weight : real,
  179.13 -     theory_const_rel_weight : real,
  179.14 -     theory_const_irrel_weight : real,
  179.15 -     chained_const_irrel_weight : real,
  179.16 -     intro_bonus : real,
  179.17 -     elim_bonus : real,
  179.18 -     simp_bonus : real,
  179.19 -     local_bonus : real,
  179.20 -     assum_bonus : real,
  179.21 -     chained_bonus : real,
  179.22 -     max_imperfect : real,
  179.23 -     max_imperfect_exp : real,
  179.24 -     threshold_divisor : real,
  179.25 -     ridiculous_threshold : real}
  179.26 -
  179.27    type prover_problem =
  179.28 -    {state : Proof.state,
  179.29 +    {comment : string,
  179.30 +     state : Proof.state,
  179.31       goal : thm,
  179.32       subgoal : int,
  179.33       subgoal_count : int,
  179.34 @@ -88,8 +69,7 @@
  179.35    val problem_prefix : string Config.T
  179.36    val completish : bool Config.T
  179.37    val atp_full_names : bool Config.T
  179.38 -  val smt_builtin_facts : bool Config.T
  179.39 -  val smt_builtin_trans : bool Config.T
  179.40 +  val smt_builtins : bool Config.T
  179.41    val smt_triggers : bool Config.T
  179.42    val smt_weights : bool Config.T
  179.43    val smt_weight_min_facts : int Config.T
  179.44 @@ -117,14 +97,9 @@
  179.45      Proof.context -> string -> string option
  179.46    val remotify_prover_if_not_installed :
  179.47      Proof.context -> string -> string option
  179.48 -  val default_max_facts_of_prover : Proof.context -> bool -> string -> int
  179.49 +  val default_max_facts_of_prover : Proof.context -> string -> int
  179.50    val is_unit_equality : term -> bool
  179.51    val is_appropriate_prop_of_prover : Proof.context -> string -> term -> bool
  179.52 -  val is_built_in_const_of_prover :
  179.53 -    Proof.context -> string -> string * typ -> term list -> bool * term list
  179.54 -  val atp_relevance_fudge : relevance_fudge
  179.55 -  val smt_relevance_fudge : relevance_fudge
  179.56 -  val relevance_fudge_of_prover : Proof.context -> string -> relevance_fudge
  179.57    val weight_smt_fact :
  179.58      Proof.context -> int -> ((string * stature) * thm) * int
  179.59      -> (string * stature) * (int option * thm)
  179.60 @@ -160,27 +135,18 @@
  179.61  (** The Sledgehammer **)
  179.62  
  179.63  (* Empty string means create files in Isabelle's temporary files directory. *)
  179.64 -val dest_dir =
  179.65 -  Attrib.setup_config_string @{binding sledgehammer_dest_dir} (K "")
  179.66 -val problem_prefix =
  179.67 -  Attrib.setup_config_string @{binding sledgehammer_problem_prefix} (K "prob")
  179.68 -val completish =
  179.69 -  Attrib.setup_config_bool @{binding sledgehammer_completish} (K false)
  179.70 +val dest_dir = Attrib.setup_config_string @{binding sledgehammer_dest_dir} (K "")
  179.71 +val problem_prefix = Attrib.setup_config_string @{binding sledgehammer_problem_prefix} (K "prob")
  179.72 +val completish = Attrib.setup_config_bool @{binding sledgehammer_completish} (K false)
  179.73  
  179.74  (* In addition to being easier to read, readable names are often much shorter,
  179.75     especially if types are mangled in names. This makes a difference for some
  179.76     provers (e.g., E). For these reason, short names are enabled by default. *)
  179.77 -val atp_full_names =
  179.78 -  Attrib.setup_config_bool @{binding sledgehammer_atp_full_names} (K false)
  179.79 +val atp_full_names = Attrib.setup_config_bool @{binding sledgehammer_atp_full_names} (K false)
  179.80  
  179.81 -val smt_builtin_facts =
  179.82 -  Attrib.setup_config_bool @{binding sledgehammer_smt_builtin_facts} (K true)
  179.83 -val smt_builtin_trans =
  179.84 -  Attrib.setup_config_bool @{binding sledgehammer_smt_builtin_trans} (K true)
  179.85 -val smt_triggers =
  179.86 -  Attrib.setup_config_bool @{binding sledgehammer_smt_triggers} (K true)
  179.87 -val smt_weights =
  179.88 -  Attrib.setup_config_bool @{binding sledgehammer_smt_weights} (K true)
  179.89 +val smt_builtins = Attrib.setup_config_bool @{binding sledgehammer_smt_builtins} (K true)
  179.90 +val smt_triggers = Attrib.setup_config_bool @{binding sledgehammer_smt_triggers} (K true)
  179.91 +val smt_weights = Attrib.setup_config_bool @{binding sledgehammer_smt_weights} (K true)
  179.92  val smt_weight_min_facts =
  179.93    Attrib.setup_config_int @{binding sledgehammer_smt_weight_min_facts} (K 20)
  179.94  
  179.95 @@ -240,15 +206,14 @@
  179.96  
  179.97  val reconstructor_default_max_facts = 20
  179.98  
  179.99 -fun slice_max_facts (_, (_, ( ((max_facts, _), _, _, _, _), _))) = max_facts
 179.100 +fun slice_max_facts (_, ( ((max_facts, _), _, _, _, _), _)) = max_facts
 179.101  
 179.102 -fun default_max_facts_of_prover ctxt slice name =
 179.103 +fun default_max_facts_of_prover ctxt name =
 179.104    let val thy = Proof_Context.theory_of ctxt in
 179.105      if is_reconstructor name then
 179.106        reconstructor_default_max_facts
 179.107      else if is_atp thy name then
 179.108 -      fold (Integer.max o slice_max_facts)
 179.109 -           (get_slices slice (#best_slices (get_atp thy name ()) ctxt)) 0
 179.110 +      fold (Integer.max o slice_max_facts) (#best_slices (get_atp thy name ()) ctxt) 0
 179.111      else (* is_smt_prover ctxt name *)
 179.112        SMT_Solver.default_max_relevant ctxt name
 179.113    end
 179.114 @@ -275,68 +240,6 @@
 179.115  fun is_appropriate_prop_of_prover ctxt name =
 179.116    if is_unit_equational_atp ctxt name then is_unit_equality else K true
 179.117  
 179.118 -val atp_irrelevant_const_tab =
 179.119 -  Symtab.make (map (rpair ()) atp_irrelevant_consts)
 179.120 -
 179.121 -fun is_built_in_const_of_prover ctxt name =
 179.122 -  if is_smt_prover ctxt name andalso Config.get ctxt smt_builtin_facts then
 179.123 -    let val ctxt = ctxt |> select_smt_solver name in
 179.124 -      fn x => fn ts =>
 179.125 -         if SMT_Builtin.is_builtin_num_ext ctxt (list_comb (Const x, ts)) then
 179.126 -           (true, [])
 179.127 -         else if SMT_Builtin.is_builtin_fun_ext ctxt x ts then
 179.128 -           (true, ts)
 179.129 -         else
 179.130 -           (false, ts)
 179.131 -    end
 179.132 -  else
 179.133 -    fn (s, _) => fn ts => (Symtab.defined atp_irrelevant_const_tab s, ts)
 179.134 -
 179.135 -(* FUDGE *)
 179.136 -val atp_relevance_fudge =
 179.137 -  {local_const_multiplier = 1.5,
 179.138 -   worse_irrel_freq = 100.0,
 179.139 -   higher_order_irrel_weight = 1.05,
 179.140 -   abs_rel_weight = 0.5,
 179.141 -   abs_irrel_weight = 2.0,
 179.142 -   theory_const_rel_weight = 0.5,
 179.143 -   theory_const_irrel_weight = 0.25,
 179.144 -   chained_const_irrel_weight = 0.25,
 179.145 -   intro_bonus = 0.15,
 179.146 -   elim_bonus = 0.15,
 179.147 -   simp_bonus = 0.15,
 179.148 -   local_bonus = 0.55,
 179.149 -   assum_bonus = 1.05,
 179.150 -   chained_bonus = 1.5,
 179.151 -   max_imperfect = 11.5,
 179.152 -   max_imperfect_exp = 1.0,
 179.153 -   threshold_divisor = 2.0,
 179.154 -   ridiculous_threshold = 0.01}
 179.155 -
 179.156 -(* FUDGE (FIXME) *)
 179.157 -val smt_relevance_fudge =
 179.158 -  {local_const_multiplier = #local_const_multiplier atp_relevance_fudge,
 179.159 -   worse_irrel_freq = #worse_irrel_freq atp_relevance_fudge,
 179.160 -   higher_order_irrel_weight = #higher_order_irrel_weight atp_relevance_fudge,
 179.161 -   abs_rel_weight = #abs_rel_weight atp_relevance_fudge,
 179.162 -   abs_irrel_weight = #abs_irrel_weight atp_relevance_fudge,
 179.163 -   theory_const_rel_weight = #theory_const_rel_weight atp_relevance_fudge,
 179.164 -   theory_const_irrel_weight = #theory_const_irrel_weight atp_relevance_fudge,
 179.165 -   chained_const_irrel_weight = #chained_const_irrel_weight atp_relevance_fudge,
 179.166 -   intro_bonus = #intro_bonus atp_relevance_fudge,
 179.167 -   elim_bonus = #elim_bonus atp_relevance_fudge,
 179.168 -   simp_bonus = #simp_bonus atp_relevance_fudge,
 179.169 -   local_bonus = #local_bonus atp_relevance_fudge,
 179.170 -   assum_bonus = #assum_bonus atp_relevance_fudge,
 179.171 -   chained_bonus = #chained_bonus atp_relevance_fudge,
 179.172 -   max_imperfect = #max_imperfect atp_relevance_fudge,
 179.173 -   max_imperfect_exp = #max_imperfect_exp atp_relevance_fudge,
 179.174 -   threshold_divisor = #threshold_divisor atp_relevance_fudge,
 179.175 -   ridiculous_threshold = #ridiculous_threshold atp_relevance_fudge}
 179.176 -
 179.177 -fun relevance_fudge_of_prover ctxt name =
 179.178 -  if is_smt_prover ctxt name then smt_relevance_fudge else atp_relevance_fudge
 179.179 -
 179.180  fun supported_provers ctxt =
 179.181    let
 179.182      val thy = Proof_Context.theory_of ctxt
 179.183 @@ -383,28 +286,9 @@
 179.184     preplay_timeout : Time.time option,
 179.185     expect : string}
 179.186  
 179.187 -type relevance_fudge =
 179.188 -  {local_const_multiplier : real,
 179.189 -   worse_irrel_freq : real,
 179.190 -   higher_order_irrel_weight : real,
 179.191 -   abs_rel_weight : real,
 179.192 -   abs_irrel_weight : real,
 179.193 -   theory_const_rel_weight : real,
 179.194 -   theory_const_irrel_weight : real,
 179.195 -   chained_const_irrel_weight : real,
 179.196 -   intro_bonus : real,
 179.197 -   elim_bonus : real,
 179.198 -   simp_bonus : real,
 179.199 -   local_bonus : real,
 179.200 -   assum_bonus : real,
 179.201 -   chained_bonus : real,
 179.202 -   max_imperfect : real,
 179.203 -   max_imperfect_exp : real,
 179.204 -   threshold_divisor : real,
 179.205 -   ridiculous_threshold : real}
 179.206 -
 179.207  type prover_problem =
 179.208 -  {state : Proof.state,
 179.209 +  {comment : string,
 179.210 +   state : Proof.state,
 179.211     goal : thm,
 179.212     subgoal : int,
 179.213     subgoal_count : int,
 179.214 @@ -681,7 +565,7 @@
 179.215                      max_new_mono_instances, isar_proofs, isar_compress,
 179.216                      isar_try0, slice, timeout, preplay_timeout, ...})
 179.217          minimize_command
 179.218 -        ({state, goal, subgoal, subgoal_count, factss, ...} : prover_problem) =
 179.219 +        ({comment, state, goal, subgoal, subgoal_count, factss, ...} : prover_problem) =
 179.220    let
 179.221      val thy = Proof.theory_of state
 179.222      val ctxt = Proof.context_of state
 179.223 @@ -734,15 +618,15 @@
 179.224        let
 179.225          (* If slicing is disabled, we expand the last slice to fill the entire
 179.226             time available. *)
 179.227 -        val actual_slices = get_slices slice (best_slices ctxt)
 179.228 +        val all_slices = best_slices ctxt
 179.229 +        val actual_slices = get_slices slice all_slices
 179.230 +        fun max_facts_of_slices f slices = fold (Integer.max o slice_max_facts o f) slices 0
 179.231          val num_actual_slices = length actual_slices
 179.232          val max_fact_factor =
 179.233 -          case max_facts of
 179.234 -            NONE => 1.0
 179.235 -          | SOME max =>
 179.236 -            Real.fromInt max
 179.237 -            / Real.fromInt (fold (Integer.max o slice_max_facts)
 179.238 -                                 actual_slices 0)
 179.239 +          Real.fromInt (case max_facts of
 179.240 +              NONE => max_facts_of_slices I all_slices
 179.241 +            | SOME max => max)
 179.242 +          / Real.fromInt (max_facts_of_slices snd actual_slices)
 179.243          fun monomorphize_facts facts =
 179.244            let
 179.245              val ctxt =
 179.246 @@ -845,7 +729,8 @@
 179.247              val _ =
 179.248                atp_problem
 179.249                |> lines_of_atp_problem format ord ord_info
 179.250 -              |> cons ("% " ^ command ^ "\n")
 179.251 +              |> cons ("% " ^ command ^ "\n" ^
 179.252 +                (if comment = "" then "" else "% " ^ comment ^ "\n"))
 179.253                |> File.write_list prob_path
 179.254              val ((output, run_time), (atp_proof, outcome)) =
 179.255                time_limit generous_slice_timeout Isabelle_System.bash_output
 179.256 @@ -1028,10 +913,8 @@
 179.257  val is_boring_builtin_typ =
 179.258    not o exists_subtype (member (op =) [@{typ nat}, @{typ int}, HOLogic.realT])
 179.259  
 179.260 -fun smt_filter_loop name
 179.261 -                    ({debug, verbose, overlord, max_mono_iters,
 179.262 -                      max_new_mono_instances, timeout, slice, ...} : params)
 179.263 -                    state goal i =
 179.264 +fun smt_filter_loop name ({debug, overlord, max_mono_iters, max_new_mono_instances, timeout, slice,
 179.265 +      ...} : params) state goal i =
 179.266    let
 179.267      fun repair_context ctxt =
 179.268        ctxt |> select_smt_solver name
 179.269 @@ -1044,7 +927,7 @@
 179.270                   I)
 179.271             |> Config.put SMT_Config.infer_triggers
 179.272                           (Config.get ctxt smt_triggers)
 179.273 -           |> not (Config.get ctxt smt_builtin_trans)
 179.274 +           |> not (Config.get ctxt smt_builtins)
 179.275                ? (SMT_Builtin.filter_builtins is_boring_builtin_typ
 179.276                   #> Config.put SMT_Config.datatypes false)
 179.277             |> repair_monomorph_context max_mono_iters default_max_mono_iters
   180.1 --- a/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML	Mon Nov 11 17:34:44 2013 +0100
   180.2 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML	Mon Nov 11 17:44:21 2013 +0100
   180.3 @@ -640,7 +640,7 @@
   180.4  fun isar_proof_would_be_a_good_idea preplay =
   180.5    case preplay of
   180.6      Played (reconstr, _) => reconstr = SMT
   180.7 -  | Trust_Playable _ => true
   180.8 +  | Trust_Playable _ => false
   180.9    | Failed_to_Play _ => true
  180.10  
  180.11  fun proof_text ctxt isar_proofs isar_params num_chained
   181.1 --- a/src/HOL/Tools/Sledgehammer/sledgehammer_run.ML	Mon Nov 11 17:34:44 2013 +0100
   181.2 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_run.ML	Mon Nov 11 17:44:21 2013 +0100
   181.3 @@ -68,21 +68,21 @@
   181.4  fun launch_prover (params as {debug, verbose, spy, blocking, max_facts, slice,
   181.5                                timeout, expect, ...})
   181.6          mode output_result minimize_command only learn
   181.7 -        {state, goal, subgoal, subgoal_count, factss as (_, facts) :: _} name =
   181.8 +        {comment, state, goal, subgoal, subgoal_count, factss as (_, facts) :: _} name =
   181.9    let
  181.10      val ctxt = Proof.context_of state
  181.11  
  181.12      val hard_timeout = time_mult 3.0 (timeout |> the_default one_day)
  181.13 -    val _ = spying spy (fn () => (state, subgoal, name, "launched"));
  181.14 +    val _ = spying spy (fn () => (state, subgoal, name, "Launched"));
  181.15      val birth_time = Time.now ()
  181.16      val death_time = Time.+ (birth_time, hard_timeout)
  181.17 -    val max_facts = max_facts |> the_default (default_max_facts_of_prover ctxt slice name)
  181.18 +    val max_facts = max_facts |> the_default (default_max_facts_of_prover ctxt name)
  181.19      val num_facts = length facts |> not only ? Integer.min max_facts
  181.20  
  181.21      fun desc () = prover_description ctxt params name num_facts subgoal subgoal_count goal
  181.22  
  181.23      val problem =
  181.24 -      {state = state, goal = goal, subgoal = subgoal,
  181.25 +      {comment = comment, state = state, goal = goal, subgoal = subgoal,
  181.26         subgoal_count = subgoal_count,
  181.27         factss = factss
  181.28           |> map (apsnd ((not (is_ho_atp ctxt name)
  181.29 @@ -100,13 +100,33 @@
  181.30                    " proof (of " ^ string_of_int (length facts) ^ "): ") "."
  181.31        |> Output.urgent_message
  181.32  
  181.33 -    fun spying_str_of_res ({outcome = NONE, used_facts, ...} : prover_result) =
  181.34 -        let val num_used_facts = length used_facts in
  181.35 -          "success: Found proof with " ^ string_of_int num_used_facts ^ " fact" ^
  181.36 -          plural_s num_used_facts
  181.37 +    fun spying_str_of_res ({outcome = NONE, used_facts, used_from, ...} : prover_result) =
  181.38 +        let
  181.39 +          val num_used_facts = length used_facts
  181.40 +
  181.41 +          fun find_indices facts =
  181.42 +            tag_list 1 facts
  181.43 +            |> map (fn (j, fact) => fact |> apsnd (K j))
  181.44 +            |> filter_used_facts false used_facts
  181.45 +            |> map (prefix "@" o string_of_int o snd)
  181.46 +
  181.47 +          fun filter_info (fact_filter, facts) =
  181.48 +            let
  181.49 +              val indices = find_indices facts
  181.50 +              val unknowns = replicate (num_used_facts - length indices) "?"
  181.51 +            in (commas (indices @ unknowns), fact_filter) end
  181.52 +
  181.53 +          val filter_infos =
  181.54 +            map filter_info (("actual", used_from) :: factss)
  181.55 +            |> AList.group (op =)
  181.56 +            |> map (fn (indices, fact_filters) => commas fact_filters ^ ": " ^ indices)
  181.57 +        in
  181.58 +          "Success: Found proof with " ^ string_of_int num_used_facts ^
  181.59 +          " of " ^ string_of_int num_facts ^ " fact" ^ plural_s num_facts ^
  181.60 +          (if num_used_facts = 0 then "" else ": " ^ commas filter_infos)
  181.61          end
  181.62        | spying_str_of_res {outcome = SOME failure, ...} =
  181.63 -        "failure: " ^ string_of_atp_failure failure
  181.64 +        "Failure: " ^ string_of_atp_failure failure
  181.65  
  181.66      fun really_go () =
  181.67        problem
  181.68 @@ -209,8 +229,7 @@
  181.69        val _ = Proof.assert_backward state
  181.70        val print =
  181.71          if mode = Normal andalso is_none output_result then Output.urgent_message else K ()
  181.72 -      val state =
  181.73 -        state |> Proof.map_context (Config.put SMT_Config.verbose debug)
  181.74 +      val state = state |> Proof.map_context (Config.put SMT_Config.verbose debug)
  181.75        val ctxt = Proof.context_of state
  181.76        val {facts = chained, goal, ...} = Proof.goal state
  181.77        val (_, hyp_ts, concl_t) = strip_subgoal goal i ctxt
  181.78 @@ -225,11 +244,9 @@
  181.79                  SOME name => error ("No such prover: " ^ name ^ ".")
  181.80                | NONE => ()
  181.81        val _ = print "Sledgehammering..."
  181.82 -      val _ =
  181.83 -        spying spy (fn () => (state, i, "***", "starting " ^ @{make_string} mode ^ " mode"))
  181.84 -      val (smts, (ueq_atps, full_atps)) =
  181.85 -        provers |> List.partition (is_smt_prover ctxt)
  181.86 -                ||> List.partition (is_unit_equational_atp ctxt)
  181.87 +      val _ = spying spy (fn () => (state, i, "***", "Starting " ^ @{make_string} mode ^ " mode"))
  181.88 +
  181.89 +      val (ueq_atps, full_provers) = List.partition (is_unit_equational_atp ctxt) provers
  181.90  
  181.91        val spying_str_of_factss =
  181.92          commas o map (fn (filter, facts) => filter ^ ": " ^ string_of_int (length facts))
  181.93 @@ -240,18 +257,16 @@
  181.94              case max_facts of
  181.95                SOME n => n
  181.96              | NONE =>
  181.97 -              0 |> fold (Integer.max o default_max_facts_of_prover ctxt slice)
  181.98 -                        provers
  181.99 +              0 |> fold (Integer.max o default_max_facts_of_prover ctxt) provers
 181.100                  |> mode = Auto_Try ? (fn n => n div auto_try_max_facts_divisor)
 181.101            val _ = spying spy (fn () => (state, i, label ^ "s",
 181.102 -            "filtering " ^ string_of_int (length all_facts) ^ " facts"));
 181.103 +            "Filtering " ^ string_of_int (length all_facts) ^ " facts"));
 181.104          in
 181.105            all_facts
 181.106            |> (case is_appropriate_prop of
 181.107                  SOME is_app => filter (is_app o prop_of o snd)
 181.108                | NONE => I)
 181.109 -          |> relevant_facts ctxt params (hd provers) max_max_facts fact_override
 181.110 -                            hyp_ts concl_t
 181.111 +          |> relevant_facts ctxt params (hd provers) max_max_facts fact_override hyp_ts concl_t
 181.112            |> tap (fn factss =>
 181.113                       if verbose then
 181.114                         label ^ plural_s (length provers) ^ ": " ^
 181.115 @@ -259,16 +274,15 @@
 181.116                         |> print
 181.117                       else
 181.118                         ())
 181.119 -          |> spy ? tap (fn factss =>
 181.120 -            spying spy (fn () =>
 181.121 -              (state, i, label ^ "s", "selected facts: " ^ spying_str_of_factss factss)))
 181.122 +          |> spy ? tap (fn factss => spying spy (fn () =>
 181.123 +            (state, i, label ^ "s", "Selected facts: " ^ spying_str_of_factss factss)))
 181.124          end
 181.125  
 181.126        fun launch_provers state label is_appropriate_prop provers =
 181.127          let
 181.128            val factss = get_factss label is_appropriate_prop provers
 181.129            val problem =
 181.130 -            {state = state, goal = goal, subgoal = i, subgoal_count = n,
 181.131 +            {comment = "", state = state, goal = goal, subgoal = i, subgoal_count = n,
 181.132               factss = factss}
 181.133            fun learn prover =
 181.134              mash_learn_proof ctxt params prover (prop_of goal) all_facts
 181.135 @@ -287,30 +301,26 @@
 181.136              |> max_outcome_code |> rpair state
 181.137          end
 181.138  
 181.139 -      fun launch_atps label is_appropriate_prop atps accum =
 181.140 -        if null atps then
 181.141 +      fun maybe_launch_provers label is_appropriate_prop provers_to_launch accum =
 181.142 +        if null provers_to_launch then
 181.143            accum
 181.144          else if is_some is_appropriate_prop andalso
 181.145                  not (the is_appropriate_prop concl_t) then
 181.146 -          (if verbose orelse length atps = length provers then
 181.147 +          (if verbose orelse length provers_to_launch = length provers then
 181.148               "Goal outside the scope of " ^
 181.149 -             space_implode " " (serial_commas "and" (map quote atps)) ^ "."
 181.150 +             space_implode " " (serial_commas "and" (map quote provers_to_launch)) ^ "."
 181.151               |> Output.urgent_message
 181.152             else
 181.153               ();
 181.154             accum)
 181.155          else
 181.156 -          launch_provers state label is_appropriate_prop atps
 181.157 +          launch_provers state label is_appropriate_prop provers_to_launch
 181.158  
 181.159 -      fun launch_smts accum =
 181.160 -        if null smts then accum else launch_provers state "SMT solver" NONE smts
 181.161 -
 181.162 -      val launch_full_atps = launch_atps "ATP" NONE full_atps
 181.163 -
 181.164 -      val launch_ueq_atps = launch_atps "Unit-equational provers" (SOME is_unit_equality) ueq_atps
 181.165 +      val launch_full_provers = maybe_launch_provers "ATP/SMT" NONE full_provers
 181.166 +      val launch_ueq_atps = maybe_launch_provers "Unit-equational provers" (SOME is_unit_equality) ueq_atps
 181.167  
 181.168        fun launch_atps_and_smt_solvers p =
 181.169 -        [launch_full_atps, launch_smts, launch_ueq_atps]
 181.170 +        [launch_full_provers, launch_ueq_atps]
 181.171          |> Par_List.map (fn f => fst (f p))
 181.172          handle ERROR msg => (print ("Error: " ^ msg); error msg)
 181.173  
 181.174 @@ -318,14 +328,9 @@
 181.175          accum |> (mode = Normal orelse outcome_code <> someN) ? f
 181.176      in
 181.177        (unknownN, state)
 181.178 -      |> (if mode = Auto_Try then
 181.179 -            launch_full_atps
 181.180 -          else if blocking then
 181.181 -            launch_atps_and_smt_solvers #> max_outcome_code #> rpair state
 181.182 -          else
 181.183 -            (fn p => (Future.fork (tap (fn () => launch_atps_and_smt_solvers p)); p)))
 181.184 -      handle TimeLimit.TimeOut =>
 181.185 -             (print "Sledgehammer ran out of time."; (unknownN, state))
 181.186 +      |> (if blocking then launch_full_provers
 181.187 +          else (fn p => (Future.fork (tap (fn () => launch_full_provers p)); p)))
 181.188 +      handle TimeLimit.TimeOut => (print "Sledgehammer ran out of time."; (unknownN, state))
 181.189      end
 181.190      |> `(fn (outcome_code, _) => outcome_code = someN)
 181.191  
   182.1 --- a/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Mon Nov 11 17:34:44 2013 +0100
   182.2 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Mon Nov 11 17:44:21 2013 +0100
   182.3 @@ -97,44 +97,29 @@
   182.4  (* FIXME: Similar yet different code in "mirabelle.ML". The code here has a few
   182.5     fixes that seem to be missing over there; or maybe the two code portions are
   182.6     not doing the same? *)
   182.7 -fun fold_body_thms outer_name (map_plain_name, map_inclass_name) =
   182.8 +fun fold_body_thm outer_name (map_plain_name, map_inclass_name) =
   182.9    let
  182.10 -    fun app map_name n (PBody {thms, ...}) =
  182.11 -      thms |> fold (fn (_, (name, _, body)) => fn accum =>
  182.12 -          let
  182.13 -            val collect = union (op =) o the_list o map_name
  182.14 -            (* The "name = outer_name" case caters for the uncommon case where
  182.15 -               the proved theorem occurs in its own proof (e.g.,
  182.16 -               "Transitive_Closure.trancl_into_trancl"). *)
  182.17 -            val (anonymous, enter_class) =
  182.18 -              if name = "" orelse (n = 1 andalso name = outer_name) then
  182.19 -                (true, false)
  182.20 -              else if n = 1 andalso map_inclass_name name = SOME outer_name then
  182.21 -                (true, true)
  182.22 -              else
  182.23 -                (false, false)
  182.24 -            val accum =
  182.25 -              accum |> (if n = 1 andalso not anonymous then collect name else I)
  182.26 -            val n = n + (if anonymous then 0 else 1)
  182.27 -          in
  182.28 -            accum
  182.29 -            |> (if n <= 1 then
  182.30 -                  app (if enter_class then map_inclass_name else map_name) n
  182.31 -                      (Future.join body)
  182.32 -                else
  182.33 -                  I)
  182.34 -          end)
  182.35 -  in fold (app map_plain_name 0) end
  182.36 +    fun app_thm map_name (_, (name, _, body)) accum =
  182.37 +      let
  182.38 +        val (anonymous, enter_class) =
  182.39 +          (* The "name = outer_name" case caters for the uncommon case where the proved theorem
  182.40 +             occurs in its own proof (e.g., "Transitive_Closure.trancl_into_trancl"). *)
  182.41 +          if name = "" orelse name = outer_name then (true, false)
  182.42 +          else if map_inclass_name name = SOME outer_name then (true, true)
  182.43 +          else (false, false)
  182.44 +      in
  182.45 +        if anonymous then
  182.46 +          accum |> app_body (if enter_class then map_inclass_name else map_name) (Future.join body)
  182.47 +        else
  182.48 +          accum |> union (op =) (the_list (map_name name))
  182.49 +      end
  182.50 +    and app_body map_name (PBody {thms, ...}) = fold (app_thm map_name) thms
  182.51 +  in app_body map_plain_name end
  182.52  
  182.53  fun thms_in_proof name_tabs th =
  182.54 -  let
  182.55 -    val map_names =
  182.56 -      case name_tabs of
  182.57 -        SOME p => pairself Symtab.lookup p
  182.58 -      | NONE => `I SOME
  182.59 -    val names =
  182.60 -      fold_body_thms (Thm.get_name_hint th) map_names [Thm.proof_body_of th] []
  182.61 -  in names end
  182.62 +  let val map_names = (case name_tabs of SOME p => pairself Symtab.lookup p | NONE => `I SOME) in
  182.63 +    fold_body_thm (Thm.get_name_hint th) map_names (Proofterm.strip_thm (Thm.proof_body_of th)) []
  182.64 +  end
  182.65  
  182.66  fun thms_of_name ctxt name =
  182.67    let
  182.68 @@ -162,7 +147,7 @@
  182.69  fun hackish_string_of_term ctxt =
  182.70    with_vanilla_print_mode (Syntax.string_of_term ctxt) #> simplify_spaces
  182.71  
  182.72 -val spying_version = "b"
  182.73 +val spying_version = "c"
  182.74  
  182.75  fun spying false _ = ()
  182.76    | spying true f =
   183.1 --- a/src/HOL/Tools/groebner.ML	Mon Nov 11 17:34:44 2013 +0100
   183.2 +++ b/src/HOL/Tools/groebner.ML	Mon Nov 11 17:44:21 2013 +0100
   183.3 @@ -21,11 +21,6 @@
   183.4  structure Groebner : GROEBNER =
   183.5  struct
   183.6  
   183.7 -fun is_comb ct =
   183.8 -  (case Thm.term_of ct of
   183.9 -    _ $ _ => true
  183.10 -  | _ => false);
  183.11 -
  183.12  val concl = Thm.cprop_of #> Thm.dest_arg;
  183.13  
  183.14  fun is_binop ct ct' =
  183.15 @@ -37,8 +32,6 @@
  183.16    if is_binop ct ct' then Thm.dest_binop ct'
  183.17    else raise CTERM ("dest_binary: bad binop", [ct, ct'])
  183.18  
  183.19 -fun inst_thm inst = Thm.instantiate ([], inst);
  183.20 -
  183.21  val rat_0 = Rat.zero;
  183.22  val rat_1 = Rat.one;
  183.23  val minus_rat = Rat.neg;
  183.24 @@ -77,10 +70,6 @@
  183.25      n1 < n2 orelse n1 = n2 andalso lexorder m1 m2
  183.26      end;
  183.27  
  183.28 -fun morder_le m1 m2 = morder_lt m1 m2 orelse (m1 = m2);
  183.29 -
  183.30 -fun morder_gt m1 m2 = morder_lt m2 m1;
  183.31 -
  183.32  (* Arithmetic on canonical polynomials. *)
  183.33  
  183.34  fun grob_neg l = map (fn (c,m) => (minus_rat c,m)) l;
  183.35 @@ -125,33 +114,9 @@
  183.36  
  183.37  fun grob_pow vars l n =
  183.38    if n < 0 then error "grob_pow: negative power"
  183.39 -  else if n = 0 then [(rat_1,map (fn v => 0) vars)]
  183.40 +  else if n = 0 then [(rat_1,map (K 0) vars)]
  183.41    else grob_mul l (grob_pow vars l (n - 1));
  183.42  
  183.43 -fun degree vn p =
  183.44 - case p of
  183.45 -  [] => error "Zero polynomial"
  183.46 -| [(c,ns)] => nth ns vn
  183.47 -| (c,ns)::p' => Int.max (nth ns vn, degree vn p');
  183.48 -
  183.49 -fun head_deg vn p = let val d = degree vn p in
  183.50 - (d,fold (fn (c,r) => fn q => grob_add q [(c, map_index (fn (i,n) => if i = vn then 0 else n) r)]) (filter (fn (c,ns) => c <>/ rat_0 andalso nth ns vn = d) p) []) end;
  183.51 -
  183.52 -val is_zerop = forall (fn (c,ns) => c =/ rat_0 andalso forall (curry (op =) 0) ns);
  183.53 -val grob_pdiv =
  183.54 - let fun pdiv_aux vn (n,a) p k s =
  183.55 -  if is_zerop s then (k,s) else
  183.56 -  let val (m,b) = head_deg vn s
  183.57 -  in if m < n then (k,s) else
  183.58 -     let val p' = grob_mul p [(rat_1, map_index (fn (i,v) => if i = vn then m - n else 0)
  183.59 -                                                (snd (hd s)))]
  183.60 -     in if a = b then pdiv_aux vn (n,a) p k (grob_sub s p')
  183.61 -        else pdiv_aux vn (n,a) p (k + 1) (grob_sub (grob_mul a s) (grob_mul b p'))
  183.62 -     end
  183.63 -  end
  183.64 - in fn vn => fn s => fn p => pdiv_aux vn (head_deg vn p) p 0 s
  183.65 - end;
  183.66 -
  183.67  (* Monomial division operation. *)
  183.68  
  183.69  fun mdiv (c1,m1) (c2,m2) =
  183.70 @@ -160,7 +125,7 @@
  183.71  
  183.72  (* Lowest common multiple of two monomials. *)
  183.73  
  183.74 -fun mlcm (c1,m1) (c2,m2) = (rat_1, ListPair.map Int.max (m1, m2));
  183.75 +fun mlcm (_,m1) (_,m2) = (rat_1, ListPair.map Int.max (m1, m2));
  183.76  
  183.77  (* Reduce monomial cm by polynomial pol, returning replacement for cm.  *)
  183.78  
  183.79 @@ -200,8 +165,8 @@
  183.80  
  183.81  fun spoly cm ph1 ph2 =
  183.82    case (ph1,ph2) of
  183.83 -    (([],h),p) => ([],h)
  183.84 -  | (p,([],h)) => ([],h)
  183.85 +    (([],h),_) => ([],h)
  183.86 +  | (_,([],h)) => ([],h)
  183.87    | ((cm1::ptl1,his1),(cm2::ptl2,his2)) =>
  183.88          (grob_sub (grob_cmul (mdiv cm cm1) ptl1)
  183.89                    (grob_cmul (mdiv cm cm2) ptl2),
  183.90 @@ -218,12 +183,12 @@
  183.91  
  183.92  (* The most popular heuristic is to order critical pairs by LCM monomial.    *)
  183.93  
  183.94 -fun forder ((c1,m1),_) ((c2,m2),_) = morder_lt m1 m2;
  183.95 +fun forder ((_,m1),_) ((_,m2),_) = morder_lt m1 m2;
  183.96  
  183.97  fun poly_lt  p q =
  183.98    case (p,q) of
  183.99 -    (p,[]) => false
 183.100 -  | ([],q) => true
 183.101 +    (_,[]) => false
 183.102 +  | ([],_) => true
 183.103    | ((c1,m1)::o1,(c2,m2)::o2) =>
 183.104          c1 </ c2 orelse
 183.105          c1 =/ c2 andalso ((morder_lt m1 m2) orelse m1 = m2 andalso poly_lt o1 o2);
 183.106 @@ -234,7 +199,7 @@
 183.107  fun poly_eq p1 p2 =
 183.108    eq_list (fn ((c1, m1), (c2, m2)) => c1 =/ c2 andalso (m1: int list) = m2) (p1, p2);
 183.109  
 183.110 -fun memx ((p1,h1),(p2,h2)) ppairs =
 183.111 +fun memx ((p1,_),(p2,_)) ppairs =
 183.112    not (exists (fn ((q1,_),(q2,_)) => poly_eq p1 q1 andalso poly_eq p2 q2) ppairs);
 183.113  
 183.114  (* Buchberger's second criterion.                                            *)
 183.115 @@ -277,7 +242,7 @@
 183.116   case pairs of
 183.117     [] => basis
 183.118   | (l,(p1,p2))::opairs =>
 183.119 -   let val (sph as (sp,hist)) = monic (reduce basis (spoly l p1 p2))
 183.120 +   let val (sph as (sp,_)) = monic (reduce basis (spoly l p1 p2))
 183.121     in
 183.122      if null sp orelse criterion2 basis (l,(p1,p2)) opairs
 183.123      then grobner_basis basis opairs
 183.124 @@ -324,7 +289,7 @@
 183.125  
 183.126  fun grobner_refute pols =
 183.127    let val gb = grobner pols in
 183.128 -  snd(find (fn (p,h) => length p = 1 andalso forall (fn x=> x=0) (snd(hd p))) gb)
 183.129 +  snd(find (fn (p,_) => length p = 1 andalso forall (fn x=> x=0) (snd(hd p))) gb)
 183.130    end;
 183.131  
 183.132  (* Turn proof into a certificate as sum of multipliers.                      *)
 183.133 @@ -366,8 +331,8 @@
 183.134  
 183.135  fun grobner_strong vars pols pol =
 183.136      let val vars' = @{cterm "True"}::vars
 183.137 -        val grob_z = [(rat_1,1::(map (fn x => 0) vars))]
 183.138 -        val grob_1 = [(rat_1,(map (fn x => 0) vars'))]
 183.139 +        val grob_z = [(rat_1,1::(map (K 0) vars))]
 183.140 +        val grob_1 = [(rat_1,(map (K 0) vars'))]
 183.141          fun augment p= map (fn (c,m) => (c,0::m)) p
 183.142          val pols' = map augment pols
 183.143          val pol' = augment pol
 183.144 @@ -387,7 +352,7 @@
 183.145  
 183.146  fun refute_disj rfn tm =
 183.147   case term_of tm of
 183.148 -  Const(@{const_name HOL.disj},_)$l$r =>
 183.149 +  Const(@{const_name HOL.disj},_)$_$_ =>
 183.150     Drule.compose
 183.151      (refute_disj rfn (Thm.dest_arg tm), 2,
 183.152        Drule.compose (refute_disj rfn (Thm.dest_arg1 tm), 2, disjE))
 183.153 @@ -398,7 +363,7 @@
 183.154  
 183.155  fun is_neg t =
 183.156      case term_of t of
 183.157 -      (Const(@{const_name Not},_)$p) => true
 183.158 +      (Const(@{const_name Not},_)$_) => true
 183.159      | _  => false;
 183.160  fun is_eq t =
 183.161   case term_of t of
 183.162 @@ -423,7 +388,7 @@
 183.163  val strip_exists =
 183.164   let fun h (acc, t) =
 183.165        case term_of t of
 183.166 -       Const (@{const_name Ex}, _) $ Abs (x, T, p) =>
 183.167 +       Const (@{const_name Ex}, _) $ Abs _ =>
 183.168          h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc))
 183.169       | _ => (acc,t)
 183.170   in fn t => h ([],t)
 183.171 @@ -435,10 +400,7 @@
 183.172  | _ => false;
 183.173  
 183.174  val mk_object_eq = fn th => th COMP meta_eq_to_obj_eq;
 183.175 -val bool_simps = @{thms bool_simps};
 183.176  val nnf_simps = @{thms nnf_simps};
 183.177 -fun nnf_conv ctxt =
 183.178 -  Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps bool_simps addsimps nnf_simps)
 183.179  
 183.180  fun weak_dnf_conv ctxt =
 183.181    Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms weak_dnf_simps});
 183.182 @@ -484,12 +446,10 @@
 183.183  
 183.184  fun fold1 f = foldr1 (uncurry f);
 183.185  
 183.186 -val list_conj = fold1 (fn c => fn c' => Thm.apply (Thm.apply @{cterm HOL.conj} c) c') ;
 183.187 -
 183.188  fun mk_conj_tab th =
 183.189   let fun h acc th =
 183.190     case prop_of th of
 183.191 -   @{term "Trueprop"}$(@{term HOL.conj}$p$q) =>
 183.192 +   @{term "Trueprop"}$(@{term HOL.conj}$_$_) =>
 183.193       h (h acc (th RS conjunct2)) (th RS conjunct1)
 183.194    | @{term "Trueprop"}$p => (p,th)::acc
 183.195  in fold (Termtab.insert Thm.eq_thm) (h [] th) Termtab.empty end;
 183.196 @@ -567,8 +527,7 @@
 183.197   | Var ((s,_),_) => s
 183.198   | _ => "x"
 183.199   fun mk_eq s t = Thm.apply (Thm.apply @{cterm "op == :: bool => _"} s) t
 183.200 - fun mkeq s t = Thm.apply @{cterm Trueprop} (Thm.apply (Thm.apply @{cterm "op = :: bool => _"} s) t)
 183.201 - fun mk_exists v th = Drule.arg_cong_rule (ext (ctyp_of_term v))
 183.202 +  fun mk_exists v th = Drule.arg_cong_rule (ext (ctyp_of_term v))
 183.203     (Thm.abstract_rule (getname v) v th)
 183.204   fun simp_ex_conv ctxt =
 183.205     Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms(39)})
 183.206 @@ -585,8 +544,8 @@
 183.207  (** main **)
 183.208  
 183.209  fun ring_and_ideal_conv
 183.210 -  {vars, semiring = (sr_ops, sr_rules), ring = (r_ops, r_rules),
 183.211 -   field = (f_ops, f_rules), idom, ideal}
 183.212 +  {vars = _, semiring = (sr_ops, _), ring = (r_ops, _),
 183.213 +   field = (f_ops, _), idom, ideal}
 183.214    dest_const mk_const ring_eq_conv ring_normalize_conv =
 183.215  let
 183.216    val [add_pat, mul_pat, pow_pat, zero_tm, one_tm] = sr_ops;
 183.217 @@ -612,7 +571,6 @@
 183.218         else raise CTERM ("ring_dest_neg", [t])
 183.219      end
 183.220  
 183.221 - val ring_mk_neg = fn tm => Thm.apply (ring_neg_tm) (tm);
 183.222   fun field_dest_inv t =
 183.223      let val (l,r) = Thm.dest_comb t in
 183.224          if Term.could_unify(term_of l, term_of field_inv_tm) then r
 183.225 @@ -621,11 +579,9 @@
 183.226   val ring_dest_add = dest_binary ring_add_tm;
 183.227   val ring_mk_add = mk_binop ring_add_tm;
 183.228   val ring_dest_sub = dest_binary ring_sub_tm;
 183.229 - val ring_mk_sub = mk_binop ring_sub_tm;
 183.230   val ring_dest_mul = dest_binary ring_mul_tm;
 183.231   val ring_mk_mul = mk_binop ring_mul_tm;
 183.232   val field_dest_div = dest_binary field_div_tm;
 183.233 - val field_mk_div = mk_binop field_div_tm;
 183.234   val ring_dest_pow = dest_binary ring_pow_tm;
 183.235   val ring_mk_pow = mk_binop ring_pow_tm ;
 183.236   fun grobvars tm acc =
 183.237 @@ -652,7 +608,7 @@
 183.238       [(rat_1,map (fn i => if i aconvc tm then 1 else 0) vars)])
 183.239  handle  CTERM _ =>
 183.240   ((let val x = dest_const tm
 183.241 - in if x =/ rat_0 then [] else [(x,map (fn v => 0) vars)]
 183.242 + in if x =/ rat_0 then [] else [(x,map (K 0) vars)]
 183.243   end)
 183.244   handle ERROR _ =>
 183.245    ((grob_neg(grobify_term vars (ring_dest_neg tm)))
 183.246 @@ -732,7 +688,7 @@
 183.247          Conv.fconv_rule
 183.248            ((Conv.arg_conv #> Conv.arg_conv) (Conv.binop_conv ring_normalize_conv)) th1
 183.249        val conc = th2 |> concl |> Thm.dest_arg
 183.250 -      val (l,r) = conc |> dest_eq
 183.251 +      val (l,_) = conc |> dest_eq
 183.252      in Thm.implies_intr (Thm.apply cTrp tm)
 183.253                      (Thm.equal_elim (Drule.arg_cong_rule cTrp (eqF_intr th2))
 183.254                             (Thm.reflexive l |> mk_object_eq))
 183.255 @@ -756,9 +712,9 @@
 183.256         val th2 = funpow deg (idom_rule ctxt o HOLogic.conj_intr th1) neq_01
 183.257        in (vars,l,cert,th2)
 183.258        end)
 183.259 -    val cert_pos = map (fn (i,p) => (i,filter (fn (c,m) => c >/ rat_0) p)) cert
 183.260 +    val cert_pos = map (fn (i,p) => (i,filter (fn (c,_) => c >/ rat_0) p)) cert
 183.261      val cert_neg = map (fn (i,p) => (i,map (fn (c,m) => (minus_rat c,m))
 183.262 -                                            (filter (fn (c,m) => c </ rat_0) p))) cert
 183.263 +                                            (filter (fn (c,_) => c </ rat_0) p))) cert
 183.264      val  herts_pos = map (fn (i,p) => (i,holify_polynomial vars p)) cert_pos
 183.265      val  herts_neg = map (fn (i,p) => (i,holify_polynomial vars p)) cert_neg
 183.266      fun thm_fn pols =
 183.267 @@ -772,7 +728,7 @@
 183.268      val th4 =
 183.269        Conv.fconv_rule ((Conv.arg_conv o Conv.arg_conv o Conv.binop_conv) ring_normalize_conv)
 183.270          (neq_rule l th3)
 183.271 -    val (l,r) = dest_eq(Thm.dest_arg(concl th4))
 183.272 +    val (l, _) = dest_eq(Thm.dest_arg(concl th4))
 183.273     in Thm.implies_intr (Thm.apply cTrp tm)
 183.274                          (Thm.equal_elim (Drule.arg_cong_rule cTrp (eqF_intr th4))
 183.275                     (Thm.reflexive l |> mk_object_eq))
 183.276 @@ -873,7 +829,6 @@
 183.277        (Drule.binop_cong_rule @{cterm HOL.conj} th1
 183.278          (Thm.reflexive (Thm.dest_arg (Thm.rhs_of th2))))
 183.279    val v = Thm.dest_arg1(Thm.dest_arg1(Thm.rhs_of th3))
 183.280 -  val vars' = (remove op aconvc v vars) @ [v]
 183.281    val th4 = Conv.fconv_rule (Conv.arg_conv (simp_ex_conv ctxt)) (mk_exists v th3)
 183.282    val th5 = ex_eq_conv (mk_eq tm (fold mk_ex (remove op aconvc v vars) (Thm.lhs_of th4)))
 183.283   in Thm.transitive th5 (fold mk_exists (remove op aconvc v vars) th4)
 183.284 @@ -961,23 +916,12 @@
 183.285  | SOME tm =>
 183.286    (case Semiring_Normalizer.match ctxt tm of
 183.287      NONE => NONE
 183.288 -  | SOME (res as (theory, {is_const, dest_const,
 183.289 +  | SOME (res as (theory, {is_const = _, dest_const,
 183.290            mk_const, conv = ring_eq_conv})) =>
 183.291       SOME (ring_and_ideal_conv theory
 183.292            dest_const (mk_const (ctyp_of_term tm)) (ring_eq_conv ctxt)
 183.293            (Semiring_Normalizer.semiring_normalize_wrapper ctxt res)))
 183.294  
 183.295 -fun ring_solve ctxt form =
 183.296 -  (case try (find_term 0 (* FIXME !? *)) form of
 183.297 -    NONE => Thm.reflexive form
 183.298 -  | SOME tm =>
 183.299 -      (case Semiring_Normalizer.match ctxt tm of
 183.300 -        NONE => Thm.reflexive form
 183.301 -      | SOME (res as (theory, {is_const, dest_const, mk_const, conv = ring_eq_conv})) =>
 183.302 -        #ring_conv (ring_and_ideal_conv theory
 183.303 -          dest_const (mk_const (ctyp_of_term tm)) (ring_eq_conv ctxt)
 183.304 -          (Semiring_Normalizer.semiring_normalize_wrapper ctxt res)) ctxt form));
 183.305 -
 183.306  fun presimplify ctxt add_thms del_thms =
 183.307    asm_full_simp_tac (put_simpset HOL_basic_ss ctxt
 183.308      addsimps (Algebra_Simplification.get ctxt)
 183.309 @@ -1014,7 +958,7 @@
 183.310   | SOME thy =>
 183.311    let
 183.312     fun poly_exists_tac {asms = asms, concl = concl, prems = prems,
 183.313 -            params = params, context = ctxt, schematics = scs} =
 183.314 +            params = _, context = ctxt, schematics = _} =
 183.315      let
 183.316       val (evs,bod) = strip_exists (Thm.dest_arg concl)
 183.317       val ps = map_filter (try (lhs o Thm.dest_arg)) asms
   184.1 --- a/src/HOL/Tools/group_cancel.ML	Mon Nov 11 17:34:44 2013 +0100
   184.2 +++ b/src/HOL/Tools/group_cancel.ML	Mon Nov 11 17:44:21 2013 +0100
   184.3 @@ -25,7 +25,7 @@
   184.4  val sub1 = @{lemma "(A::'a::ab_group_add) == k + a ==> A - b == k + (a - b)"
   184.5        by (simp only: add_diff_eq)}
   184.6  val sub2 = @{lemma "(B::'a::ab_group_add) == k + b ==> a - B == - k + (a - b)"
   184.7 -      by (simp only: diff_minus minus_add add_ac)}
   184.8 +      by (simp only: minus_add diff_conv_add_uminus add_ac)}
   184.9  val neg1 = @{lemma "(A::'a::ab_group_add) == k + a ==> - A == - k + - a"
  184.10        by (simp only: minus_add_distrib)}
  184.11  val rule0 = @{lemma "(a::'a::comm_monoid_add) == a + 0"
   185.1 --- a/src/HOL/Tools/int_arith.ML	Mon Nov 11 17:34:44 2013 +0100
   185.2 +++ b/src/HOL/Tools/int_arith.ML	Mon Nov 11 17:44:21 2013 +0100
   185.3 @@ -87,9 +87,9 @@
   185.4  val setup =
   185.5    Lin_Arith.add_inj_thms [@{thm zle_int} RS iffD2, @{thm int_int_eq} RS iffD2]
   185.6    #> Lin_Arith.add_lessD @{thm zless_imp_add1_zle}
   185.7 -  #> Lin_Arith.add_simps (@{thms simp_thms} @ @{thms arith_simps} @ @{thms rel_simps}
   185.8 -      @ @{thms pred_numeral_simps}
   185.9 -      @ @{thms arith_special} @ @{thms int_arith_rules})
  185.10 +  #> Lin_Arith.add_simps @{thms of_nat_simps of_int_simps}
  185.11 +  #> Lin_Arith.add_simps
  185.12 +      [@{thm of_int_numeral}, @{thm nat_0}, @{thm nat_1}, @{thm diff_nat_numeral}, @{thm nat_numeral}]
  185.13    #> Lin_Arith.add_simprocs [zero_one_idom_simproc]
  185.14    #> Lin_Arith.set_number_of number_of
  185.15    #> Lin_Arith.add_inj_const (@{const_name of_nat}, HOLogic.natT --> HOLogic.intT)
   186.1 --- a/src/HOL/Tools/lin_arith.ML	Mon Nov 11 17:34:44 2013 +0100
   186.2 +++ b/src/HOL/Tools/lin_arith.ML	Mon Nov 11 17:44:21 2013 +0100
   186.3 @@ -791,37 +791,16 @@
   186.4     Most of the work is done by the cancel tactics. *)
   186.5  
   186.6  val init_arith_data =
   186.7 -  Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, number_of, ...} =>
   186.8 -   {add_mono_thms = @{thms add_mono_thms_linordered_semiring} @
   186.9 -      @{thms add_mono_thms_linordered_field} @ add_mono_thms,
  186.10 -    mult_mono_thms = @{thm mult_strict_left_mono} :: @{thm mult_left_mono} ::
  186.11 -      @{lemma "a = b ==> c*a = c*b" by (rule arg_cong)} :: mult_mono_thms,
  186.12 +  Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, number_of, ...} =>
  186.13 +   {add_mono_thms = @{thms add_mono_thms_linordered_semiring}
  186.14 +      @ @{thms add_mono_thms_linordered_field} @ add_mono_thms,
  186.15 +    mult_mono_thms = @{thm mult_strict_left_mono} :: @{thm mult_left_mono}
  186.16 +      :: @{lemma "a = b ==> c * a = c * b" by (rule arg_cong)} :: mult_mono_thms,
  186.17      inj_thms = inj_thms,
  186.18 -    lessD = lessD @ [@{thm "Suc_leI"}],
  186.19 -    neqE = [@{thm linorder_neqE_nat}, @{thm linorder_neqE_linordered_idom}],
  186.20 -    simpset =
  186.21 -      put_simpset HOL_basic_ss @{context}
  186.22 -      addsimps @{thms ring_distribs}
  186.23 -      addsimps [@{thm if_True}, @{thm if_False}]
  186.24 -      addsimps
  186.25 -       [@{thm add_0_left}, @{thm add_0_right},
  186.26 -        @{thm add_Suc}, @{thm add_Suc_right},
  186.27 -        @{thm nat.inject}, @{thm Suc_le_mono}, @{thm Suc_less_eq},
  186.28 -        @{thm "Zero_not_Suc"}, @{thm "Suc_not_Zero"}, @{thm "le_0_eq"}, @{thm "One_nat_def"},
  186.29 -        @{thm "order_less_irrefl"}, @{thm "zero_neq_one"}, @{thm "zero_less_one"},
  186.30 -        @{thm "zero_le_one"}, @{thm "zero_neq_one"} RS not_sym, @{thm "not_one_le_zero"},
  186.31 -        @{thm "not_one_less_zero"}]
  186.32 -      addsimprocs [@{simproc group_cancel_add}, @{simproc group_cancel_diff},
  186.33 -                   @{simproc group_cancel_eq}, @{simproc group_cancel_le},
  186.34 -                   @{simproc group_cancel_less}]
  186.35 -       (*abel_cancel helps it work in abstract algebraic domains*)
  186.36 -      addsimprocs [@{simproc nateq_cancel_sums},
  186.37 -                   @{simproc natless_cancel_sums},
  186.38 -                   @{simproc natle_cancel_sums}]
  186.39 -      |> Simplifier.add_cong @{thm if_weak_cong}
  186.40 -      |> simpset_of,
  186.41 -    number_of = number_of}) #>
  186.42 -  add_discrete_type @{type_name nat};
  186.43 +    lessD = lessD,
  186.44 +    neqE = @{thm linorder_neqE_nat} :: @{thm linorder_neqE_linordered_idom} :: neqE,
  186.45 +    simpset = put_simpset HOL_basic_ss @{context} |> Simplifier.add_cong @{thm if_weak_cong} |> simpset_of,
  186.46 +    number_of = number_of});
  186.47  
  186.48  (* FIXME !?? *)
  186.49  fun add_arith_facts ctxt =
  186.50 @@ -909,9 +888,6 @@
  186.51  
  186.52  (* context setup *)
  186.53  
  186.54 -val setup =
  186.55 -  init_arith_data;
  186.56 -
  186.57  val global_setup =
  186.58    map_theory_simpset (fn ctxt => ctxt
  186.59      addSolver (mk_solver "lin_arith" (add_arith_facts #> Fast_Arith.prems_lin_arith_tac))) #>
  186.60 @@ -924,4 +900,22 @@
  186.61            THEN' tac ctxt)))) "linear arithmetic" #>
  186.62    Arith_Data.add_tactic "linear arithmetic" gen_tac;
  186.63  
  186.64 +val setup =
  186.65 +  init_arith_data
  186.66 +  #> add_discrete_type @{type_name nat}
  186.67 +  #> add_lessD @{thm Suc_leI}
  186.68 +  #> add_simps (@{thms simp_thms} @ @{thms ring_distribs} @ [@{thm if_True}, @{thm if_False},
  186.69 +      @{thm add_0_left}, @{thm add_0_right}, @{thm order_less_irrefl},
  186.70 +      @{thm zero_neq_one}, @{thm zero_less_one}, @{thm zero_le_one},
  186.71 +      @{thm zero_neq_one} RS not_sym, @{thm not_one_le_zero}, @{thm not_one_less_zero}])
  186.72 +  #> add_simps [@{thm add_Suc}, @{thm add_Suc_right}, @{thm nat.inject},
  186.73 +      @{thm Suc_le_mono}, @{thm Suc_less_eq}, @{thm Zero_not_Suc},
  186.74 +      @{thm Suc_not_Zero}, @{thm le_0_eq}, @{thm One_nat_def}]
  186.75 +  #> add_simprocs [@{simproc group_cancel_add}, @{simproc group_cancel_diff},
  186.76 +      @{simproc group_cancel_eq}, @{simproc group_cancel_le},
  186.77 +      @{simproc group_cancel_less}]
  186.78 +     (*abel_cancel helps it work in abstract algebraic domains*)
  186.79 +  #> add_simprocs [@{simproc nateq_cancel_sums},@{simproc natless_cancel_sums},
  186.80 +      @{simproc natle_cancel_sums}];
  186.81 +
  186.82  end;
   187.1 --- a/src/HOL/Tools/numeral_simprocs.ML	Mon Nov 11 17:34:44 2013 +0100
   187.2 +++ b/src/HOL/Tools/numeral_simprocs.ML	Mon Nov 11 17:44:21 2013 +0100
   187.3 @@ -220,7 +220,7 @@
   187.4  val minus_simps = [@{thm minus_zero}, @{thm minus_one}, @{thm minus_numeral}, @{thm minus_neg_numeral}];
   187.5  
   187.6  (*To let us treat subtraction as addition*)
   187.7 -val diff_simps = [@{thm diff_minus}, @{thm minus_add_distrib}, @{thm minus_minus}];
   187.8 +val diff_simps = [@{thm diff_conv_add_uminus}, @{thm minus_add_distrib}, @{thm minus_minus}];
   187.9  
  187.10  (*To let us treat division as multiplication*)
  187.11  val divide_simps = [@{thm divide_inverse}, @{thm inverse_mult_distrib}, @{thm inverse_inverse_eq}];
  187.12 @@ -719,7 +719,7 @@
  187.13             @{thm "times_divide_eq_left"}, @{thm "times_divide_eq_right"},
  187.14             @{thm "times_divide_times_eq"},
  187.15             @{thm "divide_divide_eq_right"},
  187.16 -           @{thm "diff_minus"}, @{thm "minus_divide_left"},
  187.17 +           @{thm diff_conv_add_uminus}, @{thm "minus_divide_left"},
  187.18             @{thm "add_divide_distrib"} RS sym,
  187.19             @{thm field_divide_inverse} RS sym, @{thm inverse_divide}, 
  187.20             Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (Conv.rewr_conv (mk_meta_eq @{thm mult_commute}))))   
   188.1 --- a/src/HOL/Tools/semiring_normalizer.ML	Mon Nov 11 17:34:44 2013 +0100
   188.2 +++ b/src/HOL/Tools/semiring_normalizer.ML	Mon Nov 11 17:44:21 2013 +0100
   188.3 @@ -848,7 +848,7 @@
   188.4  val nat_exp_ss =
   188.5    simpset_of
   188.6     (put_simpset HOL_basic_ss @{context}
   188.7 -    addsimps (@{thms eval_nat_numeral} @ @{thms nat_arith} @ @{thms arith_simps} @ @{thms rel_simps})
   188.8 +    addsimps (@{thms eval_nat_numeral} @ @{thms diff_nat_numeral} @ @{thms arith_simps} @ @{thms rel_simps})
   188.9      addsimps [@{thm Let_def}, @{thm if_False}, @{thm if_True}, @{thm Nat.add_0}, @{thm add_Suc}]);
  188.10  
  188.11  fun simple_cterm_ord t u = Term_Ord.term_ord (term_of t, term_of u) = LESS;
   189.1 --- a/src/HOL/Tools/try0.ML	Mon Nov 11 17:34:44 2013 +0100
   189.2 +++ b/src/HOL/Tools/try0.ML	Mon Nov 11 17:44:21 2013 +0100
   189.3 @@ -41,10 +41,10 @@
   189.4    end
   189.5    handle TimeLimit.TimeOut => false
   189.6  
   189.7 -fun do_generic timeout_opt command pre post apply st =
   189.8 +fun do_generic timeout_opt name command pre post apply st =
   189.9    let val timer = Timer.startRealTimer () in
  189.10      if can_apply timeout_opt pre post apply st then
  189.11 -      SOME (command, Time.toMilliseconds (Timer.checkRealTimer timer))
  189.12 +      SOME (name, command, Time.toMilliseconds (Timer.checkRealTimer timer))
  189.13      else
  189.14        NONE
  189.15    end
  189.16 @@ -75,16 +75,11 @@
  189.17                      timeout_opt quad st =
  189.18    if mode <> Auto_Try orelse run_if_auto_try then
  189.19      let val attrs = attrs_text attrs quad in
  189.20 -      do_generic timeout_opt
  189.21 -                 (name ^ attrs ^
  189.22 -                  (if all_goals andalso
  189.23 -                      nprems_of (#goal (Proof.goal st)) > 1 then
  189.24 -                     " [1]"
  189.25 -                   else
  189.26 -                     ""))
  189.27 -                 I (#goal o Proof.goal)
  189.28 -                 (apply_named_method_on_first_goal (name ^ attrs)
  189.29 -                                                   (Proof.theory_of st)) st
  189.30 +      do_generic timeout_opt name
  189.31 +        ((name ^ attrs |> attrs <> "" ? enclose "(" ")") ^
  189.32 +         (if all_goals andalso nprems_of (#goal (Proof.goal st)) > 1 then "[1]" else ""))
  189.33 +        I (#goal o Proof.goal)
  189.34 +        (apply_named_method_on_first_goal (name ^ attrs) (Proof.theory_of st)) st
  189.35      end
  189.36    else
  189.37      NONE
  189.38 @@ -108,12 +103,16 @@
  189.39     ("presburger", ((false, true), no_attrs))]
  189.40  val do_methods = map do_named_method named_methods
  189.41  
  189.42 -fun time_string (s, ms) = s ^ ": " ^ string_of_int ms ^ " ms"
  189.43 +fun time_string ms = string_of_int ms ^ " ms"
  189.44 +fun tool_time_string (s, ms) = s ^ ": " ^ time_string ms
  189.45  
  189.46  fun do_try0 mode timeout_opt quad st =
  189.47    let
  189.48      val st = st |> Proof.map_context (Config.put Metis_Tactic.verbose false #>
  189.49        Config.put Lin_Arith.verbose false)
  189.50 +    fun par_map f =
  189.51 +      if mode = Normal then Par_List.map f #> map_filter I #> sort (int_ord o pairself #3)
  189.52 +      else Par_List.get_some f #> the_list
  189.53    in
  189.54      if mode = Normal then
  189.55        "Trying " ^ space_implode " " (Try.serial_commas "and"
  189.56 @@ -121,17 +120,15 @@
  189.57        |> Output.urgent_message
  189.58      else
  189.59        ();
  189.60 -    case do_methods |> Par_List.map (fn f => f mode timeout_opt quad st)
  189.61 -                    |> map_filter I |> sort (int_ord o pairself snd) of
  189.62 +    (case par_map (fn f => f mode timeout_opt quad st) do_methods of
  189.63        [] =>
  189.64        (if mode = Normal then Output.urgent_message "No proof found." else ();
  189.65         (false, (noneN, st)))
  189.66 -    | xs as (s, _) :: _ =>
  189.67 +    | xs as (name, command, _) :: _ =>
  189.68        let
  189.69 -        val xs = xs |> map (fn (s, n) => (n, hd (space_explode " " s)))
  189.70 +        val xs = xs |> map (fn (name, _, n) => (n, name))
  189.71                      |> AList.coalesce (op =)
  189.72                      |> map (swap o apsnd commas)
  189.73 -        val need_parens = exists_string (curry (op =) " ") s
  189.74          val message =
  189.75            (case mode of
  189.76               Auto_Try => "Auto Try0 found a proof"
  189.77 @@ -139,16 +136,18 @@
  189.78             | Normal => "Try this") ^ ": " ^
  189.79            Active.sendback_markup [Markup.padding_command]
  189.80                ((if nprems_of (#goal (Proof.goal st)) = 1 then "by"
  189.81 -                else "apply") ^ " " ^ (s |> need_parens ? enclose "(" ")")) ^
  189.82 -          "\n(" ^ space_implode "; " (map time_string xs) ^ ")."
  189.83 +                else "apply") ^ " " ^ command) ^
  189.84 +          (case xs of
  189.85 +            [(_, ms)] => " (" ^ time_string ms ^ ")."
  189.86 +          | xs => "\n(" ^ space_implode "; " (map tool_time_string xs) ^ ").")
  189.87        in
  189.88 -        (true, (s, st |> (if mode = Auto_Try then
  189.89 -                            Proof.goal_message
  189.90 -                                (fn () => Pretty.markup Markup.information
  189.91 -                                                        [Pretty.str message])
  189.92 -                          else
  189.93 -                            tap (fn _ => Output.urgent_message message))))
  189.94 -      end
  189.95 +        (true, (name,
  189.96 +           st |> (if mode = Auto_Try then
  189.97 +                    Proof.goal_message
  189.98 +                      (fn () => Pretty.markup Markup.information [Pretty.str message])
  189.99 +                  else
 189.100 +                    tap (fn _ => Output.urgent_message message))))
 189.101 +      end)
 189.102    end
 189.103  
 189.104  fun try0 timeout_opt = fst oo do_try0 Normal timeout_opt
   190.1 --- a/src/HOL/Topological_Spaces.thy	Mon Nov 11 17:34:44 2013 +0100
   190.2 +++ b/src/HOL/Topological_Spaces.thy	Mon Nov 11 17:44:21 2013 +0100
   190.3 @@ -2112,7 +2112,7 @@
   190.4    with dense[of b "Inf A"] obtain c where "c < Inf A" "c \<in> A"
   190.5      by (auto simp: subset_eq)
   190.6    then show False
   190.7 -    using cInf_lower[OF `c \<in> A`, of x] bnd by (metis less_imp_le not_le)
   190.8 +    using cInf_lower[OF `c \<in> A`] bnd by (metis not_le less_imp_le bdd_belowI)
   190.9  qed
  190.10  
  190.11  lemma Sup_notin_open:
  190.12 @@ -2125,7 +2125,7 @@
  190.13    with dense[of "Sup A" b] obtain c where "Sup A < c" "c \<in> A"
  190.14      by (auto simp: subset_eq)
  190.15    then show False
  190.16 -    using cSup_upper[OF `c \<in> A`, of x] bnd by (metis less_imp_le not_le)
  190.17 +    using cSup_upper[OF `c \<in> A`] bnd by (metis less_imp_le not_le bdd_aboveI)
  190.18  qed
  190.19  
  190.20  end
  190.21 @@ -2151,7 +2151,7 @@
  190.22      let ?z = "Inf (B \<inter> {x <..})"
  190.23  
  190.24      have "x \<le> ?z" "?z \<le> y"
  190.25 -      using `y \<in> B` `x < y` by (auto intro: cInf_lower[where z=x] cInf_greatest)
  190.26 +      using `y \<in> B` `x < y` by (auto intro: cInf_lower cInf_greatest)
  190.27      with `x \<in> U` `y \<in> U` have "?z \<in> U"
  190.28        by (rule *)
  190.29      moreover have "?z \<notin> B \<inter> {x <..}"
  190.30 @@ -2163,11 +2163,11 @@
  190.31        obtain a where "?z < a" "{?z ..< a} \<subseteq> A"
  190.32          using open_right[OF `open A` `?z \<in> A` `?z < y`] by auto
  190.33        moreover obtain b where "b \<in> B" "x < b" "b < min a y"
  190.34 -        using cInf_less_iff[of "B \<inter> {x <..}" x "min a y"] `?z < a` `?z < y` `x < y` `y \<in> B`
  190.35 +        using cInf_less_iff[of "B \<inter> {x <..}" "min a y"] `?z < a` `?z < y` `x < y` `y \<in> B`
  190.36          by (auto intro: less_imp_le)
  190.37        moreover have "?z \<le> b"
  190.38          using `b \<in> B` `x < b`
  190.39 -        by (intro cInf_lower[where z=x]) auto
  190.40 +        by (intro cInf_lower) auto
  190.41        moreover have "b \<in> U"
  190.42          using `x \<le> ?z` `?z \<le> b` `b < min a y`
  190.43          by (intro *[OF `x \<in> U` `y \<in> U`]) (auto simp: less_imp_le)
   191.1 --- a/src/HOL/Transcendental.thy	Mon Nov 11 17:34:44 2013 +0100
   191.2 +++ b/src/HOL/Transcendental.thy	Mon Nov 11 17:44:21 2013 +0100
   191.3 @@ -453,7 +453,7 @@
   191.4    apply simp
   191.5    apply (simp only: lemma_termdiff1 setsum_right_distrib)
   191.6    apply (rule setsum_cong [OF refl])
   191.7 -  apply (simp add: diff_minus [symmetric] less_iff_Suc_add)
   191.8 +  apply (simp add: less_iff_Suc_add)
   191.9    apply (clarify)
  191.10    apply (simp add: setsum_right_distrib lemma_realpow_diff_sumr2 mult_ac
  191.11                del: setsum_op_ivl_Suc power_Suc)
  191.12 @@ -1129,8 +1129,7 @@
  191.13    by (rule inverse_unique [symmetric], simp add: mult_exp_exp)
  191.14  
  191.15  lemma exp_diff: "exp (x - y) = exp x / exp y"
  191.16 -  unfolding diff_minus divide_inverse
  191.17 -  by (simp add: exp_add exp_minus)
  191.18 +  using exp_add [of x "- y"] by (simp add: exp_minus divide_inverse)
  191.19  
  191.20  
  191.21  subsubsection {* Properties of the Exponential Function on Reals *}
  191.22 @@ -2410,13 +2409,13 @@
  191.23    using sin_cos_minus_lemma [where x=x] by simp
  191.24  
  191.25  lemma sin_diff: "sin (x - y) = sin x * cos y - cos x * sin y"
  191.26 -  by (simp add: diff_minus sin_add)
  191.27 +  using sin_add [of x "- y"] by simp
  191.28  
  191.29  lemma sin_diff2: "sin (x - y) = cos y * sin x - sin y * cos x"
  191.30    by (simp add: sin_diff mult_commute)
  191.31  
  191.32  lemma cos_diff: "cos (x - y) = cos x * cos y + sin x * sin y"
  191.33 -  by (simp add: diff_minus cos_add)
  191.34 +  using cos_add [of x "- y"] by simp
  191.35  
  191.36  lemma cos_diff2: "cos (x - y) = cos y * cos x + sin y * sin x"
  191.37    by (simp add: cos_diff mult_commute)
  191.38 @@ -2526,8 +2525,9 @@
  191.39          by (simp add: inverse_eq_divide less_divide_eq)
  191.40      }
  191.41      note *** = this
  191.42 +    have [simp]: "\<And>x y::real. 0 < x - y \<longleftrightarrow> y < x" by arith
  191.43      from ** show ?thesis by (rule sumr_pos_lt_pair)
  191.44 -      (simp add: divide_inverse real_0_less_add_iff mult_assoc [symmetric] ***)
  191.45 +      (simp add: divide_inverse mult_assoc [symmetric] ***)
  191.46    qed
  191.47    ultimately have "0 < (\<Sum>n. - (-1 ^ n * 2 ^ (2 * n) / real (fact (2 * n))))"
  191.48      by (rule order_less_trans)
  191.49 @@ -2810,7 +2810,7 @@
  191.50  apply (cut_tac x="-x" in cos_zero_lemma, simp, simp)
  191.51  apply (force simp add: minus_equation_iff [of x])
  191.52  apply (auto simp only: odd_Suc_mult_two_ex real_of_nat_Suc distrib_right)
  191.53 -apply (auto simp add: cos_add)
  191.54 +apply (auto simp add: cos_diff cos_add)
  191.55  done
  191.56  
  191.57  (* ditto: but to a lesser extent *)
  191.58 @@ -3833,7 +3833,7 @@
  191.59                by (rule DERIV_arctan_suminf[OF `0 < r` `r < 1` `\<bar>x\<bar> < r`])
  191.60              from DERIV_add_minus[OF this DERIV_arctan]
  191.61              show "DERIV (\<lambda> x. suminf (?c x) - arctan x) x :> 0"
  191.62 -              unfolding diff_minus by auto
  191.63 +              by auto
  191.64            qed
  191.65            hence DERIV_in_rball: "\<forall> y. a \<le> y \<and> y \<le> b \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) y :> 0"
  191.66              using `-r < a` `b < r` by auto
  191.67 @@ -3922,9 +3922,10 @@
  191.68        }
  191.69        hence "\<forall> x \<in> { 0 <..< 1 }. 0 \<le> ?a x n - ?diff x n" by auto
  191.70        moreover have "\<And>x. isCont (\<lambda> x. ?a x n - ?diff x n) x"
  191.71 -        unfolding diff_minus divide_inverse
  191.72 +        unfolding diff_conv_add_uminus divide_inverse
  191.73          by (auto intro!: isCont_add isCont_rabs isCont_ident isCont_minus isCont_arctan
  191.74 -          isCont_inverse isCont_mult isCont_power isCont_const isCont_setsum)
  191.75 +          isCont_inverse isCont_mult isCont_power isCont_const isCont_setsum
  191.76 +          simp del: add_uminus_conv_diff)
  191.77        ultimately have "0 \<le> ?a 1 n - ?diff 1 n"
  191.78          by (rule LIM_less_bound)
  191.79        hence "?diff 1 n \<le> ?a 1 n" by auto
   192.1 --- a/src/HOL/Wellfounded.thy	Mon Nov 11 17:34:44 2013 +0100
   192.2 +++ b/src/HOL/Wellfounded.thy	Mon Nov 11 17:44:21 2013 +0100
   192.3 @@ -482,6 +482,11 @@
   192.4  
   192.5  lemmas accpI = accp.accI
   192.6  
   192.7 +lemma accp_eq_acc [code]:
   192.8 +  "accp r = (\<lambda>x. x \<in> Wellfounded.acc {(x, y). r x y})"
   192.9 +  by (simp add: acc_def)
  192.10 +
  192.11 +
  192.12  text {* Induction rules *}
  192.13  
  192.14  theorem accp_induct:
  192.15 @@ -855,4 +860,7 @@
  192.16  
  192.17  declare "prod.size" [no_atp]
  192.18  
  192.19 +
  192.20 +hide_const (open) acc accp
  192.21 +
  192.22  end
   193.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
   193.2 +++ b/src/HOL/Word/Bit_Bit.thy	Mon Nov 11 17:44:21 2013 +0100
   193.3 @@ -0,0 +1,73 @@
   193.4 +(*  Title:      HOL/Word/Bit_Bit.thy
   193.5 +    Author:     Author: Brian Huffman, PSU and Gerwin Klein, NICTA
   193.6 +*)
   193.7 +
   193.8 +header {* Bit operations in \<Z>\<^sub>2*}
   193.9 +
  193.10 +theory Bit_Bit
  193.11 +imports Bit_Operations "~~/src/HOL/Library/Bit"
  193.12 +begin
  193.13 +
  193.14 +instantiation bit :: bit
  193.15 +begin
  193.16 +
  193.17 +primrec bitNOT_bit where
  193.18 +  "NOT 0 = (1::bit)"
  193.19 +  | "NOT 1 = (0::bit)"
  193.20 +
  193.21 +primrec bitAND_bit where
  193.22 +  "0 AND y = (0::bit)"
  193.23 +  | "1 AND y = (y::bit)"
  193.24 +
  193.25 +primrec bitOR_bit where
  193.26 +  "0 OR y = (y::bit)"
  193.27 +  | "1 OR y = (1::bit)"
  193.28 +
  193.29 +primrec bitXOR_bit where
  193.30 +  "0 XOR y = (y::bit)"
  193.31 +  | "1 XOR y = (NOT y :: bit)"
  193.32 +
  193.33 +instance  ..
  193.34 +
  193.35 +end
  193.36 +
  193.37 +lemmas bit_simps =
  193.38 +  bitNOT_bit.simps bitAND_bit.simps bitOR_bit.simps bitXOR_bit.simps
  193.39 +
  193.40 +lemma bit_extra_simps [simp]: 
  193.41 +  "x AND 0 = (0::bit)"
  193.42 +  "x AND 1 = (x::bit)"
  193.43 +  "x OR 1 = (1::bit)"
  193.44 +  "x OR 0 = (x::bit)"
  193.45 +  "x XOR 1 = NOT (x::bit)"
  193.46 +  "x XOR 0 = (x::bit)"
  193.47 +  by (cases x, auto)+
  193.48 +
  193.49 +lemma bit_ops_comm: 
  193.50 +  "(x::bit) AND y = y AND x"
  193.51 +  "(x::bit) OR y = y OR x"
  193.52 +  "(x::bit) XOR y = y XOR x"
  193.53 +  by (cases y, auto)+
  193.54 +
  193.55 +lemma bit_ops_same [simp]: 
  193.56 +  "(x::bit) AND x = x"
  193.57 +  "(x::bit) OR x = x"
  193.58 +  "(x::bit) XOR x = 0"
  193.59 +  by (cases x, auto)+
  193.60 +
  193.61 +lemma bit_not_not [simp]: "NOT (NOT (x::bit)) = x"
  193.62 +  by (cases x) auto
  193.63 +
  193.64 +lemma bit_or_def: "(b::bit) OR c = NOT (NOT b AND NOT c)"
  193.65 +  by (induct b, simp_all)
  193.66 +
  193.67 +lemma bit_xor_def: "(b::bit) XOR c = (b AND NOT c) OR (NOT b AND c)"
  193.68 +  by (induct b, simp_all)
  193.69 +
  193.70 +lemma bit_NOT_eq_1_iff [simp]: "NOT (b::bit) = 1 \<longleftrightarrow> b = 0"
  193.71 +  by (induct b, simp_all)
  193.72 +
  193.73 +lemma bit_AND_eq_1_iff [simp]: "(a::bit) AND b = 1 \<longleftrightarrow> a = 1 \<and> b = 1"
  193.74 +  by (induct a, simp_all)
  193.75 +
  193.76 +end
   194.1 --- a/src/HOL/Word/Bit_Int.thy	Mon Nov 11 17:34:44 2013 +0100
   194.2 +++ b/src/HOL/Word/Bit_Int.thy	Mon Nov 11 17:44:21 2013 +0100
   194.3 @@ -9,7 +9,7 @@
   194.4  header {* Bitwise Operations on Binary Integers *}
   194.5  
   194.6  theory Bit_Int
   194.7 -imports Bit_Representation Bit_Operations
   194.8 +imports Bit_Representation Bit_Bit
   194.9  begin
  194.10  
  194.11  subsection {* Logical operations *}
   195.1 --- a/src/HOL/Word/Bit_Operations.thy	Mon Nov 11 17:34:44 2013 +0100
   195.2 +++ b/src/HOL/Word/Bit_Operations.thy	Mon Nov 11 17:44:21 2013 +0100
   195.3 @@ -8,8 +8,6 @@
   195.4  imports "~~/src/HOL/Library/Bit"
   195.5  begin
   195.6  
   195.7 -subsection {* Abstract syntactic bit operations *}
   195.8 -
   195.9  class bit =
  195.10    fixes bitNOT :: "'a \<Rightarrow> 'a"       ("NOT _" [70] 71)
  195.11      and bitAND :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "AND" 64)
  195.12 @@ -37,69 +35,5 @@
  195.13  class bitss = bits +
  195.14    fixes msb      :: "'a \<Rightarrow> bool"
  195.15  
  195.16 -  
  195.17 -subsection {* Bitwise operations on @{typ bit} *}
  195.18 -
  195.19 -instantiation bit :: bit
  195.20 -begin
  195.21 -
  195.22 -primrec bitNOT_bit where
  195.23 -  "NOT 0 = (1::bit)"
  195.24 -  | "NOT 1 = (0::bit)"
  195.25 -
  195.26 -primrec bitAND_bit where
  195.27 -  "0 AND y = (0::bit)"
  195.28 -  | "1 AND y = (y::bit)"
  195.29 -
  195.30 -primrec bitOR_bit where
  195.31 -  "0 OR y = (y::bit)"
  195.32 -  | "1 OR y = (1::bit)"
  195.33 -
  195.34 -primrec bitXOR_bit where
  195.35 -  "0 XOR y = (y::bit)"
  195.36 -  | "1 XOR y = (NOT y :: bit)"
  195.37 -
  195.38 -instance  ..
  195.39 -
  195.40  end
  195.41  
  195.42 -lemmas bit_simps =
  195.43 -  bitNOT_bit.simps bitAND_bit.simps bitOR_bit.simps bitXOR_bit.simps
  195.44 -
  195.45 -lemma bit_extra_simps [simp]: 
  195.46 -  "x AND 0 = (0::bit)"
  195.47 -  "x AND 1 = (x::bit)"
  195.48 -  "x OR 1 = (1::bit)"
  195.49 -  "x OR 0 = (x::bit)"
  195.50 -  "x XOR 1 = NOT (x::bit)"
  195.51 -  "x XOR 0 = (x::bit)"
  195.52 -  by (cases x, auto)+
  195.53 -
  195.54 -lemma bit_ops_comm: 
  195.55 -  "(x::bit) AND y = y AND x"
  195.56 -  "(x::bit) OR y = y OR x"
  195.57 -  "(x::bit) XOR y = y XOR x"
  195.58 -  by (cases y, auto)+
  195.59 -
  195.60 -lemma bit_ops_same [simp]: 
  195.61 -  "(x::bit) AND x = x"
  195.62 -  "(x::bit) OR x = x"
  195.63 -  "(x::bit) XOR x = 0"
  195.64 -  by (cases x, auto)+
  195.65 -
  195.66 -lemma bit_not_not [simp]: "NOT (NOT (x::bit)) = x"
  195.67 -  by (cases x) auto
  195.68 -
  195.69 -lemma bit_or_def: "(b::bit) OR c = NOT (NOT b AND NOT c)"
  195.70 -  by (induct b, simp_all)
  195.71 -
  195.72 -lemma bit_xor_def: "(b::bit) XOR c = (b AND NOT c) OR (NOT b AND c)"
  195.73 -  by (induct b, simp_all)
  195.74 -
  195.75 -lemma bit_NOT_eq_1_iff [simp]: "NOT (b::bit) = 1 \<longleftrightarrow> b = 0"
  195.76 -  by (induct b, simp_all)
  195.77 -
  195.78 -lemma bit_AND_eq_1_iff [simp]: "(a::bit) AND b = 1 \<longleftrightarrow> a = 1 \<and> b = 1"
  195.79 -  by (induct a, simp_all)
  195.80 -
  195.81 -end
   196.1 --- a/src/HOL/Word/Bit_Representation.thy	Mon Nov 11 17:34:44 2013 +0100
   196.2 +++ b/src/HOL/Word/Bit_Representation.thy	Mon Nov 11 17:44:21 2013 +0100
   196.3 @@ -636,12 +636,12 @@
   196.4    unfolding no_sbintr_alt2 by (drule sb_inc_lem') simp
   196.5  
   196.6  lemma sb_dec_lem:
   196.7 -  "(0::int) <= - (2^k) + a ==> (a + 2^k) mod (2 * 2 ^ k) <= - (2 ^ k) + a"
   196.8 -  by (rule int_mod_le' [where n = "2 ^ (Suc k)" and b = "a + 2 ^ k", simplified])
   196.9 +  "(0::int) \<le> - (2 ^ k) + a \<Longrightarrow> (a + 2 ^ k) mod (2 * 2 ^ k) \<le> - (2 ^ k) + a"
  196.10 +  using int_mod_le'[where n = "2 ^ (Suc k)" and b = "a + 2 ^ k"] by simp
  196.11  
  196.12  lemma sb_dec_lem':
  196.13 -  "(2::int) ^ k <= a ==> (a + 2 ^ k) mod (2 * 2 ^ k) <= - (2 ^ k) + a"
  196.14 -  by (rule iffD1 [OF diff_le_eq', THEN sb_dec_lem, simplified])
  196.15 +  "(2::int) ^ k \<le> a \<Longrightarrow> (a + 2 ^ k) mod (2 * 2 ^ k) \<le> - (2 ^ k) + a"
  196.16 +  by (rule sb_dec_lem) simp
  196.17  
  196.18  lemma sbintrunc_dec:
  196.19    "x >= (2 ^ n) ==> x - 2 ^ (Suc n) >= sbintrunc n x"
   197.1 --- a/src/HOL/Word/Misc_Numeric.thy	Mon Nov 11 17:34:44 2013 +0100
   197.2 +++ b/src/HOL/Word/Misc_Numeric.thy	Mon Nov 11 17:44:21 2013 +0100
   197.3 @@ -8,10 +8,6 @@
   197.4  imports Main Parity
   197.5  begin
   197.6  
   197.7 -lemma zmod_zsub_self [simp]: (* FIXME move to Divides.thy *) 
   197.8 -  "((b :: int) - a) mod a = b mod a"
   197.9 -  by (simp add: mod_diff_right_eq)
  197.10 -
  197.11  declare iszero_0 [iff]
  197.12  
  197.13  lemma min_pm [simp]: "min a b + (a - b) = (a :: nat)" by arith
   198.1 --- a/src/HOL/Word/Word.thy	Mon Nov 11 17:34:44 2013 +0100
   198.2 +++ b/src/HOL/Word/Word.thy	Mon Nov 11 17:44:21 2013 +0100
   198.3 @@ -8,6 +8,7 @@
   198.4  imports
   198.5    Type_Length
   198.6    "~~/src/HOL/Library/Boolean_Algebra"
   198.7 +  Bit_Bit
   198.8    Bool_List_Representation
   198.9    Misc_Typedef
  198.10    Word_Miscellaneous
  198.11 @@ -505,10 +506,6 @@
  198.12  definition max_word :: "'a::len word" -- "Largest representable machine integer." where
  198.13    "max_word = word_of_int (2 ^ len_of TYPE('a) - 1)"
  198.14  
  198.15 -primrec of_bool :: "bool \<Rightarrow> 'a::len word" where
  198.16 -  "of_bool False = 0"
  198.17 -| "of_bool True = 1"
  198.18 -
  198.19  (* FIXME: only provide one theorem name *)
  198.20  lemmas of_nth_def = word_set_bits_def
  198.21  
  198.22 @@ -4238,7 +4235,7 @@
  198.23  
  198.24  lemma max_word_max [simp,intro!]: "n \<le> max_word"
  198.25    by (cases n rule: word_int_cases)
  198.26 -     (simp add: max_word_def word_le_def int_word_uint int_mod_eq')
  198.27 +    (simp add: max_word_def word_le_def int_word_uint int_mod_eq' del: minus_mod_self1)
  198.28    
  198.29  lemma word_of_int_2p_len: "word_of_int (2 ^ len_of TYPE('a)) = (0::'a::len0 word)"
  198.30    by (subst word_uint.Abs_norm [symmetric]) simp
   199.1 --- a/src/HOL/Word/WordBitwise.thy	Mon Nov 11 17:34:44 2013 +0100
   199.2 +++ b/src/HOL/Word/WordBitwise.thy	Mon Nov 11 17:44:21 2013 +0100
   199.3 @@ -65,7 +65,7 @@
   199.4  
   199.5  lemma bl_word_sub:
   199.6    "to_bl (x - y) = to_bl (x + (- y))"
   199.7 -  by (simp add: diff_def)
   199.8 +  by simp
   199.9  
  199.10  lemma rbl_word_1:
  199.11    "rev (to_bl (1 :: ('a :: len0) word))
   200.1 --- a/src/HOL/ex/Dedekind_Real.thy	Mon Nov 11 17:34:44 2013 +0100
   200.2 +++ b/src/HOL/ex/Dedekind_Real.thy	Mon Nov 11 17:44:21 2013 +0100
   200.3 @@ -1506,7 +1506,6 @@
   200.4  instance real :: linorder
   200.5    by (intro_classes, rule real_le_linear)
   200.6  
   200.7 -
   200.8  lemma real_le_eq_diff: "(x \<le> y) = (x-y \<le> (0::real))"
   200.9  apply (cases x, cases y) 
  200.10  apply (auto simp add: real_le real_zero_def real_diff_def real_add real_minus
  200.11 @@ -1520,14 +1519,14 @@
  200.12    have "z + x - (z + y) = (z + -z) + (x - y)" 
  200.13      by (simp add: algebra_simps) 
  200.14    with le show ?thesis 
  200.15 -    by (simp add: real_le_eq_diff[of x] real_le_eq_diff[of "z+x"] diff_minus)
  200.16 +    by (simp add: real_le_eq_diff[of x] real_le_eq_diff[of "z+x"])
  200.17  qed
  200.18  
  200.19  lemma real_sum_gt_zero_less: "(0 < S + (-W::real)) ==> (W < S)"
  200.20 -by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus)
  200.21 +by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S])
  200.22  
  200.23  lemma real_less_sum_gt_zero: "(W < S) ==> (0 < S + (-W::real))"
  200.24 -by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus)
  200.25 +by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S])
  200.26  
  200.27  lemma real_mult_order: "[| 0 < x; 0 < y |] ==> (0::real) < x * y"
  200.28  apply (cases x, cases y)
  200.29 @@ -1543,7 +1542,7 @@
  200.30  apply (rule real_sum_gt_zero_less)
  200.31  apply (drule real_less_sum_gt_zero [of x y])
  200.32  apply (drule real_mult_order, assumption)
  200.33 -apply (simp add: distrib_left)
  200.34 +apply (simp add: algebra_simps)
  200.35  done
  200.36  
  200.37  instantiation real :: distrib_lattice
  200.38 @@ -1657,7 +1656,6 @@
  200.39  lemma real_less_all_real2: "~ 0 < y ==> \<forall>x. y < real_of_preal x"
  200.40  by (blast intro!: real_less_all_preal linorder_not_less [THEN iffD1])
  200.41  
  200.42 -
  200.43  subsection {* Completeness of Positive Reals *}
  200.44  
  200.45  text {*
  200.46 @@ -1759,107 +1757,23 @@
  200.47  qed
  200.48  
  200.49  text {*
  200.50 -  \medskip Completeness properties using @{text "isUb"}, @{text "isLub"} etc.
  200.51 -*}
  200.52 -
  200.53 -lemma posreals_complete:
  200.54 -  assumes positive_S: "\<forall>x \<in> S. 0 < x"
  200.55 -    and not_empty_S: "\<exists>x. x \<in> S"
  200.56 -    and upper_bound_Ex: "\<exists>u. isUb (UNIV::real set) S u"
  200.57 -  shows "\<exists>t. isLub (UNIV::real set) S t"
  200.58 -proof
  200.59 -  let ?pS = "{w. real_of_preal w \<in> S}"
  200.60 -
  200.61 -  obtain u where "isUb UNIV S u" using upper_bound_Ex ..
  200.62 -  hence sup: "\<forall>x \<in> S. x \<le> u" by (simp add: isUb_def setle_def)
  200.63 -
  200.64 -  obtain x where x_in_S: "x \<in> S" using not_empty_S ..
  200.65 -  hence x_gt_zero: "0 < x" using positive_S by simp
  200.66 -  have  "x \<le> u" using sup and x_in_S ..
  200.67 -  hence "0 < u" using x_gt_zero by arith
  200.68 -
  200.69 -  then obtain pu where u_is_pu: "u = real_of_preal pu"
  200.70 -    by (auto simp add: real_gt_zero_preal_Ex)
  200.71 -
  200.72 -  have pS_less_pu: "\<forall>pa \<in> ?pS. pa \<le> pu"
  200.73 -  proof
  200.74 -    fix pa
  200.75 -    assume "pa \<in> ?pS"
  200.76 -    then obtain a where a: "a \<in> S" "a = real_of_preal pa"
  200.77 -      by simp
  200.78 -    then have "a \<le> u" using sup by simp
  200.79 -    with a show "pa \<le> pu"
  200.80 -      using sup and u_is_pu by (simp add: real_of_preal_le_iff)
  200.81 -  qed
  200.82 -
  200.83 -  have "\<forall>y \<in> S. y \<le> real_of_preal (psup ?pS)"
  200.84 -  proof
  200.85 -    fix y
  200.86 -    assume y_in_S: "y \<in> S"
  200.87 -    hence "0 < y" using positive_S by simp
  200.88 -    then obtain py where y_is_py: "y = real_of_preal py"
  200.89 -      by (auto simp add: real_gt_zero_preal_Ex)
  200.90 -    hence py_in_pS: "py \<in> ?pS" using y_in_S by simp
  200.91 -    with pS_less_pu have "py \<le> psup ?pS"
  200.92 -      by (rule preal_psup_le)
  200.93 -    thus "y \<le> real_of_preal (psup ?pS)"
  200.94 -      using y_is_py by (simp add: real_of_preal_le_iff)
  200.95 -  qed
  200.96 -
  200.97 -  moreover {
  200.98 -    fix x
  200.99 -    assume x_ub_S: "\<forall>y\<in>S. y \<le> x"
 200.100 -    have "real_of_preal (psup ?pS) \<le> x"
 200.101 -    proof -
 200.102 -      obtain "s" where s_in_S: "s \<in> S" using not_empty_S ..
 200.103 -      hence s_pos: "0 < s" using positive_S by simp
 200.104 -
 200.105 -      hence "\<exists> ps. s = real_of_preal ps" by (simp add: real_gt_zero_preal_Ex)
 200.106 -      then obtain "ps" where s_is_ps: "s = real_of_preal ps" ..
 200.107 -      hence ps_in_pS: "ps \<in> {w. real_of_preal w \<in> S}" using s_in_S by simp
 200.108 -
 200.109 -      from x_ub_S have "s \<le> x" using s_in_S ..
 200.110 -      hence "0 < x" using s_pos by simp
 200.111 -      hence "\<exists> px. x = real_of_preal px" by (simp add: real_gt_zero_preal_Ex)
 200.112 -      then obtain "px" where x_is_px: "x = real_of_preal px" ..
 200.113 -
 200.114 -      have "\<forall>pe \<in> ?pS. pe \<le> px"
 200.115 -      proof
 200.116 -        fix pe
 200.117 -        assume "pe \<in> ?pS"
 200.118 -        hence "real_of_preal pe \<in> S" by simp
 200.119 -        hence "real_of_preal pe \<le> x" using x_ub_S by simp
 200.120 -        thus "pe \<le> px" using x_is_px by (simp add: real_of_preal_le_iff)
 200.121 -      qed
 200.122 -
 200.123 -      moreover have "?pS \<noteq> {}" using ps_in_pS by auto
 200.124 -      ultimately have "(psup ?pS) \<le> px" by (simp add: psup_le_ub)
 200.125 -      thus "real_of_preal (psup ?pS) \<le> x" using x_is_px by (simp add: real_of_preal_le_iff)
 200.126 -    qed
 200.127 -  }
 200.128 -  ultimately show "isLub UNIV S (real_of_preal (psup ?pS))"
 200.129 -    by (simp add: isLub_def leastP_def isUb_def setle_def setge_def)
 200.130 -qed
 200.131 -
 200.132 -text {*
 200.133 -  \medskip reals Completeness (again!)
 200.134 +  \medskip Completeness
 200.135  *}
 200.136  
 200.137  lemma reals_complete:
 200.138 +  fixes S :: "real set"
 200.139    assumes notempty_S: "\<exists>X. X \<in> S"
 200.140 -    and exists_Ub: "\<exists>Y. isUb (UNIV::real set) S Y"
 200.141 -  shows "\<exists>t. isLub (UNIV :: real set) S t"
 200.142 +    and exists_Ub: "bdd_above S"
 200.143 +  shows "\<exists>x. (\<forall>s\<in>S. s \<le> x) \<and> (\<forall>y. (\<forall>s\<in>S. s \<le> y) \<longrightarrow> x \<le> y)"
 200.144  proof -
 200.145    obtain X where X_in_S: "X \<in> S" using notempty_S ..
 200.146 -  obtain Y where Y_isUb: "isUb (UNIV::real set) S Y"
 200.147 -    using exists_Ub ..
 200.148 +  obtain Y where Y_isUb: "\<forall>s\<in>S. s \<le> Y"
 200.149 +    using exists_Ub by (auto simp: bdd_above_def)
 200.150    let ?SHIFT = "{z. \<exists>x \<in>S. z = x + (-X) + 1} \<inter> {x. 0 < x}"
 200.151  
 200.152    {
 200.153      fix x
 200.154 -    assume "isUb (UNIV::real set) S x"
 200.155 -    hence S_le_x: "\<forall> y \<in> S. y <= x"
 200.156 -      by (simp add: isUb_def setle_def)
 200.157 +    assume S_le_x: "\<forall>s\<in>S. s \<le> x"
 200.158      {
 200.159        fix s
 200.160        assume "s \<in> {z. \<exists>x\<in>S. z = x + - X + 1}"
 200.161 @@ -1868,86 +1782,74 @@
 200.162        then have "x1 \<le> x" using S_le_x by simp
 200.163        with x1 have "s \<le> x + - X + 1" by arith
 200.164      }
 200.165 -    then have "isUb (UNIV::real set) ?SHIFT (x + (-X) + 1)"
 200.166 -      by (auto simp add: isUb_def setle_def)
 200.167 +    then have "\<forall>s\<in>?SHIFT. s \<le> x + (-X) + 1"
 200.168 +      by auto
 200.169    } note S_Ub_is_SHIFT_Ub = this
 200.170  
 200.171 -  hence "isUb UNIV ?SHIFT (Y + (-X) + 1)" using Y_isUb by simp
 200.172 -  hence "\<exists>Z. isUb UNIV ?SHIFT Z" ..
 200.173 +  have *: "\<forall>s\<in>?SHIFT. s \<le> Y + (-X) + 1" using Y_isUb by (rule S_Ub_is_SHIFT_Ub)
 200.174 +  have "\<forall>s\<in>?SHIFT. s < Y + (-X) + 2"
 200.175 +  proof
 200.176 +    fix s assume "s\<in>?SHIFT"
 200.177 +    with * have "s \<le> Y + (-X) + 1" by simp
 200.178 +    also have "\<dots> < Y + (-X) + 2" by simp
 200.179 +    finally show "s < Y + (-X) + 2" .
 200.180 +  qed
 200.181    moreover have "\<forall>y \<in> ?SHIFT. 0 < y" by auto
 200.182    moreover have shifted_not_empty: "\<exists>u. u \<in> ?SHIFT"
 200.183      using X_in_S and Y_isUb by auto
 200.184 -  ultimately obtain t where t_is_Lub: "isLub UNIV ?SHIFT t"
 200.185 -    using posreals_complete [of ?SHIFT] by blast
 200.186 +  ultimately obtain t where t_is_Lub: "\<forall>y. (\<exists>x\<in>?SHIFT. y < x) = (y < t)"
 200.187 +    using posreal_complete [of ?SHIFT] unfolding bdd_above_def by blast
 200.188  
 200.189    show ?thesis
 200.190    proof
 200.191 -    show "isLub UNIV S (t + X + (-1))"
 200.192 -    proof (rule isLubI2)
 200.193 -      {
 200.194 -        fix x
 200.195 -        assume "isUb (UNIV::real set) S x"
 200.196 -        hence "isUb (UNIV::real set) (?SHIFT) (x + (-X) + 1)"
 200.197 -          using S_Ub_is_SHIFT_Ub by simp
 200.198 -        hence "t \<le> (x + (-X) + 1)"
 200.199 -          using t_is_Lub by (simp add: isLub_le_isUb)
 200.200 -        hence "t + X + -1 \<le> x" by arith
 200.201 -      }
 200.202 -      then show "(t + X + -1) <=* Collect (isUb UNIV S)"
 200.203 -        by (simp add: setgeI)
 200.204 +    show "(\<forall>s\<in>S. s \<le> (t + X + (-1))) \<and> (\<forall>y. (\<forall>s\<in>S. s \<le> y) \<longrightarrow> (t + X + (-1)) \<le> y)"
 200.205 +    proof safe
 200.206 +      fix x
 200.207 +      assume "\<forall>s\<in>S. s \<le> x"
 200.208 +      hence "\<forall>s\<in>?SHIFT. s \<le> x + (-X) + 1"
 200.209 +        using S_Ub_is_SHIFT_Ub by simp
 200.210 +      then have "\<not> x + (-X) + 1 < t"
 200.211 +        by (subst t_is_Lub[rule_format, symmetric]) (simp add: not_less)
 200.212 +      thus "t + X + -1 \<le> x" by arith
 200.213      next
 200.214 -      show "isUb UNIV S (t + X + -1)"
 200.215 -      proof -
 200.216 -        {
 200.217 -          fix y
 200.218 -          assume y_in_S: "y \<in> S"
 200.219 -          have "y \<le> t + X + -1"
 200.220 -          proof -
 200.221 -            obtain "u" where u_in_shift: "u \<in> ?SHIFT" using shifted_not_empty ..
 200.222 -            hence "\<exists> x \<in> S. u = x + - X + 1" by simp
 200.223 -            then obtain "x" where x_and_u: "u = x + - X + 1" ..
 200.224 -            have u_le_t: "u \<le> t" using u_in_shift and t_is_Lub by (simp add: isLubD2)
 200.225 +      fix y
 200.226 +      assume y_in_S: "y \<in> S"
 200.227 +      obtain "u" where u_in_shift: "u \<in> ?SHIFT" using shifted_not_empty ..
 200.228 +      hence "\<exists> x \<in> S. u = x + - X + 1" by simp
 200.229 +      then obtain "x" where x_and_u: "u = x + - X + 1" ..
 200.230 +      have u_le_t: "u \<le> t"
 200.231 +      proof (rule dense_le)
 200.232 +        fix x assume "x < u" then have "x < t"
 200.233 +          using u_in_shift t_is_Lub by auto
 200.234 +        then show "x \<le> t"  by simp
 200.235 +      qed
 200.236  
 200.237 -            show ?thesis
 200.238 -            proof cases
 200.239 -              assume "y \<le> x"
 200.240 -              moreover have "x = u + X + - 1" using x_and_u by arith
 200.241 -              moreover have "u + X + - 1  \<le> t + X + -1" using u_le_t by arith
 200.242 -              ultimately show "y  \<le> t + X + -1" by arith
 200.243 -            next
 200.244 -              assume "~(y \<le> x)"
 200.245 -              hence x_less_y: "x < y" by arith
 200.246 +      show "y \<le> t + X + -1"
 200.247 +      proof cases
 200.248 +        assume "y \<le> x"
 200.249 +        moreover have "x = u + X + - 1" using x_and_u by arith
 200.250 +        moreover have "u + X + - 1  \<le> t + X + -1" using u_le_t by arith
 200.251 +        ultimately show "y  \<le> t + X + -1" by arith
 200.252 +      next
 200.253 +        assume "~(y \<le> x)"
 200.254 +        hence x_less_y: "x < y" by arith
 200.255  
 200.256 -              have "x + (-X) + 1 \<in> ?SHIFT" using x_and_u and u_in_shift by simp
 200.257 -              hence "0 < x + (-X) + 1" by simp
 200.258 -              hence "0 < y + (-X) + 1" using x_less_y by arith
 200.259 -              hence "y + (-X) + 1 \<in> ?SHIFT" using y_in_S by simp
 200.260 -              hence "y + (-X) + 1 \<le> t" using t_is_Lub  by (simp add: isLubD2)
 200.261 -              thus ?thesis by simp
 200.262 -            qed
 200.263 -          qed
 200.264 -        }
 200.265 -        then show ?thesis by (simp add: isUb_def setle_def)
 200.266 +        have "x + (-X) + 1 \<in> ?SHIFT" using x_and_u and u_in_shift by simp
 200.267 +        hence "0 < x + (-X) + 1" by simp
 200.268 +        hence "0 < y + (-X) + 1" using x_less_y by arith
 200.269 +        hence *: "y + (-X) + 1 \<in> ?SHIFT" using y_in_S by simp
 200.270 +        have "y + (-X) + 1 \<le> t"
 200.271 +        proof (rule dense_le)
 200.272 +          fix x assume "x < y + (-X) + 1" then have "x < t"
 200.273 +            using * t_is_Lub by auto
 200.274 +          then show "x \<le> t"  by simp
 200.275 +        qed
 200.276 +        thus ?thesis by simp
 200.277        qed
 200.278      qed
 200.279    qed
 200.280  qed
 200.281  
 200.282 -text{*A version of the same theorem without all those predicates!*}
 200.283 -lemma reals_complete2:
 200.284 -  fixes S :: "(real set)"
 200.285 -  assumes "\<exists>y. y\<in>S" and "\<exists>(x::real). \<forall>y\<in>S. y \<le> x"
 200.286 -  shows "\<exists>x. (\<forall>y\<in>S. y \<le> x) & 
 200.287 -               (\<forall>z. ((\<forall>y\<in>S. y \<le> z) --> x \<le> z))"
 200.288 -proof -
 200.289 -  have "\<exists>x. isLub UNIV S x" 
 200.290 -    by (rule reals_complete)
 200.291 -       (auto simp add: isLub_def isUb_def leastP_def setle_def setge_def assms)
 200.292 -  thus ?thesis
 200.293 -    by (metis UNIV_I isLub_isUb isLub_le_isUb isUbD isUb_def setleI)
 200.294 -qed
 200.295 -
 200.296 -
 200.297  subsection {* The Archimedean Property of the Reals *}
 200.298  
 200.299  theorem reals_Archimedean:
 200.300 @@ -1969,34 +1871,30 @@
 200.301        by (rule mult_right_mono)
 200.302      thus "x * of_nat (Suc n) \<le> 1" by (simp del: of_nat_Suc)
 200.303    qed
 200.304 -  hence "{z. \<exists>n. z = x * (of_nat (Suc n))} *<= 1"
 200.305 -    by (simp add: setle_def del: of_nat_Suc, safe, rule spec)
 200.306 -  hence "isUb (UNIV::real set) {z. \<exists>n. z = x * (of_nat (Suc n))} 1"
 200.307 -    by (simp add: isUbI)
 200.308 -  hence "\<exists>Y. isUb (UNIV::real set) {z. \<exists>n. z = x* (of_nat (Suc n))} Y" ..
 200.309 -  moreover have "\<exists>X. X \<in> {z. \<exists>n. z = x* (of_nat (Suc n))}" by auto
 200.310 -  ultimately have "\<exists>t. isLub UNIV {z. \<exists>n. z = x * of_nat (Suc n)} t"
 200.311 -    by (simp add: reals_complete)
 200.312 -  then obtain "t" where
 200.313 -    t_is_Lub: "isLub UNIV {z. \<exists>n. z = x * of_nat (Suc n)} t" ..
 200.314 +  hence 2: "bdd_above {z. \<exists>n. z = x * (of_nat (Suc n))}"
 200.315 +    by (auto intro!: bdd_aboveI[of _ 1])
 200.316 +  have 1: "\<exists>X. X \<in> {z. \<exists>n. z = x* (of_nat (Suc n))}" by auto
 200.317 +  obtain t where
 200.318 +    upper: "\<And>z. z \<in> {z. \<exists>n. z = x * of_nat (Suc n)} \<Longrightarrow> z \<le> t" and
 200.319 +    least: "\<And>y. (\<And>a. a \<in> {z. \<exists>n. z = x * of_nat (Suc n)} \<Longrightarrow> a \<le> y) \<Longrightarrow> t \<le> y"
 200.320 +    using reals_complete[OF 1 2] by auto
 200.321  
 200.322 -  have "\<forall>n::nat. x * of_nat n \<le> t + - x"
 200.323 -  proof
 200.324 -    fix n
 200.325 -    from t_is_Lub have "x * of_nat (Suc n) \<le> t"
 200.326 -      by (simp add: isLubD2)
 200.327 -    hence  "x * (of_nat n) + x \<le> t"
 200.328 -      by (simp add: distrib_left)
 200.329 -    thus  "x * (of_nat n) \<le> t + - x" by arith
 200.330 +
 200.331 +  have "t \<le> t + - x"
 200.332 +  proof (rule least)
 200.333 +    fix a assume a: "a \<in> {z. \<exists>n. z = x * (of_nat (Suc n))}"
 200.334 +    have "\<forall>n::nat. x * of_nat n \<le> t + - x"
 200.335 +    proof
 200.336 +      fix n
 200.337 +      have "x * of_nat (Suc n) \<le> t"
 200.338 +        by (simp add: upper)
 200.339 +      hence  "x * (of_nat n) + x \<le> t"
 200.340 +        by (simp add: distrib_left)
 200.341 +      thus  "x * (of_nat n) \<le> t + - x" by arith
 200.342 +    qed    hence "\<forall>m. x * of_nat (Suc m) \<le> t + - x" by (simp del: of_nat_Suc)
 200.343 +    with a show "a \<le> t + - x"
 200.344 +      by auto
 200.345    qed
 200.346 -
 200.347 -  hence "\<forall>m. x * of_nat (Suc m) \<le> t + - x" by (simp del: of_nat_Suc)
 200.348 -  hence "{z. \<exists>n. z = x * (of_nat (Suc n))}  *<= (t + - x)"
 200.349 -    by (auto simp add: setle_def)
 200.350 -  hence "isUb (UNIV::real set) {z. \<exists>n. z = x * (of_nat (Suc n))} (t + (-x))"
 200.351 -    by (simp add: isUbI)
 200.352 -  hence "t \<le> t + - x"
 200.353 -    using t_is_Lub by (simp add: isLub_le_isUb)
 200.354    thus False using x_pos by arith
 200.355  qed
 200.356  
   201.1 --- a/src/HOL/ex/Gauge_Integration.thy	Mon Nov 11 17:34:44 2013 +0100
   201.2 +++ b/src/HOL/ex/Gauge_Integration.thy	Mon Nov 11 17:44:21 2013 +0100
   201.3 @@ -511,9 +511,9 @@
   201.4    case False
   201.5    then have "inverse (z - x) * (f z - f x - f' x * (z - x)) = (f z - f x) / (z - x) - f' x"
   201.6      apply (subst mult_commute)
   201.7 -    apply (simp add: distrib_right diff_minus)
   201.8 +    apply (simp add: left_diff_distrib)
   201.9      apply (simp add: mult_assoc divide_inverse)
  201.10 -    apply (simp add: distrib_right)
  201.11 +    apply (simp add: ring_distribs)
  201.12      done
  201.13    moreover from False `\<bar>z - x\<bar> < s` have "\<bar>(f z - f x) / (z - x) - f' x\<bar> < e / 2"
  201.14      by (rule P)
   202.1 --- a/src/Provers/splitter.ML	Mon Nov 11 17:34:44 2013 +0100
   202.2 +++ b/src/Provers/splitter.ML	Mon Nov 11 17:44:21 2013 +0100
   202.3 @@ -79,6 +79,8 @@
   202.4    fold add_thm splits []
   202.5  end;
   202.6  
   202.7 +val abss = fold (Term.abs o pair "");
   202.8 +
   202.9  (* ------------------------------------------------------------------------- *)
  202.10  (* mk_case_split_tac                                                         *)
  202.11  (* ------------------------------------------------------------------------- *)
  202.12 @@ -100,31 +102,36 @@
  202.13    (Syntax.read_prop_global Pure.thy "P(%x. Q(x)) == P(%x. R(x))")
  202.14    (fn {prems, ...} => rewrite_goals_tac prems THEN rtac reflexive_thm 1)
  202.15  
  202.16 +val _ $ _ $ (_ $ (_ $ abs_lift) $ _) = prop_of lift;
  202.17 +
  202.18  val trlift = lift RS transitive_thm;
  202.19 -val _ $ (P $ _) $ _ = concl_of trlift;
  202.20  
  202.21  
  202.22  (************************************************************************
  202.23     Set up term for instantiation of P in the lift-theorem
  202.24  
  202.25 -   Ts    : types of parameters (i.e. variables bound by meta-quantifiers)
  202.26     t     : lefthand side of meta-equality in subgoal
  202.27             the lift theorem is applied to (see select)
  202.28     pos   : "path" leading to abstraction, coded as a list
  202.29     T     : type of body of P(...)
  202.30 -   maxi  : maximum index of Vars
  202.31  *************************************************************************)
  202.32  
  202.33 -fun mk_cntxt Ts t pos T maxi =
  202.34 -  let fun var (t,i) = Var(("X",i),type_of1(Ts,t));
  202.35 -      fun down [] t i = Bound 0
  202.36 -        | down (p::ps) t i =
  202.37 -            let val (h,ts) = strip_comb t
  202.38 -                val v1 = ListPair.map var (take p ts, i upto (i+p-1))
  202.39 -                val u::us = drop p ts
  202.40 -                val v2 = ListPair.map var (us, (i+p) upto (i+length(ts)-2))
  202.41 -      in list_comb(h,v1@[down ps u (i+length ts)]@v2) end;
  202.42 -  in Abs("", T, down (rev pos) t maxi) end;
  202.43 +fun mk_cntxt t pos T =
  202.44 +  let
  202.45 +    fun down [] t = (Bound 0, t)
  202.46 +      | down (p :: ps) t =
  202.47 +          let
  202.48 +            val (h, ts) = strip_comb t
  202.49 +            val (ts1, u :: ts2) = chop p ts
  202.50 +            val (u1, u2) = down ps u
  202.51 +          in
  202.52 +            (list_comb (incr_boundvars 1 h,
  202.53 +               map (incr_boundvars 1) ts1 @ u1 ::
  202.54 +               map (incr_boundvars 1) ts2),
  202.55 +             u2)
  202.56 +          end;
  202.57 +    val (u1, u2) = down (rev pos) t
  202.58 +  in (Abs ("", T, u1), u2) end;
  202.59  
  202.60  
  202.61  (************************************************************************
  202.62 @@ -301,15 +308,18 @@
  202.63               the split theorem is applied to (see cmap)
  202.64     T,U,pos : see mk_split_pack
  202.65     state   : current proof state
  202.66 -   lift    : the lift theorem
  202.67     i       : no. of subgoal
  202.68  **************************************************************)
  202.69  
  202.70  fun inst_lift Ts t (T, U, pos) state i =
  202.71    let
  202.72      val cert = cterm_of (Thm.theory_of_thm state);
  202.73 -    val cntxt = mk_cntxt Ts t pos (T --> U) (Thm.maxidx_of trlift);
  202.74 -  in cterm_instantiate [(cert P, cert cntxt)] trlift
  202.75 +    val (cntxt, u) = mk_cntxt t pos (T --> U);
  202.76 +    val trlift' = Thm.lift_rule (Thm.cprem_of state i)
  202.77 +      (Thm.rename_boundvars abs_lift u trlift);
  202.78 +    val (P, _) = strip_comb (fst (Logic.dest_equals
  202.79 +      (Logic.strip_assums_concl (Thm.prop_of trlift'))));
  202.80 +  in cterm_instantiate [(cert P, cert (abss Ts cntxt))] trlift'
  202.81    end;
  202.82  
  202.83  
  202.84 @@ -333,7 +343,6 @@
  202.85        (Logic.strip_assums_concl (Thm.prop_of thm'))));
  202.86      val cert = cterm_of (Thm.theory_of_thm state);
  202.87      val cntxt = mk_cntxt_splitthm t tt TB;
  202.88 -    val abss = fold (fn T => fn t => Abs ("", T, t));
  202.89    in cterm_instantiate [(cert P, cert (abss Ts cntxt))] thm'
  202.90    end;
  202.91  
  202.92 @@ -348,7 +357,7 @@
  202.93  fun split_tac [] i = no_tac
  202.94    | split_tac splits i =
  202.95    let val cmap = cmap_of_split_thms splits
  202.96 -      fun lift_tac Ts t p st = rtac (inst_lift Ts t p st i) i st
  202.97 +      fun lift_tac Ts t p st = compose_tac (false, inst_lift Ts t p st i, 2) i st
  202.98        fun lift_split_tac state =
  202.99              let val (Ts, t, splits) = select cmap state i
 202.100              in case splits of