Merged with mainline changes.
authorThomas Sewell <tsewell@nicta.com.au>
Fri, 11 Sep 2009 15:57:16 +1000
changeset 32747 8b9ced1051e2
parent 32746 2f1701a6d4e8 (current diff)
parent 32552 4d4ee06e9420 (diff)
child 32748 887c68b70f7d
Merged with mainline changes.
lib/Tools/codegen
lib/Tools/jedit
lib/scripts/mirabelle
src/HOL/Library/Legacy_GCD.thy
src/HOL/Library/Pocklington.thy
src/HOL/Library/Primes.thy
src/HOL/NewNumberTheory/Binomial.thy
src/HOL/NewNumberTheory/Cong.thy
src/HOL/NewNumberTheory/Fib.thy
src/HOL/NewNumberTheory/MiscAlgebra.thy
src/HOL/NewNumberTheory/ROOT.ML
src/HOL/NewNumberTheory/Residues.thy
src/HOL/NewNumberTheory/UniqueFactorization.thy
src/HOL/NumberTheory/BijectionRel.thy
src/HOL/NumberTheory/Chinese.thy
src/HOL/NumberTheory/Euler.thy
src/HOL/NumberTheory/EulerFermat.thy
src/HOL/NumberTheory/EvenOdd.thy
src/HOL/NumberTheory/Factorization.thy
src/HOL/NumberTheory/Fib.thy
src/HOL/NumberTheory/Finite2.thy
src/HOL/NumberTheory/Gauss.thy
src/HOL/NumberTheory/Int2.thy
src/HOL/NumberTheory/IntFact.thy
src/HOL/NumberTheory/IntPrimes.thy
src/HOL/NumberTheory/Quadratic_Reciprocity.thy
src/HOL/NumberTheory/ROOT.ML
src/HOL/NumberTheory/Residues.thy
src/HOL/NumberTheory/WilsonBij.thy
src/HOL/NumberTheory/WilsonRuss.thy
src/HOL/NumberTheory/document/root.tex
src/HOL/Tools/ComputeFloat.thy
src/HOL/Tools/ComputeHOL.thy
src/HOL/Tools/ComputeNumeral.thy
src/HOL/ex/Mirabelle.thy
src/HOL/ex/mirabelle.ML
src/Pure/Isar/isar.scala
src/Pure/Tools/isabelle_syntax.scala
--- a/Admin/isatest/annomaly.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/Admin/isatest/annomaly.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -20,7 +20,7 @@
   val isabelleHome =
       case OS.Process.getEnv "ISABELLE_HOME"
        of  NONE => OS.FileSys.getDir ()
-	 | SOME home => mkAbsolute home
+         | SOME home => mkAbsolute home
 
   fun noparent [] = []
     | noparent (n :: ns) =
@@ -33,12 +33,12 @@
 
   fun rewrite defPrefix name =
       let val abs = mkAbsolute name
-	  val rel = OS.Path.mkRelative { path = abs, relativeTo = isabelleHome }
-	  val exists = (OS.FileSys.access(abs, nil)
-			handle OS.SysErr _ => false)
+          val rel = OS.Path.mkRelative { path = abs, relativeTo = isabelleHome }
+          val exists = (OS.FileSys.access(abs, nil)
+                        handle OS.SysErr _ => false)
       in  if exists andalso rel <> ""
-	  then isabellePath (toArcs rel)
-	  else defPrefix @ noparent (toArcs name)
+          then isabellePath (toArcs rel)
+          else defPrefix @ noparent (toArcs name)
       end handle OS.Path.Path => defPrefix @ noparent (toArcs name)
 
 in
@@ -49,10 +49,10 @@
         (* should we have different files for different line numbers? *
         val arcs = if line <= 1 then arcs
                    else arcs @ [ "l." ^ Int.toString line ]
-	*)
-	val arcs = if t = "structure Isabelle =\nstruct\nend;"
-		      andalso name = "ML"
-		   then ["empty_Isabelle", "empty" ] else arcs
+        *)
+        val arcs = if t = "structure Isabelle =\nstruct\nend;"
+                      andalso name = "ML"
+                   then ["empty_Isabelle", "empty" ] else arcs
         val _    = AnnoMaLy.nameNextStream arcs
     in  smlnj_use_text tune str_of_pos name_space (line, name) p v t  end;
 
--- a/Admin/isatest/isatest-makeall	Fri Sep 11 15:56:51 2009 +1000
+++ b/Admin/isatest/isatest-makeall	Fri Sep 11 15:57:16 2009 +1000
@@ -10,6 +10,8 @@
 # max time until test is aborted (in sec)
 MAXTIME=28800
 
+PUBLISH_TEST=/home/isabelle-repository/repos/testtool/publish_test.py
+
 ## diagnostics
 
 PRG="$(basename "$0")"
@@ -80,7 +82,7 @@
         NICE=""
         ;;
 
-    macbroy21)
+    macbroy22)
         MFLAGS="-k"
         NICE=""
         ;;
@@ -120,6 +122,8 @@
   TOOL="$ISABELLE_TOOL makeall $MFLAGS all"
 fi
 
+IDENT=$(cat "$DISTPREFIX/ISABELLE_IDENT")
+
 # main test loop
 
 log "starting [$@]"
@@ -159,10 +163,16 @@
     then
         # test log and cleanup
         echo ------------------- test successful --- `date` --- $HOSTNAME >> $TESTLOG 2>&1
+        if [ -x $PUBLISH_TEST ]; then
+            $PUBLISH_TEST -r $IDENT -s "SUCCESS" -a log $TESTLOG
+        fi
         gzip -f $TESTLOG
     else
         # test log
         echo ------------------- test FAILED --- `date` --- $HOSTNAME >> $TESTLOG 2>&1
+        if [ -x $PUBLISH_TEST ]; then
+             $PUBLISH_TEST -r $IDENT -s "FAIL" -a log $TESTLOG
+        fi
 
         # error log
         echo "Test for platform ${SHORT} failed. Log file attached." >> $ERRORLOG
--- a/Admin/isatest/isatest-makedist	Fri Sep 11 15:56:51 2009 +1000
+++ b/Admin/isatest/isatest-makedist	Fri Sep 11 15:57:16 2009 +1000
@@ -94,13 +94,13 @@
 $SSH sunbroy2 "$MAKEALL $HOME/settings/sun-poly"
 # give test some time to copy settings and start
 sleep 15
-$SSH macbroy21 "$MAKEALL $HOME/settings/at-poly"
+$SSH macbroy22 "$MAKEALL $HOME/settings/at-poly"
 sleep 15
 $SSH macbroy20 "$MAKEALL $HOME/settings/at-poly-5.1-para-e"
 sleep 15
 #$SSH macbroy24 "$MAKEALL -l HOL proofterms $HOME/settings/at-sml-dev-p"
 #sleep 15
-$SSH macbroy22 "$MAKEALL $HOME/settings/at64-poly-5.1-para"
+$SSH macbroy21 "$MAKEALL $HOME/settings/at64-poly-5.1-para"
 sleep 15
 $SSH macbroy23 -l HOL images "$MAKEALL $HOME/settings/at-sml-dev-e"
 sleep 15
--- a/Admin/isatest/isatest-stats	Fri Sep 11 15:56:51 2009 +1000
+++ b/Admin/isatest/isatest-stats	Fri Sep 11 15:57:16 2009 +1000
@@ -6,7 +6,7 @@
 
 THIS=$(cd "$(dirname "$0")"; pwd -P)
 
-PLATFORMS="at-poly at64-poly mac-poly-M4 mac-poly64-M4 mac-poly-M8 at-poly-5.1-para-e at64-poly-5.1-para at-mac-poly-5.1-para afp at-sml-dev"
+PLATFORMS="at-poly at64-poly mac-poly-M4 mac-poly64-M4 mac-poly-M8 mac-poly64-M8 at-poly-5.1-para-e at64-poly-5.1-para at-mac-poly-5.1-para afp at-sml-dev"
 
 ISABELLE_SESSIONS="\
   HOL-Plain \
--- a/NEWS	Fri Sep 11 15:56:51 2009 +1000
+++ b/NEWS	Fri Sep 11 15:57:16 2009 +1000
@@ -18,18 +18,34 @@
 
 *** HOL ***
 
-* New proof method "sos" (sum of squares) for nonlinear real arithmetic
-(originally due to John Harison). It requires Library/Sum_Of_Squares.
-It is not a complete decision procedure but works well in practice
-on quantifier-free real arithmetic with +, -, *, ^, =, <= and <,
-i.e. boolean combinations of equalities and inequalities between
-polynomials. It makes use of external semidefinite programming solvers.
-For more information and examples see Library/Sum_Of_Squares.
-
-* Set.UNIV and Set.empty are mere abbreviations for top and bot.  INCOMPATIBILITY.
-
-* More convenient names for set intersection and union.  INCOMPATIBILITY:
-
+* Reorganization of number theory:
+  * former session NumberTheory now named Old_Number_Theory; former session NewNumberTheory
+    named NumberTheory;
+  * split off prime number ingredients from theory GCD to theory Number_Theory/Primes;
+  * moved legacy theories Legacy_GCD and Primes from Library/ to Old_Number_Theory/;
+  * moved theory Pocklington from Library/ to Old_Number_Theory/;
+  * removed various references to Old_Number_Theory from HOL distribution.
+INCOMPATIBILITY.
+
+* New testing tool "Mirabelle" for automated (proof) tools. Applies
+several tools and tactics like sledgehammer, metis, or quickcheck, to
+every proof step in a theory. To be used in batch mode via the
+"mirabelle" utility.
+
+* New proof method "sos" (sum of squares) for nonlinear real
+arithmetic (originally due to John Harison). It requires
+Library/Sum_Of_Squares.  It is not a complete decision procedure but
+works well in practice on quantifier-free real arithmetic with +, -,
+*, ^, =, <= and <, i.e. boolean combinations of equalities and
+inequalities between polynomials. It makes use of external
+semidefinite programming solvers.  For more information and examples
+see src/HOL/Library/Sum_Of_Squares.
+
+* Set.UNIV and Set.empty are mere abbreviations for top and bot.
+INCOMPATIBILITY.
+
+* More convenient names for set intersection and union.
+INCOMPATIBILITY:
     Set.Int ~>  Set.inter
     Set.Un ~>   Set.union
 
@@ -39,8 +55,6 @@
     etc.
   INCOMPATIBILITY.
 
-* New quickcheck implementation using new code generator.
-
 * New class "boolean_algebra".
 
 * Refinements to lattices classes:
@@ -77,15 +91,16 @@
 multiplicative monoids retains syntax "^" and is now defined generic
 in class power.  INCOMPATIBILITY.
 
-* Relation composition "R O S" now has a "more standard" argument order,
-i.e., "R O S = {(x,z). EX y. (x,y) : R & (y,z) : S }".
+* Relation composition "R O S" now has a "more standard" argument
+order, i.e., "R O S = {(x,z). EX y. (x,y) : R & (y,z) : S }".
 INCOMPATIBILITY: Rewrite propositions with "S O R" --> "R O S". Proofs
-may occationally break, since the O_assoc rule was not rewritten like this.
-Fix using O_assoc[symmetric].
-The same applies to the curried version "R OO S".
-
-* GCD now has functions Gcd/GCD and Lcm/LCM for the gcd and lcm of finite and
-infinite sets. It is shown that they form a complete lattice.
+may occationally break, since the O_assoc rule was not rewritten like
+this.  Fix using O_assoc[symmetric].  The same applies to the curried
+version "R OO S".
+
+* Theory GCD now has functions Gcd/GCD and Lcm/LCM for the gcd and lcm
+of finite and infinite sets. It is shown that they form a complete
+lattice.
 
 * ML antiquotation @{code_datatype} inserts definition of a datatype
 generated by the code generator; see Predicate.thy for an example.
@@ -93,9 +108,9 @@
 * New method "linarith" invokes existing linear arithmetic decision
 procedure only.
 
-* Implementation of quickcheck using generic code generator; default
-generators are provided for all suitable HOL types, records and
-datatypes.
+* New implementation of quickcheck uses generic code generator;
+default generators are provided for all suitable HOL types, records
+and datatypes.
 
 * Constants Set.Pow and Set.image now with authentic syntax;
 object-logic definitions Set.Pow_def and Set.image_def.
@@ -106,6 +121,9 @@
 Suc_eq_add_numeral_1_left -> Suc_eq_plus1_left
 Suc_plus1 -> Suc_eq_plus1
 
+* Moved theorems:
+Wellfounded.in_inv_image -> Relation.in_inv_image
+
 * New sledgehammer option "Full Types" in Proof General settings menu.
 Causes full type information to be output to the ATPs.  This slows
 ATPs down considerably but eliminates a source of unsound "proofs"
@@ -117,17 +135,13 @@
     DatatypePackage ~> Datatype
     InductivePackage ~> Inductive
 
-    etc.
-
 INCOMPATIBILITY.
 
 * NewNumberTheory: Jeremy Avigad's new version of part of
 NumberTheory.  If possible, use NewNumberTheory, not NumberTheory.
 
-* Simplified interfaces of datatype module.  INCOMPATIBILITY.
-
-* Abbreviation "arbitrary" of "undefined" has disappeared; use
-"undefined" directly.  INCOMPATIBILITY.
+* Discontinued abbreviation "arbitrary" of constant
+"undefined". INCOMPATIBILITY, use "undefined" directly.
 
 * New evaluator "approximate" approximates an real valued term using
 the same method as the approximation method.
@@ -151,10 +165,14 @@
 * PARALLEL_CHOICE and PARALLEL_GOALS provide basic support for
 parallel tactical reasoning.
 
-* Tactical FOCUS is similar to SUBPROOF, but allows the body tactic to
-introduce new subgoals and schematic variables.  FOCUS_PARAMS is
-similar, but focuses on the parameter prefix only, leaving subgoal
-premises unchanged.
+* Tacticals Subgoal.FOCUS, Subgoal.FOCUS_PREMS, Subgoal.FOCUS_PARAMS
+are similar to SUBPROOF, but are slightly more flexible: only the
+specified parts of the subgoal are imported into the context, and the
+body tactic may introduce new subgoals and schematic variables.
+
+* Old tactical METAHYPS, which does not observe the proof context, has
+been renamed to Old_Goals.METAHYPS and awaits deletion.  Use SUBPROOF
+or Subgoal.FOCUS etc.
 
 * Renamed functor TableFun to Table, and GraphFun to Graph.  (Since
 functors have their own ML name space there is no point to mark them
@@ -175,6 +193,10 @@
 or even Display.pretty_thm_without_context as last resort.
 INCOMPATIBILITY.
 
+* Discontinued Display.pretty_ctyp/cterm etc.  INCOMPATIBILITY, use
+Syntax.pretty_typ/term directly, preferably with proper context
+instead of global theory.
+
 
 *** System ***
 
--- a/bin/isabelle	Fri Sep 11 15:56:51 2009 +1000
+++ b/bin/isabelle	Fri Sep 11 15:57:16 2009 +1000
@@ -17,7 +17,7 @@
 ISABELLE_HOME="$(cd "$(dirname "$0")"; cd "$(pwd -P)"; cd ..; pwd)"
 source "$ISABELLE_HOME/lib/scripts/getsettings" || exit 2
 
-ORIG_IFS="$IFS"; IFS=":"; declare -a TOOLS=($ISABELLE_TOOLS); IFS="$ORIG_IFS"
+splitarray ":" "$ISABELLE_TOOLS"; TOOLS=("${SPLITARRAY[@]}")
 
 
 ## diagnostics
--- a/bin/isabelle-process	Fri Sep 11 15:56:51 2009 +1000
+++ b/bin/isabelle-process	Fri Sep 11 15:57:16 2009 +1000
@@ -160,7 +160,7 @@
     INFILE=""
     ISA_PATH=""
 
-    ORIG_IFS="$IFS"; IFS=":"; declare -a PATHS=($ISABELLE_PATH); IFS="$ORIG_IFS"
+    splitarray ":" "$ISABELLE_PATH"; PATHS=("${SPLITARRAY[@]}")
     for DIR in "${PATHS[@]}"
     do
       DIR="$DIR/$ML_IDENTIFIER"
--- a/doc-src/rail.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/doc-src/rail.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -99,7 +99,7 @@
       |> (if markup = "" then I else enclose ("\\" ^ markup ^ "{") "}")
       |> (if ! ThyOutput.quotes then quote else I)
       |> (if ! ThyOutput.display then enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}"
-	  else hyper o enclose "\\mbox{\\isa{" "}}")), style)
+          else hyper o enclose "\\mbox{\\isa{" "}}")), style)
   else ("Bad " ^ kind ^ " " ^ name, false)
   end;
 end;
@@ -147,8 +147,8 @@
   ) >> (Identifier o enclose "\\isa{" "}" o Output.output o implode) ||
   scan_link >> (decode_link ctxt) >>
     (fn (txt, style) =>
-	if style then Special_Identifier(txt)
-	else Identifier(txt))
+        if style then Special_Identifier(txt)
+        else Identifier(txt))
 end;
 
 fun scan_anot ctxt =
@@ -169,12 +169,12 @@
     val text_sq =
       Scan.repeat (
         Scan.one (fn s =>
-	  s <> "\n" andalso
-	  s <> "\t" andalso
-	  s <> "'" andalso
-	  s <> "\\" andalso
-	  Symbol.is_regular s) ||
-	($$ "\\" |-- $$ "'")
+          s <> "\n" andalso
+          s <> "\t" andalso
+          s <> "'" andalso
+          s <> "\\" andalso
+          Symbol.is_regular s) ||
+        ($$ "\\" |-- $$ "'")
       ) >> implode
   fun quoted scan = $$ "'" |-- scan --| $$ "'";
   in
@@ -305,9 +305,9 @@
   parse_body2 -- ($$$ "*" |-- !!! "body4e expected" (parse_body4e)) >>
     (fn (body1, body2) =>
       if is_empty body2 then
-	add_body(PLUS, new_empty_body, rev_body body1)
+        add_body(PLUS, new_empty_body, rev_body body1)
       else
-	add_body(BAR, new_empty_body, add_body (PLUS, body1, rev_body body2)) ) ||
+        add_body(BAR, new_empty_body, add_body (PLUS, body1, rev_body body2)) ) ||
   parse_body2 -- ($$$ "+" |-- !!! "body4e expected" (parse_body4e)) >>
     (fn (body1, body2) => new_body (PLUS, body1, rev_body body2) ) ||
   parse_body2e
@@ -365,36 +365,36 @@
 fun position_body (body as Body(kind, text, annot, id, bodies), ystart) =
   let fun max (x,y) = if x > y then x else y
     fun set_body_position (Body(kind, text, annot, id, bodies), ystart, ynext) =
-	  Body_Pos(kind, text, annot, id, bodies, ystart, ynext)
+          Body_Pos(kind, text, annot, id, bodies, ystart, ynext)
     fun pos_bodies_cat ([],_,ynext,liste) = (liste, ynext)
       | pos_bodies_cat (x::xs, ystart, ynext, liste) =
-	  if is_kind_of CR x then
-	      (case set_body_position(x, ystart, ynext+1) of
-		body as Body_Pos(_,_,_,_,_,_,ynext1) =>
-		  pos_bodies_cat(xs, ynext1, max(ynext,ynext1), liste@[body])
-	      )
-	  else
-	      (case position_body(x, ystart) of
-		body as Body_Pos(_,_,_,_,_,_,ynext1) =>
-		  pos_bodies_cat(xs, ystart, max(ynext,ynext1), liste@[body])
-	      )
+          if is_kind_of CR x then
+              (case set_body_position(x, ystart, ynext+1) of
+                body as Body_Pos(_,_,_,_,_,_,ynext1) =>
+                  pos_bodies_cat(xs, ynext1, max(ynext,ynext1), liste@[body])
+              )
+          else
+              (case position_body(x, ystart) of
+                body as Body_Pos(_,_,_,_,_,_,ynext1) =>
+                  pos_bodies_cat(xs, ystart, max(ynext,ynext1), liste@[body])
+              )
     fun pos_bodies_bar_plus ([],_,ynext,liste) = (liste, ynext)
       | pos_bodies_bar_plus (x::xs, ystart, ynext, liste) =
-	  (case position_body(x, ystart) of
-	    body as Body_Pos(_,_,_,_,_,_,ynext1) =>
-	      pos_bodies_bar_plus(xs, ynext1, max(ynext,ynext1), liste@[body])
-	  )
+          (case position_body(x, ystart) of
+            body as Body_Pos(_,_,_,_,_,_,ynext1) =>
+              pos_bodies_bar_plus(xs, ynext1, max(ynext,ynext1), liste@[body])
+          )
   in
   (case kind of
     CAT => (case pos_bodies_cat(bodies,ystart,ystart+1,[]) of
-	      (bodiesPos, ynext) =>
-		Body_Pos(kind, text, annot, id, bodiesPos, ystart, ynext))
+              (bodiesPos, ynext) =>
+                Body_Pos(kind, text, annot, id, bodiesPos, ystart, ynext))
   | BAR => (case pos_bodies_bar_plus(bodies,ystart,ystart+1,[]) of
-	      (bodiesPos, ynext) =>
-		Body_Pos(kind, text, annot, id, bodiesPos, ystart, ynext))
+              (bodiesPos, ynext) =>
+                Body_Pos(kind, text, annot, id, bodiesPos, ystart, ynext))
   | PLUS => (case pos_bodies_bar_plus(bodies,ystart,ystart+1,[]) of
-	      (bodiesPos, ynext) =>
-		Body_Pos(kind, text, annot, id, bodiesPos, ystart, ynext))
+              (bodiesPos, ynext) =>
+                Body_Pos(kind, text, annot, id, bodiesPos, ystart, ynext))
   | CR => set_body_position(body, ystart, ystart+3)
   | EMPTY => set_body_position(body, ystart, ystart+1)
   | ANNOTE => set_body_position(body, ystart, ystart+1)
@@ -406,15 +406,15 @@
 fun format_body (Body_Pos(EMPTY,_,_,_,_,_,_), _) = ""
   | format_body (Body_Pos(CAT,_,_,_,bodies,_,_), cent) =
     let fun format_bodies([]) = ""
-	  | format_bodies(x::xs) = format_body (x, "") ^ format_bodies(xs)
+          | format_bodies(x::xs) = format_body (x, "") ^ format_bodies(xs)
     in
       format_bodies(bodies)
     end
   | format_body (Body_Pos(BAR,_,_,_,bodies,_,_),cent) =
     let fun format_bodies([]) = "\\rail@endbar\n"
-	  | format_bodies(x::xs) =
-	      "\\rail@nextbar{" ^ string_of_int(getystart(x)) ^"}\n" ^
-	      format_body(x, "") ^ format_bodies(xs)
+          | format_bodies(x::xs) =
+              "\\rail@nextbar{" ^ string_of_int(getystart(x)) ^"}\n" ^
+              format_body(x, "") ^ format_bodies(xs)
     in
       "\\rail@bar\n" ^ format_body(hd(bodies), "") ^ format_bodies(tl(bodies))
     end
--- a/etc/components	Fri Sep 11 15:56:51 2009 +1000
+++ b/etc/components	Fri Sep 11 15:57:16 2009 +1000
@@ -11,6 +11,7 @@
 src/LCF
 src/Sequents
 #misc components
+src/Tools/Code
 src/HOL/Tools/ATP_Manager
+src/HOL/Mirabelle
 src/HOL/Library/Sum_Of_Squares
-
--- a/etc/settings	Fri Sep 11 15:56:51 2009 +1000
+++ b/etc/settings	Fri Sep 11 15:57:16 2009 +1000
@@ -173,7 +173,7 @@
 
 # The pdf file viewer
 if [ $(uname -s) = Darwin ]; then
-  PDF_VIEWER=open
+  PDF_VIEWER="open -W -n"
 else
   PDF_VIEWER=xpdf
 fi
@@ -207,22 +207,6 @@
 
 
 ###
-### jEdit
-###
-
-JEDIT_HOME=$(choosefrom \
-  "$ISABELLE_HOME/contrib/jedit" \
-  "$ISABELLE_HOME/../jedit" \
-  "/usr/local/jedit" \
-  "/usr/share/jedit" \
-  "/opt/jedit" \
-  "")
-
-JEDIT_JAVA_OPTIONS=""
-#JEDIT_JAVA_OPTIONS="-server -Xms128m -Xmx512m"
-JEDIT_OPTIONS="-reuseview -noserver -nobackground"
-
-###
 ### External reasoning tools
 ###
 
--- a/lib/Tools/codegen	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,40 +0,0 @@
-#!/usr/bin/env bash
-#
-# Author: Florian Haftmann, TUM
-#
-# DESCRIPTION: issue code generation from shell
-
-
-## diagnostics
-
-PRG="$(basename "$0")"
-
-function usage()
-{
-  echo
-  echo "Usage: isabelle $PRG IMAGE THY CMD"
-  echo
-  echo "  Issues code generation using image IMAGE,"
-  echo "  theory THY,"
-  echo "  with Isar command 'export_code CMD'"
-  echo
-  exit 1
-}
-
-
-## process command line
-
-[ "$#" -lt 2 -o "$1" = "-?" ] && usage
-
-IMAGE="$1"; shift
-THY="$1"; shift
-CMD="$1"
-
-
-## main
-
-CODE_CMD=$(echo $CMD | perl -pe 's/\\/\\\\/g; s/"/\\\"/g')
-CTXT_CMD="ML_Context.eval_in (SOME (ProofContext.init (theory \"HOL\"))) false Position.none \"Code_Target.shell_command thyname cmd\";"
-FULL_CMD="val thyname = \"$THY\"; val cmd = \"$CODE_CMD\"; $CTXT_CMD"
-
-"$ISABELLE" -q -e "$FULL_CMD" "$IMAGE" || exit 1
--- a/lib/Tools/doc	Fri Sep 11 15:56:51 2009 +1000
+++ b/lib/Tools/doc	Fri Sep 11 15:57:16 2009 +1000
@@ -34,7 +34,7 @@
 
 ## main
 
-ORIG_IFS="$IFS"; IFS=":"; declare -a DOCS=($ISABELLE_DOCS); IFS="$ORIG_IFS"
+splitarray ":" "$ISABELLE_DOCS"; DOCS=("${SPLITARRAY[@]}")
 
 if [ -z "$DOC" ]; then
   for DIR in "${DOCS[@]}"
--- a/lib/Tools/document	Fri Sep 11 15:56:51 2009 +1000
+++ b/lib/Tools/document	Fri Sep 11 15:57:16 2009 +1000
@@ -53,7 +53,7 @@
       OUTFORMAT="$OPTARG"
       ;;
     t)
-      ORIG_IFS="$IFS"; IFS=","; TAGS=($OPTARG); IFS="$ORIG_IFS"
+      splitarray "," "$OPTARG"; TAGS=("${SPLITARRAY[@]}")
       ;;
     \?)
       usage
--- a/lib/Tools/findlogics	Fri Sep 11 15:56:51 2009 +1000
+++ b/lib/Tools/findlogics	Fri Sep 11 15:57:16 2009 +1000
@@ -25,7 +25,7 @@
 declare -a LOGICS=()
 declare -a ISABELLE_PATHS=()
 
-ORIG_IFS="$IFS"; IFS=":"; ISABELLE_PATHS=($ISABELLE_PATH); IFS=$ORIG_IFS
+splitarray ":" "$ISABELLE_PATH"; ISABELLE_PATHS=("${SPLITARRAY[@]}")
 
 for DIR in "${ISABELLE_PATHS[@]}"
 do
@@ -34,7 +34,7 @@
   do
     if [ -f "$FILE" ]; then
       NAME=$(basename "$FILE")
-      LOGICS+=("$NAME")
+      LOGICS["${#LOGICS[@]}"]="$NAME"
     fi
   done
 done
--- a/lib/Tools/jedit	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,24 +0,0 @@
-#!/usr/bin/env bash
-#
-# Author: Makarius
-#
-# DESCRIPTION: Isabelle/jEdit interface wrapper
-
-
-## diagnostics
-
-function fail()
-{
-  echo "$1" >&2
-  exit 2
-}
-
-
-## main
-
-[ -z "$JEDIT_HOME" ] && fail "Missing Isabelle/jEdit installation (JEDIT_HOME)"
-
-INTERFACE="$JEDIT_HOME/interface"
-[ ! -x "$INTERFACE" ] && fail "Bad interface script: \"$INTERFACE\""
-
-exec "$INTERFACE" "$@"
--- a/lib/Tools/makeall	Fri Sep 11 15:56:51 2009 +1000
+++ b/lib/Tools/makeall	Fri Sep 11 15:57:16 2009 +1000
@@ -34,7 +34,7 @@
 echo "Started at $(date) ($ML_IDENTIFIER on $(hostname))"
 . "$ISABELLE_HOME/lib/scripts/timestart.bash"
 
-ORIG_IFS="$IFS"; IFS=":"; declare -a COMPONENTS=($ISABELLE_COMPONENTS); IFS="$ORIG_IFS"
+splitarray ":" "$ISABELLE_COMPONENTS"; COMPONENTS=("${SPLITARRAY[@]}")
 
 for DIR in "${COMPONENTS[@]}"
 do
--- a/lib/scripts/getsettings	Fri Sep 11 15:56:51 2009 +1000
+++ b/lib/scripts/getsettings	Fri Sep 11 15:57:16 2009 +1000
@@ -68,6 +68,17 @@
   done
 }
 
+#arrays
+function splitarray ()
+{
+  SPLITARRAY=()
+  local IFS="$1"; shift
+  for X in $*
+  do
+    SPLITARRAY["${#SPLITARRAY[@]}"]="$X"
+  done
+}
+
 #nested components
 ISABELLE_COMPONENTS=""
 function init_component ()
--- a/lib/scripts/mirabelle	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,129 +0,0 @@
-#!/usr/bin/perl -w
-
-use strict;
-use File::Basename;
-
-# Taken from http://www.skywayradio.com/tech/perl/trim_blanks.html
-sub trim {
-    my @out = @_;
-    for (@out) {
-        s/^\s+//;
-        s/\s+$//;
-    }
-    return wantarray ? @out : $out[0];
-}
-
-sub quote {
-    my $str = pop;
-    return "\"" . $str . "\"";
-}
-
-sub print_usage_and_quit {
-    print STDERR "Usage: mirabelle actions file1.thy...\n" .
-                 "  actions: action1:...:actionN\n" .
-                 "  action: name or name[key1=value1,...,keyM=valueM]\n";
-    exit 1;
-}
-
-my $num_args = $#ARGV + 1;
-if ($num_args < 2) {
-    print_usage_and_quit();
-}
-
-my @action_names;
-my @action_settings;
-
-foreach (split(/:/, $ARGV[0])) {
-    my %settings;
-
-    $_ =~ /([^[]*)(?:\[(.*)\])?/;
-    my ($name, $settings_str) = ($1, $2 || "");
-    my @setting_strs = split(/,/, $settings_str);
-    foreach (@setting_strs) {
-        $_ =~ /(.*)=(.*)/;
-	    my $key = $1;
-	    my $value = $2;
-	    $settings{trim($key)} = trim($value);
-    }
-
-    push @action_names, trim($name);
-    push @action_settings, \%settings;
-}
-
-my $output_path = "/tmp/mirabelle"; # FIXME: generate path
-my $mirabellesetup_thy_name = $output_path . "/MirabelleSetup";
-my $mirabellesetup_file = $mirabellesetup_thy_name . ".thy";
-my $mirabelle_log_file = $output_path . "/mirabelle.log";
-
-mkdir $output_path, 0755;
-
-open(FILE, ">$mirabellesetup_file")
-    || die "Could not create file '$mirabellesetup_file'";
-
-my $invoke_lines;
-
-for my $i (0 .. $#action_names) { 
-    my $settings_str = "";
-    my $settings = $action_settings[$i];
-    my $key;
-    my $value;
-
-    while (($key, $value) = each(%$settings)) {
-        $settings_str .= "(" . quote ($key) . ", " . quote ($value) . "), ";
-    }
-    $settings_str =~ s/, $//;
-
-    $invoke_lines .= "setup {* Mirabelle.invoke \"$action_names[$i]\" ";
-    $invoke_lines .= "[$settings_str] *}\n"
-}
-
-print FILE <<EOF;
-theory MirabelleSetup
-imports Mirabelle
-begin
-
-setup {* Mirabelle.set_logfile "$mirabelle_log_file" *}
-
-$invoke_lines
-
-end
-EOF
-
-my $root_text = "";
-my @new_thy_files;
-
-for my $i (1 .. $num_args - 1) {
-    my $old_thy_file = $ARGV[$i];
-    my ($base, $dir, $ext) = fileparse($old_thy_file, "\.thy");
-    my $new_thy_name = $base . "Mirabelle";
-    my $new_thy_file = $dir . $new_thy_name . $ext;
-
-    open(OLD_FILE, "<$old_thy_file")
-        || die "Cannot open file $old_thy_file";
-    my @lines = <OLD_FILE>;
-    close(OLD_FILE);
-
-    my $thy_text = join("", @lines);
-    my $old_len = length($thy_text);
-    $thy_text =~ s/\btheory\b[^\n]*\s*\bimports\s/theory $new_thy_name\nimports "$mirabellesetup_thy_name" /gm;
-    die "No 'imports' found" if length($thy_text) == $old_len;
-
-    open(NEW_FILE, ">$new_thy_file");
-    print NEW_FILE $thy_text;
-    close(NEW_FILE);
-
-    $root_text .= "use_thy \"" . $dir . $new_thy_name . "\";\n";
-
-    push @new_thy_files, $new_thy_file;
-}
-
-my $root_file = "ROOT_mirabelle.ML";
-open(ROOT_FILE, ">$root_file") || die "Cannot open file $root_file";
-print ROOT_FILE $root_text;
-close(ROOT_FILE);
-
-system "isabelle-process -e 'use \"ROOT_mirabelle.ML\";' -f -q HOL";
-
-# unlink $mirabellesetup_file;
-unlink $root_file;
-unlink @new_thy_files;
--- a/src/FOL/fologic.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/FOL/fologic.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -6,28 +6,28 @@
 
 signature FOLOGIC =
 sig
-  val oT		: typ
-  val mk_Trueprop	: term -> term
-  val dest_Trueprop	: term -> term
-  val not		: term
-  val conj		: term
-  val disj		: term
-  val imp		: term
-  val iff		: term
-  val mk_conj		: term * term -> term
-  val mk_disj		: term * term -> term
-  val mk_imp		: term * term -> term
-  val dest_imp	       	: term -> term*term
-  val dest_conj         : term -> term list
-  val mk_iff		: term * term -> term
-  val dest_iff	       	: term -> term*term
-  val all_const		: typ -> term
-  val mk_all		: term * term -> term
-  val exists_const	: typ -> term
-  val mk_exists		: term * term -> term
-  val eq_const		: typ -> term
-  val mk_eq		: term * term -> term
-  val dest_eq 		: term -> term*term
+  val oT: typ
+  val mk_Trueprop: term -> term
+  val dest_Trueprop: term -> term
+  val not: term
+  val conj: term
+  val disj: term
+  val imp: term
+  val iff: term
+  val mk_conj: term * term -> term
+  val mk_disj: term * term -> term
+  val mk_imp: term * term -> term
+  val dest_imp: term -> term * term
+  val dest_conj: term -> term list
+  val mk_iff: term * term -> term
+  val dest_iff: term -> term * term
+  val all_const: typ -> term
+  val mk_all: term * term -> term
+  val exists_const: typ -> term
+  val mk_exists: term * term -> term
+  val eq_const: typ -> term
+  val mk_eq: term * term -> term
+  val dest_eq: term -> term * term
   val mk_binop: string -> term * term -> term
   val mk_binrel: string -> term * term -> term
   val dest_bin: string -> typ -> term -> term * term
@@ -46,7 +46,8 @@
 fun dest_Trueprop (Const ("Trueprop", _) $ P) = P
   | dest_Trueprop t = raise TERM ("dest_Trueprop", [t]);
 
-(** Logical constants **)
+
+(* Logical constants *)
 
 val not = Const ("Not", oT --> oT);
 val conj = Const("op &", [oT,oT]--->oT);
@@ -80,6 +81,7 @@
 fun exists_const T = Const ("Ex", [T --> oT] ---> oT);
 fun mk_exists (Free(x,T),P) = exists_const T $ (absfree (x,T,P));
 
+
 (* binary oprations and relations *)
 
 fun mk_binop c (t, u) =
@@ -97,5 +99,4 @@
       else raise TERM ("dest_bin " ^ c, [tm])
   | dest_bin c _ tm = raise TERM ("dest_bin " ^ c, [tm]);
 
-
 end;
--- a/src/FOL/intprover.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/FOL/intprover.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -79,8 +79,7 @@
 (*One safe or unsafe step. *)
 fun step_tac i = FIRST [safe_tac, inst_step_tac i, biresolve_tac haz_brls i];
 
-fun step_dup_tac i = FIRST [safe_tac, inst_step_tac i, 
-			    biresolve_tac haz_dup_brls i];
+fun step_dup_tac i = FIRST [safe_tac, inst_step_tac i, biresolve_tac haz_dup_brls i];
 
 (*Dumb but fast*)
 val fast_tac = SELECT_GOAL (DEPTH_SOLVE (step_tac 1));
--- a/src/FOLP/hypsubst.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/FOLP/hypsubst.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -27,7 +27,7 @@
   val inspect_pair        : bool -> term * term -> thm
   end;
 
-functor HypsubstFun(Data: HYPSUBST_DATA): HYPSUBST = 
+functor HypsubstFun(Data: HYPSUBST_DATA): HYPSUBST =
 struct
 
 local open Data in
@@ -43,13 +43,13 @@
     but how could we check for this?*)
 fun inspect_pair bnd (t,u) =
   case (Envir.eta_contract t, Envir.eta_contract u) of
-       (Bound i, _) => if loose(i,u) then raise Match 
+       (Bound i, _) => if loose(i,u) then raise Match
                        else sym         (*eliminates t*)
-     | (_, Bound i) => if loose(i,t) then raise Match 
+     | (_, Bound i) => if loose(i,t) then raise Match
                        else asm_rl      (*eliminates u*)
-     | (Free _, _) => if bnd orelse Logic.occs(t,u) then raise Match 
+     | (Free _, _) => if bnd orelse Logic.occs(t,u) then raise Match
                       else sym          (*eliminates t*)
-     | (_, Free _) => if bnd orelse Logic.occs(u,t) then raise Match 
+     | (_, Free _) => if bnd orelse Logic.occs(u,t) then raise Match
                       else asm_rl       (*eliminates u*)
      | _ => raise Match;
 
@@ -58,7 +58,7 @@
    the rule asm_rl (resp. sym). *)
 fun eq_var bnd =
   let fun eq_var_aux k (Const("all",_) $ Abs(_,_,t)) = eq_var_aux k t
-        | eq_var_aux k (Const("==>",_) $ A $ B) = 
+        | eq_var_aux k (Const("==>",_) $ A $ B) =
               ((k, inspect_pair bnd (dest_eq A))
                       (*Exception Match comes from inspect_pair or dest_eq*)
                handle Match => eq_var_aux (k+1) B)
@@ -70,13 +70,13 @@
 fun gen_hyp_subst_tac bnd = SUBGOAL(fn (Bi,i) =>
       let val n = length(Logic.strip_assums_hyp Bi) - 1
           val (k,symopt) = eq_var bnd Bi
-      in 
+      in
          DETERM
            (EVERY [REPEAT_DETERM_N k (etac rev_mp i),
-		   etac revcut_rl i,
-		   REPEAT_DETERM_N (n-k) (etac rev_mp i),
-		   etac (symopt RS subst) i,
-		   REPEAT_DETERM_N n (rtac imp_intr i)])
+                   etac revcut_rl i,
+                   REPEAT_DETERM_N (n-k) (etac rev_mp i),
+                   etac (symopt RS subst) i,
+                   REPEAT_DETERM_N n (rtac imp_intr i)])
       end
       handle THM _ => no_tac | EQ_VAR => no_tac);
 
--- a/src/FOLP/simp.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/FOLP/simp.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -52,7 +52,7 @@
   val tracing   : bool ref
 end;
 
-functor SimpFun (Simp_data: SIMP_DATA) : SIMP = 
+functor SimpFun (Simp_data: SIMP_DATA) : SIMP =
 struct
 
 local open Simp_data in
@@ -74,12 +74,12 @@
   Similar to match_from_nat_tac, but the net does not contain numbers;
   rewrite rules are not ordered.*)
 fun net_tac net =
-  SUBGOAL(fn (prem,i) => 
+  SUBGOAL(fn (prem,i) =>
           resolve_tac (Net.unify_term net (Logic.strip_assums_concl prem)) i);
 
 (*match subgoal i against possible theorems indexed by lhs in the net*)
 fun lhs_net_tac net =
-  SUBGOAL(fn (prem,i) => 
+  SUBGOAL(fn (prem,i) =>
           biresolve_tac (Net.unify_term net
                        (lhs_of (Logic.strip_assums_concl prem))) i);
 
@@ -110,7 +110,7 @@
 
 (*Get the norm constants from norm_thms*)
 val norms =
-  let fun norm thm = 
+  let fun norm thm =
       case lhs_of(concl_of thm) of
           Const(n,_)$_ => n
         | _ => error "No constant in lhs of a norm_thm"
@@ -144,7 +144,7 @@
 (**** Adding "NORM" tags ****)
 
 (*get name of the constant from conclusion of a congruence rule*)
-fun cong_const cong = 
+fun cong_const cong =
     case head_of (lhs_of (concl_of cong)) of
         Const(c,_) => c
       | _ => ""                 (*a placeholder distinct from const names*);
@@ -156,9 +156,9 @@
 fun add_hidden_vars ccs =
   let fun add_hvars tm hvars = case tm of
               Abs(_,_,body) => OldTerm.add_term_vars(body,hvars)
-            | _$_ => let val (f,args) = strip_comb tm 
+            | _$_ => let val (f,args) = strip_comb tm
                      in case f of
-                            Const(c,T) => 
+                            Const(c,T) =>
                                 if member (op =) ccs c
                                 then fold_rev add_hvars args hvars
                                 else OldTerm.add_term_vars (tm, hvars)
@@ -202,13 +202,13 @@
     val hvs = map (#1 o dest_Var) hvars
     val refl1_tac = refl_tac 1
     fun norm_step_tac st = st |>
-	 (case head_of(rhs_of_eq 1 st) of
-	    Var(ixn,_) => if ixn mem hvs then refl1_tac
-			  else resolve_tac normI_thms 1 ORELSE refl1_tac
-	  | Const _ => resolve_tac normI_thms 1 ORELSE
-		       resolve_tac congs 1 ORELSE refl1_tac
-	  | Free _ => resolve_tac congs 1 ORELSE refl1_tac
-	  | _ => refl1_tac)
+         (case head_of(rhs_of_eq 1 st) of
+            Var(ixn,_) => if ixn mem hvs then refl1_tac
+                          else resolve_tac normI_thms 1 ORELSE refl1_tac
+          | Const _ => resolve_tac normI_thms 1 ORELSE
+                       resolve_tac congs 1 ORELSE refl1_tac
+          | Free _ => resolve_tac congs 1 ORELSE refl1_tac
+          | _ => refl1_tac)
     val add_norm_tac = DEPTH_FIRST (has_fewer_prems nops) norm_step_tac
     val SOME(thm'',_) = Seq.pull(add_norm_tac thm')
 in thm'' end;
@@ -246,9 +246,9 @@
 (** Insertion of congruences and rewrites **)
 
 (*insert a thm in a thm net*)
-fun insert_thm_warn th net = 
+fun insert_thm_warn th net =
   Net.insert_term Thm.eq_thm_prop (concl_of th, th) net
-  handle Net.INSERT => 
+  handle Net.INSERT =>
     (writeln ("Duplicate rewrite or congruence rule:\n" ^
         Display.string_of_thm_without_context th); net);
 
@@ -272,9 +272,9 @@
 (** Deletion of congruences and rewrites **)
 
 (*delete a thm from a thm net*)
-fun delete_thm_warn th net = 
+fun delete_thm_warn th net =
   Net.delete_term Thm.eq_thm_prop (concl_of th, th) net
-  handle Net.DELETE => 
+  handle Net.DELETE =>
     (writeln ("No such rewrite or congruence rule:\n" ^
         Display.string_of_thm_without_context th); net);
 
@@ -337,17 +337,17 @@
     in find_if(tm,0) end;
 
 fun IF1_TAC cong_tac i =
-    let fun seq_try (ifth::ifths,ifc::ifcs) thm = 
+    let fun seq_try (ifth::ifths,ifc::ifcs) thm =
                 (COND (if_rewritable ifc i) (DETERM(rtac ifth i))
                         (seq_try(ifths,ifcs))) thm
               | seq_try([],_) thm = no_tac thm
         and try_rew thm = (seq_try(case_rews,case_consts) ORELSE one_subt) thm
         and one_subt thm =
                 let val test = has_fewer_prems (nprems_of thm + 1)
-                    fun loop thm = 
-			COND test no_tac
+                    fun loop thm =
+                        COND test no_tac
                           ((try_rew THEN DEPTH_FIRST test (refl_tac i))
-			   ORELSE (refl_tac i THEN loop)) thm
+                           ORELSE (refl_tac i THEN loop)) thm
                 in (cong_tac THEN loop) thm end
     in COND (may_match(case_consts,i)) try_rew no_tac end;
 
@@ -381,12 +381,12 @@
 
 (*print lhs of conclusion of subgoal i*)
 fun pr_goal_lhs i st =
-    writeln (Syntax.string_of_term_global (Thm.theory_of_thm st) 
+    writeln (Syntax.string_of_term_global (Thm.theory_of_thm st)
              (lhs_of (prepare_goal i st)));
 
 (*print conclusion of subgoal i*)
 fun pr_goal_concl i st =
-    writeln (Syntax.string_of_term_global (Thm.theory_of_thm st) (prepare_goal i st)) 
+    writeln (Syntax.string_of_term_global (Thm.theory_of_thm st) (prepare_goal i st))
 
 (*print subgoals i to j (inclusive)*)
 fun pr_goals (i,j) st =
@@ -439,7 +439,7 @@
         then writeln (cat_lines
           ("Adding rewrites:" :: map Display.string_of_thm_without_context new_rws))
         else ();
-        (ss,thm,anet',anet::ats,cs) 
+        (ss,thm,anet',anet::ats,cs)
     end;
 
 fun rew(seq,thm,ss,anet,ats,cs, more) = case Seq.pull seq of
@@ -492,7 +492,7 @@
 
 fun EXEC_TAC(ss,fl) (SS{auto_tac,cong_net,simp_net,...}) =
 let val cong_tac = net_tac cong_net
-in fn i => 
+in fn i =>
     (fn thm =>
      if i <= 0 orelse nprems_of thm < i then Seq.empty
      else Seq.single(execute(ss,fl,auto_tac,cong_tac,simp_net,i,thm)))
--- a/src/HOL/Algebra/Divisibility.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Algebra/Divisibility.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -2656,25 +2656,7 @@
   shows "(x \<in> carrier G \<and> x gcdof a b) =
          greatest (division_rel G) x (Lower (division_rel G) {a, b})"
 unfolding isgcd_def greatest_def Lower_def elem_def
-proof (simp, safe)
-  fix xa
-  assume r1[rule_format]: "\<forall>x. (x = a \<or> x = b) \<and> x \<in> carrier G \<longrightarrow> xa divides x"
-  assume r2[rule_format]: "\<forall>y\<in>carrier G. y divides a \<and> y divides b \<longrightarrow> y divides x"
-
-  assume "xa \<in> carrier G"  "x divides a"  "x divides b"
-  with carr
-  show "xa divides x"
-      by (fast intro: r1 r2)
-next
-  fix a' y
-  assume r1[rule_format]:
-         "\<forall>xa\<in>{l. \<forall>x. (x = a \<or> x = b) \<and> x \<in> carrier G \<longrightarrow> l divides x} \<inter> carrier G.
-           xa divides x"
-  assume "y \<in> carrier G"  "y divides a"  "y divides b"
-  with carr
-       show "y divides x"
-       by (fast intro: r1)
-qed (simp, simp)
+by auto
 
 lemma lcmof_leastUpper:
   fixes G (structure)
@@ -2682,25 +2664,7 @@
   shows "(x \<in> carrier G \<and> x lcmof a b) =
          least (division_rel G) x (Upper (division_rel G) {a, b})"
 unfolding islcm_def least_def Upper_def elem_def
-proof (simp, safe)
-  fix xa
-  assume r1[rule_format]: "\<forall>x. (x = a \<or> x = b) \<and> x \<in> carrier G \<longrightarrow> x divides xa"
-  assume r2[rule_format]: "\<forall>y\<in>carrier G. a divides y \<and> b divides y \<longrightarrow> x divides y"
-
-  assume "xa \<in> carrier G"  "a divides x"  "b divides x"
-  with carr
-  show "x divides xa"
-      by (fast intro: r1 r2)
-next
-  fix a' y
-  assume r1[rule_format]:
-         "\<forall>xa\<in>{l. \<forall>x. (x = a \<or> x = b) \<and> x \<in> carrier G \<longrightarrow> x divides l} \<inter> carrier G.
-           x divides xa"
-  assume "y \<in> carrier G"  "a divides y"  "b divides y"
-  with carr
-       show "x divides y"
-       by (fast intro: r1)
-qed (simp, simp)
+by auto
 
 lemma somegcd_meet:
   fixes G (structure)
--- a/src/HOL/Algebra/Exponent.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Algebra/Exponent.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -1,16 +1,13 @@
 (*  Title:      HOL/Algebra/Exponent.thy
-    ID:         $Id$
     Author:     Florian Kammueller, with new proofs by L C Paulson
 
     exponent p s   yields the greatest power of p that divides s.
 *)
 
 theory Exponent
-imports Main Primes Binomial
+imports Main "~~/src/HOL/Old_Number_Theory/Primes" Binomial
 begin
 
-hide (open) const GCD.gcd GCD.coprime GCD.prime
-
 section {*Sylow's Theorem*}
 
 subsection {*The Combinatorial Argument Underlying the First Sylow Theorem*}
--- a/src/HOL/Algebra/IntRing.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Algebra/IntRing.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -4,7 +4,7 @@
 *)
 
 theory IntRing
-imports QuotRing Lattice Int Primes
+imports QuotRing Lattice Int "~~/src/HOL/Old_Number_Theory/Primes"
 begin
 
 
--- a/src/HOL/Algebra/ROOT.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Algebra/ROOT.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -5,7 +5,7 @@
 *)
 
 (* Preliminaries from set and number theory *)
-no_document use_thys ["FuncSet", "Primes", "Binomial", "Permutation"];
+no_document use_thys ["FuncSet", "~~/src/HOL/Old_Number_Theory/Primes", "Binomial", "Permutation"];
 
 
 (*** New development, based on explicit structures ***)
--- a/src/HOL/Algebra/UnivPoly.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Algebra/UnivPoly.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -592,15 +592,14 @@
         proof (cases "n = k")
           case True
           then have "\<zero> = (\<Oplus>i \<in> {..<n} \<union> {n}. ?s i)"
-            by (simp cong: R.finsum_cong add: ivl_disj_int_singleton Pi_def)
+            by (simp cong: R.finsum_cong add: Pi_def)
           also from True have "... = (\<Oplus>i \<in> {..k}. ?s i)"
             by (simp only: ivl_disj_un_singleton)
           finally show ?thesis .
         next
           case False with n_le_k have n_less_k: "n < k" by arith
           with neq have "\<zero> = (\<Oplus>i \<in> {..<n} \<union> {n}. ?s i)"
-            by (simp add: R.finsum_Un_disjoint f1 f2
-              ivl_disj_int_singleton Pi_def del: Un_insert_right)
+            by (simp add: R.finsum_Un_disjoint f1 f2 Pi_def del: Un_insert_right)
           also have "... = (\<Oplus>i \<in> {..n}. ?s i)"
             by (simp only: ivl_disj_un_singleton)
           also from n_less_k neq have "... = (\<Oplus>i \<in> {..n} \<union> {n<..k}. ?s i)"
@@ -817,15 +816,9 @@
 text {* Degree and polynomial operations *}
 
 lemma deg_add [simp]:
-  assumes R: "p \<in> carrier P" "q \<in> carrier P"
-  shows "deg R (p \<oplus>\<^bsub>P\<^esub> q) <= max (deg R p) (deg R q)"
-proof (cases "deg R p <= deg R q")
-  case True show ?thesis
-    by (rule deg_aboveI) (simp_all add: True R deg_aboveD)
-next
-  case False show ?thesis
-    by (rule deg_aboveI) (simp_all add: False R deg_aboveD)
-qed
+  "p \<in> carrier P \<Longrightarrow> q \<in> carrier P \<Longrightarrow>
+  deg R (p \<oplus>\<^bsub>P\<^esub> q) <= max (deg R p) (deg R q)"
+by(rule deg_aboveI)(simp_all add: deg_aboveD)
 
 lemma deg_monom_le:
   "a \<in> carrier R ==> deg R (monom P a n) <= n"
@@ -945,8 +938,7 @@
     also have "...= (\<Oplus>i \<in> {deg R p} \<union> {deg R p <.. deg R p + deg R q}. ?s i)"
       by (simp only: ivl_disj_un_singleton)
     also have "... = coeff P p (deg R p) \<otimes> coeff P q (deg R q)"
-      by (simp cong: R.finsum_cong
-	add: ivl_disj_int_singleton deg_aboveD R Pi_def)
+      by (simp cong: R.finsum_cong add: deg_aboveD R Pi_def)
     finally have "(\<Oplus>i \<in> {.. deg R p + deg R q}. ?s i)
       = coeff P p (deg R p) \<otimes> coeff P q (deg R q)" .
     with nz show "(\<Oplus>i \<in> {.. deg R p + deg R q}. ?s i) ~= \<zero>"
@@ -989,8 +981,7 @@
     have "... = coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..<k} \<union> {k}. ?s i) k"
       by (simp only: ivl_disj_un_singleton)
     also have "... = coeff P p k"
-      by (simp cong: R.finsum_cong
-	add: ivl_disj_int_singleton coeff_finsum deg_aboveD R RR Pi_def)
+      by (simp cong: R.finsum_cong add: coeff_finsum deg_aboveD R RR Pi_def)
     finally show ?thesis .
   next
     case False
@@ -998,8 +989,7 @@
           coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..<deg R p} \<union> {deg R p}. ?s i) k"
       by (simp only: ivl_disj_un_singleton)
     also from False have "... = coeff P p k"
-      by (simp cong: R.finsum_cong
-	add: ivl_disj_int_singleton coeff_finsum deg_aboveD R Pi_def)
+      by (simp cong: R.finsum_cong add: coeff_finsum deg_aboveD R Pi_def)
     finally show ?thesis .
   qed
 qed (simp_all add: R Pi_def)
--- a/src/HOL/Algebra/abstract/Ring2.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Algebra/abstract/Ring2.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -241,7 +241,7 @@
 proof (induct n)
   case 0 show ?case by simp
 next
-  case Suc thus ?case by (simp add: add_assoc) 
+  case Suc thus ?case by (simp add: add_assoc)
 qed
 
 lemma natsum_cong [cong]:
@@ -269,21 +269,21 @@
 
 ML {*
   local
-    val lhss = 
+    val lhss =
         ["t + u::'a::ring",
-	 "t - u::'a::ring",
-	 "t * u::'a::ring",
-	 "- t::'a::ring"];
-    fun proc ss t = 
+         "t - u::'a::ring",
+         "t * u::'a::ring",
+         "- t::'a::ring"];
+    fun proc ss t =
       let val rew = Goal.prove (Simplifier.the_context ss) [] []
             (HOLogic.mk_Trueprop
               (HOLogic.mk_eq (t, Var (("x", Term.maxidx_of_term t + 1), fastype_of t))))
                 (fn _ => simp_tac (Simplifier.inherit_context ss ring_ss) 1)
             |> mk_meta_eq;
           val (t', u) = Logic.dest_equals (Thm.prop_of rew);
-      in if t' aconv u 
+      in if t' aconv u
         then NONE
-        else SOME rew 
+        else SOME rew
     end;
   in
     val ring_simproc = Simplifier.simproc @{theory} "ring" lhss (K proc);
@@ -305,7 +305,7 @@
 declare one_not_zero [simp]
 
 lemma zero_not_one [simp]:
-  "0 ~= (1::'a::domain)" 
+  "0 ~= (1::'a::domain)"
 by (rule not_sym) simp
 
 lemma integral_iff: (* not by default a simp rule! *)
@@ -322,7 +322,7 @@
 *)
 (*
 lemma bug: "(b::'a::ring) - (b - a) = a" by simp
-   simproc for rings cannot prove "(a::'a::ring) - (a - b) = b" 
+   simproc for rings cannot prove "(a::'a::ring) - (a - b) = b"
 *)
 lemma m_lcancel:
   assumes prem: "(a::'a::domain) ~= 0" shows conc: "(a * b = a * c) = (b = c)"
@@ -330,8 +330,8 @@
   assume eq: "a * b = a * c"
   then have "a * (b - c) = 0" by simp
   then have "a = 0 | (b - c) = 0" by (simp only: integral_iff)
-  with prem have "b - c = 0" by auto 
-  then have "b = b - (b - c)" by simp 
+  with prem have "b - c = 0" by auto
+  then have "b = b - (b - c)" by simp
   also have "b - (b - c) = c" by simp
   finally show "b = c" .
 next
@@ -386,7 +386,7 @@
 qed
 
 
-lemma unit_mult: 
+lemma unit_mult:
   "!!a::'a::ring. [| a dvd 1; b dvd 1 |] ==> a * b dvd 1"
   apply (unfold dvd_def)
   apply clarify
--- a/src/HOL/Algebra/poly/UnivPoly2.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Algebra/poly/UnivPoly2.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -563,11 +563,7 @@
 
 lemma deg_add [simp]:
   "deg ((p::'a::ring up) + q) <= max (deg p) (deg q)"
-proof (cases "deg p <= deg q")
-  case True show ?thesis by (rule deg_aboveI) (simp add: True deg_aboveD) 
-next
-  case False show ?thesis by (rule deg_aboveI) (simp add: False deg_aboveD)
-qed
+by (rule deg_aboveI) (simp add: deg_aboveD)
 
 lemma deg_monom_ring:
   "deg (monom a n::'a::ring up) <= n"
@@ -678,8 +674,7 @@
     also have "... = setsum ?s ({deg p} Un {deg p <.. deg p + deg q})"
       by (simp only: ivl_disj_un_singleton)
     also have "... = coeff p (deg p) * coeff q (deg q)" 
-      by (simp add: setsum_Un_disjoint ivl_disj_int_singleton 
-        setsum_0 deg_aboveD)
+      by (simp add: setsum_Un_disjoint setsum_0 deg_aboveD)
     finally have "setsum ?s {.. deg p + deg q} 
       = coeff p (deg p) * coeff q (deg q)" .
     with nz show "setsum ?s {.. deg p + deg q} ~= 0"
@@ -723,8 +718,7 @@
     have "... = coeff (setsum ?s ({..<k} Un {k})) k"
       by (simp only: ivl_disj_un_singleton)
     also have "... = coeff p k"
-      by (simp add: setsum_Un_disjoint ivl_disj_int_singleton 
-        setsum_0 coeff_natsum deg_aboveD)
+      by (simp add: setsum_Un_disjoint setsum_0 coeff_natsum deg_aboveD)
     finally show ?thesis .
   next
     case False
@@ -732,8 +726,7 @@
           coeff (setsum ?s ({..<deg p} Un {deg p})) k"
       by (simp only: ivl_disj_un_singleton)
     also from False have "... = coeff p k"
-      by (simp add: setsum_Un_disjoint ivl_disj_int_singleton 
-        setsum_0 coeff_natsum deg_aboveD)
+      by (simp add: setsum_Un_disjoint setsum_0 coeff_natsum deg_aboveD)
     finally show ?thesis .
   qed
 qed
--- a/src/HOL/Auth/Event.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Auth/Event.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -139,9 +139,11 @@
 
 text{*Elimination rules: derive contradictions from old Says events containing
   items known to be fresh*}
+lemmas Says_imp_parts_knows_Spy = 
+       Says_imp_knows_Spy [THEN parts.Inj, THEN revcut_rl, standard] 
+
 lemmas knows_Spy_partsEs =
-     Says_imp_knows_Spy [THEN parts.Inj, THEN revcut_rl, standard] 
-     parts.Body [THEN revcut_rl, standard]
+     Says_imp_parts_knows_Spy parts.Body [THEN revcut_rl, standard]
 
 lemmas Says_imp_analz_Spy = Says_imp_knows_Spy [THEN analz.Inj]
 
--- a/src/HOL/Auth/KerberosIV.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Auth/KerberosIV.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -899,7 +899,6 @@
 apply (frule_tac [7] Says_ticket_parts)
 apply (simp_all (no_asm_simp))
 apply blast
-atp_minimize [atp=spass] Crypt_imp_invKey_keysFor invKey_K new_keys_not_used
 apply (metis Crypt_imp_invKey_keysFor invKey_K new_keys_not_used)
 apply (clarify)
 apply (frule Says_Tgs_message_form, assumption)
@@ -1316,7 +1315,6 @@
 txt{*K4*}
 apply blast
 txt{*Level 8: K5*}
-atp_minimize [atp=e] Tgs_not_bad authKeysI less_SucI mem_def nat_add_commute servK_notin_authKeysD spies_partsEs(1)
 apply (blast dest: servK_notin_authKeysD Says_Kas_message_form intro: less_SucI)
 txt{*Oops1*}
 apply (blast dest!: unique_authKeys intro: less_SucI)
--- a/src/HOL/Auth/KerberosV.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Auth/KerberosV.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -697,9 +697,7 @@
 txt{*K4*}
 apply (force dest!: Crypt_imp_keysFor, clarify)
 txt{*K6*}
-apply (drule  Says_imp_spies [THEN parts.Inj, THEN parts.Fst])
-apply (drule  Says_imp_spies [THEN parts.Inj, THEN parts.Snd])
-apply (blast dest!: unique_CryptKey)
+apply (metis Says_imp_spies Says_ticket_parts analz.Fst analz.Inj analz_conj_parts unique_CryptKey)
 done
 
 text{*Needs a unicity theorem, hence moved here*}
@@ -841,13 +839,10 @@
 apply (erule kerbV.induct, analz_mono_contra)
 apply (frule_tac [7] Says_ticket_parts)
 apply (frule_tac [5] Says_ticket_parts, simp_all, blast)
-txt{*K4 splits into distinct subcases*}
-apply auto
-txt{*servK can't have been enclosed in two certificates*}
- prefer 2 apply (blast dest: unique_CryptKey)
-txt{*servK is fresh and so could not have been used, by
-   @{text new_keys_not_used}*}
-apply (force dest!: Crypt_imp_invKey_keysFor simp add: AKcryptSK_def)
+txt{*K4*}
+apply (metis Auth_fresh_not_AKcryptSK Crypt_imp_invKey_keysFor Says_ticket_analz
+         analz.Fst invKey_K new_keys_not_analzd parts.Fst Says_imp_parts_knows_Spy
+         unique_CryptKey)
 done
 
 text{*Long term keys are not issued as servKeys*}
@@ -981,9 +976,7 @@
 txt{*K4*}
 apply (blast dest!: authK_not_AKcryptSK)
 txt{*Oops1*}
-apply clarify
-apply simp
-apply (blast dest!: AKcryptSK_analz_insert)
+apply (metis AKcryptSK_analz_insert insert_Key_singleton)
 done
 
 text{* First simplification law for analz: no session keys encrypt
@@ -1039,8 +1032,8 @@
         \<in> set evs;  authK \<in> symKeys;
          Key authK \<in> analz (spies evs); evs \<in> kerbV \<rbrakk>
       \<Longrightarrow> Key servK \<in> analz (spies evs)"
-apply (force dest: Says_imp_spies [THEN analz.Inj, THEN analz.Fst, THEN analz.Decrypt, THEN analz.Fst])
-done
+  by (metis Says_imp_analz_Spy analz.Fst analz_Decrypt')
+
 
 text{*lemma @{text servK_notin_authKeysD} not needed in version V*}
 
@@ -1112,16 +1105,16 @@
 apply (frule_tac [5] Says_ticket_analz)
 apply (safe del: impI conjI impCE)
 apply (simp_all add: less_SucI new_keys_not_analzd Says_Kas_message_form Says_Tgs_message_form analz_insert_eq not_parts_not_analz analz_insert_freshK1 analz_insert_freshK2 analz_insert_freshK3_bis pushes)
-txt{*Fake*}
-apply spy_analz
-txt{*K2*}
-apply (blast intro: parts_insertI less_SucI)
-txt{*K4*}
-apply (blast dest: authTicket_authentic Confidentiality_Kas)
-txt{*Oops1*}
+    txt{*Fake*}
+    apply spy_analz
+   txt{*K2*}
+   apply (blast intro: parts_insertI less_SucI)
+  txt{*K4*}
+  apply (blast dest: authTicket_authentic Confidentiality_Kas)
+ txt{*Oops1*}
  apply (blast dest: Says_Kas_message_form Says_Tgs_message_form intro: less_SucI)
 txt{*Oops2*}
-  apply (blast dest: Says_imp_spies [THEN parts.Inj] Key_unique_SesKey intro: less_SucI)
+apply (metis Suc_le_eq linorder_linear linorder_not_le msg.simps(2) unique_servKeys)
 done
 
 
@@ -1270,17 +1263,7 @@
          Key authK \<notin> analz (spies evs); Key servK \<notin> analz (spies evs);
          A \<notin> bad;  B \<notin> bad; evs \<in> kerbV \<rbrakk>
       \<Longrightarrow> Says B A (Crypt servK (Number T3)) \<in> set evs"
-apply (frule authK_authentic)
-apply assumption+
-apply (frule servK_authentic)
-prefer 2 apply (blast dest: authK_authentic Says_Kas_message_form)
-apply assumption+
-apply clarify
-apply (blast dest: K4_imp_K2 Key_unique_SesKey intro!: Says_K6)
-(*Single command proof: much slower!
-apply (blast dest: authK_authentic servK_authentic Says_Kas_message_form Key_unique_SesKey K4_imp_K2 intro!: Says_K6)
-*)
-done
+  by (metis authK_authentic Oops_range_spies1 Says_K6 servK_authentic u_K4_imp_K2 unique_authKeys)
 
 lemma A_authenticates_B_r:
      "\<lbrakk> Crypt servK (Number T3) \<in> parts (spies evs);
@@ -1301,8 +1284,7 @@
 apply (erule_tac [9] exE)
 apply (frule_tac [9] K4_imp_K2)
 apply assumption+
-apply (blast dest: Key_unique_SesKey intro!: Says_K6 dest: Confidentiality_Tgs
-)
+apply (blast dest: Key_unique_SesKey intro!: Says_K6 dest: Confidentiality_Tgs)
 done
 
 
@@ -1478,7 +1460,7 @@
 ...expands as follows - including extra exE because of new form of lemmas*)
 apply (frule K3_imp_K2, assumption, assumption, erule exE, erule exE)
 apply (case_tac "Key authK \<in> analz (spies evs5)")
-apply (drule Says_imp_knows_Spy [THEN analz.Inj, THEN analz.Fst, THEN analz_Decrypt', THEN analz.Fst], assumption, assumption, simp)
+ apply (metis Says_imp_analz_Spy analz.Fst analz_Decrypt')
 apply (frule K3_imp_K2, assumption, assumption, erule exE, erule exE)
 apply (drule Says_imp_knows_Spy [THEN parts.Inj, THEN parts.Fst])
 apply (frule servK_authentic_ter, blast, assumption+)
--- a/src/HOL/Auth/Kerberos_BAN.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Auth/Kerberos_BAN.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -288,15 +288,8 @@
                   on evs)"
 apply (unfold before_def)
 apply (erule rev_mp)
-apply (erule bankerberos.induct, simp_all)
-txt{*We need this simplification only for Message 2*}
-apply (simp (no_asm) add: takeWhile_tail)
-apply auto
-txt{*Two subcases of Message 2. Subcase: used before*}
-apply (blast dest: used_evs_rev [THEN equalityD2, THEN contra_subsetD] 
-                   used_takeWhile_used)
-txt{*subcase: CT before*}
-apply (fastsimp dest!: set_evs_rev [THEN equalityD2, THEN contra_subsetD, THEN takeWhile_void])
+apply (erule bankerberos.induct, simp_all add: takeWhile_tail)
+apply (metis length_rev set_rev takeWhile_void used_evs_rev)
 done
 
 
@@ -492,6 +485,7 @@
 txt{*BK3*}
 apply (blast dest: Kab_authentic unique_session_keys)
 done
+
 lemma lemma_B [rule_format]:
      "\<lbrakk> B \<notin> bad;  evs \<in> bankerberos \<rbrakk>
       \<Longrightarrow> Key K \<notin> analz (spies evs) \<longrightarrow>
@@ -585,9 +579,8 @@
 txt{*BK2*}
 apply (blast intro: parts_insertI less_SucI)
 txt{*BK3*}
-apply (case_tac "Aa \<in> bad")
- prefer 2 apply (blast dest: Kab_authentic unique_session_keys)
-apply (blast dest: Says_imp_spies [THEN analz.Inj] Crypt_Spy_analz_bad elim!: MPair_analz intro: less_SucI)
+apply (metis Crypt_Spy_analz_bad Kab_authentic Says_imp_analz_Spy 
+          Says_imp_parts_knows_Spy analz.Snd less_SucI unique_session_keys)
 txt{*Oops: PROOF FAILS if unsafe intro below*}
 apply (blast dest: unique_session_keys intro!: less_SucI)
 done
--- a/src/HOL/Auth/NS_Shared.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Auth/NS_Shared.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -273,11 +273,11 @@
 apply (simp_all add: analz_insert_eq analz_insert_freshK pushes split_ifs, spy_analz)
 txt{*NS2*}
 apply blast
-txt{*NS3, Server sub-case*}
+txt{*NS3*}
 apply (blast dest!: Crypt_Spy_analz_bad A_trusts_NS2
 	     dest:  Says_imp_knows_Spy analz.Inj unique_session_keys)
-txt{*NS3, Spy sub-case; also Oops*}
-apply (blast dest: unique_session_keys)+
+txt{*Oops*}
+apply (blast dest: unique_session_keys)
 done
 
 
@@ -318,9 +318,7 @@
     @{term "Crypt K (Nonce NB) \<in> parts (spies evs2)"} *} 
 apply (force dest!: Crypt_imp_keysFor)
 txt{*NS4*}
-apply (blast dest: B_trusts_NS3
-	           Says_imp_knows_Spy [THEN analz.Inj]
-                   Crypt_Spy_analz_bad unique_session_keys)
+apply (metis B_trusts_NS3 Crypt_Spy_analz_bad Says_imp_analz_Spy Says_imp_parts_knows_Spy analz.Fst unique_session_keys)
 done
 
 text{*This version no longer assumes that K is secure*}
@@ -349,9 +347,7 @@
 txt{*NS2*}
 apply (blast dest!: new_keys_not_used Crypt_imp_keysFor)
 txt{*NS4*}
-apply (blast dest: B_trusts_NS3
-	     dest: Says_imp_knows_Spy [THEN analz.Inj]
-                   unique_session_keys Crypt_Spy_analz_bad)
+apply (metis B_trusts_NS3 Crypt_Spy_analz_bad Says_imp_analz_Spy Says_imp_parts_knows_Spy analz.Fst unique_session_keys)
 done
 
 
@@ -475,18 +471,15 @@
 apply (erule rev_mp)
 apply (erule rev_mp)
 apply (erule ns_shared.induct, analz_mono_contra)
-apply (frule_tac [5] Says_S_message_form)
 apply (simp_all)
 txt{*Fake*}
 apply blast
 txt{*NS2*}
 apply (force dest!: Crypt_imp_keysFor)
-txt{*NS3, much quicker having installed @{term Says_S_message_form} before simplication*}
-apply fastsimp
+txt{*NS3*}
+apply (metis NS3_msg_in_parts_spies parts_cut_eq)
 txt{*NS5, the most important case, can be solved by unicity*}
-apply (case_tac "Aa \<in> bad")
-apply (force dest!: Says_imp_spies [THEN analz.Inj, THEN analz.Decrypt, THEN analz.Snd, THEN analz.Snd, THEN analz.Fst])
-apply (blast dest: A_trusts_NS2 unique_session_keys)
+apply (metis A_trusts_NS2 Crypt_Spy_analz_bad Says_imp_analz_Spy Says_imp_parts_knows_Spy analz.Fst analz.Snd unique_session_keys)
 done
 
 lemma A_Issues_B:
--- a/src/HOL/Auth/Recur.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Auth/Recur.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -419,15 +419,10 @@
 apply spy_analz
 txt{*RA2*}
 apply blast 
-txt{*RA3 remains*}
+txt{*RA3*}
 apply (simp add: parts_insert_spies)
-txt{*Now we split into two cases.  A single blast could do it, but it would take
-  a CPU minute.*}
-apply (safe del: impCE)
-txt{*RA3, case 1: use lemma previously proved by induction*}
-apply (blast elim: rev_notE [OF _ respond_Spy_not_see_session_key])
-txt{*RA3, case 2: K is an old key*}
-apply (blast dest: resp_analz_insert dest: Key_in_parts_respond)
+apply (metis Key_in_parts_respond parts.Body parts.Fst resp_analz_insert 
+             respond_Spy_not_see_session_key usedI)
 txt{*RA4*}
 apply blast 
 done
--- a/src/HOL/Bali/Example.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Bali/Example.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -1167,7 +1167,6 @@
 apply    (simp,rule assigned.select_convs)
 apply   (simp)
 apply  simp
-apply  blast
 apply simp
 apply (simp add: intersect_ts_def)
 done
--- a/src/HOL/Complete_Lattice.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Complete_Lattice.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -76,11 +76,11 @@
 
 lemma sup_bot [simp]:
   "x \<squnion> bot = x"
-  using bot_least [of x] by (simp add: le_iff_sup sup_commute)
+  using bot_least [of x] by (simp add: sup_commute)
 
 lemma inf_top [simp]:
   "x \<sqinter> top = x"
-  using top_greatest [of x] by (simp add: le_iff_inf inf_commute)
+  using top_greatest [of x] by (simp add: inf_commute)
 
 definition SUPR :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
   "SUPR A f = \<Squnion> (f ` A)"
--- a/src/HOL/Decision_Procs/Approximation.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Decision_Procs/Approximation.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -1904,7 +1904,7 @@
 	show "0 < real x * 2/3" using * by auto
 	show "real ?max + 1 \<le> real x * 2/3" using * up
 	  by (cases "0 < real x * real (lapprox_posrat prec 2 3) - 1",
-	      auto simp add: real_of_float_max max_def)
+	      auto simp add: real_of_float_max)
       qed
       finally have "real (?lb_horner (Float 1 -1)) + real (?lb_horner ?max)
 	\<le> ln (real x)"
--- a/src/HOL/Decision_Procs/Ferrack.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Decision_Procs/Ferrack.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -512,7 +512,7 @@
   assumes g0: "numgcd t = 0"
   shows "Inum bs t = 0"
   using g0[simplified numgcd_def] 
-  by (induct t rule: numgcdh.induct, auto simp add: natabs0 max_def maxcoeff_pos)
+  by (induct t rule: numgcdh.induct, auto simp add: natabs0 maxcoeff_pos)
 
 lemma numgcdh_pos: assumes gp: "g \<ge> 0" shows "numgcdh t g \<ge> 0"
   using gp
--- a/src/HOL/Extraction/Euclid.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Extraction/Euclid.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Extraction/Euclid.thy
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
                 Freek Wiedijk, Radboud University Nijmegen
                 Stefan Berghofer, TU Muenchen
@@ -8,7 +7,7 @@
 header {* Euclid's theorem *}
 
 theory Euclid
-imports "~~/src/HOL/NumberTheory/Factorization" Util Efficient_Nat
+imports "~~/src/HOL/Old_Number_Theory/Factorization" Util Efficient_Nat
 begin
 
 text {*
--- a/src/HOL/Extraction/ROOT.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Extraction/ROOT.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Extraction/ROOT.ML
-    ID:         $Id$
 
 Examples for program extraction in Higher-Order Logic.
 *)
@@ -8,5 +7,5 @@
   warning "HOL proof terms required for running extraction examples"
 else
   (Proofterm.proofs := 2;
-   no_document use_thys ["Efficient_Nat", "~~/src/HOL/NumberTheory/Factorization"];
+   no_document use_thys ["Efficient_Nat", "~~/src/HOL/Old_Number_Theory/Factorization"];
    use_thys ["Greatest_Common_Divisor", "Warshall", "Higman", "Pigeonhole", "Euclid"]);
--- a/src/HOL/Finite_Set.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Finite_Set.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -2966,11 +2966,11 @@
 
 lemma dual_max:
   "ord.max (op \<ge>) = min"
-  by (auto simp add: ord.max_def_raw min_def_raw expand_fun_eq)
+  by (auto simp add: ord.max_def_raw expand_fun_eq)
 
 lemma dual_min:
   "ord.min (op \<ge>) = max"
-  by (auto simp add: ord.min_def_raw max_def_raw expand_fun_eq)
+  by (auto simp add: ord.min_def_raw expand_fun_eq)
 
 lemma strict_below_fold1_iff:
   assumes "finite A" and "A \<noteq> {}"
--- a/src/HOL/GCD.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/GCD.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -1,11 +1,9 @@
-(*  Title:      GCD.thy
-    Authors:    Christophe Tabacznyj, Lawrence C. Paulson, Amine Chaieb,
+(*  Authors:    Christophe Tabacznyj, Lawrence C. Paulson, Amine Chaieb,
                 Thomas M. Rasmussen, Jeremy Avigad, Tobias Nipkow
 
 
-This file deals with the functions gcd and lcm, and properties of
-primes. Definitions and lemmas are proved uniformly for the natural
-numbers and integers.
+This file deals with the functions gcd and lcm.  Definitions and
+lemmas are proved uniformly for the natural numbers and integers.
 
 This file combines and revises a number of prior developments.
 
@@ -52,11 +50,6 @@
 
 end
 
-class prime = one +
-
-fixes
-  prime :: "'a \<Rightarrow> bool"
-
 
 (* definitions for the natural numbers *)
 
@@ -80,20 +73,6 @@
 end
 
 
-instantiation nat :: prime
-
-begin
-
-definition
-  prime_nat :: "nat \<Rightarrow> bool"
-where
-  [code del]: "prime_nat p = (1 < p \<and> (\<forall>m. m dvd p --> m = 1 \<or> m = p))"
-
-instance proof qed
-
-end
-
-
 (* definitions for the integers *)
 
 instantiation int :: gcd
@@ -115,28 +94,13 @@
 end
 
 
-instantiation int :: prime
-
-begin
-
-definition
-  prime_int :: "int \<Rightarrow> bool"
-where
-  [code del]: "prime_int p = prime (nat p)"
-
-instance proof qed
-
-end
-
-
 subsection {* Set up Transfer *}
 
 
 lemma transfer_nat_int_gcd:
   "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> gcd (nat x) (nat y) = nat (gcd x y)"
   "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> lcm (nat x) (nat y) = nat (lcm x y)"
-  "(x::int) >= 0 \<Longrightarrow> prime (nat x) = prime x"
-  unfolding gcd_int_def lcm_int_def prime_int_def
+  unfolding gcd_int_def lcm_int_def
   by auto
 
 lemma transfer_nat_int_gcd_closures:
@@ -150,8 +114,7 @@
 lemma transfer_int_nat_gcd:
   "gcd (int x) (int y) = int (gcd x y)"
   "lcm (int x) (int y) = int (lcm x y)"
-  "prime (int x) = prime x"
-  by (unfold gcd_int_def lcm_int_def prime_int_def, auto)
+  by (unfold gcd_int_def lcm_int_def, auto)
 
 lemma transfer_int_nat_gcd_closures:
   "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> gcd x y >= 0"
@@ -1003,20 +966,6 @@
   apply (auto simp add: gcd_mult_cancel_int)
 done
 
-lemma prime_odd_nat: "prime (p::nat) \<Longrightarrow> p > 2 \<Longrightarrow> odd p"
-  unfolding prime_nat_def
-  apply (subst even_mult_two_ex)
-  apply clarify
-  apply (drule_tac x = 2 in spec)
-  apply auto
-done
-
-lemma prime_odd_int: "prime (p::int) \<Longrightarrow> p > 2 \<Longrightarrow> odd p"
-  unfolding prime_int_def
-  apply (frule prime_odd_nat)
-  apply (auto simp add: even_nat_def)
-done
-
 lemma coprime_common_divisor_nat: "coprime (a::nat) b \<Longrightarrow> x dvd a \<Longrightarrow>
     x dvd b \<Longrightarrow> x = 1"
   apply (subgoal_tac "x dvd gcd a b")
@@ -1753,329 +1702,4 @@
   show ?thesis by(simp add: Gcd_def fold_set gcd_commute_int)
 qed
 
-
-subsection {* Primes *}
-
-(* FIXME Is there a better way to handle these, rather than making them elim rules? *)
-
-lemma prime_ge_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= 0"
-  by (unfold prime_nat_def, auto)
-
-lemma prime_gt_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p > 0"
-  by (unfold prime_nat_def, auto)
-
-lemma prime_ge_1_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= 1"
-  by (unfold prime_nat_def, auto)
-
-lemma prime_gt_1_nat [elim]: "prime (p::nat) \<Longrightarrow> p > 1"
-  by (unfold prime_nat_def, auto)
-
-lemma prime_ge_Suc_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= Suc 0"
-  by (unfold prime_nat_def, auto)
-
-lemma prime_gt_Suc_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p > Suc 0"
-  by (unfold prime_nat_def, auto)
-
-lemma prime_ge_2_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= 2"
-  by (unfold prime_nat_def, auto)
-
-lemma prime_ge_0_int [elim]: "prime (p::int) \<Longrightarrow> p >= 0"
-  by (unfold prime_int_def prime_nat_def) auto
-
-lemma prime_gt_0_int [elim]: "prime (p::int) \<Longrightarrow> p > 0"
-  by (unfold prime_int_def prime_nat_def, auto)
-
-lemma prime_ge_1_int [elim]: "prime (p::int) \<Longrightarrow> p >= 1"
-  by (unfold prime_int_def prime_nat_def, auto)
-
-lemma prime_gt_1_int [elim]: "prime (p::int) \<Longrightarrow> p > 1"
-  by (unfold prime_int_def prime_nat_def, auto)
-
-lemma prime_ge_2_int [elim]: "prime (p::int) \<Longrightarrow> p >= 2"
-  by (unfold prime_int_def prime_nat_def, auto)
-
-
-lemma prime_int_altdef: "prime (p::int) = (1 < p \<and> (\<forall>m \<ge> 0. m dvd p \<longrightarrow>
-    m = 1 \<or> m = p))"
-  using prime_nat_def [transferred]
-    apply (case_tac "p >= 0")
-    by (blast, auto simp add: prime_ge_0_int)
-
-lemma prime_imp_coprime_nat: "prime (p::nat) \<Longrightarrow> \<not> p dvd n \<Longrightarrow> coprime p n"
-  apply (unfold prime_nat_def)
-  apply (metis gcd_dvd1_nat gcd_dvd2_nat)
-  done
-
-lemma prime_imp_coprime_int: "prime (p::int) \<Longrightarrow> \<not> p dvd n \<Longrightarrow> coprime p n"
-  apply (unfold prime_int_altdef)
-  apply (metis gcd_dvd1_int gcd_dvd2_int gcd_ge_0_int)
-  done
-
-lemma prime_dvd_mult_nat: "prime (p::nat) \<Longrightarrow> p dvd m * n \<Longrightarrow> p dvd m \<or> p dvd n"
-  by (blast intro: coprime_dvd_mult_nat prime_imp_coprime_nat)
-
-lemma prime_dvd_mult_int: "prime (p::int) \<Longrightarrow> p dvd m * n \<Longrightarrow> p dvd m \<or> p dvd n"
-  by (blast intro: coprime_dvd_mult_int prime_imp_coprime_int)
-
-lemma prime_dvd_mult_eq_nat [simp]: "prime (p::nat) \<Longrightarrow>
-    p dvd m * n = (p dvd m \<or> p dvd n)"
-  by (rule iffI, rule prime_dvd_mult_nat, auto)
-
-lemma prime_dvd_mult_eq_int [simp]: "prime (p::int) \<Longrightarrow>
-    p dvd m * n = (p dvd m \<or> p dvd n)"
-  by (rule iffI, rule prime_dvd_mult_int, auto)
-
-lemma not_prime_eq_prod_nat: "(n::nat) > 1 \<Longrightarrow> ~ prime n \<Longrightarrow>
-    EX m k. n = m * k & 1 < m & m < n & 1 < k & k < n"
-  unfolding prime_nat_def dvd_def apply auto
-  by(metis mult_commute linorder_neq_iff linorder_not_le mult_1 n_less_n_mult_m one_le_mult_iff less_imp_le_nat)
-
-lemma not_prime_eq_prod_int: "(n::int) > 1 \<Longrightarrow> ~ prime n \<Longrightarrow>
-    EX m k. n = m * k & 1 < m & m < n & 1 < k & k < n"
-  unfolding prime_int_altdef dvd_def
-  apply auto
-  by(metis div_mult_self1_is_id div_mult_self2_is_id int_div_less_self int_one_le_iff_zero_less zero_less_mult_pos zless_le)
-
-lemma prime_dvd_power_nat [rule_format]: "prime (p::nat) -->
-    n > 0 --> (p dvd x^n --> p dvd x)"
-  by (induct n rule: nat_induct, auto)
-
-lemma prime_dvd_power_int [rule_format]: "prime (p::int) -->
-    n > 0 --> (p dvd x^n --> p dvd x)"
-  apply (induct n rule: nat_induct, auto)
-  apply (frule prime_ge_0_int)
-  apply auto
-done
-
-subsubsection{* Make prime naively executable *}
-
-lemma zero_not_prime_nat [simp]: "~prime (0::nat)"
-  by (simp add: prime_nat_def)
-
-lemma zero_not_prime_int [simp]: "~prime (0::int)"
-  by (simp add: prime_int_def)
-
-lemma one_not_prime_nat [simp]: "~prime (1::nat)"
-  by (simp add: prime_nat_def)
-
-lemma Suc_0_not_prime_nat [simp]: "~prime (Suc 0)"
-  by (simp add: prime_nat_def One_nat_def)
-
-lemma one_not_prime_int [simp]: "~prime (1::int)"
-  by (simp add: prime_int_def)
-
-lemma prime_nat_code[code]:
- "prime(p::nat) = (p > 1 & (ALL n : {1<..<p}. ~(n dvd p)))"
-apply(simp add: Ball_def)
-apply (metis less_not_refl prime_nat_def dvd_triv_right not_prime_eq_prod_nat)
-done
-
-lemma prime_nat_simp:
- "prime(p::nat) = (p > 1 & (list_all (%n. ~ n dvd p) [2..<p]))"
-apply(simp only:prime_nat_code list_ball_code greaterThanLessThan_upt)
-apply(simp add:nat_number One_nat_def)
-done
-
-lemmas prime_nat_simp_number_of[simp] = prime_nat_simp[of "number_of m", standard]
-
-lemma prime_int_code[code]:
-  "prime(p::int) = (p > 1 & (ALL n : {1<..<p}. ~(n dvd p)))" (is "?L = ?R")
-proof
-  assume "?L" thus "?R"
-    by (clarsimp simp: prime_gt_1_int) (metis int_one_le_iff_zero_less prime_int_altdef zless_le)
-next
-    assume "?R" thus "?L" by (clarsimp simp:Ball_def) (metis dvdI not_prime_eq_prod_int)
-qed
-
-lemma prime_int_simp:
-  "prime(p::int) = (p > 1 & (list_all (%n. ~ n dvd p) [2..p - 1]))"
-apply(simp only:prime_int_code list_ball_code greaterThanLessThan_upto)
-apply simp
-done
-
-lemmas prime_int_simp_number_of[simp] = prime_int_simp[of "number_of m", standard]
-
-declare successor_int_def[simp]
-
-lemma two_is_prime_nat [simp]: "prime (2::nat)"
-by simp
-
-lemma two_is_prime_int [simp]: "prime (2::int)"
-by simp
-
-text{* A bit of regression testing: *}
-
-lemma "prime(97::nat)"
-by simp
-
-lemma "prime(97::int)"
-by simp
-
-lemma "prime(997::nat)"
-by eval
-
-lemma "prime(997::int)"
-by eval
-
-
-lemma prime_imp_power_coprime_nat: "prime (p::nat) \<Longrightarrow> ~ p dvd a \<Longrightarrow> coprime a (p^m)"
-  apply (rule coprime_exp_nat)
-  apply (subst gcd_commute_nat)
-  apply (erule (1) prime_imp_coprime_nat)
-done
-
-lemma prime_imp_power_coprime_int: "prime (p::int) \<Longrightarrow> ~ p dvd a \<Longrightarrow> coprime a (p^m)"
-  apply (rule coprime_exp_int)
-  apply (subst gcd_commute_int)
-  apply (erule (1) prime_imp_coprime_int)
-done
-
-lemma primes_coprime_nat: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"
-  apply (rule prime_imp_coprime_nat, assumption)
-  apply (unfold prime_nat_def, auto)
-done
-
-lemma primes_coprime_int: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"
-  apply (rule prime_imp_coprime_int, assumption)
-  apply (unfold prime_int_altdef, clarify)
-  apply (drule_tac x = q in spec)
-  apply (drule_tac x = p in spec)
-  apply auto
-done
-
-lemma primes_imp_powers_coprime_nat: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow> p ~= q \<Longrightarrow> coprime (p^m) (q^n)"
-  by (rule coprime_exp2_nat, rule primes_coprime_nat)
-
-lemma primes_imp_powers_coprime_int: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow> p ~= q \<Longrightarrow> coprime (p^m) (q^n)"
-  by (rule coprime_exp2_int, rule primes_coprime_int)
-
-lemma prime_factor_nat: "n \<noteq> (1::nat) \<Longrightarrow> \<exists> p. prime p \<and> p dvd n"
-  apply (induct n rule: nat_less_induct)
-  apply (case_tac "n = 0")
-  using two_is_prime_nat apply blast
-  apply (case_tac "prime n")
-  apply blast
-  apply (subgoal_tac "n > 1")
-  apply (frule (1) not_prime_eq_prod_nat)
-  apply (auto intro: dvd_mult dvd_mult2)
-done
-
-(* An Isar version:
-
-lemma prime_factor_b_nat:
-  fixes n :: nat
-  assumes "n \<noteq> 1"
-  shows "\<exists>p. prime p \<and> p dvd n"
-
-using `n ~= 1`
-proof (induct n rule: less_induct_nat)
-  fix n :: nat
-  assume "n ~= 1" and
-    ih: "\<forall>m<n. m \<noteq> 1 \<longrightarrow> (\<exists>p. prime p \<and> p dvd m)"
-  thus "\<exists>p. prime p \<and> p dvd n"
-  proof -
-  {
-    assume "n = 0"
-    moreover note two_is_prime_nat
-    ultimately have ?thesis
-      by (auto simp del: two_is_prime_nat)
-  }
-  moreover
-  {
-    assume "prime n"
-    hence ?thesis by auto
-  }
-  moreover
-  {
-    assume "n ~= 0" and "~ prime n"
-    with `n ~= 1` have "n > 1" by auto
-    with `~ prime n` and not_prime_eq_prod_nat obtain m k where
-      "n = m * k" and "1 < m" and "m < n" by blast
-    with ih obtain p where "prime p" and "p dvd m" by blast
-    with `n = m * k` have ?thesis by auto
-  }
-  ultimately show ?thesis by blast
-  qed
-qed
-
-*)
-
-text {* One property of coprimality is easier to prove via prime factors. *}
-
-lemma prime_divprod_pow_nat:
-  assumes p: "prime (p::nat)" and ab: "coprime a b" and pab: "p^n dvd a * b"
-  shows "p^n dvd a \<or> p^n dvd b"
-proof-
-  {assume "n = 0 \<or> a = 1 \<or> b = 1" with pab have ?thesis
-      apply (cases "n=0", simp_all)
-      apply (cases "a=1", simp_all) done}
-  moreover
-  {assume n: "n \<noteq> 0" and a: "a\<noteq>1" and b: "b\<noteq>1"
-    then obtain m where m: "n = Suc m" by (cases n, auto)
-    from n have "p dvd p^n" by (intro dvd_power, auto)
-    also note pab
-    finally have pab': "p dvd a * b".
-    from prime_dvd_mult_nat[OF p pab']
-    have "p dvd a \<or> p dvd b" .
-    moreover
-    {assume pa: "p dvd a"
-      have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)
-      from coprime_common_divisor_nat [OF ab, OF pa] p have "\<not> p dvd b" by auto
-      with p have "coprime b p"
-        by (subst gcd_commute_nat, intro prime_imp_coprime_nat)
-      hence pnb: "coprime (p^n) b"
-        by (subst gcd_commute_nat, rule coprime_exp_nat)
-      from coprime_divprod_nat[OF pnba pnb] have ?thesis by blast }
-    moreover
-    {assume pb: "p dvd b"
-      have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)
-      from coprime_common_divisor_nat [OF ab, of p] pb p have "\<not> p dvd a"
-        by auto
-      with p have "coprime a p"
-        by (subst gcd_commute_nat, intro prime_imp_coprime_nat)
-      hence pna: "coprime (p^n) a"
-        by (subst gcd_commute_nat, rule coprime_exp_nat)
-      from coprime_divprod_nat[OF pab pna] have ?thesis by blast }
-    ultimately have ?thesis by blast}
-  ultimately show ?thesis by blast
-qed
-
-subsection {* Infinitely many primes *}
-
-lemma next_prime_bound: "\<exists>(p::nat). prime p \<and> n < p \<and> p <= fact n + 1"
-proof-
-  have f1: "fact n + 1 \<noteq> 1" using fact_ge_one_nat [of n] by arith 
-  from prime_factor_nat [OF f1]
-      obtain p where "prime p" and "p dvd fact n + 1" by auto
-  hence "p \<le> fact n + 1" 
-    by (intro dvd_imp_le, auto)
-  {assume "p \<le> n"
-    from `prime p` have "p \<ge> 1" 
-      by (cases p, simp_all)
-    with `p <= n` have "p dvd fact n" 
-      by (intro dvd_fact_nat)
-    with `p dvd fact n + 1` have "p dvd fact n + 1 - fact n"
-      by (rule dvd_diff_nat)
-    hence "p dvd 1" by simp
-    hence "p <= 1" by auto
-    moreover from `prime p` have "p > 1" by auto
-    ultimately have False by auto}
-  hence "n < p" by arith
-  with `prime p` and `p <= fact n + 1` show ?thesis by auto
-qed
-
-lemma bigger_prime: "\<exists>p. prime p \<and> p > (n::nat)" 
-using next_prime_bound by auto
-
-lemma primes_infinite: "\<not> (finite {(p::nat). prime p})"
-proof
-  assume "finite {(p::nat). prime p}"
-  with Max_ge have "(EX b. (ALL x : {(p::nat). prime p}. x <= b))"
-    by auto
-  then obtain b where "ALL (x::nat). prime x \<longrightarrow> x <= b"
-    by auto
-  with bigger_prime [of b] show False by auto
-qed
-
-
 end
--- a/src/HOL/HOL.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/HOL.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -29,6 +29,7 @@
   "~~/src/Tools/induct.ML"
   ("~~/src/Tools/induct_tacs.ML")
   ("Tools/recfun_codegen.ML")
+  "~~/src/Tools/more_conv.ML"
 begin
 
 setup {* Intuitionistic.method_setup @{binding iprover} *}
@@ -1886,7 +1887,7 @@
 *}
 
 setup {*
-  Code.add_const_alias @{thm equals_alias_cert}
+  Nbe.add_const_alias @{thm equals_alias_cert}
 *}
 
 hide (open) const eq
--- a/src/HOL/HoareParallel/Graph.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/HoareParallel/Graph.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -203,11 +203,11 @@
   apply(rule_tac x = "(take m path)@patha" in exI)
   apply(subgoal_tac "\<not>(length path\<le>m)")
    prefer 2 apply arith
-  apply(simp add: min_def)
+  apply(simp)
   apply(rule conjI)
    apply(subgoal_tac "\<not>(m + length patha - 1 < m)")
     prefer 2 apply arith
-   apply(simp add: nth_append min_def)
+   apply(simp add: nth_append)
   apply(rule conjI)
    apply(case_tac "m")
     apply force
@@ -236,7 +236,7 @@
     apply(erule_tac x = "m - 1" in allE)
     apply(simp add: nth_list_update)
    apply(force simp add: nth_list_update)
-  apply(simp add: nth_append min_def)
+  apply(simp add: nth_append)
   apply(rotate_tac -4)
   apply(erule_tac x = "i - m" in allE)
   apply(subgoal_tac "Suc (i - m)=(Suc i - m)" )
@@ -248,8 +248,7 @@
  apply(rule_tac x = "take (Suc m) path" in exI)
  apply(subgoal_tac "\<not>(length path\<le>Suc m)" )
   prefer 2 apply arith
- apply(simp add: min_def)
- apply clarify
+ apply clarsimp
  apply(erule_tac x = "i" in allE)
  apply simp
  apply clarify
--- a/src/HOL/Import/HOL/ROOT.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Import/HOL/ROOT.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -1,8 +1,4 @@
-(*  Title:      HOL/Import/HOL/ROOT.ML
-    ID:         $Id$
-    Author:     Sebastian Skalberg (TU Muenchen)
-*)
 
-use_thy "Primes";
+use_thy "~~/src/HOL/Old_Number_Theory/Primes";
 setmp_noncritical quick_and_dirty true use_thy "HOL4Prob";
 setmp_noncritical quick_and_dirty true use_thy "HOL4";
--- a/src/HOL/Import/HOL4Compat.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Import/HOL4Compat.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -3,7 +3,7 @@
 *)
 
 theory HOL4Compat
-imports HOL4Setup Complex_Main Primes ContNotDenum
+imports HOL4Setup Complex_Main "~~/src/HOL/Old_Number_Theory/Primes" ContNotDenum
 begin
 
 no_notation differentiable (infixl "differentiable" 60)
--- a/src/HOL/Import/hol4rews.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Import/hol4rews.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -531,7 +531,7 @@
 	val _ = app (fn (hol,(internal,isa,opt_ty)) =>
 			(out ("\n  " ^ (trans_string hol) ^ " > " ^ (trans_string (follow_cname isa thy)));
 			 case opt_ty of
-			     SOME ty => out (" :: \"" ^ (Display.string_of_ctyp (ctyp_of thy ty)) ^ "\"")
+			     SOME ty => out (" :: \"" ^ Syntax.string_of_typ_global thy ty ^ "\"")
 			   | NONE => ())) constmaps
 	val _ = if null constmaps
 		then ()
--- a/src/HOL/Import/proof_kernel.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Import/proof_kernel.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -199,12 +199,12 @@
         val ct  = (cterm_of thy (HOLogic.dest_Trueprop t)
                            handle TERM _ => ct)
     in
-        quote(
+        quote (
         PrintMode.setmp [] (
         Library.setmp show_brackets false (
         Library.setmp show_all_types true (
         Library.setmp Syntax.ambiguity_is_error false (
-        Library.setmp show_sorts true Display.string_of_cterm))))
+        Library.setmp show_sorts true (Syntax.string_of_term_global thy o Thm.term_of)))))
         ct)
     end
 
@@ -226,7 +226,8 @@
           | G _ = raise SMART_STRING
         fun F n =
             let
-                val str = Library.setmp show_brackets false (G n Display.string_of_cterm) ct
+                val str =
+                  Library.setmp show_brackets false (G n (Syntax.string_of_term ctxt o term_of)) ct
                 val u = Syntax.parse_term ctxt str
                   |> TypeInfer.constrain T |> Syntax.check_term ctxt
             in
@@ -234,8 +235,9 @@
                 then quote str
                 else F (n+1)
             end
-            handle ERROR mesg => F (n+1)
-                 | SMART_STRING => error ("smart_string failed for: "^(G 0 Display.string_of_cterm ct))
+            handle ERROR mesg => F (n + 1)
+              | SMART_STRING =>
+                  error ("smart_string failed for: "^ G 0 (Syntax.string_of_term ctxt o term_of) ct)
     in
       PrintMode.setmp [] (Library.setmp Syntax.ambiguity_is_error true F) 0
     end
@@ -243,8 +245,7 @@
 
 val smart_string_of_thm = smart_string_of_cterm o cprop_of
 
-fun prth th = writeln (PrintMode.setmp [] Display.string_of_thm_without_context th)
-fun prc ct = writeln (PrintMode.setmp [] Display.string_of_cterm ct)
+fun prth th = writeln (PrintMode.setmp [] Display.string_of_thm_without_context th);
 fun prin t = writeln (PrintMode.setmp []
   (fn () => Syntax.string_of_term (ML_Context.the_local_context ()) t) ());
 fun pth (HOLThm(ren,thm)) =
@@ -1939,16 +1940,17 @@
                     then
                         let
                             val p1 = quotename constname
-                            val p2 = Display.string_of_ctyp (ctyp_of thy'' ctype)
+                            val p2 = Syntax.string_of_typ_global thy'' ctype
                             val p3 = string_of_mixfix csyn
                             val p4 = smart_string_of_cterm crhs
                         in
-                            add_dump ("constdefs\n  " ^p1^ " :: \"" ^p2^ "\" "^p3^ "\n  " ^p4) thy''
+                          add_dump ("constdefs\n  " ^ p1 ^ " :: \"" ^ p2 ^ "\" "^ p3 ^ "\n  " ^ p4) thy''
                         end
                     else
-                        (add_dump ("consts\n  " ^ (quotename constname) ^ " :: \"" ^ Display.string_of_ctyp (ctyp_of thy'' ctype) ^
-                                   "\" " ^ (string_of_mixfix csyn) ^ "\n\ndefs\n  " ^ (quotename thmname) ^ ": " ^ (smart_string_of_cterm crhs))
-                                  thy'')
+                        add_dump ("consts\n  " ^ quotename constname ^ " :: \"" ^
+                          Syntax.string_of_typ_global thy'' ctype ^
+                          "\" " ^ string_of_mixfix csyn ^ "\n\ndefs\n  " ^
+                          quotename thmname ^ ": " ^ smart_string_of_cterm crhs) thy''
         val hth = case Shuffler.set_prop thy22 (HOLogic.mk_Trueprop tm24) [("",th)] of
                       SOME (_,res) => HOLThm(rens_of linfo,res)
                     | NONE => raise ERR "new_definition" "Bad conclusion"
@@ -2008,8 +2010,9 @@
                                                           in
                                                               ((cname,cT,mk_syn thy cname)::cs,p)
                                                           end) (([],HOLogic.dest_Trueprop (concl_of th)),names)
-                               val str = Library.foldl (fn (acc,(c,T,csyn)) =>
-                                                   acc ^ "\n  " ^ (quotename c) ^ " :: \"" ^ Display.string_of_ctyp (ctyp_of thy T) ^ "\" " ^ (string_of_mixfix csyn)) ("consts",consts)
+                               val str = Library.foldl (fn (acc, (c, T, csyn)) =>
+                                   acc ^ "\n  " ^ quotename c ^ " :: \"" ^
+                                   Syntax.string_of_typ_global thy T ^ "\" " ^ string_of_mixfix csyn) ("consts", consts)
                                val thy' = add_dump str thy
                                val _ = ImportRecorder.add_consts consts
                            in
@@ -2137,7 +2140,7 @@
 fun add_dump_constdefs thy defname constname rhs ty =
     let
         val n = quotename constname
-        val t = Display.string_of_ctyp (ctyp_of thy ty)
+        val t = Syntax.string_of_typ_global thy ty
         val syn = string_of_mixfix (mk_syn thy constname)
         (*val eq = smart_string_of_cterm (cterm_of thy (Const(rhs, ty)))*)
         val eq = quote (constname ^ " == "^rhs)
@@ -2224,7 +2227,7 @@
               ("  apply (rule light_ex_imp_nonempty[where t="^
               (proc_prop (cterm_of thy4 t))^"])\n"^
               ("  by (import " ^ thyname ^ " " ^ (quotename thmname) ^ ")"))) thy4
-            val str_aty = Display.string_of_ctyp (ctyp_of thy aty)
+            val str_aty = Syntax.string_of_typ_global thy aty
             val thy = add_dump_syntax thy rep_name
             val thy = add_dump_syntax thy abs_name
             val thy = add_dump ("lemmas " ^ (quote (thmname^"_@intern")) ^
--- a/src/HOL/Import/shuffler.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Import/shuffler.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -57,7 +57,6 @@
 fun print_sign_exn sign e = (print_sign_exn_unit sign e; raise e)
 
 val string_of_thm = PrintMode.setmp [] Display.string_of_thm_without_context;
-val string_of_cterm = PrintMode.setmp [] Display.string_of_cterm;
 
 fun mk_meta_eq th =
     (case concl_of th of
@@ -304,13 +303,14 @@
                 val lhs = #1 (Logic.dest_equals (prop_of (final init)))
             in
                 if not (lhs aconv origt)
-                then (writeln "Something is utterly wrong: (orig,lhs,frozen type,t,tinst)";
-                      writeln (Display.string_of_cterm (cterm_of thy origt));
-                      writeln (Display.string_of_cterm (cterm_of thy lhs));
-                      writeln (Display.string_of_cterm (cterm_of thy typet));
-                      writeln (Display.string_of_cterm (cterm_of thy t));
-                      app (fn (n,T) => writeln (n ^ ": " ^ (Display.string_of_ctyp (ctyp_of thy T)))) Tinst;
-                      writeln "done")
+                then
+                  writeln (cat_lines
+                    (["Something is utterly wrong: (orig, lhs, frozen type, t, tinst)",
+                      Syntax.string_of_term_global thy origt,
+                      Syntax.string_of_term_global thy lhs,
+                      Syntax.string_of_term_global thy typet,
+                      Syntax.string_of_term_global thy t] @
+                      map (fn (n, T) => n ^ ": " ^ Syntax.string_of_typ_global thy T) Tinst))
                 else ()
             end
     in
@@ -366,13 +366,14 @@
                 val lhs = #1 (Logic.dest_equals (prop_of (final init)))
             in
                 if not (lhs aconv origt)
-                then (writeln "Something is utterly wrong: (orig,lhs,frozen type,t,tinst)";
-                      writeln (Display.string_of_cterm (cterm_of thy origt));
-                      writeln (Display.string_of_cterm (cterm_of thy lhs));
-                      writeln (Display.string_of_cterm (cterm_of thy typet));
-                      writeln (Display.string_of_cterm (cterm_of thy t));
-                      app (fn (n,T) => writeln (n ^ ": " ^ (Display.string_of_ctyp (ctyp_of thy T)))) Tinst;
-                      writeln "done")
+                then
+                  writeln (cat_lines
+                    (["Something is utterly wrong: (orig, lhs, frozen type, t, tinst)",
+                      Syntax.string_of_term_global thy origt,
+                      Syntax.string_of_term_global thy lhs,
+                      Syntax.string_of_term_global thy typet,
+                      Syntax.string_of_term_global thy t] @
+                      map (fn (n, T) => n ^ ": " ^ Syntax.string_of_typ_global thy T) Tinst))
                 else ()
             end
     in
@@ -407,9 +408,8 @@
                       end
                     | _ => NONE)
             else NONE
-          | _ => (error ("Bad eta_expand argument" ^ (string_of_cterm (cterm_of thy t))); NONE)
-    end
-    handle e => (writeln "eta_expand internal error"; OldGoals.print_exn e)
+          | _ => error ("Bad eta_expand argument" ^ Syntax.string_of_term_global thy t)
+    end;
 
 fun mk_tfree s = TFree("'"^s,[])
 fun mk_free s t = Free (s,t)
--- a/src/HOL/Int.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Int.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -266,7 +266,7 @@
 proof  
   fix k :: int
   show "abs k = sup k (- k)"
-    by (auto simp add: sup_int_def zabs_def max_def less_minus_self_iff [symmetric])
+    by (auto simp add: sup_int_def zabs_def less_minus_self_iff [symmetric])
 qed
 
 lemma zless_imp_add1_zle: "w < z \<Longrightarrow> w + (1\<Colon>int) \<le> z"
@@ -1487,21 +1487,6 @@
        add_special diff_special eq_special less_special le_special
 
 
-lemma min_max_01: "min (0::int) 1 = 0 & min (1::int) 0 = 0 &
-                   max (0::int) 1 = 1 & max (1::int) 0 = 1"
-by(simp add:min_def max_def)
-
-lemmas min_max_special[simp] =
- min_max_01
- max_def[of "0::int" "number_of v", standard, simp]
- min_def[of "0::int" "number_of v", standard, simp]
- max_def[of "number_of u" "0::int", standard, simp]
- min_def[of "number_of u" "0::int", standard, simp]
- max_def[of "1::int" "number_of v", standard, simp]
- min_def[of "1::int" "number_of v", standard, simp]
- max_def[of "number_of u" "1::int", standard, simp]
- min_def[of "number_of u" "1::int", standard, simp]
-
 text {* Legacy theorems *}
 
 lemmas zle_int = of_nat_le_iff [where 'a=int]
--- a/src/HOL/IsaMakefile	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/IsaMakefile	Fri Sep 11 15:57:16 2009 +1000
@@ -31,11 +31,12 @@
   HOL-Matrix \
   HOL-MetisExamples \
   HOL-MicroJava \
+  HOL-Mirabelle \
   HOL-Modelcheck \
   HOL-NanoJava \
   HOL-Nominal-Examples \
-  HOL-NewNumberTheory \
-  HOL-NumberTheory \
+  HOL-Number_Theory \
+  HOL-Old_Number_Theory \
   HOL-Prolog \
   HOL-SET-Protocol \
   HOL-SizeChange \
@@ -108,6 +109,7 @@
   $(SRC)/Tools/random_word.ML \
   $(SRC)/Tools/value.ML \
   $(SRC)/Tools/Code_Generator.thy \
+  $(SRC)/Tools/more_conv.ML \
   HOL.thy \
   Tools/hologic.ML \
   Tools/recfun_codegen.ML \
@@ -281,11 +283,13 @@
   Complex.thy \
   Deriv.thy \
   Fact.thy \
+  GCD.thy \
   Integration.thy \
   Lim.thy \
   Limits.thy \
   Ln.thy \
   Log.thy \
+  Lubs.thy \
   MacLaurin.thy \
   NatTransfer.thy \
   NthRoot.thy \
@@ -293,9 +297,7 @@
   Series.thy \
   Taylor.thy \
   Transcendental.thy \
-  GCD.thy \
   Parity.thy \
-  Lubs.thy \
   PReal.thy \
   Rational.thy \
   RComplete.thy \
@@ -329,10 +331,10 @@
   Library/Finite_Cartesian_Product.thy Library/FrechetDeriv.thy		\
   Library/Fraction_Field.thy Library/Fundamental_Theorem_Algebra.thy	\
   Library/Inner_Product.thy Library/Kleene_Algebra.thy			\
-  Library/Lattice_Syntax.thy Library/Legacy_GCD.thy			\
+  Library/Lattice_Syntax.thy			\
   Library/Library.thy Library/List_Prefix.thy Library/List_Set.thy	\
   Library/State_Monad.thy Library/Nat_Int_Bij.thy Library/Multiset.thy	\
-  Library/Permutation.thy Library/Primes.thy Library/Pocklington.thy	\
+  Library/Permutation.thy	\
   Library/Quotient.thy Library/Quicksort.thy Library/Nat_Infinity.thy	\
   Library/Word.thy Library/README.html Library/Continuity.thy		\
   Library/Order_Relation.thy Library/Nested_Environment.thy		\
@@ -486,38 +488,39 @@
 	@cd Import/HOLLight; $(ISABELLE_TOOL) usedir -b $(OUT)/HOL HOLLight
 
 
-## HOL-NewNumberTheory
+## HOL-Number_Theory
 
-HOL-NewNumberTheory: HOL $(LOG)/HOL-NewNumberTheory.gz
+HOL-Number_Theory: HOL $(LOG)/HOL-Number_Theory.gz
 
-$(LOG)/HOL-NewNumberTheory.gz: $(OUT)/HOL $(ALGEBRA_DEPENDENCIES) \
+$(LOG)/HOL-Number_Theory.gz: $(OUT)/HOL $(ALGEBRA_DEPENDENCIES) \
   Library/Multiset.thy \
-  NewNumberTheory/Binomial.thy \
-  NewNumberTheory/Cong.thy \
-  NewNumberTheory/Fib.thy \
-  NewNumberTheory/MiscAlgebra.thy \
-  NewNumberTheory/Residues.thy \
-  NewNumberTheory/UniqueFactorization.thy  \
-  NewNumberTheory/ROOT.ML
-	@$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL NewNumberTheory
+  Number_Theory/Binomial.thy \
+  Number_Theory/Cong.thy \
+  Number_Theory/Fib.thy \
+  Number_Theory/MiscAlgebra.thy \
+  Number_Theory/Number_Theory.thy \
+  Number_Theory/Residues.thy \
+  Number_Theory/UniqueFactorization.thy  \
+  Number_Theory/ROOT.ML
+	@$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL Number_Theory
 
 
-## HOL-NumberTheory
+## HOL-Old_Number_Theory
 
-HOL-NumberTheory: HOL $(LOG)/HOL-NumberTheory.gz
+HOL-Old_Number_Theory: HOL $(LOG)/HOL-Old_Number_Theory.gz
 
-$(LOG)/HOL-NumberTheory.gz: $(OUT)/HOL Library/Permutation.thy		\
-  Library/Primes.thy NumberTheory/Fib.thy				\
-  NumberTheory/Factorization.thy NumberTheory/BijectionRel.thy		\
-  NumberTheory/Chinese.thy NumberTheory/EulerFermat.thy			\
-  NumberTheory/IntFact.thy NumberTheory/IntPrimes.thy			\
-  NumberTheory/WilsonBij.thy NumberTheory/WilsonRuss.thy		\
-  NumberTheory/Finite2.thy NumberTheory/Int2.thy			\
-  NumberTheory/EvenOdd.thy NumberTheory/Residues.thy			\
-  NumberTheory/Euler.thy NumberTheory/Gauss.thy				\
-  NumberTheory/Quadratic_Reciprocity.thy Library/Infinite_Set.thy	\
-  NumberTheory/ROOT.ML
-	@$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL NumberTheory
+$(LOG)/HOL-Old_Number_Theory.gz: $(OUT)/HOL Library/Permutation.thy		\
+  Old_Number_Theory/Primes.thy Old_Number_Theory/Fib.thy				\
+  Old_Number_Theory/Factorization.thy Old_Number_Theory/BijectionRel.thy		\
+  Old_Number_Theory/Chinese.thy Old_Number_Theory/EulerFermat.thy			\
+  Old_Number_Theory/IntFact.thy Old_Number_Theory/IntPrimes.thy			\
+  Old_Number_Theory/WilsonBij.thy Old_Number_Theory/WilsonRuss.thy		\
+  Old_Number_Theory/Finite2.thy Old_Number_Theory/Int2.thy			\
+  Old_Number_Theory/EvenOdd.thy Old_Number_Theory/Residues.thy			\
+  Old_Number_Theory/Euler.thy Old_Number_Theory/Gauss.thy				\
+  Old_Number_Theory/Quadratic_Reciprocity.thy Library/Infinite_Set.thy	\
+  Old_Number_Theory/Legacy_GCD.thy Old_Number_Theory/Pocklington.thy Old_Number_Theory/ROOT.ML
+	@$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL Old_Number_Theory
 
 
 ## HOL-Hoare
@@ -572,7 +575,7 @@
   Library/FuncSet.thy \
   Library/Multiset.thy \
   Library/Permutation.thy \
-  Library/Primes.thy \
+  Number_Theory/Primes.thy \
   Algebra/AbelCoset.thy \
   Algebra/Bij.thy \
   Algebra/Congruence.thy \
@@ -875,7 +878,7 @@
 HOL-ex: HOL $(LOG)/HOL-ex.gz
 
 $(LOG)/HOL-ex.gz: $(OUT)/HOL Library/Commutative_Ring.thy		\
-  Library/Primes.thy ex/Abstract_NAT.thy ex/Antiquote.thy		\
+  Number_Theory/Primes.thy ex/Abstract_NAT.thy ex/Antiquote.thy		\
   ex/Arith_Examples.thy				\
   ex/Arithmetic_Series_Complex.thy ex/BT.thy ex/BinEx.thy		\
   ex/Binary.thy ex/CTL.thy ex/Chinese.thy ex/Classical.thy		\
@@ -900,8 +903,7 @@
   ex/Sudoku.thy ex/Tarski.thy \
   ex/Termination.thy ex/Unification.thy ex/document/root.bib		\
   ex/document/root.tex ex/set.thy ex/svc_funcs.ML ex/svc_test.thy \
-  ex/Predicate_Compile.thy ex/predicate_compile.ML ex/Predicate_Compile_ex.thy \
-  ex/Mirabelle.thy ex/mirabelle.ML
+  ex/Predicate_Compile.thy ex/predicate_compile.ML ex/Predicate_Compile_ex.thy
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL ex
 
 
@@ -939,7 +941,7 @@
 
 HOL-Matrix: HOL $(LOG)/HOL-Matrix.gz
 
-$(LOG)/HOL-Matrix.gz: $(OUT)/HOL				\
+$(LOG)/HOL-Matrix.gz: $(OUT)/HOL					\
   $(SRC)/Tools/Compute_Oracle/Compute_Oracle.thy			\
   $(SRC)/Tools/Compute_Oracle/am_compiler.ML				\
   $(SRC)/Tools/Compute_Oracle/am_interpreter.ML				\
@@ -947,8 +949,8 @@
   $(SRC)/Tools/Compute_Oracle/linker.ML					\
   $(SRC)/Tools/Compute_Oracle/am_ghc.ML					\
   $(SRC)/Tools/Compute_Oracle/am_sml.ML					\
-  $(SRC)/Tools/Compute_Oracle/compute.ML	\
-  Tools/ComputeFloat.thy Tools/float_arith.ML \
+  $(SRC)/Tools/Compute_Oracle/compute.ML Matrix/ComputeFloat.thy	\
+  Matrix/ComputeHOL.thy Matrix/ComputeNumeral.thy Tools/float_arith.ML	\
   Matrix/Matrix.thy Matrix/SparseMatrix.thy Matrix/LP.thy		\
   Matrix/document/root.tex Matrix/ROOT.ML Matrix/cplex/Cplex.thy	\
   Matrix/cplex/CplexMatrixConverter.ML Matrix/cplex/Cplex_tools.ML	\
@@ -1118,6 +1120,20 @@
 	@cd NSA; $(ISABELLE_TOOL) usedir $(OUT)/HOL-NSA Examples
 
 
+## HOL-Mirabelle
+
+HOL-Mirabelle: HOL $(LOG)/HOL-Mirabelle.gz
+
+$(LOG)/HOL-Mirabelle.gz: $(OUT)/HOL Mirabelle/Mirabelle_Test.thy \
+  Mirabelle/Mirabelle.thy Mirabelle/Tools/mirabelle.ML Mirabelle/ROOT.ML \
+  Mirabelle/Tools/mirabelle_arith.ML \
+  Mirabelle/Tools/mirabelle_metis.ML \
+  Mirabelle/Tools/mirabelle_quickcheck.ML \
+  Mirabelle/Tools/mirabelle_refute.ML \
+  Mirabelle/Tools/mirabelle_sledgehammer.ML
+	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Mirabelle
+
+
 ## clean
 
 clean:
@@ -1139,4 +1155,5 @@
 		$(LOG)/TLA-Memory.gz $(LOG)/HOL-Library.gz		\
 		$(LOG)/HOL-Unix.gz $(OUT)/HOL-Word $(LOG)/HOL-Word.gz	\
 		$(LOG)/HOL-Word-Examples.gz $(OUT)/HOL-NSA		\
-		$(LOG)/HOL-NSA.gz $(LOG)/HOL-NSA-Examples.gz
+		$(LOG)/HOL-NSA.gz $(LOG)/HOL-NSA-Examples.gz            \
+                $(LOG)/HOL-Mirabelle.gz
--- a/src/HOL/Isar_examples/Mutilated_Checkerboard.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Isar_examples/Mutilated_Checkerboard.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -113,7 +113,7 @@
 lemma evnodd_insert: "evnodd (insert (i, j) C) b =
     (if (i + j) mod 2 = b
       then insert (i, j) (evnodd C b) else evnodd C b)"
-  by (simp add: evnodd_def) blast
+  by (simp add: evnodd_def)
 
 
 subsection {* Dominoes *}
--- a/src/HOL/Isar_examples/ROOT.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Isar_examples/ROOT.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -4,7 +4,7 @@
 Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 *)
 
-no_document use_thys ["../NumberTheory/Primes", "../NumberTheory/Fibonacci"];
+no_document use_thys ["../Old_Number_Theory/Primes", "../Old_Number_Theory/Fibonacci"];
 
 use_thys [
   "Basic_Logic",
--- a/src/HOL/Lattices.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Lattices.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -125,10 +125,10 @@
 lemma inf_left_idem[simp]: "x \<sqinter> (x \<sqinter> y) = x \<sqinter> y"
   by (rule antisym) (auto intro: le_infI2)
 
-lemma inf_absorb1: "x \<sqsubseteq> y \<Longrightarrow> x \<sqinter> y = x"
+lemma inf_absorb1[simp]: "x \<sqsubseteq> y \<Longrightarrow> x \<sqinter> y = x"
   by (rule antisym) auto
 
-lemma inf_absorb2: "y \<sqsubseteq> x \<Longrightarrow> x \<sqinter> y = y"
+lemma inf_absorb2[simp]: "y \<sqsubseteq> x \<Longrightarrow> x \<sqinter> y = y"
   by (rule antisym) auto
 
 lemma inf_left_commute: "x \<sqinter> (y \<sqinter> z) = y \<sqinter> (x \<sqinter> z)"
@@ -153,10 +153,10 @@
 lemma sup_left_idem[simp]: "x \<squnion> (x \<squnion> y) = x \<squnion> y"
   by (rule antisym) (auto intro: le_supI2)
 
-lemma sup_absorb1: "y \<sqsubseteq> x \<Longrightarrow> x \<squnion> y = x"
+lemma sup_absorb1[simp]: "y \<sqsubseteq> x \<Longrightarrow> x \<squnion> y = x"
   by (rule antisym) auto
 
-lemma sup_absorb2: "x \<sqsubseteq> y \<Longrightarrow> x \<squnion> y = y"
+lemma sup_absorb2[simp]: "x \<sqsubseteq> y \<Longrightarrow> x \<squnion> y = y"
   by (rule antisym) auto
 
 lemma sup_left_commute: "x \<squnion> (y \<squnion> z) = y \<squnion> (x \<squnion> z)"
@@ -199,7 +199,7 @@
 shows "x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)"
 proof-
   have "x \<squnion> (y \<sqinter> z) = (x \<squnion> (x \<sqinter> z)) \<squnion> (y \<sqinter> z)" by(simp add:sup_inf_absorb)
-  also have "\<dots> = x \<squnion> (z \<sqinter> (x \<squnion> y))" by(simp add:D inf_commute sup_assoc)
+  also have "\<dots> = x \<squnion> (z \<sqinter> (x \<squnion> y))" by(simp add:D inf_commute sup_assoc del:sup_absorb1)
   also have "\<dots> = ((x \<squnion> y) \<sqinter> x) \<squnion> ((x \<squnion> y) \<sqinter> z)"
     by(simp add:inf_sup_absorb inf_commute)
   also have "\<dots> = (x \<squnion> y) \<sqinter> (x \<squnion> z)" by(simp add:D)
@@ -211,7 +211,7 @@
 shows "x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)"
 proof-
   have "x \<sqinter> (y \<squnion> z) = (x \<sqinter> (x \<squnion> z)) \<sqinter> (y \<squnion> z)" by(simp add:inf_sup_absorb)
-  also have "\<dots> = x \<sqinter> (z \<squnion> (x \<sqinter> y))" by(simp add:D sup_commute inf_assoc)
+  also have "\<dots> = x \<sqinter> (z \<squnion> (x \<sqinter> y))" by(simp add:D sup_commute inf_assoc del:inf_absorb1)
   also have "\<dots> = ((x \<sqinter> y) \<squnion> x) \<sqinter> ((x \<sqinter> y) \<squnion> z)"
     by(simp add:sup_inf_absorb sup_commute)
   also have "\<dots> = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" by(simp add:D)
@@ -413,12 +413,11 @@
 subsection {* @{const min}/@{const max} on linear orders as
   special case of @{const inf}/@{const sup} *}
 
-sublocale linorder < min_max!: distrib_lattice less_eq less "Orderings.ord.min less_eq" "Orderings.ord.max less_eq"
+sublocale linorder < min_max!: distrib_lattice less_eq less min max
 proof
   fix x y z
-  show "Orderings.ord.max less_eq x (Orderings.ord.min less_eq y z) =
-    Orderings.ord.min less_eq (Orderings.ord.max less_eq x y) (Orderings.ord.max less_eq x z)"
-  unfolding min_def max_def by auto
+  show "max x (min y z) = min (max x y) (max x z)"
+    by (auto simp add: min_def max_def)
 qed (auto simp add: min_def max_def not_le less_imp_le)
 
 lemma inf_min: "inf = (min \<Colon> 'a\<Colon>{lower_semilattice, linorder} \<Rightarrow> 'a \<Rightarrow> 'a)"
@@ -514,9 +513,6 @@
   inf_compl_bot sup_compl_top diff_eq)
 
 
-text {* redundant bindings *}
-
-
 no_notation
   less_eq  (infix "\<sqsubseteq>" 50) and
   less (infix "\<sqsubset>" 50) and
--- a/src/HOL/Library/Abstract_Rat.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Library/Abstract_Rat.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -189,14 +189,9 @@
   have "\<exists> a b a' b'. x = (a,b) \<and> y = (a',b')" by auto
   then obtain a b a' b' where xy[simp]: "x = (a,b)" "y=(a',b')" by blast
   assume H: ?lhs 
-  {assume "a = 0 \<or> b = 0 \<or> a' = 0 \<or> b' = 0" hence ?rhs
-      using na nb H
-      apply (simp add: INum_def split_def isnormNum_def)
-      apply (cases "a = 0", simp_all)
-      apply (cases "b = 0", simp_all)
-      apply (cases "a' = 0", simp_all)
-      apply (cases "a' = 0", simp_all add: of_int_eq_0_iff)
-      done}
+  {assume "a = 0 \<or> b = 0 \<or> a' = 0 \<or> b' = 0"
+    hence ?rhs using na nb H
+      by (simp add: INum_def split_def isnormNum_def split: split_if_asm)}
   moreover
   { assume az: "a \<noteq> 0" and bz: "b \<noteq> 0" and a'z: "a'\<noteq>0" and b'z: "b'\<noteq>0"
     from az bz a'z b'z na nb have pos: "b > 0" "b' > 0" by (simp_all add: isnormNum_def)
@@ -517,10 +512,7 @@
     have n0: "isnormNum 0\<^sub>N" by simp
     show ?thesis using nx ny 
       apply (simp only: isnormNum_unique[where ?'a = 'a, OF  Nmul_normN[OF nx ny] n0, symmetric] Nmul[where ?'a = 'a])
-      apply (simp add: INum_def split_def isnormNum_def fst_conv snd_conv)
-      apply (cases "a=0",simp_all)
-      apply (cases "a'=0",simp_all)
-      done
+      by (simp add: INum_def split_def isnormNum_def fst_conv snd_conv split: split_if_asm)
   }
 qed
 lemma Nneg_Nneg[simp]: "~\<^sub>N (~\<^sub>N c) = c"
--- a/src/HOL/Library/Continuity.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Library/Continuity.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -156,7 +156,7 @@
  apply (rule up_chainI)
  apply simp
 apply (drule Un_absorb1)
-apply (auto simp add: nat_not_singleton)
+apply (auto split:split_if_asm)
 done
 
 
@@ -184,8 +184,7 @@
  apply (rule down_chainI)
  apply simp
 apply (drule Int_absorb1)
-apply auto
-apply (auto simp add: nat_not_singleton)
+apply (auto split:split_if_asm)
 done
 
 
--- a/src/HOL/Library/Convex_Euclidean_Space.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Library/Convex_Euclidean_Space.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -3355,9 +3355,8 @@
   qed(auto intro!:path_connected_singleton) next
   case False hence *:"{x::real^'n. norm(x - a) = r} = (\<lambda>x. a + r *\<^sub>R x) ` {x. norm x = 1}" unfolding not_le apply -apply(rule set_ext,rule)
     unfolding image_iff apply(rule_tac x="(1/r) *\<^sub>R (x - a)" in bexI) unfolding mem_Collect_eq norm_scaleR by (auto simp add: scaleR_right_diff_distrib)
-  have ***:"\<And>xa. (if xa = 0 then 0 else 1) \<noteq> 1 \<Longrightarrow> xa = 0" apply(rule ccontr) by auto
   have **:"{x::real^'n. norm x = 1} = (\<lambda>x. (1/norm x) *\<^sub>R x) ` (UNIV - {0})" apply(rule set_ext,rule)
-    unfolding image_iff apply(rule_tac x=x in bexI) unfolding mem_Collect_eq by(auto intro!: ***)
+    unfolding image_iff apply(rule_tac x=x in bexI) unfolding mem_Collect_eq by(auto split:split_if_asm)
   have "continuous_on (UNIV - {0}) (\<lambda>x::real^'n. 1 / norm x)" unfolding o_def continuous_on_eq_continuous_within
     apply(rule, rule continuous_at_within_inv[unfolded o_def inverse_eq_divide]) apply(rule continuous_at_within)
     apply(rule continuous_at_norm[unfolded o_def]) by auto
--- a/src/HOL/Library/Euclidean_Space.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Library/Euclidean_Space.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -1055,28 +1055,6 @@
 lemma norm_triangle_lt: "norm(x::real ^'n::finite) + norm(y) < e ==> norm(x + y) < e"
   by (metis basic_trans_rules(21) norm_triangle_ineq)
 
-lemma setsum_delta:
-  assumes fS: "finite S"
-  shows "setsum (\<lambda>k. if k=a then b k else 0) S = (if a \<in> S then b a else 0)"
-proof-
-  let ?f = "(\<lambda>k. if k=a then b k else 0)"
-  {assume a: "a \<notin> S"
-    hence "\<forall> k\<in> S. ?f k = 0" by simp
-    hence ?thesis  using a by simp}
-  moreover
-  {assume a: "a \<in> S"
-    let ?A = "S - {a}"
-    let ?B = "{a}"
-    have eq: "S = ?A \<union> ?B" using a by blast
-    have dj: "?A \<inter> ?B = {}" by simp
-    from fS have fAB: "finite ?A" "finite ?B" by auto
-    have "setsum ?f S = setsum ?f ?A + setsum ?f ?B"
-      using setsum_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
-      by simp
-    then have ?thesis  using a by simp}
-  ultimately show ?thesis by blast
-qed
-
 lemma component_le_norm: "\<bar>x$i\<bar> <= norm (x::real ^ 'n::finite)"
   apply (simp add: norm_vector_def)
   apply (rule member_le_setL2, simp_all)
@@ -2079,13 +2057,6 @@
 lemma matrix_add_ldistrib: "(A ** (B + C)) = (A \<star> B) + (A \<star> C)"
   by (vector matrix_matrix_mult_def setsum_addf[symmetric] ring_simps)
 
-lemma setsum_delta':
-  assumes fS: "finite S" shows
-  "setsum (\<lambda>k. if a = k then b k else 0) S =
-     (if a\<in> S then b a else 0)"
-  using setsum_delta[OF fS, of a b, symmetric]
-  by (auto intro: setsum_cong)
-
 lemma matrix_mul_lid:
   fixes A :: "'a::semiring_1 ^ 'm ^ 'n::finite"
   shows "mat 1 ** A = A"
@@ -3926,14 +3897,6 @@
   shows "finite s \<and> card s \<le> card t"
   by (metis exchange_lemma[OF f i sp] hassize_def finite_subset card_mono)
 
-lemma finite_Atleast_Atmost[simp]: "finite {f x |x. x\<in> {(i::'a::finite_intvl_succ) .. j}}"
-proof-
-  have eq: "{f x |x. x\<in> {i .. j}} = f ` {i .. j}" by auto
-  show ?thesis unfolding eq
-    apply (rule finite_imageI)
-    apply (rule finite_intvl)
-    done
-qed
 
 lemma finite_Atleast_Atmost_nat[simp]: "finite {f x |x. x\<in> (UNIV::'a::finite set)}"
 proof-
--- a/src/HOL/Library/Formal_Power_Series.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Library/Formal_Power_Series.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -633,8 +633,7 @@
 	    by (auto simp add: inverse_eq_divide power_divide)
 
 	  from k have kn: "k > n"
-	    apply (simp add: leastP_def setge_def fps_sum_rep_nth)
-	    by (cases "k \<le> n", auto)
+	    by (simp add: leastP_def setge_def fps_sum_rep_nth split:split_if_asm)
 	  then have "dist (?s n) a < (1/2)^n" unfolding dth
 	    by (auto intro: power_strict_decreasing)
 	  also have "\<dots> <= (1/2)^n0" using nn0
@@ -1244,10 +1243,9 @@
     {assume n0: "n \<noteq> 0"
       then have u: "{0} \<union> ({1} \<union> {2..n}) = {0..n}" "{1}\<union>{2..n} = {1..n}"
 	"{0..n - 1}\<union>{n} = {0..n}"
-	apply (simp_all add: expand_set_eq) by presburger+
+	by (auto simp: expand_set_eq)
       have d: "{0} \<inter> ({1} \<union> {2..n}) = {}" "{1} \<inter> {2..n} = {}"
-	"{0..n - 1}\<inter>{n} ={}" using n0
-	by (simp_all add: expand_set_eq, presburger+)
+	"{0..n - 1}\<inter>{n} ={}" using n0 by simp_all
       have f: "finite {0}" "finite {1}" "finite {2 .. n}"
 	"finite {0 .. n - 1}" "finite {n}" by simp_all
     have "((1 - ?X) * ?sa) $ n = setsum (\<lambda>i. (1 - ?X)$ i * ?sa $ (n - i)) {0 .. n}"
@@ -2503,6 +2501,29 @@
   then show ?thesis unfolding fps_inv_right[OF c0 c1] by simp
 qed
 
+lemma fps_ginv_deriv:
+  assumes a0:"a$0 = (0::'a::{field})" and a1: "a$1 \<noteq> 0"
+  shows "fps_deriv (fps_ginv b a) = (fps_deriv b / fps_deriv a) oo fps_ginv X a"
+proof-
+  let ?ia = "fps_ginv b a"
+  let ?iXa = "fps_ginv X a"
+  let ?d = "fps_deriv"
+  let ?dia = "?d ?ia"
+  have iXa0: "?iXa $ 0 = 0" by (simp add: fps_ginv_def)
+  have da0: "?d a $ 0 \<noteq> 0" using a1 by simp
+  from fps_ginv[OF a0 a1, of b] have "?d (?ia oo a) = fps_deriv b" by simp
+  then have "(?d ?ia oo a) * ?d a = ?d b" unfolding fps_compose_deriv[OF a0] .
+  then have "(?d ?ia oo a) * ?d a * inverse (?d a) = ?d b * inverse (?d a)" by simp
+  then have "(?d ?ia oo a) * (inverse (?d a) * ?d a) = ?d b / ?d a" 
+    by (simp add: fps_divide_def)
+  then have "(?d ?ia oo a) oo ?iXa =  (?d b / ?d a) oo ?iXa "
+    unfolding inverse_mult_eq_1[OF da0] by simp
+  then have "?d ?ia oo (a oo ?iXa) =  (?d b / ?d a) oo ?iXa"
+    unfolding fps_compose_assoc[OF iXa0 a0] .
+  then show ?thesis unfolding fps_inv_ginv[symmetric]
+    unfolding fps_inv_right[OF a0 a1] by simp
+qed
+
 subsection{* Elementary series *}
 
 subsubsection{* Exponential series *}
--- a/src/HOL/Library/Fundamental_Theorem_Algebra.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Library/Fundamental_Theorem_Algebra.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -577,11 +577,12 @@
 next
   case (pCons c cs)
   {assume c0: "c = 0"
-    from pCons.hyps pCons.prems c0 have ?case apply auto
+    from pCons.hyps pCons.prems c0 have ?case
+      apply (auto)
       apply (rule_tac x="k+1" in exI)
       apply (rule_tac x="a" in exI, clarsimp)
       apply (rule_tac x="q" in exI)
-      by (auto simp add: power_Suc)}
+      by (auto)}
   moreover
   {assume c0: "c\<noteq>0"
     hence ?case apply-
--- a/src/HOL/Library/Legacy_GCD.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,787 +0,0 @@
-(*  Title:      HOL/GCD.thy
-    Author:     Christophe Tabacznyj and Lawrence C Paulson
-    Copyright   1996  University of Cambridge
-*)
-
-header {* The Greatest Common Divisor *}
-
-theory Legacy_GCD
-imports Main
-begin
-
-text {*
-  See \cite{davenport92}. \bigskip
-*}
-
-subsection {* Specification of GCD on nats *}
-
-definition
-  is_gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where -- {* @{term gcd} as a relation *}
-  [code del]: "is_gcd m n p \<longleftrightarrow> p dvd m \<and> p dvd n \<and>
-    (\<forall>d. d dvd m \<longrightarrow> d dvd n \<longrightarrow> d dvd p)"
-
-text {* Uniqueness *}
-
-lemma is_gcd_unique: "is_gcd a b m \<Longrightarrow> is_gcd a b n \<Longrightarrow> m = n"
-  by (simp add: is_gcd_def) (blast intro: dvd_anti_sym)
-
-text {* Connection to divides relation *}
-
-lemma is_gcd_dvd: "is_gcd a b m \<Longrightarrow> k dvd a \<Longrightarrow> k dvd b \<Longrightarrow> k dvd m"
-  by (auto simp add: is_gcd_def)
-
-text {* Commutativity *}
-
-lemma is_gcd_commute: "is_gcd m n k = is_gcd n m k"
-  by (auto simp add: is_gcd_def)
-
-
-subsection {* GCD on nat by Euclid's algorithm *}
-
-fun
-  gcd  :: "nat => nat => nat"
-where
-  "gcd m n = (if n = 0 then m else gcd n (m mod n))"
-lemma gcd_induct [case_names "0" rec]:
-  fixes m n :: nat
-  assumes "\<And>m. P m 0"
-    and "\<And>m n. 0 < n \<Longrightarrow> P n (m mod n) \<Longrightarrow> P m n"
-  shows "P m n"
-proof (induct m n rule: gcd.induct)
-  case (1 m n) with assms show ?case by (cases "n = 0") simp_all
-qed
-
-lemma gcd_0 [simp, algebra]: "gcd m 0 = m"
-  by simp
-
-lemma gcd_0_left [simp,algebra]: "gcd 0 m = m"
-  by simp
-
-lemma gcd_non_0: "n > 0 \<Longrightarrow> gcd m n = gcd n (m mod n)"
-  by simp
-
-lemma gcd_1 [simp, algebra]: "gcd m (Suc 0) = Suc 0"
-  by simp
-
-lemma nat_gcd_1_right [simp, algebra]: "gcd m 1 = 1"
-  unfolding One_nat_def by (rule gcd_1)
-
-declare gcd.simps [simp del]
-
-text {*
-  \medskip @{term "gcd m n"} divides @{text m} and @{text n}.  The
-  conjunctions don't seem provable separately.
-*}
-
-lemma gcd_dvd1 [iff, algebra]: "gcd m n dvd m"
-  and gcd_dvd2 [iff, algebra]: "gcd m n dvd n"
-  apply (induct m n rule: gcd_induct)
-     apply (simp_all add: gcd_non_0)
-  apply (blast dest: dvd_mod_imp_dvd)
-  done
-
-text {*
-  \medskip Maximality: for all @{term m}, @{term n}, @{term k}
-  naturals, if @{term k} divides @{term m} and @{term k} divides
-  @{term n} then @{term k} divides @{term "gcd m n"}.
-*}
-
-lemma gcd_greatest: "k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd gcd m n"
-  by (induct m n rule: gcd_induct) (simp_all add: gcd_non_0 dvd_mod)
-
-text {*
-  \medskip Function gcd yields the Greatest Common Divisor.
-*}
-
-lemma is_gcd: "is_gcd m n (gcd m n) "
-  by (simp add: is_gcd_def gcd_greatest)
-
-
-subsection {* Derived laws for GCD *}
-
-lemma gcd_greatest_iff [iff, algebra]: "k dvd gcd m n \<longleftrightarrow> k dvd m \<and> k dvd n"
-  by (blast intro!: gcd_greatest intro: dvd_trans)
-
-lemma gcd_zero[algebra]: "gcd m n = 0 \<longleftrightarrow> m = 0 \<and> n = 0"
-  by (simp only: dvd_0_left_iff [symmetric] gcd_greatest_iff)
-
-lemma gcd_commute: "gcd m n = gcd n m"
-  apply (rule is_gcd_unique)
-   apply (rule is_gcd)
-  apply (subst is_gcd_commute)
-  apply (simp add: is_gcd)
-  done
-
-lemma gcd_assoc: "gcd (gcd k m) n = gcd k (gcd m n)"
-  apply (rule is_gcd_unique)
-   apply (rule is_gcd)
-  apply (simp add: is_gcd_def)
-  apply (blast intro: dvd_trans)
-  done
-
-lemma gcd_1_left [simp, algebra]: "gcd (Suc 0) m = Suc 0"
-  by (simp add: gcd_commute)
-
-lemma nat_gcd_1_left [simp, algebra]: "gcd 1 m = 1"
-  unfolding One_nat_def by (rule gcd_1_left)
-
-text {*
-  \medskip Multiplication laws
-*}
-
-lemma gcd_mult_distrib2: "k * gcd m n = gcd (k * m) (k * n)"
-    -- {* \cite[page 27]{davenport92} *}
-  apply (induct m n rule: gcd_induct)
-   apply simp
-  apply (case_tac "k = 0")
-   apply (simp_all add: mod_geq gcd_non_0 mod_mult_distrib2)
-  done
-
-lemma gcd_mult [simp, algebra]: "gcd k (k * n) = k"
-  apply (rule gcd_mult_distrib2 [of k 1 n, simplified, symmetric])
-  done
-
-lemma gcd_self [simp, algebra]: "gcd k k = k"
-  apply (rule gcd_mult [of k 1, simplified])
-  done
-
-lemma relprime_dvd_mult: "gcd k n = 1 ==> k dvd m * n ==> k dvd m"
-  apply (insert gcd_mult_distrib2 [of m k n])
-  apply simp
-  apply (erule_tac t = m in ssubst)
-  apply simp
-  done
-
-lemma relprime_dvd_mult_iff: "gcd k n = 1 ==> (k dvd m * n) = (k dvd m)"
-  by (auto intro: relprime_dvd_mult dvd_mult2)
-
-lemma gcd_mult_cancel: "gcd k n = 1 ==> gcd (k * m) n = gcd m n"
-  apply (rule dvd_anti_sym)
-   apply (rule gcd_greatest)
-    apply (rule_tac n = k in relprime_dvd_mult)
-     apply (simp add: gcd_assoc)
-     apply (simp add: gcd_commute)
-    apply (simp_all add: mult_commute)
-  done
-
-
-text {* \medskip Addition laws *}
-
-lemma gcd_add1 [simp, algebra]: "gcd (m + n) n = gcd m n"
-  by (cases "n = 0") (auto simp add: gcd_non_0)
-
-lemma gcd_add2 [simp, algebra]: "gcd m (m + n) = gcd m n"
-proof -
-  have "gcd m (m + n) = gcd (m + n) m" by (rule gcd_commute)
-  also have "... = gcd (n + m) m" by (simp add: add_commute)
-  also have "... = gcd n m" by simp
-  also have  "... = gcd m n" by (rule gcd_commute)
-  finally show ?thesis .
-qed
-
-lemma gcd_add2' [simp, algebra]: "gcd m (n + m) = gcd m n"
-  apply (subst add_commute)
-  apply (rule gcd_add2)
-  done
-
-lemma gcd_add_mult[algebra]: "gcd m (k * m + n) = gcd m n"
-  by (induct k) (simp_all add: add_assoc)
-
-lemma gcd_dvd_prod: "gcd m n dvd m * n" 
-  using mult_dvd_mono [of 1] by auto
-
-text {*
-  \medskip Division by gcd yields rrelatively primes.
-*}
-
-lemma div_gcd_relprime:
-  assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
-  shows "gcd (a div gcd a b) (b div gcd a b) = 1"
-proof -
-  let ?g = "gcd a b"
-  let ?a' = "a div ?g"
-  let ?b' = "b div ?g"
-  let ?g' = "gcd ?a' ?b'"
-  have dvdg: "?g dvd a" "?g dvd b" by simp_all
-  have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by simp_all
-  from dvdg dvdg' obtain ka kb ka' kb' where
-      kab: "a = ?g * ka" "b = ?g * kb" "?a' = ?g' * ka'" "?b' = ?g' * kb'"
-    unfolding dvd_def by blast
-  then have "?g * ?a' = (?g * ?g') * ka'" "?g * ?b' = (?g * ?g') * kb'" by simp_all
-  then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
-    by (auto simp add: dvd_mult_div_cancel [OF dvdg(1)]
-      dvd_mult_div_cancel [OF dvdg(2)] dvd_def)
-  have "?g \<noteq> 0" using nz by (simp add: gcd_zero)
-  then have gp: "?g > 0" by simp
-  from gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
-  with dvd_mult_cancel1 [OF gp] show "?g' = 1" by simp
-qed
-
-
-lemma gcd_unique: "d dvd a\<and>d dvd b \<and> (\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"
-proof(auto)
-  assume H: "d dvd a" "d dvd b" "\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d"
-  from H(3)[rule_format] gcd_dvd1[of a b] gcd_dvd2[of a b] 
-  have th: "gcd a b dvd d" by blast
-  from dvd_anti_sym[OF th gcd_greatest[OF H(1,2)]]  show "d = gcd a b" by blast 
-qed
-
-lemma gcd_eq: assumes H: "\<forall>d. d dvd x \<and> d dvd y \<longleftrightarrow> d dvd u \<and> d dvd v"
-  shows "gcd x y = gcd u v"
-proof-
-  from H have "\<forall>d. d dvd x \<and> d dvd y \<longleftrightarrow> d dvd gcd u v" by simp
-  with gcd_unique[of "gcd u v" x y]  show ?thesis by auto
-qed
-
-lemma ind_euclid: 
-  assumes c: " \<forall>a b. P (a::nat) b \<longleftrightarrow> P b a" and z: "\<forall>a. P a 0" 
-  and add: "\<forall>a b. P a b \<longrightarrow> P a (a + b)" 
-  shows "P a b"
-proof(induct n\<equiv>"a+b" arbitrary: a b rule: nat_less_induct)
-  fix n a b
-  assume H: "\<forall>m < n. \<forall>a b. m = a + b \<longrightarrow> P a b" "n = a + b"
-  have "a = b \<or> a < b \<or> b < a" by arith
-  moreover {assume eq: "a= b"
-    from add[rule_format, OF z[rule_format, of a]] have "P a b" using eq by simp}
-  moreover
-  {assume lt: "a < b"
-    hence "a + b - a < n \<or> a = 0"  using H(2) by arith
-    moreover
-    {assume "a =0" with z c have "P a b" by blast }
-    moreover
-    {assume ab: "a + b - a < n"
-      have th0: "a + b - a = a + (b - a)" using lt by arith
-      from add[rule_format, OF H(1)[rule_format, OF ab th0]]
-      have "P a b" by (simp add: th0[symmetric])}
-    ultimately have "P a b" by blast}
-  moreover
-  {assume lt: "a > b"
-    hence "b + a - b < n \<or> b = 0"  using H(2) by arith
-    moreover
-    {assume "b =0" with z c have "P a b" by blast }
-    moreover
-    {assume ab: "b + a - b < n"
-      have th0: "b + a - b = b + (a - b)" using lt by arith
-      from add[rule_format, OF H(1)[rule_format, OF ab th0]]
-      have "P b a" by (simp add: th0[symmetric])
-      hence "P a b" using c by blast }
-    ultimately have "P a b" by blast}
-ultimately  show "P a b" by blast
-qed
-
-lemma bezout_lemma: 
-  assumes ex: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x = b * y + d \<or> b * x = a * y + d)"
-  shows "\<exists>d x y. d dvd a \<and> d dvd a + b \<and> (a * x = (a + b) * y + d \<or> (a + b) * x = a * y + d)"
-using ex
-apply clarsimp
-apply (rule_tac x="d" in exI, simp add: dvd_add)
-apply (case_tac "a * x = b * y + d" , simp_all)
-apply (rule_tac x="x + y" in exI)
-apply (rule_tac x="y" in exI)
-apply algebra
-apply (rule_tac x="x" in exI)
-apply (rule_tac x="x + y" in exI)
-apply algebra
-done
-
-lemma bezout_add: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x = b * y + d \<or> b * x = a * y + d)"
-apply(induct a b rule: ind_euclid)
-apply blast
-apply clarify
-apply (rule_tac x="a" in exI, simp add: dvd_add)
-apply clarsimp
-apply (rule_tac x="d" in exI)
-apply (case_tac "a * x = b * y + d", simp_all add: dvd_add)
-apply (rule_tac x="x+y" in exI)
-apply (rule_tac x="y" in exI)
-apply algebra
-apply (rule_tac x="x" in exI)
-apply (rule_tac x="x+y" in exI)
-apply algebra
-done
-
-lemma bezout: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x - b * y = d \<or> b * x - a * y = d)"
-using bezout_add[of a b]
-apply clarsimp
-apply (rule_tac x="d" in exI, simp)
-apply (rule_tac x="x" in exI)
-apply (rule_tac x="y" in exI)
-apply auto
-done
-
-
-text {* We can get a stronger version with a nonzeroness assumption. *}
-lemma divides_le: "m dvd n ==> m <= n \<or> n = (0::nat)" by (auto simp add: dvd_def)
-
-lemma bezout_add_strong: assumes nz: "a \<noteq> (0::nat)"
-  shows "\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d"
-proof-
-  from nz have ap: "a > 0" by simp
- from bezout_add[of a b] 
- have "(\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d) \<or> (\<exists>d x y. d dvd a \<and> d dvd b \<and> b * x = a * y + d)" by blast
- moreover
- {fix d x y assume H: "d dvd a" "d dvd b" "a * x = b * y + d"
-   from H have ?thesis by blast }
- moreover
- {fix d x y assume H: "d dvd a" "d dvd b" "b * x = a * y + d"
-   {assume b0: "b = 0" with H  have ?thesis by simp}
-   moreover 
-   {assume b: "b \<noteq> 0" hence bp: "b > 0" by simp
-     from divides_le[OF H(2)] b have "d < b \<or> d = b" using le_less by blast
-     moreover
-     {assume db: "d=b"
-       from prems have ?thesis apply simp
-	 apply (rule exI[where x = b], simp)
-	 apply (rule exI[where x = b])
-	by (rule exI[where x = "a - 1"], simp add: diff_mult_distrib2)}
-    moreover
-    {assume db: "d < b" 
-	{assume "x=0" hence ?thesis  using prems by simp }
-	moreover
-	{assume x0: "x \<noteq> 0" hence xp: "x > 0" by simp
-	  
-	  from db have "d \<le> b - 1" by simp
-	  hence "d*b \<le> b*(b - 1)" by simp
-	  with xp mult_mono[of "1" "x" "d*b" "b*(b - 1)"]
-	  have dble: "d*b \<le> x*b*(b - 1)" using bp by simp
-	  from H (3) have "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x*b*(b - 1)" by algebra
-	  hence "a * ((b - 1) * y) = d + x*b*(b - 1) - d*b" using bp by simp
-	  hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)" 
-	    by (simp only: diff_add_assoc[OF dble, of d, symmetric])
-	  hence "a * ((b - 1) * y) = b*(x*(b - 1) - d) + d"
-	    by (simp only: diff_mult_distrib2 add_commute mult_ac)
-	  hence ?thesis using H(1,2)
-	    apply -
-	    apply (rule exI[where x=d], simp)
-	    apply (rule exI[where x="(b - 1) * y"])
-	    by (rule exI[where x="x*(b - 1) - d"], simp)}
-	ultimately have ?thesis by blast}
-    ultimately have ?thesis by blast}
-  ultimately have ?thesis by blast}
- ultimately show ?thesis by blast
-qed
-
-
-lemma bezout_gcd: "\<exists>x y. a * x - b * y = gcd a b \<or> b * x - a * y = gcd a b"
-proof-
-  let ?g = "gcd a b"
-  from bezout[of a b] obtain d x y where d: "d dvd a" "d dvd b" "a * x - b * y = d \<or> b * x - a * y = d" by blast
-  from d(1,2) have "d dvd ?g" by simp
-  then obtain k where k: "?g = d*k" unfolding dvd_def by blast
-  from d(3) have "(a * x - b * y)*k = d*k \<or> (b * x - a * y)*k = d*k" by blast 
-  hence "a * x * k - b * y*k = d*k \<or> b * x * k - a * y*k = d*k" 
-    by (algebra add: diff_mult_distrib)
-  hence "a * (x * k) - b * (y*k) = ?g \<or> b * (x * k) - a * (y*k) = ?g" 
-    by (simp add: k mult_assoc)
-  thus ?thesis by blast
-qed
-
-lemma bezout_gcd_strong: assumes a: "a \<noteq> 0" 
-  shows "\<exists>x y. a * x = b * y + gcd a b"
-proof-
-  let ?g = "gcd a b"
-  from bezout_add_strong[OF a, of b]
-  obtain d x y where d: "d dvd a" "d dvd b" "a * x = b * y + d" by blast
-  from d(1,2) have "d dvd ?g" by simp
-  then obtain k where k: "?g = d*k" unfolding dvd_def by blast
-  from d(3) have "a * x * k = (b * y + d) *k " by algebra
-  hence "a * (x * k) = b * (y*k) + ?g" by (algebra add: k)
-  thus ?thesis by blast
-qed
-
-lemma gcd_mult_distrib: "gcd(a * c) (b * c) = c * gcd a b"
-by(simp add: gcd_mult_distrib2 mult_commute)
-
-lemma gcd_bezout: "(\<exists>x y. a * x - b * y = d \<or> b * x - a * y = d) \<longleftrightarrow> gcd a b dvd d"
-  (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
-  let ?g = "gcd a b"
-  {assume H: ?rhs then obtain k where k: "d = ?g*k" unfolding dvd_def by blast
-    from bezout_gcd[of a b] obtain x y where xy: "a * x - b * y = ?g \<or> b * x - a * y = ?g"
-      by blast
-    hence "(a * x - b * y)*k = ?g*k \<or> (b * x - a * y)*k = ?g*k" by auto
-    hence "a * x*k - b * y*k = ?g*k \<or> b * x * k - a * y*k = ?g*k" 
-      by (simp only: diff_mult_distrib)
-    hence "a * (x*k) - b * (y*k) = d \<or> b * (x * k) - a * (y*k) = d"
-      by (simp add: k[symmetric] mult_assoc)
-    hence ?lhs by blast}
-  moreover
-  {fix x y assume H: "a * x - b * y = d \<or> b * x - a * y = d"
-    have dv: "?g dvd a*x" "?g dvd b * y" "?g dvd b*x" "?g dvd a * y"
-      using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all
-    from dvd_diff_nat[OF dv(1,2)] dvd_diff_nat[OF dv(3,4)] H
-    have ?rhs by auto}
-  ultimately show ?thesis by blast
-qed
-
-lemma gcd_bezout_sum: assumes H:"a * x + b * y = d" shows "gcd a b dvd d"
-proof-
-  let ?g = "gcd a b"
-    have dv: "?g dvd a*x" "?g dvd b * y" 
-      using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all
-    from dvd_add[OF dv] H
-    show ?thesis by auto
-qed
-
-lemma gcd_mult': "gcd b (a * b) = b"
-by (simp add: gcd_mult mult_commute[of a b]) 
-
-lemma gcd_add: "gcd(a + b) b = gcd a b" 
-  "gcd(b + a) b = gcd a b" "gcd a (a + b) = gcd a b" "gcd a (b + a) = gcd a b"
-apply (simp_all add: gcd_add1)
-by (simp add: gcd_commute gcd_add1)
-
-lemma gcd_sub: "b <= a ==> gcd(a - b) b = gcd a b" "a <= b ==> gcd a (b - a) = gcd a b"
-proof-
-  {fix a b assume H: "b \<le> (a::nat)"
-    hence th: "a - b + b = a" by arith
-    from gcd_add(1)[of "a - b" b] th  have "gcd(a - b) b = gcd a b" by simp}
-  note th = this
-{
-  assume ab: "b \<le> a"
-  from th[OF ab] show "gcd (a - b)  b = gcd a b" by blast
-next
-  assume ab: "a \<le> b"
-  from th[OF ab] show "gcd a (b - a) = gcd a b" 
-    by (simp add: gcd_commute)}
-qed
-
-
-subsection {* LCM defined by GCD *}
-
-
-definition
-  lcm :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-where
-  lcm_def: "lcm m n = m * n div gcd m n"
-
-lemma prod_gcd_lcm:
-  "m * n = gcd m n * lcm m n"
-  unfolding lcm_def by (simp add: dvd_mult_div_cancel [OF gcd_dvd_prod])
-
-lemma lcm_0 [simp]: "lcm m 0 = 0"
-  unfolding lcm_def by simp
-
-lemma lcm_1 [simp]: "lcm m 1 = m"
-  unfolding lcm_def by simp
-
-lemma lcm_0_left [simp]: "lcm 0 n = 0"
-  unfolding lcm_def by simp
-
-lemma lcm_1_left [simp]: "lcm 1 m = m"
-  unfolding lcm_def by simp
-
-lemma dvd_pos:
-  fixes n m :: nat
-  assumes "n > 0" and "m dvd n"
-  shows "m > 0"
-using assms by (cases m) auto
-
-lemma lcm_least:
-  assumes "m dvd k" and "n dvd k"
-  shows "lcm m n dvd k"
-proof (cases k)
-  case 0 then show ?thesis by auto
-next
-  case (Suc _) then have pos_k: "k > 0" by auto
-  from assms dvd_pos [OF this] have pos_mn: "m > 0" "n > 0" by auto
-  with gcd_zero [of m n] have pos_gcd: "gcd m n > 0" by simp
-  from assms obtain p where k_m: "k = m * p" using dvd_def by blast
-  from assms obtain q where k_n: "k = n * q" using dvd_def by blast
-  from pos_k k_m have pos_p: "p > 0" by auto
-  from pos_k k_n have pos_q: "q > 0" by auto
-  have "k * k * gcd q p = k * gcd (k * q) (k * p)"
-    by (simp add: mult_ac gcd_mult_distrib2)
-  also have "\<dots> = k * gcd (m * p * q) (n * q * p)"
-    by (simp add: k_m [symmetric] k_n [symmetric])
-  also have "\<dots> = k * p * q * gcd m n"
-    by (simp add: mult_ac gcd_mult_distrib2)
-  finally have "(m * p) * (n * q) * gcd q p = k * p * q * gcd m n"
-    by (simp only: k_m [symmetric] k_n [symmetric])
-  then have "p * q * m * n * gcd q p = p * q * k * gcd m n"
-    by (simp add: mult_ac)
-  with pos_p pos_q have "m * n * gcd q p = k * gcd m n"
-    by simp
-  with prod_gcd_lcm [of m n]
-  have "lcm m n * gcd q p * gcd m n = k * gcd m n"
-    by (simp add: mult_ac)
-  with pos_gcd have "lcm m n * gcd q p = k" by simp
-  then show ?thesis using dvd_def by auto
-qed
-
-lemma lcm_dvd1 [iff]:
-  "m dvd lcm m n"
-proof (cases m)
-  case 0 then show ?thesis by simp
-next
-  case (Suc _)
-  then have mpos: "m > 0" by simp
-  show ?thesis
-  proof (cases n)
-    case 0 then show ?thesis by simp
-  next
-    case (Suc _)
-    then have npos: "n > 0" by simp
-    have "gcd m n dvd n" by simp
-    then obtain k where "n = gcd m n * k" using dvd_def by auto
-    then have "m * n div gcd m n = m * (gcd m n * k) div gcd m n" by (simp add: mult_ac)
-    also have "\<dots> = m * k" using mpos npos gcd_zero by simp
-    finally show ?thesis by (simp add: lcm_def)
-  qed
-qed
-
-lemma lcm_dvd2 [iff]: 
-  "n dvd lcm m n"
-proof (cases n)
-  case 0 then show ?thesis by simp
-next
-  case (Suc _)
-  then have npos: "n > 0" by simp
-  show ?thesis
-  proof (cases m)
-    case 0 then show ?thesis by simp
-  next
-    case (Suc _)
-    then have mpos: "m > 0" by simp
-    have "gcd m n dvd m" by simp
-    then obtain k where "m = gcd m n * k" using dvd_def by auto
-    then have "m * n div gcd m n = (gcd m n * k) * n div gcd m n" by (simp add: mult_ac)
-    also have "\<dots> = n * k" using mpos npos gcd_zero by simp
-    finally show ?thesis by (simp add: lcm_def)
-  qed
-qed
-
-lemma gcd_add1_eq: "gcd (m + k) k = gcd (m + k) m"
-  by (simp add: gcd_commute)
-
-lemma gcd_diff2: "m \<le> n ==> gcd n (n - m) = gcd n m"
-  apply (subgoal_tac "n = m + (n - m)")
-  apply (erule ssubst, rule gcd_add1_eq, simp)  
-  done
-
-
-subsection {* GCD and LCM on integers *}
-
-definition
-  zgcd :: "int \<Rightarrow> int \<Rightarrow> int" where
-  "zgcd i j = int (gcd (nat (abs i)) (nat (abs j)))"
-
-lemma zgcd_zdvd1 [iff,simp, algebra]: "zgcd i j dvd i"
-by (simp add: zgcd_def int_dvd_iff)
-
-lemma zgcd_zdvd2 [iff,simp, algebra]: "zgcd i j dvd j"
-by (simp add: zgcd_def int_dvd_iff)
-
-lemma zgcd_pos: "zgcd i j \<ge> 0"
-by (simp add: zgcd_def)
-
-lemma zgcd0 [simp,algebra]: "(zgcd i j = 0) = (i = 0 \<and> j = 0)"
-by (simp add: zgcd_def gcd_zero)
-
-lemma zgcd_commute: "zgcd i j = zgcd j i"
-unfolding zgcd_def by (simp add: gcd_commute)
-
-lemma zgcd_zminus [simp, algebra]: "zgcd (- i) j = zgcd i j"
-unfolding zgcd_def by simp
-
-lemma zgcd_zminus2 [simp, algebra]: "zgcd i (- j) = zgcd i j"
-unfolding zgcd_def by simp
-
-  (* should be solved by algebra*)
-lemma zrelprime_dvd_mult: "zgcd i j = 1 \<Longrightarrow> i dvd k * j \<Longrightarrow> i dvd k"
-  unfolding zgcd_def
-proof -
-  assume "int (gcd (nat \<bar>i\<bar>) (nat \<bar>j\<bar>)) = 1" "i dvd k * j"
-  then have g: "gcd (nat \<bar>i\<bar>) (nat \<bar>j\<bar>) = 1" by simp
-  from `i dvd k * j` obtain h where h: "k*j = i*h" unfolding dvd_def by blast
-  have th: "nat \<bar>i\<bar> dvd nat \<bar>k\<bar> * nat \<bar>j\<bar>"
-    unfolding dvd_def
-    by (rule_tac x= "nat \<bar>h\<bar>" in exI, simp add: h nat_abs_mult_distrib [symmetric])
-  from relprime_dvd_mult [OF g th] obtain h' where h': "nat \<bar>k\<bar> = nat \<bar>i\<bar> * h'"
-    unfolding dvd_def by blast
-  from h' have "int (nat \<bar>k\<bar>) = int (nat \<bar>i\<bar> * h')" by simp
-  then have "\<bar>k\<bar> = \<bar>i\<bar> * int h'" by (simp add: int_mult)
-  then show ?thesis
-    apply (subst abs_dvd_iff [symmetric])
-    apply (subst dvd_abs_iff [symmetric])
-    apply (unfold dvd_def)
-    apply (rule_tac x = "int h'" in exI, simp)
-    done
-qed
-
-lemma int_nat_abs: "int (nat (abs x)) = abs x" by arith
-
-lemma zgcd_greatest:
-  assumes "k dvd m" and "k dvd n"
-  shows "k dvd zgcd m n"
-proof -
-  let ?k' = "nat \<bar>k\<bar>"
-  let ?m' = "nat \<bar>m\<bar>"
-  let ?n' = "nat \<bar>n\<bar>"
-  from `k dvd m` and `k dvd n` have dvd': "?k' dvd ?m'" "?k' dvd ?n'"
-    unfolding zdvd_int by (simp_all only: int_nat_abs abs_dvd_iff dvd_abs_iff)
-  from gcd_greatest [OF dvd'] have "int (nat \<bar>k\<bar>) dvd zgcd m n"
-    unfolding zgcd_def by (simp only: zdvd_int)
-  then have "\<bar>k\<bar> dvd zgcd m n" by (simp only: int_nat_abs)
-  then show "k dvd zgcd m n" by simp
-qed
-
-lemma div_zgcd_relprime:
-  assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
-  shows "zgcd (a div (zgcd a b)) (b div (zgcd a b)) = 1"
-proof -
-  from nz have nz': "nat \<bar>a\<bar> \<noteq> 0 \<or> nat \<bar>b\<bar> \<noteq> 0" by arith 
-  let ?g = "zgcd a b"
-  let ?a' = "a div ?g"
-  let ?b' = "b div ?g"
-  let ?g' = "zgcd ?a' ?b'"
-  have dvdg: "?g dvd a" "?g dvd b" by (simp_all add: zgcd_zdvd1 zgcd_zdvd2)
-  have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by (simp_all add: zgcd_zdvd1 zgcd_zdvd2)
-  from dvdg dvdg' obtain ka kb ka' kb' where
-   kab: "a = ?g*ka" "b = ?g*kb" "?a' = ?g'*ka'" "?b' = ?g' * kb'"
-    unfolding dvd_def by blast
-  then have "?g* ?a' = (?g * ?g') * ka'" "?g* ?b' = (?g * ?g') * kb'" by simp_all
-  then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
-    by (auto simp add: zdvd_mult_div_cancel [OF dvdg(1)]
-      zdvd_mult_div_cancel [OF dvdg(2)] dvd_def)
-  have "?g \<noteq> 0" using nz by simp
-  then have gp: "?g \<noteq> 0" using zgcd_pos[where i="a" and j="b"] by arith
-  from zgcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
-  with zdvd_mult_cancel1 [OF gp] have "\<bar>?g'\<bar> = 1" by simp
-  with zgcd_pos show "?g' = 1" by simp
-qed
-
-lemma zgcd_0 [simp, algebra]: "zgcd m 0 = abs m"
-  by (simp add: zgcd_def abs_if)
-
-lemma zgcd_0_left [simp, algebra]: "zgcd 0 m = abs m"
-  by (simp add: zgcd_def abs_if)
-
-lemma zgcd_non_0: "0 < n ==> zgcd m n = zgcd n (m mod n)"
-  apply (frule_tac b = n and a = m in pos_mod_sign)
-  apply (simp del: pos_mod_sign add: zgcd_def abs_if nat_mod_distrib)
-  apply (auto simp add: gcd_non_0 nat_mod_distrib [symmetric] zmod_zminus1_eq_if)
-  apply (frule_tac a = m in pos_mod_bound)
-  apply (simp del: pos_mod_bound add: nat_diff_distrib gcd_diff2 nat_le_eq_zle)
-  done
-
-lemma zgcd_eq: "zgcd m n = zgcd n (m mod n)"
-  apply (case_tac "n = 0", simp add: DIVISION_BY_ZERO)
-  apply (auto simp add: linorder_neq_iff zgcd_non_0)
-  apply (cut_tac m = "-m" and n = "-n" in zgcd_non_0, auto)
-  done
-
-lemma zgcd_1 [simp, algebra]: "zgcd m 1 = 1"
-  by (simp add: zgcd_def abs_if)
-
-lemma zgcd_0_1_iff [simp, algebra]: "zgcd 0 m = 1 \<longleftrightarrow> \<bar>m\<bar> = 1"
-  by (simp add: zgcd_def abs_if)
-
-lemma zgcd_greatest_iff[algebra]: "k dvd zgcd m n = (k dvd m \<and> k dvd n)"
-  by (simp add: zgcd_def abs_if int_dvd_iff dvd_int_iff nat_dvd_iff)
-
-lemma zgcd_1_left [simp, algebra]: "zgcd 1 m = 1"
-  by (simp add: zgcd_def gcd_1_left)
-
-lemma zgcd_assoc: "zgcd (zgcd k m) n = zgcd k (zgcd m n)"
-  by (simp add: zgcd_def gcd_assoc)
-
-lemma zgcd_left_commute: "zgcd k (zgcd m n) = zgcd m (zgcd k n)"
-  apply (rule zgcd_commute [THEN trans])
-  apply (rule zgcd_assoc [THEN trans])
-  apply (rule zgcd_commute [THEN arg_cong])
-  done
-
-lemmas zgcd_ac = zgcd_assoc zgcd_commute zgcd_left_commute
-  -- {* addition is an AC-operator *}
-
-lemma zgcd_zmult_distrib2: "0 \<le> k ==> k * zgcd m n = zgcd (k * m) (k * n)"
-  by (simp del: minus_mult_right [symmetric]
-      add: minus_mult_right nat_mult_distrib zgcd_def abs_if
-          mult_less_0_iff gcd_mult_distrib2 [symmetric] zmult_int [symmetric])
-
-lemma zgcd_zmult_distrib2_abs: "zgcd (k * m) (k * n) = abs k * zgcd m n"
-  by (simp add: abs_if zgcd_zmult_distrib2)
-
-lemma zgcd_self [simp]: "0 \<le> m ==> zgcd m m = m"
-  by (cut_tac k = m and m = 1 and n = 1 in zgcd_zmult_distrib2, simp_all)
-
-lemma zgcd_zmult_eq_self [simp]: "0 \<le> k ==> zgcd k (k * n) = k"
-  by (cut_tac k = k and m = 1 and n = n in zgcd_zmult_distrib2, simp_all)
-
-lemma zgcd_zmult_eq_self2 [simp]: "0 \<le> k ==> zgcd (k * n) k = k"
-  by (cut_tac k = k and m = n and n = 1 in zgcd_zmult_distrib2, simp_all)
-
-
-definition "zlcm i j = int (lcm(nat(abs i)) (nat(abs j)))"
-
-lemma dvd_zlcm_self1[simp, algebra]: "i dvd zlcm i j"
-by(simp add:zlcm_def dvd_int_iff)
-
-lemma dvd_zlcm_self2[simp, algebra]: "j dvd zlcm i j"
-by(simp add:zlcm_def dvd_int_iff)
-
-
-lemma dvd_imp_dvd_zlcm1:
-  assumes "k dvd i" shows "k dvd (zlcm i j)"
-proof -
-  have "nat(abs k) dvd nat(abs i)" using `k dvd i`
-    by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric])
-  thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans)
-qed
-
-lemma dvd_imp_dvd_zlcm2:
-  assumes "k dvd j" shows "k dvd (zlcm i j)"
-proof -
-  have "nat(abs k) dvd nat(abs j)" using `k dvd j`
-    by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric])
-  thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans)
-qed
-
-
-lemma zdvd_self_abs1: "(d::int) dvd (abs d)"
-by (case_tac "d <0", simp_all)
-
-lemma zdvd_self_abs2: "(abs (d::int)) dvd d"
-by (case_tac "d<0", simp_all)
-
-(* lcm a b is positive for positive a and b *)
-
-lemma lcm_pos: 
-  assumes mpos: "m > 0"
-  and npos: "n>0"
-  shows "lcm m n > 0"
-proof(rule ccontr, simp add: lcm_def gcd_zero)
-assume h:"m*n div gcd m n = 0"
-from mpos npos have "gcd m n \<noteq> 0" using gcd_zero by simp
-hence gcdp: "gcd m n > 0" by simp
-with h
-have "m*n < gcd m n"
-  by (cases "m * n < gcd m n") (auto simp add: div_if[OF gcdp, where m="m*n"])
-moreover 
-have "gcd m n dvd m" by simp
- with mpos dvd_imp_le have t1:"gcd m n \<le> m" by simp
- with npos have t1:"gcd m n *n \<le> m*n" by simp
- have "gcd m n \<le> gcd m n*n" using npos by simp
- with t1 have "gcd m n \<le> m*n" by arith
-ultimately show "False" by simp
-qed
-
-lemma zlcm_pos: 
-  assumes anz: "a \<noteq> 0"
-  and bnz: "b \<noteq> 0" 
-  shows "0 < zlcm a b"
-proof-
-  let ?na = "nat (abs a)"
-  let ?nb = "nat (abs b)"
-  have nap: "?na >0" using anz by simp
-  have nbp: "?nb >0" using bnz by simp
-  have "0 < lcm ?na ?nb" by (rule lcm_pos[OF nap nbp])
-  thus ?thesis by (simp add: zlcm_def)
-qed
-
-lemma zgcd_code [code]:
-  "zgcd k l = \<bar>if l = 0 then k else zgcd l (\<bar>k\<bar> mod \<bar>l\<bar>)\<bar>"
-  by (simp add: zgcd_def gcd.simps [of "nat \<bar>k\<bar>"] nat_mod_distrib)
-
-end
--- a/src/HOL/Library/Library.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Library/Library.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -43,11 +43,9 @@
   OptionalSugar
   Option_ord
   Permutation
-  Pocklington
   Poly_Deriv
   Polynomial
   Preorder
-  Primes
   Product_Vector
   Quicksort
   Quotient
--- a/src/HOL/Library/Multiset.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Library/Multiset.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -331,7 +331,7 @@
 
 lemma multiset_inter_count:
   "count (A #\<inter> B) x = min (count A x) (count B x)"
-by (simp add: multiset_inter_def min_def)
+by (simp add: multiset_inter_def)
 
 lemma multiset_inter_commute: "A #\<inter> B = B #\<inter> A"
 by (simp add: multiset_eq_conv_count_eq multiset_inter_count
@@ -353,7 +353,7 @@
 by (simp add: multiset_eq_conv_count_eq multiset_inter_count)
 
 lemma multiset_union_diff_commute: "B #\<inter> C = {#} \<Longrightarrow> A + B - C = A - C + B"
-apply (simp add: multiset_eq_conv_count_eq multiset_inter_count min_def
+apply (simp add: multiset_eq_conv_count_eq multiset_inter_count
     split: split_if_asm)
 apply clarsimp
 apply (erule_tac x = a in allE)
--- a/src/HOL/Library/Permutations.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Library/Permutations.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -843,9 +843,7 @@
 	unfolding permutes_def by metis+
       from eq have "(Fun.swap a b id o p) a  = (Fun.swap a c id o q) a" by simp
       hence bc: "b = c"
-	apply (simp add: permutes_def pa qa o_def fun_upd_def swap_def id_def cong del: if_weak_cong)
-	apply (cases "a = b", auto)
-	by (cases "b = c", auto)
+	by (simp add: permutes_def pa qa o_def fun_upd_def swap_def id_def cong del: if_weak_cong split: split_if_asm)
       from eq[unfolded bc] have "(\<lambda>p. Fun.swap a c id o p) (Fun.swap a c id o p) = (\<lambda>p. Fun.swap a c id o p) (Fun.swap a c id o q)" by simp
       hence "p = q" unfolding o_assoc swap_id_idempotent
 	by (simp add: o_def)
--- a/src/HOL/Library/Pocklington.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1263 +0,0 @@
-(*  Title:      HOL/Library/Pocklington.thy
-    Author:     Amine Chaieb
-*)
-
-header {* Pocklington's Theorem for Primes *}
-
-theory Pocklington
-imports Main Primes
-begin
-
-definition modeq:: "nat => nat => nat => bool"    ("(1[_ = _] '(mod _'))")
-  where "[a = b] (mod p) == ((a mod p) = (b mod p))"
-
-definition modneq:: "nat => nat => nat => bool"    ("(1[_ \<noteq> _] '(mod _'))")
-  where "[a \<noteq> b] (mod p) == ((a mod p) \<noteq> (b mod p))"
-
-lemma modeq_trans:
-  "\<lbrakk> [a = b] (mod p); [b = c] (mod p) \<rbrakk> \<Longrightarrow> [a = c] (mod p)"
-  by (simp add:modeq_def)
-
-
-lemma nat_mod_lemma: assumes xyn: "[x = y] (mod n)" and xy:"y \<le> x"
-  shows "\<exists>q. x = y + n * q"
-using xyn xy unfolding modeq_def using nat_mod_eq_lemma by blast
-
-lemma nat_mod[algebra]: "[x = y] (mod n) \<longleftrightarrow> (\<exists>q1 q2. x + n * q1 = y + n * q2)"
-unfolding modeq_def nat_mod_eq_iff ..
-
-(* Lemmas about previously defined terms.                                    *)
-
-lemma prime: "prime p \<longleftrightarrow> p \<noteq> 0 \<and> p\<noteq>1 \<and> (\<forall>m. 0 < m \<and> m < p \<longrightarrow> coprime p m)"
-  (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
-  {assume "p=0 \<or> p=1" hence ?thesis using prime_0 prime_1 by (cases "p=0", simp_all)}
-  moreover
-  {assume p0: "p\<noteq>0" "p\<noteq>1"
-    {assume H: "?lhs"
-      {fix m assume m: "m > 0" "m < p"
-	{assume "m=1" hence "coprime p m" by simp}
-	moreover
-	{assume "p dvd m" hence "p \<le> m" using dvd_imp_le m by blast with m(2)
-	  have "coprime p m" by simp}
-	ultimately have "coprime p m" using prime_coprime[OF H, of m] by blast}
-      hence ?rhs using p0 by auto}
-    moreover
-    { assume H: "\<forall>m. 0 < m \<and> m < p \<longrightarrow> coprime p m"
-      from prime_factor[OF p0(2)] obtain q where q: "prime q" "q dvd p" by blast
-      from prime_ge_2[OF q(1)] have q0: "q > 0" by arith
-      from dvd_imp_le[OF q(2)] p0 have qp: "q \<le> p" by arith
-      {assume "q = p" hence ?lhs using q(1) by blast}
-      moreover
-      {assume "q\<noteq>p" with qp have qplt: "q < p" by arith
-	from H[rule_format, of q] qplt q0 have "coprime p q" by arith
-	with coprime_prime[of p q q] q have False by simp hence ?lhs by blast}
-      ultimately have ?lhs by blast}
-    ultimately have ?thesis by blast}
-  ultimately show ?thesis  by (cases"p=0 \<or> p=1", auto)
-qed
-
-lemma finite_number_segment: "card { m. 0 < m \<and> m < n } = n - 1"
-proof-
-  have "{ m. 0 < m \<and> m < n } = {1..<n}" by auto
-  thus ?thesis by simp
-qed
-
-lemma coprime_mod: assumes n: "n \<noteq> 0" shows "coprime (a mod n) n \<longleftrightarrow> coprime a n"
-  using n dvd_mod_iff[of _ n a] by (auto simp add: coprime)
-
-(* Congruences.                                                              *)
-
-lemma cong_mod_01[simp,presburger]:
-  "[x = y] (mod 0) \<longleftrightarrow> x = y" "[x = y] (mod 1)" "[x = 0] (mod n) \<longleftrightarrow> n dvd x"
-  by (simp_all add: modeq_def, presburger)
-
-lemma cong_sub_cases:
-  "[x = y] (mod n) \<longleftrightarrow> (if x <= y then [y - x = 0] (mod n) else [x - y = 0] (mod n))"
-apply (auto simp add: nat_mod)
-apply (rule_tac x="q2" in exI)
-apply (rule_tac x="q1" in exI, simp)
-apply (rule_tac x="q2" in exI)
-apply (rule_tac x="q1" in exI, simp)
-apply (rule_tac x="q1" in exI)
-apply (rule_tac x="q2" in exI, simp)
-apply (rule_tac x="q1" in exI)
-apply (rule_tac x="q2" in exI, simp)
-done
-
-lemma cong_mult_lcancel: assumes an: "coprime a n" and axy:"[a * x = a * y] (mod n)"
-  shows "[x = y] (mod n)"
-proof-
-  {assume "a = 0" with an axy coprime_0'[of n] have ?thesis by (simp add: modeq_def) }
-  moreover
-  {assume az: "a\<noteq>0"
-    {assume xy: "x \<le> y" hence axy': "a*x \<le> a*y" by simp
-      with axy cong_sub_cases[of "a*x" "a*y" n]  have "[a*(y - x) = 0] (mod n)"
-	by (simp only: if_True diff_mult_distrib2)
-      hence th: "n dvd a*(y -x)" by simp
-      from coprime_divprod[OF th] an have "n dvd y - x"
-	by (simp add: coprime_commute)
-      hence ?thesis using xy cong_sub_cases[of x y n] by simp}
-    moreover
-    {assume H: "\<not>x \<le> y" hence xy: "y \<le> x"  by arith
-      from H az have axy': "\<not> a*x \<le> a*y" by auto
-      with axy H cong_sub_cases[of "a*x" "a*y" n]  have "[a*(x - y) = 0] (mod n)"
-	by (simp only: if_False diff_mult_distrib2)
-      hence th: "n dvd a*(x - y)" by simp
-      from coprime_divprod[OF th] an have "n dvd x - y"
-	by (simp add: coprime_commute)
-      hence ?thesis using xy cong_sub_cases[of x y n] by simp}
-    ultimately have ?thesis by blast}
-  ultimately show ?thesis by blast
-qed
-
-lemma cong_mult_rcancel: assumes an: "coprime a n" and axy:"[x*a = y*a] (mod n)"
-  shows "[x = y] (mod n)"
-  using cong_mult_lcancel[OF an axy[unfolded mult_commute[of _a]]] .
-
-lemma cong_refl: "[x = x] (mod n)" by (simp add: modeq_def)
-
-lemma eq_imp_cong: "a = b \<Longrightarrow> [a = b] (mod n)" by (simp add: cong_refl)
-
-lemma cong_commute: "[x = y] (mod n) \<longleftrightarrow> [y = x] (mod n)"
-  by (auto simp add: modeq_def)
-
-lemma cong_trans[trans]: "[x = y] (mod n) \<Longrightarrow> [y = z] (mod n) \<Longrightarrow> [x = z] (mod n)"
-  by (simp add: modeq_def)
-
-lemma cong_add: assumes xx': "[x = x'] (mod n)" and yy':"[y = y'] (mod n)"
-  shows "[x + y = x' + y'] (mod n)"
-proof-
-  have "(x + y) mod n = (x mod n + y mod n) mod n"
-    by (simp add: mod_add_left_eq[of x y n] mod_add_right_eq[of "x mod n" y n])
-  also have "\<dots> = (x' mod n + y' mod n) mod n" using xx' yy' modeq_def by simp
-  also have "\<dots> = (x' + y') mod n"
-    by (simp add: mod_add_left_eq[of x' y' n] mod_add_right_eq[of "x' mod n" y' n])
-  finally show ?thesis unfolding modeq_def .
-qed
-
-lemma cong_mult: assumes xx': "[x = x'] (mod n)" and yy':"[y = y'] (mod n)"
-  shows "[x * y = x' * y'] (mod n)"
-proof-
-  have "(x * y) mod n = (x mod n) * (y mod n) mod n"
-    by (simp add: mod_mult_left_eq[of x y n] mod_mult_right_eq[of "x mod n" y n])
-  also have "\<dots> = (x' mod n) * (y' mod n) mod n" using xx'[unfolded modeq_def] yy'[unfolded modeq_def] by simp
-  also have "\<dots> = (x' * y') mod n"
-    by (simp add: mod_mult_left_eq[of x' y' n] mod_mult_right_eq[of "x' mod n" y' n])
-  finally show ?thesis unfolding modeq_def .
-qed
-
-lemma cong_exp: "[x = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"
-  by (induct k, auto simp add: cong_refl cong_mult)
-lemma cong_sub: assumes xx': "[x = x'] (mod n)" and yy': "[y = y'] (mod n)"
-  and yx: "y <= x" and yx': "y' <= x'"
-  shows "[x - y = x' - y'] (mod n)"
-proof-
-  { fix x a x' a' y b y' b'
-    have "(x::nat) + a = x' + a' \<Longrightarrow> y + b = y' + b' \<Longrightarrow> y <= x \<Longrightarrow> y' <= x'
-      \<Longrightarrow> (x - y) + (a + b') = (x' - y') + (a' + b)" by arith}
-  note th = this
-  from xx' yy' obtain q1 q2 q1' q2' where q12: "x + n*q1 = x'+n*q2"
-    and q12': "y + n*q1' = y'+n*q2'" unfolding nat_mod by blast+
-  from th[OF q12 q12' yx yx']
-  have "(x - y) + n*(q1 + q2') = (x' - y') + n*(q2 + q1')"
-    by (simp add: right_distrib)
-  thus ?thesis unfolding nat_mod by blast
-qed
-
-lemma cong_mult_lcancel_eq: assumes an: "coprime a n"
-  shows "[a * x = a * y] (mod n) \<longleftrightarrow> [x = y] (mod n)" (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  assume H: "?rhs" from cong_mult[OF cong_refl[of a n] H] show ?lhs .
-next
-  assume H: "?lhs" hence H': "[x*a = y*a] (mod n)" by (simp add: mult_commute)
-  from cong_mult_rcancel[OF an H'] show ?rhs  .
-qed
-
-lemma cong_mult_rcancel_eq: assumes an: "coprime a n"
-  shows "[x * a = y * a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
-using cong_mult_lcancel_eq[OF an, of x y] by (simp add: mult_commute)
-
-lemma cong_add_lcancel_eq: "[a + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)"
-  by (simp add: nat_mod)
-
-lemma cong_add_rcancel_eq: "[x + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
-  by (simp add: nat_mod)
-
-lemma cong_add_rcancel: "[x + a = y + a] (mod n) \<Longrightarrow> [x = y] (mod n)"
-  by (simp add: nat_mod)
-
-lemma cong_add_lcancel: "[a + x = a + y] (mod n) \<Longrightarrow> [x = y] (mod n)"
-  by (simp add: nat_mod)
-
-lemma cong_add_lcancel_eq_0: "[a + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
-  by (simp add: nat_mod)
-
-lemma cong_add_rcancel_eq_0: "[x + a = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
-  by (simp add: nat_mod)
-
-lemma cong_imp_eq: assumes xn: "x < n" and yn: "y < n" and xy: "[x = y] (mod n)"
-  shows "x = y"
-  using xy[unfolded modeq_def mod_less[OF xn] mod_less[OF yn]] .
-
-lemma cong_divides_modulus: "[x = y] (mod m) \<Longrightarrow> n dvd m ==> [x = y] (mod n)"
-  apply (auto simp add: nat_mod dvd_def)
-  apply (rule_tac x="k*q1" in exI)
-  apply (rule_tac x="k*q2" in exI)
-  by simp
-
-lemma cong_0_divides: "[x = 0] (mod n) \<longleftrightarrow> n dvd x" by simp
-
-lemma cong_1_divides:"[x = 1] (mod n) ==> n dvd x - 1"
-  apply (cases "x\<le>1", simp_all)
-  using cong_sub_cases[of x 1 n] by auto
-
-lemma cong_divides: "[x = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"
-apply (auto simp add: nat_mod dvd_def)
-apply (rule_tac x="k + q1 - q2" in exI, simp add: add_mult_distrib2 diff_mult_distrib2)
-apply (rule_tac x="k + q2 - q1" in exI, simp add: add_mult_distrib2 diff_mult_distrib2)
-done
-
-lemma cong_coprime: assumes xy: "[x = y] (mod n)"
-  shows "coprime n x \<longleftrightarrow> coprime n y"
-proof-
-  {assume "n=0" hence ?thesis using xy by simp}
-  moreover
-  {assume nz: "n \<noteq> 0"
-  have "coprime n x \<longleftrightarrow> coprime (x mod n) n"
-    by (simp add: coprime_mod[OF nz, of x] coprime_commute[of n x])
-  also have "\<dots> \<longleftrightarrow> coprime (y mod n) n" using xy[unfolded modeq_def] by simp
-  also have "\<dots> \<longleftrightarrow> coprime y n" by (simp add: coprime_mod[OF nz, of y])
-  finally have ?thesis by (simp add: coprime_commute) }
-ultimately show ?thesis by blast
-qed
-
-lemma cong_mod: "~(n = 0) \<Longrightarrow> [a mod n = a] (mod n)" by (simp add: modeq_def)
-
-lemma mod_mult_cong: "~(a = 0) \<Longrightarrow> ~(b = 0)
-  \<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)"
-  by (simp add: modeq_def mod_mult2_eq mod_add_left_eq)
-
-lemma cong_mod_mult: "[x = y] (mod n) \<Longrightarrow> m dvd n \<Longrightarrow> [x = y] (mod m)"
-  apply (auto simp add: nat_mod dvd_def)
-  apply (rule_tac x="k*q1" in exI)
-  apply (rule_tac x="k*q2" in exI, simp)
-  done
-
-(* Some things when we know more about the order.                            *)
-
-lemma cong_le: "y <= x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)"
-  using nat_mod_lemma[of x y n]
-  apply auto
-  apply (simp add: nat_mod)
-  apply (rule_tac x="q" in exI)
-  apply (rule_tac x="q + q" in exI)
-  by (auto simp: algebra_simps)
-
-lemma cong_to_1: "[a = 1] (mod n) \<longleftrightarrow> a = 0 \<and> n = 1 \<or> (\<exists>m. a = 1 + m * n)"
-proof-
-  {assume "n = 0 \<or> n = 1\<or> a = 0 \<or> a = 1" hence ?thesis
-      apply (cases "n=0", simp_all add: cong_commute)
-      apply (cases "n=1", simp_all add: cong_commute modeq_def)
-      apply arith
-      by (cases "a=1", simp_all add: modeq_def cong_commute)}
-  moreover
-  {assume n: "n\<noteq>0" "n\<noteq>1" and a:"a\<noteq>0" "a \<noteq> 1" hence a': "a \<ge> 1" by simp
-    hence ?thesis using cong_le[OF a', of n] by auto }
-  ultimately show ?thesis by auto
-qed
-
-(* Some basic theorems about solving congruences.                            *)
-
-
-lemma cong_solve: assumes an: "coprime a n" shows "\<exists>x. [a * x = b] (mod n)"
-proof-
-  {assume "a=0" hence ?thesis using an by (simp add: modeq_def)}
-  moreover
-  {assume az: "a\<noteq>0"
-  from bezout_add_strong[OF az, of n]
-  obtain d x y where dxy: "d dvd a" "d dvd n" "a*x = n*y + d" by blast
-  from an[unfolded coprime, rule_format, of d] dxy(1,2) have d1: "d = 1" by blast
-  hence "a*x*b = (n*y + 1)*b" using dxy(3) by simp
-  hence "a*(x*b) = n*(y*b) + b" by algebra
-  hence "a*(x*b) mod n = (n*(y*b) + b) mod n" by simp
-  hence "a*(x*b) mod n = b mod n" by (simp add: mod_add_left_eq)
-  hence "[a*(x*b) = b] (mod n)" unfolding modeq_def .
-  hence ?thesis by blast}
-ultimately  show ?thesis by blast
-qed
-
-lemma cong_solve_unique: assumes an: "coprime a n" and nz: "n \<noteq> 0"
-  shows "\<exists>!x. x < n \<and> [a * x = b] (mod n)"
-proof-
-  let ?P = "\<lambda>x. x < n \<and> [a * x = b] (mod n)"
-  from cong_solve[OF an] obtain x where x: "[a*x = b] (mod n)" by blast
-  let ?x = "x mod n"
-  from x have th: "[a * ?x = b] (mod n)"
-    by (simp add: modeq_def mod_mult_right_eq[of a x n])
-  from mod_less_divisor[ of n x] nz th have Px: "?P ?x" by simp
-  {fix y assume Py: "y < n" "[a * y = b] (mod n)"
-    from Py(2) th have "[a * y = a*?x] (mod n)" by (simp add: modeq_def)
-    hence "[y = ?x] (mod n)" by (simp add: cong_mult_lcancel_eq[OF an])
-    with mod_less[OF Py(1)] mod_less_divisor[ of n x] nz
-    have "y = ?x" by (simp add: modeq_def)}
-  with Px show ?thesis by blast
-qed
-
-lemma cong_solve_unique_nontrivial:
-  assumes p: "prime p" and pa: "coprime p a" and x0: "0 < x" and xp: "x < p"
-  shows "\<exists>!y. 0 < y \<and> y < p \<and> [x * y = a] (mod p)"
-proof-
-  from p have p1: "p > 1" using prime_ge_2[OF p] by arith
-  hence p01: "p \<noteq> 0" "p \<noteq> 1" by arith+
-  from pa have ap: "coprime a p" by (simp add: coprime_commute)
-  from prime_coprime[OF p, of x] dvd_imp_le[of p x] x0 xp have px:"coprime x p"
-    by (auto simp add: coprime_commute)
-  from cong_solve_unique[OF px p01(1)]
-  obtain y where y: "y < p" "[x * y = a] (mod p)" "\<forall>z. z < p \<and> [x * z = a] (mod p) \<longrightarrow> z = y" by blast
-  {assume y0: "y = 0"
-    with y(2) have th: "p dvd a" by (simp add: cong_commute[of 0 a p])
-    with p coprime_prime[OF pa, of p] have False by simp}
-  with y show ?thesis unfolding Ex1_def using neq0_conv by blast
-qed
-lemma cong_unique_inverse_prime:
-  assumes p: "prime p" and x0: "0 < x" and xp: "x < p"
-  shows "\<exists>!y. 0 < y \<and> y < p \<and> [x * y = 1] (mod p)"
-  using cong_solve_unique_nontrivial[OF p coprime_1[of p] x0 xp] .
-
-(* Forms of the Chinese remainder theorem.                                   *)
-
-lemma cong_chinese:
-  assumes ab: "coprime a b" and  xya: "[x = y] (mod a)"
-  and xyb: "[x = y] (mod b)"
-  shows "[x = y] (mod a*b)"
-  using ab xya xyb
-  by (simp add: cong_sub_cases[of x y a] cong_sub_cases[of x y b]
-    cong_sub_cases[of x y "a*b"])
-(cases "x \<le> y", simp_all add: divides_mul[of a _ b])
-
-lemma chinese_remainder_unique:
-  assumes ab: "coprime a b" and az: "a \<noteq> 0" and bz: "b\<noteq>0"
-  shows "\<exists>!x. x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"
-proof-
-  from az bz have abpos: "a*b > 0" by simp
-  from chinese_remainder[OF ab az bz] obtain x q1 q2 where
-    xq12: "x = m + q1 * a" "x = n + q2 * b" by blast
-  let ?w = "x mod (a*b)"
-  have wab: "?w < a*b" by (simp add: mod_less_divisor[OF abpos])
-  from xq12(1) have "?w mod a = ((m + q1 * a) mod (a*b)) mod a" by simp
-  also have "\<dots> = m mod a" apply (simp add: mod_mult2_eq)
-    apply (subst mod_add_left_eq)
-    by simp
-  finally have th1: "[?w = m] (mod a)" by (simp add: modeq_def)
-  from xq12(2) have "?w mod b = ((n + q2 * b) mod (a*b)) mod b" by simp
-  also have "\<dots> = ((n + q2 * b) mod (b*a)) mod b" by (simp add: mult_commute)
-  also have "\<dots> = n mod b" apply (simp add: mod_mult2_eq)
-    apply (subst mod_add_left_eq)
-    by simp
-  finally have th2: "[?w = n] (mod b)" by (simp add: modeq_def)
-  {fix y assume H: "y < a*b" "[y = m] (mod a)" "[y = n] (mod b)"
-    with th1 th2 have H': "[y = ?w] (mod a)" "[y = ?w] (mod b)"
-      by (simp_all add: modeq_def)
-    from cong_chinese[OF ab H'] mod_less[OF H(1)] mod_less[OF wab]
-    have "y = ?w" by (simp add: modeq_def)}
-  with th1 th2 wab show ?thesis by blast
-qed
-
-lemma chinese_remainder_coprime_unique:
-  assumes ab: "coprime a b" and az: "a \<noteq> 0" and bz: "b \<noteq> 0"
-  and ma: "coprime m a" and nb: "coprime n b"
-  shows "\<exists>!x. coprime x (a * b) \<and> x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"
-proof-
-  let ?P = "\<lambda>x. x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"
-  from chinese_remainder_unique[OF ab az bz]
-  obtain x where x: "x < a * b" "[x = m] (mod a)" "[x = n] (mod b)"
-    "\<forall>y. ?P y \<longrightarrow> y = x" by blast
-  from ma nb cong_coprime[OF x(2)] cong_coprime[OF x(3)]
-  have "coprime x a" "coprime x b" by (simp_all add: coprime_commute)
-  with coprime_mul[of x a b] have "coprime x (a*b)" by simp
-  with x show ?thesis by blast
-qed
-
-(* Euler totient function.                                                   *)
-
-definition phi_def: "\<phi> n = card { m. 0 < m \<and> m <= n \<and> coprime m n }"
-
-lemma phi_0[simp]: "\<phi> 0 = 0"
-  unfolding phi_def by auto
-
-lemma phi_finite[simp]: "finite ({ m. 0 < m \<and> m <= n \<and> coprime m n })"
-proof-
-  have "{ m. 0 < m \<and> m <= n \<and> coprime m n } \<subseteq> {0..n}" by auto
-  thus ?thesis by (auto intro: finite_subset)
-qed
-
-declare coprime_1[presburger]
-lemma phi_1[simp]: "\<phi> 1 = 1"
-proof-
-  {fix m
-    have "0 < m \<and> m <= 1 \<and> coprime m 1 \<longleftrightarrow> m = 1" by presburger }
-  thus ?thesis by (simp add: phi_def)
-qed
-
-lemma [simp]: "\<phi> (Suc 0) = Suc 0" using phi_1 by simp
-
-lemma phi_alt: "\<phi>(n) = card { m. coprime m n \<and> m < n}"
-proof-
-  {assume "n=0 \<or> n=1" hence ?thesis by (cases "n=0", simp_all)}
-  moreover
-  {assume n: "n\<noteq>0" "n\<noteq>1"
-    {fix m
-      from n have "0 < m \<and> m <= n \<and> coprime m n \<longleftrightarrow> coprime m n \<and> m < n"
-	apply (cases "m = 0", simp_all)
-	apply (cases "m = 1", simp_all)
-	apply (cases "m = n", auto)
-	done }
-    hence ?thesis unfolding phi_def by simp}
-  ultimately show ?thesis by auto
-qed
-
-lemma phi_finite_lemma[simp]: "finite {m. coprime m n \<and>  m < n}" (is "finite ?S")
-  by (rule finite_subset[of "?S" "{0..n}"], auto)
-
-lemma phi_another: assumes n: "n\<noteq>1"
-  shows "\<phi> n = card {m. 0 < m \<and> m < n \<and> coprime m n }"
-proof-
-  {fix m
-    from n have "0 < m \<and> m < n \<and> coprime m n \<longleftrightarrow> coprime m n \<and> m < n"
-      by (cases "m=0", auto)}
-  thus ?thesis unfolding phi_alt by auto
-qed
-
-lemma phi_limit: "\<phi> n \<le> n"
-proof-
-  have "{ m. coprime m n \<and> m < n} \<subseteq> {0 ..<n}" by auto
-  with card_mono[of "{0 ..<n}" "{ m. coprime m n \<and> m < n}"]
-  show ?thesis unfolding phi_alt by auto
-qed
-
-lemma stupid[simp]: "{m. (0::nat) < m \<and> m < n} = {1..<n}"
-  by auto
-
-lemma phi_limit_strong: assumes n: "n\<noteq>1"
-  shows "\<phi>(n) \<le> n - 1"
-proof-
-  show ?thesis
-    unfolding phi_another[OF n] finite_number_segment[of n, symmetric]
-    by (rule card_mono[of "{m. 0 < m \<and> m < n}" "{m. 0 < m \<and> m < n \<and> coprime m n}"], auto)
-qed
-
-lemma phi_lowerbound_1_strong: assumes n: "n \<ge> 1"
-  shows "\<phi>(n) \<ge> 1"
-proof-
-  let ?S = "{ m. 0 < m \<and> m <= n \<and> coprime m n }"
-  from card_0_eq[of ?S] n have "\<phi> n \<noteq> 0" unfolding phi_alt
-    apply auto
-    apply (cases "n=1", simp_all)
-    apply (rule exI[where x=1], simp)
-    done
-  thus ?thesis by arith
-qed
-
-lemma phi_lowerbound_1: "2 <= n ==> 1 <= \<phi>(n)"
-  using phi_lowerbound_1_strong[of n] by auto
-
-lemma phi_lowerbound_2: assumes n: "3 <= n" shows "2 <= \<phi> (n)"
-proof-
-  let ?S = "{ m. 0 < m \<and> m <= n \<and> coprime m n }"
-  have inS: "{1, n - 1} \<subseteq> ?S" using n coprime_plus1[of "n - 1"]
-    by (auto simp add: coprime_commute)
-  from n have c2: "card {1, n - 1} = 2" by (auto simp add: card_insert_if)
-  from card_mono[of ?S "{1, n - 1}", simplified inS c2] show ?thesis
-    unfolding phi_def by auto
-qed
-
-lemma phi_prime: "\<phi> n = n - 1 \<and> n\<noteq>0 \<and> n\<noteq>1 \<longleftrightarrow> prime n"
-proof-
-  {assume "n=0 \<or> n=1" hence ?thesis by (cases "n=1", simp_all)}
-  moreover
-  {assume n: "n\<noteq>0" "n\<noteq>1"
-    let ?S = "{m. 0 < m \<and> m < n}"
-    have fS: "finite ?S" by simp
-    let ?S' = "{m. 0 < m \<and> m < n \<and> coprime m n}"
-    have fS':"finite ?S'" apply (rule finite_subset[of ?S' ?S]) by auto
-    {assume H: "\<phi> n = n - 1 \<and> n\<noteq>0 \<and> n\<noteq>1"
-      hence ceq: "card ?S' = card ?S"
-      using n finite_number_segment[of n] phi_another[OF n(2)] by simp
-      {fix m assume m: "0 < m" "m < n" "\<not> coprime m n"
-	hence mS': "m \<notin> ?S'" by auto
-	have "insert m ?S' \<le> ?S" using m by auto
-	from m have "card (insert m ?S') \<le> card ?S"
-	  by - (rule card_mono[of ?S "insert m ?S'"], auto)
-	hence False
-	  unfolding card_insert_disjoint[of "?S'" m, OF fS' mS'] ceq
-	  by simp }
-      hence "\<forall>m. 0 <m \<and> m < n \<longrightarrow> coprime m n" by blast
-      hence "prime n" unfolding prime using n by (simp add: coprime_commute)}
-    moreover
-    {assume H: "prime n"
-      hence "?S = ?S'" unfolding prime using n
-	by (auto simp add: coprime_commute)
-      hence "card ?S = card ?S'" by simp
-      hence "\<phi> n = n - 1" unfolding phi_another[OF n(2)] by simp}
-    ultimately have ?thesis using n by blast}
-  ultimately show ?thesis by (cases "n=0") blast+
-qed
-
-(* Multiplicativity property.                                                *)
-
-lemma phi_multiplicative: assumes ab: "coprime a b"
-  shows "\<phi> (a * b) = \<phi> a * \<phi> b"
-proof-
-  {assume "a = 0 \<or> b = 0 \<or> a = 1 \<or> b = 1"
-    hence ?thesis
-      by (cases "a=0", simp, cases "b=0", simp, cases"a=1", simp_all) }
-  moreover
-  {assume a: "a\<noteq>0" "a\<noteq>1" and b: "b\<noteq>0" "b\<noteq>1"
-    hence ab0: "a*b \<noteq> 0" by simp
-    let ?S = "\<lambda>k. {m. coprime m k \<and> m < k}"
-    let ?f = "\<lambda>x. (x mod a, x mod b)"
-    have eq: "?f ` (?S (a*b)) = (?S a \<times> ?S b)"
-    proof-
-      {fix x assume x:"x \<in> ?S (a*b)"
-	hence x': "coprime x (a*b)" "x < a*b" by simp_all
-	hence xab: "coprime x a" "coprime x b" by (simp_all add: coprime_mul_eq)
-	from mod_less_divisor a b have xab':"x mod a < a" "x mod b < b" by auto
-	from xab xab' have "?f x \<in> (?S a \<times> ?S b)"
-	  by (simp add: coprime_mod[OF a(1)] coprime_mod[OF b(1)])}
-      moreover
-      {fix x y assume x: "x \<in> ?S a" and y: "y \<in> ?S b"
-	hence x': "coprime x a" "x < a" and y': "coprime y b" "y < b" by simp_all
-	from chinese_remainder_coprime_unique[OF ab a(1) b(1) x'(1) y'(1)]
-	obtain z where z: "coprime z (a * b)" "z < a * b" "[z = x] (mod a)"
-	  "[z = y] (mod b)" by blast
-	hence "(x,y) \<in> ?f ` (?S (a*b))"
-	  using y'(2) mod_less_divisor[of b y] x'(2) mod_less_divisor[of a x]
-	  by (auto simp add: image_iff modeq_def)}
-      ultimately show ?thesis by auto
-    qed
-    have finj: "inj_on ?f (?S (a*b))"
-      unfolding inj_on_def
-    proof(clarify)
-      fix x y assume H: "coprime x (a * b)" "x < a * b" "coprime y (a * b)"
-	"y < a * b" "x mod a = y mod a" "x mod b = y mod b"
-      hence cp: "coprime x a" "coprime x b" "coprime y a" "coprime y b"
-	by (simp_all add: coprime_mul_eq)
-      from chinese_remainder_coprime_unique[OF ab a(1) b(1) cp(3,4)] H
-      show "x = y" unfolding modeq_def by blast
-    qed
-    from card_image[OF finj, unfolded eq] have ?thesis
-      unfolding phi_alt by simp }
-  ultimately show ?thesis by auto
-qed
-
-(* Fermat's Little theorem / Fermat-Euler theorem.                           *)
-
-
-lemma nproduct_mod:
-  assumes fS: "finite S" and n0: "n \<noteq> 0"
-  shows "[setprod (\<lambda>m. a(m) mod n) S = setprod a S] (mod n)"
-proof-
-  have th1:"[1 = 1] (mod n)" by (simp add: modeq_def)
-  from cong_mult
-  have th3:"\<forall>x1 y1 x2 y2.
-    [x1 = x2] (mod n) \<and> [y1 = y2] (mod n) \<longrightarrow> [x1 * y1 = x2 * y2] (mod n)"
-    by blast
-  have th4:"\<forall>x\<in>S. [a x mod n = a x] (mod n)" by (simp add: modeq_def)
-  from fold_image_related[where h="(\<lambda>m. a(m) mod n)" and g=a, OF th1 th3 fS, OF th4] show ?thesis unfolding setprod_def by (simp add: fS)
-qed
-
-lemma nproduct_cmul:
-  assumes fS:"finite S"
-  shows "setprod (\<lambda>m. (c::'a::{comm_monoid_mult})* a(m)) S = c ^ (card S) * setprod a S"
-unfolding setprod_timesf setprod_constant[OF fS, of c] ..
-
-lemma coprime_nproduct:
-  assumes fS: "finite S" and Sn: "\<forall>x\<in>S. coprime n (a x)"
-  shows "coprime n (setprod a S)"
-  using fS unfolding setprod_def by (rule finite_subset_induct)
-    (insert Sn, auto simp add: coprime_mul)
-
-lemma fermat_little: assumes an: "coprime a n"
-  shows "[a ^ (\<phi> n) = 1] (mod n)"
-proof-
-  {assume "n=0" hence ?thesis by simp}
-  moreover
-  {assume "n=1" hence ?thesis by (simp add: modeq_def)}
-  moreover
-  {assume nz: "n \<noteq> 0" and n1: "n \<noteq> 1"
-    let ?S = "{m. coprime m n \<and> m < n}"
-    let ?P = "\<Prod> ?S"
-    have fS: "finite ?S" by simp
-    have cardfS: "\<phi> n = card ?S" unfolding phi_alt ..
-    {fix m assume m: "m \<in> ?S"
-      hence "coprime m n" by simp
-      with coprime_mul[of n a m] an have "coprime (a*m) n"
-	by (simp add: coprime_commute)}
-    hence Sn: "\<forall>m\<in> ?S. coprime (a*m) n " by blast
-    from coprime_nproduct[OF fS, of n "\<lambda>m. m"] have nP:"coprime ?P n"
-      by (simp add: coprime_commute)
-    have Paphi: "[?P*a^ (\<phi> n) = ?P*1] (mod n)"
-    proof-
-      let ?h = "\<lambda>m. m mod n"
-      {fix m assume mS: "m\<in> ?S"
-	hence "?h m \<in> ?S" by simp}
-      hence hS: "?h ` ?S = ?S"by (auto simp add: image_iff)
-      have "a\<noteq>0" using an n1 nz apply- apply (rule ccontr) by simp
-      hence inj: "inj_on (op * a) ?S" unfolding inj_on_def by simp
-
-      have eq0: "fold_image op * (?h \<circ> op * a) 1 {m. coprime m n \<and> m < n} =
-     fold_image op * (\<lambda>m. m) 1 {m. coprime m n \<and> m < n}"
-      proof (rule fold_image_eq_general[where h="?h o (op * a)"])
-	show "finite ?S" using fS .
-      next
-	{fix y assume yS: "y \<in> ?S" hence y: "coprime y n" "y < n" by simp_all
-	  from cong_solve_unique[OF an nz, of y]
-	  obtain x where x:"x < n" "[a * x = y] (mod n)" "\<forall>z. z < n \<and> [a * z = y] (mod n) \<longrightarrow> z=x" by blast
-	  from cong_coprime[OF x(2)] y(1)
-	  have xm: "coprime x n" by (simp add: coprime_mul_eq coprime_commute)
-	  {fix z assume "z \<in> ?S" "(?h \<circ> op * a) z = y"
-	    hence z: "coprime z n" "z < n" "(?h \<circ> op * a) z = y" by simp_all
-	    from x(3)[rule_format, of z] z(2,3) have "z=x"
-	      unfolding modeq_def mod_less[OF y(2)] by simp}
-	  with xm x(1,2) have "\<exists>!x. x \<in> ?S \<and> (?h \<circ> op * a) x = y"
-	    unfolding modeq_def mod_less[OF y(2)] by auto }
-	thus "\<forall>y\<in>{m. coprime m n \<and> m < n}.
-       \<exists>!x. x \<in> {m. coprime m n \<and> m < n} \<and> ((\<lambda>m. m mod n) \<circ> op * a) x = y" by blast
-      next
-	{fix x assume xS: "x\<in> ?S"
-	  hence x: "coprime x n" "x < n" by simp_all
-	  with an have "coprime (a*x) n"
-	    by (simp add: coprime_mul_eq[of n a x] coprime_commute)
-	  hence "?h (a*x) \<in> ?S" using nz
-	    by (simp add: coprime_mod[OF nz] mod_less_divisor)}
-	thus " \<forall>x\<in>{m. coprime m n \<and> m < n}.
-       ((\<lambda>m. m mod n) \<circ> op * a) x \<in> {m. coprime m n \<and> m < n} \<and>
-       ((\<lambda>m. m mod n) \<circ> op * a) x = ((\<lambda>m. m mod n) \<circ> op * a) x" by simp
-      qed
-      from nproduct_mod[OF fS nz, of "op * a"]
-      have "[(setprod (op *a) ?S) = (setprod (?h o (op * a)) ?S)] (mod n)"
-	unfolding o_def
-	by (simp add: cong_commute)
-      also have "[setprod (?h o (op * a)) ?S = ?P ] (mod n)"
-	using eq0 fS an by (simp add: setprod_def modeq_def o_def)
-      finally show "[?P*a^ (\<phi> n) = ?P*1] (mod n)"
-	unfolding cardfS mult_commute[of ?P "a^ (card ?S)"]
-	  nproduct_cmul[OF fS, symmetric] mult_1_right by simp
-    qed
-    from cong_mult_lcancel[OF nP Paphi] have ?thesis . }
-  ultimately show ?thesis by blast
-qed
-
-lemma fermat_little_prime: assumes p: "prime p" and ap: "coprime a p"
-  shows "[a^ (p - 1) = 1] (mod p)"
-  using fermat_little[OF ap] p[unfolded phi_prime[symmetric]]
-by simp
-
-
-(* Lucas's theorem.                                                          *)
-
-lemma lucas_coprime_lemma:
-  assumes m: "m\<noteq>0" and am: "[a^m = 1] (mod n)"
-  shows "coprime a n"
-proof-
-  {assume "n=1" hence ?thesis by simp}
-  moreover
-  {assume "n = 0" hence ?thesis using am m exp_eq_1[of a m] by simp}
-  moreover
-  {assume n: "n\<noteq>0" "n\<noteq>1"
-    from m obtain m' where m': "m = Suc m'" by (cases m, blast+)
-    {fix d
-      assume d: "d dvd a" "d dvd n"
-      from n have n1: "1 < n" by arith
-      from am mod_less[OF n1] have am1: "a^m mod n = 1" unfolding modeq_def by simp
-      from dvd_mult2[OF d(1), of "a^m'"] have dam:"d dvd a^m" by (simp add: m')
-      from dvd_mod_iff[OF d(2), of "a^m"] dam am1
-      have "d = 1" by simp }
-    hence ?thesis unfolding coprime by auto
-  }
-  ultimately show ?thesis by blast
-qed
-
-lemma lucas_weak:
-  assumes n: "n \<ge> 2" and an:"[a^(n - 1) = 1] (mod n)"
-  and nm: "\<forall>m. 0 <m \<and> m < n - 1 \<longrightarrow> \<not> [a^m = 1] (mod n)"
-  shows "prime n"
-proof-
-  from n have n1: "n \<noteq> 1" "n\<noteq>0" "n - 1 \<noteq> 0" "n - 1 > 0" "n - 1 < n" by arith+
-  from lucas_coprime_lemma[OF n1(3) an] have can: "coprime a n" .
-  from fermat_little[OF can] have afn: "[a ^ \<phi> n = 1] (mod n)" .
-  {assume "\<phi> n \<noteq> n - 1"
-    with phi_limit_strong[OF n1(1)] phi_lowerbound_1[OF n]
-    have c:"\<phi> n > 0 \<and> \<phi> n < n - 1" by arith
-    from nm[rule_format, OF c] afn have False ..}
-  hence "\<phi> n = n - 1" by blast
-  with phi_prime[of n] n1(1,2) show ?thesis by simp
-qed
-
-lemma nat_exists_least_iff: "(\<exists>(n::nat). P n) \<longleftrightarrow> (\<exists>n. P n \<and> (\<forall>m < n. \<not> P m))"
-  (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  assume ?rhs thus ?lhs by blast
-next
-  assume H: ?lhs then obtain n where n: "P n" by blast
-  let ?x = "Least P"
-  {fix m assume m: "m < ?x"
-    from not_less_Least[OF m] have "\<not> P m" .}
-  with LeastI_ex[OF H] show ?rhs by blast
-qed
-
-lemma nat_exists_least_iff': "(\<exists>(n::nat). P n) \<longleftrightarrow> (P (Least P) \<and> (\<forall>m < (Least P). \<not> P m))"
-  (is "?lhs \<longleftrightarrow> ?rhs")
-proof-
-  {assume ?rhs hence ?lhs by blast}
-  moreover
-  { assume H: ?lhs then obtain n where n: "P n" by blast
-    let ?x = "Least P"
-    {fix m assume m: "m < ?x"
-      from not_less_Least[OF m] have "\<not> P m" .}
-    with LeastI_ex[OF H] have ?rhs by blast}
-  ultimately show ?thesis by blast
-qed
-
-lemma power_mod: "((x::nat) mod m)^n mod m = x^n mod m"
-proof(induct n)
-  case 0 thus ?case by simp
-next
-  case (Suc n)
-  have "(x mod m)^(Suc n) mod m = ((x mod m) * (((x mod m) ^ n) mod m)) mod m"
-    by (simp add: mod_mult_right_eq[symmetric])
-  also have "\<dots> = ((x mod m) * (x^n mod m)) mod m" using Suc.hyps by simp
-  also have "\<dots> = x^(Suc n) mod m"
-    by (simp add: mod_mult_left_eq[symmetric] mod_mult_right_eq[symmetric])
-  finally show ?case .
-qed
-
-lemma lucas:
-  assumes n2: "n \<ge> 2" and an1: "[a^(n - 1) = 1] (mod n)"
-  and pn: "\<forall>p. prime p \<and> p dvd n - 1 \<longrightarrow> \<not> [a^((n - 1) div p) = 1] (mod n)"
-  shows "prime n"
-proof-
-  from n2 have n01: "n\<noteq>0" "n\<noteq>1" "n - 1 \<noteq> 0" by arith+
-  from mod_less_divisor[of n 1] n01 have onen: "1 mod n = 1" by simp
-  from lucas_coprime_lemma[OF n01(3) an1] cong_coprime[OF an1]
-  have an: "coprime a n" "coprime (a^(n - 1)) n" by (simp_all add: coprime_commute)
-  {assume H0: "\<exists>m. 0 < m \<and> m < n - 1 \<and> [a ^ m = 1] (mod n)" (is "EX m. ?P m")
-    from H0[unfolded nat_exists_least_iff[of ?P]] obtain m where
-      m: "0 < m" "m < n - 1" "[a ^ m = 1] (mod n)" "\<forall>k <m. \<not>?P k" by blast
-    {assume nm1: "(n - 1) mod m > 0"
-      from mod_less_divisor[OF m(1)] have th0:"(n - 1) mod m < m" by blast
-      let ?y = "a^ ((n - 1) div m * m)"
-      note mdeq = mod_div_equality[of "(n - 1)" m]
-      from coprime_exp[OF an(1)[unfolded coprime_commute[of a n]],
-	of "(n - 1) div m * m"]
-      have yn: "coprime ?y n" by (simp add: coprime_commute)
-      have "?y mod n = (a^m)^((n - 1) div m) mod n"
-	by (simp add: algebra_simps power_mult)
-      also have "\<dots> = (a^m mod n)^((n - 1) div m) mod n"
-	using power_mod[of "a^m" n "(n - 1) div m"] by simp
-      also have "\<dots> = 1" using m(3)[unfolded modeq_def onen] onen
-	by (simp add: power_Suc0)
-      finally have th3: "?y mod n = 1"  .
-      have th2: "[?y * a ^ ((n - 1) mod m) = ?y* 1] (mod n)"
-	using an1[unfolded modeq_def onen] onen
-	  mod_div_equality[of "(n - 1)" m, symmetric]
-	by (simp add:power_add[symmetric] modeq_def th3 del: One_nat_def)
-      from cong_mult_lcancel[of ?y n "a^((n - 1) mod m)" 1, OF yn th2]
-      have th1: "[a ^ ((n - 1) mod m) = 1] (mod n)"  .
-      from m(4)[rule_format, OF th0] nm1
-	less_trans[OF mod_less_divisor[OF m(1), of "n - 1"] m(2)] th1
-      have False by blast }
-    hence "(n - 1) mod m = 0" by auto
-    then have mn: "m dvd n - 1" by presburger
-    then obtain r where r: "n - 1 = m*r" unfolding dvd_def by blast
-    from n01 r m(2) have r01: "r\<noteq>0" "r\<noteq>1" by - (rule ccontr, simp)+
-    from prime_factor[OF r01(2)] obtain p where p: "prime p" "p dvd r" by blast
-    hence th: "prime p \<and> p dvd n - 1" unfolding r by (auto intro: dvd_mult)
-    have "(a ^ ((n - 1) div p)) mod n = (a^(m*r div p)) mod n" using r
-      by (simp add: power_mult)
-    also have "\<dots> = (a^(m*(r div p))) mod n" using div_mult1_eq[of m r p] p(2)[unfolded dvd_eq_mod_eq_0] by simp
-    also have "\<dots> = ((a^m)^(r div p)) mod n" by (simp add: power_mult)
-    also have "\<dots> = ((a^m mod n)^(r div p)) mod n" using power_mod[of "a^m" "n" "r div p" ] ..
-    also have "\<dots> = 1" using m(3) onen by (simp add: modeq_def power_Suc0)
-    finally have "[(a ^ ((n - 1) div p))= 1] (mod n)"
-      using onen by (simp add: modeq_def)
-    with pn[rule_format, OF th] have False by blast}
-  hence th: "\<forall>m. 0 < m \<and> m < n - 1 \<longrightarrow> \<not> [a ^ m = 1] (mod n)" by blast
-  from lucas_weak[OF n2 an1 th] show ?thesis .
-qed
-
-(* Definition of the order of a number mod n (0 in non-coprime case).        *)
-
-definition "ord n a = (if coprime n a then Least (\<lambda>d. d > 0 \<and> [a ^d = 1] (mod n)) else 0)"
-
-(* This has the expected properties.                                         *)
-
-lemma coprime_ord:
-  assumes na: "coprime n a"
-  shows "ord n a > 0 \<and> [a ^(ord n a) = 1] (mod n) \<and> (\<forall>m. 0 < m \<and> m < ord n a \<longrightarrow> \<not> [a^ m = 1] (mod n))"
-proof-
-  let ?P = "\<lambda>d. 0 < d \<and> [a ^ d = 1] (mod n)"
-  from euclid[of a] obtain p where p: "prime p" "a < p" by blast
-  from na have o: "ord n a = Least ?P" by (simp add: ord_def)
-  {assume "n=0 \<or> n=1" with na have "\<exists>m>0. ?P m" apply auto apply (rule exI[where x=1]) by (simp  add: modeq_def)}
-  moreover
-  {assume "n\<noteq>0 \<and> n\<noteq>1" hence n2:"n \<ge> 2" by arith
-    from na have na': "coprime a n" by (simp add: coprime_commute)
-    from phi_lowerbound_1[OF n2] fermat_little[OF na']
-    have ex: "\<exists>m>0. ?P m" by - (rule exI[where x="\<phi> n"], auto) }
-  ultimately have ex: "\<exists>m>0. ?P m" by blast
-  from nat_exists_least_iff'[of ?P] ex na show ?thesis
-    unfolding o[symmetric] by auto
-qed
-(* With the special value 0 for non-coprime case, it's more convenient.      *)
-lemma ord_works:
- "[a ^ (ord n a) = 1] (mod n) \<and> (\<forall>m. 0 < m \<and> m < ord n a \<longrightarrow> ~[a^ m = 1] (mod n))"
-apply (cases "coprime n a")
-using coprime_ord[of n a]
-by (blast, simp add: ord_def modeq_def)
-
-lemma ord: "[a^(ord n a) = 1] (mod n)" using ord_works by blast
-lemma ord_minimal: "0 < m \<Longrightarrow> m < ord n a \<Longrightarrow> ~[a^m = 1] (mod n)"
-  using ord_works by blast
-lemma ord_eq_0: "ord n a = 0 \<longleftrightarrow> ~coprime n a"
-by (cases "coprime n a", simp add: neq0_conv coprime_ord, simp add: neq0_conv ord_def)
-
-lemma ord_divides:
- "[a ^ d = 1] (mod n) \<longleftrightarrow> ord n a dvd d" (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  assume rh: ?rhs
-  then obtain k where "d = ord n a * k" unfolding dvd_def by blast
-  hence "[a ^ d = (a ^ (ord n a) mod n)^k] (mod n)"
-    by (simp add : modeq_def power_mult power_mod)
-  also have "[(a ^ (ord n a) mod n)^k = 1] (mod n)"
-    using ord[of a n, unfolded modeq_def]
-    by (simp add: modeq_def power_mod power_Suc0)
-  finally  show ?lhs .
-next
-  assume lh: ?lhs
-  { assume H: "\<not> coprime n a"
-    hence o: "ord n a = 0" by (simp add: ord_def)
-    {assume d: "d=0" with o H have ?rhs by (simp add: modeq_def)}
-    moreover
-    {assume d0: "d\<noteq>0" then obtain d' where d': "d = Suc d'" by (cases d, auto)
-      from H[unfolded coprime]
-      obtain p where p: "p dvd n" "p dvd a" "p \<noteq> 1" by auto
-      from lh[unfolded nat_mod]
-      obtain q1 q2 where q12:"a ^ d + n * q1 = 1 + n * q2" by blast
-      hence "a ^ d + n * q1 - n * q2 = 1" by simp
-      with dvd_diff_nat [OF dvd_add [OF divides_rexp[OF p(2), of d'] dvd_mult2[OF p(1), of q1]] dvd_mult2[OF p(1), of q2]] d' have "p dvd 1" by simp
-      with p(3) have False by simp
-      hence ?rhs ..}
-    ultimately have ?rhs by blast}
-  moreover
-  {assume H: "coprime n a"
-    let ?o = "ord n a"
-    let ?q = "d div ord n a"
-    let ?r = "d mod ord n a"
-    from cong_exp[OF ord[of a n], of ?q]
-    have eqo: "[(a^?o)^?q = 1] (mod n)"  by (simp add: modeq_def power_Suc0)
-    from H have onz: "?o \<noteq> 0" by (simp add: ord_eq_0)
-    hence op: "?o > 0" by simp
-    from mod_div_equality[of d "ord n a"] lh
-    have "[a^(?o*?q + ?r) = 1] (mod n)" by (simp add: modeq_def mult_commute)
-    hence "[(a^?o)^?q * (a^?r) = 1] (mod n)"
-      by (simp add: modeq_def power_mult[symmetric] power_add[symmetric])
-    hence th: "[a^?r = 1] (mod n)"
-      using eqo mod_mult_left_eq[of "(a^?o)^?q" "a^?r" n]
-      apply (simp add: modeq_def del: One_nat_def)
-      by (simp add: mod_mult_left_eq[symmetric])
-    {assume r: "?r = 0" hence ?rhs by (simp add: dvd_eq_mod_eq_0)}
-    moreover
-    {assume r: "?r \<noteq> 0"
-      with mod_less_divisor[OF op, of d] have r0o:"?r >0 \<and> ?r < ?o" by simp
-      from conjunct2[OF ord_works[of a n], rule_format, OF r0o] th
-      have ?rhs by blast}
-    ultimately have ?rhs by blast}
-  ultimately  show ?rhs by blast
-qed
-
-lemma order_divides_phi: "coprime n a \<Longrightarrow> ord n a dvd \<phi> n"
-using ord_divides fermat_little coprime_commute by simp
-lemma order_divides_expdiff:
-  assumes na: "coprime n a"
-  shows "[a^d = a^e] (mod n) \<longleftrightarrow> [d = e] (mod (ord n a))"
-proof-
-  {fix n a d e
-    assume na: "coprime n a" and ed: "(e::nat) \<le> d"
-    hence "\<exists>c. d = e + c" by arith
-    then obtain c where c: "d = e + c" by arith
-    from na have an: "coprime a n" by (simp add: coprime_commute)
-    from coprime_exp[OF na, of e]
-    have aen: "coprime (a^e) n" by (simp add: coprime_commute)
-    from coprime_exp[OF na, of c]
-    have acn: "coprime (a^c) n" by (simp add: coprime_commute)
-    have "[a^d = a^e] (mod n) \<longleftrightarrow> [a^(e + c) = a^(e + 0)] (mod n)"
-      using c by simp
-    also have "\<dots> \<longleftrightarrow> [a^e* a^c = a^e *a^0] (mod n)" by (simp add: power_add)
-    also have  "\<dots> \<longleftrightarrow> [a ^ c = 1] (mod n)"
-      using cong_mult_lcancel_eq[OF aen, of "a^c" "a^0"] by simp
-    also  have "\<dots> \<longleftrightarrow> ord n a dvd c" by (simp only: ord_divides)
-    also have "\<dots> \<longleftrightarrow> [e + c = e + 0] (mod ord n a)"
-      using cong_add_lcancel_eq[of e c 0 "ord n a", simplified cong_0_divides]
-      by simp
-    finally have "[a^d = a^e] (mod n) \<longleftrightarrow> [d = e] (mod (ord n a))"
-      using c by simp }
-  note th = this
-  have "e \<le> d \<or> d \<le> e" by arith
-  moreover
-  {assume ed: "e \<le> d" from th[OF na ed] have ?thesis .}
-  moreover
-  {assume de: "d \<le> e"
-    from th[OF na de] have ?thesis by (simp add: cong_commute) }
-  ultimately show ?thesis by blast
-qed
-
-(* Another trivial primality characterization.                               *)
-
-lemma prime_prime_factor:
-  "prime n \<longleftrightarrow> n \<noteq> 1\<and> (\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n)"
-proof-
-  {assume n: "n=0 \<or> n=1" hence ?thesis using prime_0 two_is_prime by auto}
-  moreover
-  {assume n: "n\<noteq>0" "n\<noteq>1"
-    {assume pn: "prime n"
-
-      from pn[unfolded prime_def] have "\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n"
-	using n
-	apply (cases "n = 0 \<or> n=1",simp)
-	by (clarsimp, erule_tac x="p" in allE, auto)}
-    moreover
-    {assume H: "\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n"
-      from n have n1: "n > 1" by arith
-      {fix m assume m: "m dvd n" "m\<noteq>1"
-	from prime_factor[OF m(2)] obtain p where
-	  p: "prime p" "p dvd m" by blast
-	from dvd_trans[OF p(2) m(1)] p(1) H have "p = n" by blast
-	with p(2) have "n dvd m"  by simp
-	hence "m=n"  using dvd_anti_sym[OF m(1)] by simp }
-      with n1 have "prime n"  unfolding prime_def by auto }
-    ultimately have ?thesis using n by blast}
-  ultimately       show ?thesis by auto
-qed
-
-lemma prime_divisor_sqrt:
-  "prime n \<longleftrightarrow> n \<noteq> 1 \<and> (\<forall>d. d dvd n \<and> d^2 \<le> n \<longrightarrow> d = 1)"
-proof-
-  {assume "n=0 \<or> n=1" hence ?thesis using prime_0 prime_1
-    by (auto simp add: nat_power_eq_0_iff)}
-  moreover
-  {assume n: "n\<noteq>0" "n\<noteq>1"
-    hence np: "n > 1" by arith
-    {fix d assume d: "d dvd n" "d^2 \<le> n" and H: "\<forall>m. m dvd n \<longrightarrow> m=1 \<or> m=n"
-      from H d have d1n: "d = 1 \<or> d=n" by blast
-      {assume dn: "d=n"
-	have "n^2 > n*1" using n
-	  by (simp add: power2_eq_square mult_less_cancel1)
-	with dn d(2) have "d=1" by simp}
-      with d1n have "d = 1" by blast  }
-    moreover
-    {fix d assume d: "d dvd n" and H: "\<forall>d'. d' dvd n \<and> d'^2 \<le> n \<longrightarrow> d' = 1"
-      from d n have "d \<noteq> 0" apply - apply (rule ccontr) by simp
-      hence dp: "d > 0" by simp
-      from d[unfolded dvd_def] obtain e where e: "n= d*e" by blast
-      from n dp e have ep:"e > 0" by simp
-      have "d^2 \<le> n \<or> e^2 \<le> n" using dp ep
-	by (auto simp add: e power2_eq_square mult_le_cancel_left)
-      moreover
-      {assume h: "d^2 \<le> n"
-	from H[rule_format, of d] h d have "d = 1" by blast}
-      moreover
-      {assume h: "e^2 \<le> n"
-	from e have "e dvd n" unfolding dvd_def by (simp add: mult_commute)
-	with H[rule_format, of e] h have "e=1" by simp
-	with e have "d = n" by simp}
-      ultimately have "d=1 \<or> d=n"  by blast}
-    ultimately have ?thesis unfolding prime_def using np n(2) by blast}
-  ultimately show ?thesis by auto
-qed
-lemma prime_prime_factor_sqrt:
-  "prime n \<longleftrightarrow> n \<noteq> 0 \<and> n \<noteq> 1 \<and> \<not> (\<exists>p. prime p \<and> p dvd n \<and> p^2 \<le> n)"
-  (is "?lhs \<longleftrightarrow>?rhs")
-proof-
-  {assume "n=0 \<or> n=1" hence ?thesis using prime_0 prime_1 by auto}
-  moreover
-  {assume n: "n\<noteq>0" "n\<noteq>1"
-    {assume H: ?lhs
-      from H[unfolded prime_divisor_sqrt] n
-      have ?rhs  apply clarsimp by (erule_tac x="p" in allE, simp add: prime_1)
-    }
-    moreover
-    {assume H: ?rhs
-      {fix d assume d: "d dvd n" "d^2 \<le> n" "d\<noteq>1"
-	from prime_factor[OF d(3)]
-	obtain p where p: "prime p" "p dvd d" by blast
-	from n have np: "n > 0" by arith
-	from d(1) n have "d \<noteq> 0" by - (rule ccontr, auto)
-	hence dp: "d > 0" by arith
-	from mult_mono[OF dvd_imp_le[OF p(2) dp] dvd_imp_le[OF p(2) dp]] d(2)
-	have "p^2 \<le> n" unfolding power2_eq_square by arith
-	with H n p(1) dvd_trans[OF p(2) d(1)] have False  by blast}
-      with n prime_divisor_sqrt  have ?lhs by auto}
-    ultimately have ?thesis by blast }
-  ultimately show ?thesis by (cases "n=0 \<or> n=1", auto)
-qed
-(* Pocklington theorem. *)
-
-lemma pocklington_lemma:
-  assumes n: "n \<ge> 2" and nqr: "n - 1 = q*r" and an: "[a^ (n - 1) = 1] (mod n)"
-  and aq:"\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a^ ((n - 1) div p) - 1) n"
-  and pp: "prime p" and pn: "p dvd n"
-  shows "[p = 1] (mod q)"
-proof-
-  from pp prime_0 prime_1 have p01: "p \<noteq> 0" "p \<noteq> 1" by - (rule ccontr, simp)+
-  from cong_1_divides[OF an, unfolded nqr, unfolded dvd_def]
-  obtain k where k: "a ^ (q * r) - 1 = n*k" by blast
-  from pn[unfolded dvd_def] obtain l where l: "n = p*l" by blast
-  {assume a0: "a = 0"
-    hence "a^ (n - 1) = 0" using n by (simp add: power_0_left)
-    with n an mod_less[of 1 n]  have False by (simp add: power_0_left modeq_def)}
-  hence a0: "a\<noteq>0" ..
-  from n nqr have aqr0: "a ^ (q * r) \<noteq> 0" using a0 by (simp add: neq0_conv)
-  hence "(a ^ (q * r) - 1) + 1  = a ^ (q * r)" by simp
-  with k l have "a ^ (q * r) = p*l*k + 1" by simp
-  hence "a ^ (r * q) + p * 0 = 1 + p * (l*k)" by (simp add: mult_ac)
-  hence odq: "ord p (a^r) dvd q"
-    unfolding ord_divides[symmetric] power_mult[symmetric] nat_mod  by blast
-  from odq[unfolded dvd_def] obtain d where d: "q = ord p (a^r) * d" by blast
-  {assume d1: "d \<noteq> 1"
-    from prime_factor[OF d1] obtain P where P: "prime P" "P dvd d" by blast
-    from d dvd_mult[OF P(2), of "ord p (a^r)"] have Pq: "P dvd q" by simp
-    from aq P(1) Pq have caP:"coprime (a^ ((n - 1) div P) - 1) n" by blast
-    from Pq obtain s where s: "q = P*s" unfolding dvd_def by blast
-    have P0: "P \<noteq> 0" using P(1) prime_0 by - (rule ccontr, simp)
-    from P(2) obtain t where t: "d = P*t" unfolding dvd_def by blast
-    from d s t P0  have s': "ord p (a^r) * t = s" by algebra
-    have "ord p (a^r) * t*r = r * ord p (a^r) * t" by algebra
-    hence exps: "a^(ord p (a^r) * t*r) = ((a ^ r) ^ ord p (a^r)) ^ t"
-      by (simp only: power_mult)
-    have "[((a ^ r) ^ ord p (a^r)) ^ t= 1^t] (mod p)"
-      by (rule cong_exp, rule ord)
-    then have th: "[((a ^ r) ^ ord p (a^r)) ^ t= 1] (mod p)"
-      by (simp add: power_Suc0)
-    from cong_1_divides[OF th] exps have pd0: "p dvd a^(ord p (a^r) * t*r) - 1" by simp
-    from nqr s s' have "(n - 1) div P = ord p (a^r) * t*r" using P0 by simp
-    with caP have "coprime (a^(ord p (a^r) * t*r) - 1) n" by simp
-    with p01 pn pd0 have False unfolding coprime by auto}
-  hence d1: "d = 1" by blast
-  hence o: "ord p (a^r) = q" using d by simp
-  from pp phi_prime[of p] have phip: " \<phi> p = p - 1" by simp
-  {fix d assume d: "d dvd p" "d dvd a" "d \<noteq> 1"
-    from pp[unfolded prime_def] d have dp: "d = p" by blast
-    from n have n12:"Suc (n - 2) = n - 1" by arith
-    with divides_rexp[OF d(2)[unfolded dp], of "n - 2"]
-    have th0: "p dvd a ^ (n - 1)" by simp
-    from n have n0: "n \<noteq> 0" by simp
-    from d(2) an n12[symmetric] have a0: "a \<noteq> 0"
-      by - (rule ccontr, simp add: modeq_def)
-    have th1: "a^ (n - 1) \<noteq> 0" using n d(2) dp a0 by (auto simp add: neq0_conv)
-    from coprime_minus1[OF th1, unfolded coprime]
-      dvd_trans[OF pn cong_1_divides[OF an]] th0 d(3) dp
-    have False by auto}
-  hence cpa: "coprime p a" using coprime by auto
-  from coprime_exp[OF cpa, of r] coprime_commute
-  have arp: "coprime (a^r) p" by blast
-  from fermat_little[OF arp, simplified ord_divides] o phip
-  have "q dvd (p - 1)" by simp
-  then obtain d where d:"p - 1 = q * d" unfolding dvd_def by blast
-  from prime_0 pp have p0:"p \<noteq> 0" by -  (rule ccontr, auto)
-  from p0 d have "p + q * 0 = 1 + q * d" by simp
-  with nat_mod[of p 1 q, symmetric]
-  show ?thesis by blast
-qed
-
-lemma pocklington:
-  assumes n: "n \<ge> 2" and nqr: "n - 1 = q*r" and sqr: "n \<le> q^2"
-  and an: "[a^ (n - 1) = 1] (mod n)"
-  and aq:"\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a^ ((n - 1) div p) - 1) n"
-  shows "prime n"
-unfolding prime_prime_factor_sqrt[of n]
-proof-
-  let ?ths = "n \<noteq> 0 \<and> n \<noteq> 1 \<and> \<not> (\<exists>p. prime p \<and> p dvd n \<and> p\<twosuperior> \<le> n)"
-  from n have n01: "n\<noteq>0" "n\<noteq>1" by arith+
-  {fix p assume p: "prime p" "p dvd n" "p^2 \<le> n"
-    from p(3) sqr have "p^(Suc 1) \<le> q^(Suc 1)" by (simp add: power2_eq_square)
-    hence pq: "p \<le> q" unfolding exp_mono_le .
-    from pocklington_lemma[OF n nqr an aq p(1,2)]  cong_1_divides
-    have th: "q dvd p - 1" by blast
-    have "p - 1 \<noteq> 0"using prime_ge_2[OF p(1)] by arith
-    with divides_ge[OF th] pq have False by arith }
-  with n01 show ?ths by blast
-qed
-
-(* Variant for application, to separate the exponentiation.                  *)
-lemma pocklington_alt:
-  assumes n: "n \<ge> 2" and nqr: "n - 1 = q*r" and sqr: "n \<le> q^2"
-  and an: "[a^ (n - 1) = 1] (mod n)"
-  and aq:"\<forall>p. prime p \<and> p dvd q \<longrightarrow> (\<exists>b. [a^((n - 1) div p) = b] (mod n) \<and> coprime (b - 1) n)"
-  shows "prime n"
-proof-
-  {fix p assume p: "prime p" "p dvd q"
-    from aq[rule_format] p obtain b where
-      b: "[a^((n - 1) div p) = b] (mod n)" "coprime (b - 1) n" by blast
-    {assume a0: "a=0"
-      from n an have "[0 = 1] (mod n)" unfolding a0 power_0_left by auto
-      hence False using n by (simp add: modeq_def dvd_eq_mod_eq_0[symmetric])}
-    hence a0: "a\<noteq> 0" ..
-    hence a1: "a \<ge> 1" by arith
-    from one_le_power[OF a1] have ath: "1 \<le> a ^ ((n - 1) div p)" .
-    {assume b0: "b = 0"
-      from p(2) nqr have "(n - 1) mod p = 0"
-	apply (simp only: dvd_eq_mod_eq_0[symmetric]) by (rule dvd_mult2, simp)
-      with mod_div_equality[of "n - 1" p]
-      have "(n - 1) div p * p= n - 1" by auto
-      hence eq: "(a^((n - 1) div p))^p = a^(n - 1)"
-	by (simp only: power_mult[symmetric])
-      from prime_ge_2[OF p(1)] have pS: "Suc (p - 1) = p" by arith
-      from b(1) have d: "n dvd a^((n - 1) div p)" unfolding b0 cong_0_divides .
-      from divides_rexp[OF d, of "p - 1"] pS eq cong_divides[OF an] n
-      have False by simp}
-    then have b0: "b \<noteq> 0" ..
-    hence b1: "b \<ge> 1" by arith
-    from cong_coprime[OF cong_sub[OF b(1) cong_refl[of 1] ath b1]] b(2) nqr
-    have "coprime (a ^ ((n - 1) div p) - 1) n" by (simp add: coprime_commute)}
-  hence th: "\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a ^ ((n - 1) div p) - 1) n "
-    by blast
-  from pocklington[OF n nqr sqr an th] show ?thesis .
-qed
-
-(* Prime factorizations.                                                     *)
-
-definition "primefact ps n = (foldr op * ps  1 = n \<and> (\<forall>p\<in> set ps. prime p))"
-
-lemma primefact: assumes n: "n \<noteq> 0"
-  shows "\<exists>ps. primefact ps n"
-using n
-proof(induct n rule: nat_less_induct)
-  fix n assume H: "\<forall>m<n. m \<noteq> 0 \<longrightarrow> (\<exists>ps. primefact ps m)" and n: "n\<noteq>0"
-  let ?ths = "\<exists>ps. primefact ps n"
-  {assume "n = 1"
-    hence "primefact [] n" by (simp add: primefact_def)
-    hence ?ths by blast }
-  moreover
-  {assume n1: "n \<noteq> 1"
-    with n have n2: "n \<ge> 2" by arith
-    from prime_factor[OF n1] obtain p where p: "prime p" "p dvd n" by blast
-    from p(2) obtain m where m: "n = p*m" unfolding dvd_def by blast
-    from n m have m0: "m > 0" "m\<noteq>0" by auto
-    from prime_ge_2[OF p(1)] have "1 < p" by arith
-    with m0 m have mn: "m < n" by auto
-    from H[rule_format, OF mn m0(2)] obtain ps where ps: "primefact ps m" ..
-    from ps m p(1) have "primefact (p#ps) n" by (simp add: primefact_def)
-    hence ?ths by blast}
-  ultimately show ?ths by blast
-qed
-
-lemma primefact_contains:
-  assumes pf: "primefact ps n" and p: "prime p" and pn: "p dvd n"
-  shows "p \<in> set ps"
-  using pf p pn
-proof(induct ps arbitrary: p n)
-  case Nil thus ?case by (auto simp add: primefact_def)
-next
-  case (Cons q qs p n)
-  from Cons.prems[unfolded primefact_def]
-  have q: "prime q" "q * foldr op * qs 1 = n" "\<forall>p \<in>set qs. prime p"  and p: "prime p" "p dvd q * foldr op * qs 1" by simp_all
-  {assume "p dvd q"
-    with p(1) q(1) have "p = q" unfolding prime_def by auto
-    hence ?case by simp}
-  moreover
-  { assume h: "p dvd foldr op * qs 1"
-    from q(3) have pqs: "primefact qs (foldr op * qs 1)"
-      by (simp add: primefact_def)
-    from Cons.hyps[OF pqs p(1) h] have ?case by simp}
-  ultimately show ?case using prime_divprod[OF p] by blast
-qed
-
-lemma primefact_variant: "primefact ps n \<longleftrightarrow> foldr op * ps 1 = n \<and> list_all prime ps" by (auto simp add: primefact_def list_all_iff)
-
-(* Variant of Lucas theorem.                                                 *)
-
-lemma lucas_primefact:
-  assumes n: "n \<ge> 2" and an: "[a^(n - 1) = 1] (mod n)"
-  and psn: "foldr op * ps 1 = n - 1"
-  and psp: "list_all (\<lambda>p. prime p \<and> \<not> [a^((n - 1) div p) = 1] (mod n)) ps"
-  shows "prime n"
-proof-
-  {fix p assume p: "prime p" "p dvd n - 1" "[a ^ ((n - 1) div p) = 1] (mod n)"
-    from psn psp have psn1: "primefact ps (n - 1)"
-      by (auto simp add: list_all_iff primefact_variant)
-    from p(3) primefact_contains[OF psn1 p(1,2)] psp
-    have False by (induct ps, auto)}
-  with lucas[OF n an] show ?thesis by blast
-qed
-
-(* Variant of Pocklington theorem.                                           *)
-
-lemma mod_le: assumes n: "n \<noteq> (0::nat)" shows "m mod n \<le> m"
-proof-
-    from mod_div_equality[of m n]
-    have "\<exists>x. x + m mod n = m" by blast
-    then show ?thesis by auto
-qed
-
-
-lemma pocklington_primefact:
-  assumes n: "n \<ge> 2" and qrn: "q*r = n - 1" and nq2: "n \<le> q^2"
-  and arnb: "(a^r) mod n = b" and psq: "foldr op * ps 1 = q"
-  and bqn: "(b^q) mod n = 1"
-  and psp: "list_all (\<lambda>p. prime p \<and> coprime ((b^(q div p)) mod n - 1) n) ps"
-  shows "prime n"
-proof-
-  from bqn psp qrn
-  have bqn: "a ^ (n - 1) mod n = 1"
-    and psp: "list_all (\<lambda>p. prime p \<and> coprime (a^(r *(q div p)) mod n - 1) n) ps"  unfolding arnb[symmetric] power_mod
-    by (simp_all add: power_mult[symmetric] algebra_simps)
-  from n  have n0: "n > 0" by arith
-  from mod_div_equality[of "a^(n - 1)" n]
-    mod_less_divisor[OF n0, of "a^(n - 1)"]
-  have an1: "[a ^ (n - 1) = 1] (mod n)"
-    unfolding nat_mod bqn
-    apply -
-    apply (rule exI[where x="0"])
-    apply (rule exI[where x="a^(n - 1) div n"])
-    by (simp add: algebra_simps)
-  {fix p assume p: "prime p" "p dvd q"
-    from psp psq have pfpsq: "primefact ps q"
-      by (auto simp add: primefact_variant list_all_iff)
-    from psp primefact_contains[OF pfpsq p]
-    have p': "coprime (a ^ (r * (q div p)) mod n - 1) n"
-      by (simp add: list_all_iff)
-    from prime_ge_2[OF p(1)] have p01: "p \<noteq> 0" "p \<noteq> 1" "p =Suc(p - 1)" by arith+
-    from div_mult1_eq[of r q p] p(2)
-    have eq1: "r* (q div p) = (n - 1) div p"
-      unfolding qrn[symmetric] dvd_eq_mod_eq_0 by (simp add: mult_commute)
-    have ath: "\<And>a (b::nat). a <= b \<Longrightarrow> a \<noteq> 0 ==> 1 <= a \<and> 1 <= b" by arith
-    from n0 have n00: "n \<noteq> 0" by arith
-    from mod_le[OF n00]
-    have th10: "a ^ ((n - 1) div p) mod n \<le> a ^ ((n - 1) div p)" .
-    {assume "a ^ ((n - 1) div p) mod n = 0"
-      then obtain s where s: "a ^ ((n - 1) div p) = n*s"
-	unfolding mod_eq_0_iff by blast
-      hence eq0: "(a^((n - 1) div p))^p = (n*s)^p" by simp
-      from qrn[symmetric] have qn1: "q dvd n - 1" unfolding dvd_def by auto
-      from dvd_trans[OF p(2) qn1] div_mod_equality'[of "n - 1" p]
-      have npp: "(n - 1) div p * p = n - 1" by (simp add: dvd_eq_mod_eq_0)
-      with eq0 have "a^ (n - 1) = (n*s)^p"
-	by (simp add: power_mult[symmetric])
-      hence "1 = (n*s)^(Suc (p - 1)) mod n" using bqn p01 by simp
-      also have "\<dots> = 0" by (simp add: mult_assoc)
-      finally have False by simp }
-      then have th11: "a ^ ((n - 1) div p) mod n \<noteq> 0" by auto
-    have th1: "[a ^ ((n - 1) div p) mod n = a ^ ((n - 1) div p)] (mod n)"
-      unfolding modeq_def by simp
-    from cong_sub[OF th1 cong_refl[of 1]]  ath[OF th10 th11]
-    have th: "[a ^ ((n - 1) div p) mod n - 1 = a ^ ((n - 1) div p) - 1] (mod n)"
-      by blast
-    from cong_coprime[OF th] p'[unfolded eq1]
-    have "coprime (a ^ ((n - 1) div p) - 1) n" by (simp add: coprime_commute) }
-  with pocklington[OF n qrn[symmetric] nq2 an1]
-  show ?thesis by blast
-qed
-
-end
--- a/src/HOL/Library/Primes.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,828 +0,0 @@
-(*  Title:      HOL/Library/Primes.thy
-    Author:     Amine Chaieb, Christophe Tabacznyj and Lawrence C Paulson
-    Copyright   1996  University of Cambridge
-*)
-
-header {* Primality on nat *}
-
-theory Primes
-imports Complex_Main Legacy_GCD
-begin
-
-hide (open) const GCD.gcd GCD.coprime GCD.prime
-
-definition
-  coprime :: "nat => nat => bool" where
-  "coprime m n \<longleftrightarrow> gcd m n = 1"
-
-definition
-  prime :: "nat \<Rightarrow> bool" where
-  [code del]: "prime p \<longleftrightarrow> (1 < p \<and> (\<forall>m. m dvd p --> m = 1 \<or> m = p))"
-
-
-lemma two_is_prime: "prime 2"
-  apply (auto simp add: prime_def)
-  apply (case_tac m)
-   apply (auto dest!: dvd_imp_le)
-  done
-
-lemma prime_imp_relprime: "prime p ==> \<not> p dvd n ==> gcd p n = 1"
-  apply (auto simp add: prime_def)
-  apply (metis One_nat_def gcd_dvd1 gcd_dvd2)
-  done
-
-text {*
-  This theorem leads immediately to a proof of the uniqueness of
-  factorization.  If @{term p} divides a product of primes then it is
-  one of those primes.
-*}
-
-lemma prime_dvd_mult: "prime p ==> p dvd m * n ==> p dvd m \<or> p dvd n"
-  by (blast intro: relprime_dvd_mult prime_imp_relprime)
-
-lemma prime_dvd_square: "prime p ==> p dvd m^Suc (Suc 0) ==> p dvd m"
-  by (auto dest: prime_dvd_mult)
-
-lemma prime_dvd_power_two: "prime p ==> p dvd m\<twosuperior> ==> p dvd m"
-  by (rule prime_dvd_square) (simp_all add: power2_eq_square)
-
-
-lemma exp_eq_1:"(x::nat)^n = 1 \<longleftrightarrow> x = 1 \<or> n = 0"
-by (induct n, auto)
-
-lemma exp_mono_lt: "(x::nat) ^ (Suc n) < y ^ (Suc n) \<longleftrightarrow> x < y"
-by(metis linorder_not_less not_less0 power_le_imp_le_base power_less_imp_less_base)
-
-lemma exp_mono_le: "(x::nat) ^ (Suc n) \<le> y ^ (Suc n) \<longleftrightarrow> x \<le> y"
-by (simp only: linorder_not_less[symmetric] exp_mono_lt)
-
-lemma exp_mono_eq: "(x::nat) ^ Suc n = y ^ Suc n \<longleftrightarrow> x = y"
-using power_inject_base[of x n y] by auto
-
-
-lemma even_square: assumes e: "even (n::nat)" shows "\<exists>x. n ^ 2 = 4*x"
-proof-
-  from e have "2 dvd n" by presburger
-  then obtain k where k: "n = 2*k" using dvd_def by auto
-  hence "n^2 = 4* (k^2)" by (simp add: power2_eq_square)
-  thus ?thesis by blast
-qed
-
-lemma odd_square: assumes e: "odd (n::nat)" shows "\<exists>x. n ^ 2 = 4*x + 1"
-proof-
-  from e have np: "n > 0" by presburger
-  from e have "2 dvd (n - 1)" by presburger
-  then obtain k where "n - 1 = 2*k" using dvd_def by auto
-  hence k: "n = 2*k + 1"  using e by presburger 
-  hence "n^2 = 4* (k^2 + k) + 1" by algebra   
-  thus ?thesis by blast
-qed
-
-lemma diff_square: "(x::nat)^2 - y^2 = (x+y)*(x - y)" 
-proof-
-  have "x \<le> y \<or> y \<le> x" by (rule nat_le_linear)
-  moreover
-  {assume le: "x \<le> y"
-    hence "x ^2 \<le> y^2" by (simp only: numeral_2_eq_2 exp_mono_le Let_def)
-    with le have ?thesis by simp }
-  moreover
-  {assume le: "y \<le> x"
-    hence le2: "y ^2 \<le> x^2" by (simp only: numeral_2_eq_2 exp_mono_le Let_def)
-    from le have "\<exists>z. y + z = x" by presburger
-    then obtain z where z: "x = y + z" by blast 
-    from le2 have "\<exists>z. x^2 = y^2 + z" by presburger
-    then obtain z2 where z2: "x^2 = y^2 + z2"  by blast
-    from z z2 have ?thesis apply simp by algebra }
-  ultimately show ?thesis by blast  
-qed
-
-text {* Elementary theory of divisibility *}
-lemma divides_ge: "(a::nat) dvd b \<Longrightarrow> b = 0 \<or> a \<le> b" unfolding dvd_def by auto
-lemma divides_antisym: "(x::nat) dvd y \<and> y dvd x \<longleftrightarrow> x = y"
-  using dvd_anti_sym[of x y] by auto
-
-lemma divides_add_revr: assumes da: "(d::nat) dvd a" and dab:"d dvd (a + b)"
-  shows "d dvd b"
-proof-
-  from da obtain k where k:"a = d*k" by (auto simp add: dvd_def)
-  from dab obtain k' where k': "a + b = d*k'" by (auto simp add: dvd_def)
-  from k k' have "b = d *(k' - k)" by (simp add : diff_mult_distrib2)
-  thus ?thesis unfolding dvd_def by blast
-qed
-
-declare nat_mult_dvd_cancel_disj[presburger]
-lemma nat_mult_dvd_cancel_disj'[presburger]: 
-  "(m\<Colon>nat)*k dvd n*k \<longleftrightarrow> k = 0 \<or> m dvd n" unfolding mult_commute[of m k] mult_commute[of n k] by presburger 
-
-lemma divides_mul_l: "(a::nat) dvd b ==> (c * a) dvd (c * b)"
-  by presburger
-
-lemma divides_mul_r: "(a::nat) dvd b ==> (a * c) dvd (b * c)" by presburger
-lemma divides_cases: "(n::nat) dvd m ==> m = 0 \<or> m = n \<or> 2 * n <= m" 
-  by (auto simp add: dvd_def)
-
-lemma divides_div_not: "(x::nat) = (q * n) + r \<Longrightarrow> 0 < r \<Longrightarrow> r < n ==> ~(n dvd x)"
-proof(auto simp add: dvd_def)
-  fix k assume H: "0 < r" "r < n" "q * n + r = n * k"
-  from H(3) have r: "r = n* (k -q)" by(simp add: diff_mult_distrib2 mult_commute)
-  {assume "k - q = 0" with r H(1) have False by simp}
-  moreover
-  {assume "k - q \<noteq> 0" with r have "r \<ge> n" by auto
-    with H(2) have False by simp}
-  ultimately show False by blast
-qed
-lemma divides_exp: "(x::nat) dvd y ==> x ^ n dvd y ^ n"
-  by (auto simp add: power_mult_distrib dvd_def)
-
-lemma divides_exp2: "n \<noteq> 0 \<Longrightarrow> (x::nat) ^ n dvd y \<Longrightarrow> x dvd y" 
-  by (induct n ,auto simp add: dvd_def)
-
-fun fact :: "nat \<Rightarrow> nat" where
-  "fact 0 = 1"
-| "fact (Suc n) = Suc n * fact n"	
-
-lemma fact_lt: "0 < fact n" by(induct n, simp_all)
-lemma fact_le: "fact n \<ge> 1" using fact_lt[of n] by simp 
-lemma fact_mono: assumes le: "m \<le> n" shows "fact m \<le> fact n"
-proof-
-  from le have "\<exists>i. n = m+i" by presburger
-  then obtain i where i: "n = m+i" by blast 
-  have "fact m \<le> fact (m + i)"
-  proof(induct m)
-    case 0 thus ?case using fact_le[of i] by simp
-  next
-    case (Suc m)
-    have "fact (Suc m) = Suc m * fact m" by simp
-    have th1: "Suc m \<le> Suc (m + i)" by simp
-    from mult_le_mono[of "Suc m" "Suc (m+i)" "fact m" "fact (m+i)", OF th1 Suc.hyps]
-    show ?case by simp
-  qed
-  thus ?thesis using i by simp
-qed
-
-lemma divides_fact: "1 <= p \<Longrightarrow> p <= n ==> p dvd fact n"
-proof(induct n arbitrary: p)
-  case 0 thus ?case by simp
-next
-  case (Suc n p)
-  from Suc.prems have "p = Suc n \<or> p \<le> n" by presburger 
-  moreover
-  {assume "p = Suc n" hence ?case  by (simp only: fact.simps dvd_triv_left)}
-  moreover
-  {assume "p \<le> n"
-    with Suc.prems(1) Suc.hyps have th: "p dvd fact n" by simp
-    from dvd_mult[OF th] have ?case by (simp only: fact.simps) }
-  ultimately show ?case by blast
-qed
-
-declare dvd_triv_left[presburger]
-declare dvd_triv_right[presburger]
-lemma divides_rexp: 
-  "x dvd y \<Longrightarrow> (x::nat) dvd (y^(Suc n))" by (simp add: dvd_mult2[of x y])
-
-text {* Coprimality *}
-
-lemma coprime: "coprime a b \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> d = 1)"
-using gcd_unique[of 1 a b, simplified] by (auto simp add: coprime_def)
-lemma coprime_commute: "coprime a b \<longleftrightarrow> coprime b a" by (simp add: coprime_def gcd_commute)
-
-lemma coprime_bezout: "coprime a b \<longleftrightarrow> (\<exists>x y. a * x - b * y = 1 \<or> b * x - a * y = 1)"
-using coprime_def gcd_bezout by auto
-
-lemma coprime_divprod: "d dvd a * b  \<Longrightarrow> coprime d a \<Longrightarrow> d dvd b"
-  using relprime_dvd_mult_iff[of d a b] by (auto simp add: coprime_def mult_commute)
-
-lemma coprime_1[simp]: "coprime a 1" by (simp add: coprime_def)
-lemma coprime_1'[simp]: "coprime 1 a" by (simp add: coprime_def)
-lemma coprime_Suc0[simp]: "coprime a (Suc 0)" by (simp add: coprime_def)
-lemma coprime_Suc0'[simp]: "coprime (Suc 0) a" by (simp add: coprime_def)
-
-lemma gcd_coprime: 
-  assumes z: "gcd a b \<noteq> 0" and a: "a = a' * gcd a b" and b: "b = b' * gcd a b" 
-  shows    "coprime a' b'"
-proof-
-  let ?g = "gcd a b"
-  {assume bz: "a = 0" from b bz z a have ?thesis by (simp add: gcd_zero coprime_def)}
-  moreover 
-  {assume az: "a\<noteq> 0" 
-    from z have z': "?g > 0" by simp
-    from bezout_gcd_strong[OF az, of b] 
-    obtain x y where xy: "a*x = b*y + ?g" by blast
-    from xy a b have "?g * a'*x = ?g * (b'*y + 1)" by (simp add: algebra_simps)
-    hence "?g * (a'*x) = ?g * (b'*y + 1)" by (simp add: mult_assoc)
-    hence "a'*x = (b'*y + 1)"
-      by (simp only: nat_mult_eq_cancel1[OF z']) 
-    hence "a'*x - b'*y = 1" by simp
-    with coprime_bezout[of a' b'] have ?thesis by auto}
-  ultimately show ?thesis by blast
-qed
-lemma coprime_0: "coprime d 0 \<longleftrightarrow> d = 1" by (simp add: coprime_def)
-lemma coprime_mul: assumes da: "coprime d a" and db: "coprime d b"
-  shows "coprime d (a * b)"
-proof-
-  from da have th: "gcd a d = 1" by (simp add: coprime_def gcd_commute)
-  from gcd_mult_cancel[of a d b, OF th] db[unfolded coprime_def] have "gcd d (a*b) = 1"
-    by (simp add: gcd_commute)
-  thus ?thesis unfolding coprime_def .
-qed
-lemma coprime_lmul2: assumes dab: "coprime d (a * b)" shows "coprime d b"
-using prems unfolding coprime_bezout
-apply clarsimp
-apply (case_tac "d * x - a * b * y = Suc 0 ", simp_all)
-apply (rule_tac x="x" in exI)
-apply (rule_tac x="a*y" in exI)
-apply (simp add: mult_ac)
-apply (rule_tac x="a*x" in exI)
-apply (rule_tac x="y" in exI)
-apply (simp add: mult_ac)
-done
-
-lemma coprime_rmul2: "coprime d (a * b) \<Longrightarrow> coprime d a"
-unfolding coprime_bezout
-apply clarsimp
-apply (case_tac "d * x - a * b * y = Suc 0 ", simp_all)
-apply (rule_tac x="x" in exI)
-apply (rule_tac x="b*y" in exI)
-apply (simp add: mult_ac)
-apply (rule_tac x="b*x" in exI)
-apply (rule_tac x="y" in exI)
-apply (simp add: mult_ac)
-done
-lemma coprime_mul_eq: "coprime d (a * b) \<longleftrightarrow> coprime d a \<and>  coprime d b"
-  using coprime_rmul2[of d a b] coprime_lmul2[of d a b] coprime_mul[of d a b] 
-  by blast
-
-lemma gcd_coprime_exists:
-  assumes nz: "gcd a b \<noteq> 0" 
-  shows "\<exists>a' b'. a = a' * gcd a b \<and> b = b' * gcd a b \<and> coprime a' b'"
-proof-
-  let ?g = "gcd a b"
-  from gcd_dvd1[of a b] gcd_dvd2[of a b] 
-  obtain a' b' where "a = ?g*a'"  "b = ?g*b'" unfolding dvd_def by blast
-  hence ab': "a = a'*?g" "b = b'*?g" by algebra+
-  from ab' gcd_coprime[OF nz ab'] show ?thesis by blast
-qed
-
-lemma coprime_exp: "coprime d a ==> coprime d (a^n)" 
-  by(induct n, simp_all add: coprime_mul)
-
-lemma coprime_exp_imp: "coprime a b ==> coprime (a ^n) (b ^n)"
-  by (induct n, simp_all add: coprime_mul_eq coprime_commute coprime_exp)
-lemma coprime_refl[simp]: "coprime n n \<longleftrightarrow> n = 1" by (simp add: coprime_def)
-lemma coprime_plus1[simp]: "coprime (n + 1) n"
-  apply (simp add: coprime_bezout)
-  apply (rule exI[where x=1])
-  apply (rule exI[where x=1])
-  apply simp
-  done
-lemma coprime_minus1: "n \<noteq> 0 ==> coprime (n - 1) n"
-  using coprime_plus1[of "n - 1"] coprime_commute[of "n - 1" n] by auto
-
-lemma bezout_gcd_pow: "\<exists>x y. a ^n * x - b ^ n * y = gcd a b ^ n \<or> b ^ n * x - a ^ n * y = gcd a b ^ n"
-proof-
-  let ?g = "gcd a b"
-  {assume z: "?g = 0" hence ?thesis 
-      apply (cases n, simp)
-      apply arith
-      apply (simp only: z power_0_Suc)
-      apply (rule exI[where x=0])
-      apply (rule exI[where x=0])
-      by simp}
-  moreover
-  {assume z: "?g \<noteq> 0"
-    from gcd_dvd1[of a b] gcd_dvd2[of a b] obtain a' b' where
-      ab': "a = a'*?g" "b = b'*?g" unfolding dvd_def by (auto simp add: mult_ac)
-    hence ab'': "?g*a' = a" "?g * b' = b" by algebra+
-    from coprime_exp_imp[OF gcd_coprime[OF z ab'], unfolded coprime_bezout, of n]
-    obtain x y where "a'^n * x - b'^n * y = 1 \<or> b'^n * x - a'^n * y = 1"  by blast
-    hence "?g^n * (a'^n * x - b'^n * y) = ?g^n \<or> ?g^n*(b'^n * x - a'^n * y) = ?g^n"
-      using z by auto 
-    then have "a^n * x - b^n * y = ?g^n \<or> b^n * x - a^n * y = ?g^n"
-      using z ab'' by (simp only: power_mult_distrib[symmetric] 
-	diff_mult_distrib2 mult_assoc[symmetric])
-    hence  ?thesis by blast }
-  ultimately show ?thesis by blast
-qed
-
-lemma gcd_exp: "gcd (a^n) (b^n) = gcd a b^n"
-proof-
-  let ?g = "gcd (a^n) (b^n)"
-  let ?gn = "gcd a b^n"
-  {fix e assume H: "e dvd a^n" "e dvd b^n"
-    from bezout_gcd_pow[of a n b] obtain x y 
-      where xy: "a ^ n * x - b ^ n * y = ?gn \<or> b ^ n * x - a ^ n * y = ?gn" by blast
-    from dvd_diff_nat [OF dvd_mult2[OF H(1), of x] dvd_mult2[OF H(2), of y]]
-      dvd_diff_nat [OF dvd_mult2[OF H(2), of x] dvd_mult2[OF H(1), of y]] xy
-    have "e dvd ?gn" by (cases "a ^ n * x - b ^ n * y = gcd a b ^ n", simp_all)}
-  hence th:  "\<forall>e. e dvd a^n \<and> e dvd b^n \<longrightarrow> e dvd ?gn" by blast
-  from divides_exp[OF gcd_dvd1[of a b], of n] divides_exp[OF gcd_dvd2[of a b], of n] th
-    gcd_unique have "?gn = ?g" by blast thus ?thesis by simp 
-qed
-
-lemma coprime_exp2:  "coprime (a ^ Suc n) (b^ Suc n) \<longleftrightarrow> coprime a b"
-by (simp only: coprime_def gcd_exp exp_eq_1) simp
-
-lemma division_decomp: assumes dc: "(a::nat) dvd b * c"
-  shows "\<exists>b' c'. a = b' * c' \<and> b' dvd b \<and> c' dvd c"
-proof-
-  let ?g = "gcd a b"
-  {assume "?g = 0" with dc have ?thesis apply (simp add: gcd_zero)
-      apply (rule exI[where x="0"])
-      by (rule exI[where x="c"], simp)}
-  moreover
-  {assume z: "?g \<noteq> 0"
-    from gcd_coprime_exists[OF z]
-    obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'" by blast
-    from gcd_dvd2[of a b] have thb: "?g dvd b" .
-    from ab'(1) have "a' dvd a"  unfolding dvd_def by blast  
-    with dc have th0: "a' dvd b*c" using dvd_trans[of a' a "b*c"] by simp
-    from dc ab'(1,2) have "a'*?g dvd (b'*?g) *c" by auto
-    hence "?g*a' dvd ?g * (b' * c)" by (simp add: mult_assoc)
-    with z have th_1: "a' dvd b'*c" by simp
-    from coprime_divprod[OF th_1 ab'(3)] have thc: "a' dvd c" . 
-    from ab' have "a = ?g*a'" by algebra
-    with thb thc have ?thesis by blast }
-  ultimately show ?thesis by blast
-qed
-
-lemma nat_power_eq_0_iff: "(m::nat) ^ n = 0 \<longleftrightarrow> n \<noteq> 0 \<and> m = 0" by (induct n, auto)
-
-lemma divides_rev: assumes ab: "(a::nat) ^ n dvd b ^n" and n:"n \<noteq> 0" shows "a dvd b"
-proof-
-  let ?g = "gcd a b"
-  from n obtain m where m: "n = Suc m" by (cases n, simp_all)
-  {assume "?g = 0" with ab n have ?thesis by (simp add: gcd_zero)}
-  moreover
-  {assume z: "?g \<noteq> 0"
-    hence zn: "?g ^ n \<noteq> 0" using n by (simp add: neq0_conv)
-    from gcd_coprime_exists[OF z] 
-    obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'" by blast
-    from ab have "(a' * ?g) ^ n dvd (b' * ?g)^n" by (simp add: ab'(1,2)[symmetric])
-    hence "?g^n*a'^n dvd ?g^n *b'^n" by (simp only: power_mult_distrib mult_commute)
-    with zn z n have th0:"a'^n dvd b'^n" by (auto simp add: nat_power_eq_0_iff)
-    have "a' dvd a'^n" by (simp add: m)
-    with th0 have "a' dvd b'^n" using dvd_trans[of a' "a'^n" "b'^n"] by simp
-    hence th1: "a' dvd b'^m * b'" by (simp add: m mult_commute)
-    from coprime_divprod[OF th1 coprime_exp[OF ab'(3), of m]]
-    have "a' dvd b'" .
-    hence "a'*?g dvd b'*?g" by simp
-    with ab'(1,2)  have ?thesis by simp }
-  ultimately show ?thesis by blast
-qed
-
-lemma divides_mul: assumes mr: "m dvd r" and nr: "n dvd r" and mn:"coprime m n" 
-  shows "m * n dvd r"
-proof-
-  from mr nr obtain m' n' where m': "r = m*m'" and n': "r = n*n'"
-    unfolding dvd_def by blast
-  from mr n' have "m dvd n'*n" by (simp add: mult_commute)
-  hence "m dvd n'" using relprime_dvd_mult_iff[OF mn[unfolded coprime_def]] by simp
-  then obtain k where k: "n' = m*k" unfolding dvd_def by blast
-  from n' k show ?thesis unfolding dvd_def by auto
-qed
-
-
-text {* A binary form of the Chinese Remainder Theorem. *}
-
-lemma chinese_remainder: assumes ab: "coprime a b" and a:"a \<noteq> 0" and b:"b \<noteq> 0"
-  shows "\<exists>x q1 q2. x = u + q1 * a \<and> x = v + q2 * b"
-proof-
-  from bezout_add_strong[OF a, of b] bezout_add_strong[OF b, of a]
-  obtain d1 x1 y1 d2 x2 y2 where dxy1: "d1 dvd a" "d1 dvd b" "a * x1 = b * y1 + d1" 
-    and dxy2: "d2 dvd b" "d2 dvd a" "b * x2 = a * y2 + d2" by blast
-  from gcd_unique[of 1 a b, simplified ab[unfolded coprime_def], simplified] 
-    dxy1(1,2) dxy2(1,2) have d12: "d1 = 1" "d2 =1" by auto
-  let ?x = "v * a * x1 + u * b * x2"
-  let ?q1 = "v * x1 + u * y2"
-  let ?q2 = "v * y1 + u * x2"
-  from dxy2(3)[simplified d12] dxy1(3)[simplified d12] 
-  have "?x = u + ?q1 * a" "?x = v + ?q2 * b" by algebra+ 
-  thus ?thesis by blast
-qed
-
-text {* Primality *}
-
-text {* A few useful theorems about primes *}
-
-lemma prime_0[simp]: "~prime 0" by (simp add: prime_def)
-lemma prime_1[simp]: "~ prime 1"  by (simp add: prime_def)
-lemma prime_Suc0[simp]: "~ prime (Suc 0)"  by (simp add: prime_def)
-
-lemma prime_ge_2: "prime p ==> p \<ge> 2" by (simp add: prime_def)
-lemma prime_factor: assumes n: "n \<noteq> 1" shows "\<exists> p. prime p \<and> p dvd n"
-using n
-proof(induct n rule: nat_less_induct)
-  fix n
-  assume H: "\<forall>m<n. m \<noteq> 1 \<longrightarrow> (\<exists>p. prime p \<and> p dvd m)" "n \<noteq> 1"
-  let ?ths = "\<exists>p. prime p \<and> p dvd n"
-  {assume "n=0" hence ?ths using two_is_prime by auto}
-  moreover
-  {assume nz: "n\<noteq>0" 
-    {assume "prime n" hence ?ths by - (rule exI[where x="n"], simp)}
-    moreover
-    {assume n: "\<not> prime n"
-      with nz H(2) 
-      obtain k where k:"k dvd n" "k \<noteq> 1" "k \<noteq> n" by (auto simp add: prime_def) 
-      from dvd_imp_le[OF k(1)] nz k(3) have kn: "k < n" by simp
-      from H(1)[rule_format, OF kn k(2)] obtain p where p: "prime p" "p dvd k" by blast
-      from dvd_trans[OF p(2) k(1)] p(1) have ?ths by blast}
-    ultimately have ?ths by blast}
-  ultimately show ?ths by blast
-qed
-
-lemma prime_factor_lt: assumes p: "prime p" and n: "n \<noteq> 0" and npm:"n = p * m"
-  shows "m < n"
-proof-
-  {assume "m=0" with n have ?thesis by simp}
-  moreover
-  {assume m: "m \<noteq> 0"
-    from npm have mn: "m dvd n" unfolding dvd_def by auto
-    from npm m have "n \<noteq> m" using p by auto
-    with dvd_imp_le[OF mn] n have ?thesis by simp}
-  ultimately show ?thesis by blast
-qed
-
-lemma euclid_bound: "\<exists>p. prime p \<and> n < p \<and>  p <= Suc (fact n)"
-proof-
-  have f1: "fact n + 1 \<noteq> 1" using fact_le[of n] by arith 
-  from prime_factor[OF f1] obtain p where p: "prime p" "p dvd fact n + 1" by blast
-  from dvd_imp_le[OF p(2)] have pfn: "p \<le> fact n + 1" by simp
-  {assume np: "p \<le> n"
-    from p(1) have p1: "p \<ge> 1" by (cases p, simp_all)
-    from divides_fact[OF p1 np] have pfn': "p dvd fact n" .
-    from divides_add_revr[OF pfn' p(2)] p(1) have False by simp}
-  hence "n < p" by arith
-  with p(1) pfn show ?thesis by auto
-qed
-
-lemma euclid: "\<exists>p. prime p \<and> p > n" using euclid_bound by auto
-
-lemma primes_infinite: "\<not> (finite {p. prime p})"
-apply(simp add: finite_nat_set_iff_bounded_le)
-apply (metis euclid linorder_not_le)
-done
-
-lemma coprime_prime: assumes ab: "coprime a b"
-  shows "~(prime p \<and> p dvd a \<and> p dvd b)"
-proof
-  assume "prime p \<and> p dvd a \<and> p dvd b"
-  thus False using ab gcd_greatest[of p a b] by (simp add: coprime_def)
-qed
-lemma coprime_prime_eq: "coprime a b \<longleftrightarrow> (\<forall>p. ~(prime p \<and> p dvd a \<and> p dvd b))" 
-  (is "?lhs = ?rhs")
-proof-
-  {assume "?lhs" with coprime_prime  have ?rhs by blast}
-  moreover
-  {assume r: "?rhs" and c: "\<not> ?lhs"
-    then obtain g where g: "g\<noteq>1" "g dvd a" "g dvd b" unfolding coprime_def by blast
-    from prime_factor[OF g(1)] obtain p where p: "prime p" "p dvd g" by blast
-    from dvd_trans [OF p(2) g(2)] dvd_trans [OF p(2) g(3)] 
-    have "p dvd a" "p dvd b" . with p(1) r have False by blast}
-  ultimately show ?thesis by blast
-qed
-
-lemma prime_coprime: assumes p: "prime p" 
-  shows "n = 1 \<or> p dvd n \<or> coprime p n"
-using p prime_imp_relprime[of p n] by (auto simp add: coprime_def)
-
-lemma prime_coprime_strong: "prime p \<Longrightarrow> p dvd n \<or> coprime p n"
-  using prime_coprime[of p n] by auto
-
-declare  coprime_0[simp]
-
-lemma coprime_0'[simp]: "coprime 0 d \<longleftrightarrow> d = 1" by (simp add: coprime_commute[of 0 d])
-lemma coprime_bezout_strong: assumes ab: "coprime a b" and b: "b \<noteq> 1"
-  shows "\<exists>x y. a * x = b * y + 1"
-proof-
-  from ab b have az: "a \<noteq> 0" by - (rule ccontr, auto)
-  from bezout_gcd_strong[OF az, of b] ab[unfolded coprime_def]
-  show ?thesis by auto
-qed
-
-lemma bezout_prime: assumes p: "prime p"  and pa: "\<not> p dvd a"
-  shows "\<exists>x y. a*x = p*y + 1"
-proof-
-  from p have p1: "p \<noteq> 1" using prime_1 by blast 
-  from prime_coprime[OF p, of a] p1 pa have ap: "coprime a p" 
-    by (auto simp add: coprime_commute)
-  from coprime_bezout_strong[OF ap p1] show ?thesis . 
-qed
-lemma prime_divprod: assumes p: "prime p" and pab: "p dvd a*b"
-  shows "p dvd a \<or> p dvd b"
-proof-
-  {assume "a=1" hence ?thesis using pab by simp }
-  moreover
-  {assume "p dvd a" hence ?thesis by blast}
-  moreover
-  {assume pa: "coprime p a" from coprime_divprod[OF pab pa]  have ?thesis .. }
-  ultimately show ?thesis using prime_coprime[OF p, of a] by blast
-qed
-
-lemma prime_divprod_eq: assumes p: "prime p"
-  shows "p dvd a*b \<longleftrightarrow> p dvd a \<or> p dvd b"
-using p prime_divprod dvd_mult dvd_mult2 by auto
-
-lemma prime_divexp: assumes p:"prime p" and px: "p dvd x^n"
-  shows "p dvd x"
-using px
-proof(induct n)
-  case 0 thus ?case by simp
-next
-  case (Suc n) 
-  hence th: "p dvd x*x^n" by simp
-  {assume H: "p dvd x^n"
-    from Suc.hyps[OF H] have ?case .}
-  with prime_divprod[OF p th] show ?case by blast
-qed
-
-lemma prime_divexp_n: "prime p \<Longrightarrow> p dvd x^n \<Longrightarrow> p^n dvd x^n"
-  using prime_divexp[of p x n] divides_exp[of p x n] by blast
-
-lemma coprime_prime_dvd_ex: assumes xy: "\<not>coprime x y"
-  shows "\<exists>p. prime p \<and> p dvd x \<and> p dvd y"
-proof-
-  from xy[unfolded coprime_def] obtain g where g: "g \<noteq> 1" "g dvd x" "g dvd y" 
-    by blast
-  from prime_factor[OF g(1)] obtain p where p: "prime p" "p dvd g" by blast
-  from g(2,3) dvd_trans[OF p(2)] p(1) show ?thesis by auto
-qed
-lemma coprime_sos: assumes xy: "coprime x y" 
-  shows "coprime (x * y) (x^2 + y^2)"
-proof-
-  {assume c: "\<not> coprime (x * y) (x^2 + y^2)"
-    from coprime_prime_dvd_ex[OF c] obtain p 
-      where p: "prime p" "p dvd x*y" "p dvd x^2 + y^2" by blast
-    {assume px: "p dvd x"
-      from dvd_mult[OF px, of x] p(3) 
-        obtain r s where "x * x = p * r" and "x^2 + y^2 = p * s"
-          by (auto elim!: dvdE)
-        then have "y^2 = p * (s - r)" 
-          by (auto simp add: power2_eq_square diff_mult_distrib2)
-        then have "p dvd y^2" ..
-      with prime_divexp[OF p(1), of y 2] have py: "p dvd y" .
-      from p(1) px py xy[unfolded coprime, rule_format, of p] prime_1  
-      have False by simp }
-    moreover
-    {assume py: "p dvd y"
-      from dvd_mult[OF py, of y] p(3)
-        obtain r s where "y * y = p * r" and "x^2 + y^2 = p * s"
-          by (auto elim!: dvdE)
-        then have "x^2 = p * (s - r)" 
-          by (auto simp add: power2_eq_square diff_mult_distrib2)
-        then have "p dvd x^2" ..
-      with prime_divexp[OF p(1), of x 2] have px: "p dvd x" .
-      from p(1) px py xy[unfolded coprime, rule_format, of p] prime_1  
-      have False by simp }
-    ultimately have False using prime_divprod[OF p(1,2)] by blast}
-  thus ?thesis by blast
-qed
-
-lemma distinct_prime_coprime: "prime p \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"
-  unfolding prime_def coprime_prime_eq by blast
-
-lemma prime_coprime_lt: assumes p: "prime p" and x: "0 < x" and xp: "x < p"
-  shows "coprime x p"
-proof-
-  {assume c: "\<not> coprime x p"
-    then obtain g where g: "g \<noteq> 1" "g dvd x" "g dvd p" unfolding coprime_def by blast
-  from dvd_imp_le[OF g(2)] x xp have gp: "g < p" by arith
-  from g(2) x have "g \<noteq> 0" by - (rule ccontr, simp)
-  with g gp p[unfolded prime_def] have False by blast}
-thus ?thesis by blast
-qed
-
-lemma even_dvd[simp]: "even (n::nat) \<longleftrightarrow> 2 dvd n" by presburger
-lemma prime_odd: "prime p \<Longrightarrow> p = 2 \<or> odd p" unfolding prime_def by auto
-
-
-text {* One property of coprimality is easier to prove via prime factors. *}
-
-lemma prime_divprod_pow: 
-  assumes p: "prime p" and ab: "coprime a b" and pab: "p^n dvd a * b"
-  shows "p^n dvd a \<or> p^n dvd b"
-proof-
-  {assume "n = 0 \<or> a = 1 \<or> b = 1" with pab have ?thesis 
-      apply (cases "n=0", simp_all)
-      apply (cases "a=1", simp_all) done}
-  moreover
-  {assume n: "n \<noteq> 0" and a: "a\<noteq>1" and b: "b\<noteq>1" 
-    then obtain m where m: "n = Suc m" by (cases n, auto)
-    from divides_exp2[OF n pab] have pab': "p dvd a*b" .
-    from prime_divprod[OF p pab'] 
-    have "p dvd a \<or> p dvd b" .
-    moreover
-    {assume pa: "p dvd a"
-      have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)
-      from coprime_prime[OF ab, of p] p pa have "\<not> p dvd b" by blast
-      with prime_coprime[OF p, of b] b 
-      have cpb: "coprime b p" using coprime_commute by blast 
-      from coprime_exp[OF cpb] have pnb: "coprime (p^n) b" 
-	by (simp add: coprime_commute)
-      from coprime_divprod[OF pnba pnb] have ?thesis by blast }
-    moreover
-    {assume pb: "p dvd b"
-      have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)
-      from coprime_prime[OF ab, of p] p pb have "\<not> p dvd a" by blast
-      with prime_coprime[OF p, of a] a
-      have cpb: "coprime a p" using coprime_commute by blast 
-      from coprime_exp[OF cpb] have pnb: "coprime (p^n) a" 
-	by (simp add: coprime_commute)
-      from coprime_divprod[OF pab pnb] have ?thesis by blast }
-    ultimately have ?thesis by blast}
-  ultimately show ?thesis by blast
-qed
-
-lemma nat_mult_eq_one: "(n::nat) * m = 1 \<longleftrightarrow> n = 1 \<and> m = 1" (is "?lhs \<longleftrightarrow> ?rhs")
-proof
-  assume H: "?lhs"
-  hence "n dvd 1" "m dvd 1" unfolding dvd_def by (auto simp add: mult_commute)
-  thus ?rhs by auto
-next
-  assume ?rhs then show ?lhs by auto
-qed
-  
-lemma power_Suc0[simp]: "Suc 0 ^ n = Suc 0" 
-  unfolding One_nat_def[symmetric] power_one ..
-lemma coprime_pow: assumes ab: "coprime a b" and abcn: "a * b = c ^n"
-  shows "\<exists>r s. a = r^n  \<and> b = s ^n"
-  using ab abcn
-proof(induct c arbitrary: a b rule: nat_less_induct)
-  fix c a b
-  assume H: "\<forall>m<c. \<forall>a b. coprime a b \<longrightarrow> a * b = m ^ n \<longrightarrow> (\<exists>r s. a = r ^ n \<and> b = s ^ n)" "coprime a b" "a * b = c ^ n" 
-  let ?ths = "\<exists>r s. a = r^n  \<and> b = s ^n"
-  {assume n: "n = 0"
-    with H(3) power_one have "a*b = 1" by simp
-    hence "a = 1 \<and> b = 1" by simp
-    hence ?ths 
-      apply -
-      apply (rule exI[where x=1])
-      apply (rule exI[where x=1])
-      using power_one[of  n]
-      by simp}
-  moreover
-  {assume n: "n \<noteq> 0" then obtain m where m: "n = Suc m" by (cases n, auto)
-    {assume c: "c = 0"
-      with H(3) m H(2) have ?ths apply simp 
-	apply (cases "a=0", simp_all) 
-	apply (rule exI[where x="0"], simp)
-	apply (rule exI[where x="0"], simp)
-	done}
-    moreover
-    {assume "c=1" with H(3) power_one have "a*b = 1" by simp 
-	hence "a = 1 \<and> b = 1" by simp
-	hence ?ths 
-      apply -
-      apply (rule exI[where x=1])
-      apply (rule exI[where x=1])
-      using power_one[of  n]
-      by simp}
-  moreover
-  {assume c: "c\<noteq>1" "c \<noteq> 0"
-    from prime_factor[OF c(1)] obtain p where p: "prime p" "p dvd c" by blast
-    from prime_divprod_pow[OF p(1) H(2), unfolded H(3), OF divides_exp[OF p(2), of n]] 
-    have pnab: "p ^ n dvd a \<or> p^n dvd b" . 
-    from p(2) obtain l where l: "c = p*l" unfolding dvd_def by blast
-    have pn0: "p^n \<noteq> 0" using n prime_ge_2 [OF p(1)] by (simp add: neq0_conv)
-    {assume pa: "p^n dvd a"
-      then obtain k where k: "a = p^n * k" unfolding dvd_def by blast
-      from l have "l dvd c" by auto
-      with dvd_imp_le[of l c] c have "l \<le> c" by auto
-      moreover {assume "l = c" with l c  have "p = 1" by simp with p have False by simp}
-      ultimately have lc: "l < c" by arith
-      from coprime_lmul2 [OF H(2)[unfolded k coprime_commute[of "p^n*k" b]]]
-      have kb: "coprime k b" by (simp add: coprime_commute) 
-      from H(3) l k pn0 have kbln: "k * b = l ^ n" 
-	by (auto simp add: power_mult_distrib)
-      from H(1)[rule_format, OF lc kb kbln]
-      obtain r s where rs: "k = r ^n" "b = s^n" by blast
-      from k rs(1) have "a = (p*r)^n" by (simp add: power_mult_distrib)
-      with rs(2) have ?ths by blast }
-    moreover
-    {assume pb: "p^n dvd b"
-      then obtain k where k: "b = p^n * k" unfolding dvd_def by blast
-      from l have "l dvd c" by auto
-      with dvd_imp_le[of l c] c have "l \<le> c" by auto
-      moreover {assume "l = c" with l c  have "p = 1" by simp with p have False by simp}
-      ultimately have lc: "l < c" by arith
-      from coprime_lmul2 [OF H(2)[unfolded k coprime_commute[of "p^n*k" a]]]
-      have kb: "coprime k a" by (simp add: coprime_commute) 
-      from H(3) l k pn0 n have kbln: "k * a = l ^ n" 
-	by (simp add: power_mult_distrib mult_commute)
-      from H(1)[rule_format, OF lc kb kbln]
-      obtain r s where rs: "k = r ^n" "a = s^n" by blast
-      from k rs(1) have "b = (p*r)^n" by (simp add: power_mult_distrib)
-      with rs(2) have ?ths by blast }
-    ultimately have ?ths using pnab by blast}
-  ultimately have ?ths by blast}
-ultimately show ?ths by blast
-qed
-
-text {* More useful lemmas. *}
-lemma prime_product: 
-  assumes "prime (p * q)"
-  shows "p = 1 \<or> q = 1"
-proof -
-  from assms have 
-    "1 < p * q" and P: "\<And>m. m dvd p * q \<Longrightarrow> m = 1 \<or> m = p * q"
-    unfolding prime_def by auto
-  from `1 < p * q` have "p \<noteq> 0" by (cases p) auto
-  then have Q: "p = p * q \<longleftrightarrow> q = 1" by auto
-  have "p dvd p * q" by simp
-  then have "p = 1 \<or> p = p * q" by (rule P)
-  then show ?thesis by (simp add: Q)
-qed
-
-lemma prime_exp: "prime (p^n) \<longleftrightarrow> prime p \<and> n = 1"
-proof(induct n)
-  case 0 thus ?case by simp
-next
-  case (Suc n)
-  {assume "p = 0" hence ?case by simp}
-  moreover
-  {assume "p=1" hence ?case by simp}
-  moreover
-  {assume p: "p \<noteq> 0" "p\<noteq>1"
-    {assume pp: "prime (p^Suc n)"
-      hence "p = 1 \<or> p^n = 1" using prime_product[of p "p^n"] by simp
-      with p have n: "n = 0" 
-	by (simp only: exp_eq_1 ) simp
-      with pp have "prime p \<and> Suc n = 1" by simp}
-    moreover
-    {assume n: "prime p \<and> Suc n = 1" hence "prime (p^Suc n)" by simp}
-    ultimately have ?case by blast}
-  ultimately show ?case by blast
-qed
-
-lemma prime_power_mult: 
-  assumes p: "prime p" and xy: "x * y = p ^ k"
-  shows "\<exists>i j. x = p ^i \<and> y = p^ j"
-  using xy
-proof(induct k arbitrary: x y)
-  case 0 thus ?case apply simp by (rule exI[where x="0"], simp)
-next
-  case (Suc k x y)
-  from Suc.prems have pxy: "p dvd x*y" by auto
-  from prime_divprod[OF p pxy] have pxyc: "p dvd x \<or> p dvd y" .
-  from p have p0: "p \<noteq> 0" by - (rule ccontr, simp) 
-  {assume px: "p dvd x"
-    then obtain d where d: "x = p*d" unfolding dvd_def by blast
-    from Suc.prems d  have "p*d*y = p^Suc k" by simp
-    hence th: "d*y = p^k" using p0 by simp
-    from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "y = p^j" by blast
-    with d have "x = p^Suc i" by simp 
-    with ij(2) have ?case by blast}
-  moreover 
-  {assume px: "p dvd y"
-    then obtain d where d: "y = p*d" unfolding dvd_def by blast
-    from Suc.prems d  have "p*d*x = p^Suc k" by (simp add: mult_commute)
-    hence th: "d*x = p^k" using p0 by simp
-    from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "x = p^j" by blast
-    with d have "y = p^Suc i" by simp 
-    with ij(2) have ?case by blast}
-  ultimately show ?case  using pxyc by blast
-qed
-
-lemma prime_power_exp: assumes p: "prime p" and n:"n \<noteq> 0" 
-  and xn: "x^n = p^k" shows "\<exists>i. x = p^i"
-  using n xn
-proof(induct n arbitrary: k)
-  case 0 thus ?case by simp
-next
-  case (Suc n k) hence th: "x*x^n = p^k" by simp
-  {assume "n = 0" with prems have ?case apply simp 
-      by (rule exI[where x="k"],simp)}
-  moreover
-  {assume n: "n \<noteq> 0"
-    from prime_power_mult[OF p th] 
-    obtain i j where ij: "x = p^i" "x^n = p^j"by blast
-    from Suc.hyps[OF n ij(2)] have ?case .}
-  ultimately show ?case by blast
-qed
-
-lemma divides_primepow: assumes p: "prime p" 
-  shows "d dvd p^k \<longleftrightarrow> (\<exists> i. i \<le> k \<and> d = p ^i)"
-proof
-  assume H: "d dvd p^k" then obtain e where e: "d*e = p^k" 
-    unfolding dvd_def  apply (auto simp add: mult_commute) by blast
-  from prime_power_mult[OF p e] obtain i j where ij: "d = p^i" "e=p^j" by blast
-  from prime_ge_2[OF p] have p1: "p > 1" by arith
-  from e ij have "p^(i + j) = p^k" by (simp add: power_add)
-  hence "i + j = k" using power_inject_exp[of p "i+j" k, OF p1] by simp 
-  hence "i \<le> k" by arith
-  with ij(1) show "\<exists>i\<le>k. d = p ^ i" by blast
-next
-  {fix i assume H: "i \<le> k" "d = p^i"
-    hence "\<exists>j. k = i + j" by arith
-    then obtain j where j: "k = i + j" by blast
-    hence "p^k = p^j*d" using H(2) by (simp add: power_add)
-    hence "d dvd p^k" unfolding dvd_def by auto}
-  thus "\<exists>i\<le>k. d = p ^ i \<Longrightarrow> d dvd p ^ k" by blast
-qed
-
-lemma coprime_divisors: "d dvd a \<Longrightarrow> e dvd b \<Longrightarrow> coprime a b \<Longrightarrow> coprime d e"
-  by (auto simp add: dvd_def coprime)
-
-declare power_Suc0[simp del]
-declare even_dvd[simp del]
-
-end
--- a/src/HOL/Library/Sum_Of_Squares.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Library/Sum_Of_Squares.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -38,13 +38,16 @@
 (*
 lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x \<Longrightarrow> a < 0" by sos
 
-lemma "a1 >= 0 & a2 >= 0 \<and> (a1 * a1 + a2 * a2 = b1 * b1 + b2 * b2 + 2) \<and> (a1 * b1 + a2 * b2 = 0) --> a1 * a2 - b1 * b2 >= (0::real)" by sos
+lemma "a1 >= 0 & a2 >= 0 \<and> (a1 * a1 + a2 * a2 = b1 * b1 + b2 * b2 + 2) \<and>
+  (a1 * b1 + a2 * b2 = 0) --> a1 * a2 - b1 * b2 >= (0::real)" by sos
 
 lemma "(3::real) * x + 7 * a < 4 & 3 < 2 * x --> a < 0" by sos
 
-lemma "(0::real) <= x & x <= 1 & 0 <= y & y <= 1  --> x^2 + y^2 < 1 |(x - 1)^2 + y^2 < 1 | x^2 + (y - 1)^2 < 1 | (x - 1)^2 + (y - 1)^2 < 1" by sos
+lemma "(0::real) <= x & x <= 1 & 0 <= y & y <= 1 -->
+  x^2 + y^2 < 1 |(x - 1)^2 + y^2 < 1 | x^2 + (y - 1)^2 < 1 | (x - 1)^2 + (y - 1)^2 < 1" by sos
 
-lemma "(0::real) <= x & 0 <= y & 0 <= z & x + y + z <= 3 --> x * y + x * z + y * z >= 3 * x * y * z" by sos
+lemma "(0::real) <= x & 0 <= y & 0 <= z & x + y + z <= 3 -->
+  x * y + x * z + y * z >= 3 * x * y * z" by sos
 
 lemma "((x::real)^2 + y^2 + z^2 = 1) --> (x + y + z)^2 <= 3" by sos
 
@@ -55,30 +58,27 @@
 lemma "(x::real) > 1 & y > 1 --> x * y > x + y - 1" by sos; 
 
 lemma "abs(x) <= 1 --> abs(64 * x^7 - 112 * x^5 + 56 * x^3 - 7 * x) <= (1::real)" by sos  
-*)
-(* ------------------------------------------------------------------------- *)
-(* One component of denominator in dodecahedral example.                     *)
-(* ------------------------------------------------------------------------- *)
-(*
-lemma "2 <= x & x <= 125841 / 50000 & 2 <= y & y <= 125841 / 50000 & 2 <= z & z <= 125841 / 50000 --> 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z) >= (0::real)" by sos;
-*)
-(* ------------------------------------------------------------------------- *)
-(* Over a larger but simpler interval.                                       *)
-(* ------------------------------------------------------------------------- *)
-(*
-lemma "(2::real) <= x & x <= 4 & 2 <= y & y <= 4 & 2 <= z & z <= 4 --> 0 <= 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z)" by sos
-*)
-(* ------------------------------------------------------------------------- *)
-(* We can do 12. I think 12 is a sharp bound; see PP's certificate.          *)
-(* ------------------------------------------------------------------------- *)
-(*
-lemma "2 <= (x::real) & x <= 4 & 2 <= y & y <= 4 & 2 <= z & z <= 4 --> 12 <= 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z)" by sos
-*)
+
+
+text {* One component of denominator in dodecahedral example. *}
+
+lemma "2 <= x & x <= 125841 / 50000 & 2 <= y & y <= 125841 / 50000 & 2 <= z &
+  z <= 125841 / 50000 --> 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z) >= (0::real)" by sos
+
+
+text {* Over a larger but simpler interval. *}
 
-(* ------------------------------------------------------------------------- *)
-(* Inequality from sci.math (see "Leon-Sotelo, por favor").                  *)
-(* ------------------------------------------------------------------------- *)
-(*
+lemma "(2::real) <= x & x <= 4 & 2 <= y & y <= 4 & 2 <= z &
+  z <= 4 --> 0 <= 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z)" by sos
+
+text {* We can do 12. I think 12 is a sharp bound; see PP's certificate. *}
+
+lemma "2 <= (x::real) & x <= 4 & 2 <= y & y <= 4 & 2 <= z & z <= 4 -->
+  12 <= 2 * (x * z + x * y + y * z) - (x * x + y * y + z * z)" by sos
+
+
+text {* Inequality from sci.math (see "Leon-Sotelo, por favor"). *}
+
 lemma "0 <= (x::real) & 0 <= y & (x * y = 1) --> x + y <= x^2 + y^2" by sos 
 
 lemma "0 <= (x::real) & 0 <= y & (x * y = 1) --> x * y * (x + y) <= x^2 + y^2" by sos 
@@ -100,7 +100,6 @@
 lemma "(0::real) < 1 + (1 + x)^2 * (abs x)" by sos
 
 
-
 lemma "abs ((1::real) + x^2) = (1::real) + x^2" by sos
 lemma "(3::real) * x + 7 * a < 4 \<and> 3 < 2 * x \<longrightarrow> a < 0" by sos
 
@@ -110,25 +109,25 @@
 lemma "(b::real)^2 < 4 * a * c --> ~(a * x^2 + b * x + c = 0)" by sos
 lemma "((a::real) * x^2 + b * x + c = 0) --> b^2 >= 4 * a * c" by sos
 lemma "(0::real) <= b & 0 <= c & 0 <= x & 0 <= y & (x^2 = c) & (y^2 = a^2 * c + b) --> a * c <= y * x" by sos
-lemma "abs(x - z) <= e & abs(y - z) <= e & 0 <= u & 0 <= v & (u + v = 1) --> abs((u * x + v * y) - z) <= (e::real)" by sos
-*)
+lemma "abs(x - z) <= e & abs(y - z) <= e & 0 <= u & 0 <= v & (u + v = 1) -->
+  abs((u * x + v * y) - z) <= (e::real)" by sos
+
 (*
-lemma "((x::real) - y - 2 * x^4 = 0) & 0 <= x & x <= 2 & 0 <= y & y <= 3 --> y^2 - 7 * y - 12 * x + 17 >= 0" by sos *) (* Too hard?*)
-(*
+lemma "((x::real) - y - 2 * x^4 = 0) & 0 <= x & x <= 2 & 0 <= y & y <= 3 -->
+  y^2 - 7 * y - 12 * x + 17 >= 0" by sos  -- {* Too hard?*}
+*)
+
 lemma "(0::real) <= x --> (1 + x + x^2)/(1 + x^2) <= 1 + x"
-apply sos
-done
+  by sos
 
 lemma "(0::real) <= x --> 1 - x <= 1 / (1 + x + x^2)"
-apply sos
-done
+  by sos
 
 lemma "(x::real) <= 1 / 2 --> - x - 2 * x^2 <= - x / (1 - x)"
-apply sos
-done 
+  by sos
 
-lemma "4*r^2 = p^2 - 4*q & r >= (0::real) & x^2 + p*x + q = 0 --> 2*(x::real) = - p + 2*r | 2*x = -p - 2*r" by sos
+lemma "4*r^2 = p^2 - 4*q & r >= (0::real) & x^2 + p*x + q = 0 -->
+  2*(x::real) = - p + 2*r | 2*x = -p - 2*r" by sos
 *)
 
 end
-
--- a/src/HOL/Library/Sum_Of_Squares/neos_csdp_client	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Library/Sum_Of_Squares/neos_csdp_client	Fri Sep 11 15:57:16 2009 +1000
@@ -9,9 +9,10 @@
 NEOS_HOST="neos.mcs.anl.gov"
 NEOS_PORT=3332
 
+neos=xmlrpclib.Server("http://%s:%d" % (NEOS_HOST, NEOS_PORT))
+
 jobNumber = 0
 password = ""
-neos = None
 inputfile = None
 outputfile = None
 # interrupt handler
@@ -34,8 +35,6 @@
   sys.stderr.write("Usage: neos_csdp_client <input_filename> <output_filename>\n")
   sys.exit(19)
 
-neos=xmlrpclib.Server("http://%s:%d" % (NEOS_HOST, NEOS_PORT))
-
 xml_pre = "<document>\n<category>sdp</category>\n<solver>csdp</solver>\n<inputMethod>SPARSE_SDPA</inputMethod>\n<dat><![CDATA["
 xml_post = "]]></dat>\n</document>\n"
 xml = xml_pre
@@ -74,9 +73,9 @@
 if len(result) > 1:
   solution = result[1].strip()
   if solution != "":
-    output = open(sys.argv[2],"w")
-    output.write(solution)
-    output.close()
+    outputfile = open(sys.argv[2],"w")
+    outputfile.write(solution)
+    outputfile.close()
 
 # extract return code
 p = re.compile(r"^Error: Command exited with non-zero status (\d+)$", re.MULTILINE)
--- a/src/HOL/Library/Univ_Poly.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Library/Univ_Poly.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -820,37 +820,24 @@
     done
 qed
 
-lemma (in semiring_0) pnormalize_sing: "(pnormalize [x] = [x]) \<longleftrightarrow> x \<noteq> 0" by simp
+lemma (in semiring_0) pnormalize_sing: "(pnormalize [x] = [x]) \<longleftrightarrow> x \<noteq> 0"
+by simp
 lemma (in semiring_0) pnormalize_pair: "y \<noteq> 0 \<longleftrightarrow> (pnormalize [x, y] = [x, y])" by simp
 lemma (in semiring_0) pnormal_cons: "pnormal p \<Longrightarrow> pnormal (c#p)"
   unfolding pnormal_def by simp
 lemma (in semiring_0) pnormal_tail: "p\<noteq>[] \<Longrightarrow> pnormal (c#p) \<Longrightarrow> pnormal p"
-  unfolding pnormal_def
-  apply (cases "pnormalize p = []", auto)
-  by (cases "c = 0", auto)
+  unfolding pnormal_def by(auto split: split_if_asm)
 
 
 lemma (in semiring_0) pnormal_last_nonzero: "pnormal p ==> last p \<noteq> 0"
-proof(induct p)
-  case Nil thus ?case by (simp add: pnormal_def)
-next
-  case (Cons a as) thus ?case
-    apply (simp add: pnormal_def)
-    apply (cases "pnormalize as = []", simp_all)
-    apply (cases "as = []", simp_all)
-    apply (cases "a=0", simp_all)
-    apply (cases "a=0", simp_all)
-    done
-qed
+by(induct p) (simp_all add: pnormal_def split: split_if_asm)
 
 lemma (in semiring_0) pnormal_length: "pnormal p \<Longrightarrow> 0 < length p"
   unfolding pnormal_def length_greater_0_conv by blast
 
 lemma (in semiring_0) pnormal_last_length: "\<lbrakk>0 < length p ; last p \<noteq> 0\<rbrakk> \<Longrightarrow> pnormal p"
-  apply (induct p, auto)
-  apply (case_tac "p = []", auto)
-  apply (simp add: pnormal_def)
-  by (rule pnormal_cons, auto)
+by (induct p) (auto simp: pnormal_def  split: split_if_asm)
+
 
 lemma (in semiring_0) pnormal_id: "pnormal p \<longleftrightarrow> (0 < length p \<and> last p \<noteq> 0)"
   using pnormal_last_length pnormal_length pnormal_last_nonzero by blast
@@ -918,9 +905,7 @@
 qed
 
 lemma (in semiring_0) pnormalize_eq: "last p \<noteq> 0 \<Longrightarrow> pnormalize p = p"
-  apply (induct p, auto)
-  apply (case_tac p, auto)+
-  done
+by (induct p) (auto split: split_if_asm)
 
 lemma (in semiring_0) last_pnormalize: "pnormalize p \<noteq> [] \<Longrightarrow> last (pnormalize p) \<noteq> 0"
   by (induct p, auto)
--- a/src/HOL/Library/Word.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Library/Word.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -1008,12 +1008,7 @@
   fix xs
   assume "length (norm_signed (\<zero>#xs)) = Suc (length xs)"
   thus "norm_signed (\<zero>#xs) = \<zero>#xs"
-    apply (simp add: norm_signed_Cons)
-    apply safe
-    apply simp_all
-    apply (rule norm_unsigned_equal)
-    apply assumption
-    done
+    by (simp add: norm_signed_Cons norm_unsigned_equal split: split_if_asm)
 next
   fix xs
   assume "length (norm_signed (\<one>#xs)) = Suc (length xs)"
@@ -1519,9 +1514,7 @@
 proof -
   have "((2::int) ^ (length w1 - 1)) + (2 ^ (length w2 - 1)) \<le>
       2 ^ (max (length w1) (length w2) - 1) + 2 ^ (max (length w1) (length w2) - 1)"
-    apply (cases "length w1 \<le> length w2")
-    apply (auto simp add: max_def)
-    done
+    by (auto simp:max_def)
   also have "... = 2 ^ max (length w1) (length w2)"
   proof -
     from lw
@@ -2173,16 +2166,16 @@
     apply (subst bv_sliceI [where ?j = i and ?i = j and ?w = w and ?w1.0 = "w1 @ w2" and ?w2.0 = w3 and ?w3.0 = "w4 @ w5"])
     apply simp_all
     apply (rule w_def)
-    apply (simp add: w_defs min_def)
-    apply (simp add: w_defs min_def)
+    apply (simp add: w_defs)
+    apply (simp add: w_defs)
     apply (subst bv_sliceI [where ?j = k and ?i = l and ?w = w and ?w1.0 = w1 and ?w2.0 = "w2 @ w3 @ w4" and ?w3.0 = w5])
     apply simp_all
     apply (rule w_def)
-    apply (simp add: w_defs min_def)
-    apply (simp add: w_defs min_def)
+    apply (simp add: w_defs)
+    apply (simp add: w_defs)
     apply (subst bv_sliceI [where ?j = "i-k" and ?i = "j-k" and ?w = "w2 @ w3 @ w4" and ?w1.0 = w2 and ?w2.0 = w3 and ?w3.0 = w4])
     apply simp_all
-    apply (simp_all add: w_defs min_def)
+    apply (simp_all add: w_defs)
     done
 qed
 
--- a/src/HOL/Library/normarith.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Library/normarith.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -15,7 +15,7 @@
 structure NormArith : NORM_ARITH = 
 struct
 
- open Conv Thm Conv2;
+ open Conv Thm;
  val bool_eq = op = : bool *bool -> bool
   fun dest_ratconst t = case term_of t of
    Const(@{const_name divide}, _)$a$b => Rat.rat_of_quotient(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
@@ -322,6 +322,7 @@
      val ths = map_filter (fn (v,t) => if v =/ Rat.zero then NONE 
                                      else SOME(norm_cmul_rule v t))
                             (v ~~ nubs) 
+     fun end_itlist f xs = split_last xs |> uncurry (fold_rev f)
     in inequality_canon_rule ctxt (end_itlist norm_add_rule ths)
     end
    val ges' = map_filter (try compute_ineq) (fold_rev (append o consider) destfuns []) @
@@ -332,9 +333,9 @@
   in RealArith.real_linear_prover translator
         (map (fn t => instantiate ([(tv_n, ctyp_of_term t)],[]) pth_zero)
             zerodests,
-        map (fconv_rule (once_depth_conv (norm_canon_conv) then_conv
+        map (fconv_rule (try_conv (More_Conv.top_sweep_conv (K norm_canon_conv) ctxt) then_conv
                        arg_conv (arg_conv real_poly_conv))) ges',
-        map (fconv_rule (once_depth_conv (norm_canon_conv) then_conv 
+        map (fconv_rule (try_conv (More_Conv.top_sweep_conv (K norm_canon_conv) ctxt) then_conv 
                        arg_conv (arg_conv real_poly_conv))) gts)
   end
 in val real_vector_combo_prover = real_vector_combo_prover
@@ -353,6 +354,7 @@
   val ntms = fold_rev find_normedterms (map (dest_arg o concl) (ges @ gts)) []
   val lctab = vector_lincombs (map snd (filter (not o fst) ntms))
   val (fxns, ctxt') = Variable.variant_fixes (replicate (length lctab) "x") ctxt
+  fun instantiate_cterm' ty tms = Drule.cterm_rule (Drule.instantiate' ty tms)
   fun mk_norm t = capply (instantiate_cterm' [SOME (ctyp_of_term t)] [] @{cpat "norm :: (?'a :: real_normed_vector) => real"}) t
   fun mk_equals l r = capply (capply (instantiate_cterm' [SOME (ctyp_of_term l)] [] @{cpat "op == :: ?'a =>_"}) l) r
   val asl = map2 (fn (t,_) => fn n => assume (mk_equals (mk_norm t) (cterm_of (ProofContext.theory_of ctxt') (Free(n,@{typ real}))))) lctab fxns
--- a/src/HOL/Library/positivstellensatz.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Library/positivstellensatz.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -81,82 +81,6 @@
 structure Ctermfunc = FuncFun(type key = cterm val ord = (fn (s,t) => TermOrd.fast_term_ord(term_of s, term_of t)));
 
 structure Ratfunc = FuncFun(type key = Rat.rat val ord = Rat.ord);
-    (* Some conversions-related stuff which has been forbidden entrance into Pure/conv.ML*)
-structure Conv2 = 
-struct
- open Conv
-fun instantiate_cterm' ty tms = Drule.cterm_rule (Drule.instantiate' ty tms)
-fun is_comb t = case (term_of t) of _$_ => true | _ => false;
-fun is_abs t = case (term_of t) of Abs _ => true | _ => false;
-
-fun end_itlist f l =
- case l of 
-   []     => error "end_itlist"
- | [x]    => x
- | (h::t) => f h (end_itlist f t);
-
- fun absc cv ct = case term_of ct of 
- Abs (v,_, _) => 
-  let val (x,t) = Thm.dest_abs (SOME v) ct
-  in Thm.abstract_rule ((fst o dest_Free o term_of) x) x (cv t)
-  end
- | _ => all_conv ct;
-
-fun cache_conv conv =
- let 
-  val tab = ref Termtab.empty
-  fun cconv t =  
-    case Termtab.lookup (!tab) (term_of t) of
-     SOME th => th
-   | NONE => let val th = conv t
-             in ((tab := Termtab.insert Thm.eq_thm (term_of t, th) (!tab)); th) end
- in cconv end;
-fun is_binop ct ct' = ct aconvc (Thm.dest_fun (Thm.dest_fun ct'))
-  handle CTERM _ => false;
-
-local
- fun thenqc conv1 conv2 tm =
-   case try conv1 tm of
-    SOME th1 => (case try conv2 (Thm.rhs_of th1) of SOME th2 => Thm.transitive th1 th2 | NONE => th1)
-  | NONE => conv2 tm
-
- fun thencqc conv1 conv2 tm =
-    let val th1 = conv1 tm 
-    in (case try conv2 (Thm.rhs_of th1) of SOME th2 => Thm.transitive th1 th2 | NONE => th1)
-    end
- fun comb_qconv conv tm =
-   let val (l,r) = Thm.dest_comb tm 
-   in (case try conv l of 
-        SOME th1 => (case try conv r of SOME th2 => Thm.combination th1 th2 
-                                      | NONE => Drule.fun_cong_rule th1 r)
-      | NONE => Drule.arg_cong_rule l (conv r))
-   end
- fun repeatqc conv tm = thencqc conv (repeatqc conv) tm 
- fun sub_qconv conv tm =  if is_abs tm then absc conv tm else comb_qconv conv tm 
- fun once_depth_qconv conv tm =
-      (conv else_conv (sub_qconv (once_depth_qconv conv))) tm
- fun depth_qconv conv tm =
-    thenqc (sub_qconv (depth_qconv conv))
-           (repeatqc conv) tm
- fun redepth_qconv conv tm =
-    thenqc (sub_qconv (redepth_qconv conv))
-           (thencqc conv (redepth_qconv conv)) tm
- fun top_depth_qconv conv tm =
-    thenqc (repeatqc conv)
-           (thencqc (sub_qconv (top_depth_qconv conv))
-                    (thencqc conv (top_depth_qconv conv))) tm
- fun top_sweep_qconv conv tm =
-    thenqc (repeatqc conv)
-           (sub_qconv (top_sweep_qconv conv)) tm
-in 
-val (once_depth_conv, depth_conv, rdepth_conv, top_depth_conv, top_sweep_conv) = 
-  (fn c => try_conv (once_depth_qconv c),
-   fn c => try_conv (depth_qconv c),
-   fn c => try_conv (redepth_qconv c),
-   fn c => try_conv (top_depth_qconv c),
-   fn c => try_conv (top_sweep_qconv c));
-end;
-end;
 
 
     (* Some useful derived rules *)
@@ -373,15 +297,6 @@
 fun instantiate_cterm' ty tms = Drule.cterm_rule (Drule.instantiate' ty tms)
 fun is_comb t = case (term_of t) of _$_ => true | _ => false;
 
-fun cache_conv conv =
- let 
-  val tab = ref Termtab.empty
-  fun cconv t =  
-    case Termtab.lookup (!tab) (term_of t) of
-     SOME th => th
-   | NONE => let val th = conv t
-             in ((tab := Termtab.insert Thm.eq_thm (term_of t, th) (!tab)); th) end
- in cconv end;
 fun is_binop ct ct' = ct aconvc (Thm.dest_fun (Thm.dest_fun ct'))
   handle CTERM _ => false;
 
@@ -571,7 +486,7 @@
    val nnf_norm_conv' = 
      nnf_conv then_conv 
      literals_conv [@{term "op &"}, @{term "op |"}] [] 
-     (cache_conv 
+     (More_Conv.cache_conv 
        (first_conv [real_lt_conv, real_le_conv, 
                     real_eq_conv, real_not_lt_conv, 
                     real_not_le_conv, real_not_eq_conv, all_conv]))
--- a/src/HOL/Lim.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Lim.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -544,8 +544,7 @@
     case True thus ?thesis using `0 < s` by auto
   next
     case False hence "s / 2 \<ge> (x - b) / 2" by auto
-    from inf_absorb2[OF this, unfolded inf_real_def]
-    have "?x = (x + b) / 2" by auto
+    hence "?x = (x + b) / 2" by(simp add:field_simps)
     thus ?thesis using `b < x` by auto
   qed
   hence "0 \<le> f ?x" using all_le `?x < x` by auto
--- a/src/HOL/List.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/List.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -881,10 +881,8 @@
 lemma set_filter [simp]: "set (filter P xs) = {x. x : set xs \<and> P x}"
 by (induct xs) auto
 
-lemma set_upt [simp]: "set[i..<j] = {k. i \<le> k \<and> k < j}"
-apply (induct j, simp_all)
-apply (erule ssubst, auto)
-done
+lemma set_upt [simp]: "set[i..<j] = {i..<j}"
+by (induct j) (simp_all add: atLeastLessThanSuc)
 
 
 lemma split_list: "x : set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs"
@@ -2394,6 +2392,30 @@
         nth_Cons_number_of [simp] 
 
 
+subsubsection {* @{text upto}: interval-list on @{typ int} *}
+
+(* FIXME make upto tail recursive? *)
+
+function upto :: "int \<Rightarrow> int \<Rightarrow> int list" ("(1[_../_])") where
+"upto i j = (if i \<le> j then i # [i+1..j] else [])"
+by auto
+termination
+by(relation "measure(%(i::int,j). nat(j - i + 1))") auto
+
+declare upto.simps[code, simp del]
+
+lemmas upto_rec_number_of[simp] =
+  upto.simps[of "number_of m" "number_of n", standard]
+
+lemma upto_empty[simp]: "j < i \<Longrightarrow> [i..j] = []"
+by(simp add: upto.simps)
+
+lemma set_upto[simp]: "set[i..j] = {i..j}"
+apply(induct i j rule:upto.induct)
+apply(simp add: upto.simps simp_from_to)
+done
+
+
 subsubsection {* @{text "distinct"} and @{text remdups} *}
 
 lemma distinct_append [simp]:
@@ -2448,6 +2470,12 @@
 lemma distinct_upt[simp]: "distinct[i..<j]"
 by (induct j) auto
 
+lemma distinct_upto[simp]: "distinct[i..j]"
+apply(induct i j rule:upto.induct)
+apply(subst upto.simps)
+apply(simp)
+done
+
 lemma distinct_take[simp]: "distinct xs \<Longrightarrow> distinct (take i xs)"
 apply(induct xs arbitrary: i)
  apply simp
@@ -3091,6 +3119,12 @@
 lemma sorted_upt[simp]: "sorted[i..<j]"
 by (induct j) (simp_all add:sorted_append)
 
+lemma sorted_upto[simp]: "sorted[i..j]"
+apply(induct i j rule:upto.induct)
+apply(subst upto.simps)
+apply(simp add:sorted_Cons)
+done
+
 
 subsubsection {* @{text sorted_list_of_set} *}
 
@@ -3124,89 +3158,6 @@
 end
 
 
-subsubsection {* @{text upto}: the generic interval-list *}
-
-class finite_intvl_succ = linorder +
-fixes successor :: "'a \<Rightarrow> 'a"
-assumes finite_intvl: "finite{a..b}"
-and successor_incr: "a < successor a"
-and ord_discrete: "\<not>(\<exists>x. a < x & x < successor a)"
-
-context finite_intvl_succ
-begin
-
-definition
- upto :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list" ("(1[_../_])") where
-"upto i j == sorted_list_of_set {i..j}"
-
-lemma upto[simp]: "set[a..b] = {a..b} & sorted[a..b] & distinct[a..b]"
-by(simp add:upto_def finite_intvl)
-
-lemma insert_intvl: "i \<le> j \<Longrightarrow> insert i {successor i..j} = {i..j}"
-apply(insert successor_incr[of i])
-apply(auto simp: atLeastAtMost_def atLeast_def atMost_def)
-apply(metis ord_discrete less_le not_le)
-done
-
-lemma sorted_list_of_set_rec: "i \<le> j \<Longrightarrow>
-  sorted_list_of_set {i..j} = i # sorted_list_of_set {successor i..j}"
-apply(simp add:sorted_list_of_set_def upto_def)
-apply (rule the1_equality[OF finite_sorted_distinct_unique])
- apply (simp add:finite_intvl)
-apply(rule the1I2[OF finite_sorted_distinct_unique])
- apply (simp add:finite_intvl)
-apply (simp add: sorted_Cons insert_intvl Ball_def)
-apply (metis successor_incr leD less_imp_le order_trans)
-done
-
-lemma sorted_list_of_set_rec2: "i \<le> j \<Longrightarrow>
-  sorted_list_of_set {i..successor j} =
-  sorted_list_of_set {i..j} @ [successor j]"
-apply(simp add:sorted_list_of_set_def upto_def)
-apply (rule the1_equality[OF finite_sorted_distinct_unique])
- apply (simp add:finite_intvl)
-apply(rule the1I2[OF finite_sorted_distinct_unique])
- apply (simp add:finite_intvl)
-apply (simp add: sorted_append Ball_def expand_set_eq)
-apply(rule conjI)
-apply (metis eq_iff leD linear not_leE ord_discrete order_trans successor_incr)
-apply (metis leD linear order_trans successor_incr)
-done
-
-lemma upto_rec[code]: "[i..j] = (if i \<le> j then i # [successor i..j] else [])"
-by(simp add: upto_def sorted_list_of_set_rec)
-
-lemma upto_empty[simp]: "j < i \<Longrightarrow> [i..j] = []"
-by(simp add: upto_rec)
-
-lemma upto_rec2: "i \<le> j \<Longrightarrow> [i..successor j] = [i..j] @ [successor j]"
-by(simp add: upto_def sorted_list_of_set_rec2)
-
-end
-
-text{* The integers are an instance of the above class: *}
-
-instantiation int:: finite_intvl_succ
-begin
-
-definition
-successor_int_def[simp]: "successor = (%i\<Colon>int. i+1)"
-
-instance
-by intro_classes (simp_all add: successor_int_def)
-
-end
-
-text{* Now @{term"[i..j::int]"} is defined for integers. *}
-
-hide (open) const successor
-
-lemma upto_rec2_int: "(i::int) \<le> j \<Longrightarrow> [i..j+1] = [i..j] @ [j+1]"
-by(rule upto_rec2[where 'a = int, simplified successor_int_def])
-
-lemmas upto_rec_number_of_int[simp] = upto_rec[of "number_of m :: int" "number_of n", standard]
-
-
 subsubsection {* @{text lists}: the list-forming operator over sets *}
 
 inductive_set
@@ -3829,9 +3780,7 @@
   "{n<..<m} = set [Suc n..<m]"
 by auto
 
-lemma atLeastLessThan_upt [code_unfold]:
-  "{n..<m} = set [n..<m]"
-by auto
+lemmas atLeastLessThan_upt [code_unfold] = set_upt [symmetric]
 
 lemma greaterThanAtMost_upt [code_unfold]:
   "{n<..m} = set [Suc n..<Suc m]"
@@ -3880,12 +3829,123 @@
   "{i<..j::int} = set [i+1..j]"
 by auto
 
-lemma atLeastAtMost_upto [code_unfold]:
-  "{i..j::int} = set [i..j]"
-by auto
+lemmas atLeastAtMost_upto [code_unfold] = set_upto[symmetric]
 
 lemma setsum_set_upto_conv_listsum [code_unfold]:
   "setsum f (set [i..j::int]) = listsum (map f [i..j])"
 by (rule setsum_set_distinct_conv_listsum) simp
 
+
+text {* Optimized code for @{text"\<forall>i\<in>{a..b::int}"} and @{text"\<forall>n:{a..<b::nat}"}
+and similiarly for @{text"\<exists>"}. *}
+
+function all_from_to_nat :: "(nat \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
+"all_from_to_nat P i j =
+ (if i < j then if P i then all_from_to_nat P (i+1) j else False
+  else True)"
+by auto
+termination
+by (relation "measure(%(P,i,j). j - i)") auto
+
+declare all_from_to_nat.simps[simp del]
+
+lemma all_from_to_nat_iff_ball:
+  "all_from_to_nat P i j = (ALL n : {i ..< j}. P n)"
+proof(induct P i j rule:all_from_to_nat.induct)
+  case (1 P i j)
+  let ?yes = "i < j & P i"
+  show ?case
+  proof (cases)
+    assume ?yes
+    hence "all_from_to_nat P i j = (P i & all_from_to_nat P (i+1) j)"
+      by(simp add: all_from_to_nat.simps)
+    also have "... = (P i & (ALL n : {i+1 ..< j}. P n))" using `?yes` 1 by simp
+    also have "... = (ALL n : {i ..< j}. P n)" (is "?L = ?R")
+    proof
+      assume L: ?L
+      show ?R
+      proof clarify
+	fix n assume n: "n : {i..<j}"
+	show "P n"
+	proof cases
+	  assume "n = i" thus "P n" using L by simp
+	next
+	  assume "n ~= i"
+	  hence "i+1 <= n" using n by auto
+	  thus "P n" using L n by simp
+	qed
+      qed
+    next
+      assume R: ?R thus ?L using `?yes` 1 by auto
+    qed
+    finally show ?thesis .
+  next
+    assume "~?yes" thus ?thesis by(auto simp add: all_from_to_nat.simps)
+  qed
+qed
+
+
+lemma list_all_iff_all_from_to_nat[code_unfold]:
+  "list_all P [i..<j] = all_from_to_nat P i j"
+by(simp add: all_from_to_nat_iff_ball list_all_iff)
+
+lemma list_ex_iff_not_all_from_to_not_nat[code_unfold]:
+  "list_ex P [i..<j] = (~all_from_to_nat (%x. ~P x) i j)"
+by(simp add: all_from_to_nat_iff_ball list_ex_iff)
+
+
+function all_from_to_int :: "(int \<Rightarrow> bool) \<Rightarrow> int \<Rightarrow> int \<Rightarrow> bool" where
+"all_from_to_int P i j =
+ (if i <= j then if P i then all_from_to_int P (i+1) j else False
+  else True)"
+by auto
+termination
+by (relation "measure(%(P,i,j). nat(j - i + 1))") auto
+
+declare all_from_to_int.simps[simp del]
+
+lemma all_from_to_int_iff_ball:
+  "all_from_to_int P i j = (ALL n : {i .. j}. P n)"
+proof(induct P i j rule:all_from_to_int.induct)
+  case (1 P i j)
+  let ?yes = "i <= j & P i"
+  show ?case
+  proof (cases)
+    assume ?yes
+    hence "all_from_to_int P i j = (P i & all_from_to_int P (i+1) j)"
+      by(simp add: all_from_to_int.simps)
+    also have "... = (P i & (ALL n : {i+1 .. j}. P n))" using `?yes` 1 by simp
+    also have "... = (ALL n : {i .. j}. P n)" (is "?L = ?R")
+    proof
+      assume L: ?L
+      show ?R
+      proof clarify
+	fix n assume n: "n : {i..j}"
+	show "P n"
+	proof cases
+	  assume "n = i" thus "P n" using L by simp
+	next
+	  assume "n ~= i"
+	  hence "i+1 <= n" using n by auto
+	  thus "P n" using L n by simp
+	qed
+      qed
+    next
+      assume R: ?R thus ?L using `?yes` 1 by auto
+    qed
+    finally show ?thesis .
+  next
+    assume "~?yes" thus ?thesis by(auto simp add: all_from_to_int.simps)
+  qed
+qed
+
+lemma list_all_iff_all_from_to_int[code_unfold]:
+  "list_all P [i..j] = all_from_to_int P i j"
+by(simp add: all_from_to_int_iff_ball list_all_iff)
+
+lemma list_ex_iff_not_all_from_to_not_int[code_unfold]:
+  "list_ex P [i..j] = (~ all_from_to_int (%x. ~P x) i j)"
+by(simp add: all_from_to_int_iff_ball list_ex_iff)
+
+
 end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Matrix/ComputeFloat.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,568 @@
+(*  Title:  HOL/Tools/ComputeFloat.thy
+    Author: Steven Obua
+*)
+
+header {* Floating Point Representation of the Reals *}
+
+theory ComputeFloat
+imports Complex_Main
+uses "~~/src/Tools/float.ML" ("~~/src/HOL/Tools/float_arith.ML")
+begin
+
+definition
+  pow2 :: "int \<Rightarrow> real" where
+  "pow2 a = (if (0 <= a) then (2^(nat a)) else (inverse (2^(nat (-a)))))"
+
+definition
+  float :: "int * int \<Rightarrow> real" where
+  "float x = real (fst x) * pow2 (snd x)"
+
+lemma pow2_0[simp]: "pow2 0 = 1"
+by (simp add: pow2_def)
+
+lemma pow2_1[simp]: "pow2 1 = 2"
+by (simp add: pow2_def)
+
+lemma pow2_neg: "pow2 x = inverse (pow2 (-x))"
+by (simp add: pow2_def)
+
+lemma pow2_add1: "pow2 (1 + a) = 2 * (pow2 a)"
+proof -
+  have h: "! n. nat (2 + int n) - Suc 0 = nat (1 + int n)" by arith
+  have g: "! a b. a - -1 = a + (1::int)" by arith
+  have pos: "! n. pow2 (int n + 1) = 2 * pow2 (int n)"
+    apply (auto, induct_tac n)
+    apply (simp_all add: pow2_def)
+    apply (rule_tac m1="2" and n1="nat (2 + int na)" in ssubst[OF realpow_num_eq_if])
+    by (auto simp add: h)
+  show ?thesis
+  proof (induct a)
+    case (1 n)
+    from pos show ?case by (simp add: algebra_simps)
+  next
+    case (2 n)
+    show ?case
+      apply (auto)
+      apply (subst pow2_neg[of "- int n"])
+      apply (subst pow2_neg[of "-1 - int n"])
+      apply (auto simp add: g pos)
+      done
+  qed
+qed
+
+lemma pow2_add: "pow2 (a+b) = (pow2 a) * (pow2 b)"
+proof (induct b)
+  case (1 n)
+  show ?case
+  proof (induct n)
+    case 0
+    show ?case by simp
+  next
+    case (Suc m)
+    show ?case by (auto simp add: algebra_simps pow2_add1 prems)
+  qed
+next
+  case (2 n)
+  show ?case
+  proof (induct n)
+    case 0
+    show ?case
+      apply (auto)
+      apply (subst pow2_neg[of "a + -1"])
+      apply (subst pow2_neg[of "-1"])
+      apply (simp)
+      apply (insert pow2_add1[of "-a"])
+      apply (simp add: algebra_simps)
+      apply (subst pow2_neg[of "-a"])
+      apply (simp)
+      done
+    case (Suc m)
+    have a: "int m - (a + -2) =  1 + (int m - a + 1)" by arith
+    have b: "int m - -2 = 1 + (int m + 1)" by arith
+    show ?case
+      apply (auto)
+      apply (subst pow2_neg[of "a + (-2 - int m)"])
+      apply (subst pow2_neg[of "-2 - int m"])
+      apply (auto simp add: algebra_simps)
+      apply (subst a)
+      apply (subst b)
+      apply (simp only: pow2_add1)
+      apply (subst pow2_neg[of "int m - a + 1"])
+      apply (subst pow2_neg[of "int m + 1"])
+      apply auto
+      apply (insert prems)
+      apply (auto simp add: algebra_simps)
+      done
+  qed
+qed
+
+lemma "float (a, e) + float (b, e) = float (a + b, e)"
+by (simp add: float_def algebra_simps)
+
+definition
+  int_of_real :: "real \<Rightarrow> int" where
+  "int_of_real x = (SOME y. real y = x)"
+
+definition
+  real_is_int :: "real \<Rightarrow> bool" where
+  "real_is_int x = (EX (u::int). x = real u)"
+
+lemma real_is_int_def2: "real_is_int x = (x = real (int_of_real x))"
+by (auto simp add: real_is_int_def int_of_real_def)
+
+lemma float_transfer: "real_is_int ((real a)*(pow2 c)) \<Longrightarrow> float (a, b) = float (int_of_real ((real a)*(pow2 c)), b - c)"
+by (simp add: float_def real_is_int_def2 pow2_add[symmetric])
+
+lemma pow2_int: "pow2 (int c) = 2^c"
+by (simp add: pow2_def)
+
+lemma float_transfer_nat: "float (a, b) = float (a * 2^c, b - int c)"
+by (simp add: float_def pow2_int[symmetric] pow2_add[symmetric])
+
+lemma real_is_int_real[simp]: "real_is_int (real (x::int))"
+by (auto simp add: real_is_int_def int_of_real_def)
+
+lemma int_of_real_real[simp]: "int_of_real (real x) = x"
+by (simp add: int_of_real_def)
+
+lemma real_int_of_real[simp]: "real_is_int x \<Longrightarrow> real (int_of_real x) = x"
+by (auto simp add: int_of_real_def real_is_int_def)
+
+lemma real_is_int_add_int_of_real: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> (int_of_real (a+b)) = (int_of_real a) + (int_of_real b)"
+by (auto simp add: int_of_real_def real_is_int_def)
+
+lemma real_is_int_add[simp]: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> real_is_int (a+b)"
+apply (subst real_is_int_def2)
+apply (simp add: real_is_int_add_int_of_real real_int_of_real)
+done
+
+lemma int_of_real_sub: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> (int_of_real (a-b)) = (int_of_real a) - (int_of_real b)"
+by (auto simp add: int_of_real_def real_is_int_def)
+
+lemma real_is_int_sub[simp]: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> real_is_int (a-b)"
+apply (subst real_is_int_def2)
+apply (simp add: int_of_real_sub real_int_of_real)
+done
+
+lemma real_is_int_rep: "real_is_int x \<Longrightarrow> ?! (a::int). real a = x"
+by (auto simp add: real_is_int_def)
+
+lemma int_of_real_mult:
+  assumes "real_is_int a" "real_is_int b"
+  shows "(int_of_real (a*b)) = (int_of_real a) * (int_of_real b)"
+proof -
+  from prems have a: "?! (a'::int). real a' = a" by (rule_tac real_is_int_rep, auto)
+  from prems have b: "?! (b'::int). real b' = b" by (rule_tac real_is_int_rep, auto)
+  from a obtain a'::int where a':"a = real a'" by auto
+  from b obtain b'::int where b':"b = real b'" by auto
+  have r: "real a' * real b' = real (a' * b')" by auto
+  show ?thesis
+    apply (simp add: a' b')
+    apply (subst r)
+    apply (simp only: int_of_real_real)
+    done
+qed
+
+lemma real_is_int_mult[simp]: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> real_is_int (a*b)"
+apply (subst real_is_int_def2)
+apply (simp add: int_of_real_mult)
+done
+
+lemma real_is_int_0[simp]: "real_is_int (0::real)"
+by (simp add: real_is_int_def int_of_real_def)
+
+lemma real_is_int_1[simp]: "real_is_int (1::real)"
+proof -
+  have "real_is_int (1::real) = real_is_int(real (1::int))" by auto
+  also have "\<dots> = True" by (simp only: real_is_int_real)
+  ultimately show ?thesis by auto
+qed
+
+lemma real_is_int_n1: "real_is_int (-1::real)"
+proof -
+  have "real_is_int (-1::real) = real_is_int(real (-1::int))" by auto
+  also have "\<dots> = True" by (simp only: real_is_int_real)
+  ultimately show ?thesis by auto
+qed
+
+lemma real_is_int_number_of[simp]: "real_is_int ((number_of \<Colon> int \<Rightarrow> real) x)"
+proof -
+  have neg1: "real_is_int (-1::real)"
+  proof -
+    have "real_is_int (-1::real) = real_is_int(real (-1::int))" by auto
+    also have "\<dots> = True" by (simp only: real_is_int_real)
+    ultimately show ?thesis by auto
+  qed
+
+  {
+    fix x :: int
+    have "real_is_int ((number_of \<Colon> int \<Rightarrow> real) x)"
+      unfolding number_of_eq
+      apply (induct x)
+      apply (induct_tac n)
+      apply (simp)
+      apply (simp)
+      apply (induct_tac n)
+      apply (simp add: neg1)
+    proof -
+      fix n :: nat
+      assume rn: "(real_is_int (of_int (- (int (Suc n)))))"
+      have s: "-(int (Suc (Suc n))) = -1 + - (int (Suc n))" by simp
+      show "real_is_int (of_int (- (int (Suc (Suc n)))))"
+        apply (simp only: s of_int_add)
+        apply (rule real_is_int_add)
+        apply (simp add: neg1)
+        apply (simp only: rn)
+        done
+    qed
+  }
+  note Abs_Bin = this
+  {
+    fix x :: int
+    have "? u. x = u"
+      apply (rule exI[where x = "x"])
+      apply (simp)
+      done
+  }
+  then obtain u::int where "x = u" by auto
+  with Abs_Bin show ?thesis by auto
+qed
+
+lemma int_of_real_0[simp]: "int_of_real (0::real) = (0::int)"
+by (simp add: int_of_real_def)
+
+lemma int_of_real_1[simp]: "int_of_real (1::real) = (1::int)"
+proof -
+  have 1: "(1::real) = real (1::int)" by auto
+  show ?thesis by (simp only: 1 int_of_real_real)
+qed
+
+lemma int_of_real_number_of[simp]: "int_of_real (number_of b) = number_of b"
+proof -
+  have "real_is_int (number_of b)" by simp
+  then have uu: "?! u::int. number_of b = real u" by (auto simp add: real_is_int_rep)
+  then obtain u::int where u:"number_of b = real u" by auto
+  have "number_of b = real ((number_of b)::int)"
+    by (simp add: number_of_eq real_of_int_def)
+  have ub: "number_of b = real ((number_of b)::int)"
+    by (simp add: number_of_eq real_of_int_def)
+  from uu u ub have unb: "u = number_of b"
+    by blast
+  have "int_of_real (number_of b) = u" by (simp add: u)
+  with unb show ?thesis by simp
+qed
+
+lemma float_transfer_even: "even a \<Longrightarrow> float (a, b) = float (a div 2, b+1)"
+  apply (subst float_transfer[where a="a" and b="b" and c="-1", simplified])
+  apply (simp_all add: pow2_def even_def real_is_int_def algebra_simps)
+  apply (auto)
+proof -
+  fix q::int
+  have a:"b - (-1\<Colon>int) = (1\<Colon>int) + b" by arith
+  show "(float (q, (b - (-1\<Colon>int)))) = (float (q, ((1\<Colon>int) + b)))"
+    by (simp add: a)
+qed
+
+lemma int_div_zdiv: "int (a div b) = (int a) div (int b)"
+by (rule zdiv_int)
+
+lemma int_mod_zmod: "int (a mod b) = (int a) mod (int b)"
+by (rule zmod_int)
+
+lemma abs_div_2_less: "a \<noteq> 0 \<Longrightarrow> a \<noteq> -1 \<Longrightarrow> abs((a::int) div 2) < abs a"
+by arith
+
+function norm_float :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
+  "norm_float a b = (if a \<noteq> 0 \<and> even a then norm_float (a div 2) (b + 1)
+    else if a = 0 then (0, 0) else (a, b))"
+by auto
+
+termination by (relation "measure (nat o abs o fst)")
+  (auto intro: abs_div_2_less)
+
+lemma norm_float: "float x = float (split norm_float x)"
+proof -
+  {
+    fix a b :: int
+    have norm_float_pair: "float (a, b) = float (norm_float a b)"
+    proof (induct a b rule: norm_float.induct)
+      case (1 u v)
+      show ?case
+      proof cases
+        assume u: "u \<noteq> 0 \<and> even u"
+        with prems have ind: "float (u div 2, v + 1) = float (norm_float (u div 2) (v + 1))" by auto
+        with u have "float (u,v) = float (u div 2, v+1)" by (simp add: float_transfer_even)
+        then show ?thesis
+          apply (subst norm_float.simps)
+          apply (simp add: ind)
+          done
+      next
+        assume "~(u \<noteq> 0 \<and> even u)"
+        then show ?thesis
+          by (simp add: prems float_def)
+      qed
+    qed
+  }
+  note helper = this
+  have "? a b. x = (a,b)" by auto
+  then obtain a b where "x = (a, b)" by blast
+  then show ?thesis by (simp add: helper)
+qed
+
+lemma float_add_l0: "float (0, e) + x = x"
+  by (simp add: float_def)
+
+lemma float_add_r0: "x + float (0, e) = x"
+  by (simp add: float_def)
+
+lemma float_add:
+  "float (a1, e1) + float (a2, e2) =
+  (if e1<=e2 then float (a1+a2*2^(nat(e2-e1)), e1)
+  else float (a1*2^(nat (e1-e2))+a2, e2))"
+  apply (simp add: float_def algebra_simps)
+  apply (auto simp add: pow2_int[symmetric] pow2_add[symmetric])
+  done
+
+lemma float_add_assoc1:
+  "(x + float (y1, e1)) + float (y2, e2) = (float (y1, e1) + float (y2, e2)) + x"
+  by simp
+
+lemma float_add_assoc2:
+  "(float (y1, e1) + x) + float (y2, e2) = (float (y1, e1) + float (y2, e2)) + x"
+  by simp
+
+lemma float_add_assoc3:
+  "float (y1, e1) + (x + float (y2, e2)) = (float (y1, e1) + float (y2, e2)) + x"
+  by simp
+
+lemma float_add_assoc4:
+  "float (y1, e1) + (float (y2, e2) + x) = (float (y1, e1) + float (y2, e2)) + x"
+  by simp
+
+lemma float_mult_l0: "float (0, e) * x = float (0, 0)"
+  by (simp add: float_def)
+
+lemma float_mult_r0: "x * float (0, e) = float (0, 0)"
+  by (simp add: float_def)
+
+definition 
+  lbound :: "real \<Rightarrow> real"
+where
+  "lbound x = min 0 x"
+
+definition
+  ubound :: "real \<Rightarrow> real"
+where
+  "ubound x = max 0 x"
+
+lemma lbound: "lbound x \<le> x"   
+  by (simp add: lbound_def)
+
+lemma ubound: "x \<le> ubound x"
+  by (simp add: ubound_def)
+
+lemma float_mult:
+  "float (a1, e1) * float (a2, e2) =
+  (float (a1 * a2, e1 + e2))"
+  by (simp add: float_def pow2_add)
+
+lemma float_minus:
+  "- (float (a,b)) = float (-a, b)"
+  by (simp add: float_def)
+
+lemma zero_less_pow2:
+  "0 < pow2 x"
+proof -
+  {
+    fix y
+    have "0 <= y \<Longrightarrow> 0 < pow2 y"
+      by (induct y, induct_tac n, simp_all add: pow2_add)
+  }
+  note helper=this
+  show ?thesis
+    apply (case_tac "0 <= x")
+    apply (simp add: helper)
+    apply (subst pow2_neg)
+    apply (simp add: helper)
+    done
+qed
+
+lemma zero_le_float:
+  "(0 <= float (a,b)) = (0 <= a)"
+  apply (auto simp add: float_def)
+  apply (auto simp add: zero_le_mult_iff zero_less_pow2)
+  apply (insert zero_less_pow2[of b])
+  apply (simp_all)
+  done
+
+lemma float_le_zero:
+  "(float (a,b) <= 0) = (a <= 0)"
+  apply (auto simp add: float_def)
+  apply (auto simp add: mult_le_0_iff)
+  apply (insert zero_less_pow2[of b])
+  apply auto
+  done
+
+lemma float_abs:
+  "abs (float (a,b)) = (if 0 <= a then (float (a,b)) else (float (-a,b)))"
+  apply (auto simp add: abs_if)
+  apply (simp_all add: zero_le_float[symmetric, of a b] float_minus)
+  done
+
+lemma float_zero:
+  "float (0, b) = 0"
+  by (simp add: float_def)
+
+lemma float_pprt:
+  "pprt (float (a, b)) = (if 0 <= a then (float (a,b)) else (float (0, b)))"
+  by (auto simp add: zero_le_float float_le_zero float_zero)
+
+lemma pprt_lbound: "pprt (lbound x) = float (0, 0)"
+  apply (simp add: float_def)
+  apply (rule pprt_eq_0)
+  apply (simp add: lbound_def)
+  done
+
+lemma nprt_ubound: "nprt (ubound x) = float (0, 0)"
+  apply (simp add: float_def)
+  apply (rule nprt_eq_0)
+  apply (simp add: ubound_def)
+  done
+
+lemma float_nprt:
+  "nprt (float (a, b)) = (if 0 <= a then (float (0,b)) else (float (a, b)))"
+  by (auto simp add: zero_le_float float_le_zero float_zero)
+
+lemma norm_0_1: "(0::_::number_ring) = Numeral0 & (1::_::number_ring) = Numeral1"
+  by auto
+
+lemma add_left_zero: "0 + a = (a::'a::comm_monoid_add)"
+  by simp
+
+lemma add_right_zero: "a + 0 = (a::'a::comm_monoid_add)"
+  by simp
+
+lemma mult_left_one: "1 * a = (a::'a::semiring_1)"
+  by simp
+
+lemma mult_right_one: "a * 1 = (a::'a::semiring_1)"
+  by simp
+
+lemma int_pow_0: "(a::int)^(Numeral0) = 1"
+  by simp
+
+lemma int_pow_1: "(a::int)^(Numeral1) = a"
+  by simp
+
+lemma zero_eq_Numeral0_nring: "(0::'a::number_ring) = Numeral0"
+  by simp
+
+lemma one_eq_Numeral1_nring: "(1::'a::number_ring) = Numeral1"
+  by simp
+
+lemma zero_eq_Numeral0_nat: "(0::nat) = Numeral0"
+  by simp
+
+lemma one_eq_Numeral1_nat: "(1::nat) = Numeral1"
+  by simp
+
+lemma zpower_Pls: "(z::int)^Numeral0 = Numeral1"
+  by simp
+
+lemma zpower_Min: "(z::int)^((-1)::nat) = Numeral1"
+proof -
+  have 1:"((-1)::nat) = 0"
+    by simp
+  show ?thesis by (simp add: 1)
+qed
+
+lemma fst_cong: "a=a' \<Longrightarrow> fst (a,b) = fst (a',b)"
+  by simp
+
+lemma snd_cong: "b=b' \<Longrightarrow> snd (a,b) = snd (a,b')"
+  by simp
+
+lemma lift_bool: "x \<Longrightarrow> x=True"
+  by simp
+
+lemma nlift_bool: "~x \<Longrightarrow> x=False"
+  by simp
+
+lemma not_false_eq_true: "(~ False) = True" by simp
+
+lemma not_true_eq_false: "(~ True) = False" by simp
+
+lemmas binarith =
+  normalize_bin_simps
+  pred_bin_simps succ_bin_simps
+  add_bin_simps minus_bin_simps mult_bin_simps
+
+lemma int_eq_number_of_eq:
+  "(((number_of v)::int)=(number_of w)) = iszero ((number_of (v + uminus w))::int)"
+  by (rule eq_number_of_eq)
+
+lemma int_iszero_number_of_Pls: "iszero (Numeral0::int)"
+  by (simp only: iszero_number_of_Pls)
+
+lemma int_nonzero_number_of_Min: "~(iszero ((-1)::int))"
+  by simp
+
+lemma int_iszero_number_of_Bit0: "iszero ((number_of (Int.Bit0 w))::int) = iszero ((number_of w)::int)"
+  by simp
+
+lemma int_iszero_number_of_Bit1: "\<not> iszero ((number_of (Int.Bit1 w))::int)"
+  by simp
+
+lemma int_less_number_of_eq_neg: "(((number_of x)::int) < number_of y) = neg ((number_of (x + (uminus y)))::int)"
+  unfolding neg_def number_of_is_id by simp
+
+lemma int_not_neg_number_of_Pls: "\<not> (neg (Numeral0::int))"
+  by simp
+
+lemma int_neg_number_of_Min: "neg (-1::int)"
+  by simp
+
+lemma int_neg_number_of_Bit0: "neg ((number_of (Int.Bit0 w))::int) = neg ((number_of w)::int)"
+  by simp
+
+lemma int_neg_number_of_Bit1: "neg ((number_of (Int.Bit1 w))::int) = neg ((number_of w)::int)"
+  by simp
+
+lemma int_le_number_of_eq: "(((number_of x)::int) \<le> number_of y) = (\<not> neg ((number_of (y + (uminus x)))::int))"
+  unfolding neg_def number_of_is_id by (simp add: not_less)
+
+lemmas intarithrel =
+  int_eq_number_of_eq
+  lift_bool[OF int_iszero_number_of_Pls] nlift_bool[OF int_nonzero_number_of_Min] int_iszero_number_of_Bit0
+  lift_bool[OF int_iszero_number_of_Bit1] int_less_number_of_eq_neg nlift_bool[OF int_not_neg_number_of_Pls] lift_bool[OF int_neg_number_of_Min]
+  int_neg_number_of_Bit0 int_neg_number_of_Bit1 int_le_number_of_eq
+
+lemma int_number_of_add_sym: "((number_of v)::int) + number_of w = number_of (v + w)"
+  by simp
+
+lemma int_number_of_diff_sym: "((number_of v)::int) - number_of w = number_of (v + (uminus w))"
+  by simp
+
+lemma int_number_of_mult_sym: "((number_of v)::int) * number_of w = number_of (v * w)"
+  by simp
+
+lemma int_number_of_minus_sym: "- ((number_of v)::int) = number_of (uminus v)"
+  by simp
+
+lemmas intarith = int_number_of_add_sym int_number_of_minus_sym int_number_of_diff_sym int_number_of_mult_sym
+
+lemmas natarith = add_nat_number_of diff_nat_number_of mult_nat_number_of eq_nat_number_of less_nat_number_of
+
+lemmas powerarith = nat_number_of zpower_number_of_even
+  zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring]
+  zpower_Pls zpower_Min
+
+lemmas floatarith[simplified norm_0_1] = float_add float_add_l0 float_add_r0 float_mult float_mult_l0 float_mult_r0 
+          float_minus float_abs zero_le_float float_pprt float_nprt pprt_lbound nprt_ubound
+
+(* for use with the compute oracle *)
+lemmas arith = binarith intarith intarithrel natarith powerarith floatarith not_false_eq_true not_true_eq_false
+
+use "~~/src/HOL/Tools/float_arith.ML"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Matrix/ComputeHOL.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,191 @@
+theory ComputeHOL
+imports Complex_Main "~~/src/Tools/Compute_Oracle/Compute_Oracle"
+begin
+
+lemma Trueprop_eq_eq: "Trueprop X == (X == True)" by (simp add: atomize_eq)
+lemma meta_eq_trivial: "x == y \<Longrightarrow> x == y" by simp
+lemma meta_eq_imp_eq: "x == y \<Longrightarrow> x = y" by auto
+lemma eq_trivial: "x = y \<Longrightarrow> x = y" by auto
+lemma bool_to_true: "x :: bool \<Longrightarrow> x == True"  by simp
+lemma transmeta_1: "x = y \<Longrightarrow> y == z \<Longrightarrow> x = z" by simp
+lemma transmeta_2: "x == y \<Longrightarrow> y = z \<Longrightarrow> x = z" by simp
+lemma transmeta_3: "x == y \<Longrightarrow> y == z \<Longrightarrow> x = z" by simp
+
+
+(**** compute_if ****)
+
+lemma If_True: "If True = (\<lambda> x y. x)" by ((rule ext)+,auto)
+lemma If_False: "If False = (\<lambda> x y. y)" by ((rule ext)+, auto)
+
+lemmas compute_if = If_True If_False
+
+(**** compute_bool ****)
+
+lemma bool1: "(\<not> True) = False"  by blast
+lemma bool2: "(\<not> False) = True"  by blast
+lemma bool3: "(P \<and> True) = P" by blast
+lemma bool4: "(True \<and> P) = P" by blast
+lemma bool5: "(P \<and> False) = False" by blast
+lemma bool6: "(False \<and> P) = False" by blast
+lemma bool7: "(P \<or> True) = True" by blast
+lemma bool8: "(True \<or> P) = True" by blast
+lemma bool9: "(P \<or> False) = P" by blast
+lemma bool10: "(False \<or> P) = P" by blast
+lemma bool11: "(True \<longrightarrow> P) = P" by blast
+lemma bool12: "(P \<longrightarrow> True) = True" by blast
+lemma bool13: "(True \<longrightarrow> P) = P" by blast
+lemma bool14: "(P \<longrightarrow> False) = (\<not> P)" by blast
+lemma bool15: "(False \<longrightarrow> P) = True" by blast
+lemma bool16: "(False = False) = True" by blast
+lemma bool17: "(True = True) = True" by blast
+lemma bool18: "(False = True) = False" by blast
+lemma bool19: "(True = False) = False" by blast
+
+lemmas compute_bool = bool1 bool2 bool3 bool4 bool5 bool6 bool7 bool8 bool9 bool10 bool11 bool12 bool13 bool14 bool15 bool16 bool17 bool18 bool19
+
+
+(*** compute_pair ***)
+
+lemma compute_fst: "fst (x,y) = x" by simp
+lemma compute_snd: "snd (x,y) = y" by simp
+lemma compute_pair_eq: "((a, b) = (c, d)) = (a = c \<and> b = d)" by auto
+
+lemma prod_case_simp: "prod_case f (x,y) = f x y" by simp
+
+lemmas compute_pair = compute_fst compute_snd compute_pair_eq prod_case_simp
+
+(*** compute_option ***)
+
+lemma compute_the: "the (Some x) = x" by simp
+lemma compute_None_Some_eq: "(None = Some x) = False" by auto
+lemma compute_Some_None_eq: "(Some x = None) = False" by auto
+lemma compute_None_None_eq: "(None = None) = True" by auto
+lemma compute_Some_Some_eq: "(Some x = Some y) = (x = y)" by auto
+
+definition
+   option_case_compute :: "'b option \<Rightarrow> 'a \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a"
+where
+   "option_case_compute opt a f = option_case a f opt"
+
+lemma option_case_compute: "option_case = (\<lambda> a f opt. option_case_compute opt a f)"
+  by (simp add: option_case_compute_def)
+
+lemma option_case_compute_None: "option_case_compute None = (\<lambda> a f. a)"
+  apply (rule ext)+
+  apply (simp add: option_case_compute_def)
+  done
+
+lemma option_case_compute_Some: "option_case_compute (Some x) = (\<lambda> a f. f x)"
+  apply (rule ext)+
+  apply (simp add: option_case_compute_def)
+  done
+
+lemmas compute_option_case = option_case_compute option_case_compute_None option_case_compute_Some
+
+lemmas compute_option = compute_the compute_None_Some_eq compute_Some_None_eq compute_None_None_eq compute_Some_Some_eq compute_option_case
+
+(**** compute_list_length ****)
+
+lemma length_cons:"length (x#xs) = 1 + (length xs)"
+  by simp
+
+lemma length_nil: "length [] = 0"
+  by simp
+
+lemmas compute_list_length = length_nil length_cons
+
+(*** compute_list_case ***)
+
+definition
+  list_case_compute :: "'b list \<Rightarrow> 'a \<Rightarrow> ('b \<Rightarrow> 'b list \<Rightarrow> 'a) \<Rightarrow> 'a"
+where
+  "list_case_compute l a f = list_case a f l"
+
+lemma list_case_compute: "list_case = (\<lambda> (a::'a) f (l::'b list). list_case_compute l a f)"
+  apply (rule ext)+
+  apply (simp add: list_case_compute_def)
+  done
+
+lemma list_case_compute_empty: "list_case_compute ([]::'b list) = (\<lambda> (a::'a) f. a)"
+  apply (rule ext)+
+  apply (simp add: list_case_compute_def)
+  done
+
+lemma list_case_compute_cons: "list_case_compute (u#v) = (\<lambda> (a::'a) f. (f (u::'b) v))"
+  apply (rule ext)+
+  apply (simp add: list_case_compute_def)
+  done
+
+lemmas compute_list_case = list_case_compute list_case_compute_empty list_case_compute_cons
+
+(*** compute_list_nth ***)
+(* Of course, you will need computation with nats for this to work \<dots> *)
+
+lemma compute_list_nth: "((x#xs) ! n) = (if n = 0 then x else (xs ! (n - 1)))"
+  by (cases n, auto)
+  
+(*** compute_list ***)
+
+lemmas compute_list = compute_list_case compute_list_length compute_list_nth
+
+(*** compute_let ***)
+
+lemmas compute_let = Let_def
+
+(***********************)
+(* Everything together *)
+(***********************)
+
+lemmas compute_hol = compute_if compute_bool compute_pair compute_option compute_list compute_let
+
+ML {*
+signature ComputeHOL =
+sig
+  val prep_thms : thm list -> thm list
+  val to_meta_eq : thm -> thm
+  val to_hol_eq : thm -> thm
+  val symmetric : thm -> thm 
+  val trans : thm -> thm -> thm
+end
+
+structure ComputeHOL : ComputeHOL =
+struct
+
+local
+fun lhs_of eq = fst (Thm.dest_equals (cprop_of eq));
+in
+fun rewrite_conv [] ct = raise CTERM ("rewrite_conv", [])
+  | rewrite_conv (eq :: eqs) ct =
+      Thm.instantiate (Thm.match (lhs_of eq, ct)) eq
+      handle Pattern.MATCH => rewrite_conv eqs ct;
+end
+
+val convert_conditions = Conv.fconv_rule (Conv.prems_conv ~1 (Conv.try_conv (rewrite_conv [@{thm "Trueprop_eq_eq"}])))
+
+val eq_th = @{thm "HOL.eq_reflection"}
+val meta_eq_trivial = @{thm "ComputeHOL.meta_eq_trivial"}
+val bool_to_true = @{thm "ComputeHOL.bool_to_true"}
+
+fun to_meta_eq th = eq_th OF [th] handle THM _ => meta_eq_trivial OF [th] handle THM _ => bool_to_true OF [th]
+
+fun to_hol_eq th = @{thm "meta_eq_imp_eq"} OF [th] handle THM _ => @{thm "eq_trivial"} OF [th] 
+
+fun prep_thms ths = map (convert_conditions o to_meta_eq) ths
+
+fun symmetric th = @{thm "HOL.sym"} OF [th] handle THM _ => @{thm "Pure.symmetric"} OF [th]
+
+local
+    val trans_HOL = @{thm "HOL.trans"}
+    val trans_HOL_1 = @{thm "ComputeHOL.transmeta_1"}
+    val trans_HOL_2 = @{thm "ComputeHOL.transmeta_2"}
+    val trans_HOL_3 = @{thm "ComputeHOL.transmeta_3"}
+    fun tr [] th1 th2 = trans_HOL OF [th1, th2]
+      | tr (t::ts) th1 th2 = (t OF [th1, th2] handle THM _ => tr ts th1 th2) 
+in
+  fun trans th1 th2 = tr [trans_HOL, trans_HOL_1, trans_HOL_2, trans_HOL_3] th1 th2
+end
+
+end
+*}
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Matrix/ComputeNumeral.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,195 @@
+theory ComputeNumeral
+imports ComputeHOL ComputeFloat
+begin
+
+(* normalization of bit strings *)
+lemmas bitnorm = normalize_bin_simps
+
+(* neg for bit strings *)
+lemma neg1: "neg Int.Pls = False" by (simp add: Int.Pls_def)
+lemma neg2: "neg Int.Min = True" apply (subst Int.Min_def) by auto
+lemma neg3: "neg (Int.Bit0 x) = neg x" apply (simp add: neg_def) apply (subst Bit0_def) by auto
+lemma neg4: "neg (Int.Bit1 x) = neg x" apply (simp add: neg_def) apply (subst Bit1_def) by auto  
+lemmas bitneg = neg1 neg2 neg3 neg4
+
+(* iszero for bit strings *)
+lemma iszero1: "iszero Int.Pls = True" by (simp add: Int.Pls_def iszero_def)
+lemma iszero2: "iszero Int.Min = False" apply (subst Int.Min_def) apply (subst iszero_def) by simp
+lemma iszero3: "iszero (Int.Bit0 x) = iszero x" apply (subst Int.Bit0_def) apply (subst iszero_def)+ by auto
+lemma iszero4: "iszero (Int.Bit1 x) = False" apply (subst Int.Bit1_def) apply (subst iszero_def)+  apply simp by arith
+lemmas bitiszero = iszero1 iszero2 iszero3 iszero4
+
+(* lezero for bit strings *)
+constdefs
+  "lezero x == (x \<le> 0)"
+lemma lezero1: "lezero Int.Pls = True" unfolding Int.Pls_def lezero_def by auto
+lemma lezero2: "lezero Int.Min = True" unfolding Int.Min_def lezero_def by auto
+lemma lezero3: "lezero (Int.Bit0 x) = lezero x" unfolding Int.Bit0_def lezero_def by auto
+lemma lezero4: "lezero (Int.Bit1 x) = neg x" unfolding Int.Bit1_def lezero_def neg_def by auto
+lemmas bitlezero = lezero1 lezero2 lezero3 lezero4
+
+(* equality for bit strings *)
+lemmas biteq = eq_bin_simps
+
+(* x < y for bit strings *)
+lemmas bitless = less_bin_simps
+
+(* x \<le> y for bit strings *)
+lemmas bitle = le_bin_simps
+
+(* succ for bit strings *)
+lemmas bitsucc = succ_bin_simps
+
+(* pred for bit strings *)
+lemmas bitpred = pred_bin_simps
+
+(* unary minus for bit strings *)
+lemmas bituminus = minus_bin_simps
+
+(* addition for bit strings *)
+lemmas bitadd = add_bin_simps
+
+(* multiplication for bit strings *) 
+lemma mult_Pls_right: "x * Int.Pls = Int.Pls" by (simp add: Pls_def)
+lemma mult_Min_right: "x * Int.Min = - x" by (subst mult_commute, simp add: mult_Min)
+lemma multb0x: "(Int.Bit0 x) * y = Int.Bit0 (x * y)" by (rule mult_Bit0)
+lemma multxb0: "x * (Int.Bit0 y) = Int.Bit0 (x * y)" unfolding Bit0_def by simp
+lemma multb1: "(Int.Bit1 x) * (Int.Bit1 y) = Int.Bit1 (Int.Bit0 (x * y) + x + y)"
+  unfolding Bit0_def Bit1_def by (simp add: algebra_simps)
+lemmas bitmul = mult_Pls mult_Min mult_Pls_right mult_Min_right multb0x multxb0 multb1
+
+lemmas bitarith = bitnorm bitiszero bitneg bitlezero biteq bitless bitle bitsucc bitpred bituminus bitadd bitmul 
+
+constdefs 
+  "nat_norm_number_of (x::nat) == x"
+
+lemma nat_norm_number_of: "nat_norm_number_of (number_of w) = (if lezero w then 0 else number_of w)"
+  apply (simp add: nat_norm_number_of_def)
+  unfolding lezero_def iszero_def neg_def
+  apply (simp add: numeral_simps)
+  done
+
+(* Normalization of nat literals *)
+lemma natnorm0: "(0::nat) = number_of (Int.Pls)" by auto
+lemma natnorm1: "(1 :: nat) = number_of (Int.Bit1 Int.Pls)"  by auto 
+lemmas natnorm = natnorm0 natnorm1 nat_norm_number_of
+
+(* Suc *)
+lemma natsuc: "Suc (number_of x) = (if neg x then 1 else number_of (Int.succ x))" by (auto simp add: number_of_is_id)
+
+(* Addition for nat *)
+lemma natadd: "number_of x + ((number_of y)::nat) = (if neg x then (number_of y) else (if neg y then number_of x else (number_of (x + y))))"
+  unfolding nat_number_of_def number_of_is_id neg_def
+  by auto
+
+(* Subtraction for nat *)
+lemma natsub: "(number_of x) - ((number_of y)::nat) = 
+  (if neg x then 0 else (if neg y then number_of x else (nat_norm_number_of (number_of (x + (- y))))))"
+  unfolding nat_norm_number_of
+  by (auto simp add: number_of_is_id neg_def lezero_def iszero_def Let_def nat_number_of_def)
+
+(* Multiplication for nat *)
+lemma natmul: "(number_of x) * ((number_of y)::nat) = 
+  (if neg x then 0 else (if neg y then 0 else number_of (x * y)))"
+  unfolding nat_number_of_def number_of_is_id neg_def
+  by (simp add: nat_mult_distrib)
+
+lemma nateq: "(((number_of x)::nat) = (number_of y)) = ((lezero x \<and> lezero y) \<or> (x = y))"
+  by (auto simp add: iszero_def lezero_def neg_def number_of_is_id)
+
+lemma natless: "(((number_of x)::nat) < (number_of y)) = ((x < y) \<and> (\<not> (lezero y)))"
+  by (simp add: lezero_def numeral_simps not_le)
+
+lemma natle: "(((number_of x)::nat) \<le> (number_of y)) = (y < x \<longrightarrow> lezero x)"
+  by (auto simp add: number_of_is_id lezero_def nat_number_of_def)
+
+fun natfac :: "nat \<Rightarrow> nat"
+where
+   "natfac n = (if n = 0 then 1 else n * (natfac (n - 1)))"
+
+lemmas compute_natarith = bitarith natnorm natsuc natadd natsub natmul nateq natless natle natfac.simps
+
+lemma number_eq: "(((number_of x)::'a::{number_ring, ordered_idom}) = (number_of y)) = (x = y)"
+  unfolding number_of_eq
+  apply simp
+  done
+
+lemma number_le: "(((number_of x)::'a::{number_ring, ordered_idom}) \<le>  (number_of y)) = (x \<le> y)"
+  unfolding number_of_eq
+  apply simp
+  done
+
+lemma number_less: "(((number_of x)::'a::{number_ring, ordered_idom}) <  (number_of y)) = (x < y)"
+  unfolding number_of_eq 
+  apply simp
+  done
+
+lemma number_diff: "((number_of x)::'a::{number_ring, ordered_idom}) - number_of y = number_of (x + (- y))"
+  apply (subst diff_number_of_eq)
+  apply simp
+  done
+
+lemmas number_norm = number_of_Pls[symmetric] numeral_1_eq_1[symmetric]
+
+lemmas compute_numberarith = number_of_minus[symmetric] number_of_add[symmetric] number_diff number_of_mult[symmetric] number_norm number_eq number_le number_less
+
+lemma compute_real_of_nat_number_of: "real ((number_of v)::nat) = (if neg v then 0 else number_of v)"
+  by (simp only: real_of_nat_number_of number_of_is_id)
+
+lemma compute_nat_of_int_number_of: "nat ((number_of v)::int) = (number_of v)"
+  by simp
+
+lemmas compute_num_conversions = compute_real_of_nat_number_of compute_nat_of_int_number_of real_number_of
+
+lemmas zpowerarith = zpower_number_of_even
+  zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring]
+  zpower_Pls zpower_Min
+
+(* div, mod *)
+
+lemma adjust: "adjust b (q, r) = (if 0 \<le> r - b then (2 * q + 1, r - b) else (2 * q, r))"
+  by (auto simp only: adjust_def)
+
+lemma negateSnd: "negateSnd (q, r) = (q, -r)" 
+  by (simp add: negateSnd_def)
+
+lemma divmod: "IntDiv.divmod a b = (if 0\<le>a then
+                  if 0\<le>b then posDivAlg a b
+                  else if a=0 then (0, 0)
+                       else negateSnd (negDivAlg (-a) (-b))
+               else 
+                  if 0<b then negDivAlg a b
+                  else negateSnd (posDivAlg (-a) (-b)))"
+  by (auto simp only: IntDiv.divmod_def)
+
+lemmas compute_div_mod = div_def mod_def divmod adjust negateSnd posDivAlg.simps negDivAlg.simps
+
+
+
+(* collecting all the theorems *)
+
+lemma even_Pls: "even (Int.Pls) = True"
+  apply (unfold Pls_def even_def)
+  by simp
+
+lemma even_Min: "even (Int.Min) = False"
+  apply (unfold Min_def even_def)
+  by simp
+
+lemma even_B0: "even (Int.Bit0 x) = True"
+  apply (unfold Bit0_def)
+  by simp
+
+lemma even_B1: "even (Int.Bit1 x) = False"
+  apply (unfold Bit1_def)
+  by simp
+
+lemma even_number_of: "even ((number_of w)::int) = even w"
+  by (simp only: number_of_is_id)
+
+lemmas compute_even = even_Pls even_Min even_B0 even_B1 even_number_of
+
+lemmas compute_numeral = compute_if compute_let compute_pair compute_bool 
+                         compute_natarith compute_numberarith max_def min_def compute_num_conversions zpowerarith compute_div_mod compute_even
+
+end
--- a/src/HOL/Matrix/LP.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Matrix/LP.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Matrix/LP.thy
-    ID:         $Id$
     Author:     Steven Obua
 *)
 
--- a/src/HOL/Matrix/Matrix.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Matrix/Matrix.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Matrix/Matrix.thy
-    ID:         $Id$
     Author:     Steven Obua
 *)
 
@@ -1663,7 +1662,7 @@
 apply (simp add: times_matrix_def Rep_mult_matrix)
 apply (rule_tac j1="xa" in ssubst[OF foldseq_almostzero])
 apply (simp_all)
-by (simp add: max_def ncols)
+by (simp add: ncols)
 
 lemma one_matrix_mult_left[simp]: "nrows A <= n \<Longrightarrow> (one_matrix n) * A = (A::('a::ring_1) matrix)"
 apply (subst Rep_matrix_inject[THEN sym])
@@ -1671,7 +1670,7 @@
 apply (simp add: times_matrix_def Rep_mult_matrix)
 apply (rule_tac j1="x" in ssubst[OF foldseq_almostzero])
 apply (simp_all)
-by (simp add: max_def nrows)
+by (simp add: nrows)
 
 lemma transpose_matrix_mult: "transpose_matrix ((A::('a::comm_ring) matrix)*B) = (transpose_matrix B) * (transpose_matrix A)"
 apply (simp add: times_matrix_def)
--- a/src/HOL/Matrix/ROOT.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Matrix/ROOT.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -1,5 +1,1 @@
-(*  Title:      HOL/Matrix/ROOT.ML
-    ID:         $Id$
-*)
-
 use_thys ["SparseMatrix", "cplex/Cplex"];
--- a/src/HOL/Matrix/SparseMatrix.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Matrix/SparseMatrix.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Matrix/SparseMatrix.thy
-    ID:         $Id$
     Author:     Steven Obua
 *)
 
--- a/src/HOL/Matrix/cplex/Cplex.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Matrix/cplex/Cplex.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -1,10 +1,9 @@
 (*  Title:      HOL/Matrix/cplex/Cplex.thy
-    ID:         $Id$
     Author:     Steven Obua
 *)
 
 theory Cplex 
-imports SparseMatrix LP "~~/src/HOL/Tools/ComputeFloat" "~~/src/HOL/Tools/ComputeNumeral"
+imports SparseMatrix LP ComputeFloat ComputeNumeral
 uses "Cplex_tools.ML" "CplexMatrixConverter.ML" "FloatSparseMatrixBuilder.ML"
   "fspmlp.ML" ("matrixlp.ML")
 begin
--- a/src/HOL/Matrix/cplex/CplexMatrixConverter.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Matrix/cplex/CplexMatrixConverter.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Matrix/cplex/CplexMatrixConverter.ML
-    ID:         $Id$
     Author:     Steven Obua
 *)
 
--- a/src/HOL/Matrix/cplex/Cplex_tools.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Matrix/cplex/Cplex_tools.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Matrix/cplex/Cplex_tools.ML
-    ID:         $Id$
     Author:     Steven Obua
 *)
 
--- a/src/HOL/MetisExamples/set.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/MetisExamples/set.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -238,7 +238,7 @@
 
 lemma (*singleton_example_2:*)
      "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
-by (metis Un_absorb2 Union_insert insertI1 insert_Diff insert_Diff_single subset_eq)
+by (metis Set.subsetI Union_upper insert_code mem_def set_eq_subset)
 
 lemma singleton_example_2:
      "\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
@@ -275,8 +275,8 @@
 apply (metis emptyE)
 apply (metis insert_iff singletonE)
 apply (metis insertCI singletonE zless_le)
-apply (metis insert_iff singletonE)
-apply (metis insert_iff singletonE)
+apply (metis Collect_def Collect_mem_eq)
+apply (metis Collect_def Collect_mem_eq)
 apply (metis DiffE)
 apply (metis pair_in_Id_conv) 
 done
--- a/src/HOL/MicroJava/BV/BVNoTypeError.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/MicroJava/BV/BVNoTypeError.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -145,7 +145,7 @@
   method (G,C) (mn,fpTs) = Some (mD', rT', b') \<and> G \<turnstile> X \<preceq> Class C)"
   (is "?app ST LT = ?P ST LT")
 proof
-  assume "?P ST LT" thus "?app ST LT" by (auto simp add: min_def list_all2_def)
+  assume "?P ST LT" thus "?app ST LT" by (auto simp add: list_all2_def)
 next  
   assume app: "?app ST LT"
   hence l: "length fpTs < length ST" by simp
@@ -153,7 +153,7 @@
   proof -
     have "ST = take (length fpTs) ST @ drop (length fpTs) ST" by simp
     moreover from l have "length (take (length fpTs) ST) = length fpTs"
-      by (simp add: min_def)
+      by simp
     ultimately show ?thesis ..
   qed
   obtain apTs where
@@ -168,11 +168,11 @@
   have "ST = (rev apTs) @ X # ST'" "length apTs = length fpTs" by auto
   with app
   show "?P ST LT"
-    apply (clarsimp simp add: list_all2_def min_def)
+    apply (clarsimp simp add: list_all2_def)
     apply ((rule exI)+, (rule conjI)?)+
     apply auto
     done
-qed 
+qed
 
 lemma approx_loc_len [simp]:
   "approx_loc G hp loc LT \<Longrightarrow> length loc = length LT"
--- a/src/HOL/MicroJava/BV/Effect.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/MicroJava/BV/Effect.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -375,7 +375,7 @@
     hence "?a \<and> 0 < length (drop (length fpTs) a)" (is "?a \<and> ?l") 
       by auto
     hence "?a \<and> ?l \<and> length (rev (take (length fpTs) a)) = length fpTs" 
-      by (auto simp add: min_def)
+      by (auto)
     hence "\<exists>apTs ST. a = rev apTs @ ST \<and> length apTs = length fpTs \<and> 0 < length ST" 
       by blast
     hence "\<exists>apTs ST. a = rev apTs @ ST \<and> length apTs = length fpTs \<and> ST \<noteq> []" 
@@ -391,7 +391,7 @@
   with Pair 
   have "?app s \<Longrightarrow> ?P s" by (simp only:)
   moreover
-  have "?P s \<Longrightarrow> ?app s" by (unfold app_def) (clarsimp simp add: min_def)
+  have "?P s \<Longrightarrow> ?app s" by (unfold app_def) (clarsimp)
   ultimately
   show ?thesis by (rule iffI) 
 qed 
--- a/src/HOL/MicroJava/BV/LBVSpec.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/MicroJava/BV/LBVSpec.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -293,7 +293,7 @@
   shows "wtl (take (pc+1) is) c 0 s = wtc c pc (wtl (take pc is) c 0 s)"
 proof -
   from suc have "take (pc+1) is=(take pc is)@[is!pc]" by (simp add: take_Suc)
-  with suc wtl show ?thesis by (simp add: min_def)
+  with suc wtl show ?thesis by (simp)
 qed
 
 lemma (in lbv) wtl_all:
@@ -308,7 +308,7 @@
   with all have take: "?wtl (take pc is@i#r) \<noteq> \<top>" by simp 
   from pc have "is!pc = drop pc is ! 0" by simp
   with Cons have "is!pc = i" by simp
-  with take pc show ?thesis by (auto simp add: min_def split: split_if_asm)
+  with take pc show ?thesis by (auto)
 qed
 
 section "preserves-type"
--- a/src/HOL/MicroJava/BV/Typing_Framework_err.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/MicroJava/BV/Typing_Framework_err.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -62,15 +62,9 @@
 lemma bounded_err_stepI:
   "\<forall>p. p < n \<longrightarrow> (\<forall>s. ap p s \<longrightarrow> (\<forall>(q,s') \<in> set (step p s). q < n))
   \<Longrightarrow> bounded (err_step n ap step) n"
-apply (unfold bounded_def)
-apply clarify
-apply (simp add: err_step_def split: err.splits)
-apply (simp add: error_def)
- apply blast
-apply (simp split: split_if_asm) 
- apply (blast dest: in_map_sndD)
-apply (simp add: error_def)
-apply blast
+apply (clarsimp simp: bounded_def err_step_def split: err.splits)
+apply (simp add: error_def image_def)
+apply (blast dest: in_map_sndD)
 done
 
 
--- a/src/HOL/MicroJava/Comp/CorrCompTp.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/MicroJava/Comp/CorrCompTp.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -959,7 +959,7 @@
 apply simp
 apply (simp (no_asm_simp))+
 apply simp
-apply (simp add: max_ssize_def max_of_list_append) apply (simp (no_asm_simp) add: max_def)
+apply (simp add: max_ssize_def max_of_list_append) apply (simp (no_asm_simp))
 
   (* show check_type \<dots> *)
 apply (subgoal_tac "((mt2 @ [Some sttp2]) ! length bc2) = Some sttp2")
@@ -973,7 +973,7 @@
 apply (simp add: check_type_def)
 apply (rule states_lower, assumption)
 apply (simp (no_asm_simp) add: max_ssize_def max_of_list_append)
-apply (simp (no_asm_simp) add: max_of_list_def ssize_sto_def max_def)
+apply (simp (no_asm_simp) add: max_of_list_def ssize_sto_def)
 apply (simp (no_asm_simp))+
 done
 
@@ -988,7 +988,7 @@
 apply (drule_tac x=sttp in spec, erule exE)
 apply simp
 apply (rule check_type_mono, assumption)
-apply (simp add: max_ssize_def max_of_list_def ssize_sto_def max_def split: prod.split)
+apply (simp add: max_ssize_def max_of_list_def ssize_sto_def split: prod.split)
 done
 
   (* ********************************************************************** *)
@@ -1058,8 +1058,7 @@
 lemma bc_mt_corresp_New: "\<lbrakk>is_class cG cname \<rbrakk>
   \<Longrightarrow> bc_mt_corresp [New cname] (pushST [Class cname]) (ST, LT) cG rT mxr (Suc 0)"
 apply (simp add: bc_mt_corresp_def pushST_def wt_instr_altern_def
-    max_ssize_def max_of_list_def ssize_sto_def max_def 
-    eff_def norm_eff_def)
+    max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def)
 apply (intro strip)
 apply (rule conjI)
 apply (rule check_type_mono, assumption, simp)
@@ -1070,7 +1069,6 @@
   bc_mt_corresp [Pop] (popST (Suc 0)) (T # ST, LT) cG rT mxr (Suc 0)"
   apply (simp add: bc_mt_corresp_def popST_def wt_instr_altern_def eff_def norm_eff_def)
   apply (simp add: max_ssize_def ssize_sto_def max_of_list_def) 
-  apply (simp add: max_def)
   apply (simp add: check_type_simps)
   apply clarify
   apply (rule_tac x="(length ST)" in exI)
@@ -1082,7 +1080,7 @@
   \<Longrightarrow> bc_mt_corresp [Checkcast cname] (replST (Suc 0) (Class cname)) sttp cG rT mxr (Suc 0)"
   apply (erule exE)+
   apply (simp add: bc_mt_corresp_def replST_def wt_instr_altern_def eff_def norm_eff_def)
-  apply (simp add: max_ssize_def max_of_list_def ssize_sto_def max_def)
+  apply (simp add: max_ssize_def max_of_list_def ssize_sto_def)
   apply (simp add: check_type_simps)
   apply clarify
   apply (rule_tac x="Suc (length STo)" in exI)
@@ -1094,8 +1092,7 @@
   \<Longrightarrow> bc_mt_corresp [LitPush val] (pushST [T]) sttp cG rT mxr (Suc 0)"
 apply (subgoal_tac "\<exists> ST LT. sttp= (ST, LT)", (erule exE)+)
   apply (simp add: bc_mt_corresp_def pushST_def wt_instr_altern_def
-                 max_ssize_def max_of_list_def ssize_sto_def max_def
-                 eff_def norm_eff_def)
+              max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def)
   apply (intro strip)
   apply (rule conjI)
   apply (rule check_type_mono, assumption, simp)
@@ -1114,8 +1111,7 @@
   \<Longrightarrow> bc_mt_corresp [LitPush val] (pushST [T']) sttp cG rT mxr (Suc 0)"
 apply (subgoal_tac "\<exists> ST LT. sttp= (ST, LT)", (erule exE)+)
   apply (simp add: bc_mt_corresp_def pushST_def wt_instr_altern_def
-                 max_ssize_def max_of_list_def ssize_sto_def max_def
-                 eff_def norm_eff_def)
+              max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def)
   apply (intro strip)
   apply (rule conjI)
   apply (rule check_type_mono, assumption, simp)
@@ -1131,8 +1127,7 @@
   \<Longrightarrow> bc_mt_corresp [Load i] 
          (\<lambda>(ST, LT). pushST [ok_val (LT ! i)] (ST, LT)) (ST, LT) cG rT mxr (Suc 0)"
 apply (simp add: bc_mt_corresp_def pushST_def wt_instr_altern_def
-                 max_ssize_def max_of_list_def ssize_sto_def max_def
-                 eff_def norm_eff_def)
+            max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def)
   apply (intro strip)
   apply (rule conjI)
   apply (rule check_type_mono, assumption, simp)
@@ -1154,7 +1149,7 @@
   \<Longrightarrow> bc_mt_corresp [Store i] (storeST i T) (T # ST, LT) cG rT mxr (Suc 0)"
  apply (simp add: bc_mt_corresp_def storeST_def wt_instr_altern_def eff_def norm_eff_def)
   apply (simp add: max_ssize_def  max_of_list_def )
-  apply (simp add: ssize_sto_def) apply (simp add: max_def)
+  apply (simp add: ssize_sto_def)
   apply (intro strip)
 apply (simp add: check_type_simps)
 apply clarify
@@ -1170,7 +1165,6 @@
   apply (simp add: bc_mt_corresp_def popST_def wt_instr_altern_def eff_def norm_eff_def)
   apply (simp add: sup_state_conv)
   apply (simp add: max_ssize_def max_of_list_def ssize_sto_def)
-  apply (simp add: max_def)
  apply (intro strip)
 apply (simp add: check_type_simps)
 apply clarify
@@ -1182,8 +1176,7 @@
 lemma bc_mt_corresp_Dup: "
   bc_mt_corresp [Dup] dupST (T # ST, LT) cG rT mxr (Suc 0)"
  apply (simp add: bc_mt_corresp_def dupST_def wt_instr_altern_def
-                 max_ssize_def max_of_list_def ssize_sto_def max_def
-                 eff_def norm_eff_def)
+             max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def)
   apply (intro strip)
   apply (rule conjI)
   apply (rule check_type_mono, assumption, simp)
@@ -1196,8 +1189,7 @@
 lemma bc_mt_corresp_Dup_x1: "
   bc_mt_corresp [Dup_x1] dup_x1ST (T1 # T2 # ST, LT) cG rT mxr (Suc 0)"
   apply (simp add: bc_mt_corresp_def dup_x1ST_def wt_instr_altern_def
-                 max_ssize_def max_of_list_def ssize_sto_def max_def
-                 eff_def norm_eff_def)
+              max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def)
   apply (intro strip)
   apply (rule conjI)
   apply (rule check_type_mono, assumption, simp)
@@ -1213,7 +1205,7 @@
   bc_mt_corresp [IAdd] (replST 2 (PrimT Integer)) 
          (PrimT Integer # PrimT Integer # ST, LT) cG rT mxr (Suc 0)"
   apply (simp add: bc_mt_corresp_def replST_def wt_instr_altern_def eff_def norm_eff_def)
-  apply (simp add: max_ssize_def max_of_list_def ssize_sto_def max_def)
+  apply (simp add: max_ssize_def max_of_list_def ssize_sto_def)
   apply (simp add: check_type_simps)
   apply clarify
   apply (rule_tac x="Suc (length ST)" in exI)
@@ -1254,7 +1246,7 @@
   apply (frule widen_field, assumption+)
   apply (simp add: bc_mt_corresp_def popST_def wt_instr_altern_def eff_def norm_eff_def)
   apply (simp add: comp_field comp_subcls1 comp_widen comp_is_class)
-  apply (simp add: max_ssize_def max_of_list_def ssize_sto_def max_def)
+  apply (simp add: max_ssize_def max_of_list_def ssize_sto_def)
 
   apply (intro strip)
 apply (simp add: check_type_simps)
@@ -1305,7 +1297,7 @@
   apply (simp add: max_spec_preserves_length [THEN sym])
 
   -- "@{text check_type}"
-apply (simp add: max_ssize_def ssize_sto_def max_def)
+apply (simp add: max_ssize_def ssize_sto_def)
 apply (simp add: max_of_list_def)
 apply (subgoal_tac "(max (length pTsa + length ST) (length ST)) = (length pTsa + length ST)")
 apply simp
@@ -1316,7 +1308,7 @@
 apply (simp only: comp_is_type)
 apply (frule method_wf_mdecl) apply assumption apply assumption
 apply (simp add: wf_mdecl_def wf_mhead_def)
-apply (simp add: max_def)
+apply (simp)
   done
 
 
@@ -1473,7 +1465,7 @@
 apply (case_tac "sttp1", simp)
 apply (rule check_type_lower) apply assumption
 apply (simp (no_asm_simp) add: max_ssize_def ssize_sto_def)
-apply (simp (no_asm_simp) add: max_of_list_def max_def)
+apply (simp (no_asm_simp) add: max_of_list_def)
 
   (* subgoals \<exists> ... *)
 apply (rule surj_pair)+
--- a/src/HOL/MicroJava/J/TypeRel.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/MicroJava/J/TypeRel.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -97,6 +97,14 @@
   qed
 qed
 
+text {* Code generator setup (FIXME!) *}
+
+consts_code
+  "wfrec"   ("\<module>wfrec?")
+attach {*
+fun wfrec f x = f (wfrec f) x;
+*}
+
 consts
 
   method :: "'c prog \<times> cname => ( sig   \<rightharpoonup> cname \<times> ty \<times> 'c)" (* ###curry *)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mirabelle/Mirabelle.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,25 @@
+(* Title: Mirabelle.thy
+   Author: Jasmin Blanchette and Sascha Boehme
+*)
+
+theory Mirabelle
+imports Pure
+uses "Tools/mirabelle.ML"
+begin
+
+(* no multithreading, no parallel proofs *)
+ML {* Multithreading.max_threads := 1 *}
+ML {* Goal.parallel_proofs := 0 *}
+
+ML {* Toplevel.add_hook Mirabelle.step_hook *}
+
+setup Mirabelle.setup
+
+ML {*
+signature MIRABELLE_ACTION =
+sig
+  val invoke : (string * string) list -> theory -> theory
+end
+*}
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mirabelle/Mirabelle_Test.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,22 @@
+(* Title: Mirabelle_Test.thy
+   Author: Sascha Boehme
+*)
+
+header {* Simple test theory for Mirabelle actions *}
+
+theory Mirabelle_Test
+imports Main Mirabelle
+uses
+  "Tools/mirabelle_arith.ML"
+  "Tools/mirabelle_metis.ML"
+  "Tools/mirabelle_quickcheck.ML"
+  "Tools/mirabelle_refute.ML"
+  "Tools/mirabelle_sledgehammer.ML"
+begin
+
+(*
+  only perform type-checking of the actions,
+  any reasonable test would be too complicated
+*)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mirabelle/ROOT.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,1 @@
+use_thy "Mirabelle_Test";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mirabelle/Tools/mirabelle.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,215 @@
+(* Title:  mirabelle.ML
+   Author: Jasmin Blanchette and Sascha Boehme
+*)
+
+signature MIRABELLE =
+sig
+  (* configuration *)
+  val logfile : string Config.T
+  val timeout : int Config.T
+  val start_line : int Config.T
+  val end_line : int Config.T
+  val setup : theory -> theory
+
+  (* core *)
+  type init_action
+  type done_action
+  type run_action
+  type action
+  val catch : (int -> string) -> run_action -> run_action
+  val register : action -> theory -> theory
+  val step_hook : Toplevel.transition -> Toplevel.state -> Toplevel.state ->
+    unit
+
+  (* utility functions *)
+  val goal_thm_of : Proof.state -> thm
+  val can_apply : Time.time -> (Proof.context -> int -> tactic) ->
+    Proof.state -> bool
+  val theorems_in_proof_term : Thm.thm -> Thm.thm list
+  val theorems_of_sucessful_proof : Toplevel.state option -> Thm.thm list
+  val get_setting : (string * string) list -> string * string -> string
+  val get_int_setting : (string * string) list -> string * int -> int
+  val cpu_time : ('a -> 'b) -> 'a -> 'b * int
+end
+
+
+
+structure Mirabelle : MIRABELLE =
+struct
+
+(* Mirabelle configuration *)
+
+val (logfile, setup1) = Attrib.config_string "mirabelle_logfile" ""
+val (timeout, setup2) = Attrib.config_int "mirabelle_timeout" 30
+val (start_line, setup3) = Attrib.config_int "mirabelle_start_line" 0
+val (end_line, setup4) = Attrib.config_int "mirabelle_end_line" ~1
+
+val setup = setup1 #> setup2 #> setup3 #> setup4
+
+
+
+(* Mirabelle core *)
+
+
+type init_action = int -> theory -> theory
+type done_action = int -> {last: Toplevel.state, log: string -> unit} -> unit
+type run_action = int -> {pre: Proof.state, post: Toplevel.state option,
+  timeout: Time.time, log: string -> unit, pos: Position.T} -> unit
+type action = init_action * run_action * done_action
+
+structure Actions = TheoryDataFun
+(
+  type T = (int * run_action * done_action) list
+  val empty = []
+  val copy = I
+  val extend = I
+  fun merge _ = Library.merge (K true)
+)
+
+
+fun app_with f g x = (g x; ())
+  handle (exn as Exn.Interrupt) => reraise exn | exn => (f exn; ())
+
+fun catch tag f id (st as {log, ...}) =
+  let fun log_exn e = log (tag id ^ "exception:\n" ^ General.exnMessage e)
+  in app_with log_exn (f id) st end
+
+fun register (init, run, done) thy =
+  let val id = length (Actions.get thy) + 1
+  in
+    thy
+    |> init id
+    |> Actions.map (cons (id, run, done))
+  end
+
+local
+
+fun log thy s =
+  let fun append_to n = if n = "" then K () else File.append (Path.explode n)
+  in append_to (Config.get_thy thy logfile) (s ^ "\n") end
+  (* FIXME: with multithreading and parallel proofs enabled, we might need to
+     encapsulate this inside a critical section *)
+
+fun log_sep thy = log thy "------------------"
+
+fun apply_actions thy pos info (pre, post, time) actions =
+  let
+    fun apply f = f {pre=pre, post=post, timeout=time, log=log thy, pos=pos}
+    fun run (id, run, _) = (apply (run id); log_sep thy)
+  in (log thy info; log_sep thy; List.app run actions) end
+
+fun in_range _ _ NONE = true
+  | in_range l r (SOME i) = (l <= i andalso (r < 0 orelse i <= r))
+
+fun only_within_range thy pos f x =
+  let val l = Config.get_thy thy start_line and r = Config.get_thy thy end_line
+  in if in_range l r (Position.line_of pos) then f x else () end
+
+in
+
+fun run_actions tr pre post =
+  let
+    val thy = Proof.theory_of pre
+    val pos = Toplevel.pos_of tr
+    val name = Toplevel.name_of tr
+    val st = (pre, post, Time.fromSeconds (Config.get_thy thy timeout))
+
+    val str0 = string_of_int o the_default 0
+    val loc = str0 (Position.line_of pos) ^ ":" ^ str0 (Position.column_of pos)
+    val info = "\n\nat " ^ loc ^ " (" ^ name ^ "):"
+  in
+    only_within_range thy pos (apply_actions thy pos info st) (Actions.get thy)
+  end
+
+fun done_actions st =
+  let
+    val thy = Toplevel.theory_of st
+    val _ = log thy "\n\n";
+  in
+    thy
+    |> Actions.get
+    |> List.app (fn (id, _, done) => done id {last=st, log=log thy})
+  end
+
+end
+
+val whitelist = ["apply", "by", "proof"]
+
+fun step_hook tr pre post =
+ (* FIXME: might require wrapping into "interruptible" *)
+  if can (Proof.assert_backward o Toplevel.proof_of) pre andalso
+     member (op =) whitelist (Toplevel.name_of tr)
+  then run_actions tr (Toplevel.proof_of pre) (SOME post)
+  else if not (Toplevel.is_toplevel pre) andalso Toplevel.is_toplevel post
+  then done_actions pre
+  else ()   (* FIXME: add theory_hook here *)
+
+
+
+(* Mirabelle utility functions *)
+
+val goal_thm_of = snd o snd o Proof.get_goal
+
+fun can_apply time tac st =
+  let
+    val (ctxt, (facts, goal)) = Proof.get_goal st
+    val full_tac = HEADGOAL (Method.insert_tac facts THEN' tac ctxt)
+  in
+    (case TimeLimit.timeLimit time (Seq.pull o full_tac) goal of
+      SOME (thm, _) => true
+    | NONE => false)
+  end
+
+local
+
+fun fold_body_thms f =
+  let
+    fun app n (PBody {thms, ...}) = thms |> fold (fn (i, (name, prop, body)) =>
+      fn (x, seen) =>
+        if Inttab.defined seen i then (x, seen)
+        else
+          let
+            val body' = Future.join body
+            val (x', seen') = app (n + (if name = "" then 0 else 1)) body'
+              (x, Inttab.update (i, ()) seen)
+        in (x' |> n = 0 ? f (name, prop, body'), seen') end)
+  in fn bodies => fn x => #1 (fold (app 0) bodies (x, Inttab.empty)) end
+
+in
+
+fun theorems_in_proof_term thm =
+  let
+    val all_thms = PureThy.all_thms_of (Thm.theory_of_thm thm)
+    fun collect (s, _, _) = if s <> "" then insert (op =) s else I
+    fun member_of xs (x, y) = if member (op =) xs x then SOME y else NONE
+    fun resolve_thms names = map_filter (member_of names) all_thms
+  in
+    resolve_thms (fold_body_thms collect [Thm.proof_body_of thm] [])
+  end
+
+end
+
+fun theorems_of_sucessful_proof state =
+  (case state of
+    NONE => []
+  | SOME st =>
+      if not (Toplevel.is_proof st) then []
+      else theorems_in_proof_term (goal_thm_of (Toplevel.proof_of st)))
+
+fun get_setting settings (key, default) =
+  the_default default (AList.lookup (op =) settings key)
+
+fun get_int_setting settings (key, default) =
+  (case Option.map Int.fromString (AList.lookup (op =) settings key) of
+    SOME (SOME i) => i
+  | SOME NONE => error ("bad option: " ^ key)
+  | NONE => default)
+
+fun cpu_time f x =
+  let
+    val start = start_timing ()
+    val result = Exn.capture (fn () => f x) ()
+    val time = Time.toMilliseconds (#cpu (end_timing start))
+  in (Exn.release result, time) end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mirabelle/Tools/mirabelle_arith.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,21 @@
+(* Title:  mirabelle_arith.ML
+   Author: Jasmin Blanchette and Sascha Boehme
+*)
+
+structure Mirabelle_Arith : MIRABELLE_ACTION =
+struct
+
+fun arith_tag id = "#" ^ string_of_int id ^ " arith: "
+
+fun init _ = I
+fun done _ _ = ()
+
+fun run id {pre, timeout, log, ...} =
+  if Mirabelle.can_apply timeout Arith_Data.arith_tac pre
+  then log (arith_tag id ^ "succeeded")
+  else ()
+  handle TimeLimit.TimeOut => log (arith_tag id ^ "timeout")
+
+fun invoke _ = Mirabelle.register (init, Mirabelle.catch arith_tag run, done)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mirabelle/Tools/mirabelle_metis.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,33 @@
+(* Title:  mirabelle_metis.ML
+   Author: Jasmin Blanchette and Sascha Boehme
+*)
+
+structure Mirabelle_Metis : MIRABELLE_ACTION =
+struct
+
+fun metis_tag id = "#" ^ string_of_int id ^ " metis: "
+
+fun init _ = I
+fun done _ _ = ()
+
+fun run id {pre, post, timeout, log} =
+  let
+    val thms = Mirabelle.theorems_of_sucessful_proof post
+    val names = map Thm.get_name thms
+    val add_info = if null names then I else suffix (":\n" ^ commas names)
+
+    val facts = Facts.props (ProofContext.facts_of (Proof.context_of pre))
+
+    fun metis ctxt = MetisTools.metis_tac ctxt (thms @ facts)
+  in
+    (if Mirabelle.can_apply timeout metis pre then "succeeded" else "failed")
+    |> prefix (metis_tag id)
+    |> add_info
+    |> log
+  end
+  handle TimeLimit.TimeOut => log (metis_tag id ^ "timeout")
+       | ERROR msg => log (metis_tag id ^ "error: " ^ msg)
+
+fun invoke _ = Mirabelle.register (init, Mirabelle.catch metis_tag run, done)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mirabelle/Tools/mirabelle_quickcheck.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,28 @@
+(* Title:  mirabelle_quickcheck.ML
+   Author: Jasmin Blanchette and Sascha Boehme
+*)
+
+structure Mirabelle_Quickcheck : MIRABELLE_ACTION =
+struct
+
+fun qc_tag id = "#" ^ string_of_int id ^ " quickcheck: "
+
+fun init _ = I
+fun done _ _ = ()
+
+fun run args id {pre, timeout, log, ...} =
+  let
+    val has_valid_key = member (op =) ["iterations", "size", "generator"] o fst
+    val quickcheck = Quickcheck.quickcheck (filter has_valid_key args) 1
+  in
+    (case TimeLimit.timeLimit timeout quickcheck pre of
+      NONE => log (qc_tag id ^ "no counterexample")
+    | SOME _ => log (qc_tag id ^ "counterexample found"))
+  end
+  handle TimeLimit.TimeOut => log (qc_tag id ^ "timeout")
+       | ERROR msg => log (qc_tag id ^ "error: " ^ msg)
+
+fun invoke args =
+  Mirabelle.register (init, Mirabelle.catch qc_tag (run args), done)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mirabelle/Tools/mirabelle_refute.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,39 @@
+(* Title:  mirabelle_refute.ML
+   Author: Jasmin Blanchette and Sascha Boehme
+*)
+
+structure Mirabelle_Refute : MIRABELLE_ACTION =
+struct
+
+
+(* FIXME:
+fun refute_action args timeout {pre=st, ...} = 
+  let
+    val subgoal = 0
+    val thy     = Proof.theory_of st
+    val thm = goal_thm_of st
+
+    val refute = Refute.refute_subgoal thy args thm
+    val _ = TimeLimit.timeLimit timeout refute subgoal
+  in
+    val writ_log = Substring.full (the (Symtab.lookup tab "writeln"))
+    val warn_log = Substring.full (the (Symtab.lookup tab "warning"))
+
+    val r =
+      if Substring.isSubstring "model found" writ_log
+      then
+        if Substring.isSubstring "spurious" warn_log
+        then SOME "potential counterexample"
+        else SOME "real counterexample (bug?)"
+      else
+        if Substring.isSubstring "time limit" writ_log
+        then SOME "no counterexample (timeout)"
+        else if Substring.isSubstring "Search terminated" writ_log
+        then SOME "no counterexample (normal termination)"
+        else SOME "no counterexample (unknown)"
+  in r end
+*)
+
+fun invoke args = I (*Mirabelle.register ("refute", refute_action args)*)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,365 @@
+(* Title:  mirabelle_sledgehammer.ML
+   Author: Jasmin Blanchette and Sascha Boehme and Tobias Nipkow
+*)
+
+structure Mirabelle_Sledgehammer : MIRABELLE_ACTION =
+struct
+
+val proverK = "prover"
+val prover_timeoutK = "prover_timeout"
+val keepK = "keep"
+val full_typesK = "full_types"
+val minimizeK = "minimize"
+val minimize_timeoutK = "minimize_timeout"
+
+fun sh_tag id = "#" ^ string_of_int id ^ " sledgehammer: "
+fun minimize_tag id = "#" ^ string_of_int id ^ " minimize (sledgehammer): "
+fun metis_tag id = "#" ^ string_of_int id ^ " metis (sledgehammer): "
+
+val separator = "-----"
+
+
+datatype sh_data = ShData of {
+  calls: int,
+  success: int,
+  time_isa: int,
+  time_atp: int,
+  time_atp_fail: int}
+
+datatype me_data = MeData of {
+  calls: int,
+  success: int,
+  time: int,
+  timeout: int,
+  lemmas: int,
+  posns: Position.T list
+  }
+
+
+(* The first me_data component is only used if "minimize" is on.
+   Then it records how metis behaves with un-minimized lemmas.
+*)
+datatype data = Data of sh_data * me_data * me_data
+
+fun make_sh_data (sh_calls, sh_success, sh_time_isa, sh_time_atp, sh_time_atp_fail) =
+  ShData{calls=sh_calls, success=sh_success, time_isa=sh_time_isa,
+    time_atp=sh_time_atp, time_atp_fail=sh_time_atp_fail}
+
+fun make_me_data (calls, success, time, timeout, lemmas, posns) =
+  MeData{calls=calls, success=success, time=time, timeout=timeout, lemmas=lemmas, posns=posns}
+
+val empty_data = Data(make_sh_data (0, 0, 0, 0, 0), make_me_data(0, 0, 0, 0, 0, []), make_me_data(0, 0, 0, 0, 0, []))
+
+fun map_sh_data f
+  (Data (ShData{calls, success, time_isa, time_atp, time_atp_fail}, meda0, meda)) =
+  Data (make_sh_data (f (calls, success, time_isa, time_atp, time_atp_fail)),
+        meda0, meda)
+
+fun map_me_data0 f (Data (shda, MeData{calls,success,time,timeout,lemmas,posns}, meda)) =
+  Data(shda, make_me_data(f (calls,success,time,timeout,lemmas,posns)), meda)
+
+fun map_me_data f (Data (shda, meda0, MeData{calls,success,time,timeout,lemmas,posns})) =
+  Data(shda, meda0, make_me_data(f (calls,success,time,timeout,lemmas,posns)))
+
+val inc_sh_calls = map_sh_data (fn (sh_calls, sh_success, sh_time_isa, sh_time_atp, sh_time_atp_fail)
+  => (sh_calls + 1, sh_success, sh_time_isa, sh_time_atp, sh_time_atp_fail))
+
+val inc_sh_success = map_sh_data (fn (sh_calls, sh_success, sh_time_isa, sh_time_atp, sh_time_atp_fail)
+  => (sh_calls, sh_success + 1, sh_time_isa, sh_time_atp, sh_time_atp_fail))
+
+fun inc_sh_time_isa t = map_sh_data (fn (sh_calls, sh_success, sh_time_isa, sh_time_atp, sh_time_atp_fail)
+  => (sh_calls, sh_success, sh_time_isa + t, sh_time_atp, sh_time_atp_fail))
+
+fun inc_sh_time_atp t = map_sh_data (fn (sh_calls, sh_success, sh_time_isa, sh_time_atp, sh_time_atp_fail)
+  => (sh_calls, sh_success, sh_time_isa, sh_time_atp + t, sh_time_atp_fail))
+
+fun inc_sh_time_atp_fail t = map_sh_data (fn (sh_calls, sh_success, sh_time_isa, sh_time_atp, sh_time_atp_fail)
+  => (sh_calls, sh_success, sh_time_isa, sh_time_atp, sh_time_atp_fail + t))
+
+val inc_metis_calls = map_me_data
+ (fn (calls, success, time, timeout, lemmas,posns)
+  => (calls + 1, success, time, timeout, lemmas,posns))
+
+val inc_metis_success = map_me_data
+ (fn (calls,success,time,timeout,lemmas,posns)
+  => (calls, success + 1, time, timeout, lemmas,posns))
+
+fun inc_metis_time t = map_me_data
+ (fn (calls,success,time,timeout,lemmas,posns)
+  => (calls, success, time + t, timeout, lemmas,posns))
+
+val inc_metis_timeout = map_me_data
+ (fn (calls,success,time,timeout,lemmas,posns)
+  => (calls, success, time, timeout + 1, lemmas,posns))
+
+fun inc_metis_lemmas n = map_me_data
+ (fn (calls,success,time,timeout,lemmas,posns)
+  => (calls, success, time, timeout, lemmas + n, posns))
+
+fun inc_metis_posns pos = map_me_data
+ (fn (calls,success,time,timeout,lemmas,posns)
+  => (calls, success, time, timeout, lemmas, pos::posns))
+
+val inc_metis_calls0 = map_me_data0 
+(fn (calls, success, time, timeout, lemmas,posns)
+  => (calls + 1, success, time, timeout, lemmas,posns))
+
+val inc_metis_success0 = map_me_data0
+ (fn (calls,success,time,timeout,lemmas,posns)
+  => (calls, success + 1, time, timeout, lemmas,posns))
+
+fun inc_metis_time0 t = map_me_data0
+ (fn (calls,success,time,timeout,lemmas,posns)
+  => (calls, success, time + t, timeout, lemmas,posns))
+
+val inc_metis_timeout0 = map_me_data0
+ (fn (calls,success,time,timeout,lemmas,posns)
+  => (calls, success, time, timeout + 1, lemmas,posns))
+
+fun inc_metis_lemmas0 n = map_me_data0
+ (fn (calls,success,time,timeout,lemmas,posns)
+  => (calls, success, time, timeout, lemmas + n, posns))
+
+fun inc_metis_posns0 pos = map_me_data0
+ (fn (calls,success,time,timeout,lemmas,posns)
+  => (calls, success, time, timeout, lemmas, pos::posns))
+
+local
+
+val str = string_of_int
+val str3 = Real.fmt (StringCvt.FIX (SOME 3))
+fun percentage a b = string_of_int (a * 100 div b)
+fun time t = Real.fromInt t / 1000.0
+fun avg_time t n =
+  if n > 0 then (Real.fromInt t / 1000.0) / Real.fromInt n else 0.0
+
+fun log_sh_data log sh_calls sh_success sh_time_isa sh_time_atp sh_time_atp_fail =
+ (log ("Total number of sledgehammer calls: " ^ str sh_calls);
+  log ("Number of successful sledgehammer calls: " ^ str sh_success);
+  log ("Success rate: " ^ percentage sh_success sh_calls ^ "%");
+  log ("Total time for sledgehammer calls (Isabelle): " ^ str3 (time sh_time_isa));
+  log ("Total time for successful sledgehammer calls (ATP): " ^ str3 (time sh_time_atp));
+  log ("Total time for failed sledgehammer calls (ATP): " ^ str3 (time sh_time_atp_fail));
+  log ("Average time for sledgehammer calls (Isabelle): " ^
+    str3 (avg_time sh_time_isa sh_calls));
+  log ("Average time for successful sledgehammer calls (ATP): " ^
+    str3 (avg_time sh_time_atp sh_success));
+  log ("Average time for failed sledgehammer calls (ATP): " ^
+    str3 (avg_time sh_time_atp_fail (sh_calls - sh_success)))
+  )
+
+
+fun str_of_pos pos =
+  let val str0 = string_of_int o the_default 0
+  in str0 (Position.line_of pos) ^ ":" ^ str0 (Position.column_of pos) end
+
+fun log_metis_data log tag sh_calls sh_success metis_calls metis_success metis_time
+    metis_timeout metis_lemmas metis_posns =
+ (log ("Total number of " ^ tag ^ "metis calls: " ^ str metis_calls);
+  log ("Number of successful " ^ tag ^ "metis calls: " ^ str metis_success);
+  log ("Number of " ^ tag ^ "metis timeouts: " ^ str metis_timeout);
+  log ("Number of " ^ tag ^ "metis exceptions: " ^
+    str (sh_success - metis_success - metis_timeout));
+  log ("Success rate: " ^ percentage metis_success sh_calls ^ "%");
+  log ("Number of " ^ tag ^ "metis lemmas: " ^ str metis_lemmas);
+  log ("Total time for successful metis calls: " ^ str3 (time metis_time));
+  log ("Average time for successful metis calls: " ^
+    str3 (avg_time metis_time metis_success));
+  if tag=""
+  then log ("Proved: " ^ space_implode " " (map str_of_pos metis_posns))
+  else ()
+ )
+in
+
+fun log_data id log (Data (ShData{calls=sh_calls, success=sh_success, time_isa=sh_time_isa, time_atp=sh_time_atp, time_atp_fail=sh_time_atp_fail}, MeData{calls=metis_calls0,
+    success=metis_success0, time=metis_time0, timeout=metis_timeout0, lemmas=metis_lemmas0,posns=metis_posns0}, MeData{calls=metis_calls,
+    success=metis_success, time=metis_time, timeout=metis_timeout, lemmas=metis_lemmas,posns=metis_posns})) =
+  if sh_calls > 0
+  then
+   (log ("\n\n\nReport #" ^ string_of_int id ^ ":\n");
+    log_sh_data log sh_calls sh_success sh_time_isa sh_time_atp sh_time_atp_fail;
+    log "";
+    if metis_calls > 0 then log_metis_data log "" sh_calls sh_success metis_calls
+      metis_success metis_time metis_timeout metis_lemmas  metis_posns else ();
+    log "";
+    if metis_calls0 > 0
+      then log_metis_data log "unminimized " sh_calls sh_success metis_calls0
+              metis_success0 metis_time0 metis_timeout0 metis_lemmas0 metis_posns0
+      else ()
+   )
+  else ()
+
+end
+
+
+(* Warning: we implicitly assume single-threaded execution here! *)
+val data = ref ([] : (int * data) list)
+
+fun init id thy = (change data (cons (id, empty_data)); thy)
+fun done id {log, ...} =
+  AList.lookup (op =) (!data) id
+  |> Option.map (log_data id log)
+  |> K ()
+
+fun change_data id f = (change data (AList.map_entry (op =) id f); ())
+
+
+fun get_atp thy args =
+  AList.lookup (op =) args proverK
+  |> the_default (hd (space_explode " " (AtpManager.get_atps ())))
+  |> (fn name => (name, the (AtpManager.get_prover name thy)))
+
+local
+
+fun safe init done f x =
+  let
+    val y = init x
+    val z = Exn.capture f y
+    val _ = done y
+  in Exn.release z end
+
+fun init_sh NONE = !AtpWrapper.destdir
+  | init_sh (SOME path) =
+      let
+        (* Warning: we implicitly assume single-threaded execution here! *)
+        val old = !AtpWrapper.destdir
+        val _ = AtpWrapper.destdir := path
+      in old end
+
+fun done_sh path = AtpWrapper.destdir := path
+
+datatype sh_result =
+  SH_OK of int * int * string list |
+  SH_FAIL of int * int |
+  SH_ERROR
+
+fun run_sh (prover_name, prover) timeout st _ =
+  let
+    val atp = prover timeout NONE NONE prover_name 1
+    val ((success, (message, thm_names), time_atp, _, _, _), time_isa) =
+      Mirabelle.cpu_time atp (Proof.get_goal st)
+  in
+    if success then (message, SH_OK (time_isa, time_atp, thm_names))
+    else (message, SH_FAIL(time_isa, time_atp))
+  end
+  handle ResHolClause.TOO_TRIVIAL => ("trivial", SH_OK (0, 0, []))
+       | ERROR msg => ("error: " ^ msg, SH_ERROR)
+
+fun thms_of_name ctxt name =
+  let
+    val lex = OuterKeyword.get_lexicons
+    val get = maps (ProofContext.get_fact ctxt o fst)
+  in
+    Source.of_string name
+    |> Symbol.source {do_recover=false}
+    |> OuterLex.source {do_recover=SOME false} lex Position.start
+    |> OuterLex.source_proper
+    |> Source.source OuterLex.stopper (SpecParse.xthms1 >> get) NONE
+    |> Source.exhaust
+  end
+
+in
+
+fun run_sledgehammer args named_thms id {pre=st, log, ...} =
+  let
+    val _ = change_data id inc_sh_calls
+    val atp as (prover_name, _) = get_atp (Proof.theory_of st) args
+    val dir = AList.lookup (op =) args keepK
+    val timeout = Mirabelle.get_int_setting args (prover_timeoutK, 30)
+    val (msg, result) = safe init_sh done_sh (run_sh atp timeout st) dir
+  in
+    case result of
+      SH_OK (time_isa, time_atp, names) =>
+        let
+          val _ = change_data id inc_sh_success
+          val _ = change_data id (inc_sh_time_isa time_isa)
+          val _ = change_data id (inc_sh_time_atp time_atp)
+
+          fun get_thms name = (name, thms_of_name (Proof.context_of st) name)
+          val _ = named_thms := SOME (map get_thms names)
+        in
+          log (sh_tag id ^ "succeeded (" ^ string_of_int time_isa ^ "+" ^
+            string_of_int time_atp ^ ") [" ^ prover_name ^ "]:\n" ^ msg)
+        end
+    | SH_FAIL (time_isa, time_atp) =>
+        let
+          val _ = change_data id (inc_sh_time_isa time_isa)
+          val _ = change_data id (inc_sh_time_atp_fail time_atp)
+        in log (sh_tag id ^ "failed: " ^ msg) end
+    | SH_ERROR => log (sh_tag id ^ "failed: " ^ msg)
+  end
+
+end
+
+
+fun run_minimize args named_thms id {pre=st, log, ...} =
+  let
+    val (prover_name, prover) = get_atp (Proof.theory_of st) args
+    val minimize = AtpMinimal.minimalize prover prover_name
+    val timeout =
+      AList.lookup (op =) args minimize_timeoutK
+      |> Option.map (fst o read_int o explode)
+      |> the_default 5
+    val _ = log separator
+  in
+    (case minimize timeout st (these (!named_thms)) of
+      (SOME named_thms', msg) =>
+        if length named_thms' = length (these (!named_thms))
+        then log (minimize_tag id ^ "already minimal")
+        else
+         (named_thms := SOME named_thms';
+          log (minimize_tag id ^ "succeeded:\n" ^ msg))
+    | (NONE, msg) => log (minimize_tag id ^ "failed: " ^ msg))
+  end
+
+
+fun run_metis (inc_metis_calls, inc_metis_success, inc_metis_time, inc_metis_timeout, inc_metis_lemmas, inc_metis_posns) args named_thms id {pre=st, timeout, log, pos, ...} =
+  let
+    fun metis thms ctxt = MetisTools.metis_tac ctxt thms
+    fun apply_metis thms = Mirabelle.can_apply timeout (metis thms) st
+
+    fun with_time (false, t) = "failed (" ^ string_of_int t ^ ")"
+      | with_time (true, t) = (change_data id inc_metis_success;
+          change_data id (inc_metis_time t);
+          change_data id (inc_metis_posns pos);
+          "succeeded (" ^ string_of_int t ^ ")")
+    fun timed_metis thms = with_time (Mirabelle.cpu_time apply_metis thms)
+      handle TimeLimit.TimeOut => (change_data id inc_metis_timeout; "timeout")
+           | ERROR msg => "error: " ^ msg
+
+    val _ = log separator
+    val _ = change_data id inc_metis_calls
+    val _ = change_data id (inc_metis_lemmas (length named_thms))
+  in
+    maps snd named_thms
+    |> timed_metis
+    |> log o prefix (metis_tag id) 
+  end
+
+fun sledgehammer_action args id (st as {log, ...}) =
+  let
+    val metis_fns = (inc_metis_calls, inc_metis_success, inc_metis_time, inc_metis_timeout, inc_metis_lemmas, inc_metis_posns)
+    val metis0_fns = (inc_metis_calls0, inc_metis_success0, inc_metis_time0, inc_metis_timeout0, inc_metis_lemmas0, inc_metis_posns0)
+    val named_thms = ref (NONE : (string * thm list) list option)
+
+    fun if_enabled k f =
+      if AList.defined (op =) args k andalso is_some (!named_thms)
+      then f id st else ()
+
+    val _ = Mirabelle.catch sh_tag (run_sledgehammer args named_thms) id st
+    val _ = if_enabled minimizeK
+      (Mirabelle.catch metis_tag (run_metis metis0_fns args (these (!named_thms))))
+    val _ = if is_some (!named_thms)
+      then Mirabelle.catch minimize_tag (run_minimize args named_thms) id st
+      else ()
+    val _ = if is_some (!named_thms)
+      then Mirabelle.catch metis_tag (run_metis metis_fns args (these (!named_thms))) id st
+      else ()
+  in () end
+
+fun invoke args =
+  let
+    val _ = AtpManager.set_full_types (AList.defined (op =) args full_typesK)
+  in Mirabelle.register (init, sledgehammer_action args, done) end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mirabelle/doc/options.txt	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,9 @@
+Options for sledgehammer:
+
+  * prover=NAME: name of the external prover to call
+  * prover_timeout=TIME: timeout for invoked ATP
+  * keep=PATH: path where to keep temporary files created by sledgehammer 
+  * full_types: enable full-typed encoding
+  * minimize: enable minimization of theorem set found by sledgehammer
+  * minimize_timeout=TIME: timeout for each minimization step
+  * metis: apply metis with the theorems found by sledgehammer
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mirabelle/etc/settings	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,8 @@
+MIRABELLE_HOME="$COMPONENT"
+
+MIRABELLE_LOGIC=HOL
+MIRABELLE_THEORY=Main
+MIRABELLE_OUTPUT_PATH=/tmp/mirabelle
+MIRABELLE_TIMEOUT=30
+
+ISABELLE_TOOLS="$ISABELLE_TOOLS:$COMPONENT/lib/Tools"
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mirabelle/lib/Tools/mirabelle	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,97 @@
+#!/usr/bin/env bash
+#
+# Author: Sascha Boehme
+#
+# DESCRIPTION: testing tool for automated proof tools
+
+
+PRG="$(basename "$0")"
+
+function print_action_names() {
+  TOOLS="$MIRABELLE_HOME/Tools/mirabelle_*.ML"
+  for NAME in $TOOLS
+  do
+    echo $NAME | sed 's/.*mirabelle_\(.*\)\.ML/    \1/'
+  done
+}
+
+function usage() {
+  out="$MIRABELLE_OUTPUT_PATH"
+  timeout="$MIRABELLE_TIMEOUT"
+  echo
+  echo "Usage: isabelle $PRG [OPTIONS] ACTIONS FILES"
+  echo
+  echo "  Options are:"
+  echo "    -L LOGIC     parent logic to use (default $ISABELLE_LOGIC)"
+  echo "    -T THEORY    parent theory to use (default $MIRABELLE_THEORY)"
+  echo "    -O DIR       output directory for test data (default $out)"
+  echo "    -t TIMEOUT   timeout for each action in seconds (default $timeout)"
+  echo "    -q           be quiet (suppress output of Isabelle process)" 
+  echo
+  echo "  Apply the given actions (i.e., automated proof tools)"
+  echo "  at all proof steps in the given theory files."
+  echo
+  echo "  ACTIONS is a colon-separated list of actions, where each action is"
+  echo "  either NAME or NAME[OPTION,...,OPTION]. Available actions are:"
+  print_action_names
+  echo
+  echo "  A list of available OPTIONs can be found here:"
+  echo "    $MIRABELLE_HOME/doc/options.txt"
+  echo
+  echo "  FILES is a list of theory files, where each file is either NAME.thy"
+  echo "  or NAME.thy[START:END] and START and END are numbers indicating the"
+  echo "  range the given actions are to be applied."
+  echo
+  exit 1
+}
+
+
+## process command line
+
+# options
+
+while getopts "L:T:O:t:q?" OPT
+do
+  case "$OPT" in
+    L)
+      MIRABELLE_LOGIC="$OPTARG"
+      ;;
+    T)
+      MIRABELLE_THEORY="$OPTARG"
+      ;;
+    O)
+      MIRABELLE_OUTPUT_PATH="$OPTARG"
+      ;;
+    t)
+      MIRABELLE_TIMEOUT="$OPTARG"
+      ;;
+    q)
+      MIRABELLE_QUIET="true"
+      ;;
+    \?)
+      usage
+      ;;
+  esac
+done
+
+export MIRABELLE_QUIET
+
+shift $(($OPTIND - 1))
+
+export MIRABELLE_ACTIONS="$1"
+
+shift
+
+
+# setup
+
+mkdir -p "$MIRABELLE_OUTPUT_PATH"
+
+
+## main
+
+for FILE in "$@"
+do
+  perl -w "$MIRABELLE_HOME/lib/scripts/mirabelle.pl" "$FILE"
+done
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mirabelle/lib/scripts/mirabelle.pl	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,147 @@
+#
+# Author: Jasmin Blanchette and Sascha Boehme
+#
+# Testing tool for automated proof tools.
+#
+
+use File::Basename;
+
+# environment
+
+my $isabelle_home = $ENV{'ISABELLE_HOME'};
+my $mirabelle_home = $ENV{'MIRABELLE_HOME'};
+my $mirabelle_logic = $ENV{'MIRABELLE_LOGIC'};
+my $mirabelle_theory = $ENV{'MIRABELLE_THEORY'};
+my $output_path = $ENV{'MIRABELLE_OUTPUT_PATH'};
+my $timeout = $ENV{'MIRABELLE_TIMEOUT'};
+my $be_quiet = $ENV{'MIRABELLE_QUIET'};
+my $actions = $ENV{'MIRABELLE_ACTIONS'};
+
+my $mirabelle_thy = $mirabelle_home . "/Mirabelle";
+
+
+# arguments
+
+my $thy_file = $ARGV[0];
+my $start_line = "0";
+my $end_line = "~1";
+if ($thy_file =~ /^(.*)\[([0-9]+)\:(~?[0-9]+)\]$/) {
+  $thy_file = $1;
+  $start_line = $2;
+  $end_line = $3;
+}
+my ($thy_name, $path, $ext) = fileparse($thy_file, ".thy");
+my $new_thy_name = $thy_name . "_Mirabelle";
+my $new_thy_file = $output_path . "/" . $new_thy_name . $ext;
+
+
+# setup
+
+my $setup_thy_name = $thy_name . "_Setup";
+my $setup_file = $output_path . "/" . $setup_thy_name . ".thy";
+my $log_file = $output_path . "/" . $thy_name . ".log";
+
+my @action_files;
+my @action_names;
+foreach (split(/:/, $actions)) {
+  if (m/([^[]*)/) {
+    push @action_files, "\"$mirabelle_home/Tools/mirabelle_$1.ML\"";
+    push @action_names, $1;
+  }
+}
+my $tools = "";
+if ($#action_files >= 0) {
+  $tools = "uses " . join(" ", @action_files);
+}
+
+open(SETUP_FILE, ">$setup_file") || die "Could not create file '$setup_file'";
+
+print SETUP_FILE <<END;
+theory "$setup_thy_name"
+imports "$mirabelle_thy" "$mirabelle_theory"
+$tools
+begin
+
+setup {* 
+  Config.put_thy Mirabelle.logfile "$log_file" #>
+  Config.put_thy Mirabelle.timeout $timeout #>
+  Config.put_thy Mirabelle.start_line $start_line #>
+  Config.put_thy Mirabelle.end_line $end_line
+*}
+
+END
+
+foreach (split(/:/, $actions)) {
+  if (m/([^[]*)(?:\[(.*)\])?/) {
+    my ($name, $settings_str) = ($1, $2 || "");
+    $name =~ s/^([a-z])/\U$1/;
+    print SETUP_FILE "setup {* Mirabelle_$name.invoke [";
+    my $sep = "";
+    foreach (split(/,/, $settings_str)) {
+      if (m/\s*(.*)\s*=\s*(.*)\s*/) {
+        print SETUP_FILE "$sep(\"$1\", \"$2\")";
+        $sep = ", ";
+      }
+      elsif (m/\s*(.*)\s*/) {
+	print SETUP_FILE "$sep(\"$1\", \"\")";
+	$sep = ", ";
+      }
+    }
+    print SETUP_FILE "] *}\n";
+  }
+}
+
+print SETUP_FILE "\nend";
+close SETUP_FILE;
+
+
+# modify target theory file
+
+open(OLD_FILE, "<$thy_file") || die "Cannot open file '$thy_file'";
+my @lines = <OLD_FILE>;
+close(OLD_FILE);
+
+my $thy_text = join("", @lines);
+my $old_len = length($thy_text);
+$thy_text =~ s/(theory\s+)\"?$thy_name\"?/$1"$new_thy_name"/g;
+$thy_text =~ s/(imports)(\s+)/$1 "$setup_thy_name"$2/g;
+die "No 'imports' found" if length($thy_text) == $old_len;
+
+open(NEW_FILE, ">$new_thy_file") || die "Cannot create file '$new_thy_file'";
+print NEW_FILE $thy_text;
+close(NEW_FILE);
+
+my $root_file = "$output_path/ROOT_$thy_name.ML";
+open(ROOT_FILE, ">$root_file") || die "Cannot create file '$root_file'";
+print ROOT_FILE "use_thy \"$output_path/$new_thy_name\";\n";
+close(ROOT_FILE);
+
+
+# run isabelle
+
+open(LOG_FILE, ">$log_file");
+print LOG_FILE "Run of $new_thy_file with:\n";
+foreach $name (@action_names) {
+  print LOG_FILE "  $name\n";
+}
+close(LOG_FILE);
+
+my $quiet = "";
+if (defined $be_quiet and $be_quiet ne "") {
+  $quiet = " > /dev/null 2>&1";
+}
+
+print "Mirabelle: $thy_file\n" if ($quiet ne "");
+
+my $result = system "\"$ENV{'ISABELLE_PROCESS'}\" " .
+  "-e 'use \"$root_file\";' -q $mirabelle_logic" . $quiet;
+
+print "Finished:  $thy_file\n" if ($quiet ne "");
+
+
+# cleanup
+
+unlink $root_file;
+unlink $setup_file;
+
+exit $result;
--- a/src/HOL/NSA/Examples/NSPrimes.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/NSA/Examples/NSPrimes.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -7,7 +7,7 @@
 header{*The Nonstandard Primes as an Extension of the Prime Numbers*}
 
 theory NSPrimes
-imports "~~/src/HOL/NumberTheory/Factorization" Hyperreal
+imports "~~/src/HOL/Old_Number_Theory/Factorization" Hyperreal
 begin
 
 text{*These can be used to derive an alternative proof of the infinitude of
--- a/src/HOL/Nat.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Nat.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -1512,7 +1512,7 @@
 by (simp split add: nat_diff_split)
 
 lemma min_diff: "min (m - (i::nat)) (n - i) = min m n - i"
-unfolding min_def by auto
+by auto
 
 lemma inj_on_diff_nat: 
   assumes k_le_n: "\<forall>n \<in> N. k \<le> (n::nat)"
@@ -1588,9 +1588,6 @@
 lemma zero_induct: "P k ==> (!!n. P (Suc n) ==> P n) ==> P 0"
   using inc_induct[of 0 k P] by blast
 
-lemma nat_not_singleton: "(\<forall>x. x = (0::nat)) = False"
-  by auto
-
 (*The others are
       i - j - k = i - (j + k),
       k \<le> j ==> j - k + i = j + i - k,
--- a/src/HOL/NewNumberTheory/Binomial.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,373 +0,0 @@
-(*  Title:      Binomial.thy
-    Authors:    Lawrence C. Paulson, Jeremy Avigad, Tobias Nipkow
-
-
-Defines the "choose" function, and establishes basic properties.
-
-The original theory "Binomial" was by Lawrence C. Paulson, based on
-the work of Andy Gordon and Florian Kammueller. The approach here,
-which derives the definition of binomial coefficients in terms of the
-factorial function, is due to Jeremy Avigad. The binomial theorem was
-formalized by Tobias Nipkow.
-
-*)
-
-
-header {* Binomial *}
-
-theory Binomial
-imports Cong Fact
-begin
-
-
-subsection {* Main definitions *}
-
-class binomial =
-
-fixes 
-  binomial :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "choose" 65)
-
-(* definitions for the natural numbers *)
-
-instantiation nat :: binomial
-
-begin 
-
-fun
-  binomial_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-where
-  "binomial_nat n k =
-   (if k = 0 then 1 else
-    if n = 0 then 0 else
-      (binomial (n - 1) k) + (binomial (n - 1) (k - 1)))"
-
-instance proof qed
-
-end
-
-(* definitions for the integers *)
-
-instantiation int :: binomial
-
-begin 
-
-definition
-  binomial_int :: "int => int \<Rightarrow> int"
-where
-  "binomial_int n k = (if n \<ge> 0 \<and> k \<ge> 0 then int (binomial (nat n) (nat k))
-      else 0)"
-instance proof qed
-
-end
-
-
-subsection {* Set up Transfer *}
-
-lemma transfer_nat_int_binomial:
-  "(n::int) >= 0 \<Longrightarrow> k >= 0 \<Longrightarrow> binomial (nat n) (nat k) = 
-      nat (binomial n k)"
-  unfolding binomial_int_def 
-  by auto
-
-lemma transfer_nat_int_binomial_closure:
-  "n >= (0::int) \<Longrightarrow> k >= 0 \<Longrightarrow> binomial n k >= 0"
-  by (auto simp add: binomial_int_def)
-
-declare TransferMorphism_nat_int[transfer add return: 
-    transfer_nat_int_binomial transfer_nat_int_binomial_closure]
-
-lemma transfer_int_nat_binomial:
-  "binomial (int n) (int k) = int (binomial n k)"
-  unfolding fact_int_def binomial_int_def by auto
-
-lemma transfer_int_nat_binomial_closure:
-  "is_nat n \<Longrightarrow> is_nat k \<Longrightarrow> binomial n k >= 0"
-  by (auto simp add: binomial_int_def)
-
-declare TransferMorphism_int_nat[transfer add return: 
-    transfer_int_nat_binomial transfer_int_nat_binomial_closure]
-
-
-subsection {* Binomial coefficients *}
-
-lemma choose_zero_nat [simp]: "(n::nat) choose 0 = 1"
-  by simp
-
-lemma choose_zero_int [simp]: "n \<ge> 0 \<Longrightarrow> (n::int) choose 0 = 1"
-  by (simp add: binomial_int_def)
-
-lemma zero_choose_nat [rule_format,simp]: "ALL (k::nat) > n. n choose k = 0"
-  by (induct n rule: induct'_nat, auto)
-
-lemma zero_choose_int [rule_format,simp]: "(k::int) > n \<Longrightarrow> n choose k = 0"
-  unfolding binomial_int_def apply (case_tac "n < 0")
-  apply force
-  apply (simp del: binomial_nat.simps)
-done
-
-lemma choose_reduce_nat: "(n::nat) > 0 \<Longrightarrow> 0 < k \<Longrightarrow>
-    (n choose k) = ((n - 1) choose k) + ((n - 1) choose (k - 1))"
-  by simp
-
-lemma choose_reduce_int: "(n::int) > 0 \<Longrightarrow> 0 < k \<Longrightarrow>
-    (n choose k) = ((n - 1) choose k) + ((n - 1) choose (k - 1))"
-  unfolding binomial_int_def apply (subst choose_reduce_nat)
-    apply (auto simp del: binomial_nat.simps 
-      simp add: nat_diff_distrib)
-done
-
-lemma choose_plus_one_nat: "((n::nat) + 1) choose (k + 1) = 
-    (n choose (k + 1)) + (n choose k)"
-  by (simp add: choose_reduce_nat)
-
-lemma choose_Suc_nat: "(Suc n) choose (Suc k) = 
-    (n choose (Suc k)) + (n choose k)"
-  by (simp add: choose_reduce_nat One_nat_def)
-
-lemma choose_plus_one_int: "n \<ge> 0 \<Longrightarrow> k \<ge> 0 \<Longrightarrow> ((n::int) + 1) choose (k + 1) = 
-    (n choose (k + 1)) + (n choose k)"
-  by (simp add: binomial_int_def choose_plus_one_nat nat_add_distrib del: binomial_nat.simps)
-
-declare binomial_nat.simps [simp del]
-
-lemma choose_self_nat [simp]: "((n::nat) choose n) = 1"
-  by (induct n rule: induct'_nat, auto simp add: choose_plus_one_nat)
-
-lemma choose_self_int [simp]: "n \<ge> 0 \<Longrightarrow> ((n::int) choose n) = 1"
-  by (auto simp add: binomial_int_def)
-
-lemma choose_one_nat [simp]: "(n::nat) choose 1 = n"
-  by (induct n rule: induct'_nat, auto simp add: choose_reduce_nat)
-
-lemma choose_one_int [simp]: "n \<ge> 0 \<Longrightarrow> (n::int) choose 1 = n"
-  by (auto simp add: binomial_int_def)
-
-lemma plus_one_choose_self_nat [simp]: "(n::nat) + 1 choose n = n + 1"
-  apply (induct n rule: induct'_nat, force)
-  apply (case_tac "n = 0")
-  apply auto
-  apply (subst choose_reduce_nat)
-  apply (auto simp add: One_nat_def)  
-  (* natdiff_cancel_numerals introduces Suc *)
-done
-
-lemma Suc_choose_self_nat [simp]: "(Suc n) choose n = Suc n"
-  using plus_one_choose_self_nat by (simp add: One_nat_def)
-
-lemma plus_one_choose_self_int [rule_format, simp]: 
-    "(n::int) \<ge> 0 \<longrightarrow> n + 1 choose n = n + 1"
-   by (auto simp add: binomial_int_def nat_add_distrib)
-
-(* bounded quantification doesn't work with the unicode characters? *)
-lemma choose_pos_nat [rule_format]: "ALL k <= (n::nat). 
-    ((n::nat) choose k) > 0"
-  apply (induct n rule: induct'_nat) 
-  apply force
-  apply clarify
-  apply (case_tac "k = 0")
-  apply force
-  apply (subst choose_reduce_nat)
-  apply auto
-done
-
-lemma choose_pos_int: "n \<ge> 0 \<Longrightarrow> k >= 0 \<Longrightarrow> k \<le> n \<Longrightarrow>
-    ((n::int) choose k) > 0"
-  by (auto simp add: binomial_int_def choose_pos_nat)
-
-lemma binomial_induct [rule_format]: "(ALL (n::nat). P n n) \<longrightarrow> 
-    (ALL n. P (n + 1) 0) \<longrightarrow> (ALL n. (ALL k < n. P n k \<longrightarrow> P n (k + 1) \<longrightarrow>
-    P (n + 1) (k + 1))) \<longrightarrow> (ALL k <= n. P n k)"
-  apply (induct n rule: induct'_nat)
-  apply auto
-  apply (case_tac "k = 0")
-  apply auto
-  apply (case_tac "k = n + 1")
-  apply auto
-  apply (drule_tac x = n in spec) back back 
-  apply (drule_tac x = "k - 1" in spec) back back back
-  apply auto
-done
-
-lemma choose_altdef_aux_nat: "(k::nat) \<le> n \<Longrightarrow> 
-    fact k * fact (n - k) * (n choose k) = fact n"
-  apply (rule binomial_induct [of _ k n])
-  apply auto
-proof -
-  fix k :: nat and n
-  assume less: "k < n"
-  assume ih1: "fact k * fact (n - k) * (n choose k) = fact n"
-  hence one: "fact (k + 1) * fact (n - k) * (n choose k) = (k + 1) * fact n"
-    by (subst fact_plus_one_nat, auto)
-  assume ih2: "fact (k + 1) * fact (n - (k + 1)) * (n choose (k + 1)) = 
-      fact n"
-  with less have "fact (k + 1) * fact ((n - (k + 1)) + 1) * 
-      (n choose (k + 1)) = (n - k) * fact n"
-    by (subst (2) fact_plus_one_nat, auto)
-  with less have two: "fact (k + 1) * fact (n - k) * (n choose (k + 1)) = 
-      (n - k) * fact n" by simp
-  have "fact (k + 1) * fact (n - k) * (n + 1 choose (k + 1)) =
-      fact (k + 1) * fact (n - k) * (n choose (k + 1)) + 
-      fact (k + 1) * fact (n - k) * (n choose k)" 
-    by (subst choose_reduce_nat, auto simp add: ring_simps)
-  also note one
-  also note two
-  also with less have "(n - k) * fact n + (k + 1) * fact n= fact (n + 1)" 
-    apply (subst fact_plus_one_nat)
-    apply (subst left_distrib [symmetric])
-    apply simp
-    done
-  finally show "fact (k + 1) * fact (n - k) * (n + 1 choose (k + 1)) = 
-    fact (n + 1)" .
-qed
-
-lemma choose_altdef_nat: "(k::nat) \<le> n \<Longrightarrow> 
-    n choose k = fact n div (fact k * fact (n - k))"
-  apply (frule choose_altdef_aux_nat)
-  apply (erule subst)
-  apply (simp add: mult_ac)
-done
-
-
-lemma choose_altdef_int: 
-  assumes "(0::int) <= k" and "k <= n"
-  shows "n choose k = fact n div (fact k * fact (n - k))"
-  
-  apply (subst tsub_eq [symmetric], rule prems)
-  apply (rule choose_altdef_nat [transferred])
-  using prems apply auto
-done
-
-lemma choose_dvd_nat: "(k::nat) \<le> n \<Longrightarrow> fact k * fact (n - k) dvd fact n"
-  unfolding dvd_def apply (frule choose_altdef_aux_nat)
-  (* why don't blast and auto get this??? *)
-  apply (rule exI)
-  apply (erule sym)
-done
-
-lemma choose_dvd_int: 
-  assumes "(0::int) <= k" and "k <= n"
-  shows "fact k * fact (n - k) dvd fact n"
- 
-  apply (subst tsub_eq [symmetric], rule prems)
-  apply (rule choose_dvd_nat [transferred])
-  using prems apply auto
-done
-
-(* generalizes Tobias Nipkow's proof to any commutative semiring *)
-theorem binomial: "(a+b::'a::{comm_ring_1,power})^n = 
-  (SUM k=0..n. (of_nat (n choose k)) * a^k * b^(n-k))" (is "?P n")
-proof (induct n rule: induct'_nat)
-  show "?P 0" by simp
-next
-  fix n
-  assume ih: "?P n"
-  have decomp: "{0..n+1} = {0} Un {n+1} Un {1..n}"
-    by auto
-  have decomp2: "{0..n} = {0} Un {1..n}"
-    by auto
-  have decomp3: "{1..n+1} = {n+1} Un {1..n}"
-    by auto
-  have "(a+b)^(n+1) = 
-      (a+b) * (SUM k=0..n. of_nat (n choose k) * a^k * b^(n-k))"
-    using ih by (simp add: power_plus_one)
-  also have "... =  a*(SUM k=0..n. of_nat (n choose k) * a^k * b^(n-k)) +
-                   b*(SUM k=0..n. of_nat (n choose k) * a^k * b^(n-k))"
-    by (rule distrib)
-  also have "... = (SUM k=0..n. of_nat (n choose k) * a^(k+1) * b^(n-k)) +
-                  (SUM k=0..n. of_nat (n choose k) * a^k * b^(n-k+1))"
-    by (subst (1 2) power_plus_one, simp add: setsum_right_distrib mult_ac)
-  also have "... = (SUM k=0..n. of_nat (n choose k) * a^k * b^(n+1-k)) +
-                  (SUM k=1..n+1. of_nat (n choose (k - 1)) * a^k * b^(n+1-k))"
-    by (simp add:setsum_shift_bounds_cl_Suc_ivl Suc_diff_le 
-             power_Suc ring_simps One_nat_def del:setsum_cl_ivl_Suc)
-  also have "... = a^(n+1) + b^(n+1) +
-                  (SUM k=1..n. of_nat (n choose (k - 1)) * a^k * b^(n+1-k)) +
-                  (SUM k=1..n. of_nat (n choose k) * a^k * b^(n+1-k))"
-    by (simp add: decomp2 decomp3)
-  also have
-      "... = a^(n+1) + b^(n+1) + 
-         (SUM k=1..n. of_nat(n+1 choose k) * a^k * b^(n+1-k))"
-    by (auto simp add: ring_simps setsum_addf [symmetric]
-      choose_reduce_nat)
-  also have "... = (SUM k=0..n+1. of_nat (n+1 choose k) * a^k * b^(n+1-k))"
-    using decomp by (simp add: ring_simps)
-  finally show "?P (n + 1)" by simp
-qed
-
-lemma set_explicit: "{S. S = T \<and> P S} = (if P T then {T} else {})"
-  by auto
-
-lemma card_subsets_nat [rule_format]:
-  fixes S :: "'a set"
-  assumes "finite S"
-  shows "ALL k. card {T. T \<le> S \<and> card T = k} = card S choose k" 
-      (is "?P S")
-using `finite S`
-proof (induct set: finite)
-  show "?P {}" by (auto simp add: set_explicit)
-  next fix x :: "'a" and F
-  assume iassms: "finite F" "x ~: F"
-  assume ih: "?P F"
-  show "?P (insert x F)" (is "ALL k. ?Q k")
-  proof
-    fix k
-    show "card {T. T \<subseteq> (insert x F) \<and> card T = k} = 
-        card (insert x F) choose k" (is "?Q k")
-    proof (induct k rule: induct'_nat)
-      from iassms have "{T. T \<le> (insert x F) \<and> card T = 0} = {{}}"
-        apply auto
-        apply (subst (asm) card_0_eq)
-        apply (auto elim: finite_subset)
-        done
-      thus "?Q 0" 
-        by auto
-      next fix k
-      show "?Q (k + 1)"
-      proof -
-        from iassms have fin: "finite (insert x F)" by auto
-        hence "{ T. T \<subseteq> insert x F \<and> card T = k + 1} =
-          {T. T \<le> F & card T = k + 1} Un 
-          {T. T \<le> insert x F & x : T & card T = k + 1}"
-          by (auto intro!: subsetI)
-        with iassms fin have "card ({T. T \<le> insert x F \<and> card T = k + 1}) = 
-          card ({T. T \<subseteq> F \<and> card T = k + 1}) + 
-          card ({T. T \<subseteq> insert x F \<and> x : T \<and> card T = k + 1})"
-          apply (subst card_Un_disjoint [symmetric])
-          apply auto
-          (* note: nice! Didn't have to say anything here *)
-          done
-        also from ih have "card ({T. T \<subseteq> F \<and> card T = k + 1}) = 
-          card F choose (k+1)" by auto
-        also have "card ({T. T \<subseteq> insert x F \<and> x : T \<and> card T = k + 1}) =
-          card ({T. T <= F & card T = k})"
-        proof -
-          let ?f = "%T. T Un {x}"
-          from iassms have "inj_on ?f {T. T <= F & card T = k}"
-            unfolding inj_on_def by (auto intro!: subsetI)
-          hence "card ({T. T <= F & card T = k}) = 
-            card(?f ` {T. T <= F & card T = k})"
-            by (rule card_image [symmetric])
-          also from iassms fin have "?f ` {T. T <= F & card T = k} = 
-            {T. T \<subseteq> insert x F \<and> x : T \<and> card T = k + 1}"
-            unfolding image_def 
-            (* I can't figure out why this next line takes so long *)
-            apply auto
-            apply (frule (1) finite_subset, force)
-            apply (rule_tac x = "xa - {x}" in exI)
-            apply (subst card_Diff_singleton)
-            apply (auto elim: finite_subset)
-            done
-          finally show ?thesis by (rule sym)
-        qed
-        also from ih have "card ({T. T <= F & card T = k}) = card F choose k"
-          by auto
-        finally have "card ({T. T \<le> insert x F \<and> card T = k + 1}) = 
-          card F choose (k + 1) + (card F choose k)".
-        with iassms choose_plus_one_nat show ?thesis
-          by auto
-      qed
-    qed
-  qed
-qed
-
-end
--- a/src/HOL/NewNumberTheory/Cong.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1091 +0,0 @@
-(*  Title:      HOL/Library/Cong.thy
-    ID:         
-    Authors:    Christophe Tabacznyj, Lawrence C. Paulson, Amine Chaieb,
-                Thomas M. Rasmussen, Jeremy Avigad
-
-
-Defines congruence (notation: [x = y] (mod z)) for natural numbers and
-integers.
-
-This file combines and revises a number of prior developments.
-
-The original theories "GCD" and "Primes" were by Christophe Tabacznyj
-and Lawrence C. Paulson, based on \cite{davenport92}. They introduced
-gcd, lcm, and prime for the natural numbers.
-
-The original theory "IntPrimes" was by Thomas M. Rasmussen, and
-extended gcd, lcm, primes to the integers. Amine Chaieb provided
-another extension of the notions to the integers, and added a number
-of results to "Primes" and "GCD". 
-
-The original theory, "IntPrimes", by Thomas M. Rasmussen, defined and
-developed the congruence relations on the integers. The notion was
-extended to the natural numbers by Chiaeb. Jeremy Avigad combined
-these, revised and tidied them, made the development uniform for the
-natural numbers and the integers, and added a number of new theorems.
-
-*)
-
-
-header {* Congruence *}
-
-theory Cong
-imports GCD
-begin
-
-subsection {* Turn off One_nat_def *}
-
-lemma induct'_nat [case_names zero plus1, induct type: nat]: 
-    "\<lbrakk> P (0::nat); !!n. P n \<Longrightarrow> P (n + 1)\<rbrakk> \<Longrightarrow> P n"
-by (erule nat_induct) (simp add:One_nat_def)
-
-lemma cases_nat [case_names zero plus1, cases type: nat]: 
-    "P (0::nat) \<Longrightarrow> (!!n. P (n + 1)) \<Longrightarrow> P n"
-by(metis induct'_nat)
-
-lemma power_plus_one [simp]: "(x::'a::power)^(n + 1) = x * x^n"
-by (simp add: One_nat_def)
-
-lemma power_eq_one_eq_nat [simp]: 
-  "((x::nat)^m = 1) = (m = 0 | x = 1)"
-by (induct m, auto)
-
-lemma card_insert_if' [simp]: "finite A \<Longrightarrow>
-  card (insert x A) = (if x \<in> A then (card A) else (card A) + 1)"
-by (auto simp add: insert_absorb)
-
-(* why wasn't card_insert_if a simp rule? *)
-declare card_insert_disjoint [simp del]
-
-lemma nat_1' [simp]: "nat 1 = 1"
-by simp
-
-(* For those annoying moments where Suc reappears, use Suc_eq_plus1 *)
-
-declare nat_1 [simp del]
-declare add_2_eq_Suc [simp del] 
-declare add_2_eq_Suc' [simp del]
-
-
-declare mod_pos_pos_trivial [simp]
-
-
-subsection {* Main definitions *}
-
-class cong =
-
-fixes 
-  cong :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" ("(1[_ = _] '(mod _'))")
-
-begin
-
-abbreviation
-  notcong :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" ("(1[_ \<noteq> _] '(mod _'))")
-where
-  "notcong x y m == (~cong x y m)" 
-
-end
-
-(* definitions for the natural numbers *)
-
-instantiation nat :: cong
-
-begin 
-
-definition 
-  cong_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool"
-where 
-  "cong_nat x y m = ((x mod m) = (y mod m))"
-
-instance proof qed
-
-end
-
-
-(* definitions for the integers *)
-
-instantiation int :: cong
-
-begin 
-
-definition 
-  cong_int :: "int \<Rightarrow> int \<Rightarrow> int \<Rightarrow> bool"
-where 
-  "cong_int x y m = ((x mod m) = (y mod m))"
-
-instance proof qed
-
-end
-
-
-subsection {* Set up Transfer *}
-
-
-lemma transfer_nat_int_cong:
-  "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> m >= 0 \<Longrightarrow> 
-    ([(nat x) = (nat y)] (mod (nat m))) = ([x = y] (mod m))"
-  unfolding cong_int_def cong_nat_def 
-  apply (auto simp add: nat_mod_distrib [symmetric])
-  apply (subst (asm) eq_nat_nat_iff)
-  apply (case_tac "m = 0", force, rule pos_mod_sign, force)+
-  apply assumption
-done
-
-declare TransferMorphism_nat_int[transfer add return: 
-    transfer_nat_int_cong]
-
-lemma transfer_int_nat_cong:
-  "[(int x) = (int y)] (mod (int m)) = [x = y] (mod m)"
-  apply (auto simp add: cong_int_def cong_nat_def)
-  apply (auto simp add: zmod_int [symmetric])
-done
-
-declare TransferMorphism_int_nat[transfer add return: 
-    transfer_int_nat_cong]
-
-
-subsection {* Congruence *}
-
-(* was zcong_0, etc. *)
-lemma cong_0_nat [simp, presburger]: "([(a::nat) = b] (mod 0)) = (a = b)"
-  by (unfold cong_nat_def, auto)
-
-lemma cong_0_int [simp, presburger]: "([(a::int) = b] (mod 0)) = (a = b)"
-  by (unfold cong_int_def, auto)
-
-lemma cong_1_nat [simp, presburger]: "[(a::nat) = b] (mod 1)"
-  by (unfold cong_nat_def, auto)
-
-lemma cong_Suc_0_nat [simp, presburger]: "[(a::nat) = b] (mod Suc 0)"
-  by (unfold cong_nat_def, auto simp add: One_nat_def)
-
-lemma cong_1_int [simp, presburger]: "[(a::int) = b] (mod 1)"
-  by (unfold cong_int_def, auto)
-
-lemma cong_refl_nat [simp]: "[(k::nat) = k] (mod m)"
-  by (unfold cong_nat_def, auto)
-
-lemma cong_refl_int [simp]: "[(k::int) = k] (mod m)"
-  by (unfold cong_int_def, auto)
-
-lemma cong_sym_nat: "[(a::nat) = b] (mod m) \<Longrightarrow> [b = a] (mod m)"
-  by (unfold cong_nat_def, auto)
-
-lemma cong_sym_int: "[(a::int) = b] (mod m) \<Longrightarrow> [b = a] (mod m)"
-  by (unfold cong_int_def, auto)
-
-lemma cong_sym_eq_nat: "[(a::nat) = b] (mod m) = [b = a] (mod m)"
-  by (unfold cong_nat_def, auto)
-
-lemma cong_sym_eq_int: "[(a::int) = b] (mod m) = [b = a] (mod m)"
-  by (unfold cong_int_def, auto)
-
-lemma cong_trans_nat [trans]:
-    "[(a::nat) = b] (mod m) \<Longrightarrow> [b = c] (mod m) \<Longrightarrow> [a = c] (mod m)"
-  by (unfold cong_nat_def, auto)
-
-lemma cong_trans_int [trans]:
-    "[(a::int) = b] (mod m) \<Longrightarrow> [b = c] (mod m) \<Longrightarrow> [a = c] (mod m)"
-  by (unfold cong_int_def, auto)
-
-lemma cong_add_nat:
-    "[(a::nat) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a + c = b + d] (mod m)"
-  apply (unfold cong_nat_def)
-  apply (subst (1 2) mod_add_eq)
-  apply simp
-done
-
-lemma cong_add_int:
-    "[(a::int) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a + c = b + d] (mod m)"
-  apply (unfold cong_int_def)
-  apply (subst (1 2) mod_add_left_eq)
-  apply (subst (1 2) mod_add_right_eq)
-  apply simp
-done
-
-lemma cong_diff_int:
-    "[(a::int) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a - c = b - d] (mod m)"
-  apply (unfold cong_int_def)
-  apply (subst (1 2) mod_diff_eq)
-  apply simp
-done
-
-lemma cong_diff_aux_int:
-  "(a::int) >= c \<Longrightarrow> b >= d \<Longrightarrow> [(a::int) = b] (mod m) \<Longrightarrow> 
-      [c = d] (mod m) \<Longrightarrow> [tsub a c = tsub b d] (mod m)"
-  apply (subst (1 2) tsub_eq)
-  apply (auto intro: cong_diff_int)
-done;
-
-lemma cong_diff_nat:
-  assumes "(a::nat) >= c" and "b >= d" and "[a = b] (mod m)" and
-    "[c = d] (mod m)"
-  shows "[a - c = b - d] (mod m)"
-
-  using prems by (rule cong_diff_aux_int [transferred]);
-
-lemma cong_mult_nat:
-    "[(a::nat) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a * c = b * d] (mod m)"
-  apply (unfold cong_nat_def)
-  apply (subst (1 2) mod_mult_eq)
-  apply simp
-done
-
-lemma cong_mult_int:
-    "[(a::int) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a * c = b * d] (mod m)"
-  apply (unfold cong_int_def)
-  apply (subst (1 2) zmod_zmult1_eq)
-  apply (subst (1 2) mult_commute)
-  apply (subst (1 2) zmod_zmult1_eq)
-  apply simp
-done
-
-lemma cong_exp_nat: "[(x::nat) = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"
-  apply (induct k)
-  apply (auto simp add: cong_refl_nat cong_mult_nat)
-done
-
-lemma cong_exp_int: "[(x::int) = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"
-  apply (induct k)
-  apply (auto simp add: cong_refl_int cong_mult_int)
-done
-
-lemma cong_setsum_nat [rule_format]: 
-    "(ALL x: A. [((f x)::nat) = g x] (mod m)) \<longrightarrow> 
-      [(SUM x:A. f x) = (SUM x:A. g x)] (mod m)"
-  apply (case_tac "finite A")
-  apply (induct set: finite)
-  apply (auto intro: cong_add_nat)
-done
-
-lemma cong_setsum_int [rule_format]:
-    "(ALL x: A. [((f x)::int) = g x] (mod m)) \<longrightarrow> 
-      [(SUM x:A. f x) = (SUM x:A. g x)] (mod m)"
-  apply (case_tac "finite A")
-  apply (induct set: finite)
-  apply (auto intro: cong_add_int)
-done
-
-lemma cong_setprod_nat [rule_format]: 
-    "(ALL x: A. [((f x)::nat) = g x] (mod m)) \<longrightarrow> 
-      [(PROD x:A. f x) = (PROD x:A. g x)] (mod m)"
-  apply (case_tac "finite A")
-  apply (induct set: finite)
-  apply (auto intro: cong_mult_nat)
-done
-
-lemma cong_setprod_int [rule_format]: 
-    "(ALL x: A. [((f x)::int) = g x] (mod m)) \<longrightarrow> 
-      [(PROD x:A. f x) = (PROD x:A. g x)] (mod m)"
-  apply (case_tac "finite A")
-  apply (induct set: finite)
-  apply (auto intro: cong_mult_int)
-done
-
-lemma cong_scalar_nat: "[(a::nat)= b] (mod m) \<Longrightarrow> [a * k = b * k] (mod m)"
-  by (rule cong_mult_nat, simp_all)
-
-lemma cong_scalar_int: "[(a::int)= b] (mod m) \<Longrightarrow> [a * k = b * k] (mod m)"
-  by (rule cong_mult_int, simp_all)
-
-lemma cong_scalar2_nat: "[(a::nat)= b] (mod m) \<Longrightarrow> [k * a = k * b] (mod m)"
-  by (rule cong_mult_nat, simp_all)
-
-lemma cong_scalar2_int: "[(a::int)= b] (mod m) \<Longrightarrow> [k * a = k * b] (mod m)"
-  by (rule cong_mult_int, simp_all)
-
-lemma cong_mult_self_nat: "[(a::nat) * m = 0] (mod m)"
-  by (unfold cong_nat_def, auto)
-
-lemma cong_mult_self_int: "[(a::int) * m = 0] (mod m)"
-  by (unfold cong_int_def, auto)
-
-lemma cong_eq_diff_cong_0_int: "[(a::int) = b] (mod m) = [a - b = 0] (mod m)"
-  apply (rule iffI)
-  apply (erule cong_diff_int [of a b m b b, simplified])
-  apply (erule cong_add_int [of "a - b" 0 m b b, simplified])
-done
-
-lemma cong_eq_diff_cong_0_aux_int: "a >= b \<Longrightarrow>
-    [(a::int) = b] (mod m) = [tsub a b = 0] (mod m)"
-  by (subst tsub_eq, assumption, rule cong_eq_diff_cong_0_int)
-
-lemma cong_eq_diff_cong_0_nat:
-  assumes "(a::nat) >= b"
-  shows "[a = b] (mod m) = [a - b = 0] (mod m)"
-
-  using prems by (rule cong_eq_diff_cong_0_aux_int [transferred])
-
-lemma cong_diff_cong_0'_nat: 
-  "[(x::nat) = y] (mod n) \<longleftrightarrow> 
-    (if x <= y then [y - x = 0] (mod n) else [x - y = 0] (mod n))"
-  apply (case_tac "y <= x")
-  apply (frule cong_eq_diff_cong_0_nat [where m = n])
-  apply auto [1]
-  apply (subgoal_tac "x <= y")
-  apply (frule cong_eq_diff_cong_0_nat [where m = n])
-  apply (subst cong_sym_eq_nat)
-  apply auto
-done
-
-lemma cong_altdef_nat: "(a::nat) >= b \<Longrightarrow> [a = b] (mod m) = (m dvd (a - b))"
-  apply (subst cong_eq_diff_cong_0_nat, assumption)
-  apply (unfold cong_nat_def)
-  apply (simp add: dvd_eq_mod_eq_0 [symmetric])
-done
-
-lemma cong_altdef_int: "[(a::int) = b] (mod m) = (m dvd (a - b))"
-  apply (subst cong_eq_diff_cong_0_int)
-  apply (unfold cong_int_def)
-  apply (simp add: dvd_eq_mod_eq_0 [symmetric])
-done
-
-lemma cong_abs_int: "[(x::int) = y] (mod abs m) = [x = y] (mod m)"
-  by (simp add: cong_altdef_int)
-
-lemma cong_square_int:
-   "\<lbrakk> prime (p::int); 0 < a; [a * a = 1] (mod p) \<rbrakk>
-    \<Longrightarrow> [a = 1] (mod p) \<or> [a = - 1] (mod p)"
-  apply (simp only: cong_altdef_int)
-  apply (subst prime_dvd_mult_eq_int [symmetric], assumption)
-  (* any way around this? *)
-  apply (subgoal_tac "a * a - 1 = (a - 1) * (a - -1)")
-  apply (auto simp add: ring_simps)
-done
-
-lemma cong_mult_rcancel_int:
-  "coprime k (m::int) \<Longrightarrow> [a * k = b * k] (mod m) = [a = b] (mod m)"
-  apply (subst (1 2) cong_altdef_int)
-  apply (subst left_diff_distrib [symmetric])
-  apply (rule coprime_dvd_mult_iff_int)
-  apply (subst gcd_commute_int, assumption)
-done
-
-lemma cong_mult_rcancel_nat:
-  assumes  "coprime k (m::nat)"
-  shows "[a * k = b * k] (mod m) = [a = b] (mod m)"
-
-  apply (rule cong_mult_rcancel_int [transferred])
-  using prems apply auto
-done
-
-lemma cong_mult_lcancel_nat:
-  "coprime k (m::nat) \<Longrightarrow> [k * a = k * b ] (mod m) = [a = b] (mod m)"
-  by (simp add: mult_commute cong_mult_rcancel_nat)
-
-lemma cong_mult_lcancel_int:
-  "coprime k (m::int) \<Longrightarrow> [k * a = k * b] (mod m) = [a = b] (mod m)"
-  by (simp add: mult_commute cong_mult_rcancel_int)
-
-(* was zcong_zgcd_zmult_zmod *)
-lemma coprime_cong_mult_int:
-  "[(a::int) = b] (mod m) \<Longrightarrow> [a = b] (mod n) \<Longrightarrow> coprime m n
-    \<Longrightarrow> [a = b] (mod m * n)"
-  apply (simp only: cong_altdef_int)
-  apply (erule (2) divides_mult_int)
-done
-
-lemma coprime_cong_mult_nat:
-  assumes "[(a::nat) = b] (mod m)" and "[a = b] (mod n)" and "coprime m n"
-  shows "[a = b] (mod m * n)"
-
-  apply (rule coprime_cong_mult_int [transferred])
-  using prems apply auto
-done
-
-lemma cong_less_imp_eq_nat: "0 \<le> (a::nat) \<Longrightarrow>
-    a < m \<Longrightarrow> 0 \<le> b \<Longrightarrow> b < m \<Longrightarrow> [a = b] (mod m) \<Longrightarrow> a = b"
-  by (auto simp add: cong_nat_def mod_pos_pos_trivial)
-
-lemma cong_less_imp_eq_int: "0 \<le> (a::int) \<Longrightarrow>
-    a < m \<Longrightarrow> 0 \<le> b \<Longrightarrow> b < m \<Longrightarrow> [a = b] (mod m) \<Longrightarrow> a = b"
-  by (auto simp add: cong_int_def mod_pos_pos_trivial)
-
-lemma cong_less_unique_nat:
-    "0 < (m::nat) \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))"
-  apply auto
-  apply (rule_tac x = "a mod m" in exI)
-  apply (unfold cong_nat_def, auto)
-done
-
-lemma cong_less_unique_int:
-    "0 < (m::int) \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))"
-  apply auto
-  apply (rule_tac x = "a mod m" in exI)
-  apply (unfold cong_int_def, auto simp add: mod_pos_pos_trivial)
-done
-
-lemma cong_iff_lin_int: "([(a::int) = b] (mod m)) = (\<exists>k. b = a + m * k)"
-  apply (auto simp add: cong_altdef_int dvd_def ring_simps)
-  apply (rule_tac [!] x = "-k" in exI, auto)
-done
-
-lemma cong_iff_lin_nat: "([(a::nat) = b] (mod m)) = 
-    (\<exists>k1 k2. b + k1 * m = a + k2 * m)"
-  apply (rule iffI)
-  apply (case_tac "b <= a")
-  apply (subst (asm) cong_altdef_nat, assumption)
-  apply (unfold dvd_def, auto)
-  apply (rule_tac x = k in exI)
-  apply (rule_tac x = 0 in exI)
-  apply (auto simp add: ring_simps)
-  apply (subst (asm) cong_sym_eq_nat)
-  apply (subst (asm) cong_altdef_nat)
-  apply force
-  apply (unfold dvd_def, auto)
-  apply (rule_tac x = 0 in exI)
-  apply (rule_tac x = k in exI)
-  apply (auto simp add: ring_simps)
-  apply (unfold cong_nat_def)
-  apply (subgoal_tac "a mod m = (a + k2 * m) mod m")
-  apply (erule ssubst)back
-  apply (erule subst)
-  apply auto
-done
-
-lemma cong_gcd_eq_int: "[(a::int) = b] (mod m) \<Longrightarrow> gcd a m = gcd b m"
-  apply (subst (asm) cong_iff_lin_int, auto)
-  apply (subst add_commute) 
-  apply (subst (2) gcd_commute_int)
-  apply (subst mult_commute)
-  apply (subst gcd_add_mult_int)
-  apply (rule gcd_commute_int)
-done
-
-lemma cong_gcd_eq_nat: 
-  assumes "[(a::nat) = b] (mod m)"
-  shows "gcd a m = gcd b m"
-
-  apply (rule cong_gcd_eq_int [transferred])
-  using prems apply auto
-done
-
-lemma cong_imp_coprime_nat: "[(a::nat) = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow> 
-    coprime b m"
-  by (auto simp add: cong_gcd_eq_nat)
-
-lemma cong_imp_coprime_int: "[(a::int) = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow> 
-    coprime b m"
-  by (auto simp add: cong_gcd_eq_int)
-
-lemma cong_cong_mod_nat: "[(a::nat) = b] (mod m) = 
-    [a mod m = b mod m] (mod m)"
-  by (auto simp add: cong_nat_def)
-
-lemma cong_cong_mod_int: "[(a::int) = b] (mod m) = 
-    [a mod m = b mod m] (mod m)"
-  by (auto simp add: cong_int_def)
-
-lemma cong_minus_int [iff]: "[(a::int) = b] (mod -m) = [a = b] (mod m)"
-  by (subst (1 2) cong_altdef_int, auto)
-
-lemma cong_zero_nat [iff]: "[(a::nat) = b] (mod 0) = (a = b)"
-  by (auto simp add: cong_nat_def)
-
-lemma cong_zero_int [iff]: "[(a::int) = b] (mod 0) = (a = b)"
-  by (auto simp add: cong_int_def)
-
-(*
-lemma mod_dvd_mod_int:
-    "0 < (m::int) \<Longrightarrow> m dvd b \<Longrightarrow> (a mod b mod m) = (a mod m)"
-  apply (unfold dvd_def, auto)
-  apply (rule mod_mod_cancel)
-  apply auto
-done
-
-lemma mod_dvd_mod:
-  assumes "0 < (m::nat)" and "m dvd b"
-  shows "(a mod b mod m) = (a mod m)"
-
-  apply (rule mod_dvd_mod_int [transferred])
-  using prems apply auto
-done
-*)
-
-lemma cong_add_lcancel_nat: 
-    "[(a::nat) + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)" 
-  by (simp add: cong_iff_lin_nat)
-
-lemma cong_add_lcancel_int: 
-    "[(a::int) + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)" 
-  by (simp add: cong_iff_lin_int)
-
-lemma cong_add_rcancel_nat: "[(x::nat) + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
-  by (simp add: cong_iff_lin_nat)
-
-lemma cong_add_rcancel_int: "[(x::int) + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
-  by (simp add: cong_iff_lin_int)
-
-lemma cong_add_lcancel_0_nat: "[(a::nat) + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" 
-  by (simp add: cong_iff_lin_nat)
-
-lemma cong_add_lcancel_0_int: "[(a::int) + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" 
-  by (simp add: cong_iff_lin_int)
-
-lemma cong_add_rcancel_0_nat: "[x + (a::nat) = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" 
-  by (simp add: cong_iff_lin_nat)
-
-lemma cong_add_rcancel_0_int: "[x + (a::int) = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" 
-  by (simp add: cong_iff_lin_int)
-
-lemma cong_dvd_modulus_nat: "[(x::nat) = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow> 
-    [x = y] (mod n)"
-  apply (auto simp add: cong_iff_lin_nat dvd_def)
-  apply (rule_tac x="k1 * k" in exI)
-  apply (rule_tac x="k2 * k" in exI)
-  apply (simp add: ring_simps)
-done
-
-lemma cong_dvd_modulus_int: "[(x::int) = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow> 
-    [x = y] (mod n)"
-  by (auto simp add: cong_altdef_int dvd_def)
-
-lemma cong_dvd_eq_nat: "[(x::nat) = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"
-  by (unfold cong_nat_def, auto simp add: dvd_eq_mod_eq_0)
-
-lemma cong_dvd_eq_int: "[(x::int) = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"
-  by (unfold cong_int_def, auto simp add: dvd_eq_mod_eq_0)
-
-lemma cong_mod_nat: "(n::nat) ~= 0 \<Longrightarrow> [a mod n = a] (mod n)" 
-  by (simp add: cong_nat_def)
-
-lemma cong_mod_int: "(n::int) ~= 0 \<Longrightarrow> [a mod n = a] (mod n)" 
-  by (simp add: cong_int_def)
-
-lemma mod_mult_cong_nat: "(a::nat) ~= 0 \<Longrightarrow> b ~= 0 
-    \<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)"
-  by (simp add: cong_nat_def mod_mult2_eq  mod_add_left_eq)
-
-lemma neg_cong_int: "([(a::int) = b] (mod m)) = ([-a = -b] (mod m))"
-  apply (simp add: cong_altdef_int)
-  apply (subst dvd_minus_iff [symmetric])
-  apply (simp add: ring_simps)
-done
-
-lemma cong_modulus_neg_int: "([(a::int) = b] (mod m)) = ([a = b] (mod -m))"
-  by (auto simp add: cong_altdef_int)
-
-lemma mod_mult_cong_int: "(a::int) ~= 0 \<Longrightarrow> b ~= 0 
-    \<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)"
-  apply (case_tac "b > 0")
-  apply (simp add: cong_int_def mod_mod_cancel mod_add_left_eq)
-  apply (subst (1 2) cong_modulus_neg_int)
-  apply (unfold cong_int_def)
-  apply (subgoal_tac "a * b = (-a * -b)")
-  apply (erule ssubst)
-  apply (subst zmod_zmult2_eq)
-  apply (auto simp add: mod_add_left_eq) 
-done
-
-lemma cong_to_1_nat: "([(a::nat) = 1] (mod n)) \<Longrightarrow> (n dvd (a - 1))"
-  apply (case_tac "a = 0")
-  apply force
-  apply (subst (asm) cong_altdef_nat)
-  apply auto
-done
-
-lemma cong_0_1_nat: "[(0::nat) = 1] (mod n) = (n = 1)"
-  by (unfold cong_nat_def, auto)
-
-lemma cong_0_1_int: "[(0::int) = 1] (mod n) = ((n = 1) | (n = -1))"
-  by (unfold cong_int_def, auto simp add: zmult_eq_1_iff)
-
-lemma cong_to_1'_nat: "[(a::nat) = 1] (mod n) \<longleftrightarrow> 
-    a = 0 \<and> n = 1 \<or> (\<exists>m. a = 1 + m * n)"
-  apply (case_tac "n = 1")
-  apply auto [1]
-  apply (drule_tac x = "a - 1" in spec)
-  apply force
-  apply (case_tac "a = 0")
-  apply (auto simp add: cong_0_1_nat) [1]
-  apply (rule iffI)
-  apply (drule cong_to_1_nat)
-  apply (unfold dvd_def)
-  apply auto [1]
-  apply (rule_tac x = k in exI)
-  apply (auto simp add: ring_simps) [1]
-  apply (subst cong_altdef_nat)
-  apply (auto simp add: dvd_def)
-done
-
-lemma cong_le_nat: "(y::nat) <= x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)"
-  apply (subst cong_altdef_nat)
-  apply assumption
-  apply (unfold dvd_def, auto simp add: ring_simps)
-  apply (rule_tac x = k in exI)
-  apply auto
-done
-
-lemma cong_solve_nat: "(a::nat) \<noteq> 0 \<Longrightarrow> EX x. [a * x = gcd a n] (mod n)"
-  apply (case_tac "n = 0")
-  apply force
-  apply (frule bezout_nat [of a n], auto)
-  apply (rule exI, erule ssubst)
-  apply (rule cong_trans_nat)
-  apply (rule cong_add_nat)
-  apply (subst mult_commute)
-  apply (rule cong_mult_self_nat)
-  prefer 2
-  apply simp
-  apply (rule cong_refl_nat)
-  apply (rule cong_refl_nat)
-done
-
-lemma cong_solve_int: "(a::int) \<noteq> 0 \<Longrightarrow> EX x. [a * x = gcd a n] (mod n)"
-  apply (case_tac "n = 0")
-  apply (case_tac "a \<ge> 0")
-  apply auto
-  apply (rule_tac x = "-1" in exI)
-  apply auto
-  apply (insert bezout_int [of a n], auto)
-  apply (rule exI)
-  apply (erule subst)
-  apply (rule cong_trans_int)
-  prefer 2
-  apply (rule cong_add_int)
-  apply (rule cong_refl_int)
-  apply (rule cong_sym_int)
-  apply (rule cong_mult_self_int)
-  apply simp
-  apply (subst mult_commute)
-  apply (rule cong_refl_int)
-done
-  
-lemma cong_solve_dvd_nat: 
-  assumes a: "(a::nat) \<noteq> 0" and b: "gcd a n dvd d"
-  shows "EX x. [a * x = d] (mod n)"
-proof -
-  from cong_solve_nat [OF a] obtain x where 
-      "[a * x = gcd a n](mod n)"
-    by auto
-  hence "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)" 
-    by (elim cong_scalar2_nat)
-  also from b have "(d div gcd a n) * gcd a n = d"
-    by (rule dvd_div_mult_self)
-  also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)"
-    by auto
-  finally show ?thesis
-    by auto
-qed
-
-lemma cong_solve_dvd_int: 
-  assumes a: "(a::int) \<noteq> 0" and b: "gcd a n dvd d"
-  shows "EX x. [a * x = d] (mod n)"
-proof -
-  from cong_solve_int [OF a] obtain x where 
-      "[a * x = gcd a n](mod n)"
-    by auto
-  hence "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)" 
-    by (elim cong_scalar2_int)
-  also from b have "(d div gcd a n) * gcd a n = d"
-    by (rule dvd_div_mult_self)
-  also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)"
-    by auto
-  finally show ?thesis
-    by auto
-qed
-
-lemma cong_solve_coprime_nat: "coprime (a::nat) n \<Longrightarrow> 
-    EX x. [a * x = 1] (mod n)"
-  apply (case_tac "a = 0")
-  apply force
-  apply (frule cong_solve_nat [of a n])
-  apply auto
-done
-
-lemma cong_solve_coprime_int: "coprime (a::int) n \<Longrightarrow> 
-    EX x. [a * x = 1] (mod n)"
-  apply (case_tac "a = 0")
-  apply auto
-  apply (case_tac "n \<ge> 0")
-  apply auto
-  apply (subst cong_int_def, auto)
-  apply (frule cong_solve_int [of a n])
-  apply auto
-done
-
-lemma coprime_iff_invertible_nat: "m > (1::nat) \<Longrightarrow> coprime a m = 
-    (EX x. [a * x = 1] (mod m))"
-  apply (auto intro: cong_solve_coprime_nat)
-  apply (unfold cong_nat_def, auto intro: invertible_coprime_nat)
-done
-
-lemma coprime_iff_invertible_int: "m > (1::int) \<Longrightarrow> coprime a m = 
-    (EX x. [a * x = 1] (mod m))"
-  apply (auto intro: cong_solve_coprime_int)
-  apply (unfold cong_int_def)
-  apply (auto intro: invertible_coprime_int)
-done
-
-lemma coprime_iff_invertible'_int: "m > (1::int) \<Longrightarrow> coprime a m = 
-    (EX x. 0 <= x & x < m & [a * x = 1] (mod m))"
-  apply (subst coprime_iff_invertible_int)
-  apply auto
-  apply (auto simp add: cong_int_def)
-  apply (rule_tac x = "x mod m" in exI)
-  apply (auto simp add: mod_mult_right_eq [symmetric])
-done
-
-
-lemma cong_cong_lcm_nat: "[(x::nat) = y] (mod a) \<Longrightarrow>
-    [x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)"
-  apply (case_tac "y \<le> x")
-  apply (auto simp add: cong_altdef_nat lcm_least_nat) [1]
-  apply (rule cong_sym_nat)
-  apply (subst (asm) (1 2) cong_sym_eq_nat)
-  apply (auto simp add: cong_altdef_nat lcm_least_nat)
-done
-
-lemma cong_cong_lcm_int: "[(x::int) = y] (mod a) \<Longrightarrow>
-    [x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)"
-  by (auto simp add: cong_altdef_int lcm_least_int) [1]
-
-lemma cong_cong_coprime_nat: "coprime a b \<Longrightarrow> [(x::nat) = y] (mod a) \<Longrightarrow>
-    [x = y] (mod b) \<Longrightarrow> [x = y] (mod a * b)"
-  apply (frule (1) cong_cong_lcm_nat)back
-  apply (simp add: lcm_nat_def)
-done
-
-lemma cong_cong_coprime_int: "coprime a b \<Longrightarrow> [(x::int) = y] (mod a) \<Longrightarrow>
-    [x = y] (mod b) \<Longrightarrow> [x = y] (mod a * b)"
-  apply (frule (1) cong_cong_lcm_int)back
-  apply (simp add: lcm_altdef_int cong_abs_int abs_mult [symmetric])
-done
-
-lemma cong_cong_setprod_coprime_nat [rule_format]: "finite A \<Longrightarrow>
-    (ALL i:A. (ALL j:A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>
-    (ALL i:A. [(x::nat) = y] (mod m i)) \<longrightarrow>
-      [x = y] (mod (PROD i:A. m i))"
-  apply (induct set: finite)
-  apply auto
-  apply (rule cong_cong_coprime_nat)
-  apply (subst gcd_commute_nat)
-  apply (rule setprod_coprime_nat)
-  apply auto
-done
-
-lemma cong_cong_setprod_coprime_int [rule_format]: "finite A \<Longrightarrow>
-    (ALL i:A. (ALL j:A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>
-    (ALL i:A. [(x::int) = y] (mod m i)) \<longrightarrow>
-      [x = y] (mod (PROD i:A. m i))"
-  apply (induct set: finite)
-  apply auto
-  apply (rule cong_cong_coprime_int)
-  apply (subst gcd_commute_int)
-  apply (rule setprod_coprime_int)
-  apply auto
-done
-
-lemma binary_chinese_remainder_aux_nat: 
-  assumes a: "coprime (m1::nat) m2"
-  shows "EX b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and>
-    [b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)"
-proof -
-  from cong_solve_coprime_nat [OF a]
-      obtain x1 where one: "[m1 * x1 = 1] (mod m2)"
-    by auto
-  from a have b: "coprime m2 m1" 
-    by (subst gcd_commute_nat)
-  from cong_solve_coprime_nat [OF b]
-      obtain x2 where two: "[m2 * x2 = 1] (mod m1)"
-    by auto
-  have "[m1 * x1 = 0] (mod m1)"
-    by (subst mult_commute, rule cong_mult_self_nat)
-  moreover have "[m2 * x2 = 0] (mod m2)"
-    by (subst mult_commute, rule cong_mult_self_nat)
-  moreover note one two
-  ultimately show ?thesis by blast
-qed
-
-lemma binary_chinese_remainder_aux_int: 
-  assumes a: "coprime (m1::int) m2"
-  shows "EX b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and>
-    [b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)"
-proof -
-  from cong_solve_coprime_int [OF a]
-      obtain x1 where one: "[m1 * x1 = 1] (mod m2)"
-    by auto
-  from a have b: "coprime m2 m1" 
-    by (subst gcd_commute_int)
-  from cong_solve_coprime_int [OF b]
-      obtain x2 where two: "[m2 * x2 = 1] (mod m1)"
-    by auto
-  have "[m1 * x1 = 0] (mod m1)"
-    by (subst mult_commute, rule cong_mult_self_int)
-  moreover have "[m2 * x2 = 0] (mod m2)"
-    by (subst mult_commute, rule cong_mult_self_int)
-  moreover note one two
-  ultimately show ?thesis by blast
-qed
-
-lemma binary_chinese_remainder_nat:
-  assumes a: "coprime (m1::nat) m2"
-  shows "EX x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)"
-proof -
-  from binary_chinese_remainder_aux_nat [OF a] obtain b1 b2
-    where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" and
-          "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)"
-    by blast
-  let ?x = "u1 * b1 + u2 * b2"
-  have "[?x = u1 * 1 + u2 * 0] (mod m1)"
-    apply (rule cong_add_nat)
-    apply (rule cong_scalar2_nat)
-    apply (rule `[b1 = 1] (mod m1)`)
-    apply (rule cong_scalar2_nat)
-    apply (rule `[b2 = 0] (mod m1)`)
-    done
-  hence "[?x = u1] (mod m1)" by simp
-  have "[?x = u1 * 0 + u2 * 1] (mod m2)"
-    apply (rule cong_add_nat)
-    apply (rule cong_scalar2_nat)
-    apply (rule `[b1 = 0] (mod m2)`)
-    apply (rule cong_scalar2_nat)
-    apply (rule `[b2 = 1] (mod m2)`)
-    done
-  hence "[?x = u2] (mod m2)" by simp
-  with `[?x = u1] (mod m1)` show ?thesis by blast
-qed
-
-lemma binary_chinese_remainder_int:
-  assumes a: "coprime (m1::int) m2"
-  shows "EX x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)"
-proof -
-  from binary_chinese_remainder_aux_int [OF a] obtain b1 b2
-    where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" and
-          "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)"
-    by blast
-  let ?x = "u1 * b1 + u2 * b2"
-  have "[?x = u1 * 1 + u2 * 0] (mod m1)"
-    apply (rule cong_add_int)
-    apply (rule cong_scalar2_int)
-    apply (rule `[b1 = 1] (mod m1)`)
-    apply (rule cong_scalar2_int)
-    apply (rule `[b2 = 0] (mod m1)`)
-    done
-  hence "[?x = u1] (mod m1)" by simp
-  have "[?x = u1 * 0 + u2 * 1] (mod m2)"
-    apply (rule cong_add_int)
-    apply (rule cong_scalar2_int)
-    apply (rule `[b1 = 0] (mod m2)`)
-    apply (rule cong_scalar2_int)
-    apply (rule `[b2 = 1] (mod m2)`)
-    done
-  hence "[?x = u2] (mod m2)" by simp
-  with `[?x = u1] (mod m1)` show ?thesis by blast
-qed
-
-lemma cong_modulus_mult_nat: "[(x::nat) = y] (mod m * n) \<Longrightarrow> 
-    [x = y] (mod m)"
-  apply (case_tac "y \<le> x")
-  apply (simp add: cong_altdef_nat)
-  apply (erule dvd_mult_left)
-  apply (rule cong_sym_nat)
-  apply (subst (asm) cong_sym_eq_nat)
-  apply (simp add: cong_altdef_nat) 
-  apply (erule dvd_mult_left)
-done
-
-lemma cong_modulus_mult_int: "[(x::int) = y] (mod m * n) \<Longrightarrow> 
-    [x = y] (mod m)"
-  apply (simp add: cong_altdef_int) 
-  apply (erule dvd_mult_left)
-done
-
-lemma cong_less_modulus_unique_nat: 
-    "[(x::nat) = y] (mod m) \<Longrightarrow> x < m \<Longrightarrow> y < m \<Longrightarrow> x = y"
-  by (simp add: cong_nat_def)
-
-lemma binary_chinese_remainder_unique_nat:
-  assumes a: "coprime (m1::nat) m2" and
-         nz: "m1 \<noteq> 0" "m2 \<noteq> 0"
-  shows "EX! x. x < m1 * m2 \<and> [x = u1] (mod m1) \<and> [x = u2] (mod m2)"
-proof -
-  from binary_chinese_remainder_nat [OF a] obtain y where 
-      "[y = u1] (mod m1)" and "[y = u2] (mod m2)"
-    by blast
-  let ?x = "y mod (m1 * m2)"
-  from nz have less: "?x < m1 * m2"
-    by auto   
-  have one: "[?x = u1] (mod m1)"
-    apply (rule cong_trans_nat)
-    prefer 2
-    apply (rule `[y = u1] (mod m1)`)
-    apply (rule cong_modulus_mult_nat)
-    apply (rule cong_mod_nat)
-    using nz apply auto
-    done
-  have two: "[?x = u2] (mod m2)"
-    apply (rule cong_trans_nat)
-    prefer 2
-    apply (rule `[y = u2] (mod m2)`)
-    apply (subst mult_commute)
-    apply (rule cong_modulus_mult_nat)
-    apply (rule cong_mod_nat)
-    using nz apply auto
-    done
-  have "ALL z. z < m1 * m2 \<and> [z = u1] (mod m1) \<and> [z = u2] (mod m2) \<longrightarrow>
-      z = ?x"
-  proof (clarify)
-    fix z
-    assume "z < m1 * m2"
-    assume "[z = u1] (mod m1)" and  "[z = u2] (mod m2)"
-    have "[?x = z] (mod m1)"
-      apply (rule cong_trans_nat)
-      apply (rule `[?x = u1] (mod m1)`)
-      apply (rule cong_sym_nat)
-      apply (rule `[z = u1] (mod m1)`)
-      done
-    moreover have "[?x = z] (mod m2)"
-      apply (rule cong_trans_nat)
-      apply (rule `[?x = u2] (mod m2)`)
-      apply (rule cong_sym_nat)
-      apply (rule `[z = u2] (mod m2)`)
-      done
-    ultimately have "[?x = z] (mod m1 * m2)"
-      by (auto intro: coprime_cong_mult_nat a)
-    with `z < m1 * m2` `?x < m1 * m2` show "z = ?x"
-      apply (intro cong_less_modulus_unique_nat)
-      apply (auto, erule cong_sym_nat)
-      done
-  qed  
-  with less one two show ?thesis
-    by auto
- qed
-
-lemma chinese_remainder_aux_nat:
-  fixes A :: "'a set" and
-        m :: "'a \<Rightarrow> nat"
-  assumes fin: "finite A" and
-          cop: "ALL i : A. (ALL j : A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"
-  shows "EX b. (ALL i : A. 
-      [b i = 1] (mod m i) \<and> [b i = 0] (mod (PROD j : A - {i}. m j)))"
-proof (rule finite_set_choice, rule fin, rule ballI)
-  fix i
-  assume "i : A"
-  with cop have "coprime (PROD j : A - {i}. m j) (m i)"
-    by (intro setprod_coprime_nat, auto)
-  hence "EX x. [(PROD j : A - {i}. m j) * x = 1] (mod m i)"
-    by (elim cong_solve_coprime_nat)
-  then obtain x where "[(PROD j : A - {i}. m j) * x = 1] (mod m i)"
-    by auto
-  moreover have "[(PROD j : A - {i}. m j) * x = 0] 
-    (mod (PROD j : A - {i}. m j))"
-    by (subst mult_commute, rule cong_mult_self_nat)
-  ultimately show "\<exists>a. [a = 1] (mod m i) \<and> [a = 0] 
-      (mod setprod m (A - {i}))"
-    by blast
-qed
-
-lemma chinese_remainder_nat:
-  fixes A :: "'a set" and
-        m :: "'a \<Rightarrow> nat" and
-        u :: "'a \<Rightarrow> nat"
-  assumes 
-        fin: "finite A" and
-        cop: "ALL i:A. (ALL j : A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"
-  shows "EX x. (ALL i:A. [x = u i] (mod m i))"
-proof -
-  from chinese_remainder_aux_nat [OF fin cop] obtain b where
-    bprop: "ALL i:A. [b i = 1] (mod m i) \<and> 
-      [b i = 0] (mod (PROD j : A - {i}. m j))"
-    by blast
-  let ?x = "SUM i:A. (u i) * (b i)"
-  show "?thesis"
-  proof (rule exI, clarify)
-    fix i
-    assume a: "i : A"
-    show "[?x = u i] (mod m i)" 
-    proof -
-      from fin a have "?x = (SUM j:{i}. u j * b j) + 
-          (SUM j:A-{i}. u j * b j)"
-        by (subst setsum_Un_disjoint [symmetric], auto intro: setsum_cong)
-      hence "[?x = u i * b i + (SUM j:A-{i}. u j * b j)] (mod m i)"
-        by auto
-      also have "[u i * b i + (SUM j:A-{i}. u j * b j) =
-                  u i * 1 + (SUM j:A-{i}. u j * 0)] (mod m i)"
-        apply (rule cong_add_nat)
-        apply (rule cong_scalar2_nat)
-        using bprop a apply blast
-        apply (rule cong_setsum_nat)
-        apply (rule cong_scalar2_nat)
-        using bprop apply auto
-        apply (rule cong_dvd_modulus_nat)
-        apply (drule (1) bspec)
-        apply (erule conjE)
-        apply assumption
-        apply (rule dvd_setprod)
-        using fin a apply auto
-        done
-      finally show ?thesis
-        by simp
-    qed
-  qed
-qed
-
-lemma coprime_cong_prod_nat [rule_format]: "finite A \<Longrightarrow> 
-    (ALL i: A. (ALL j: A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>
-      (ALL i: A. [(x::nat) = y] (mod m i)) \<longrightarrow>
-         [x = y] (mod (PROD i:A. m i))" 
-  apply (induct set: finite)
-  apply auto
-  apply (erule (1) coprime_cong_mult_nat)
-  apply (subst gcd_commute_nat)
-  apply (rule setprod_coprime_nat)
-  apply auto
-done
-
-lemma chinese_remainder_unique_nat:
-  fixes A :: "'a set" and
-        m :: "'a \<Rightarrow> nat" and
-        u :: "'a \<Rightarrow> nat"
-  assumes 
-        fin: "finite A" and
-         nz: "ALL i:A. m i \<noteq> 0" and
-        cop: "ALL i:A. (ALL j : A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"
-  shows "EX! x. x < (PROD i:A. m i) \<and> (ALL i:A. [x = u i] (mod m i))"
-proof -
-  from chinese_remainder_nat [OF fin cop] obtain y where
-      one: "(ALL i:A. [y = u i] (mod m i))" 
-    by blast
-  let ?x = "y mod (PROD i:A. m i)"
-  from fin nz have prodnz: "(PROD i:A. m i) \<noteq> 0"
-    by auto
-  hence less: "?x < (PROD i:A. m i)"
-    by auto
-  have cong: "ALL i:A. [?x = u i] (mod m i)"
-    apply auto
-    apply (rule cong_trans_nat)
-    prefer 2
-    using one apply auto
-    apply (rule cong_dvd_modulus_nat)
-    apply (rule cong_mod_nat)
-    using prodnz apply auto
-    apply (rule dvd_setprod)
-    apply (rule fin)
-    apply assumption
-    done
-  have unique: "ALL z. z < (PROD i:A. m i) \<and> 
-      (ALL i:A. [z = u i] (mod m i)) \<longrightarrow> z = ?x"
-  proof (clarify)
-    fix z
-    assume zless: "z < (PROD i:A. m i)"
-    assume zcong: "(ALL i:A. [z = u i] (mod m i))"
-    have "ALL i:A. [?x = z] (mod m i)"
-      apply clarify     
-      apply (rule cong_trans_nat)
-      using cong apply (erule bspec)
-      apply (rule cong_sym_nat)
-      using zcong apply auto
-      done
-    with fin cop have "[?x = z] (mod (PROD i:A. m i))"
-      by (intro coprime_cong_prod_nat, auto)
-    with zless less show "z = ?x"
-      apply (intro cong_less_modulus_unique_nat)
-      apply (auto, erule cong_sym_nat)
-      done
-  qed 
-  from less cong unique show ?thesis
-    by blast  
-qed
-
-end
--- a/src/HOL/NewNumberTheory/Fib.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,319 +0,0 @@
-(*  Title:      Fib.thy
-    Authors:    Lawrence C. Paulson, Jeremy Avigad
-
-
-Defines the fibonacci function.
-
-The original "Fib" is due to Lawrence C. Paulson, and was adapted by
-Jeremy Avigad.
-*)
-
-
-header {* Fib *}
-
-theory Fib
-imports Binomial
-begin
-
-
-subsection {* Main definitions *}
-
-class fib =
-
-fixes 
-  fib :: "'a \<Rightarrow> 'a"
-
-
-(* definition for the natural numbers *)
-
-instantiation nat :: fib
-
-begin 
-
-fun 
-  fib_nat :: "nat \<Rightarrow> nat"
-where
-  "fib_nat n =
-   (if n = 0 then 0 else
-   (if n = 1 then 1 else
-     fib (n - 1) + fib (n - 2)))"
-
-instance proof qed
-
-end
-
-(* definition for the integers *)
-
-instantiation int :: fib
-
-begin 
-
-definition
-  fib_int :: "int \<Rightarrow> int"
-where  
-  "fib_int n = (if n >= 0 then int (fib (nat n)) else 0)"
-
-instance proof qed
-
-end
-
-
-subsection {* Set up Transfer *}
-
-
-lemma transfer_nat_int_fib:
-  "(x::int) >= 0 \<Longrightarrow> fib (nat x) = nat (fib x)"
-  unfolding fib_int_def by auto
-
-lemma transfer_nat_int_fib_closure:
-  "n >= (0::int) \<Longrightarrow> fib n >= 0"
-  by (auto simp add: fib_int_def)
-
-declare TransferMorphism_nat_int[transfer add return: 
-    transfer_nat_int_fib transfer_nat_int_fib_closure]
-
-lemma transfer_int_nat_fib:
-  "fib (int n) = int (fib n)"
-  unfolding fib_int_def by auto
-
-lemma transfer_int_nat_fib_closure:
-  "is_nat n \<Longrightarrow> fib n >= 0"
-  unfolding fib_int_def by auto
-
-declare TransferMorphism_int_nat[transfer add return: 
-    transfer_int_nat_fib transfer_int_nat_fib_closure]
-
-
-subsection {* Fibonacci numbers *}
-
-lemma fib_0_nat [simp]: "fib (0::nat) = 0"
-  by simp
-
-lemma fib_0_int [simp]: "fib (0::int) = 0"
-  unfolding fib_int_def by simp
-
-lemma fib_1_nat [simp]: "fib (1::nat) = 1"
-  by simp
-
-lemma fib_Suc_0_nat [simp]: "fib (Suc 0) = Suc 0"
-  by simp
-
-lemma fib_1_int [simp]: "fib (1::int) = 1"
-  unfolding fib_int_def by simp
-
-lemma fib_reduce_nat: "(n::nat) >= 2 \<Longrightarrow> fib n = fib (n - 1) + fib (n - 2)"
-  by simp
-
-declare fib_nat.simps [simp del]
-
-lemma fib_reduce_int: "(n::int) >= 2 \<Longrightarrow> fib n = fib (n - 1) + fib (n - 2)"
-  unfolding fib_int_def
-  by (auto simp add: fib_reduce_nat nat_diff_distrib)
-
-lemma fib_neg_int [simp]: "(n::int) < 0 \<Longrightarrow> fib n = 0"
-  unfolding fib_int_def by auto
-
-lemma fib_2_nat [simp]: "fib (2::nat) = 1"
-  by (subst fib_reduce_nat, auto)
-
-lemma fib_2_int [simp]: "fib (2::int) = 1"
-  by (subst fib_reduce_int, auto)
-
-lemma fib_plus_2_nat: "fib ((n::nat) + 2) = fib (n + 1) + fib n"
-  by (subst fib_reduce_nat, auto simp add: One_nat_def)
-(* the need for One_nat_def is due to the natdiff_cancel_numerals
-   procedure *)
-
-lemma fib_induct_nat: "P (0::nat) \<Longrightarrow> P (1::nat) \<Longrightarrow> 
-    (!!n. P n \<Longrightarrow> P (n + 1) \<Longrightarrow> P (n + 2)) \<Longrightarrow> P n"
-  apply (atomize, induct n rule: nat_less_induct)
-  apply auto
-  apply (case_tac "n = 0", force)
-  apply (case_tac "n = 1", force)
-  apply (subgoal_tac "n >= 2")
-  apply (frule_tac x = "n - 1" in spec)
-  apply (drule_tac x = "n - 2" in spec)
-  apply (drule_tac x = "n - 2" in spec)
-  apply auto
-  apply (auto simp add: One_nat_def) (* again, natdiff_cancel *)
-done
-
-lemma fib_add_nat: "fib ((n::nat) + k + 1) = fib (k + 1) * fib (n + 1) + 
-    fib k * fib n"
-  apply (induct n rule: fib_induct_nat)
-  apply auto
-  apply (subst fib_reduce_nat)
-  apply (auto simp add: ring_simps)
-  apply (subst (1 3 5) fib_reduce_nat)
-  apply (auto simp add: ring_simps Suc_eq_plus1)
-(* hmmm. Why doesn't "n + (1 + (1 + k))" simplify to "n + k + 2"? *)
-  apply (subgoal_tac "n + (k + 2) = n + (1 + (1 + k))")
-  apply (erule ssubst) back back
-  apply (erule ssubst) back 
-  apply auto
-done
-
-lemma fib_add'_nat: "fib (n + Suc k) = fib (Suc k) * fib (Suc n) + 
-    fib k * fib n"
-  using fib_add_nat by (auto simp add: One_nat_def)
-
-
-(* transfer from nats to ints *)
-lemma fib_add_int [rule_format]: "(n::int) >= 0 \<Longrightarrow> k >= 0 \<Longrightarrow>
-    fib (n + k + 1) = fib (k + 1) * fib (n + 1) + 
-    fib k * fib n "
-
-  by (rule fib_add_nat [transferred])
-
-lemma fib_neq_0_nat: "(n::nat) > 0 \<Longrightarrow> fib n ~= 0"
-  apply (induct n rule: fib_induct_nat)
-  apply (auto simp add: fib_plus_2_nat)
-done
-
-lemma fib_gr_0_nat: "(n::nat) > 0 \<Longrightarrow> fib n > 0"
-  by (frule fib_neq_0_nat, simp)
-
-lemma fib_gr_0_int: "(n::int) > 0 \<Longrightarrow> fib n > 0"
-  unfolding fib_int_def by (simp add: fib_gr_0_nat)
-
-text {*
-  \medskip Concrete Mathematics, page 278: Cassini's identity.  The proof is
-  much easier using integers, not natural numbers!
-*}
-
-lemma fib_Cassini_aux_int: "fib (int n + 2) * fib (int n) - 
-    (fib (int n + 1))^2 = (-1)^(n + 1)"
-  apply (induct n)
-  apply (auto simp add: ring_simps power2_eq_square fib_reduce_int
-      power_add)
-done
-
-lemma fib_Cassini_int: "n >= 0 \<Longrightarrow> fib (n + 2) * fib n - 
-    (fib (n + 1))^2 = (-1)^(nat n + 1)"
-  by (insert fib_Cassini_aux_int [of "nat n"], auto)
-
-(*
-lemma fib_Cassini'_int: "n >= 0 \<Longrightarrow> fib (n + 2) * fib n = 
-    (fib (n + 1))^2 + (-1)^(nat n + 1)"
-  by (frule fib_Cassini_int, simp) 
-*)
-
-lemma fib_Cassini'_int: "n >= 0 \<Longrightarrow> fib ((n::int) + 2) * fib n =
-  (if even n then tsub ((fib (n + 1))^2) 1
-   else (fib (n + 1))^2 + 1)"
-  apply (frule fib_Cassini_int, auto simp add: pos_int_even_equiv_nat_even)
-  apply (subst tsub_eq)
-  apply (insert fib_gr_0_int [of "n + 1"], force)
-  apply auto
-done
-
-lemma fib_Cassini_nat: "fib ((n::nat) + 2) * fib n =
-  (if even n then (fib (n + 1))^2 - 1
-   else (fib (n + 1))^2 + 1)"
-
-  by (rule fib_Cassini'_int [transferred, of n], auto)
-
-
-text {* \medskip Toward Law 6.111 of Concrete Mathematics *}
-
-lemma coprime_fib_plus_1_nat: "coprime (fib (n::nat)) (fib (n + 1))"
-  apply (induct n rule: fib_induct_nat)
-  apply auto
-  apply (subst (2) fib_reduce_nat)
-  apply (auto simp add: Suc_eq_plus1) (* again, natdiff_cancel *)
-  apply (subst add_commute, auto)
-  apply (subst gcd_commute_nat, auto simp add: ring_simps)
-done
-
-lemma coprime_fib_Suc_nat: "coprime (fib n) (fib (Suc n))"
-  using coprime_fib_plus_1_nat by (simp add: One_nat_def)
-
-lemma coprime_fib_plus_1_int: 
-    "n >= 0 \<Longrightarrow> coprime (fib (n::int)) (fib (n + 1))"
-  by (erule coprime_fib_plus_1_nat [transferred])
-
-lemma gcd_fib_add_nat: "gcd (fib (m::nat)) (fib (n + m)) = gcd (fib m) (fib n)"
-  apply (simp add: gcd_commute_nat [of "fib m"])
-  apply (rule cases_nat [of _ m])
-  apply simp
-  apply (subst add_assoc [symmetric])
-  apply (simp add: fib_add_nat)
-  apply (subst gcd_commute_nat)
-  apply (subst mult_commute)
-  apply (subst gcd_add_mult_nat)
-  apply (subst gcd_commute_nat)
-  apply (rule gcd_mult_cancel_nat)
-  apply (rule coprime_fib_plus_1_nat)
-done
-
-lemma gcd_fib_add_int [rule_format]: "m >= 0 \<Longrightarrow> n >= 0 \<Longrightarrow> 
-    gcd (fib (m::int)) (fib (n + m)) = gcd (fib m) (fib n)"
-  by (erule gcd_fib_add_nat [transferred])
-
-lemma gcd_fib_diff_nat: "(m::nat) \<le> n \<Longrightarrow> 
-    gcd (fib m) (fib (n - m)) = gcd (fib m) (fib n)"
-  by (simp add: gcd_fib_add_nat [symmetric, of _ "n-m"])
-
-lemma gcd_fib_diff_int: "0 <= (m::int) \<Longrightarrow> m \<le> n \<Longrightarrow> 
-    gcd (fib m) (fib (n - m)) = gcd (fib m) (fib n)"
-  by (simp add: gcd_fib_add_int [symmetric, of _ "n-m"])
-
-lemma gcd_fib_mod_nat: "0 < (m::nat) \<Longrightarrow> 
-    gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
-proof (induct n rule: less_induct)
-  case (less n)
-  from less.prems have pos_m: "0 < m" .
-  show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
-  proof (cases "m < n")
-    case True note m_n = True
-    then have m_n': "m \<le> n" by auto
-    with pos_m have pos_n: "0 < n" by auto
-    with pos_m m_n have diff: "n - m < n" by auto
-    have "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib ((n - m) mod m))"
-    by (simp add: mod_if [of n]) (insert m_n, auto)
-    also have "\<dots> = gcd (fib m)  (fib (n - m))" 
-      by (simp add: less.hyps diff pos_m)
-    also have "\<dots> = gcd (fib m) (fib n)" by (simp add: gcd_fib_diff_nat m_n')
-    finally show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)" .
-  next
-    case False then show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
-    by (cases "m = n") auto
-  qed
-qed
-
-lemma gcd_fib_mod_int: 
-  assumes "0 < (m::int)" and "0 <= n"
-  shows "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
-
-  apply (rule gcd_fib_mod_nat [transferred])
-  using prems apply auto
-done
-
-lemma fib_gcd_nat: "fib (gcd (m::nat) n) = gcd (fib m) (fib n)"  
-    -- {* Law 6.111 *}
-  apply (induct m n rule: gcd_nat_induct)
-  apply (simp_all add: gcd_non_0_nat gcd_commute_nat gcd_fib_mod_nat)
-done
-
-lemma fib_gcd_int: "m >= 0 \<Longrightarrow> n >= 0 \<Longrightarrow>
-    fib (gcd (m::int) n) = gcd (fib m) (fib n)"
-  by (erule fib_gcd_nat [transferred])
-
-lemma atMost_plus_one_nat: "{..(k::nat) + 1} = insert (k + 1) {..k}" 
-  by auto
-
-theorem fib_mult_eq_setsum_nat:
-    "fib ((n::nat) + 1) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
-  apply (induct n)
-  apply (auto simp add: atMost_plus_one_nat fib_plus_2_nat ring_simps)
-done
-
-theorem fib_mult_eq_setsum'_nat:
-    "fib (Suc n) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
-  using fib_mult_eq_setsum_nat by (simp add: One_nat_def)
-
-theorem fib_mult_eq_setsum_int [rule_format]:
-    "n >= 0 \<Longrightarrow> fib ((n::int) + 1) * fib n = (\<Sum>k \<in> {0..n}. fib k * fib k)"
-  by (erule fib_mult_eq_setsum_nat [transferred])
-
-end
--- a/src/HOL/NewNumberTheory/MiscAlgebra.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,355 +0,0 @@
-(*  Title:      MiscAlgebra.thy
-    Author:     Jeremy Avigad
-
-These are things that can be added to the Algebra library.
-*)
-
-theory MiscAlgebra
-imports
-  "~~/src/HOL/Algebra/Ring"
-  "~~/src/HOL/Algebra/FiniteProduct"
-begin;
-
-(* finiteness stuff *)
-
-lemma bounded_set1_int [intro]: "finite {(x::int). a < x & x < b & P x}" 
-  apply (subgoal_tac "{x. a < x & x < b & P x} <= {a<..<b}")
-  apply (erule finite_subset)
-  apply auto
-done
-
-
-(* The rest is for the algebra libraries *)
-
-(* These go in Group.thy. *)
-
-(*
-  Show that the units in any monoid give rise to a group.
-
-  The file Residues.thy provides some infrastructure to use
-  facts about the unit group within the ring locale.
-*)
-
-
-constdefs 
-  units_of :: "('a, 'b) monoid_scheme => 'a monoid"
-  "units_of G == (| carrier = Units G,
-     Group.monoid.mult = Group.monoid.mult G,
-     one  = one G |)";
-
-(*
-
-lemma (in monoid) Units_mult_closed [intro]:
-  "x : Units G ==> y : Units G ==> x \<otimes> y : Units G"
-  apply (unfold Units_def)
-  apply (clarsimp)
-  apply (rule_tac x = "xaa \<otimes> xa" in bexI)
-  apply auto
-  apply (subst m_assoc)
-  apply auto
-  apply (subst (2) m_assoc [symmetric])
-  apply auto
-  apply (subst m_assoc)
-  apply auto
-  apply (subst (2) m_assoc [symmetric])
-  apply auto
-done
-
-*)
-
-lemma (in monoid) units_group: "group(units_of G)"
-  apply (unfold units_of_def)
-  apply (rule groupI)
-  apply auto
-  apply (subst m_assoc)
-  apply auto
-  apply (rule_tac x = "inv x" in bexI)
-  apply auto
-done
-
-lemma (in comm_monoid) units_comm_group: "comm_group(units_of G)"
-  apply (rule group.group_comm_groupI)
-  apply (rule units_group)
-  apply (insert prems)
-  apply (unfold units_of_def Units_def comm_monoid_def comm_monoid_axioms_def)
-  apply auto;
-done;
-
-lemma units_of_carrier: "carrier (units_of G) = Units G"
-  by (unfold units_of_def, auto)
-
-lemma units_of_mult: "mult(units_of G) = mult G"
-  by (unfold units_of_def, auto)
-
-lemma units_of_one: "one(units_of G) = one G"
-  by (unfold units_of_def, auto)
-
-lemma (in monoid) units_of_inv: "x : Units G ==> 
-    m_inv (units_of G) x = m_inv G x"
-  apply (rule sym)
-  apply (subst m_inv_def)
-  apply (rule the1_equality)
-  apply (rule ex_ex1I)
-  apply (subst (asm) Units_def)
-  apply auto
-  apply (erule inv_unique)
-  apply auto
-  apply (rule Units_closed)
-  apply (simp_all only: units_of_carrier [symmetric])
-  apply (insert units_group)
-  apply auto
-  apply (subst units_of_mult [symmetric])
-  apply (subst units_of_one [symmetric])
-  apply (erule group.r_inv, assumption)
-  apply (subst units_of_mult [symmetric])
-  apply (subst units_of_one [symmetric])
-  apply (erule group.l_inv, assumption)
-done
-
-lemma (in group) inj_on_const_mult: "a: (carrier G) ==> 
-    inj_on (%x. a \<otimes> x) (carrier G)"
-  by (unfold inj_on_def, auto)
-
-lemma (in group) surj_const_mult: "a : (carrier G) ==>
-    (%x. a \<otimes> x) ` (carrier G) = (carrier G)" 
-  apply (auto simp add: image_def)
-  apply (rule_tac x = "(m_inv G a) \<otimes> x" in bexI)
-  apply auto
-(* auto should get this. I suppose we need "comm_monoid_simprules"
-   for mult_ac rewriting. *)
-  apply (subst m_assoc [symmetric])
-  apply auto
-done
-
-lemma (in group) l_cancel_one [simp]: "x : carrier G \<Longrightarrow> a : carrier G \<Longrightarrow>
-    (x \<otimes> a = x) = (a = one G)"
-  apply auto
-  apply (subst l_cancel [symmetric])
-  prefer 4
-  apply (erule ssubst)
-  apply auto
-done
-
-lemma (in group) r_cancel_one [simp]: "x : carrier G \<Longrightarrow> a : carrier G \<Longrightarrow>
-    (a \<otimes> x = x) = (a = one G)"
-  apply auto
-  apply (subst r_cancel [symmetric])
-  prefer 4
-  apply (erule ssubst)
-  apply auto
-done
-
-(* Is there a better way to do this? *)
-
-lemma (in group) l_cancel_one' [simp]: "x : carrier G \<Longrightarrow> a : carrier G \<Longrightarrow>
-    (x = x \<otimes> a) = (a = one G)"
-  by (subst eq_commute, simp)
-
-lemma (in group) r_cancel_one' [simp]: "x : carrier G \<Longrightarrow> a : carrier G \<Longrightarrow>
-    (x = a \<otimes> x) = (a = one G)"
-  by (subst eq_commute, simp)
-
-(* This should be generalized to arbitrary groups, not just commutative
-   ones, using Lagrange's theorem. *)
-
-lemma (in comm_group) power_order_eq_one:
-    assumes fin [simp]: "finite (carrier G)"
-        and a [simp]: "a : carrier G" 
-      shows "a (^) card(carrier G) = one G" 
-proof -
-  have "(\<Otimes>x:carrier G. x) = (\<Otimes>x:carrier G. a \<otimes> x)"
-    by (subst (2) finprod_reindex [symmetric], 
-      auto simp add: Pi_def inj_on_const_mult surj_const_mult)
-  also have "\<dots> = (\<Otimes>x:carrier G. a) \<otimes> (\<Otimes>x:carrier G. x)"
-    by (auto simp add: finprod_multf Pi_def)
-  also have "(\<Otimes>x:carrier G. a) = a (^) card(carrier G)"
-    by (auto simp add: finprod_const)
-  finally show ?thesis
-(* uses the preceeding lemma *)
-    by auto
-qed
-
-
-(* Miscellaneous *)
-
-lemma (in cring) field_intro2: "\<zero>\<^bsub>R\<^esub> ~= \<one>\<^bsub>R\<^esub> \<Longrightarrow> ALL x : carrier R - {\<zero>\<^bsub>R\<^esub>}.
-    x : Units R \<Longrightarrow> field R"
-  apply (unfold_locales)
-  apply (insert prems, auto)
-  apply (rule trans)
-  apply (subgoal_tac "a = (a \<otimes> b) \<otimes> inv b")
-  apply assumption
-  apply (subst m_assoc) 
-  apply (auto simp add: Units_r_inv)
-  apply (unfold Units_def)
-  apply auto
-done
-
-lemma (in monoid) inv_char: "x : carrier G \<Longrightarrow> y : carrier G \<Longrightarrow>
-  x \<otimes> y = \<one> \<Longrightarrow> y \<otimes> x = \<one> \<Longrightarrow> inv x = y"
-  apply (subgoal_tac "x : Units G")
-  apply (subgoal_tac "y = inv x \<otimes> \<one>")
-  apply simp
-  apply (erule subst)
-  apply (subst m_assoc [symmetric])
-  apply auto
-  apply (unfold Units_def)
-  apply auto
-done
-
-lemma (in comm_monoid) comm_inv_char: "x : carrier G \<Longrightarrow> y : carrier G \<Longrightarrow>
-  x \<otimes> y = \<one> \<Longrightarrow> inv x = y"
-  apply (rule inv_char)
-  apply auto
-  apply (subst m_comm, auto) 
-done
-
-lemma (in ring) inv_neg_one [simp]: "inv (\<ominus> \<one>) = \<ominus> \<one>"  
-  apply (rule inv_char)
-  apply (auto simp add: l_minus r_minus)
-done
-
-lemma (in monoid) inv_eq_imp_eq: "x : Units G \<Longrightarrow> y : Units G \<Longrightarrow> 
-    inv x = inv y \<Longrightarrow> x = y"
-  apply (subgoal_tac "inv(inv x) = inv(inv y)")
-  apply (subst (asm) Units_inv_inv)+
-  apply auto
-done
-
-lemma (in ring) Units_minus_one_closed [intro]: "\<ominus> \<one> : Units R"
-  apply (unfold Units_def)
-  apply auto
-  apply (rule_tac x = "\<ominus> \<one>" in bexI)
-  apply auto
-  apply (simp add: l_minus r_minus)
-done
-
-lemma (in monoid) inv_one [simp]: "inv \<one> = \<one>"
-  apply (rule inv_char)
-  apply auto
-done
-
-lemma (in ring) inv_eq_neg_one_eq: "x : Units R \<Longrightarrow> (inv x = \<ominus> \<one>) = (x = \<ominus> \<one>)"
-  apply auto
-  apply (subst Units_inv_inv [symmetric])
-  apply auto
-done
-
-lemma (in monoid) inv_eq_one_eq: "x : Units G \<Longrightarrow> (inv x = \<one>) = (x = \<one>)"
-  apply auto
-  apply (subst Units_inv_inv [symmetric])
-  apply auto
-done
-
-
-(* This goes in FiniteProduct *)
-
-lemma (in comm_monoid) finprod_UN_disjoint:
-  "finite I \<Longrightarrow> (ALL i:I. finite (A i)) \<longrightarrow> (ALL i:I. ALL j:I. i ~= j \<longrightarrow>
-     (A i) Int (A j) = {}) \<longrightarrow>
-      (ALL i:I. ALL x: (A i). g x : carrier G) \<longrightarrow>
-        finprod G g (UNION I A) = finprod G (%i. finprod G g (A i)) I"
-  apply (induct set: finite)
-  apply force
-  apply clarsimp
-  apply (subst finprod_Un_disjoint)
-  apply blast
-  apply (erule finite_UN_I)
-  apply blast
-  apply (fastsimp)
-  apply (auto intro!: funcsetI finprod_closed) 
-done
-
-lemma (in comm_monoid) finprod_Union_disjoint:
-  "[| finite C; (ALL A:C. finite A & (ALL x:A. f x : carrier G));
-      (ALL A:C. ALL B:C. A ~= B --> A Int B = {}) |] 
-   ==> finprod G f (Union C) = finprod G (finprod G f) C" 
-  apply (frule finprod_UN_disjoint [of C id f])
-  apply (unfold Union_def id_def, auto)
-done
-
-lemma (in comm_monoid) finprod_one [rule_format]: 
-  "finite A \<Longrightarrow> (ALL x:A. f x = \<one>) \<longrightarrow>
-     finprod G f A = \<one>"
-by (induct set: finite) auto
-
-
-(* need better simplification rules for rings *)
-(* the next one holds more generally for abelian groups *)
-
-lemma (in cring) sum_zero_eq_neg:
-  "x : carrier R \<Longrightarrow> y : carrier R \<Longrightarrow> x \<oplus> y = \<zero> \<Longrightarrow> x = \<ominus> y"
-  apply (subgoal_tac "\<ominus> y = \<zero> \<oplus> \<ominus> y") 
-  apply (erule ssubst)back
-  apply (erule subst)
-  apply (simp add: ring_simprules)+
-done
-
-(* there's a name conflict -- maybe "domain" should be
-   "integral_domain" *)
-
-lemma (in Ring.domain) square_eq_one: 
-  fixes x
-  assumes [simp]: "x : carrier R" and
-    "x \<otimes> x = \<one>"
-  shows "x = \<one> | x = \<ominus>\<one>"
-proof -
-  have "(x \<oplus> \<one>) \<otimes> (x \<oplus> \<ominus> \<one>) = x \<otimes> x \<oplus> \<ominus> \<one>"
-    by (simp add: ring_simprules)
-  also with `x \<otimes> x = \<one>` have "\<dots> = \<zero>"
-    by (simp add: ring_simprules)
-  finally have "(x \<oplus> \<one>) \<otimes> (x \<oplus> \<ominus> \<one>) = \<zero>" .
-  hence "(x \<oplus> \<one>) = \<zero> | (x \<oplus> \<ominus> \<one>) = \<zero>"
-    by (intro integral, auto)
-  thus ?thesis
-    apply auto
-    apply (erule notE)
-    apply (rule sum_zero_eq_neg)
-    apply auto
-    apply (subgoal_tac "x = \<ominus> (\<ominus> \<one>)")
-    apply (simp add: ring_simprules) 
-    apply (rule sum_zero_eq_neg)
-    apply auto
-    done
-qed
-
-lemma (in Ring.domain) inv_eq_self: "x : Units R \<Longrightarrow>
-    x = inv x \<Longrightarrow> x = \<one> | x = \<ominus> \<one>"
-  apply (rule square_eq_one)
-  apply auto
-  apply (erule ssubst)back
-  apply (erule Units_r_inv)
-done
-
-
-(*
-  The following translates theorems about groups to the facts about
-  the units of a ring. (The list should be expanded as more things are
-  needed.)
-*)
-
-lemma (in ring) finite_ring_finite_units [intro]: "finite (carrier R) \<Longrightarrow> 
-    finite (Units R)"
-  by (rule finite_subset, auto)
-
-(* this belongs with MiscAlgebra.thy *)
-lemma (in monoid) units_of_pow: 
-    "x : Units G \<Longrightarrow> x (^)\<^bsub>units_of G\<^esub> (n::nat) = x (^)\<^bsub>G\<^esub> n"
-  apply (induct n)
-  apply (auto simp add: units_group group.is_monoid  
-    monoid.nat_pow_0 monoid.nat_pow_Suc units_of_one units_of_mult
-    One_nat_def)
-done
-
-lemma (in cring) units_power_order_eq_one: "finite (Units R) \<Longrightarrow> a : Units R
-    \<Longrightarrow> a (^) card(Units R) = \<one>"
-  apply (subst units_of_carrier [symmetric])
-  apply (subst units_of_one [symmetric])
-  apply (subst units_of_pow [symmetric])
-  apply assumption
-  apply (rule comm_group.power_order_eq_one)
-  apply (rule units_comm_group)
-  apply (unfold units_of_def, auto)
-done
-
-end
--- a/src/HOL/NewNumberTheory/ROOT.ML	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-use_thys ["Fib","Residues"];
--- a/src/HOL/NewNumberTheory/Residues.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,466 +0,0 @@
-(*  Title:      HOL/Library/Residues.thy
-    ID:         
-    Author:     Jeremy Avigad
-
-    An algebraic treatment of residue rings, and resulting proofs of
-    Euler's theorem and Wilson's theorem. 
-*)
-
-header {* Residue rings *}
-
-theory Residues
-imports
-   UniqueFactorization
-   Binomial
-   MiscAlgebra
-begin
-
-
-(*
- 
-  A locale for residue rings
-
-*)
-
-constdefs 
-  residue_ring :: "int => int ring"
-  "residue_ring m == (| 
-    carrier =       {0..m - 1}, 
-    mult =          (%x y. (x * y) mod m),
-    one =           1,
-    zero =          0,
-    add =           (%x y. (x + y) mod m) |)"
-
-locale residues =
-  fixes m :: int and R (structure)
-  assumes m_gt_one: "m > 1"
-  defines "R == residue_ring m"
-
-context residues begin
-
-lemma abelian_group: "abelian_group R"
-  apply (insert m_gt_one)
-  apply (rule abelian_groupI)
-  apply (unfold R_def residue_ring_def)
-  apply (auto simp add: mod_pos_pos_trivial mod_add_right_eq [symmetric]
-    add_ac)
-  apply (case_tac "x = 0")
-  apply force
-  apply (subgoal_tac "(x + (m - x)) mod m = 0")
-  apply (erule bexI)
-  apply auto
-done
-
-lemma comm_monoid: "comm_monoid R"
-  apply (insert m_gt_one)
-  apply (unfold R_def residue_ring_def)
-  apply (rule comm_monoidI)
-  apply auto
-  apply (subgoal_tac "x * y mod m * z mod m = z * (x * y mod m) mod m")
-  apply (erule ssubst)
-  apply (subst zmod_zmult1_eq [symmetric])+
-  apply (simp_all only: mult_ac)
-done
-
-lemma cring: "cring R"
-  apply (rule cringI)
-  apply (rule abelian_group)
-  apply (rule comm_monoid)
-  apply (unfold R_def residue_ring_def, auto)
-  apply (subst mod_add_eq [symmetric])
-  apply (subst mult_commute)
-  apply (subst zmod_zmult1_eq [symmetric])
-  apply (simp add: ring_simps)
-done
-
-end
-
-sublocale residues < cring
-  by (rule cring)
-
-
-context residues begin
-
-(* These lemmas translate back and forth between internal and 
-   external concepts *)
-
-lemma res_carrier_eq: "carrier R = {0..m - 1}"
-  by (unfold R_def residue_ring_def, auto)
-
-lemma res_add_eq: "x \<oplus> y = (x + y) mod m"
-  by (unfold R_def residue_ring_def, auto)
-
-lemma res_mult_eq: "x \<otimes> y = (x * y) mod m"
-  by (unfold R_def residue_ring_def, auto)
-
-lemma res_zero_eq: "\<zero> = 0"
-  by (unfold R_def residue_ring_def, auto)
-
-lemma res_one_eq: "\<one> = 1"
-  by (unfold R_def residue_ring_def units_of_def residue_ring_def, auto)
-
-lemma res_units_eq: "Units R = { x. 0 < x & x < m & coprime x m}"
-  apply (insert m_gt_one)
-  apply (unfold Units_def R_def residue_ring_def)
-  apply auto
-  apply (subgoal_tac "x ~= 0")
-  apply auto
-  apply (rule invertible_coprime_int)
-  apply (subgoal_tac "x ~= 0")
-  apply auto
-  apply (subst (asm) coprime_iff_invertible'_int)
-  apply (rule m_gt_one)
-  apply (auto simp add: cong_int_def mult_commute)
-done
-
-lemma res_neg_eq: "\<ominus> x = (- x) mod m"
-  apply (insert m_gt_one)
-  apply (unfold R_def a_inv_def m_inv_def residue_ring_def)
-  apply auto
-  apply (rule the_equality)
-  apply auto
-  apply (subst mod_add_right_eq [symmetric])
-  apply auto
-  apply (subst mod_add_left_eq [symmetric])
-  apply auto
-  apply (subgoal_tac "y mod m = - x mod m")
-  apply simp
-  apply (subst zmod_eq_dvd_iff)
-  apply auto
-done
-
-lemma finite [iff]: "finite(carrier R)"
-  by (subst res_carrier_eq, auto)
-
-lemma finite_Units [iff]: "finite(Units R)"
-  by (subst res_units_eq, auto)
-
-(* The function a -> a mod m maps the integers to the 
-   residue classes. The following lemmas show that this mapping 
-   respects addition and multiplication on the integers. *)
-
-lemma mod_in_carrier [iff]: "a mod m : carrier R"
-  apply (unfold res_carrier_eq)
-  apply (insert m_gt_one, auto)
-done
-
-lemma add_cong: "(x mod m) \<oplus> (y mod m) = (x + y) mod m"
-  by (unfold R_def residue_ring_def, auto, arith)
-
-lemma mult_cong: "(x mod m) \<otimes> (y mod m) = (x * y) mod m"
-  apply (unfold R_def residue_ring_def, auto)
-  apply (subst zmod_zmult1_eq [symmetric])
-  apply (subst mult_commute)
-  apply (subst zmod_zmult1_eq [symmetric])
-  apply (subst mult_commute)
-  apply auto
-done  
-
-lemma zero_cong: "\<zero> = 0"
-  apply (unfold R_def residue_ring_def, auto)
-done
-
-lemma one_cong: "\<one> = 1 mod m"
-  apply (insert m_gt_one)
-  apply (unfold R_def residue_ring_def, auto)
-done
-
-(* revise algebra library to use 1? *)
-lemma pow_cong: "(x mod m) (^) n = x^n mod m"
-  apply (insert m_gt_one)
-  apply (induct n)
-  apply (auto simp add: nat_pow_def one_cong One_nat_def)
-  apply (subst mult_commute)
-  apply (rule mult_cong)
-done
-
-lemma neg_cong: "\<ominus> (x mod m) = (- x) mod m"
-  apply (rule sym)
-  apply (rule sum_zero_eq_neg)
-  apply auto
-  apply (subst add_cong)
-  apply (subst zero_cong)
-  apply auto
-done
-
-lemma (in residues) prod_cong: 
-  "finite A \<Longrightarrow> (\<Otimes> i:A. (f i) mod m) = (PROD i:A. f i) mod m"
-  apply (induct set: finite)
-  apply (auto simp: one_cong mult_cong)
-done
-
-lemma (in residues) sum_cong:
-  "finite A \<Longrightarrow> (\<Oplus> i:A. (f i) mod m) = (SUM i: A. f i) mod m"
-  apply (induct set: finite)
-  apply (auto simp: zero_cong add_cong)
-done
-
-lemma mod_in_res_units [simp]: "1 < m \<Longrightarrow> coprime a m \<Longrightarrow> 
-    a mod m : Units R"
-  apply (subst res_units_eq, auto)
-  apply (insert pos_mod_sign [of m a])
-  apply (subgoal_tac "a mod m ~= 0")
-  apply arith
-  apply auto
-  apply (subst (asm) gcd_red_int)
-  apply (subst gcd_commute_int, assumption)
-done
-
-lemma res_eq_to_cong: "((a mod m) = (b mod m)) = [a = b] (mod (m::int))" 
-  unfolding cong_int_def by auto
-
-(* Simplifying with these will translate a ring equation in R to a 
-   congruence. *)
-
-lemmas res_to_cong_simps = add_cong mult_cong pow_cong one_cong
-    prod_cong sum_cong neg_cong res_eq_to_cong
-
-(* Other useful facts about the residue ring *)
-
-lemma one_eq_neg_one: "\<one> = \<ominus> \<one> \<Longrightarrow> m = 2"
-  apply (simp add: res_one_eq res_neg_eq)
-  apply (insert m_gt_one)
-  apply (subgoal_tac "~(m > 2)")
-  apply arith
-  apply (rule notI)
-  apply (subgoal_tac "-1 mod m = m - 1")
-  apply force
-  apply (subst mod_add_self2 [symmetric])
-  apply (subst mod_pos_pos_trivial)
-  apply auto
-done
-
-end
-
-
-(* prime residues *)
-
-locale residues_prime =
-  fixes p :: int and R (structure)
-  assumes p_prime [intro]: "prime p"
-  defines "R == residue_ring p"
-
-sublocale residues_prime < residues p
-  apply (unfold R_def residues_def)
-  using p_prime apply auto
-done
-
-context residues_prime begin
-
-lemma is_field: "field R"
-  apply (rule cring.field_intro2)
-  apply (rule cring)
-  apply (auto simp add: res_carrier_eq res_one_eq res_zero_eq
-    res_units_eq)
-  apply (rule classical)
-  apply (erule notE)
-  apply (subst gcd_commute_int)
-  apply (rule prime_imp_coprime_int)
-  apply (rule p_prime)
-  apply (rule notI)
-  apply (frule zdvd_imp_le)
-  apply auto
-done
-
-lemma res_prime_units_eq: "Units R = {1..p - 1}"
-  apply (subst res_units_eq)
-  apply auto
-  apply (subst gcd_commute_int)
-  apply (rule prime_imp_coprime_int)
-  apply (rule p_prime)
-  apply (rule zdvd_not_zless)
-  apply auto
-done
-
-end
-
-sublocale residues_prime < field
-  by (rule is_field)
-
-
-(*
-  Test cases: Euler's theorem and Wilson's theorem.
-*)
-
-
-subsection{* Euler's theorem *}
-
-(* the definition of the phi function *)
-
-constdefs
-  phi :: "int => nat"
-  "phi m == card({ x. 0 < x & x < m & gcd x m = 1})" 
-
-lemma phi_zero [simp]: "phi 0 = 0"
-  apply (subst phi_def)
-(* Auto hangs here. Once again, where is the simplification rule 
-   1 == Suc 0 coming from? *)
-  apply (auto simp add: card_eq_0_iff)
-(* Add card_eq_0_iff as a simp rule? delete card_empty_imp? *)
-done
-
-lemma phi_one [simp]: "phi 1 = 0"
-  apply (auto simp add: phi_def card_eq_0_iff)
-done
-
-lemma (in residues) phi_eq: "phi m = card(Units R)"
-  by (simp add: phi_def res_units_eq)
-
-lemma (in residues) euler_theorem1: 
-  assumes a: "gcd a m = 1"
-  shows "[a^phi m = 1] (mod m)"
-proof -
-  from a m_gt_one have [simp]: "a mod m : Units R"
-    by (intro mod_in_res_units)
-  from phi_eq have "(a mod m) (^) (phi m) = (a mod m) (^) (card (Units R))"
-    by simp
-  also have "\<dots> = \<one>" 
-    by (intro units_power_order_eq_one, auto)
-  finally show ?thesis
-    by (simp add: res_to_cong_simps)
-qed
-
-(* In fact, there is a two line proof!
-
-lemma (in residues) euler_theorem1: 
-  assumes a: "gcd a m = 1"
-  shows "[a^phi m = 1] (mod m)"
-proof -
-  have "(a mod m) (^) (phi m) = \<one>"
-    by (simp add: phi_eq units_power_order_eq_one a m_gt_one)
-  thus ?thesis
-    by (simp add: res_to_cong_simps)
-qed
-
-*)
-
-(* outside the locale, we can relax the restriction m > 1 *)
-
-lemma euler_theorem:
-  assumes "m >= 0" and "gcd a m = 1"
-  shows "[a^phi m = 1] (mod m)"
-proof (cases)
-  assume "m = 0 | m = 1"
-  thus ?thesis by auto
-next
-  assume "~(m = 0 | m = 1)"
-  with prems show ?thesis
-    by (intro residues.euler_theorem1, unfold residues_def, auto)
-qed
-
-lemma (in residues_prime) phi_prime: "phi p = (nat p - 1)"
-  apply (subst phi_eq)
-  apply (subst res_prime_units_eq)
-  apply auto
-done
-
-lemma phi_prime: "prime p \<Longrightarrow> phi p = (nat p - 1)"
-  apply (rule residues_prime.phi_prime)
-  apply (erule residues_prime.intro)
-done
-
-lemma fermat_theorem:
-  assumes "prime p" and "~ (p dvd a)"
-  shows "[a^(nat p - 1) = 1] (mod p)"
-proof -
-  from prems have "[a^phi p = 1] (mod p)"
-    apply (intro euler_theorem)
-    (* auto should get this next part. matching across
-       substitutions is needed. *)
-    apply (frule prime_gt_1_int, arith)
-    apply (subst gcd_commute_int, erule prime_imp_coprime_int, assumption)
-    done
-  also have "phi p = nat p - 1"
-    by (rule phi_prime, rule prems)
-  finally show ?thesis .
-qed
-
-
-subsection {* Wilson's theorem *}
-
-lemma (in field) inv_pair_lemma: "x : Units R \<Longrightarrow> y : Units R \<Longrightarrow> 
-  {x, inv x} ~= {y, inv y} \<Longrightarrow> {x, inv x} Int {y, inv y} = {}" 
-  apply auto
-  apply (erule notE)
-  apply (erule inv_eq_imp_eq)
-  apply auto
-  apply (erule notE)
-  apply (erule inv_eq_imp_eq)
-  apply auto
-done
-
-lemma (in residues_prime) wilson_theorem1:
-  assumes a: "p > 2"
-  shows "[fact (p - 1) = - 1] (mod p)"
-proof -
-  let ?InversePairs = "{ {x, inv x} | x. x : Units R - {\<one>, \<ominus> \<one>}}" 
-  have UR: "Units R = {\<one>, \<ominus> \<one>} Un (Union ?InversePairs)"
-    by auto
-  have "(\<Otimes>i: Units R. i) = 
-    (\<Otimes>i: {\<one>, \<ominus> \<one>}. i) \<otimes> (\<Otimes>i: Union ?InversePairs. i)"
-    apply (subst UR)
-    apply (subst finprod_Un_disjoint)
-    apply (auto intro:funcsetI)
-    apply (drule sym, subst (asm) inv_eq_one_eq)
-    apply auto
-    apply (drule sym, subst (asm) inv_eq_neg_one_eq)
-    apply auto
-    done
-  also have "(\<Otimes>i: {\<one>, \<ominus> \<one>}. i) = \<ominus> \<one>"
-    apply (subst finprod_insert)
-    apply auto
-    apply (frule one_eq_neg_one)
-    apply (insert a, force)
-    done
-  also have "(\<Otimes>i:(Union ?InversePairs). i) = 
-      (\<Otimes> A: ?InversePairs. (\<Otimes> y:A. y))"
-    apply (subst finprod_Union_disjoint)
-    apply force
-    apply force
-    apply clarify
-    apply (rule inv_pair_lemma)
-    apply auto
-    done
-  also have "\<dots> = \<one>"
-    apply (rule finprod_one)
-    apply auto
-    apply (subst finprod_insert)
-    apply auto
-    apply (frule inv_eq_self)
-    apply (auto)
-    done
-  finally have "(\<Otimes>i: Units R. i) = \<ominus> \<one>"
-    by simp
-  also have "(\<Otimes>i: Units R. i) = (\<Otimes>i: Units R. i mod p)"
-    apply (rule finprod_cong')
-    apply (auto)
-    apply (subst (asm) res_prime_units_eq)
-    apply auto
-    done
-  also have "\<dots> = (PROD i: Units R. i) mod p"
-    apply (rule prod_cong)
-    apply auto
-    done
-  also have "\<dots> = fact (p - 1) mod p"
-    apply (subst fact_altdef_int)
-    apply (insert prems, force)
-    apply (subst res_prime_units_eq, rule refl)
-    done
-  finally have "fact (p - 1) mod p = \<ominus> \<one>".
-  thus ?thesis
-    by (simp add: res_to_cong_simps)
-qed
-
-lemma wilson_theorem: "prime (p::int) \<Longrightarrow> [fact (p - 1) = - 1] (mod p)"
-  apply (frule prime_gt_1_int)
-  apply (case_tac "p = 2")
-  apply (subst fact_altdef_int, simp)
-  apply (subst cong_int_def)
-  apply simp
-  apply (rule residues_prime.wilson_theorem1)
-  apply (rule residues_prime.intro)
-  apply auto
-done
-
-
-end
--- a/src/HOL/NewNumberTheory/UniqueFactorization.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,967 +0,0 @@
-(*  Title:      UniqueFactorization.thy
-    ID:         
-    Author:     Jeremy Avigad
-
-    
-    Unique factorization for the natural numbers and the integers.
-
-    Note: there were previous Isabelle formalizations of unique
-    factorization due to Thomas Marthedal Rasmussen, and, building on
-    that, by Jeremy Avigad and David Gray.  
-*)
-
-header {* UniqueFactorization *}
-
-theory UniqueFactorization
-imports Cong Multiset
-begin
-
-(* inherited from Multiset *)
-declare One_nat_def [simp del] 
-
-(* As a simp or intro rule,
-
-     prime p \<Longrightarrow> p > 0
-
-   wreaks havoc here. When the premise includes ALL x :# M. prime x, it 
-   leads to the backchaining
-
-     x > 0  
-     prime x 
-     x :# M   which is, unfortunately,
-     count M x > 0
-*)
-
-
-(* useful facts *)
-
-lemma setsum_Un2: "finite (A Un B) \<Longrightarrow> 
-    setsum f (A Un B) = setsum f (A - B) + setsum f (B - A) + 
-      setsum f (A Int B)"
-  apply (subgoal_tac "A Un B = (A - B) Un (B - A) Un (A Int B)")
-  apply (erule ssubst)
-  apply (subst setsum_Un_disjoint)
-  apply auto
-  apply (subst setsum_Un_disjoint)
-  apply auto
-done
-
-lemma setprod_Un2: "finite (A Un B) \<Longrightarrow> 
-    setprod f (A Un B) = setprod f (A - B) * setprod f (B - A) * 
-      setprod f (A Int B)"
-  apply (subgoal_tac "A Un B = (A - B) Un (B - A) Un (A Int B)")
-  apply (erule ssubst)
-  apply (subst setprod_Un_disjoint)
-  apply auto
-  apply (subst setprod_Un_disjoint)
-  apply auto
-done
- 
-(* Should this go in Multiset.thy? *)
-(* TN: No longer an intro-rule; needed only once and might get in the way *)
-lemma multiset_eqI: "[| !!x. count M x = count N x |] ==> M = N"
-  by (subst multiset_eq_conv_count_eq, blast)
-
-(* Here is a version of set product for multisets. Is it worth moving
-   to multiset.thy? If so, one should similarly define msetsum for abelian 
-   semirings, using of_nat. Also, is it worth developing bounded quantifiers 
-   "ALL i :# M. P i"? 
-*)
-
-constdefs
-  msetprod :: "('a => ('b::{power,comm_monoid_mult})) => 'a multiset => 'b"
-  "msetprod f M == setprod (%x. (f x)^(count M x)) (set_of M)"
-
-syntax
-  "_msetprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" 
-      ("(3PROD _:#_. _)" [0, 51, 10] 10)
-
-translations
-  "PROD i :# A. b" == "msetprod (%i. b) A"
-
-lemma msetprod_Un: "msetprod f (A+B) = msetprod f A * msetprod f B" 
-  apply (simp add: msetprod_def power_add)
-  apply (subst setprod_Un2)
-  apply auto
-  apply (subgoal_tac 
-      "(PROD x:set_of A - set_of B. f x ^ count A x * f x ^ count B x) =
-       (PROD x:set_of A - set_of B. f x ^ count A x)")
-  apply (erule ssubst)
-  apply (subgoal_tac 
-      "(PROD x:set_of B - set_of A. f x ^ count A x * f x ^ count B x) =
-       (PROD x:set_of B - set_of A. f x ^ count B x)")
-  apply (erule ssubst)
-  apply (subgoal_tac "(PROD x:set_of A. f x ^ count A x) = 
-    (PROD x:set_of A - set_of B. f x ^ count A x) *
-    (PROD x:set_of A Int set_of B. f x ^ count A x)")
-  apply (erule ssubst)
-  apply (subgoal_tac "(PROD x:set_of B. f x ^ count B x) = 
-    (PROD x:set_of B - set_of A. f x ^ count B x) *
-    (PROD x:set_of A Int set_of B. f x ^ count B x)")
-  apply (erule ssubst)
-  apply (subst setprod_timesf)
-  apply (force simp add: mult_ac)
-  apply (subst setprod_Un_disjoint [symmetric])
-  apply (auto intro: setprod_cong)
-  apply (subst setprod_Un_disjoint [symmetric])
-  apply (auto intro: setprod_cong)
-done
-
-
-subsection {* unique factorization: multiset version *}
-
-lemma multiset_prime_factorization_exists [rule_format]: "n > 0 --> 
-    (EX M. (ALL (p::nat) : set_of M. prime p) & n = (PROD i :# M. i))"
-proof (rule nat_less_induct, clarify)
-  fix n :: nat
-  assume ih: "ALL m < n. 0 < m --> (EX M. (ALL p : set_of M. prime p) & m = 
-      (PROD i :# M. i))"
-  assume "(n::nat) > 0"
-  then have "n = 1 | (n > 1 & prime n) | (n > 1 & ~ prime n)"
-    by arith
-  moreover 
-  {
-    assume "n = 1"
-    then have "(ALL p : set_of {#}. prime p) & n = (PROD i :# {#}. i)"
-        by (auto simp add: msetprod_def)
-  } 
-  moreover 
-  {
-    assume "n > 1" and "prime n"
-    then have "(ALL p : set_of {# n #}. prime p) & n = (PROD i :# {# n #}. i)"
-      by (auto simp add: msetprod_def)
-  } 
-  moreover 
-  {
-    assume "n > 1" and "~ prime n"
-    from prems not_prime_eq_prod_nat
-      obtain m k where "n = m * k & 1 < m & m < n & 1 < k & k < n"
-        by blast
-    with ih obtain Q R where "(ALL p : set_of Q. prime p) & m = (PROD i:#Q. i)"
-        and "(ALL p: set_of R. prime p) & k = (PROD i:#R. i)"
-      by blast
-    hence "(ALL p: set_of (Q + R). prime p) & n = (PROD i :# Q + R. i)"
-      by (auto simp add: prems msetprod_Un set_of_union)
-    then have "EX M. (ALL p : set_of M. prime p) & n = (PROD i :# M. i)"..
-  }
-  ultimately show "EX M. (ALL p : set_of M. prime p) & n = (PROD i::nat:#M. i)"
-    by blast
-qed
-
-lemma multiset_prime_factorization_unique_aux:
-  fixes a :: nat
-  assumes "(ALL p : set_of M. prime p)" and
-    "(ALL p : set_of N. prime p)" and
-    "(PROD i :# M. i) dvd (PROD i:# N. i)"
-  shows
-    "count M a <= count N a"
-proof cases
-  assume "a : set_of M"
-  with prems have a: "prime a"
-    by auto
-  with prems have "a ^ count M a dvd (PROD i :# M. i)"
-    by (auto intro: dvd_setprod simp add: msetprod_def)
-  also have "... dvd (PROD i :# N. i)"
-    by (rule prems)
-  also have "... = (PROD i : (set_of N). i ^ (count N i))"
-    by (simp add: msetprod_def)
-  also have "... = 
-      a^(count N a) * (PROD i : (set_of N - {a}). i ^ (count N i))"
-    proof (cases)
-      assume "a : set_of N"
-      hence b: "set_of N = {a} Un (set_of N - {a})"
-        by auto
-      thus ?thesis
-        by (subst (1) b, subst setprod_Un_disjoint, auto)
-    next
-      assume "a ~: set_of N" 
-      thus ?thesis
-        by auto
-    qed
-  finally have "a ^ count M a dvd 
-      a^(count N a) * (PROD i : (set_of N - {a}). i ^ (count N i))".
-  moreover have "coprime (a ^ count M a)
-      (PROD i : (set_of N - {a}). i ^ (count N i))"
-    apply (subst gcd_commute_nat)
-    apply (rule setprod_coprime_nat)
-    apply (rule primes_imp_powers_coprime_nat)
-    apply (insert prems, auto) 
-    done
-  ultimately have "a ^ count M a dvd a^(count N a)"
-    by (elim coprime_dvd_mult_nat)
-  with a show ?thesis 
-    by (intro power_dvd_imp_le, auto)
-next
-  assume "a ~: set_of M"
-  thus ?thesis by auto
-qed
-
-lemma multiset_prime_factorization_unique:
-  assumes "(ALL (p::nat) : set_of M. prime p)" and
-    "(ALL p : set_of N. prime p)" and
-    "(PROD i :# M. i) = (PROD i:# N. i)"
-  shows
-    "M = N"
-proof -
-  {
-    fix a
-    from prems have "count M a <= count N a"
-      by (intro multiset_prime_factorization_unique_aux, auto) 
-    moreover from prems have "count N a <= count M a"
-      by (intro multiset_prime_factorization_unique_aux, auto) 
-    ultimately have "count M a = count N a"
-      by auto
-  }
-  thus ?thesis by (simp add:multiset_eq_conv_count_eq)
-qed
-
-constdefs
-  multiset_prime_factorization :: "nat => nat multiset"
-  "multiset_prime_factorization n ==
-     if n > 0 then (THE M. ((ALL p : set_of M. prime p) & 
-       n = (PROD i :# M. i)))
-     else {#}"
-
-lemma multiset_prime_factorization: "n > 0 ==>
-    (ALL p : set_of (multiset_prime_factorization n). prime p) &
-       n = (PROD i :# (multiset_prime_factorization n). i)"
-  apply (unfold multiset_prime_factorization_def)
-  apply clarsimp
-  apply (frule multiset_prime_factorization_exists)
-  apply clarify
-  apply (rule theI)
-  apply (insert multiset_prime_factorization_unique, blast)+
-done
-
-
-subsection {* Prime factors and multiplicity for nats and ints *}
-
-class unique_factorization =
-
-fixes
-  multiplicity :: "'a \<Rightarrow> 'a \<Rightarrow> nat" and
-  prime_factors :: "'a \<Rightarrow> 'a set"
-
-(* definitions for the natural numbers *)
-
-instantiation nat :: unique_factorization
-
-begin
-
-definition
-  multiplicity_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
-where
-  "multiplicity_nat p n = count (multiset_prime_factorization n) p"
-
-definition
-  prime_factors_nat :: "nat \<Rightarrow> nat set"
-where
-  "prime_factors_nat n = set_of (multiset_prime_factorization n)"
-
-instance proof qed
-
-end
-
-(* definitions for the integers *)
-
-instantiation int :: unique_factorization
-
-begin
-
-definition
-  multiplicity_int :: "int \<Rightarrow> int \<Rightarrow> nat"
-where
-  "multiplicity_int p n = multiplicity (nat p) (nat n)"
-
-definition
-  prime_factors_int :: "int \<Rightarrow> int set"
-where
-  "prime_factors_int n = int ` (prime_factors (nat n))"
-
-instance proof qed
-
-end
-
-
-subsection {* Set up transfer *}
-
-lemma transfer_nat_int_prime_factors: 
-  "prime_factors (nat n) = nat ` prime_factors n"
-  unfolding prime_factors_int_def apply auto
-  by (subst transfer_int_nat_set_return_embed, assumption)
-
-lemma transfer_nat_int_prime_factors_closure: "n >= 0 \<Longrightarrow> 
-    nat_set (prime_factors n)"
-  by (auto simp add: nat_set_def prime_factors_int_def)
-
-lemma transfer_nat_int_multiplicity: "p >= 0 \<Longrightarrow> n >= 0 \<Longrightarrow>
-  multiplicity (nat p) (nat n) = multiplicity p n"
-  by (auto simp add: multiplicity_int_def)
-
-declare TransferMorphism_nat_int[transfer add return: 
-  transfer_nat_int_prime_factors transfer_nat_int_prime_factors_closure
-  transfer_nat_int_multiplicity]
-
-
-lemma transfer_int_nat_prime_factors:
-    "prime_factors (int n) = int ` prime_factors n"
-  unfolding prime_factors_int_def by auto
-
-lemma transfer_int_nat_prime_factors_closure: "is_nat n \<Longrightarrow> 
-    nat_set (prime_factors n)"
-  by (simp only: transfer_nat_int_prime_factors_closure is_nat_def)
-
-lemma transfer_int_nat_multiplicity: 
-    "multiplicity (int p) (int n) = multiplicity p n"
-  by (auto simp add: multiplicity_int_def)
-
-declare TransferMorphism_int_nat[transfer add return: 
-  transfer_int_nat_prime_factors transfer_int_nat_prime_factors_closure
-  transfer_int_nat_multiplicity]
-
-
-subsection {* Properties of prime factors and multiplicity for nats and ints *}
-
-lemma prime_factors_ge_0_int [elim]: "p : prime_factors (n::int) \<Longrightarrow> p >= 0"
-  by (unfold prime_factors_int_def, auto)
-
-lemma prime_factors_prime_nat [intro]: "p : prime_factors (n::nat) \<Longrightarrow> prime p"
-  apply (case_tac "n = 0")
-  apply (simp add: prime_factors_nat_def multiset_prime_factorization_def)
-  apply (auto simp add: prime_factors_nat_def multiset_prime_factorization)
-done
-
-lemma prime_factors_prime_int [intro]:
-  assumes "n >= 0" and "p : prime_factors (n::int)"
-  shows "prime p"
-
-  apply (rule prime_factors_prime_nat [transferred, of n p])
-  using prems apply auto
-done
-
-lemma prime_factors_gt_0_nat [elim]: "p : prime_factors x \<Longrightarrow> p > (0::nat)"
-  by (frule prime_factors_prime_nat, auto)
-
-lemma prime_factors_gt_0_int [elim]: "x >= 0 \<Longrightarrow> p : prime_factors x \<Longrightarrow> 
-    p > (0::int)"
-  by (frule (1) prime_factors_prime_int, auto)
-
-lemma prime_factors_finite_nat [iff]: "finite (prime_factors (n::nat))"
-  by (unfold prime_factors_nat_def, auto)
-
-lemma prime_factors_finite_int [iff]: "finite (prime_factors (n::int))"
-  by (unfold prime_factors_int_def, auto)
-
-lemma prime_factors_altdef_nat: "prime_factors (n::nat) = 
-    {p. multiplicity p n > 0}"
-  by (force simp add: prime_factors_nat_def multiplicity_nat_def)
-
-lemma prime_factors_altdef_int: "prime_factors (n::int) = 
-    {p. p >= 0 & multiplicity p n > 0}"
-  apply (unfold prime_factors_int_def multiplicity_int_def)
-  apply (subst prime_factors_altdef_nat)
-  apply (auto simp add: image_def)
-done
-
-lemma prime_factorization_nat: "(n::nat) > 0 \<Longrightarrow> 
-    n = (PROD p : prime_factors n. p^(multiplicity p n))"
-  by (frule multiset_prime_factorization, 
-    simp add: prime_factors_nat_def multiplicity_nat_def msetprod_def)
-
-thm prime_factorization_nat [transferred] 
-
-lemma prime_factorization_int: 
-  assumes "(n::int) > 0"
-  shows "n = (PROD p : prime_factors n. p^(multiplicity p n))"
-
-  apply (rule prime_factorization_nat [transferred, of n])
-  using prems apply auto
-done
-
-lemma neq_zero_eq_gt_zero_nat: "((x::nat) ~= 0) = (x > 0)"
-  by auto
-
-lemma prime_factorization_unique_nat: 
-    "S = { (p::nat) . f p > 0} \<Longrightarrow> finite S \<Longrightarrow> (ALL p : S. prime p) \<Longrightarrow>
-      n = (PROD p : S. p^(f p)) \<Longrightarrow>
-        S = prime_factors n & (ALL p. f p = multiplicity p n)"
-  apply (subgoal_tac "multiset_prime_factorization n = Abs_multiset
-      f")
-  apply (unfold prime_factors_nat_def multiplicity_nat_def)
-  apply (simp add: set_of_def count_def Abs_multiset_inverse multiset_def)
-  apply (unfold multiset_prime_factorization_def)
-  apply (subgoal_tac "n > 0")
-  prefer 2
-  apply force
-  apply (subst if_P, assumption)
-  apply (rule the1_equality)
-  apply (rule ex_ex1I)
-  apply (rule multiset_prime_factorization_exists, assumption)
-  apply (rule multiset_prime_factorization_unique)
-  apply force
-  apply force
-  apply force
-  unfolding set_of_def count_def msetprod_def
-  apply (subgoal_tac "f : multiset")
-  apply (auto simp only: Abs_multiset_inverse)
-  unfolding multiset_def apply force 
-done
-
-lemma prime_factors_characterization_nat: "S = {p. 0 < f (p::nat)} \<Longrightarrow> 
-    finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
-      prime_factors n = S"
-  by (rule prime_factorization_unique_nat [THEN conjunct1, symmetric],
-    assumption+)
-
-lemma prime_factors_characterization'_nat: 
-  "finite {p. 0 < f (p::nat)} \<Longrightarrow>
-    (ALL p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow>
-      prime_factors (PROD p | 0 < f p . p ^ f p) = {p. 0 < f p}"
-  apply (rule prime_factors_characterization_nat)
-  apply auto
-done
-
-(* A minor glitch:*)
-
-thm prime_factors_characterization'_nat 
-    [where f = "%x. f (int (x::nat))", 
-      transferred direction: nat "op <= (0::int)", rule_format]
-
-(*
-  Transfer isn't smart enough to know that the "0 < f p" should 
-  remain a comparison between nats. But the transfer still works. 
-*)
-
-lemma primes_characterization'_int [rule_format]: 
-    "finite {p. p >= 0 & 0 < f (p::int)} \<Longrightarrow>
-      (ALL p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow>
-        prime_factors (PROD p | p >=0 & 0 < f p . p ^ f p) = 
-          {p. p >= 0 & 0 < f p}"
-
-  apply (insert prime_factors_characterization'_nat 
-    [where f = "%x. f (int (x::nat))", 
-    transferred direction: nat "op <= (0::int)"])
-  apply auto
-done
-
-lemma prime_factors_characterization_int: "S = {p. 0 < f (p::int)} \<Longrightarrow> 
-    finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
-      prime_factors n = S"
-  apply simp
-  apply (subgoal_tac "{p. 0 < f p} = {p. 0 <= p & 0 < f p}")
-  apply (simp only:)
-  apply (subst primes_characterization'_int)
-  apply auto
-  apply (auto simp add: prime_ge_0_int)
-done
-
-lemma multiplicity_characterization_nat: "S = {p. 0 < f (p::nat)} \<Longrightarrow> 
-    finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
-      multiplicity p n = f p"
-  by (frule prime_factorization_unique_nat [THEN conjunct2, rule_format, 
-    symmetric], auto)
-
-lemma multiplicity_characterization'_nat: "finite {p. 0 < f (p::nat)} \<longrightarrow>
-    (ALL p. 0 < f p \<longrightarrow> prime p) \<longrightarrow>
-      multiplicity p (PROD p | 0 < f p . p ^ f p) = f p"
-  apply (rule impI)+
-  apply (rule multiplicity_characterization_nat)
-  apply auto
-done
-
-lemma multiplicity_characterization'_int [rule_format]: 
-  "finite {p. p >= 0 & 0 < f (p::int)} \<Longrightarrow>
-    (ALL p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow> p >= 0 \<Longrightarrow>
-      multiplicity p (PROD p | p >= 0 & 0 < f p . p ^ f p) = f p"
-
-  apply (insert multiplicity_characterization'_nat 
-    [where f = "%x. f (int (x::nat))", 
-      transferred direction: nat "op <= (0::int)", rule_format])
-  apply auto
-done
-
-lemma multiplicity_characterization_int: "S = {p. 0 < f (p::int)} \<Longrightarrow> 
-    finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
-      p >= 0 \<Longrightarrow> multiplicity p n = f p"
-  apply simp
-  apply (subgoal_tac "{p. 0 < f p} = {p. 0 <= p & 0 < f p}")
-  apply (simp only:)
-  apply (subst multiplicity_characterization'_int)
-  apply auto
-  apply (auto simp add: prime_ge_0_int)
-done
-
-lemma multiplicity_zero_nat [simp]: "multiplicity (p::nat) 0 = 0"
-  by (simp add: multiplicity_nat_def multiset_prime_factorization_def)
-
-lemma multiplicity_zero_int [simp]: "multiplicity (p::int) 0 = 0"
-  by (simp add: multiplicity_int_def) 
-
-lemma multiplicity_one_nat [simp]: "multiplicity p (1::nat) = 0"
-  by (subst multiplicity_characterization_nat [where f = "%x. 0"], auto)
-
-lemma multiplicity_one_int [simp]: "multiplicity p (1::int) = 0"
-  by (simp add: multiplicity_int_def)
-
-lemma multiplicity_prime_nat [simp]: "prime (p::nat) \<Longrightarrow> multiplicity p p = 1"
-  apply (subst multiplicity_characterization_nat
-      [where f = "(%q. if q = p then 1 else 0)"])
-  apply auto
-  apply (case_tac "x = p")
-  apply auto
-done
-
-lemma multiplicity_prime_int [simp]: "prime (p::int) \<Longrightarrow> multiplicity p p = 1"
-  unfolding prime_int_def multiplicity_int_def by auto
-
-lemma multiplicity_prime_power_nat [simp]: "prime (p::nat) \<Longrightarrow> 
-    multiplicity p (p^n) = n"
-  apply (case_tac "n = 0")
-  apply auto
-  apply (subst multiplicity_characterization_nat
-      [where f = "(%q. if q = p then n else 0)"])
-  apply auto
-  apply (case_tac "x = p")
-  apply auto
-done
-
-lemma multiplicity_prime_power_int [simp]: "prime (p::int) \<Longrightarrow> 
-    multiplicity p (p^n) = n"
-  apply (frule prime_ge_0_int)
-  apply (auto simp add: prime_int_def multiplicity_int_def nat_power_eq)
-done
-
-lemma multiplicity_nonprime_nat [simp]: "~ prime (p::nat) \<Longrightarrow> 
-    multiplicity p n = 0"
-  apply (case_tac "n = 0")
-  apply auto
-  apply (frule multiset_prime_factorization)
-  apply (auto simp add: set_of_def multiplicity_nat_def)
-done
-
-lemma multiplicity_nonprime_int [simp]: "~ prime (p::int) \<Longrightarrow> multiplicity p n = 0"
-  by (unfold multiplicity_int_def prime_int_def, auto)
-
-lemma multiplicity_not_factor_nat [simp]: 
-    "p ~: prime_factors (n::nat) \<Longrightarrow> multiplicity p n = 0"
-  by (subst (asm) prime_factors_altdef_nat, auto)
-
-lemma multiplicity_not_factor_int [simp]: 
-    "p >= 0 \<Longrightarrow> p ~: prime_factors (n::int) \<Longrightarrow> multiplicity p n = 0"
-  by (subst (asm) prime_factors_altdef_int, auto)
-
-lemma multiplicity_product_aux_nat: "(k::nat) > 0 \<Longrightarrow> l > 0 \<Longrightarrow>
-    (prime_factors k) Un (prime_factors l) = prime_factors (k * l) &
-    (ALL p. multiplicity p k + multiplicity p l = multiplicity p (k * l))"
-  apply (rule prime_factorization_unique_nat)
-  apply (simp only: prime_factors_altdef_nat)
-  apply auto
-  apply (subst power_add)
-  apply (subst setprod_timesf)
-  apply (rule arg_cong2)back back
-  apply (subgoal_tac "prime_factors k Un prime_factors l = prime_factors k Un 
-      (prime_factors l - prime_factors k)")
-  apply (erule ssubst)
-  apply (subst setprod_Un_disjoint)
-  apply auto
-  apply (subgoal_tac "(\<Prod>p\<in>prime_factors l - prime_factors k. p ^ multiplicity p k) = 
-      (\<Prod>p\<in>prime_factors l - prime_factors k. 1)")
-  apply (erule ssubst)
-  apply (simp add: setprod_1)
-  apply (erule prime_factorization_nat)
-  apply (rule setprod_cong, auto)
-  apply (subgoal_tac "prime_factors k Un prime_factors l = prime_factors l Un 
-      (prime_factors k - prime_factors l)")
-  apply (erule ssubst)
-  apply (subst setprod_Un_disjoint)
-  apply auto
-  apply (subgoal_tac "(\<Prod>p\<in>prime_factors k - prime_factors l. p ^ multiplicity p l) = 
-      (\<Prod>p\<in>prime_factors k - prime_factors l. 1)")
-  apply (erule ssubst)
-  apply (simp add: setprod_1)
-  apply (erule prime_factorization_nat)
-  apply (rule setprod_cong, auto)
-done
-
-(* transfer doesn't have the same problem here with the right 
-   choice of rules. *)
-
-lemma multiplicity_product_aux_int: 
-  assumes "(k::int) > 0" and "l > 0"
-  shows 
-    "(prime_factors k) Un (prime_factors l) = prime_factors (k * l) &
-    (ALL p >= 0. multiplicity p k + multiplicity p l = multiplicity p (k * l))"
-
-  apply (rule multiplicity_product_aux_nat [transferred, of l k])
-  using prems apply auto
-done
-
-lemma prime_factors_product_nat: "(k::nat) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> prime_factors (k * l) = 
-    prime_factors k Un prime_factors l"
-  by (rule multiplicity_product_aux_nat [THEN conjunct1, symmetric])
-
-lemma prime_factors_product_int: "(k::int) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> prime_factors (k * l) = 
-    prime_factors k Un prime_factors l"
-  by (rule multiplicity_product_aux_int [THEN conjunct1, symmetric])
-
-lemma multiplicity_product_nat: "(k::nat) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> multiplicity p (k * l) = 
-    multiplicity p k + multiplicity p l"
-  by (rule multiplicity_product_aux_nat [THEN conjunct2, rule_format, 
-      symmetric])
-
-lemma multiplicity_product_int: "(k::int) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> p >= 0 \<Longrightarrow> 
-    multiplicity p (k * l) = multiplicity p k + multiplicity p l"
-  by (rule multiplicity_product_aux_int [THEN conjunct2, rule_format, 
-      symmetric])
-
-lemma multiplicity_setprod_nat: "finite S \<Longrightarrow> (ALL x : S. f x > 0) \<Longrightarrow> 
-    multiplicity (p::nat) (PROD x : S. f x) = 
-      (SUM x : S. multiplicity p (f x))"
-  apply (induct set: finite)
-  apply auto
-  apply (subst multiplicity_product_nat)
-  apply auto
-done
-
-(* Transfer is delicate here for two reasons: first, because there is
-   an implicit quantifier over functions (f), and, second, because the 
-   product over the multiplicity should not be translated to an integer 
-   product.
-
-   The way to handle the first is to use quantifier rules for functions.
-   The way to handle the second is to turn off the offending rule.
-*)
-
-lemma transfer_nat_int_sum_prod_closure3:
-  "(SUM x : A. int (f x)) >= 0"
-  "(PROD x : A. int (f x)) >= 0"
-  apply (rule setsum_nonneg, auto)
-  apply (rule setprod_nonneg, auto)
-done
-
-declare TransferMorphism_nat_int[transfer 
-  add return: transfer_nat_int_sum_prod_closure3
-  del: transfer_nat_int_sum_prod2 (1)]
-
-lemma multiplicity_setprod_int: "p >= 0 \<Longrightarrow> finite S \<Longrightarrow> 
-  (ALL x : S. f x > 0) \<Longrightarrow> 
-    multiplicity (p::int) (PROD x : S. f x) = 
-      (SUM x : S. multiplicity p (f x))"
-
-  apply (frule multiplicity_setprod_nat
-    [where f = "%x. nat(int(nat(f x)))", 
-      transferred direction: nat "op <= (0::int)"])
-  apply auto
-  apply (subst (asm) setprod_cong)
-  apply (rule refl)
-  apply (rule if_P)
-  apply auto
-  apply (rule setsum_cong)
-  apply auto
-done
-
-declare TransferMorphism_nat_int[transfer 
-  add return: transfer_nat_int_sum_prod2 (1)]
-
-lemma multiplicity_prod_prime_powers_nat:
-    "finite S \<Longrightarrow> (ALL p : S. prime (p::nat)) \<Longrightarrow>
-       multiplicity p (PROD p : S. p ^ f p) = (if p : S then f p else 0)"
-  apply (subgoal_tac "(PROD p : S. p ^ f p) = 
-      (PROD p : S. p ^ (%x. if x : S then f x else 0) p)")
-  apply (erule ssubst)
-  apply (subst multiplicity_characterization_nat)
-  prefer 5 apply (rule refl)
-  apply (rule refl)
-  apply auto
-  apply (subst setprod_mono_one_right)
-  apply assumption
-  prefer 3
-  apply (rule setprod_cong)
-  apply (rule refl)
-  apply auto
-done
-
-(* Here the issue with transfer is the implicit quantifier over S *)
-
-lemma multiplicity_prod_prime_powers_int:
-    "(p::int) >= 0 \<Longrightarrow> finite S \<Longrightarrow> (ALL p : S. prime p) \<Longrightarrow>
-       multiplicity p (PROD p : S. p ^ f p) = (if p : S then f p else 0)"
-
-  apply (subgoal_tac "int ` nat ` S = S")
-  apply (frule multiplicity_prod_prime_powers_nat [where f = "%x. f(int x)" 
-    and S = "nat ` S", transferred])
-  apply auto
-  apply (subst prime_int_def [symmetric])
-  apply auto
-  apply (subgoal_tac "xb >= 0")
-  apply force
-  apply (rule prime_ge_0_int)
-  apply force
-  apply (subst transfer_nat_int_set_return_embed)
-  apply (unfold nat_set_def, auto)
-done
-
-lemma multiplicity_distinct_prime_power_nat: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow>
-    p ~= q \<Longrightarrow> multiplicity p (q^n) = 0"
-  apply (subgoal_tac "q^n = setprod (%x. x^n) {q}")
-  apply (erule ssubst)
-  apply (subst multiplicity_prod_prime_powers_nat)
-  apply auto
-done
-
-lemma multiplicity_distinct_prime_power_int: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow>
-    p ~= q \<Longrightarrow> multiplicity p (q^n) = 0"
-  apply (frule prime_ge_0_int [of q])
-  apply (frule multiplicity_distinct_prime_power_nat [transferred leaving: n]) 
-  prefer 4
-  apply assumption
-  apply auto
-done
-
-lemma dvd_multiplicity_nat: 
-    "(0::nat) < y \<Longrightarrow> x dvd y \<Longrightarrow> multiplicity p x <= multiplicity p y"
-  apply (case_tac "x = 0")
-  apply (auto simp add: dvd_def multiplicity_product_nat)
-done
-
-lemma dvd_multiplicity_int: 
-    "(0::int) < y \<Longrightarrow> 0 <= x \<Longrightarrow> x dvd y \<Longrightarrow> p >= 0 \<Longrightarrow> 
-      multiplicity p x <= multiplicity p y"
-  apply (case_tac "x = 0")
-  apply (auto simp add: dvd_def)
-  apply (subgoal_tac "0 < k")
-  apply (auto simp add: multiplicity_product_int)
-  apply (erule zero_less_mult_pos)
-  apply arith
-done
-
-lemma dvd_prime_factors_nat [intro]:
-    "0 < (y::nat) \<Longrightarrow> x dvd y \<Longrightarrow> prime_factors x <= prime_factors y"
-  apply (simp only: prime_factors_altdef_nat)
-  apply auto
-  apply (frule dvd_multiplicity_nat)
-  apply auto
-(* It is a shame that auto and arith don't get this. *)
-  apply (erule order_less_le_trans)back
-  apply assumption
-done
-
-lemma dvd_prime_factors_int [intro]:
-    "0 < (y::int) \<Longrightarrow> 0 <= x \<Longrightarrow> x dvd y \<Longrightarrow> prime_factors x <= prime_factors y"
-  apply (auto simp add: prime_factors_altdef_int)
-  apply (erule order_less_le_trans)
-  apply (rule dvd_multiplicity_int)
-  apply auto
-done
-
-lemma multiplicity_dvd_nat: "0 < (x::nat) \<Longrightarrow> 0 < y \<Longrightarrow> 
-    ALL p. multiplicity p x <= multiplicity p y \<Longrightarrow>
-      x dvd y"
-  apply (subst prime_factorization_nat [of x], assumption)
-  apply (subst prime_factorization_nat [of y], assumption)
-  apply (rule setprod_dvd_setprod_subset2)
-  apply force
-  apply (subst prime_factors_altdef_nat)+
-  apply auto
-(* Again, a shame that auto and arith don't get this. *)
-  apply (drule_tac x = xa in spec, auto)
-  apply (rule le_imp_power_dvd)
-  apply blast
-done
-
-lemma multiplicity_dvd_int: "0 < (x::int) \<Longrightarrow> 0 < y \<Longrightarrow> 
-    ALL p >= 0. multiplicity p x <= multiplicity p y \<Longrightarrow>
-      x dvd y"
-  apply (subst prime_factorization_int [of x], assumption)
-  apply (subst prime_factorization_int [of y], assumption)
-  apply (rule setprod_dvd_setprod_subset2)
-  apply force
-  apply (subst prime_factors_altdef_int)+
-  apply auto
-  apply (rule dvd_power_le)
-  apply auto
-  apply (drule_tac x = xa in spec)
-  apply (erule impE)
-  apply auto
-done
-
-lemma multiplicity_dvd'_nat: "(0::nat) < x \<Longrightarrow> 
-    \<forall>p. prime p \<longrightarrow> multiplicity p x \<le> multiplicity p y \<Longrightarrow> x dvd y"
-  apply (cases "y = 0")
-  apply auto
-  apply (rule multiplicity_dvd_nat, auto)
-  apply (case_tac "prime p")
-  apply auto
-done
-
-lemma multiplicity_dvd'_int: "(0::int) < x \<Longrightarrow> 0 <= y \<Longrightarrow>
-    \<forall>p. prime p \<longrightarrow> multiplicity p x \<le> multiplicity p y \<Longrightarrow> x dvd y"
-  apply (cases "y = 0")
-  apply auto
-  apply (rule multiplicity_dvd_int, auto)
-  apply (case_tac "prime p")
-  apply auto
-done
-
-lemma dvd_multiplicity_eq_nat: "0 < (x::nat) \<Longrightarrow> 0 < y \<Longrightarrow>
-    (x dvd y) = (ALL p. multiplicity p x <= multiplicity p y)"
-  by (auto intro: dvd_multiplicity_nat multiplicity_dvd_nat)
-
-lemma dvd_multiplicity_eq_int: "0 < (x::int) \<Longrightarrow> 0 < y \<Longrightarrow>
-    (x dvd y) = (ALL p >= 0. multiplicity p x <= multiplicity p y)"
-  by (auto intro: dvd_multiplicity_int multiplicity_dvd_int)
-
-lemma prime_factors_altdef2_nat: "(n::nat) > 0 \<Longrightarrow> 
-    (p : prime_factors n) = (prime p & p dvd n)"
-  apply (case_tac "prime p")
-  apply auto
-  apply (subst prime_factorization_nat [where n = n], assumption)
-  apply (rule dvd_trans) 
-  apply (rule dvd_power [where x = p and n = "multiplicity p n"])
-  apply (subst (asm) prime_factors_altdef_nat, force)
-  apply (rule dvd_setprod)
-  apply auto  
-  apply (subst prime_factors_altdef_nat)
-  apply (subst (asm) dvd_multiplicity_eq_nat)
-  apply auto
-  apply (drule spec [where x = p])
-  apply auto
-done
-
-lemma prime_factors_altdef2_int: 
-  assumes "(n::int) > 0" 
-  shows "(p : prime_factors n) = (prime p & p dvd n)"
-
-  apply (case_tac "p >= 0")
-  apply (rule prime_factors_altdef2_nat [transferred])
-  using prems apply auto
-  apply (auto simp add: prime_ge_0_int prime_factors_ge_0_int)
-done
-
-lemma multiplicity_eq_nat:
-  fixes x and y::nat 
-  assumes [arith]: "x > 0" "y > 0" and
-    mult_eq [simp]: "!!p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
-  shows "x = y"
-
-  apply (rule dvd_anti_sym)
-  apply (auto intro: multiplicity_dvd'_nat) 
-done
-
-lemma multiplicity_eq_int:
-  fixes x and y::int 
-  assumes [arith]: "x > 0" "y > 0" and
-    mult_eq [simp]: "!!p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
-  shows "x = y"
-
-  apply (rule dvd_anti_sym [transferred])
-  apply (auto intro: multiplicity_dvd'_int) 
-done
-
-
-subsection {* An application *}
-
-lemma gcd_eq_nat: 
-  assumes pos [arith]: "x > 0" "y > 0"
-  shows "gcd (x::nat) y = 
-    (PROD p: prime_factors x Un prime_factors y. 
-      p ^ (min (multiplicity p x) (multiplicity p y)))"
-proof -
-  def z == "(PROD p: prime_factors (x::nat) Un prime_factors y. 
-      p ^ (min (multiplicity p x) (multiplicity p y)))"
-  have [arith]: "z > 0"
-    unfolding z_def by (rule setprod_pos_nat, auto)
-  have aux: "!!p. prime p \<Longrightarrow> multiplicity p z = 
-      min (multiplicity p x) (multiplicity p y)"
-    unfolding z_def
-    apply (subst multiplicity_prod_prime_powers_nat)
-    apply (auto simp add: multiplicity_not_factor_nat)
-    done
-  have "z dvd x" 
-    by (intro multiplicity_dvd'_nat, auto simp add: aux)
-  moreover have "z dvd y" 
-    by (intro multiplicity_dvd'_nat, auto simp add: aux)
-  moreover have "ALL w. w dvd x & w dvd y \<longrightarrow> w dvd z"
-    apply auto
-    apply (case_tac "w = 0", auto)
-    apply (erule multiplicity_dvd'_nat)
-    apply (auto intro: dvd_multiplicity_nat simp add: aux)
-    done
-  ultimately have "z = gcd x y"
-    by (subst gcd_unique_nat [symmetric], blast)
-  thus ?thesis
-    unfolding z_def by auto
-qed
-
-lemma lcm_eq_nat: 
-  assumes pos [arith]: "x > 0" "y > 0"
-  shows "lcm (x::nat) y = 
-    (PROD p: prime_factors x Un prime_factors y. 
-      p ^ (max (multiplicity p x) (multiplicity p y)))"
-proof -
-  def z == "(PROD p: prime_factors (x::nat) Un prime_factors y. 
-      p ^ (max (multiplicity p x) (multiplicity p y)))"
-  have [arith]: "z > 0"
-    unfolding z_def by (rule setprod_pos_nat, auto)
-  have aux: "!!p. prime p \<Longrightarrow> multiplicity p z = 
-      max (multiplicity p x) (multiplicity p y)"
-    unfolding z_def
-    apply (subst multiplicity_prod_prime_powers_nat)
-    apply (auto simp add: multiplicity_not_factor_nat)
-    done
-  have "x dvd z" 
-    by (intro multiplicity_dvd'_nat, auto simp add: aux)
-  moreover have "y dvd z" 
-    by (intro multiplicity_dvd'_nat, auto simp add: aux)
-  moreover have "ALL w. x dvd w & y dvd w \<longrightarrow> z dvd w"
-    apply auto
-    apply (case_tac "w = 0", auto)
-    apply (rule multiplicity_dvd'_nat)
-    apply (auto intro: dvd_multiplicity_nat simp add: aux)
-    done
-  ultimately have "z = lcm x y"
-    by (subst lcm_unique_nat [symmetric], blast)
-  thus ?thesis
-    unfolding z_def by auto
-qed
-
-lemma multiplicity_gcd_nat: 
-  assumes [arith]: "x > 0" "y > 0"
-  shows "multiplicity (p::nat) (gcd x y) = 
-    min (multiplicity p x) (multiplicity p y)"
-
-  apply (subst gcd_eq_nat)
-  apply auto
-  apply (subst multiplicity_prod_prime_powers_nat)
-  apply auto
-done
-
-lemma multiplicity_lcm_nat: 
-  assumes [arith]: "x > 0" "y > 0"
-  shows "multiplicity (p::nat) (lcm x y) = 
-    max (multiplicity p x) (multiplicity p y)"
-
-  apply (subst lcm_eq_nat)
-  apply auto
-  apply (subst multiplicity_prod_prime_powers_nat)
-  apply auto
-done
-
-lemma gcd_lcm_distrib_nat: "gcd (x::nat) (lcm y z) = lcm (gcd x y) (gcd x z)"
-  apply (case_tac "x = 0 | y = 0 | z = 0") 
-  apply auto
-  apply (rule multiplicity_eq_nat)
-  apply (auto simp add: multiplicity_gcd_nat multiplicity_lcm_nat 
-      lcm_pos_nat)
-done
-
-lemma gcd_lcm_distrib_int: "gcd (x::int) (lcm y z) = lcm (gcd x y) (gcd x z)"
-  apply (subst (1 2 3) gcd_abs_int)
-  apply (subst lcm_abs_int)
-  apply (subst (2) abs_of_nonneg)
-  apply force
-  apply (rule gcd_lcm_distrib_nat [transferred])
-  apply auto
-done
-
-end
--- a/src/HOL/Nominal/nominal_thmdecls.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Nominal/nominal_thmdecls.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -74,7 +74,7 @@
     val lhs = Const (@{const_name "perm"}, typi --> HOLogic.boolT --> HOLogic.boolT) $ pi' $ hyp;
     val ([goal_term, pi''], ctxt') = Variable.import_terms false
       [HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, concl)), pi'] ctxt
-    val _ = Display.print_cterm (cterm_of thy goal_term)
+    val _ = writeln (Syntax.string_of_term ctxt' goal_term);
   in
     Goal.prove ctxt' [] [] goal_term
       (fn _ => prove_eqvt_tac thy orig_thm pi' pi'') |>
--- a/src/HOL/NumberTheory/BijectionRel.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,231 +0,0 @@
-(*  Title:      HOL/NumberTheory/BijectionRel.thy
-    ID:         $Id$
-    Author:     Thomas M. Rasmussen
-    Copyright   2000  University of Cambridge
-*)
-
-header {* Bijections between sets *}
-
-theory BijectionRel imports Main begin
-
-text {*
-  Inductive definitions of bijections between two different sets and
-  between the same set.  Theorem for relating the two definitions.
-
-  \bigskip
-*}
-
-inductive_set
-  bijR :: "('a => 'b => bool) => ('a set * 'b set) set"
-  for P :: "'a => 'b => bool"
-where
-  empty [simp]: "({}, {}) \<in> bijR P"
-| insert: "P a b ==> a \<notin> A ==> b \<notin> B ==> (A, B) \<in> bijR P
-    ==> (insert a A, insert b B) \<in> bijR P"
-
-text {*
-  Add extra condition to @{term insert}: @{term "\<forall>b \<in> B. \<not> P a b"}
-  (and similar for @{term A}).
-*}
-
-definition
-  bijP :: "('a => 'a => bool) => 'a set => bool" where
-  "bijP P F = (\<forall>a b. a \<in> F \<and> P a b --> b \<in> F)"
-
-definition
-  uniqP :: "('a => 'a => bool) => bool" where
-  "uniqP P = (\<forall>a b c d. P a b \<and> P c d --> (a = c) = (b = d))"
-
-definition
-  symP :: "('a => 'a => bool) => bool" where
-  "symP P = (\<forall>a b. P a b = P b a)"
-
-inductive_set
-  bijER :: "('a => 'a => bool) => 'a set set"
-  for P :: "'a => 'a => bool"
-where
-  empty [simp]: "{} \<in> bijER P"
-| insert1: "P a a ==> a \<notin> A ==> A \<in> bijER P ==> insert a A \<in> bijER P"
-| insert2: "P a b ==> a \<noteq> b ==> a \<notin> A ==> b \<notin> A ==> A \<in> bijER P
-    ==> insert a (insert b A) \<in> bijER P"
-
-
-text {* \medskip @{term bijR} *}
-
-lemma fin_bijRl: "(A, B) \<in> bijR P ==> finite A"
-  apply (erule bijR.induct)
-  apply auto
-  done
-
-lemma fin_bijRr: "(A, B) \<in> bijR P ==> finite B"
-  apply (erule bijR.induct)
-  apply auto
-  done
-
-lemma aux_induct:
-  assumes major: "finite F"
-    and subs: "F \<subseteq> A"
-    and cases: "P {}"
-      "!!F a. F \<subseteq> A ==> a \<in> A ==> a \<notin> F ==> P F ==> P (insert a F)"
-  shows "P F"
-  using major subs
-  apply (induct set: finite)
-   apply (blast intro: cases)+
-  done
-
-
-lemma inj_func_bijR_aux1:
-    "A \<subseteq> B ==> a \<notin> A ==> a \<in> B ==> inj_on f B ==> f a \<notin> f ` A"
-  apply (unfold inj_on_def)
-  apply auto
-  done
-
-lemma inj_func_bijR_aux2:
-  "\<forall>a. a \<in> A --> P a (f a) ==> inj_on f A ==> finite A ==> F <= A
-    ==> (F, f ` F) \<in> bijR P"
-  apply (rule_tac F = F and A = A in aux_induct)
-     apply (rule finite_subset)
-      apply auto
-  apply (rule bijR.insert)
-     apply (rule_tac [3] inj_func_bijR_aux1)
-        apply auto
-  done
-
-lemma inj_func_bijR:
-  "\<forall>a. a \<in> A --> P a (f a) ==> inj_on f A ==> finite A
-    ==> (A, f ` A) \<in> bijR P"
-  apply (rule inj_func_bijR_aux2)
-     apply auto
-  done
-
-
-text {* \medskip @{term bijER} *}
-
-lemma fin_bijER: "A \<in> bijER P ==> finite A"
-  apply (erule bijER.induct)
-    apply auto
-  done
-
-lemma aux1:
-  "a \<notin> A ==> a \<notin> B ==> F \<subseteq> insert a A ==> F \<subseteq> insert a B ==> a \<in> F
-    ==> \<exists>C. F = insert a C \<and> a \<notin> C \<and> C <= A \<and> C <= B"
-  apply (rule_tac x = "F - {a}" in exI)
-  apply auto
-  done
-
-lemma aux2: "a \<noteq> b ==> a \<notin> A ==> b \<notin> B ==> a \<in> F ==> b \<in> F
-    ==> F \<subseteq> insert a A ==> F \<subseteq> insert b B
-    ==> \<exists>C. F = insert a (insert b C) \<and> a \<notin> C \<and> b \<notin> C \<and> C \<subseteq> A \<and> C \<subseteq> B"
-  apply (rule_tac x = "F - {a, b}" in exI)
-  apply auto
-  done
-
-lemma aux_uniq: "uniqP P ==> P a b ==> P c d ==> (a = c) = (b = d)"
-  apply (unfold uniqP_def)
-  apply auto
-  done
-
-lemma aux_sym: "symP P ==> P a b = P b a"
-  apply (unfold symP_def)
-  apply auto
-  done
-
-lemma aux_in1:
-    "uniqP P ==> b \<notin> C ==> P b b ==> bijP P (insert b C) ==> bijP P C"
-  apply (unfold bijP_def)
-  apply auto
-  apply (subgoal_tac "b \<noteq> a")
-   prefer 2
-   apply clarify
-  apply (simp add: aux_uniq)
-  apply auto
-  done
-
-lemma aux_in2:
-  "symP P ==> uniqP P ==> a \<notin> C ==> b \<notin> C ==> a \<noteq> b ==> P a b
-    ==> bijP P (insert a (insert b C)) ==> bijP P C"
-  apply (unfold bijP_def)
-  apply auto
-  apply (subgoal_tac "aa \<noteq> a")
-   prefer 2
-   apply clarify
-  apply (subgoal_tac "aa \<noteq> b")
-   prefer 2
-   apply clarify
-  apply (simp add: aux_uniq)
-  apply (subgoal_tac "ba \<noteq> a")
-   apply auto
-  apply (subgoal_tac "P a aa")
-   prefer 2
-   apply (simp add: aux_sym)
-  apply (subgoal_tac "b = aa")
-   apply (rule_tac [2] iffD1)
-    apply (rule_tac [2] a = a and c = a and P = P in aux_uniq)
-      apply auto
-  done
-
-lemma aux_foo: "\<forall>a b. Q a \<and> P a b --> R b ==> P a b ==> Q a ==> R b"
-  apply auto
-  done
-
-lemma aux_bij: "bijP P F ==> symP P ==> P a b ==> (a \<in> F) = (b \<in> F)"
-  apply (unfold bijP_def)
-  apply (rule iffI)
-  apply (erule_tac [!] aux_foo)
-      apply simp_all
-  apply (rule iffD2)
-   apply (rule_tac P = P in aux_sym)
-   apply simp_all
-  done
-
-
-lemma aux_bijRER:
-  "(A, B) \<in> bijR P ==> uniqP P ==> symP P
-    ==> \<forall>F. bijP P F \<and> F \<subseteq> A \<and> F \<subseteq> B --> F \<in> bijER P"
-  apply (erule bijR.induct)
-   apply simp
-  apply (case_tac "a = b")
-   apply clarify
-   apply (case_tac "b \<in> F")
-    prefer 2
-    apply (simp add: subset_insert)
-   apply (cut_tac F = F and a = b and A = A and B = B in aux1)
-        prefer 6
-        apply clarify
-        apply (rule bijER.insert1)
-          apply simp_all
-   apply (subgoal_tac "bijP P C")
-    apply simp
-   apply (rule aux_in1)
-      apply simp_all
-  apply clarify
-  apply (case_tac "a \<in> F")
-   apply (case_tac [!] "b \<in> F")
-     apply (cut_tac F = F and a = a and b = b and A = A and B = B
-       in aux2)
-            apply (simp_all add: subset_insert)
-    apply clarify
-    apply (rule bijER.insert2)
-        apply simp_all
-    apply (subgoal_tac "bijP P C")
-     apply simp
-    apply (rule aux_in2)
-          apply simp_all
-   apply (subgoal_tac "b \<in> F")
-    apply (rule_tac [2] iffD1)
-     apply (rule_tac [2] a = a and F = F and P = P in aux_bij)
-       apply (simp_all (no_asm_simp))
-   apply (subgoal_tac [2] "a \<in> F")
-    apply (rule_tac [3] iffD2)
-     apply (rule_tac [3] b = b and F = F and P = P in aux_bij)
-       apply auto
-  done
-
-lemma bijR_bijER:
-  "(A, A) \<in> bijR P ==>
-    bijP P A ==> uniqP P ==> symP P ==> A \<in> bijER P"
-  apply (cut_tac A = A and B = A and P = P in aux_bijRER)
-     apply auto
-  done
-
-end
--- a/src/HOL/NumberTheory/Chinese.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,259 +0,0 @@
-(*  Title:      HOL/NumberTheory/Chinese.thy
-    ID:         $Id$
-    Author:     Thomas M. Rasmussen
-    Copyright   2000  University of Cambridge
-*)
-
-header {* The Chinese Remainder Theorem *}
-
-theory Chinese 
-imports IntPrimes
-begin
-
-text {*
-  The Chinese Remainder Theorem for an arbitrary finite number of
-  equations.  (The one-equation case is included in theory @{text
-  IntPrimes}.  Uses functions for indexing.\footnote{Maybe @{term
-  funprod} and @{term funsum} should be based on general @{term fold}
-  on indices?}
-*}
-
-
-subsection {* Definitions *}
-
-consts
-  funprod :: "(nat => int) => nat => nat => int"
-  funsum :: "(nat => int) => nat => nat => int"
-
-primrec
-  "funprod f i 0 = f i"
-  "funprod f i (Suc n) = f (Suc (i + n)) * funprod f i n"
-
-primrec
-  "funsum f i 0 = f i"
-  "funsum f i (Suc n) = f (Suc (i + n)) + funsum f i n"
-
-definition
-  m_cond :: "nat => (nat => int) => bool" where
-  "m_cond n mf =
-    ((\<forall>i. i \<le> n --> 0 < mf i) \<and>
-      (\<forall>i j. i \<le> n \<and> j \<le> n \<and> i \<noteq> j --> zgcd (mf i) (mf j) = 1))"
-
-definition
-  km_cond :: "nat => (nat => int) => (nat => int) => bool" where
-  "km_cond n kf mf = (\<forall>i. i \<le> n --> zgcd (kf i) (mf i) = 1)"
-
-definition
-  lincong_sol ::
-    "nat => (nat => int) => (nat => int) => (nat => int) => int => bool" where
-  "lincong_sol n kf bf mf x = (\<forall>i. i \<le> n --> zcong (kf i * x) (bf i) (mf i))"
-
-definition
-  mhf :: "(nat => int) => nat => nat => int" where
-  "mhf mf n i =
-    (if i = 0 then funprod mf (Suc 0) (n - Suc 0)
-     else if i = n then funprod mf 0 (n - Suc 0)
-     else funprod mf 0 (i - Suc 0) * funprod mf (Suc i) (n - Suc 0 - i))"
-
-definition
-  xilin_sol ::
-    "nat => nat => (nat => int) => (nat => int) => (nat => int) => int" where
-  "xilin_sol i n kf bf mf =
-    (if 0 < n \<and> i \<le> n \<and> m_cond n mf \<and> km_cond n kf mf then
-        (SOME x. 0 \<le> x \<and> x < mf i \<and> zcong (kf i * mhf mf n i * x) (bf i) (mf i))
-     else 0)"
-
-definition
-  x_sol :: "nat => (nat => int) => (nat => int) => (nat => int) => int" where
-  "x_sol n kf bf mf = funsum (\<lambda>i. xilin_sol i n kf bf mf * mhf mf n i) 0 n"
-
-
-text {* \medskip @{term funprod} and @{term funsum} *}
-
-lemma funprod_pos: "(\<forall>i. i \<le> n --> 0 < mf i) ==> 0 < funprod mf 0 n"
-  apply (induct n)
-   apply auto
-  apply (simp add: zero_less_mult_iff)
-  done
-
-lemma funprod_zgcd [rule_format (no_asm)]:
-  "(\<forall>i. k \<le> i \<and> i \<le> k + l --> zgcd (mf i) (mf m) = 1) -->
-    zgcd (funprod mf k l) (mf m) = 1"
-  apply (induct l)
-   apply simp_all
-  apply (rule impI)+
-  apply (subst zgcd_zmult_cancel)
-  apply auto
-  done
-
-lemma funprod_zdvd [rule_format]:
-    "k \<le> i --> i \<le> k + l --> mf i dvd funprod mf k l"
-  apply (induct l)
-   apply auto
-  apply (subgoal_tac "i = Suc (k + l)")
-   apply (simp_all (no_asm_simp))
-  done
-
-lemma funsum_mod:
-    "funsum f k l mod m = funsum (\<lambda>i. (f i) mod m) k l mod m"
-  apply (induct l)
-   apply auto
-  apply (rule trans)
-   apply (rule mod_add_eq)
-  apply simp
-  apply (rule mod_add_right_eq [symmetric])
-  done
-
-lemma funsum_zero [rule_format (no_asm)]:
-    "(\<forall>i. k \<le> i \<and> i \<le> k + l --> f i = 0) --> (funsum f k l) = 0"
-  apply (induct l)
-   apply auto
-  done
-
-lemma funsum_oneelem [rule_format (no_asm)]:
-  "k \<le> j --> j \<le> k + l -->
-    (\<forall>i. k \<le> i \<and> i \<le> k + l \<and> i \<noteq> j --> f i = 0) -->
-    funsum f k l = f j"
-  apply (induct l)
-   prefer 2
-   apply clarify
-   defer
-   apply clarify
-   apply (subgoal_tac "k = j")
-    apply (simp_all (no_asm_simp))
-  apply (case_tac "Suc (k + l) = j")
-   apply (subgoal_tac "funsum f k l = 0")
-    apply (rule_tac [2] funsum_zero)
-    apply (subgoal_tac [3] "f (Suc (k + l)) = 0")
-     apply (subgoal_tac [3] "j \<le> k + l")
-      prefer 4
-      apply arith
-     apply auto
-  done
-
-
-subsection {* Chinese: uniqueness *}
-
-lemma zcong_funprod_aux:
-  "m_cond n mf ==> km_cond n kf mf
-    ==> lincong_sol n kf bf mf x ==> lincong_sol n kf bf mf y
-    ==> [x = y] (mod mf n)"
-  apply (unfold m_cond_def km_cond_def lincong_sol_def)
-  apply (rule iffD1)
-   apply (rule_tac k = "kf n" in zcong_cancel2)
-    apply (rule_tac [3] b = "bf n" in zcong_trans)
-     prefer 4
-     apply (subst zcong_sym)
-     defer
-     apply (rule order_less_imp_le)
-     apply simp_all
-  done
-
-lemma zcong_funprod [rule_format]:
-  "m_cond n mf --> km_cond n kf mf -->
-    lincong_sol n kf bf mf x --> lincong_sol n kf bf mf y -->
-    [x = y] (mod funprod mf 0 n)"
-  apply (induct n)
-   apply (simp_all (no_asm))
-   apply (blast intro: zcong_funprod_aux)
-  apply (rule impI)+
-  apply (rule zcong_zgcd_zmult_zmod)
-    apply (blast intro: zcong_funprod_aux)
-    prefer 2
-    apply (subst zgcd_commute)
-    apply (rule funprod_zgcd)
-   apply (auto simp add: m_cond_def km_cond_def lincong_sol_def)
-  done
-
-
-subsection {* Chinese: existence *}
-
-lemma unique_xi_sol:
-  "0 < n ==> i \<le> n ==> m_cond n mf ==> km_cond n kf mf
-    ==> \<exists>!x. 0 \<le> x \<and> x < mf i \<and> [kf i * mhf mf n i * x = bf i] (mod mf i)"
-  apply (rule zcong_lineq_unique)
-   apply (tactic {* stac (thm "zgcd_zmult_cancel") 2 *})
-    apply (unfold m_cond_def km_cond_def mhf_def)
-    apply (simp_all (no_asm_simp))
-  apply safe
-    apply (tactic {* stac (thm "zgcd_zmult_cancel") 3 *})
-     apply (rule_tac [!] funprod_zgcd)
-     apply safe
-     apply simp_all
-   apply (subgoal_tac "i<n")
-    prefer 2
-    apply arith
-   apply (case_tac [2] i)
-    apply simp_all
-  done
-
-lemma x_sol_lin_aux:
-    "0 < n ==> i \<le> n ==> j \<le> n ==> j \<noteq> i ==> mf j dvd mhf mf n i"
-  apply (unfold mhf_def)
-  apply (case_tac "i = 0")
-   apply (case_tac [2] "i = n")
-    apply (simp_all (no_asm_simp))
-    apply (case_tac [3] "j < i")
-     apply (rule_tac [3] dvd_mult2)
-     apply (rule_tac [4] dvd_mult)
-     apply (rule_tac [!] funprod_zdvd)
-     apply arith
-     apply arith
-     apply arith
-     apply arith
-     apply arith
-     apply arith
-     apply arith
-     apply arith
-  done
-
-lemma x_sol_lin:
-  "0 < n ==> i \<le> n
-    ==> x_sol n kf bf mf mod mf i =
-      xilin_sol i n kf bf mf * mhf mf n i mod mf i"
-  apply (unfold x_sol_def)
-  apply (subst funsum_mod)
-  apply (subst funsum_oneelem)
-     apply auto
-  apply (subst dvd_eq_mod_eq_0 [symmetric])
-  apply (rule dvd_mult)
-  apply (rule x_sol_lin_aux)
-  apply auto
-  done
-
-
-subsection {* Chinese *}
-
-lemma chinese_remainder:
-  "0 < n ==> m_cond n mf ==> km_cond n kf mf
-    ==> \<exists>!x. 0 \<le> x \<and> x < funprod mf 0 n \<and> lincong_sol n kf bf mf x"
-  apply safe
-   apply (rule_tac [2] m = "funprod mf 0 n" in zcong_zless_imp_eq)
-       apply (rule_tac [6] zcong_funprod)
-          apply auto
-  apply (rule_tac x = "x_sol n kf bf mf mod funprod mf 0 n" in exI)
-  apply (unfold lincong_sol_def)
-  apply safe
-    apply (tactic {* stac (thm "zcong_zmod") 3 *})
-    apply (tactic {* stac (thm "mod_mult_eq") 3 *})
-    apply (tactic {* stac (thm "mod_mod_cancel") 3 *})
-      apply (tactic {* stac (thm "x_sol_lin") 4 *})
-        apply (tactic {* stac (thm "mod_mult_eq" RS sym) 6 *})
-        apply (tactic {* stac (thm "zcong_zmod" RS sym) 6 *})
-        apply (subgoal_tac [6]
-          "0 \<le> xilin_sol i n kf bf mf \<and> xilin_sol i n kf bf mf < mf i
-          \<and> [kf i * mhf mf n i * xilin_sol i n kf bf mf = bf i] (mod mf i)")
-         prefer 6
-         apply (simp add: zmult_ac)
-        apply (unfold xilin_sol_def)
-        apply (tactic {* asm_simp_tac @{simpset} 6 *})
-        apply (rule_tac [6] ex1_implies_ex [THEN someI_ex])
-        apply (rule_tac [6] unique_xi_sol)
-           apply (rule_tac [3] funprod_zdvd)
-            apply (unfold m_cond_def)
-            apply (rule funprod_pos [THEN pos_mod_sign])
-            apply (rule_tac [2] funprod_pos [THEN pos_mod_bound])
-            apply auto
-  done
-
-end
--- a/src/HOL/NumberTheory/Euler.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,304 +0,0 @@
-(*  Title:      HOL/Quadratic_Reciprocity/Euler.thy
-    ID:         $Id$
-    Authors:    Jeremy Avigad, David Gray, and Adam Kramer
-*)
-
-header {* Euler's criterion *}
-
-theory Euler imports Residues EvenOdd begin
-
-definition
-  MultInvPair :: "int => int => int => int set" where
-  "MultInvPair a p j = {StandardRes p j, StandardRes p (a * (MultInv p j))}"
-
-definition
-  SetS        :: "int => int => int set set" where
-  "SetS        a p   =  (MultInvPair a p ` SRStar p)"
-
-
-subsection {* Property for MultInvPair *}
-
-lemma MultInvPair_prop1a:
-  "[| zprime p; 2 < p; ~([a = 0](mod p));
-      X \<in> (SetS a p); Y \<in> (SetS a p);
-      ~((X \<inter> Y) = {}) |] ==> X = Y"
-  apply (auto simp add: SetS_def)
-  apply (drule StandardRes_SRStar_prop1a)+ defer 1
-  apply (drule StandardRes_SRStar_prop1a)+
-  apply (auto simp add: MultInvPair_def StandardRes_prop2 zcong_sym)
-  apply (drule notE, rule MultInv_zcong_prop1, auto)[]
-  apply (drule notE, rule MultInv_zcong_prop2, auto simp add: zcong_sym)[]
-  apply (drule MultInv_zcong_prop2, auto simp add: zcong_sym)[]
-  apply (drule MultInv_zcong_prop3, auto simp add: zcong_sym)[]
-  apply (drule MultInv_zcong_prop1, auto)[]
-  apply (drule MultInv_zcong_prop2, auto simp add: zcong_sym)[]
-  apply (drule MultInv_zcong_prop2, auto simp add: zcong_sym)[]
-  apply (drule MultInv_zcong_prop3, auto simp add: zcong_sym)[]
-  done
-
-lemma MultInvPair_prop1b:
-  "[| zprime p; 2 < p; ~([a = 0](mod p));
-      X \<in> (SetS a p); Y \<in> (SetS a p);
-      X \<noteq> Y |] ==> X \<inter> Y = {}"
-  apply (rule notnotD)
-  apply (rule notI)
-  apply (drule MultInvPair_prop1a, auto)
-  done
-
-lemma MultInvPair_prop1c: "[| zprime p; 2 < p; ~([a = 0](mod p)) |] ==>  
-    \<forall>X \<in> SetS a p. \<forall>Y \<in> SetS a p. X \<noteq> Y --> X\<inter>Y = {}"
-  by (auto simp add: MultInvPair_prop1b)
-
-lemma MultInvPair_prop2: "[| zprime p; 2 < p; ~([a = 0](mod p)) |] ==> 
-                          Union ( SetS a p) = SRStar p"
-  apply (auto simp add: SetS_def MultInvPair_def StandardRes_SRStar_prop4 
-    SRStar_mult_prop2)
-  apply (frule StandardRes_SRStar_prop3)
-  apply (rule bexI, auto)
-  done
-
-lemma MultInvPair_distinct: "[| zprime p; 2 < p; ~([a = 0] (mod p)); 
-                                ~([j = 0] (mod p)); 
-                                ~(QuadRes p a) |]  ==> 
-                             ~([j = a * MultInv p j] (mod p))"
-proof
-  assume "zprime p" and "2 < p" and "~([a = 0] (mod p))" and 
-    "~([j = 0] (mod p))" and "~(QuadRes p a)"
-  assume "[j = a * MultInv p j] (mod p)"
-  then have "[j * j = (a * MultInv p j) * j] (mod p)"
-    by (auto simp add: zcong_scalar)
-  then have a:"[j * j = a * (MultInv p j * j)] (mod p)"
-    by (auto simp add: zmult_ac)
-  have "[j * j = a] (mod p)"
-    proof -
-      from prems have b: "[MultInv p j * j = 1] (mod p)"
-        by (simp add: MultInv_prop2a)
-      from b a show ?thesis
-        by (auto simp add: zcong_zmult_prop2)
-    qed
-  then have "[j^2 = a] (mod p)"
-    by (metis  number_of_is_id power2_eq_square succ_bin_simps)
-  with prems show False
-    by (simp add: QuadRes_def)
-qed
-
-lemma MultInvPair_card_two: "[| zprime p; 2 < p; ~([a = 0] (mod p)); 
-                                ~(QuadRes p a); ~([j = 0] (mod p)) |]  ==> 
-                             card (MultInvPair a p j) = 2"
-  apply (auto simp add: MultInvPair_def)
-  apply (subgoal_tac "~ (StandardRes p j = StandardRes p (a * MultInv p j))")
-  apply auto
-  apply (metis MultInvPair_distinct Pls_def StandardRes_def aux number_of_is_id one_is_num_one)
-  done
-
-
-subsection {* Properties of SetS *}
-
-lemma SetS_finite: "2 < p ==> finite (SetS a p)"
-  by (auto simp add: SetS_def SRStar_finite [of p] finite_imageI)
-
-lemma SetS_elems_finite: "\<forall>X \<in> SetS a p. finite X"
-  by (auto simp add: SetS_def MultInvPair_def)
-
-lemma SetS_elems_card: "[| zprime p; 2 < p; ~([a = 0] (mod p)); 
-                        ~(QuadRes p a) |]  ==>
-                        \<forall>X \<in> SetS a p. card X = 2"
-  apply (auto simp add: SetS_def)
-  apply (frule StandardRes_SRStar_prop1a)
-  apply (rule MultInvPair_card_two, auto)
-  done
-
-lemma Union_SetS_finite: "2 < p ==> finite (Union (SetS a p))"
-  by (auto simp add: SetS_finite SetS_elems_finite finite_Union)
-
-lemma card_setsum_aux: "[| finite S; \<forall>X \<in> S. finite (X::int set); 
-    \<forall>X \<in> S. card X = n |] ==> setsum card S = setsum (%x. n) S"
-  by (induct set: finite) auto
-
-lemma SetS_card: "[| zprime p; 2 < p; ~([a = 0] (mod p)); ~(QuadRes p a) |] ==> 
-                  int(card(SetS a p)) = (p - 1) div 2"
-proof -
-  assume "zprime p" and "2 < p" and  "~([a = 0] (mod p))" and "~(QuadRes p a)"
-  then have "(p - 1) = 2 * int(card(SetS a p))"
-  proof -
-    have "p - 1 = int(card(Union (SetS a p)))"
-      by (auto simp add: prems MultInvPair_prop2 SRStar_card)
-    also have "... = int (setsum card (SetS a p))"
-      by (auto simp add: prems SetS_finite SetS_elems_finite
-                         MultInvPair_prop1c [of p a] card_Union_disjoint)
-    also have "... = int(setsum (%x.2) (SetS a p))"
-      using prems
-      by (auto simp add: SetS_elems_card SetS_finite SetS_elems_finite 
-        card_setsum_aux simp del: setsum_constant)
-    also have "... = 2 * int(card( SetS a p))"
-      by (auto simp add: prems SetS_finite setsum_const2)
-    finally show ?thesis .
-  qed
-  from this show ?thesis
-    by auto
-qed
-
-lemma SetS_setprod_prop: "[| zprime p; 2 < p; ~([a = 0] (mod p));
-                              ~(QuadRes p a); x \<in> (SetS a p) |] ==> 
-                          [\<Prod>x = a] (mod p)"
-  apply (auto simp add: SetS_def MultInvPair_def)
-  apply (frule StandardRes_SRStar_prop1a)
-  apply (subgoal_tac "StandardRes p x \<noteq> StandardRes p (a * MultInv p x)")
-  apply (auto simp add: StandardRes_prop2 MultInvPair_distinct)
-  apply (frule_tac m = p and x = x and y = "(a * MultInv p x)" in 
-    StandardRes_prop4)
-  apply (subgoal_tac "[x * (a * MultInv p x) = a * (x * MultInv p x)] (mod p)")
-  apply (drule_tac a = "StandardRes p x * StandardRes p (a * MultInv p x)" and
-                   b = "x * (a * MultInv p x)" and
-                   c = "a * (x * MultInv p x)" in  zcong_trans, force)
-  apply (frule_tac p = p and x = x in MultInv_prop2, auto)
-apply (metis StandardRes_SRStar_prop3 mult_1_right mult_commute zcong_sym zcong_zmult_prop1)
-  apply (auto simp add: zmult_ac)
-  done
-
-lemma aux1: "[| 0 < x; (x::int) < a; x \<noteq> (a - 1) |] ==> x < a - 1"
-  by arith
-
-lemma aux2: "[| (a::int) < c; b < c |] ==> (a \<le> b | b \<le> a)"
-  by auto
-
-lemma SRStar_d22set_prop: "2 < p \<Longrightarrow> (SRStar p) = {1} \<union> (d22set (p - 1))"
-  apply (induct p rule: d22set.induct)
-  apply auto
-  apply (simp add: SRStar_def d22set.simps)
-  apply (simp add: SRStar_def d22set.simps, clarify)
-  apply (frule aux1)
-  apply (frule aux2, auto)
-  apply (simp_all add: SRStar_def)
-  apply (simp add: d22set.simps)
-  apply (frule d22set_le)
-  apply (frule d22set_g_1, auto)
-  done
-
-lemma Union_SetS_setprod_prop1: "[| zprime p; 2 < p; ~([a = 0] (mod p)); ~(QuadRes p a) |] ==>
-                                 [\<Prod>(Union (SetS a p)) = a ^ nat ((p - 1) div 2)] (mod p)"
-proof -
-  assume "zprime p" and "2 < p" and  "~([a = 0] (mod p))" and "~(QuadRes p a)"
-  then have "[\<Prod>(Union (SetS a p)) = 
-      setprod (setprod (%x. x)) (SetS a p)] (mod p)"
-    by (auto simp add: SetS_finite SetS_elems_finite
-                       MultInvPair_prop1c setprod_Union_disjoint)
-  also have "[setprod (setprod (%x. x)) (SetS a p) = 
-      setprod (%x. a) (SetS a p)] (mod p)"
-    by (rule setprod_same_function_zcong)
-      (auto simp add: prems SetS_setprod_prop SetS_finite)
-  also (zcong_trans) have "[setprod (%x. a) (SetS a p) = 
-      a^(card (SetS a p))] (mod p)"
-    by (auto simp add: prems SetS_finite setprod_constant)
-  finally (zcong_trans) show ?thesis
-    apply (rule zcong_trans)
-    apply (subgoal_tac "card(SetS a p) = nat((p - 1) div 2)", auto)
-    apply (subgoal_tac "nat(int(card(SetS a p))) = nat((p - 1) div 2)", force)
-    apply (auto simp add: prems SetS_card)
-    done
-qed
-
-lemma Union_SetS_setprod_prop2: "[| zprime p; 2 < p; ~([a = 0](mod p)) |] ==> 
-                                    \<Prod>(Union (SetS a p)) = zfact (p - 1)"
-proof -
-  assume "zprime p" and "2 < p" and "~([a = 0](mod p))"
-  then have "\<Prod>(Union (SetS a p)) = \<Prod>(SRStar p)"
-    by (auto simp add: MultInvPair_prop2)
-  also have "... = \<Prod>({1} \<union> (d22set (p - 1)))"
-    by (auto simp add: prems SRStar_d22set_prop)
-  also have "... = zfact(p - 1)"
-  proof -
-    have "~(1 \<in> d22set (p - 1)) & finite( d22set (p - 1))"
-      by (metis d22set_fin d22set_g_1 linorder_neq_iff)
-    then have "\<Prod>({1} \<union> (d22set (p - 1))) = \<Prod>(d22set (p - 1))"
-      by auto
-    then show ?thesis
-      by (auto simp add: d22set_prod_zfact)
-  qed
-  finally show ?thesis .
-qed
-
-lemma zfact_prop: "[| zprime p; 2 < p; ~([a = 0] (mod p)); ~(QuadRes p a) |] ==>
-                   [zfact (p - 1) = a ^ nat ((p - 1) div 2)] (mod p)"
-  apply (frule Union_SetS_setprod_prop1) 
-  apply (auto simp add: Union_SetS_setprod_prop2)
-  done
-
-text {* \medskip Prove the first part of Euler's Criterion: *}
-
-lemma Euler_part1: "[| 2 < p; zprime p; ~([x = 0](mod p)); 
-    ~(QuadRes p x) |] ==> 
-      [x^(nat (((p) - 1) div 2)) = -1](mod p)"
-  by (metis Wilson_Russ number_of_is_id zcong_sym zcong_trans zfact_prop)
-
-text {* \medskip Prove another part of Euler Criterion: *}
-
-lemma aux_1: "0 < p ==> (a::int) ^ nat (p) = a * a ^ (nat (p) - 1)"
-proof -
-  assume "0 < p"
-  then have "a ^ (nat p) =  a ^ (1 + (nat p - 1))"
-    by (auto simp add: diff_add_assoc)
-  also have "... = (a ^ 1) * a ^ (nat(p) - 1)"
-    by (simp only: zpower_zadd_distrib)
-  also have "... = a * a ^ (nat(p) - 1)"
-    by auto
-  finally show ?thesis .
-qed
-
-lemma aux_2: "[| (2::int) < p; p \<in> zOdd |] ==> 0 < ((p - 1) div 2)"
-proof -
-  assume "2 < p" and "p \<in> zOdd"
-  then have "(p - 1):zEven"
-    by (auto simp add: zEven_def zOdd_def)
-  then have aux_1: "2 * ((p - 1) div 2) = (p - 1)"
-    by (auto simp add: even_div_2_prop2)
-  with `2 < p` have "1 < (p - 1)"
-    by auto
-  then have " 1 < (2 * ((p - 1) div 2))"
-    by (auto simp add: aux_1)
-  then have "0 < (2 * ((p - 1) div 2)) div 2"
-    by auto
-  then show ?thesis by auto
-qed
-
-lemma Euler_part2:
-    "[| 2 < p; zprime p; [a = 0] (mod p) |] ==> [0 = a ^ nat ((p - 1) div 2)] (mod p)"
-  apply (frule zprime_zOdd_eq_grt_2)
-  apply (frule aux_2, auto)
-  apply (frule_tac a = a in aux_1, auto)
-  apply (frule zcong_zmult_prop1, auto)
-  done
-
-text {* \medskip Prove the final part of Euler's Criterion: *}
-
-lemma aux__1: "[| ~([x = 0] (mod p)); [y ^ 2 = x] (mod p)|] ==> ~(p dvd y)"
-  by (metis dvdI power2_eq_square zcong_sym zcong_trans zcong_zero_equiv_div dvd_trans)
-
-lemma aux__2: "2 * nat((p - 1) div 2) =  nat (2 * ((p - 1) div 2))"
-  by (auto simp add: nat_mult_distrib)
-
-lemma Euler_part3: "[| 2 < p; zprime p; ~([x = 0](mod p)); QuadRes p x |] ==> 
-                      [x^(nat (((p) - 1) div 2)) = 1](mod p)"
-  apply (subgoal_tac "p \<in> zOdd")
-  apply (auto simp add: QuadRes_def)
-   prefer 2 
-   apply (metis number_of_is_id numeral_1_eq_1 zprime_zOdd_eq_grt_2)
-  apply (frule aux__1, auto)
-  apply (drule_tac z = "nat ((p - 1) div 2)" in zcong_zpower)
-  apply (auto simp add: zpower_zpower) 
-  apply (rule zcong_trans)
-  apply (auto simp add: zcong_sym [of "x ^ nat ((p - 1) div 2)"])
-  apply (metis Little_Fermat even_div_2_prop2 mult_Bit0 number_of_is_id odd_minus_one_even one_is_num_one zmult_1 aux__2)
-  done
-
-
-text {* \medskip Finally show Euler's Criterion: *}
-
-theorem Euler_Criterion: "[| 2 < p; zprime p |] ==> [(Legendre a p) =
-    a^(nat (((p) - 1) div 2))] (mod p)"
-  apply (auto simp add: Legendre_def Euler_part2)
-  apply (frule Euler_part3, auto simp add: zcong_sym)[]
-  apply (frule Euler_part1, auto simp add: zcong_sym)[]
-  done
-
-end
--- a/src/HOL/NumberTheory/EulerFermat.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,348 +0,0 @@
-(*  Title:      HOL/NumberTheory/EulerFermat.thy
-    ID:         $Id$
-    Author:     Thomas M. Rasmussen
-    Copyright   2000  University of Cambridge
-*)
-
-header {* Fermat's Little Theorem extended to Euler's Totient function *}
-
-theory EulerFermat
-imports BijectionRel IntFact
-begin
-
-text {*
-  Fermat's Little Theorem extended to Euler's Totient function. More
-  abstract approach than Boyer-Moore (which seems necessary to achieve
-  the extended version).
-*}
-
-
-subsection {* Definitions and lemmas *}
-
-inductive_set
-  RsetR :: "int => int set set"
-  for m :: int
-  where
-    empty [simp]: "{} \<in> RsetR m"
-  | insert: "A \<in> RsetR m ==> zgcd a m = 1 ==>
-      \<forall>a'. a' \<in> A --> \<not> zcong a a' m ==> insert a A \<in> RsetR m"
-
-consts
-  BnorRset :: "int * int => int set"
-
-recdef BnorRset
-  "measure ((\<lambda>(a, m). nat a) :: int * int => nat)"
-  "BnorRset (a, m) =
-   (if 0 < a then
-    let na = BnorRset (a - 1, m)
-    in (if zgcd a m = 1 then insert a na else na)
-    else {})"
-
-definition
-  norRRset :: "int => int set" where
-  "norRRset m = BnorRset (m - 1, m)"
-
-definition
-  noXRRset :: "int => int => int set" where
-  "noXRRset m x = (\<lambda>a. a * x) ` norRRset m"
-
-definition
-  phi :: "int => nat" where
-  "phi m = card (norRRset m)"
-
-definition
-  is_RRset :: "int set => int => bool" where
-  "is_RRset A m = (A \<in> RsetR m \<and> card A = phi m)"
-
-definition
-  RRset2norRR :: "int set => int => int => int" where
-  "RRset2norRR A m a =
-     (if 1 < m \<and> is_RRset A m \<and> a \<in> A then
-        SOME b. zcong a b m \<and> b \<in> norRRset m
-      else 0)"
-
-definition
-  zcongm :: "int => int => int => bool" where
-  "zcongm m = (\<lambda>a b. zcong a b m)"
-
-lemma abs_eq_1_iff [iff]: "(abs z = (1::int)) = (z = 1 \<or> z = -1)"
-  -- {* LCP: not sure why this lemma is needed now *}
-  by (auto simp add: abs_if)
-
-
-text {* \medskip @{text norRRset} *}
-
-declare BnorRset.simps [simp del]
-
-lemma BnorRset_induct:
-  assumes "!!a m. P {} a m"
-    and "!!a m. 0 < (a::int) ==> P (BnorRset (a - 1, m::int)) (a - 1) m
-      ==> P (BnorRset(a,m)) a m"
-  shows "P (BnorRset(u,v)) u v"
-  apply (rule BnorRset.induct)
-  apply safe
-   apply (case_tac [2] "0 < a")
-    apply (rule_tac [2] prems)
-     apply simp_all
-   apply (simp_all add: BnorRset.simps prems)
-  done
-
-lemma Bnor_mem_zle [rule_format]: "b \<in> BnorRset (a, m) \<longrightarrow> b \<le> a"
-  apply (induct a m rule: BnorRset_induct)
-   apply simp
-  apply (subst BnorRset.simps)
-   apply (unfold Let_def, auto)
-  done
-
-lemma Bnor_mem_zle_swap: "a < b ==> b \<notin> BnorRset (a, m)"
-  by (auto dest: Bnor_mem_zle)
-
-lemma Bnor_mem_zg [rule_format]: "b \<in> BnorRset (a, m) --> 0 < b"
-  apply (induct a m rule: BnorRset_induct)
-   prefer 2
-   apply (subst BnorRset.simps)
-   apply (unfold Let_def, auto)
-  done
-
-lemma Bnor_mem_if [rule_format]:
-    "zgcd b m = 1 --> 0 < b --> b \<le> a --> b \<in> BnorRset (a, m)"
-  apply (induct a m rule: BnorRset.induct, auto)
-   apply (subst BnorRset.simps)
-   defer
-   apply (subst BnorRset.simps)
-   apply (unfold Let_def, auto)
-  done
-
-lemma Bnor_in_RsetR [rule_format]: "a < m --> BnorRset (a, m) \<in> RsetR m"
-  apply (induct a m rule: BnorRset_induct, simp)
-  apply (subst BnorRset.simps)
-  apply (unfold Let_def, auto)
-  apply (rule RsetR.insert)
-    apply (rule_tac [3] allI)
-    apply (rule_tac [3] impI)
-    apply (rule_tac [3] zcong_not)
-       apply (subgoal_tac [6] "a' \<le> a - 1")
-        apply (rule_tac [7] Bnor_mem_zle)
-        apply (rule_tac [5] Bnor_mem_zg, auto)
-  done
-
-lemma Bnor_fin: "finite (BnorRset (a, m))"
-  apply (induct a m rule: BnorRset_induct)
-   prefer 2
-   apply (subst BnorRset.simps)
-   apply (unfold Let_def, auto)
-  done
-
-lemma norR_mem_unique_aux: "a \<le> b - 1 ==> a < (b::int)"
-  apply auto
-  done
-
-lemma norR_mem_unique:
-  "1 < m ==>
-    zgcd a m = 1 ==> \<exists>!b. [a = b] (mod m) \<and> b \<in> norRRset m"
-  apply (unfold norRRset_def)
-  apply (cut_tac a = a and m = m in zcong_zless_unique, auto)
-   apply (rule_tac [2] m = m in zcong_zless_imp_eq)
-       apply (auto intro: Bnor_mem_zle Bnor_mem_zg zcong_trans
-	 order_less_imp_le norR_mem_unique_aux simp add: zcong_sym)
-  apply (rule_tac x = b in exI, safe)
-  apply (rule Bnor_mem_if)
-    apply (case_tac [2] "b = 0")
-     apply (auto intro: order_less_le [THEN iffD2])
-   prefer 2
-   apply (simp only: zcong_def)
-   apply (subgoal_tac "zgcd a m = m")
-    prefer 2
-    apply (subst zdvd_iff_zgcd [symmetric])
-     apply (rule_tac [4] zgcd_zcong_zgcd)
-       apply (simp_all add: zcong_sym)
-  done
-
-
-text {* \medskip @{term noXRRset} *}
-
-lemma RRset_gcd [rule_format]:
-    "is_RRset A m ==> a \<in> A --> zgcd a m = 1"
-  apply (unfold is_RRset_def)
-  apply (rule RsetR.induct [where P="%A. a \<in> A --> zgcd a m = 1"], auto)
-  done
-
-lemma RsetR_zmult_mono:
-  "A \<in> RsetR m ==>
-    0 < m ==> zgcd x m = 1 ==> (\<lambda>a. a * x) ` A \<in> RsetR m"
-  apply (erule RsetR.induct, simp_all)
-  apply (rule RsetR.insert, auto)
-   apply (blast intro: zgcd_zgcd_zmult)
-  apply (simp add: zcong_cancel)
-  done
-
-lemma card_nor_eq_noX:
-  "0 < m ==>
-    zgcd x m = 1 ==> card (noXRRset m x) = card (norRRset m)"
-  apply (unfold norRRset_def noXRRset_def)
-  apply (rule card_image)
-   apply (auto simp add: inj_on_def Bnor_fin)
-  apply (simp add: BnorRset.simps)
-  done
-
-lemma noX_is_RRset:
-    "0 < m ==> zgcd x m = 1 ==> is_RRset (noXRRset m x) m"
-  apply (unfold is_RRset_def phi_def)
-  apply (auto simp add: card_nor_eq_noX)
-  apply (unfold noXRRset_def norRRset_def)
-  apply (rule RsetR_zmult_mono)
-    apply (rule Bnor_in_RsetR, simp_all)
-  done
-
-lemma aux_some:
-  "1 < m ==> is_RRset A m ==> a \<in> A
-    ==> zcong a (SOME b. [a = b] (mod m) \<and> b \<in> norRRset m) m \<and>
-      (SOME b. [a = b] (mod m) \<and> b \<in> norRRset m) \<in> norRRset m"
-  apply (rule norR_mem_unique [THEN ex1_implies_ex, THEN someI_ex])
-   apply (rule_tac [2] RRset_gcd, simp_all)
-  done
-
-lemma RRset2norRR_correct:
-  "1 < m ==> is_RRset A m ==> a \<in> A ==>
-    [a = RRset2norRR A m a] (mod m) \<and> RRset2norRR A m a \<in> norRRset m"
-  apply (unfold RRset2norRR_def, simp)
-  apply (rule aux_some, simp_all)
-  done
-
-lemmas RRset2norRR_correct1 =
-  RRset2norRR_correct [THEN conjunct1, standard]
-lemmas RRset2norRR_correct2 =
-  RRset2norRR_correct [THEN conjunct2, standard]
-
-lemma RsetR_fin: "A \<in> RsetR m ==> finite A"
-  by (induct set: RsetR) auto
-
-lemma RRset_zcong_eq [rule_format]:
-  "1 < m ==>
-    is_RRset A m ==> [a = b] (mod m) ==> a \<in> A --> b \<in> A --> a = b"
-  apply (unfold is_RRset_def)
-  apply (rule RsetR.induct [where P="%A. a \<in> A --> b \<in> A --> a = b"])
-    apply (auto simp add: zcong_sym)
-  done
-
-lemma aux:
-  "P (SOME a. P a) ==> Q (SOME a. Q a) ==>
-    (SOME a. P a) = (SOME a. Q a) ==> \<exists>a. P a \<and> Q a"
-  apply auto
-  done
-
-lemma RRset2norRR_inj:
-    "1 < m ==> is_RRset A m ==> inj_on (RRset2norRR A m) A"
-  apply (unfold RRset2norRR_def inj_on_def, auto)
-  apply (subgoal_tac "\<exists>b. ([x = b] (mod m) \<and> b \<in> norRRset m) \<and>
-      ([y = b] (mod m) \<and> b \<in> norRRset m)")
-   apply (rule_tac [2] aux)
-     apply (rule_tac [3] aux_some)
-       apply (rule_tac [2] aux_some)
-         apply (rule RRset_zcong_eq, auto)
-  apply (rule_tac b = b in zcong_trans)
-   apply (simp_all add: zcong_sym)
-  done
-
-lemma RRset2norRR_eq_norR:
-    "1 < m ==> is_RRset A m ==> RRset2norRR A m ` A = norRRset m"
-  apply (rule card_seteq)
-    prefer 3
-    apply (subst card_image)
-      apply (rule_tac RRset2norRR_inj, auto)
-     apply (rule_tac [3] RRset2norRR_correct2, auto)
-    apply (unfold is_RRset_def phi_def norRRset_def)
-    apply (auto simp add: Bnor_fin)
-  done
-
-
-lemma Bnor_prod_power_aux: "a \<notin> A ==> inj f ==> f a \<notin> f ` A"
-by (unfold inj_on_def, auto)
-
-lemma Bnor_prod_power [rule_format]:
-  "x \<noteq> 0 ==> a < m --> \<Prod>((\<lambda>a. a * x) ` BnorRset (a, m)) =
-      \<Prod>(BnorRset(a, m)) * x^card (BnorRset (a, m))"
-  apply (induct a m rule: BnorRset_induct)
-   prefer 2
-   apply (simplesubst BnorRset.simps)  --{*multiple redexes*}
-   apply (unfold Let_def, auto)
-  apply (simp add: Bnor_fin Bnor_mem_zle_swap)
-  apply (subst setprod_insert)
-    apply (rule_tac [2] Bnor_prod_power_aux)
-     apply (unfold inj_on_def)
-     apply (simp_all add: zmult_ac Bnor_fin finite_imageI
-       Bnor_mem_zle_swap)
-  done
-
-
-subsection {* Fermat *}
-
-lemma bijzcong_zcong_prod:
-    "(A, B) \<in> bijR (zcongm m) ==> [\<Prod>A = \<Prod>B] (mod m)"
-  apply (unfold zcongm_def)
-  apply (erule bijR.induct)
-   apply (subgoal_tac [2] "a \<notin> A \<and> b \<notin> B \<and> finite A \<and> finite B")
-    apply (auto intro: fin_bijRl fin_bijRr zcong_zmult)
-  done
-
-lemma Bnor_prod_zgcd [rule_format]:
-    "a < m --> zgcd (\<Prod>(BnorRset(a, m))) m = 1"
-  apply (induct a m rule: BnorRset_induct)
-   prefer 2
-   apply (subst BnorRset.simps)
-   apply (unfold Let_def, auto)
-  apply (simp add: Bnor_fin Bnor_mem_zle_swap)
-  apply (blast intro: zgcd_zgcd_zmult)
-  done
-
-theorem Euler_Fermat:
-    "0 < m ==> zgcd x m = 1 ==> [x^(phi m) = 1] (mod m)"
-  apply (unfold norRRset_def phi_def)
-  apply (case_tac "x = 0")
-   apply (case_tac [2] "m = 1")
-    apply (rule_tac [3] iffD1)
-     apply (rule_tac [3] k = "\<Prod>(BnorRset(m - 1, m))"
-       in zcong_cancel2)
-      prefer 5
-      apply (subst Bnor_prod_power [symmetric])
-        apply (rule_tac [7] Bnor_prod_zgcd, simp_all)
-  apply (rule bijzcong_zcong_prod)
-  apply (fold norRRset_def noXRRset_def)
-  apply (subst RRset2norRR_eq_norR [symmetric])
-    apply (rule_tac [3] inj_func_bijR, auto)
-     apply (unfold zcongm_def)
-     apply (rule_tac [2] RRset2norRR_correct1)
-       apply (rule_tac [5] RRset2norRR_inj)
-        apply (auto intro: order_less_le [THEN iffD2]
-	   simp add: noX_is_RRset)
-  apply (unfold noXRRset_def norRRset_def)
-  apply (rule finite_imageI)
-  apply (rule Bnor_fin)
-  done
-
-lemma Bnor_prime:
-  "\<lbrakk> zprime p; a < p \<rbrakk> \<Longrightarrow> card (BnorRset (a, p)) = nat a"
-  apply (induct a p rule: BnorRset.induct)
-  apply (subst BnorRset.simps)
-  apply (unfold Let_def, auto simp add:zless_zprime_imp_zrelprime)
-  apply (subgoal_tac "finite (BnorRset (a - 1,m))")
-   apply (subgoal_tac "a ~: BnorRset (a - 1,m)")
-    apply (auto simp add: card_insert_disjoint Suc_nat_eq_nat_zadd1)
-   apply (frule Bnor_mem_zle, arith)
-  apply (frule Bnor_fin)
-  done
-
-lemma phi_prime: "zprime p ==> phi p = nat (p - 1)"
-  apply (unfold phi_def norRRset_def)
-  apply (rule Bnor_prime, auto)
-  done
-
-theorem Little_Fermat:
-    "zprime p ==> \<not> p dvd x ==> [x^(nat (p - 1)) = 1] (mod p)"
-  apply (subst phi_prime [symmetric])
-   apply (rule_tac [2] Euler_Fermat)
-    apply (erule_tac [3] zprime_imp_zrelprime)
-    apply (unfold zprime_def, auto)
-  done
-
-end
--- a/src/HOL/NumberTheory/EvenOdd.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,256 +0,0 @@
-(*  Title:      HOL/Quadratic_Reciprocity/EvenOdd.thy
-    Authors:    Jeremy Avigad, David Gray, and Adam Kramer
-*)
-
-header {*Parity: Even and Odd Integers*}
-
-theory EvenOdd
-imports Int2
-begin
-
-definition
-  zOdd    :: "int set" where
-  "zOdd = {x. \<exists>k. x = 2 * k + 1}"
-
-definition
-  zEven   :: "int set" where
-  "zEven = {x. \<exists>k. x = 2 * k}"
-
-subsection {* Some useful properties about even and odd *}
-
-lemma zOddI [intro?]: "x = 2 * k + 1 \<Longrightarrow> x \<in> zOdd"
-  and zOddE [elim?]: "x \<in> zOdd \<Longrightarrow> (!!k. x = 2 * k + 1 \<Longrightarrow> C) \<Longrightarrow> C"
-  by (auto simp add: zOdd_def)
-
-lemma zEvenI [intro?]: "x = 2 * k \<Longrightarrow> x \<in> zEven"
-  and zEvenE [elim?]: "x \<in> zEven \<Longrightarrow> (!!k. x = 2 * k \<Longrightarrow> C) \<Longrightarrow> C"
-  by (auto simp add: zEven_def)
-
-lemma one_not_even: "~(1 \<in> zEven)"
-proof
-  assume "1 \<in> zEven"
-  then obtain k :: int where "1 = 2 * k" ..
-  then show False by arith
-qed
-
-lemma even_odd_conj: "~(x \<in> zOdd & x \<in> zEven)"
-proof -
-  {
-    fix a b
-    assume "2 * (a::int) = 2 * (b::int) + 1"
-    then have "2 * (a::int) - 2 * (b :: int) = 1"
-      by arith
-    then have "2 * (a - b) = 1"
-      by (auto simp add: zdiff_zmult_distrib)
-    moreover have "(2 * (a - b)):zEven"
-      by (auto simp only: zEven_def)
-    ultimately have False
-      by (auto simp add: one_not_even)
-  }
-  then show ?thesis
-    by (auto simp add: zOdd_def zEven_def)
-qed
-
-lemma even_odd_disj: "(x \<in> zOdd | x \<in> zEven)"
-  by (simp add: zOdd_def zEven_def) arith
-
-lemma not_odd_impl_even: "~(x \<in> zOdd) ==> x \<in> zEven"
-  using even_odd_disj by auto
-
-lemma odd_mult_odd_prop: "(x*y):zOdd ==> x \<in> zOdd"
-proof (rule classical)
-  assume "\<not> ?thesis"
-  then have "x \<in> zEven" by (rule not_odd_impl_even)
-  then obtain a where a: "x = 2 * a" ..
-  assume "x * y : zOdd"
-  then obtain b where "x * y = 2 * b + 1" ..
-  with a have "2 * a * y = 2 * b + 1" by simp
-  then have "2 * a * y - 2 * b = 1"
-    by arith
-  then have "2 * (a * y - b) = 1"
-    by (auto simp add: zdiff_zmult_distrib)
-  moreover have "(2 * (a * y - b)):zEven"
-    by (auto simp only: zEven_def)
-  ultimately have False
-    by (auto simp add: one_not_even)
-  then show ?thesis ..
-qed
-
-lemma odd_minus_one_even: "x \<in> zOdd ==> (x - 1):zEven"
-  by (auto simp add: zOdd_def zEven_def)
-
-lemma even_div_2_prop1: "x \<in> zEven ==> (x mod 2) = 0"
-  by (auto simp add: zEven_def)
-
-lemma even_div_2_prop2: "x \<in> zEven ==> (2 * (x div 2)) = x"
-  by (auto simp add: zEven_def)
-
-lemma even_plus_even: "[| x \<in> zEven; y \<in> zEven |] ==> x + y \<in> zEven"
-  apply (auto simp add: zEven_def)
-  apply (auto simp only: zadd_zmult_distrib2 [symmetric])
-  done
-
-lemma even_times_either: "x \<in> zEven ==> x * y \<in> zEven"
-  by (auto simp add: zEven_def)
-
-lemma even_minus_even: "[| x \<in> zEven; y \<in> zEven |] ==> x - y \<in> zEven"
-  apply (auto simp add: zEven_def)
-  apply (auto simp only: zdiff_zmult_distrib2 [symmetric])
-  done
-
-lemma odd_minus_odd: "[| x \<in> zOdd; y \<in> zOdd |] ==> x - y \<in> zEven"
-  apply (auto simp add: zOdd_def zEven_def)
-  apply (auto simp only: zdiff_zmult_distrib2 [symmetric])
-  done
-
-lemma even_minus_odd: "[| x \<in> zEven; y \<in> zOdd |] ==> x - y \<in> zOdd"
-  apply (auto simp add: zOdd_def zEven_def)
-  apply (rule_tac x = "k - ka - 1" in exI)
-  apply auto
-  done
-
-lemma odd_minus_even: "[| x \<in> zOdd; y \<in> zEven |] ==> x - y \<in> zOdd"
-  apply (auto simp add: zOdd_def zEven_def)
-  apply (auto simp only: zdiff_zmult_distrib2 [symmetric])
-  done
-
-lemma odd_times_odd: "[| x \<in> zOdd;  y \<in> zOdd |] ==> x * y \<in> zOdd"
-  apply (auto simp add: zOdd_def zadd_zmult_distrib zadd_zmult_distrib2)
-  apply (rule_tac x = "2 * ka * k + ka + k" in exI)
-  apply (auto simp add: zadd_zmult_distrib)
-  done
-
-lemma odd_iff_not_even: "(x \<in> zOdd) = (~ (x \<in> zEven))"
-  using even_odd_conj even_odd_disj by auto
-
-lemma even_product: "x * y \<in> zEven ==> x \<in> zEven | y \<in> zEven"
-  using odd_iff_not_even odd_times_odd by auto
-
-lemma even_diff: "x - y \<in> zEven = ((x \<in> zEven) = (y \<in> zEven))"
-proof
-  assume xy: "x - y \<in> zEven"
-  {
-    assume x: "x \<in> zEven"
-    have "y \<in> zEven"
-    proof (rule classical)
-      assume "\<not> ?thesis"
-      then have "y \<in> zOdd"
-        by (simp add: odd_iff_not_even)
-      with x have "x - y \<in> zOdd"
-        by (simp add: even_minus_odd)
-      with xy have False
-        by (auto simp add: odd_iff_not_even)
-      then show ?thesis ..
-    qed
-  } moreover {
-    assume y: "y \<in> zEven"
-    have "x \<in> zEven"
-    proof (rule classical)
-      assume "\<not> ?thesis"
-      then have "x \<in> zOdd"
-        by (auto simp add: odd_iff_not_even)
-      with y have "x - y \<in> zOdd"
-        by (simp add: odd_minus_even)
-      with xy have False
-        by (auto simp add: odd_iff_not_even)
-      then show ?thesis ..
-    qed
-  }
-  ultimately show "(x \<in> zEven) = (y \<in> zEven)"
-    by (auto simp add: odd_iff_not_even even_minus_even odd_minus_odd
-      even_minus_odd odd_minus_even)
-next
-  assume "(x \<in> zEven) = (y \<in> zEven)"
-  then show "x - y \<in> zEven"
-    by (auto simp add: odd_iff_not_even even_minus_even odd_minus_odd
-      even_minus_odd odd_minus_even)
-qed
-
-lemma neg_one_even_power: "[| x \<in> zEven; 0 \<le> x |] ==> (-1::int)^(nat x) = 1"
-proof -
-  assume "x \<in> zEven" and "0 \<le> x"
-  from `x \<in> zEven` obtain a where "x = 2 * a" ..
-  with `0 \<le> x` have "0 \<le> a" by simp
-  from `0 \<le> x` and `x = 2 * a` have "nat x = nat (2 * a)"
-    by simp
-  also from `x = 2 * a` have "nat (2 * a) = 2 * nat a"
-    by (simp add: nat_mult_distrib)
-  finally have "(-1::int)^nat x = (-1)^(2 * nat a)"
-    by simp
-  also have "... = ((-1::int)^2)^ (nat a)"
-    by (simp add: zpower_zpower [symmetric])
-  also have "(-1::int)^2 = 1"
-    by simp
-  finally show ?thesis
-    by simp
-qed
-
-lemma neg_one_odd_power: "[| x \<in> zOdd; 0 \<le> x |] ==> (-1::int)^(nat x) = -1"
-proof -
-  assume "x \<in> zOdd" and "0 \<le> x"
-  from `x \<in> zOdd` obtain a where "x = 2 * a + 1" ..
-  with `0 \<le> x` have a: "0 \<le> a" by simp
-  with `0 \<le> x` and `x = 2 * a + 1` have "nat x = nat (2 * a + 1)"
-    by simp
-  also from a have "nat (2 * a + 1) = 2 * nat a + 1"
-    by (auto simp add: nat_mult_distrib nat_add_distrib)
-  finally have "(-1::int)^nat x = (-1)^(2 * nat a + 1)"
-    by simp
-  also have "... = ((-1::int)^2)^ (nat a) * (-1)^1"
-    by (auto simp add: zpower_zpower [symmetric] zpower_zadd_distrib)
-  also have "(-1::int)^2 = 1"
-    by simp
-  finally show ?thesis
-    by simp
-qed
-
-lemma neg_one_power_parity: "[| 0 \<le> x; 0 \<le> y; (x \<in> zEven) = (y \<in> zEven) |] ==>
-    (-1::int)^(nat x) = (-1::int)^(nat y)"
-  using even_odd_disj [of x] even_odd_disj [of y]
-  by (auto simp add: neg_one_even_power neg_one_odd_power)
-
-
-lemma one_not_neg_one_mod_m: "2 < m ==> ~([1 = -1] (mod m))"
-  by (auto simp add: zcong_def zdvd_not_zless)
-
-lemma even_div_2_l: "[| y \<in> zEven; x < y |] ==> x div 2 < y div 2"
-proof -
-  assume "y \<in> zEven" and "x < y"
-  from `y \<in> zEven` obtain k where k: "y = 2 * k" ..
-  with `x < y` have "x < 2 * k" by simp
-  then have "x div 2 < k" by (auto simp add: div_prop1)
-  also have "k = (2 * k) div 2" by simp
-  finally have "x div 2 < 2 * k div 2" by simp
-  with k show ?thesis by simp
-qed
-
-lemma even_sum_div_2: "[| x \<in> zEven; y \<in> zEven |] ==> (x + y) div 2 = x div 2 + y div 2"
-  by (auto simp add: zEven_def)
-
-lemma even_prod_div_2: "[| x \<in> zEven |] ==> (x * y) div 2 = (x div 2) * y"
-  by (auto simp add: zEven_def)
-
-(* An odd prime is greater than 2 *)
-
-lemma zprime_zOdd_eq_grt_2: "zprime p ==> (p \<in> zOdd) = (2 < p)"
-  apply (auto simp add: zOdd_def zprime_def)
-  apply (drule_tac x = 2 in allE)
-  using odd_iff_not_even [of p]
-  apply (auto simp add: zOdd_def zEven_def)
-  done
-
-(* Powers of -1 and parity *)
-
-lemma neg_one_special: "finite A ==>
-    ((-1 :: int) ^ card A) * (-1 ^ card A) = 1"
-  by (induct set: finite) auto
-
-lemma neg_one_power: "(-1::int)^n = 1 | (-1::int)^n = -1"
-  by (induct n) auto
-
-lemma neg_one_power_eq_mod_m: "[| 2 < m; [(-1::int)^j = (-1::int)^k] (mod m) |]
-    ==> ((-1::int)^j = (-1::int)^k)"
-  using neg_one_power [of j] and ListMem.insert neg_one_power [of k]
-  by (auto simp add: one_not_neg_one_mod_m zcong_sym)
-
-end
--- a/src/HOL/NumberTheory/Factorization.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,341 +0,0 @@
-(*  Title:      HOL/NumberTheory/Factorization.thy
-    ID:         $Id$
-    Author:     Thomas Marthedal Rasmussen
-    Copyright   2000  University of Cambridge
-*)
-
-header {* Fundamental Theorem of Arithmetic (unique factorization into primes) *}
-
-theory Factorization
-imports Main Primes Permutation
-begin
-
-
-subsection {* Definitions *}
-
-definition
-  primel :: "nat list => bool" where
-  "primel xs = (\<forall>p \<in> set xs. prime p)"
-
-consts
-  nondec :: "nat list => bool "
-  prod :: "nat list => nat"
-  oinsert :: "nat => nat list => nat list"
-  sort :: "nat list => nat list"
-
-primrec
-  "nondec [] = True"
-  "nondec (x # xs) = (case xs of [] => True | y # ys => x \<le> y \<and> nondec xs)"
-
-primrec
-  "prod [] = Suc 0"
-  "prod (x # xs) = x * prod xs"
-
-primrec
-  "oinsert x [] = [x]"
-  "oinsert x (y # ys) = (if x \<le> y then x # y # ys else y # oinsert x ys)"
-
-primrec
-  "sort [] = []"
-  "sort (x # xs) = oinsert x (sort xs)"
-
-
-subsection {* Arithmetic *}
-
-lemma one_less_m: "(m::nat) \<noteq> m * k ==> m \<noteq> Suc 0 ==> Suc 0 < m"
-  apply (cases m)
-   apply auto
-  done
-
-lemma one_less_k: "(m::nat) \<noteq> m * k ==> Suc 0 < m * k ==> Suc 0 < k"
-  apply (cases k)
-   apply auto
-  done
-
-lemma mult_left_cancel: "(0::nat) < k ==> k * n = k * m ==> n = m"
-  apply auto
-  done
-
-lemma mn_eq_m_one: "(0::nat) < m ==> m * n = m ==> n = Suc 0"
-  apply (cases n)
-   apply auto
-  done
-
-lemma prod_mn_less_k:
-    "(0::nat) < n ==> 0 < k ==> Suc 0 < m ==> m * n = k ==> n < k"
-  apply (induct m)
-   apply auto
-  done
-
-
-subsection {* Prime list and product *}
-
-lemma prod_append: "prod (xs @ ys) = prod xs * prod ys"
-  apply (induct xs)
-   apply (simp_all add: mult_assoc)
-  done
-
-lemma prod_xy_prod:
-    "prod (x # xs) = prod (y # ys) ==> x * prod xs = y * prod ys"
-  apply auto
-  done
-
-lemma primel_append: "primel (xs @ ys) = (primel xs \<and> primel ys)"
-  apply (unfold primel_def)
-  apply auto
-  done
-
-lemma prime_primel: "prime n ==> primel [n] \<and> prod [n] = n"
-  apply (unfold primel_def)
-  apply auto
-  done
-
-lemma prime_nd_one: "prime p ==> \<not> p dvd Suc 0"
-  apply (unfold prime_def dvd_def)
-  apply auto
-  done
-
-lemma hd_dvd_prod: "prod (x # xs) = prod ys ==> x dvd (prod ys)" 
-  by (metis dvd_mult_left dvd_refl prod.simps(2))
-
-lemma primel_tl: "primel (x # xs) ==> primel xs"
-  apply (unfold primel_def)
-  apply auto
-  done
-
-lemma primel_hd_tl: "(primel (x # xs)) = (prime x \<and> primel xs)"
-  apply (unfold primel_def)
-  apply auto
-  done
-
-lemma primes_eq: "prime p ==> prime q ==> p dvd q ==> p = q"
-  apply (unfold prime_def)
-  apply auto
-  done
-
-lemma primel_one_empty: "primel xs ==> prod xs = Suc 0 ==> xs = []"
-  apply (cases xs)
-   apply (simp_all add: primel_def prime_def)
-  done
-
-lemma prime_g_one: "prime p ==> Suc 0 < p"
-  apply (unfold prime_def)
-  apply auto
-  done
-
-lemma prime_g_zero: "prime p ==> 0 < p"
-  apply (unfold prime_def)
-  apply auto
-  done
-
-lemma primel_nempty_g_one:
-    "primel xs \<Longrightarrow> xs \<noteq> [] \<Longrightarrow> Suc 0 < prod xs"
-  apply (induct xs)
-   apply simp
-  apply (fastsimp simp: primel_def prime_def elim: one_less_mult)
-  done
-
-lemma primel_prod_gz: "primel xs ==> 0 < prod xs"
-  apply (induct xs)
-   apply (auto simp: primel_def prime_def)
-  done
-
-
-subsection {* Sorting *}
-
-lemma nondec_oinsert: "nondec xs \<Longrightarrow> nondec (oinsert x xs)"
-  apply (induct xs)
-   apply simp
-   apply (case_tac xs)
-    apply (simp_all cong del: list.weak_case_cong)
-  done
-
-lemma nondec_sort: "nondec (sort xs)"
-  apply (induct xs)
-   apply simp_all
-  apply (erule nondec_oinsert)
-  done
-
-lemma x_less_y_oinsert: "x \<le> y ==> l = y # ys ==> x # l = oinsert x l"
-  apply simp_all
-  done
-
-lemma nondec_sort_eq [rule_format]: "nondec xs \<longrightarrow> xs = sort xs"
-  apply (induct xs)
-   apply safe
-    apply simp_all
-   apply (case_tac xs)
-    apply simp_all
-  apply (case_tac xs)
-   apply simp
-  apply (rule_tac y = aa and ys = list in x_less_y_oinsert)
-   apply simp_all
-  done
-
-lemma oinsert_x_y: "oinsert x (oinsert y l) = oinsert y (oinsert x l)"
-  apply (induct l)
-  apply auto
-  done
-
-
-subsection {* Permutation *}
-
-lemma perm_primel [rule_format]: "xs <~~> ys ==> primel xs --> primel ys"
-  apply (unfold primel_def)
-  apply (induct set: perm)
-     apply simp
-    apply simp
-   apply (simp (no_asm))
-   apply blast
-  apply blast
-  done
-
-lemma perm_prod: "xs <~~> ys ==> prod xs = prod ys"
-  apply (induct set: perm)
-     apply (simp_all add: mult_ac)
-  done
-
-lemma perm_subst_oinsert: "xs <~~> ys ==> oinsert a xs <~~> oinsert a ys"
-  apply (induct set: perm)
-     apply auto
-  done
-
-lemma perm_oinsert: "x # xs <~~> oinsert x xs"
-  apply (induct xs)
-   apply auto
-  done
-
-lemma perm_sort: "xs <~~> sort xs"
-  apply (induct xs)
-  apply (auto intro: perm_oinsert elim: perm_subst_oinsert)
-  done
-
-lemma perm_sort_eq: "xs <~~> ys ==> sort xs = sort ys"
-  apply (induct set: perm)
-     apply (simp_all add: oinsert_x_y)
-  done
-
-
-subsection {* Existence *}
-
-lemma ex_nondec_lemma:
-    "primel xs ==> \<exists>ys. primel ys \<and> nondec ys \<and> prod ys = prod xs"
-  apply (blast intro: nondec_sort perm_prod perm_primel perm_sort perm_sym)
-  done
-
-lemma not_prime_ex_mk:
-  "Suc 0 < n \<and> \<not> prime n ==>
-    \<exists>m k. Suc 0 < m \<and> Suc 0 < k \<and> m < n \<and> k < n \<and> n = m * k"
-  apply (unfold prime_def dvd_def)
-  apply (auto intro: n_less_m_mult_n n_less_n_mult_m one_less_m one_less_k)
-  done
-
-lemma split_primel:
-  "primel xs \<Longrightarrow> primel ys \<Longrightarrow> \<exists>l. primel l \<and> prod l = prod xs * prod ys"
-  apply (rule exI)
-  apply safe
-   apply (rule_tac [2] prod_append)
-  apply (simp add: primel_append)
-  done
-
-lemma factor_exists [rule_format]: "Suc 0 < n --> (\<exists>l. primel l \<and> prod l = n)"
-  apply (induct n rule: nat_less_induct)
-  apply (rule impI)
-  apply (case_tac "prime n")
-   apply (rule exI)
-   apply (erule prime_primel)
-  apply (cut_tac n = n in not_prime_ex_mk)
-   apply (auto intro!: split_primel)
-  done
-
-lemma nondec_factor_exists: "Suc 0 < n ==> \<exists>l. primel l \<and> nondec l \<and> prod l = n"
-  apply (erule factor_exists [THEN exE])
-  apply (blast intro!: ex_nondec_lemma)
-  done
-
-
-subsection {* Uniqueness *}
-
-lemma prime_dvd_mult_list [rule_format]:
-    "prime p ==> p dvd (prod xs) --> (\<exists>m. m:set xs \<and> p dvd m)"
-  apply (induct xs)
-   apply (force simp add: prime_def)
-   apply (force dest: prime_dvd_mult)
-  done
-
-lemma hd_xs_dvd_prod:
-  "primel (x # xs) ==> primel ys ==> prod (x # xs) = prod ys
-    ==> \<exists>m. m \<in> set ys \<and> x dvd m"
-  apply (rule prime_dvd_mult_list)
-   apply (simp add: primel_hd_tl)
-  apply (erule hd_dvd_prod)
-  done
-
-lemma prime_dvd_eq: "primel (x # xs) ==> primel ys ==> m \<in> set ys ==> x dvd m ==> x = m"
-  apply (rule primes_eq)
-    apply (auto simp add: primel_def primel_hd_tl)
-  done
-
-lemma hd_xs_eq_prod:
-  "primel (x # xs) ==>
-    primel ys ==> prod (x # xs) = prod ys ==> x \<in> set ys"
-  apply (frule hd_xs_dvd_prod)
-    apply auto
-  apply (drule prime_dvd_eq)
-     apply auto
-  done
-
-lemma perm_primel_ex:
-  "primel (x # xs) ==>
-    primel ys ==> prod (x # xs) = prod ys ==> \<exists>l. ys <~~> (x # l)"
-  apply (rule exI)
-  apply (rule perm_remove)
-  apply (erule hd_xs_eq_prod)
-   apply simp_all
-  done
-
-lemma primel_prod_less:
-  "primel (x # xs) ==>
-    primel ys ==> prod (x # xs) = prod ys ==> prod xs < prod ys"
-  by (metis less_asym linorder_neqE_nat mult_less_cancel2 nat_0_less_mult_iff
-    nat_less_le nat_mult_1 prime_def primel_hd_tl primel_prod_gz prod.simps(2))
-
-lemma prod_one_empty:
-    "primel xs ==> p * prod xs = p ==> prime p ==> xs = []"
-  apply (auto intro: primel_one_empty simp add: prime_def)
-  done
-
-lemma uniq_ex_aux:
-  "\<forall>m. m < prod ys --> (\<forall>xs ys. primel xs \<and> primel ys \<and>
-      prod xs = prod ys \<and> prod xs = m --> xs <~~> ys) ==>
-    primel list ==> primel x ==> prod list = prod x ==> prod x < prod ys
-    ==> x <~~> list"
-  apply simp
-  done
-
-lemma factor_unique [rule_format]:
-  "\<forall>xs ys. primel xs \<and> primel ys \<and> prod xs = prod ys \<and> prod xs = n
-    --> xs <~~> ys"
-  apply (induct n rule: nat_less_induct)
-  apply safe
-  apply (case_tac xs)
-   apply (force intro: primel_one_empty)
-  apply (rule perm_primel_ex [THEN exE])
-     apply simp_all
-  apply (rule perm.trans [THEN perm_sym])
-  apply assumption
-  apply (rule perm.Cons)
-  apply (case_tac "x = []")
-   apply (metis perm_prod perm_refl prime_primel primel_hd_tl primel_tl prod_one_empty)
-  apply (metis nat_0_less_mult_iff nat_mult_eq_cancel1 perm_primel perm_prod primel_prod_gz primel_prod_less primel_tl prod.simps(2))
-  done
-
-lemma perm_nondec_unique:
-    "xs <~~> ys ==> nondec xs ==> nondec ys ==> xs = ys"
-  by (metis nondec_sort_eq perm_sort_eq)
-
-theorem unique_prime_factorization [rule_format]:
-    "\<forall>n. Suc 0 < n --> (\<exists>!l. primel l \<and> nondec l \<and> prod l = n)"
-  by (metis factor_unique nondec_factor_exists perm_nondec_unique)
-
-end
--- a/src/HOL/NumberTheory/Fib.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,150 +0,0 @@
-(*  ID:         $Id$
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1997  University of Cambridge
-*)
-
-header {* The Fibonacci function *}
-
-theory Fib
-imports Primes
-begin
-
-text {*
-  Fibonacci numbers: proofs of laws taken from:
-  R. L. Graham, D. E. Knuth, O. Patashnik.  Concrete Mathematics.
-  (Addison-Wesley, 1989)
-
-  \bigskip
-*}
-
-fun fib :: "nat \<Rightarrow> nat"
-where
-         "fib 0 = 0"
-|        "fib (Suc 0) = 1"
-| fib_2: "fib (Suc (Suc n)) = fib n + fib (Suc n)"
-
-text {*
-  \medskip The difficulty in these proofs is to ensure that the
-  induction hypotheses are applied before the definition of @{term
-  fib}.  Towards this end, the @{term fib} equations are not declared
-  to the Simplifier and are applied very selectively at first.
-*}
-
-text{*We disable @{text fib.fib_2fib_2} for simplification ...*}
-declare fib_2 [simp del]
-
-text{*...then prove a version that has a more restrictive pattern.*}
-lemma fib_Suc3: "fib (Suc (Suc (Suc n))) = fib (Suc n) + fib (Suc (Suc n))"
-  by (rule fib_2)
-
-text {* \medskip Concrete Mathematics, page 280 *}
-
-lemma fib_add: "fib (Suc (n + k)) = fib (Suc k) * fib (Suc n) + fib k * fib n"
-proof (induct n rule: fib.induct)
-  case 1 show ?case by simp
-next
-  case 2 show ?case  by (simp add: fib_2)
-next
-  case 3 thus ?case by (simp add: fib_2 add_mult_distrib2)
-qed
-
-lemma fib_Suc_neq_0: "fib (Suc n) \<noteq> 0"
-  apply (induct n rule: fib.induct)
-    apply (simp_all add: fib_2)
-  done
-
-lemma fib_Suc_gr_0: "0 < fib (Suc n)"
-  by (insert fib_Suc_neq_0 [of n], simp)  
-
-lemma fib_gr_0: "0 < n ==> 0 < fib n"
-  by (case_tac n, auto simp add: fib_Suc_gr_0) 
-
-
-text {*
-  \medskip Concrete Mathematics, page 278: Cassini's identity.  The proof is
-  much easier using integers, not natural numbers!
-*}
-
-lemma fib_Cassini_int:
- "int (fib (Suc (Suc n)) * fib n) =
-  (if n mod 2 = 0 then int (fib (Suc n) * fib (Suc n)) - 1
-   else int (fib (Suc n) * fib (Suc n)) + 1)"
-proof(induct n rule: fib.induct)
-  case 1 thus ?case by (simp add: fib_2)
-next
-  case 2 thus ?case by (simp add: fib_2 mod_Suc)
-next 
-  case (3 x) 
-  have "Suc 0 \<noteq> x mod 2 \<longrightarrow> x mod 2 = 0" by presburger
-  with "3.hyps" show ?case by (simp add: fib.simps add_mult_distrib add_mult_distrib2)
-qed
-
-text{*We now obtain a version for the natural numbers via the coercion 
-   function @{term int}.*}
-theorem fib_Cassini:
- "fib (Suc (Suc n)) * fib n =
-  (if n mod 2 = 0 then fib (Suc n) * fib (Suc n) - 1
-   else fib (Suc n) * fib (Suc n) + 1)"
-  apply (rule int_int_eq [THEN iffD1]) 
-  apply (simp add: fib_Cassini_int)
-  apply (subst zdiff_int [symmetric]) 
-   apply (insert fib_Suc_gr_0 [of n], simp_all)
-  done
-
-
-text {* \medskip Toward Law 6.111 of Concrete Mathematics *}
-
-lemma gcd_fib_Suc_eq_1: "gcd (fib n) (fib (Suc n)) = Suc 0"
-  apply (induct n rule: fib.induct)
-    prefer 3
-    apply (simp add: gcd_commute fib_Suc3)
-   apply (simp_all add: fib_2)
-  done
-
-lemma gcd_fib_add: "gcd (fib m) (fib (n + m)) = gcd (fib m) (fib n)"
-  apply (simp add: gcd_commute [of "fib m"])
-  apply (case_tac m)
-   apply simp 
-  apply (simp add: fib_add)
-  apply (simp add: add_commute gcd_non_0 [OF fib_Suc_gr_0])
-  apply (simp add: gcd_non_0 [OF fib_Suc_gr_0, symmetric])
-  apply (simp add: gcd_fib_Suc_eq_1 gcd_mult_cancel)
-  done
-
-lemma gcd_fib_diff: "m \<le> n ==> gcd (fib m) (fib (n - m)) = gcd (fib m) (fib n)"
-  by (simp add: gcd_fib_add [symmetric, of _ "n-m"]) 
-
-lemma gcd_fib_mod: "0 < m ==> gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
-proof (induct n rule: less_induct)
-  case (less n)
-  from less.prems have pos_m: "0 < m" .
-  show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
-  proof (cases "m < n")
-    case True note m_n = True
-    then have m_n': "m \<le> n" by auto
-    with pos_m have pos_n: "0 < n" by auto
-    with pos_m m_n have diff: "n - m < n" by auto
-    have "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib ((n - m) mod m))"
-    by (simp add: mod_if [of n]) (insert m_n, auto)
-    also have "\<dots> = gcd (fib m) (fib (n - m))" by (simp add: less.hyps diff pos_m)
-    also have "\<dots> = gcd (fib m) (fib n)" by (simp add: gcd_fib_diff m_n')
-    finally show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)" .
-  next
-    case False then show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
-    by (cases "m = n") auto
-  qed
-qed
-
-lemma fib_gcd: "fib (gcd m n) = gcd (fib m) (fib n)"  -- {* Law 6.111 *}
-  apply (induct m n rule: gcd_induct)
-  apply (simp_all add: gcd_non_0 gcd_commute gcd_fib_mod)
-  done
-
-theorem fib_mult_eq_setsum:
-    "fib (Suc n) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
-  apply (induct n rule: fib.induct)
-    apply (auto simp add: atMost_Suc fib_2)
-  apply (simp add: add_mult_distrib add_mult_distrib2)
-  done
-
-end
--- a/src/HOL/NumberTheory/Finite2.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,223 +0,0 @@
-(*  Title:      HOL/Quadratic_Reciprocity/Finite2.thy
-    ID:         $Id$
-    Authors:    Jeremy Avigad, David Gray, and Adam Kramer
-*)
-
-header {*Finite Sets and Finite Sums*}
-
-theory Finite2
-imports Main IntFact Infinite_Set
-begin
-
-text{*
-  These are useful for combinatorial and number-theoretic counting
-  arguments.
-*}
-
-
-subsection {* Useful properties of sums and products *}
-
-lemma setsum_same_function_zcong:
-  assumes a: "\<forall>x \<in> S. [f x = g x](mod m)"
-  shows "[setsum f S = setsum g S] (mod m)"
-proof cases
-  assume "finite S"
-  thus ?thesis using a by induct (simp_all add: zcong_zadd)
-next
-  assume "infinite S" thus ?thesis by(simp add:setsum_def)
-qed
-
-lemma setprod_same_function_zcong:
-  assumes a: "\<forall>x \<in> S. [f x = g x](mod m)"
-  shows "[setprod f S = setprod g S] (mod m)"
-proof cases
-  assume "finite S"
-  thus ?thesis using a by induct (simp_all add: zcong_zmult)
-next
-  assume "infinite S" thus ?thesis by(simp add:setprod_def)
-qed
-
-lemma setsum_const: "finite X ==> setsum (%x. (c :: int)) X = c * int(card X)"
-  apply (induct set: finite)
-  apply (auto simp add: left_distrib right_distrib int_eq_of_nat)
-  done
-
-lemma setsum_const2: "finite X ==> int (setsum (%x. (c :: nat)) X) =
-    int(c) * int(card X)"
-  apply (induct set: finite)
-  apply (auto simp add: zadd_zmult_distrib2)
-  done
-
-lemma setsum_const_mult: "finite A ==> setsum (%x. c * ((f x)::int)) A =
-    c * setsum f A"
-  by (induct set: finite) (auto simp add: zadd_zmult_distrib2)
-
-
-subsection {* Cardinality of explicit finite sets *}
-
-lemma finite_surjI: "[| B \<subseteq> f ` A; finite A |] ==> finite B"
-  by (simp add: finite_subset finite_imageI)
-
-lemma bdd_nat_set_l_finite: "finite {y::nat . y < x}"
-  by (rule bounded_nat_set_is_finite) blast
-
-lemma bdd_nat_set_le_finite: "finite {y::nat . y \<le> x}"
-proof -
-  have "{y::nat . y \<le> x} = {y::nat . y < Suc x}" by auto
-  then show ?thesis by (auto simp add: bdd_nat_set_l_finite)
-qed
-
-lemma  bdd_int_set_l_finite: "finite {x::int. 0 \<le> x & x < n}"
-  apply (subgoal_tac " {(x :: int). 0 \<le> x & x < n} \<subseteq>
-      int ` {(x :: nat). x < nat n}")
-   apply (erule finite_surjI)
-   apply (auto simp add: bdd_nat_set_l_finite image_def)
-  apply (rule_tac x = "nat x" in exI, simp)
-  done
-
-lemma bdd_int_set_le_finite: "finite {x::int. 0 \<le> x & x \<le> n}"
-  apply (subgoal_tac "{x. 0 \<le> x & x \<le> n} = {x. 0 \<le> x & x < n + 1}")
-   apply (erule ssubst)
-   apply (rule bdd_int_set_l_finite)
-  apply auto
-  done
-
-lemma bdd_int_set_l_l_finite: "finite {x::int. 0 < x & x < n}"
-proof -
-  have "{x::int. 0 < x & x < n} \<subseteq> {x::int. 0 \<le> x & x < n}"
-    by auto
-  then show ?thesis by (auto simp add: bdd_int_set_l_finite finite_subset)
-qed
-
-lemma bdd_int_set_l_le_finite: "finite {x::int. 0 < x & x \<le> n}"
-proof -
-  have "{x::int. 0 < x & x \<le> n} \<subseteq> {x::int. 0 \<le> x & x \<le> n}"
-    by auto
-  then show ?thesis by (auto simp add: bdd_int_set_le_finite finite_subset)
-qed
-
-lemma card_bdd_nat_set_l: "card {y::nat . y < x} = x"
-proof (induct x)
-  case 0
-  show "card {y::nat . y < 0} = 0" by simp
-next
-  case (Suc n)
-  have "{y. y < Suc n} = insert n {y. y < n}"
-    by auto
-  then have "card {y. y < Suc n} = card (insert n {y. y < n})"
-    by auto
-  also have "... = Suc (card {y. y < n})"
-    by (rule card_insert_disjoint) (auto simp add: bdd_nat_set_l_finite)
-  finally show "card {y. y < Suc n} = Suc n"
-    using `card {y. y < n} = n` by simp
-qed
-
-lemma card_bdd_nat_set_le: "card { y::nat. y \<le> x} = Suc x"
-proof -
-  have "{y::nat. y \<le> x} = { y::nat. y < Suc x}"
-    by auto
-  then show ?thesis by (auto simp add: card_bdd_nat_set_l)
-qed
-
-lemma card_bdd_int_set_l: "0 \<le> (n::int) ==> card {y. 0 \<le> y & y < n} = nat n"
-proof -
-  assume "0 \<le> n"
-  have "inj_on (%y. int y) {y. y < nat n}"
-    by (auto simp add: inj_on_def)
-  hence "card (int ` {y. y < nat n}) = card {y. y < nat n}"
-    by (rule card_image)
-  also from `0 \<le> n` have "int ` {y. y < nat n} = {y. 0 \<le> y & y < n}"
-    apply (auto simp add: zless_nat_eq_int_zless image_def)
-    apply (rule_tac x = "nat x" in exI)
-    apply (auto simp add: nat_0_le)
-    done
-  also have "card {y. y < nat n} = nat n"
-    by (rule card_bdd_nat_set_l)
-  finally show "card {y. 0 \<le> y & y < n} = nat n" .
-qed
-
-lemma card_bdd_int_set_le: "0 \<le> (n::int) ==> card {y. 0 \<le> y & y \<le> n} =
-  nat n + 1"
-proof -
-  assume "0 \<le> n"
-  moreover have "{y. 0 \<le> y & y \<le> n} = {y. 0 \<le> y & y < n+1}" by auto
-  ultimately show ?thesis
-    using card_bdd_int_set_l [of "n + 1"]
-    by (auto simp add: nat_add_distrib)
-qed
-
-lemma card_bdd_int_set_l_le: "0 \<le> (n::int) ==>
-    card {x. 0 < x & x \<le> n} = nat n"
-proof -
-  assume "0 \<le> n"
-  have "inj_on (%x. x+1) {x. 0 \<le> x & x < n}"
-    by (auto simp add: inj_on_def)
-  hence "card ((%x. x+1) ` {x. 0 \<le> x & x < n}) =
-     card {x. 0 \<le> x & x < n}"
-    by (rule card_image)
-  also from `0 \<le> n` have "... = nat n"
-    by (rule card_bdd_int_set_l)
-  also have "(%x. x + 1) ` {x. 0 \<le> x & x < n} = {x. 0 < x & x<= n}"
-    apply (auto simp add: image_def)
-    apply (rule_tac x = "x - 1" in exI)
-    apply arith
-    done
-  finally show "card {x. 0 < x & x \<le> n} = nat n" .
-qed
-
-lemma card_bdd_int_set_l_l: "0 < (n::int) ==>
-  card {x. 0 < x & x < n} = nat n - 1"
-proof -
-  assume "0 < n"
-  moreover have "{x. 0 < x & x < n} = {x. 0 < x & x \<le> n - 1}"
-    by simp
-  ultimately show ?thesis
-    using insert card_bdd_int_set_l_le [of "n - 1"]
-    by (auto simp add: nat_diff_distrib)
-qed
-
-lemma int_card_bdd_int_set_l_l: "0 < n ==>
-    int(card {x. 0 < x & x < n}) = n - 1"
-  apply (auto simp add: card_bdd_int_set_l_l)
-  done
-
-lemma int_card_bdd_int_set_l_le: "0 \<le> n ==>
-    int(card {x. 0 < x & x \<le> n}) = n"
-  by (auto simp add: card_bdd_int_set_l_le)
-
-
-subsection {* Cardinality of finite cartesian products *}
-
-(* FIXME could be useful in general but not needed here
-lemma insert_Sigma [simp]: "(insert x A) <*> B = ({ x } <*> B) \<union> (A <*> B)"
-  by blast
- *)
-
-text {* Lemmas for counting arguments. *}
-
-lemma setsum_bij_eq: "[| finite A; finite B; f ` A \<subseteq> B; inj_on f A;
-    g ` B \<subseteq> A; inj_on g B |] ==> setsum g B = setsum (g \<circ> f) A"
-  apply (frule_tac h = g and f = f in setsum_reindex)
-  apply (subgoal_tac "setsum g B = setsum g (f ` A)")
-   apply (simp add: inj_on_def)
-  apply (subgoal_tac "card A = card B")
-   apply (drule_tac A = "f ` A" and B = B in card_seteq)
-     apply (auto simp add: card_image)
-  apply (frule_tac A = A and B = B and f = f in card_inj_on_le, auto)
-  apply (frule_tac A = B and B = A and f = g in card_inj_on_le)
-    apply auto
-  done
-
-lemma setprod_bij_eq: "[| finite A; finite B; f ` A \<subseteq> B; inj_on f A;
-    g ` B \<subseteq> A; inj_on g B |] ==> setprod g B = setprod (g \<circ> f) A"
-  apply (frule_tac h = g and f = f in setprod_reindex)
-  apply (subgoal_tac "setprod g B = setprod g (f ` A)")
-   apply (simp add: inj_on_def)
-  apply (subgoal_tac "card A = card B")
-   apply (drule_tac A = "f ` A" and B = B in card_seteq)
-     apply (auto simp add: card_image)
-  apply (frule_tac A = A and B = B and f = f in card_inj_on_le, auto)
-  apply (frule_tac A = B and B = A and f = g in card_inj_on_le, auto)
-  done
-
-end
--- a/src/HOL/NumberTheory/Gauss.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,535 +0,0 @@
-(*  Title:      HOL/Quadratic_Reciprocity/Gauss.thy
-    ID:         $Id$
-    Authors:    Jeremy Avigad, David Gray, and Adam Kramer)
-*)
-
-header {* Gauss' Lemma *}
-
-theory Gauss
-imports Euler
-begin
-
-locale GAUSS =
-  fixes p :: "int"
-  fixes a :: "int"
-
-  assumes p_prime: "zprime p"
-  assumes p_g_2: "2 < p"
-  assumes p_a_relprime: "~[a = 0](mod p)"
-  assumes a_nonzero:    "0 < a"
-begin
-
-definition
-  A :: "int set" where
-  "A = {(x::int). 0 < x & x \<le> ((p - 1) div 2)}"
-
-definition
-  B :: "int set" where
-  "B = (%x. x * a) ` A"
-
-definition
-  C :: "int set" where
-  "C = StandardRes p ` B"
-
-definition
-  D :: "int set" where
-  "D = C \<inter> {x. x \<le> ((p - 1) div 2)}"
-
-definition
-  E :: "int set" where
-  "E = C \<inter> {x. ((p - 1) div 2) < x}"
-
-definition
-  F :: "int set" where
-  "F = (%x. (p - x)) ` E"
-
-
-subsection {* Basic properties of p *}
-
-lemma p_odd: "p \<in> zOdd"
-  by (auto simp add: p_prime p_g_2 zprime_zOdd_eq_grt_2)
-
-lemma p_g_0: "0 < p"
-  using p_g_2 by auto
-
-lemma int_nat: "int (nat ((p - 1) div 2)) = (p - 1) div 2"
-  using ListMem.insert p_g_2 by (auto simp add: pos_imp_zdiv_nonneg_iff)
-
-lemma p_minus_one_l: "(p - 1) div 2 < p"
-proof -
-  have "(p - 1) div 2 \<le> (p - 1) div 1"
-    by (rule zdiv_mono2) (auto simp add: p_g_0)
-  also have "\<dots> = p - 1" by simp
-  finally show ?thesis by simp
-qed
-
-lemma p_eq: "p = (2 * (p - 1) div 2) + 1"
-  using div_mult_self1_is_id [of 2 "p - 1"] by auto
-
-
-lemma (in -) zodd_imp_zdiv_eq: "x \<in> zOdd ==> 2 * (x - 1) div 2 = 2 * ((x - 1) div 2)"
-  apply (frule odd_minus_one_even)
-  apply (simp add: zEven_def)
-  apply (subgoal_tac "2 \<noteq> 0")
-  apply (frule_tac b = "2 :: int" and a = "x - 1" in div_mult_self1_is_id)
-  apply (auto simp add: even_div_2_prop2)
-  done
-
-
-lemma p_eq2: "p = (2 * ((p - 1) div 2)) + 1"
-  apply (insert p_eq p_prime p_g_2 zprime_zOdd_eq_grt_2 [of p], auto)
-  apply (frule zodd_imp_zdiv_eq, auto)
-  done
-
-
-subsection {* Basic Properties of the Gauss Sets *}
-
-lemma finite_A: "finite (A)"
-  apply (auto simp add: A_def)
-  apply (subgoal_tac "{x. 0 < x & x \<le> (p - 1) div 2} \<subseteq> {x. 0 \<le> x & x < 1 + (p - 1) div 2}")
-  apply (auto simp add: bdd_int_set_l_finite finite_subset)
-  done
-
-lemma finite_B: "finite (B)"
-  by (auto simp add: B_def finite_A finite_imageI)
-
-lemma finite_C: "finite (C)"
-  by (auto simp add: C_def finite_B finite_imageI)
-
-lemma finite_D: "finite (D)"
-  by (auto simp add: D_def finite_Int finite_C)
-
-lemma finite_E: "finite (E)"
-  by (auto simp add: E_def finite_Int finite_C)
-
-lemma finite_F: "finite (F)"
-  by (auto simp add: F_def finite_E finite_imageI)
-
-lemma C_eq: "C = D \<union> E"
-  by (auto simp add: C_def D_def E_def)
-
-lemma A_card_eq: "card A = nat ((p - 1) div 2)"
-  apply (auto simp add: A_def)
-  apply (insert int_nat)
-  apply (erule subst)
-  apply (auto simp add: card_bdd_int_set_l_le)
-  done
-
-lemma inj_on_xa_A: "inj_on (%x. x * a) A"
-  using a_nonzero by (simp add: A_def inj_on_def)
-
-lemma A_res: "ResSet p A"
-  apply (auto simp add: A_def ResSet_def)
-  apply (rule_tac m = p in zcong_less_eq)
-  apply (insert p_g_2, auto)
-  done
-
-lemma B_res: "ResSet p B"
-  apply (insert p_g_2 p_a_relprime p_minus_one_l)
-  apply (auto simp add: B_def)
-  apply (rule ResSet_image)
-  apply (auto simp add: A_res)
-  apply (auto simp add: A_def)
-proof -
-  fix x fix y
-  assume a: "[x * a = y * a] (mod p)"
-  assume b: "0 < x"
-  assume c: "x \<le> (p - 1) div 2"
-  assume d: "0 < y"
-  assume e: "y \<le> (p - 1) div 2"
-  from a p_a_relprime p_prime a_nonzero zcong_cancel [of p a x y]
-  have "[x = y](mod p)"
-    by (simp add: zprime_imp_zrelprime zcong_def p_g_0 order_le_less)
-  with zcong_less_eq [of x y p] p_minus_one_l
-      order_le_less_trans [of x "(p - 1) div 2" p]
-      order_le_less_trans [of y "(p - 1) div 2" p] show "x = y"
-    by (simp add: prems p_minus_one_l p_g_0)
-qed
-
-lemma SR_B_inj: "inj_on (StandardRes p) B"
-  apply (auto simp add: B_def StandardRes_def inj_on_def A_def prems)
-proof -
-  fix x fix y
-  assume a: "x * a mod p = y * a mod p"
-  assume b: "0 < x"
-  assume c: "x \<le> (p - 1) div 2"
-  assume d: "0 < y"
-  assume e: "y \<le> (p - 1) div 2"
-  assume f: "x \<noteq> y"
-  from a have "[x * a = y * a](mod p)"
-    by (simp add: zcong_zmod_eq p_g_0)
-  with p_a_relprime p_prime a_nonzero zcong_cancel [of p a x y]
-  have "[x = y](mod p)"
-    by (simp add: zprime_imp_zrelprime zcong_def p_g_0 order_le_less)
-  with zcong_less_eq [of x y p] p_minus_one_l
-    order_le_less_trans [of x "(p - 1) div 2" p]
-    order_le_less_trans [of y "(p - 1) div 2" p] have "x = y"
-    by (simp add: prems p_minus_one_l p_g_0)
-  then have False
-    by (simp add: f)
-  then show "a = 0"
-    by simp
-qed
-
-lemma inj_on_pminusx_E: "inj_on (%x. p - x) E"
-  apply (auto simp add: E_def C_def B_def A_def)
-  apply (rule_tac g = "%x. -1 * (x - p)" in inj_on_inverseI)
-  apply auto
-  done
-
-lemma A_ncong_p: "x \<in> A ==> ~[x = 0](mod p)"
-  apply (auto simp add: A_def)
-  apply (frule_tac m = p in zcong_not_zero)
-  apply (insert p_minus_one_l)
-  apply auto
-  done
-
-lemma A_greater_zero: "x \<in> A ==> 0 < x"
-  by (auto simp add: A_def)
-
-lemma B_ncong_p: "x \<in> B ==> ~[x = 0](mod p)"
-  apply (auto simp add: B_def)
-  apply (frule A_ncong_p)
-  apply (insert p_a_relprime p_prime a_nonzero)
-  apply (frule_tac a = x and b = a in zcong_zprime_prod_zero_contra)
-  apply (auto simp add: A_greater_zero)
-  done
-
-lemma B_greater_zero: "x \<in> B ==> 0 < x"
-  using a_nonzero by (auto simp add: B_def mult_pos_pos A_greater_zero)
-
-lemma C_ncong_p: "x \<in> C ==>  ~[x = 0](mod p)"
-  apply (auto simp add: C_def)
-  apply (frule B_ncong_p)
-  apply (subgoal_tac "[x = StandardRes p x](mod p)")
-  defer apply (simp add: StandardRes_prop1)
-  apply (frule_tac a = x and b = "StandardRes p x" and c = 0 in zcong_trans)
-  apply auto
-  done
-
-lemma C_greater_zero: "y \<in> C ==> 0 < y"
-  apply (auto simp add: C_def)
-proof -
-  fix x
-  assume a: "x \<in> B"
-  from p_g_0 have "0 \<le> StandardRes p x"
-    by (simp add: StandardRes_lbound)
-  moreover have "~[x = 0] (mod p)"
-    by (simp add: a B_ncong_p)
-  then have "StandardRes p x \<noteq> 0"
-    by (simp add: StandardRes_prop3)
-  ultimately show "0 < StandardRes p x"
-    by (simp add: order_le_less)
-qed
-
-lemma D_ncong_p: "x \<in> D ==> ~[x = 0](mod p)"
-  by (auto simp add: D_def C_ncong_p)
-
-lemma E_ncong_p: "x \<in> E ==> ~[x = 0](mod p)"
-  by (auto simp add: E_def C_ncong_p)
-
-lemma F_ncong_p: "x \<in> F ==> ~[x = 0](mod p)"
-  apply (auto simp add: F_def)
-proof -
-  fix x assume a: "x \<in> E" assume b: "[p - x = 0] (mod p)"
-  from E_ncong_p have "~[x = 0] (mod p)"
-    by (simp add: a)
-  moreover from a have "0 < x"
-    by (simp add: a E_def C_greater_zero)
-  moreover from a have "x < p"
-    by (auto simp add: E_def C_def p_g_0 StandardRes_ubound)
-  ultimately have "~[p - x = 0] (mod p)"
-    by (simp add: zcong_not_zero)
-  from this show False by (simp add: b)
-qed
-
-lemma F_subset: "F \<subseteq> {x. 0 < x & x \<le> ((p - 1) div 2)}"
-  apply (auto simp add: F_def E_def)
-  apply (insert p_g_0)
-  apply (frule_tac x = xa in StandardRes_ubound)
-  apply (frule_tac x = x in StandardRes_ubound)
-  apply (subgoal_tac "xa = StandardRes p xa")
-  apply (auto simp add: C_def StandardRes_prop2 StandardRes_prop1)
-proof -
-  from zodd_imp_zdiv_eq p_prime p_g_2 zprime_zOdd_eq_grt_2 have
-    "2 * (p - 1) div 2 = 2 * ((p - 1) div 2)"
-    by simp
-  with p_eq2 show " !!x. [| (p - 1) div 2 < StandardRes p x; x \<in> B |]
-      ==> p - StandardRes p x \<le> (p - 1) div 2"
-    by simp
-qed
-
-lemma D_subset: "D \<subseteq> {x. 0 < x & x \<le> ((p - 1) div 2)}"
-  by (auto simp add: D_def C_greater_zero)
-
-lemma F_eq: "F = {x. \<exists>y \<in> A. ( x = p - (StandardRes p (y*a)) & (p - 1) div 2 < StandardRes p (y*a))}"
-  by (auto simp add: F_def E_def D_def C_def B_def A_def)
-
-lemma D_eq: "D = {x. \<exists>y \<in> A. ( x = StandardRes p (y*a) & StandardRes p (y*a) \<le> (p - 1) div 2)}"
-  by (auto simp add: D_def C_def B_def A_def)
-
-lemma D_leq: "x \<in> D ==> x \<le> (p - 1) div 2"
-  by (auto simp add: D_eq)
-
-lemma F_ge: "x \<in> F ==> x \<le> (p - 1) div 2"
-  apply (auto simp add: F_eq A_def)
-proof -
-  fix y
-  assume "(p - 1) div 2 < StandardRes p (y * a)"
-  then have "p - StandardRes p (y * a) < p - ((p - 1) div 2)"
-    by arith
-  also from p_eq2 have "... = 2 * ((p - 1) div 2) + 1 - ((p - 1) div 2)"
-    by auto
-  also have "2 * ((p - 1) div 2) + 1 - (p - 1) div 2 = (p - 1) div 2 + 1"
-    by arith
-  finally show "p - StandardRes p (y * a) \<le> (p - 1) div 2"
-    using zless_add1_eq [of "p - StandardRes p (y * a)" "(p - 1) div 2"] by auto
-qed
-
-lemma all_A_relprime: "\<forall>x \<in> A. zgcd x p = 1"
-  using p_prime p_minus_one_l by (auto simp add: A_def zless_zprime_imp_zrelprime)
-
-lemma A_prod_relprime: "zgcd (setprod id A) p = 1"
-by(rule all_relprime_prod_relprime[OF finite_A all_A_relprime])
-
-
-subsection {* Relationships Between Gauss Sets *}
-
-lemma B_card_eq_A: "card B = card A"
-  using finite_A by (simp add: finite_A B_def inj_on_xa_A card_image)
-
-lemma B_card_eq: "card B = nat ((p - 1) div 2)"
-  by (simp add: B_card_eq_A A_card_eq)
-
-lemma F_card_eq_E: "card F = card E"
-  using finite_E by (simp add: F_def inj_on_pminusx_E card_image)
-
-lemma C_card_eq_B: "card C = card B"
-  apply (insert finite_B)
-  apply (subgoal_tac "inj_on (StandardRes p) B")
-  apply (simp add: B_def C_def card_image)
-  apply (rule StandardRes_inj_on_ResSet)
-  apply (simp add: B_res)
-  done
-
-lemma D_E_disj: "D \<inter> E = {}"
-  by (auto simp add: D_def E_def)
-
-lemma C_card_eq_D_plus_E: "card C = card D + card E"
-  by (auto simp add: C_eq card_Un_disjoint D_E_disj finite_D finite_E)
-
-lemma C_prod_eq_D_times_E: "setprod id E * setprod id D = setprod id C"
-  apply (insert D_E_disj finite_D finite_E C_eq)
-  apply (frule setprod_Un_disjoint [of D E id])
-  apply auto
-  done
-
-lemma C_B_zcong_prod: "[setprod id C = setprod id B] (mod p)"
-  apply (auto simp add: C_def)
-  apply (insert finite_B SR_B_inj)
-  apply (frule_tac f = "StandardRes p" in setprod_reindex_id [symmetric], auto)
-  apply (rule setprod_same_function_zcong)
-  apply (auto simp add: StandardRes_prop1 zcong_sym p_g_0)
-  done
-
-lemma F_Un_D_subset: "(F \<union> D) \<subseteq> A"
-  apply (rule Un_least)
-  apply (auto simp add: A_def F_subset D_subset)
-  done
-
-lemma F_D_disj: "(F \<inter> D) = {}"
-  apply (simp add: F_eq D_eq)
-  apply (auto simp add: F_eq D_eq)
-proof -
-  fix y fix ya
-  assume "p - StandardRes p (y * a) = StandardRes p (ya * a)"
-  then have "p = StandardRes p (y * a) + StandardRes p (ya * a)"
-    by arith
-  moreover have "p dvd p"
-    by auto
-  ultimately have "p dvd (StandardRes p (y * a) + StandardRes p (ya * a))"
-    by auto
-  then have a: "[StandardRes p (y * a) + StandardRes p (ya * a) = 0] (mod p)"
-    by (auto simp add: zcong_def)
-  have "[y * a = StandardRes p (y * a)] (mod p)"
-    by (simp only: zcong_sym StandardRes_prop1)
-  moreover have "[ya * a = StandardRes p (ya * a)] (mod p)"
-    by (simp only: zcong_sym StandardRes_prop1)
-  ultimately have "[y * a + ya * a =
-    StandardRes p (y * a) + StandardRes p (ya * a)] (mod p)"
-    by (rule zcong_zadd)
-  with a have "[y * a + ya * a = 0] (mod p)"
-    apply (elim zcong_trans)
-    by (simp only: zcong_refl)
-  also have "y * a + ya * a = a * (y + ya)"
-    by (simp add: zadd_zmult_distrib2 zmult_commute)
-  finally have "[a * (y + ya) = 0] (mod p)" .
-  with p_prime a_nonzero zcong_zprime_prod_zero [of p a "y + ya"]
-    p_a_relprime
-  have a: "[y + ya = 0] (mod p)"
-    by auto
-  assume b: "y \<in> A" and c: "ya: A"
-  with A_def have "0 < y + ya"
-    by auto
-  moreover from b c A_def have "y + ya \<le> (p - 1) div 2 + (p - 1) div 2"
-    by auto
-  moreover from b c p_eq2 A_def have "y + ya < p"
-    by auto
-  ultimately show False
-    apply simp
-    apply (frule_tac m = p in zcong_not_zero)
-    apply (auto simp add: a)
-    done
-qed
-
-lemma F_Un_D_card: "card (F \<union> D) = nat ((p - 1) div 2)"
-proof -
-  have "card (F \<union> D) = card E + card D"
-    by (auto simp add: finite_F finite_D F_D_disj
-      card_Un_disjoint F_card_eq_E)
-  then have "card (F \<union> D) = card C"
-    by (simp add: C_card_eq_D_plus_E)
-  from this show "card (F \<union> D) = nat ((p - 1) div 2)"
-    by (simp add: C_card_eq_B B_card_eq)
-qed
-
-lemma F_Un_D_eq_A: "F \<union> D = A"
-  using finite_A F_Un_D_subset A_card_eq F_Un_D_card by (auto simp add: card_seteq)
-
-lemma prod_D_F_eq_prod_A:
-    "(setprod id D) * (setprod id F) = setprod id A"
-  apply (insert F_D_disj finite_D finite_F)
-  apply (frule setprod_Un_disjoint [of F D id])
-  apply (auto simp add: F_Un_D_eq_A)
-  done
-
-lemma prod_F_zcong:
-  "[setprod id F = ((-1) ^ (card E)) * (setprod id E)] (mod p)"
-proof -
-  have "setprod id F = setprod id (op - p ` E)"
-    by (auto simp add: F_def)
-  then have "setprod id F = setprod (op - p) E"
-    apply simp
-    apply (insert finite_E inj_on_pminusx_E)
-    apply (frule_tac f = "op - p" in setprod_reindex_id, auto)
-    done
-  then have one:
-    "[setprod id F = setprod (StandardRes p o (op - p)) E] (mod p)"
-    apply simp
-    apply (insert p_g_0 finite_E StandardRes_prod)
-    by (auto)
-  moreover have a: "\<forall>x \<in> E. [p - x = 0 - x] (mod p)"
-    apply clarify
-    apply (insert zcong_id [of p])
-    apply (rule_tac a = p and m = p and c = x and d = x in zcong_zdiff, auto)
-    done
-  moreover have b: "\<forall>x \<in> E. [StandardRes p (p - x) = p - x](mod p)"
-    apply clarify
-    apply (simp add: StandardRes_prop1 zcong_sym)
-    done
-  moreover have "\<forall>x \<in> E. [StandardRes p (p - x) = - x](mod p)"
-    apply clarify
-    apply (insert a b)
-    apply (rule_tac b = "p - x" in zcong_trans, auto)
-    done
-  ultimately have c:
-    "[setprod (StandardRes p o (op - p)) E = setprod (uminus) E](mod p)"
-    apply simp
-    using finite_E p_g_0
-      setprod_same_function_zcong [of E "StandardRes p o (op - p)" uminus p]
-    by auto
-  then have two: "[setprod id F = setprod (uminus) E](mod p)"
-    apply (insert one c)
-    apply (rule zcong_trans [of "setprod id F"
-                               "setprod (StandardRes p o op - p) E" p
-                               "setprod uminus E"], auto)
-    done
-  also have "setprod uminus E = (setprod id E) * (-1)^(card E)"
-    using finite_E by (induct set: finite) auto
-  then have "setprod uminus E = (-1) ^ (card E) * (setprod id E)"
-    by (simp add: zmult_commute)
-  with two show ?thesis
-    by simp
-qed
-
-
-subsection {* Gauss' Lemma *}
-
-lemma aux: "setprod id A * -1 ^ card E * a ^ card A * -1 ^ card E = setprod id A * a ^ card A"
-  by (auto simp add: finite_E neg_one_special)
-
-theorem pre_gauss_lemma:
-  "[a ^ nat((p - 1) div 2) = (-1) ^ (card E)] (mod p)"
-proof -
-  have "[setprod id A = setprod id F * setprod id D](mod p)"
-    by (auto simp add: prod_D_F_eq_prod_A zmult_commute cong del:setprod_cong)
-  then have "[setprod id A = ((-1)^(card E) * setprod id E) *
-      setprod id D] (mod p)"
-    apply (rule zcong_trans)
-    apply (auto simp add: prod_F_zcong zcong_scalar cong del: setprod_cong)
-    done
-  then have "[setprod id A = ((-1)^(card E) * setprod id C)] (mod p)"
-    apply (rule zcong_trans)
-    apply (insert C_prod_eq_D_times_E, erule subst)
-    apply (subst zmult_assoc, auto)
-    done
-  then have "[setprod id A = ((-1)^(card E) * setprod id B)] (mod p)"
-    apply (rule zcong_trans)
-    apply (simp add: C_B_zcong_prod zcong_scalar2 cong del:setprod_cong)
-    done
-  then have "[setprod id A = ((-1)^(card E) *
-    (setprod id ((%x. x * a) ` A)))] (mod p)"
-    by (simp add: B_def)
-  then have "[setprod id A = ((-1)^(card E) * (setprod (%x. x * a) A))]
-    (mod p)"
-    by (simp add:finite_A inj_on_xa_A setprod_reindex_id[symmetric] cong del:setprod_cong)
-  moreover have "setprod (%x. x * a) A =
-    setprod (%x. a) A * setprod id A"
-    using finite_A by (induct set: finite) auto
-  ultimately have "[setprod id A = ((-1)^(card E) * (setprod (%x. a) A *
-    setprod id A))] (mod p)"
-    by simp
-  then have "[setprod id A = ((-1)^(card E) * a^(card A) *
-      setprod id A)](mod p)"
-    apply (rule zcong_trans)
-    apply (simp add: zcong_scalar2 zcong_scalar finite_A setprod_constant zmult_assoc)
-    done
-  then have a: "[setprod id A * (-1)^(card E) =
-      ((-1)^(card E) * a^(card A) * setprod id A * (-1)^(card E))](mod p)"
-    by (rule zcong_scalar)
-  then have "[setprod id A * (-1)^(card E) = setprod id A *
-      (-1)^(card E) * a^(card A) * (-1)^(card E)](mod p)"
-    apply (rule zcong_trans)
-    apply (simp add: a mult_commute mult_left_commute)
-    done
-  then have "[setprod id A * (-1)^(card E) = setprod id A *
-      a^(card A)](mod p)"
-    apply (rule zcong_trans)
-    apply (simp add: aux cong del:setprod_cong)
-    done
-  with this zcong_cancel2 [of p "setprod id A" "-1 ^ card E" "a ^ card A"]
-      p_g_0 A_prod_relprime have "[-1 ^ card E = a ^ card A](mod p)"
-    by (simp add: order_less_imp_le)
-  from this show ?thesis
-    by (simp add: A_card_eq zcong_sym)
-qed
-
-theorem gauss_lemma: "(Legendre a p) = (-1) ^ (card E)"
-proof -
-  from Euler_Criterion p_prime p_g_2 have
-      "[(Legendre a p) = a^(nat (((p) - 1) div 2))] (mod p)"
-    by auto
-  moreover note pre_gauss_lemma
-  ultimately have "[(Legendre a p) = (-1) ^ (card E)] (mod p)"
-    by (rule zcong_trans)
-  moreover from p_a_relprime have "(Legendre a p) = 1 | (Legendre a p) = (-1)"
-    by (auto simp add: Legendre_def)
-  moreover have "(-1::int) ^ (card E) = 1 | (-1::int) ^ (card E) = -1"
-    by (rule neg_one_power)
-  ultimately show ?thesis
-    by (auto simp add: p_g_2 one_not_neg_one_mod_m zcong_sym)
-qed
-
-end
-
-end
--- a/src/HOL/NumberTheory/Int2.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,299 +0,0 @@
-(*  Title:      HOL/Quadratic_Reciprocity/Gauss.thy
-    ID:         $Id$
-    Authors:    Jeremy Avigad, David Gray, and Adam Kramer
-*)
-
-header {*Integers: Divisibility and Congruences*}
-
-theory Int2
-imports Finite2 WilsonRuss
-begin
-
-definition
-  MultInv :: "int => int => int" where
-  "MultInv p x = x ^ nat (p - 2)"
-
-
-subsection {* Useful lemmas about dvd and powers *}
-
-lemma zpower_zdvd_prop1:
-  "0 < n \<Longrightarrow> p dvd y \<Longrightarrow> p dvd ((y::int) ^ n)"
-  by (induct n) (auto simp add: dvd_mult2 [of p y])
-
-lemma zdvd_bounds: "n dvd m ==> m \<le> (0::int) | n \<le> m"
-proof -
-  assume "n dvd m"
-  then have "~(0 < m & m < n)"
-    using zdvd_not_zless [of m n] by auto
-  then show ?thesis by auto
-qed
-
-lemma zprime_zdvd_zmult_better: "[| zprime p;  p dvd (m * n) |] ==>
-    (p dvd m) | (p dvd n)"
-  apply (cases "0 \<le> m")
-  apply (simp add: zprime_zdvd_zmult)
-  apply (insert zprime_zdvd_zmult [of "-m" p n])
-  apply auto
-  done
-
-lemma zpower_zdvd_prop2:
-    "zprime p \<Longrightarrow> p dvd ((y::int) ^ n) \<Longrightarrow> 0 < n \<Longrightarrow> p dvd y"
-  apply (induct n)
-   apply simp
-  apply (frule zprime_zdvd_zmult_better)
-   apply simp
-  apply (force simp del:dvd_mult)
-  done
-
-lemma div_prop1: "[| 0 < z; (x::int) < y * z |] ==> x div z < y"
-proof -
-  assume "0 < z" then have modth: "x mod z \<ge> 0" by simp
-  have "(x div z) * z \<le> (x div z) * z" by simp
-  then have "(x div z) * z \<le> (x div z) * z + x mod z" using modth by arith 
-  also have "\<dots> = x"
-    by (auto simp add: zmod_zdiv_equality [symmetric] zmult_ac)
-  also assume  "x < y * z"
-  finally show ?thesis
-    by (auto simp add: prems mult_less_cancel_right, insert prems, arith)
-qed
-
-lemma div_prop2: "[| 0 < z; (x::int) < (y * z) + z |] ==> x div z \<le> y"
-proof -
-  assume "0 < z" and "x < (y * z) + z"
-  then have "x < (y + 1) * z" by (auto simp add: int_distrib)
-  then have "x div z < y + 1"
-    apply -
-    apply (rule_tac y = "y + 1" in div_prop1)
-    apply (auto simp add: prems)
-    done
-  then show ?thesis by auto
-qed
-
-lemma zdiv_leq_prop: "[| 0 < y |] ==> y * (x div y) \<le> (x::int)"
-proof-
-  assume "0 < y"
-  from zmod_zdiv_equality have "x = y * (x div y) + x mod y" by auto
-  moreover have "0 \<le> x mod y"
-    by (auto simp add: prems pos_mod_sign)
-  ultimately show ?thesis
-    by arith
-qed
-
-
-subsection {* Useful properties of congruences *}
-
-lemma zcong_eq_zdvd_prop: "[x = 0](mod p) = (p dvd x)"
-  by (auto simp add: zcong_def)
-
-lemma zcong_id: "[m = 0] (mod m)"
-  by (auto simp add: zcong_def)
-
-lemma zcong_shift: "[a = b] (mod m) ==> [a + c = b + c] (mod m)"
-  by (auto simp add: zcong_refl zcong_zadd)
-
-lemma zcong_zpower: "[x = y](mod m) ==> [x^z = y^z](mod m)"
-  by (induct z) (auto simp add: zcong_zmult)
-
-lemma zcong_eq_trans: "[| [a = b](mod m); b = c; [c = d](mod m) |] ==>
-    [a = d](mod m)"
-  apply (erule zcong_trans)
-  apply simp
-  done
-
-lemma aux1: "a - b = (c::int) ==> a = c + b"
-  by auto
-
-lemma zcong_zmult_prop1: "[a = b](mod m) ==> ([c = a * d](mod m) =
-    [c = b * d] (mod m))"
-  apply (auto simp add: zcong_def dvd_def)
-  apply (rule_tac x = "ka + k * d" in exI)
-  apply (drule aux1)+
-  apply (auto simp add: int_distrib)
-  apply (rule_tac x = "ka - k * d" in exI)
-  apply (drule aux1)+
-  apply (auto simp add: int_distrib)
-  done
-
-lemma zcong_zmult_prop2: "[a = b](mod m) ==>
-    ([c = d * a](mod m) = [c = d * b] (mod m))"
-  by (auto simp add: zmult_ac zcong_zmult_prop1)
-
-lemma zcong_zmult_prop3: "[| zprime p; ~[x = 0] (mod p);
-    ~[y = 0] (mod p) |] ==> ~[x * y = 0] (mod p)"
-  apply (auto simp add: zcong_def)
-  apply (drule zprime_zdvd_zmult_better, auto)
-  done
-
-lemma zcong_less_eq: "[| 0 < x; 0 < y; 0 < m; [x = y] (mod m);
-    x < m; y < m |] ==> x = y"
-  by (metis zcong_not zcong_sym zless_linear)
-
-lemma zcong_neg_1_impl_ne_1: "[| 2 < p; [x = -1] (mod p) |] ==>
-    ~([x = 1] (mod p))"
-proof
-  assume "2 < p" and "[x = 1] (mod p)" and "[x = -1] (mod p)"
-  then have "[1 = -1] (mod p)"
-    apply (auto simp add: zcong_sym)
-    apply (drule zcong_trans, auto)
-    done
-  then have "[1 + 1 = -1 + 1] (mod p)"
-    by (simp only: zcong_shift)
-  then have "[2 = 0] (mod p)"
-    by auto
-  then have "p dvd 2"
-    by (auto simp add: dvd_def zcong_def)
-  with prems show False
-    by (auto simp add: zdvd_not_zless)
-qed
-
-lemma zcong_zero_equiv_div: "[a = 0] (mod m) = (m dvd a)"
-  by (auto simp add: zcong_def)
-
-lemma zcong_zprime_prod_zero: "[| zprime p; 0 < a |] ==>
-    [a * b = 0] (mod p) ==> [a = 0] (mod p) | [b = 0] (mod p)"
-  by (auto simp add: zcong_zero_equiv_div zprime_zdvd_zmult)
-
-lemma zcong_zprime_prod_zero_contra: "[| zprime p; 0 < a |] ==>
-  ~[a = 0](mod p) & ~[b = 0](mod p) ==> ~[a * b = 0] (mod p)"
-  apply auto
-  apply (frule_tac a = a and b = b and p = p in zcong_zprime_prod_zero)
-  apply auto
-  done
-
-lemma zcong_not_zero: "[| 0 < x; x < m |] ==> ~[x = 0] (mod m)"
-  by (auto simp add: zcong_zero_equiv_div zdvd_not_zless)
-
-lemma zcong_zero: "[| 0 \<le> x; x < m; [x = 0](mod m) |] ==> x = 0"
-  apply (drule order_le_imp_less_or_eq, auto)
-  apply (frule_tac m = m in zcong_not_zero)
-  apply auto
-  done
-
-lemma all_relprime_prod_relprime: "[| finite A; \<forall>x \<in> A. zgcd x y = 1 |]
-    ==> zgcd (setprod id A) y = 1"
-  by (induct set: finite) (auto simp add: zgcd_zgcd_zmult)
-
-
-subsection {* Some properties of MultInv *}
-
-lemma MultInv_prop1: "[| 2 < p; [x = y] (mod p) |] ==>
-    [(MultInv p x) = (MultInv p y)] (mod p)"
-  by (auto simp add: MultInv_def zcong_zpower)
-
-lemma MultInv_prop2: "[| 2 < p; zprime p; ~([x = 0](mod p)) |] ==>
-  [(x * (MultInv p x)) = 1] (mod p)"
-proof (simp add: MultInv_def zcong_eq_zdvd_prop)
-  assume "2 < p" and "zprime p" and "~ p dvd x"
-  have "x * x ^ nat (p - 2) = x ^ (nat (p - 2) + 1)"
-    by auto
-  also from prems have "nat (p - 2) + 1 = nat (p - 2 + 1)"
-    by (simp only: nat_add_distrib)
-  also have "p - 2 + 1 = p - 1" by arith
-  finally have "[x * x ^ nat (p - 2) = x ^ nat (p - 1)] (mod p)"
-    by (rule ssubst, auto)
-  also from prems have "[x ^ nat (p - 1) = 1] (mod p)"
-    by (auto simp add: Little_Fermat)
-  finally (zcong_trans) show "[x * x ^ nat (p - 2) = 1] (mod p)" .
-qed
-
-lemma MultInv_prop2a: "[| 2 < p; zprime p; ~([x = 0](mod p)) |] ==>
-    [(MultInv p x) * x = 1] (mod p)"
-  by (auto simp add: MultInv_prop2 zmult_ac)
-
-lemma aux_1: "2 < p ==> ((nat p) - 2) = (nat (p - 2))"
-  by (simp add: nat_diff_distrib)
-
-lemma aux_2: "2 < p ==> 0 < nat (p - 2)"
-  by auto
-
-lemma MultInv_prop3: "[| 2 < p; zprime p; ~([x = 0](mod p)) |] ==>
-    ~([MultInv p x = 0](mod p))"
-  apply (auto simp add: MultInv_def zcong_eq_zdvd_prop aux_1)
-  apply (drule aux_2)
-  apply (drule zpower_zdvd_prop2, auto)
-  done
-
-lemma aux__1: "[| 2 < p; zprime p; ~([x = 0](mod p))|] ==>
-    [(MultInv p (MultInv p x)) = (x * (MultInv p x) *
-      (MultInv p (MultInv p x)))] (mod p)"
-  apply (drule MultInv_prop2, auto)
-  apply (drule_tac k = "MultInv p (MultInv p x)" in zcong_scalar, auto)
-  apply (auto simp add: zcong_sym)
-  done
-
-lemma aux__2: "[| 2 < p; zprime p; ~([x = 0](mod p))|] ==>
-    [(x * (MultInv p x) * (MultInv p (MultInv p x))) = x] (mod p)"
-  apply (frule MultInv_prop3, auto)
-  apply (insert MultInv_prop2 [of p "MultInv p x"], auto)
-  apply (drule MultInv_prop2, auto)
-  apply (drule_tac k = x in zcong_scalar2, auto)
-  apply (auto simp add: zmult_ac)
-  done
-
-lemma MultInv_prop4: "[| 2 < p; zprime p; ~([x = 0](mod p)) |] ==>
-    [(MultInv p (MultInv p x)) = x] (mod p)"
-  apply (frule aux__1, auto)
-  apply (drule aux__2, auto)
-  apply (drule zcong_trans, auto)
-  done
-
-lemma MultInv_prop5: "[| 2 < p; zprime p; ~([x = 0](mod p));
-    ~([y = 0](mod p)); [(MultInv p x) = (MultInv p y)] (mod p) |] ==>
-    [x = y] (mod p)"
-  apply (drule_tac a = "MultInv p x" and b = "MultInv p y" and
-    m = p and k = x in zcong_scalar)
-  apply (insert MultInv_prop2 [of p x], simp)
-  apply (auto simp only: zcong_sym [of "MultInv p x * x"])
-  apply (auto simp add:  zmult_ac)
-  apply (drule zcong_trans, auto)
-  apply (drule_tac a = "x * MultInv p y" and k = y in zcong_scalar, auto)
-  apply (insert MultInv_prop2a [of p y], auto simp add: zmult_ac)
-  apply (insert zcong_zmult_prop2 [of "y * MultInv p y" 1 p y x])
-  apply (auto simp add: zcong_sym)
-  done
-
-lemma MultInv_zcong_prop1: "[| 2 < p; [j = k] (mod p) |] ==>
-    [a * MultInv p j = a * MultInv p k] (mod p)"
-  by (drule MultInv_prop1, auto simp add: zcong_scalar2)
-
-lemma aux___1: "[j = a * MultInv p k] (mod p) ==>
-    [j * k = a * MultInv p k * k] (mod p)"
-  by (auto simp add: zcong_scalar)
-
-lemma aux___2: "[|2 < p; zprime p; ~([k = 0](mod p));
-    [j * k = a * MultInv p k * k] (mod p) |] ==> [j * k = a] (mod p)"
-  apply (insert MultInv_prop2a [of p k] zcong_zmult_prop2
-    [of "MultInv p k * k" 1 p "j * k" a])
-  apply (auto simp add: zmult_ac)
-  done
-
-lemma aux___3: "[j * k = a] (mod p) ==> [(MultInv p j) * j * k =
-     (MultInv p j) * a] (mod p)"
-  by (auto simp add: zmult_assoc zcong_scalar2)
-
-lemma aux___4: "[|2 < p; zprime p; ~([j = 0](mod p));
-    [(MultInv p j) * j * k = (MultInv p j) * a] (mod p) |]
-       ==> [k = a * (MultInv p j)] (mod p)"
-  apply (insert MultInv_prop2a [of p j] zcong_zmult_prop1
-    [of "MultInv p j * j" 1 p "MultInv p j * a" k])
-  apply (auto simp add: zmult_ac zcong_sym)
-  done
-
-lemma MultInv_zcong_prop2: "[| 2 < p; zprime p; ~([k = 0](mod p));
-    ~([j = 0](mod p)); [j = a * MultInv p k] (mod p) |] ==>
-    [k = a * MultInv p j] (mod p)"
-  apply (drule aux___1)
-  apply (frule aux___2, auto)
-  by (drule aux___3, drule aux___4, auto)
-
-lemma MultInv_zcong_prop3: "[| 2 < p; zprime p; ~([a = 0](mod p));
-    ~([k = 0](mod p)); ~([j = 0](mod p));
-    [a * MultInv p j = a * MultInv p k] (mod p) |] ==>
-      [j = k] (mod p)"
-  apply (auto simp add: zcong_eq_zdvd_prop [of a p])
-  apply (frule zprime_imp_zrelprime, auto)
-  apply (insert zcong_cancel2 [of p a "MultInv p j" "MultInv p k"], auto)
-  apply (drule MultInv_prop5, auto)
-  done
-
-end
--- a/src/HOL/NumberTheory/IntFact.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,96 +0,0 @@
-(*  Title:      HOL/NumberTheory/IntFact.thy
-    ID:         $Id$
-    Author:     Thomas M. Rasmussen
-    Copyright   2000  University of Cambridge
-*)
-
-header {* Factorial on integers *}
-
-theory IntFact imports IntPrimes begin
-
-text {*
-  Factorial on integers and recursively defined set including all
-  Integers from @{text 2} up to @{text a}.  Plus definition of product
-  of finite set.
-
-  \bigskip
-*}
-
-consts
-  zfact :: "int => int"
-  d22set :: "int => int set"
-
-recdef zfact  "measure ((\<lambda>n. nat n) :: int => nat)"
-  "zfact n = (if n \<le> 0 then 1 else n * zfact (n - 1))"
-
-recdef d22set  "measure ((\<lambda>a. nat a) :: int => nat)"
-  "d22set a = (if 1 < a then insert a (d22set (a - 1)) else {})"
-
-
-text {*
-  \medskip @{term d22set} --- recursively defined set including all
-  integers from @{text 2} up to @{text a}
-*}
-
-declare d22set.simps [simp del]
-
-
-lemma d22set_induct:
-  assumes "!!a. P {} a"
-    and "!!a. 1 < (a::int) ==> P (d22set (a - 1)) (a - 1) ==> P (d22set a) a"
-  shows "P (d22set u) u"
-  apply (rule d22set.induct)
-  apply safe
-   prefer 2
-   apply (case_tac "1 < a")
-    apply (rule_tac prems)
-     apply (simp_all (no_asm_simp))
-   apply (simp_all (no_asm_simp) add: d22set.simps prems)
-  done
-
-lemma d22set_g_1 [rule_format]: "b \<in> d22set a --> 1 < b"
-  apply (induct a rule: d22set_induct)
-   apply simp
-  apply (subst d22set.simps)
-  apply auto
-  done
-
-lemma d22set_le [rule_format]: "b \<in> d22set a --> b \<le> a"
-  apply (induct a rule: d22set_induct)
-  apply simp
-   apply (subst d22set.simps)
-   apply auto
-  done
-
-lemma d22set_le_swap: "a < b ==> b \<notin> d22set a"
-  by (auto dest: d22set_le)
-
-lemma d22set_mem: "1 < b \<Longrightarrow> b \<le> a \<Longrightarrow> b \<in> d22set a"
-  apply (induct a rule: d22set.induct)
-  apply auto
-   apply (simp_all add: d22set.simps)
-  done
-
-lemma d22set_fin: "finite (d22set a)"
-  apply (induct a rule: d22set_induct)
-   prefer 2
-   apply (subst d22set.simps)
-   apply auto
-  done
-
-
-declare zfact.simps [simp del]
-
-lemma d22set_prod_zfact: "\<Prod>(d22set a) = zfact a"
-  apply (induct a rule: d22set.induct)
-  apply safe
-   apply (simp add: d22set.simps zfact.simps)
-  apply (subst d22set.simps)
-  apply (subst zfact.simps)
-  apply (case_tac "1 < a")
-   prefer 2
-   apply (simp add: d22set.simps zfact.simps)
-  apply (simp add: d22set_fin d22set_le_swap)
-  done
-
-end
--- a/src/HOL/NumberTheory/IntPrimes.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,423 +0,0 @@
-(*  Title:      HOL/NumberTheory/IntPrimes.thy
-    ID:         $Id$
-    Author:     Thomas M. Rasmussen
-    Copyright   2000  University of Cambridge
-*)
-
-header {* Divisibility and prime numbers (on integers) *}
-
-theory IntPrimes
-imports Main Primes
-begin
-
-text {*
-  The @{text dvd} relation, GCD, Euclid's extended algorithm, primes,
-  congruences (all on the Integers).  Comparable to theory @{text
-  Primes}, but @{text dvd} is included here as it is not present in
-  main HOL.  Also includes extended GCD and congruences not present in
-  @{text Primes}.
-*}
-
-
-subsection {* Definitions *}
-
-consts
-  xzgcda :: "int * int * int * int * int * int * int * int => int * int * int"
-
-recdef xzgcda
-  "measure ((\<lambda>(m, n, r', r, s', s, t', t). nat r)
-    :: int * int * int * int *int * int * int * int => nat)"
-  "xzgcda (m, n, r', r, s', s, t', t) =
-	(if r \<le> 0 then (r', s', t')
-	 else xzgcda (m, n, r, r' mod r, 
-		      s, s' - (r' div r) * s, 
-		      t, t' - (r' div r) * t))"
-
-definition
-  zprime :: "int \<Rightarrow> bool" where
-  "zprime p = (1 < p \<and> (\<forall>m. 0 <= m & m dvd p --> m = 1 \<or> m = p))"
-
-definition
-  xzgcd :: "int => int => int * int * int" where
-  "xzgcd m n = xzgcda (m, n, m, n, 1, 0, 0, 1)"
-
-definition
-  zcong :: "int => int => int => bool"  ("(1[_ = _] '(mod _'))") where
-  "[a = b] (mod m) = (m dvd (a - b))"
-
-subsection {* Euclid's Algorithm and GCD *}
-
-
-lemma zrelprime_zdvd_zmult_aux:
-     "zgcd n k = 1 ==> k dvd m * n ==> 0 \<le> m ==> k dvd m"
-    by (metis abs_of_nonneg dvd_triv_right zgcd_greatest_iff zgcd_zmult_distrib2_abs zmult_1_right)
-
-lemma zrelprime_zdvd_zmult: "zgcd n k = 1 ==> k dvd m * n ==> k dvd m"
-  apply (case_tac "0 \<le> m")
-   apply (blast intro: zrelprime_zdvd_zmult_aux)
-  apply (subgoal_tac "k dvd -m")
-   apply (rule_tac [2] zrelprime_zdvd_zmult_aux, auto)
-  done
-
-lemma zgcd_geq_zero: "0 <= zgcd x y"
-  by (auto simp add: zgcd_def)
-
-text{*This is merely a sanity check on zprime, since the previous version
-      denoted the empty set.*}
-lemma "zprime 2"
-  apply (auto simp add: zprime_def) 
-  apply (frule zdvd_imp_le, simp) 
-  apply (auto simp add: order_le_less dvd_def) 
-  done
-
-lemma zprime_imp_zrelprime:
-    "zprime p ==> \<not> p dvd n ==> zgcd n p = 1"
-  apply (auto simp add: zprime_def)
-  apply (metis zgcd_geq_zero zgcd_zdvd1 zgcd_zdvd2)
-  done
-
-lemma zless_zprime_imp_zrelprime:
-    "zprime p ==> 0 < n ==> n < p ==> zgcd n p = 1"
-  apply (erule zprime_imp_zrelprime)
-  apply (erule zdvd_not_zless, assumption)
-  done
-
-lemma zprime_zdvd_zmult:
-    "0 \<le> (m::int) ==> zprime p ==> p dvd m * n ==> p dvd m \<or> p dvd n"
-  by (metis zgcd_zdvd1 zgcd_zdvd2 zgcd_pos zprime_def zrelprime_dvd_mult)
-
-lemma zgcd_zadd_zmult [simp]: "zgcd (m + n * k) n = zgcd m n"
-  apply (rule zgcd_eq [THEN trans])
-  apply (simp add: mod_add_eq)
-  apply (rule zgcd_eq [symmetric])
-  done
-
-lemma zgcd_zdvd_zgcd_zmult: "zgcd m n dvd zgcd (k * m) n"
-by (simp add: zgcd_greatest_iff)
-
-lemma zgcd_zmult_zdvd_zgcd:
-    "zgcd k n = 1 ==> zgcd (k * m) n dvd zgcd m n"
-  apply (simp add: zgcd_greatest_iff)
-  apply (rule_tac n = k in zrelprime_zdvd_zmult)
-   prefer 2
-   apply (simp add: zmult_commute)
-  apply (metis zgcd_1 zgcd_commute zgcd_left_commute)
-  done
-
-lemma zgcd_zmult_cancel: "zgcd k n = 1 ==> zgcd (k * m) n = zgcd m n"
-  by (simp add: zgcd_def nat_abs_mult_distrib gcd_mult_cancel)
-
-lemma zgcd_zgcd_zmult:
-    "zgcd k m = 1 ==> zgcd n m = 1 ==> zgcd (k * n) m = 1"
-  by (simp add: zgcd_zmult_cancel)
-
-lemma zdvd_iff_zgcd: "0 < m ==> m dvd n \<longleftrightarrow> zgcd n m = m"
-  by (metis abs_of_pos zdvd_mult_div_cancel zgcd_0 zgcd_commute zgcd_geq_zero zgcd_zdvd2 zgcd_zmult_eq_self)
-
-
-
-subsection {* Congruences *}
-
-lemma zcong_1 [simp]: "[a = b] (mod 1)"
-  by (unfold zcong_def, auto)
-
-lemma zcong_refl [simp]: "[k = k] (mod m)"
-  by (unfold zcong_def, auto)
-
-lemma zcong_sym: "[a = b] (mod m) = [b = a] (mod m)"
-  unfolding zcong_def minus_diff_eq [of a, symmetric] dvd_minus_iff ..
-
-lemma zcong_zadd:
-    "[a = b] (mod m) ==> [c = d] (mod m) ==> [a + c = b + d] (mod m)"
-  apply (unfold zcong_def)
-  apply (rule_tac s = "(a - b) + (c - d)" in subst)
-   apply (rule_tac [2] dvd_add, auto)
-  done
-
-lemma zcong_zdiff:
-    "[a = b] (mod m) ==> [c = d] (mod m) ==> [a - c = b - d] (mod m)"
-  apply (unfold zcong_def)
-  apply (rule_tac s = "(a - b) - (c - d)" in subst)
-   apply (rule_tac [2] dvd_diff, auto)
-  done
-
-lemma zcong_trans:
-  "[a = b] (mod m) ==> [b = c] (mod m) ==> [a = c] (mod m)"
-unfolding zcong_def by (auto elim!: dvdE simp add: algebra_simps)
-
-lemma zcong_zmult:
-    "[a = b] (mod m) ==> [c = d] (mod m) ==> [a * c = b * d] (mod m)"
-  apply (rule_tac b = "b * c" in zcong_trans)
-   apply (unfold zcong_def)
-  apply (metis zdiff_zmult_distrib2 dvd_mult zmult_commute)
-  apply (metis zdiff_zmult_distrib2 dvd_mult)
-  done
-
-lemma zcong_scalar: "[a = b] (mod m) ==> [a * k = b * k] (mod m)"
-  by (rule zcong_zmult, simp_all)
-
-lemma zcong_scalar2: "[a = b] (mod m) ==> [k * a = k * b] (mod m)"
-  by (rule zcong_zmult, simp_all)
-
-lemma zcong_zmult_self: "[a * m = b * m] (mod m)"
-  apply (unfold zcong_def)
-  apply (rule dvd_diff, simp_all)
-  done
-
-lemma zcong_square:
-   "[| zprime p;  0 < a;  [a * a = 1] (mod p)|]
-    ==> [a = 1] (mod p) \<or> [a = p - 1] (mod p)"
-  apply (unfold zcong_def)
-  apply (rule zprime_zdvd_zmult)
-    apply (rule_tac [3] s = "a * a - 1 + p * (1 - a)" in subst)
-     prefer 4
-     apply (simp add: zdvd_reduce)
-    apply (simp_all add: zdiff_zmult_distrib zmult_commute zdiff_zmult_distrib2)
-  done
-
-lemma zcong_cancel:
-  "0 \<le> m ==>
-    zgcd k m = 1 ==> [a * k = b * k] (mod m) = [a = b] (mod m)"
-  apply safe
-   prefer 2
-   apply (blast intro: zcong_scalar)
-  apply (case_tac "b < a")
-   prefer 2
-   apply (subst zcong_sym)
-   apply (unfold zcong_def)
-   apply (rule_tac [!] zrelprime_zdvd_zmult)
-     apply (simp_all add: zdiff_zmult_distrib)
-  apply (subgoal_tac "m dvd (-(a * k - b * k))")
-   apply simp
-  apply (subst dvd_minus_iff, assumption)
-  done
-
-lemma zcong_cancel2:
-  "0 \<le> m ==>
-    zgcd k m = 1 ==> [k * a = k * b] (mod m) = [a = b] (mod m)"
-  by (simp add: zmult_commute zcong_cancel)
-
-lemma zcong_zgcd_zmult_zmod:
-  "[a = b] (mod m) ==> [a = b] (mod n) ==> zgcd m n = 1
-    ==> [a = b] (mod m * n)"
-  apply (auto simp add: zcong_def dvd_def)
-  apply (subgoal_tac "m dvd n * ka")
-   apply (subgoal_tac "m dvd ka")
-    apply (case_tac [2] "0 \<le> ka")
-  apply (metis zdvd_mult_div_cancel dvd_refl dvd_mult_left zmult_commute zrelprime_zdvd_zmult)
-  apply (metis abs_dvd_iff abs_of_nonneg zadd_0 zgcd_0_left zgcd_commute zgcd_zadd_zmult zgcd_zdvd_zgcd_zmult zgcd_zmult_distrib2_abs zmult_1_right zmult_commute)
-  apply (metis mult_le_0_iff  zdvd_mono zdvd_mult_cancel dvd_triv_left zero_le_mult_iff zle_anti_sym zle_linear zle_refl zmult_commute zrelprime_zdvd_zmult)
-  apply (metis dvd_triv_left)
-  done
-
-lemma zcong_zless_imp_eq:
-  "0 \<le> a ==>
-    a < m ==> 0 \<le> b ==> b < m ==> [a = b] (mod m) ==> a = b"
-  apply (unfold zcong_def dvd_def, auto)
-  apply (drule_tac f = "\<lambda>z. z mod m" in arg_cong)
-  apply (metis diff_add_cancel mod_pos_pos_trivial zadd_0 zadd_commute zmod_eq_0_iff mod_add_right_eq)
-  done
-
-lemma zcong_square_zless:
-  "zprime p ==> 0 < a ==> a < p ==>
-    [a * a = 1] (mod p) ==> a = 1 \<or> a = p - 1"
-  apply (cut_tac p = p and a = a in zcong_square)
-     apply (simp add: zprime_def)
-    apply (auto intro: zcong_zless_imp_eq)
-  done
-
-lemma zcong_not:
-    "0 < a ==> a < m ==> 0 < b ==> b < a ==> \<not> [a = b] (mod m)"
-  apply (unfold zcong_def)
-  apply (rule zdvd_not_zless, auto)
-  done
-
-lemma zcong_zless_0:
-    "0 \<le> a ==> a < m ==> [a = 0] (mod m) ==> a = 0"
-  apply (unfold zcong_def dvd_def, auto)
-  apply (metis div_pos_pos_trivial linorder_not_less div_mult_self1_is_id)
-  done
-
-lemma zcong_zless_unique:
-    "0 < m ==> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))"
-  apply auto
-   prefer 2 apply (metis zcong_sym zcong_trans zcong_zless_imp_eq)
-  apply (unfold zcong_def dvd_def)
-  apply (rule_tac x = "a mod m" in exI, auto)
-  apply (metis zmult_div_cancel)
-  done
-
-lemma zcong_iff_lin: "([a = b] (mod m)) = (\<exists>k. b = a + m * k)"
-  unfolding zcong_def
-  apply (auto elim!: dvdE simp add: algebra_simps)
-  apply (rule_tac x = "-k" in exI) apply simp
-  done
-
-lemma zgcd_zcong_zgcd:
-  "0 < m ==>
-    zgcd a m = 1 ==> [a = b] (mod m) ==> zgcd b m = 1"
-  by (auto simp add: zcong_iff_lin)
-
-lemma zcong_zmod_aux:
-     "a - b = (m::int) * (a div m - b div m) + (a mod m - b mod m)"
-  by(simp add: zdiff_zmult_distrib2 add_diff_eq eq_diff_eq add_ac)
-
-lemma zcong_zmod: "[a = b] (mod m) = [a mod m = b mod m] (mod m)"
-  apply (unfold zcong_def)
-  apply (rule_tac t = "a - b" in ssubst)
-  apply (rule_tac m = m in zcong_zmod_aux)
-  apply (rule trans)
-   apply (rule_tac [2] k = m and m = "a div m - b div m" in zdvd_reduce)
-  apply (simp add: zadd_commute)
-  done
-
-lemma zcong_zmod_eq: "0 < m ==> [a = b] (mod m) = (a mod m = b mod m)"
-  apply auto
-  apply (metis pos_mod_conj zcong_zless_imp_eq zcong_zmod)
-  apply (metis zcong_refl zcong_zmod)
-  done
-
-lemma zcong_zminus [iff]: "[a = b] (mod -m) = [a = b] (mod m)"
-  by (auto simp add: zcong_def)
-
-lemma zcong_zero [iff]: "[a = b] (mod 0) = (a = b)"
-  by (auto simp add: zcong_def)
-
-lemma "[a = b] (mod m) = (a mod m = b mod m)"
-  apply (case_tac "m = 0", simp add: DIVISION_BY_ZERO)
-  apply (simp add: linorder_neq_iff)
-  apply (erule disjE)  
-   prefer 2 apply (simp add: zcong_zmod_eq)
-  txt{*Remainding case: @{term "m<0"}*}
-  apply (rule_tac t = m in zminus_zminus [THEN subst])
-  apply (subst zcong_zminus)
-  apply (subst zcong_zmod_eq, arith)
-  apply (frule neg_mod_bound [of _ a], frule neg_mod_bound [of _ b]) 
-  apply (simp add: zmod_zminus2_eq_if del: neg_mod_bound)
-  done
-
-subsection {* Modulo *}
-
-lemma zmod_zdvd_zmod:
-    "0 < (m::int) ==> m dvd b ==> (a mod b mod m) = (a mod m)"
-  by (rule mod_mod_cancel) 
-
-
-subsection {* Extended GCD *}
-
-declare xzgcda.simps [simp del]
-
-lemma xzgcd_correct_aux1:
-  "zgcd r' r = k --> 0 < r -->
-    (\<exists>sn tn. xzgcda (m, n, r', r, s', s, t', t) = (k, sn, tn))"
-  apply (rule_tac u = m and v = n and w = r' and x = r and y = s' and
-    z = s and aa = t' and ab = t in xzgcda.induct)
-  apply (subst zgcd_eq)
-  apply (subst xzgcda.simps, auto)
-  apply (case_tac "r' mod r = 0")
-   prefer 2
-   apply (frule_tac a = "r'" in pos_mod_sign, auto)
-  apply (rule exI)
-  apply (rule exI)
-  apply (subst xzgcda.simps, auto)
-  done
-
-lemma xzgcd_correct_aux2:
-  "(\<exists>sn tn. xzgcda (m, n, r', r, s', s, t', t) = (k, sn, tn)) --> 0 < r -->
-    zgcd r' r = k"
-  apply (rule_tac u = m and v = n and w = r' and x = r and y = s' and
-    z = s and aa = t' and ab = t in xzgcda.induct)
-  apply (subst zgcd_eq)
-  apply (subst xzgcda.simps)
-  apply (auto simp add: linorder_not_le)
-  apply (case_tac "r' mod r = 0")
-   prefer 2
-   apply (frule_tac a = "r'" in pos_mod_sign, auto)
-  apply (metis Pair_eq simps zle_refl)
-  done
-
-lemma xzgcd_correct:
-    "0 < n ==> (zgcd m n = k) = (\<exists>s t. xzgcd m n = (k, s, t))"
-  apply (unfold xzgcd_def)
-  apply (rule iffI)
-   apply (rule_tac [2] xzgcd_correct_aux2 [THEN mp, THEN mp])
-    apply (rule xzgcd_correct_aux1 [THEN mp, THEN mp], auto)
-  done
-
-
-text {* \medskip @{term xzgcd} linear *}
-
-lemma xzgcda_linear_aux1:
-  "(a - r * b) * m + (c - r * d) * (n::int) =
-   (a * m + c * n) - r * (b * m + d * n)"
-  by (simp add: zdiff_zmult_distrib zadd_zmult_distrib2 zmult_assoc)
-
-lemma xzgcda_linear_aux2:
-  "r' = s' * m + t' * n ==> r = s * m + t * n
-    ==> (r' mod r) = (s' - (r' div r) * s) * m + (t' - (r' div r) * t) * (n::int)"
-  apply (rule trans)
-   apply (rule_tac [2] xzgcda_linear_aux1 [symmetric])
-  apply (simp add: eq_diff_eq mult_commute)
-  done
-
-lemma order_le_neq_implies_less: "(x::'a::order) \<le> y ==> x \<noteq> y ==> x < y"
-  by (rule iffD2 [OF order_less_le conjI])
-
-lemma xzgcda_linear [rule_format]:
-  "0 < r --> xzgcda (m, n, r', r, s', s, t', t) = (rn, sn, tn) -->
-    r' = s' * m + t' * n -->  r = s * m + t * n --> rn = sn * m + tn * n"
-  apply (rule_tac u = m and v = n and w = r' and x = r and y = s' and
-    z = s and aa = t' and ab = t in xzgcda.induct)
-  apply (subst xzgcda.simps)
-  apply (simp (no_asm))
-  apply (rule impI)+
-  apply (case_tac "r' mod r = 0")
-   apply (simp add: xzgcda.simps, clarify)
-  apply (subgoal_tac "0 < r' mod r")
-   apply (rule_tac [2] order_le_neq_implies_less)
-   apply (rule_tac [2] pos_mod_sign)
-    apply (cut_tac m = m and n = n and r' = r' and r = r and s' = s' and
-      s = s and t' = t' and t = t in xzgcda_linear_aux2, auto)
-  done
-
-lemma xzgcd_linear:
-    "0 < n ==> xzgcd m n = (r, s, t) ==> r = s * m + t * n"
-  apply (unfold xzgcd_def)
-  apply (erule xzgcda_linear, assumption, auto)
-  done
-
-lemma zgcd_ex_linear:
-    "0 < n ==> zgcd m n = k ==> (\<exists>s t. k = s * m + t * n)"
-  apply (simp add: xzgcd_correct, safe)
-  apply (rule exI)+
-  apply (erule xzgcd_linear, auto)
-  done
-
-lemma zcong_lineq_ex:
-    "0 < n ==> zgcd a n = 1 ==> \<exists>x. [a * x = 1] (mod n)"
-  apply (cut_tac m = a and n = n and k = 1 in zgcd_ex_linear, safe)
-  apply (rule_tac x = s in exI)
-  apply (rule_tac b = "s * a + t * n" in zcong_trans)
-   prefer 2
-   apply simp
-  apply (unfold zcong_def)
-  apply (simp (no_asm) add: zmult_commute)
-  done
-
-lemma zcong_lineq_unique:
-  "0 < n ==>
-    zgcd a n = 1 ==> \<exists>!x. 0 \<le> x \<and> x < n \<and> [a * x = b] (mod n)"
-  apply auto
-   apply (rule_tac [2] zcong_zless_imp_eq)
-       apply (tactic {* stac (thm "zcong_cancel2" RS sym) 6 *})
-         apply (rule_tac [8] zcong_trans)
-          apply (simp_all (no_asm_simp))
-   prefer 2
-   apply (simp add: zcong_sym)
-  apply (cut_tac a = a and n = n in zcong_lineq_ex, auto)
-  apply (rule_tac x = "x * b mod n" in exI, safe)
-    apply (simp_all (no_asm_simp))
-  apply (metis zcong_scalar zcong_zmod zmod_zmult1_eq zmult_1 zmult_assoc)
-  done
-
-end
--- a/src/HOL/NumberTheory/Quadratic_Reciprocity.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,644 +0,0 @@
-(*  Title:      HOL/NumberTheory/Quadratic_Reciprocity.thy
-    ID:         $Id$
-    Authors:    Jeremy Avigad, David Gray, and Adam Kramer
-*)
-
-header {* The law of Quadratic reciprocity *}
-
-theory Quadratic_Reciprocity
-imports Gauss
-begin
-
-text {*
-  Lemmas leading up to the proof of theorem 3.3 in Niven and
-  Zuckerman's presentation.
-*}
-
-context GAUSS
-begin
-
-lemma QRLemma1: "a * setsum id A =
-  p * setsum (%x. ((x * a) div p)) A + setsum id D + setsum id E"
-proof -
-  from finite_A have "a * setsum id A = setsum (%x. a * x) A"
-    by (auto simp add: setsum_const_mult id_def)
-  also have "setsum (%x. a * x) = setsum (%x. x * a)"
-    by (auto simp add: zmult_commute)
-  also have "setsum (%x. x * a) A = setsum id B"
-    by (simp add: B_def setsum_reindex_id[OF inj_on_xa_A])
-  also have "... = setsum (%x. p * (x div p) + StandardRes p x) B"
-    by (auto simp add: StandardRes_def zmod_zdiv_equality)
-  also have "... = setsum (%x. p * (x div p)) B + setsum (StandardRes p) B"
-    by (rule setsum_addf)
-  also have "setsum (StandardRes p) B = setsum id C"
-    by (auto simp add: C_def setsum_reindex_id[OF SR_B_inj])
-  also from C_eq have "... = setsum id (D \<union> E)"
-    by auto
-  also from finite_D finite_E have "... = setsum id D + setsum id E"
-    by (rule setsum_Un_disjoint) (auto simp add: D_def E_def)
-  also have "setsum (%x. p * (x div p)) B =
-      setsum ((%x. p * (x div p)) o (%x. (x * a))) A"
-    by (auto simp add: B_def setsum_reindex inj_on_xa_A)
-  also have "... = setsum (%x. p * ((x * a) div p)) A"
-    by (auto simp add: o_def)
-  also from finite_A have "setsum (%x. p * ((x * a) div p)) A =
-    p * setsum (%x. ((x * a) div p)) A"
-    by (auto simp add: setsum_const_mult)
-  finally show ?thesis by arith
-qed
-
-lemma QRLemma2: "setsum id A = p * int (card E) - setsum id E +
-  setsum id D"
-proof -
-  from F_Un_D_eq_A have "setsum id A = setsum id (D \<union> F)"
-    by (simp add: Un_commute)
-  also from F_D_disj finite_D finite_F
-  have "... = setsum id D + setsum id F"
-    by (auto simp add: Int_commute intro: setsum_Un_disjoint)
-  also from F_def have "F = (%x. (p - x)) ` E"
-    by auto
-  also from finite_E inj_on_pminusx_E have "setsum id ((%x. (p - x)) ` E) =
-      setsum (%x. (p - x)) E"
-    by (auto simp add: setsum_reindex)
-  also from finite_E have "setsum (op - p) E = setsum (%x. p) E - setsum id E"
-    by (auto simp add: setsum_subtractf id_def)
-  also from finite_E have "setsum (%x. p) E = p * int(card E)"
-    by (intro setsum_const)
-  finally show ?thesis
-    by arith
-qed
-
-lemma QRLemma3: "(a - 1) * setsum id A =
-    p * (setsum (%x. ((x * a) div p)) A - int(card E)) + 2 * setsum id E"
-proof -
-  have "(a - 1) * setsum id A = a * setsum id A - setsum id A"
-    by (auto simp add: zdiff_zmult_distrib)
-  also note QRLemma1
-  also from QRLemma2 have "p * (\<Sum>x \<in> A. x * a div p) + setsum id D +
-     setsum id E - setsum id A =
-      p * (\<Sum>x \<in> A. x * a div p) + setsum id D +
-      setsum id E - (p * int (card E) - setsum id E + setsum id D)"
-    by auto
-  also have "... = p * (\<Sum>x \<in> A. x * a div p) -
-      p * int (card E) + 2 * setsum id E"
-    by arith
-  finally show ?thesis
-    by (auto simp only: zdiff_zmult_distrib2)
-qed
-
-lemma QRLemma4: "a \<in> zOdd ==>
-    (setsum (%x. ((x * a) div p)) A \<in> zEven) = (int(card E): zEven)"
-proof -
-  assume a_odd: "a \<in> zOdd"
-  from QRLemma3 have a: "p * (setsum (%x. ((x * a) div p)) A - int(card E)) =
-      (a - 1) * setsum id A - 2 * setsum id E"
-    by arith
-  from a_odd have "a - 1 \<in> zEven"
-    by (rule odd_minus_one_even)
-  hence "(a - 1) * setsum id A \<in> zEven"
-    by (rule even_times_either)
-  moreover have "2 * setsum id E \<in> zEven"
-    by (auto simp add: zEven_def)
-  ultimately have "(a - 1) * setsum id A - 2 * setsum id E \<in> zEven"
-    by (rule even_minus_even)
-  with a have "p * (setsum (%x. ((x * a) div p)) A - int(card E)): zEven"
-    by simp
-  hence "p \<in> zEven | (setsum (%x. ((x * a) div p)) A - int(card E)): zEven"
-    by (rule EvenOdd.even_product)
-  with p_odd have "(setsum (%x. ((x * a) div p)) A - int(card E)): zEven"
-    by (auto simp add: odd_iff_not_even)
-  thus ?thesis
-    by (auto simp only: even_diff [symmetric])
-qed
-
-lemma QRLemma5: "a \<in> zOdd ==>
-   (-1::int)^(card E) = (-1::int)^(nat(setsum (%x. ((x * a) div p)) A))"
-proof -
-  assume "a \<in> zOdd"
-  from QRLemma4 [OF this] have
-    "(int(card E): zEven) = (setsum (%x. ((x * a) div p)) A \<in> zEven)" ..
-  moreover have "0 \<le> int(card E)"
-    by auto
-  moreover have "0 \<le> setsum (%x. ((x * a) div p)) A"
-    proof (intro setsum_nonneg)
-      show "\<forall>x \<in> A. 0 \<le> x * a div p"
-      proof
-        fix x
-        assume "x \<in> A"
-        then have "0 \<le> x"
-          by (auto simp add: A_def)
-        with a_nonzero have "0 \<le> x * a"
-          by (auto simp add: zero_le_mult_iff)
-        with p_g_2 show "0 \<le> x * a div p"
-          by (auto simp add: pos_imp_zdiv_nonneg_iff)
-      qed
-    qed
-  ultimately have "(-1::int)^nat((int (card E))) =
-      (-1)^nat(((\<Sum>x \<in> A. x * a div p)))"
-    by (intro neg_one_power_parity, auto)
-  also have "nat (int(card E)) = card E"
-    by auto
-  finally show ?thesis .
-qed
-
-end
-
-lemma MainQRLemma: "[| a \<in> zOdd; 0 < a; ~([a = 0] (mod p)); zprime p; 2 < p;
-  A = {x. 0 < x & x \<le> (p - 1) div 2} |] ==>
-  (Legendre a p) = (-1::int)^(nat(setsum (%x. ((x * a) div p)) A))"
-  apply (subst GAUSS.gauss_lemma)
-  apply (auto simp add: GAUSS_def)
-  apply (subst GAUSS.QRLemma5)
-  apply (auto simp add: GAUSS_def)
-  apply (simp add: GAUSS.A_def [OF GAUSS.intro] GAUSS_def)
-  done
-
-
-subsection {* Stuff about S, S1 and S2 *}
-
-locale QRTEMP =
-  fixes p     :: "int"
-  fixes q     :: "int"
-
-  assumes p_prime: "zprime p"
-  assumes p_g_2: "2 < p"
-  assumes q_prime: "zprime q"
-  assumes q_g_2: "2 < q"
-  assumes p_neq_q:      "p \<noteq> q"
-begin
-
-definition
-  P_set :: "int set" where
-  "P_set = {x. 0 < x & x \<le> ((p - 1) div 2) }"
-
-definition
-  Q_set :: "int set" where
-  "Q_set = {x. 0 < x & x \<le> ((q - 1) div 2) }"
-  
-definition
-  S :: "(int * int) set" where
-  "S = P_set <*> Q_set"
-
-definition
-  S1 :: "(int * int) set" where
-  "S1 = { (x, y). (x, y):S & ((p * y) < (q * x)) }"
-
-definition
-  S2 :: "(int * int) set" where
-  "S2 = { (x, y). (x, y):S & ((q * x) < (p * y)) }"
-
-definition
-  f1 :: "int => (int * int) set" where
-  "f1 j = { (j1, y). (j1, y):S & j1 = j & (y \<le> (q * j) div p) }"
-
-definition
-  f2 :: "int => (int * int) set" where
-  "f2 j = { (x, j1). (x, j1):S & j1 = j & (x \<le> (p * j) div q) }"
-
-lemma p_fact: "0 < (p - 1) div 2"
-proof -
-  from p_g_2 have "2 \<le> p - 1" by arith
-  then have "2 div 2 \<le> (p - 1) div 2" by (rule zdiv_mono1, auto)
-  then show ?thesis by auto
-qed
-
-lemma q_fact: "0 < (q - 1) div 2"
-proof -
-  from q_g_2 have "2 \<le> q - 1" by arith
-  then have "2 div 2 \<le> (q - 1) div 2" by (rule zdiv_mono1, auto)
-  then show ?thesis by auto
-qed
-
-lemma pb_neq_qa: "[|1 \<le> b; b \<le> (q - 1) div 2 |] ==>
-    (p * b \<noteq> q * a)"
-proof
-  assume "p * b = q * a" and "1 \<le> b" and "b \<le> (q - 1) div 2"
-  then have "q dvd (p * b)" by (auto simp add: dvd_def)
-  with q_prime p_g_2 have "q dvd p | q dvd b"
-    by (auto simp add: zprime_zdvd_zmult)
-  moreover have "~ (q dvd p)"
-  proof
-    assume "q dvd p"
-    with p_prime have "q = 1 | q = p"
-      apply (auto simp add: zprime_def QRTEMP_def)
-      apply (drule_tac x = q and R = False in allE)
-      apply (simp add: QRTEMP_def)
-      apply (subgoal_tac "0 \<le> q", simp add: QRTEMP_def)
-      apply (insert prems)
-      apply (auto simp add: QRTEMP_def)
-      done
-    with q_g_2 p_neq_q show False by auto
-  qed
-  ultimately have "q dvd b" by auto
-  then have "q \<le> b"
-  proof -
-    assume "q dvd b"
-    moreover from prems have "0 < b" by auto
-    ultimately show ?thesis using zdvd_bounds [of q b] by auto
-  qed
-  with prems have "q \<le> (q - 1) div 2" by auto
-  then have "2 * q \<le> 2 * ((q - 1) div 2)" by arith
-  then have "2 * q \<le> q - 1"
-  proof -
-    assume "2 * q \<le> 2 * ((q - 1) div 2)"
-    with prems have "q \<in> zOdd" by (auto simp add: QRTEMP_def zprime_zOdd_eq_grt_2)
-    with odd_minus_one_even have "(q - 1):zEven" by auto
-    with even_div_2_prop2 have "(q - 1) = 2 * ((q - 1) div 2)" by auto
-    with prems show ?thesis by auto
-  qed
-  then have p1: "q \<le> -1" by arith
-  with q_g_2 show False by auto
-qed
-
-lemma P_set_finite: "finite (P_set)"
-  using p_fact by (auto simp add: P_set_def bdd_int_set_l_le_finite)
-
-lemma Q_set_finite: "finite (Q_set)"
-  using q_fact by (auto simp add: Q_set_def bdd_int_set_l_le_finite)
-
-lemma S_finite: "finite S"
-  by (auto simp add: S_def  P_set_finite Q_set_finite finite_cartesian_product)
-
-lemma S1_finite: "finite S1"
-proof -
-  have "finite S" by (auto simp add: S_finite)
-  moreover have "S1 \<subseteq> S" by (auto simp add: S1_def S_def)
-  ultimately show ?thesis by (auto simp add: finite_subset)
-qed
-
-lemma S2_finite: "finite S2"
-proof -
-  have "finite S" by (auto simp add: S_finite)
-  moreover have "S2 \<subseteq> S" by (auto simp add: S2_def S_def)
-  ultimately show ?thesis by (auto simp add: finite_subset)
-qed
-
-lemma P_set_card: "(p - 1) div 2 = int (card (P_set))"
-  using p_fact by (auto simp add: P_set_def card_bdd_int_set_l_le)
-
-lemma Q_set_card: "(q - 1) div 2 = int (card (Q_set))"
-  using q_fact by (auto simp add: Q_set_def card_bdd_int_set_l_le)
-
-lemma S_card: "((p - 1) div 2) * ((q - 1) div 2) = int (card(S))"
-  using P_set_card Q_set_card P_set_finite Q_set_finite
-  by (auto simp add: S_def zmult_int setsum_constant)
-
-lemma S1_Int_S2_prop: "S1 \<inter> S2 = {}"
-  by (auto simp add: S1_def S2_def)
-
-lemma S1_Union_S2_prop: "S = S1 \<union> S2"
-  apply (auto simp add: S_def P_set_def Q_set_def S1_def S2_def)
-proof -
-  fix a and b
-  assume "~ q * a < p * b" and b1: "0 < b" and b2: "b \<le> (q - 1) div 2"
-  with zless_linear have "(p * b < q * a) | (p * b = q * a)" by auto
-  moreover from pb_neq_qa b1 b2 have "(p * b \<noteq> q * a)" by auto
-  ultimately show "p * b < q * a" by auto
-qed
-
-lemma card_sum_S1_S2: "((p - 1) div 2) * ((q - 1) div 2) =
-    int(card(S1)) + int(card(S2))"
-proof -
-  have "((p - 1) div 2) * ((q - 1) div 2) = int (card(S))"
-    by (auto simp add: S_card)
-  also have "... = int( card(S1) + card(S2))"
-    apply (insert S1_finite S2_finite S1_Int_S2_prop S1_Union_S2_prop)
-    apply (drule card_Un_disjoint, auto)
-    done
-  also have "... = int(card(S1)) + int(card(S2))" by auto
-  finally show ?thesis .
-qed
-
-lemma aux1a: "[| 0 < a; a \<le> (p - 1) div 2;
-                             0 < b; b \<le> (q - 1) div 2 |] ==>
-                          (p * b < q * a) = (b \<le> q * a div p)"
-proof -
-  assume "0 < a" and "a \<le> (p - 1) div 2" and "0 < b" and "b \<le> (q - 1) div 2"
-  have "p * b < q * a ==> b \<le> q * a div p"
-  proof -
-    assume "p * b < q * a"
-    then have "p * b \<le> q * a" by auto
-    then have "(p * b) div p \<le> (q * a) div p"
-      by (rule zdiv_mono1) (insert p_g_2, auto)
-    then show "b \<le> (q * a) div p"
-      apply (subgoal_tac "p \<noteq> 0")
-      apply (frule div_mult_self1_is_id, force)
-      apply (insert p_g_2, auto)
-      done
-  qed
-  moreover have "b \<le> q * a div p ==> p * b < q * a"
-  proof -
-    assume "b \<le> q * a div p"
-    then have "p * b \<le> p * ((q * a) div p)"
-      using p_g_2 by (auto simp add: mult_le_cancel_left)
-    also have "... \<le> q * a"
-      by (rule zdiv_leq_prop) (insert p_g_2, auto)
-    finally have "p * b \<le> q * a" .
-    then have "p * b < q * a | p * b = q * a"
-      by (simp only: order_le_imp_less_or_eq)
-    moreover have "p * b \<noteq> q * a"
-      by (rule  pb_neq_qa) (insert prems, auto)
-    ultimately show ?thesis by auto
-  qed
-  ultimately show ?thesis ..
-qed
-
-lemma aux1b: "[| 0 < a; a \<le> (p - 1) div 2;
-                             0 < b; b \<le> (q - 1) div 2 |] ==>
-                          (q * a < p * b) = (a \<le> p * b div q)"
-proof -
-  assume "0 < a" and "a \<le> (p - 1) div 2" and "0 < b" and "b \<le> (q - 1) div 2"
-  have "q * a < p * b ==> a \<le> p * b div q"
-  proof -
-    assume "q * a < p * b"
-    then have "q * a \<le> p * b" by auto
-    then have "(q * a) div q \<le> (p * b) div q"
-      by (rule zdiv_mono1) (insert q_g_2, auto)
-    then show "a \<le> (p * b) div q"
-      apply (subgoal_tac "q \<noteq> 0")
-      apply (frule div_mult_self1_is_id, force)
-      apply (insert q_g_2, auto)
-      done
-  qed
-  moreover have "a \<le> p * b div q ==> q * a < p * b"
-  proof -
-    assume "a \<le> p * b div q"
-    then have "q * a \<le> q * ((p * b) div q)"
-      using q_g_2 by (auto simp add: mult_le_cancel_left)
-    also have "... \<le> p * b"
-      by (rule zdiv_leq_prop) (insert q_g_2, auto)
-    finally have "q * a \<le> p * b" .
-    then have "q * a < p * b | q * a = p * b"
-      by (simp only: order_le_imp_less_or_eq)
-    moreover have "p * b \<noteq> q * a"
-      by (rule  pb_neq_qa) (insert prems, auto)
-    ultimately show ?thesis by auto
-  qed
-  ultimately show ?thesis ..
-qed
-
-lemma (in -) aux2: "[| zprime p; zprime q; 2 < p; 2 < q |] ==>
-             (q * ((p - 1) div 2)) div p \<le> (q - 1) div 2"
-proof-
-  assume "zprime p" and "zprime q" and "2 < p" and "2 < q"
-  (* Set up what's even and odd *)
-  then have "p \<in> zOdd & q \<in> zOdd"
-    by (auto simp add:  zprime_zOdd_eq_grt_2)
-  then have even1: "(p - 1):zEven & (q - 1):zEven"
-    by (auto simp add: odd_minus_one_even)
-  then have even2: "(2 * p):zEven & ((q - 1) * p):zEven"
-    by (auto simp add: zEven_def)
-  then have even3: "(((q - 1) * p) + (2 * p)):zEven"
-    by (auto simp: EvenOdd.even_plus_even)
-  (* using these prove it *)
-  from prems have "q * (p - 1) < ((q - 1) * p) + (2 * p)"
-    by (auto simp add: int_distrib)
-  then have "((p - 1) * q) div 2 < (((q - 1) * p) + (2 * p)) div 2"
-    apply (rule_tac x = "((p - 1) * q)" in even_div_2_l)
-    by (auto simp add: even3, auto simp add: zmult_ac)
-  also have "((p - 1) * q) div 2 = q * ((p - 1) div 2)"
-    by (auto simp add: even1 even_prod_div_2)
-  also have "(((q - 1) * p) + (2 * p)) div 2 = (((q - 1) div 2) * p) + p"
-    by (auto simp add: even1 even2 even_prod_div_2 even_sum_div_2)
-  finally show ?thesis
-    apply (rule_tac x = " q * ((p - 1) div 2)" and
-                    y = "(q - 1) div 2" in div_prop2)
-    using prems by auto
-qed
-
-lemma aux3a: "\<forall>j \<in> P_set. int (card (f1 j)) = (q * j) div p"
-proof
-  fix j
-  assume j_fact: "j \<in> P_set"
-  have "int (card (f1 j)) = int (card {y. y \<in> Q_set & y \<le> (q * j) div p})"
-  proof -
-    have "finite (f1 j)"
-    proof -
-      have "(f1 j) \<subseteq> S" by (auto simp add: f1_def)
-      with S_finite show ?thesis by (auto simp add: finite_subset)
-    qed
-    moreover have "inj_on (%(x,y). y) (f1 j)"
-      by (auto simp add: f1_def inj_on_def)
-    ultimately have "card ((%(x,y). y) ` (f1 j)) = card  (f1 j)"
-      by (auto simp add: f1_def card_image)
-    moreover have "((%(x,y). y) ` (f1 j)) = {y. y \<in> Q_set & y \<le> (q * j) div p}"
-      using prems by (auto simp add: f1_def S_def Q_set_def P_set_def image_def)
-    ultimately show ?thesis by (auto simp add: f1_def)
-  qed
-  also have "... = int (card {y. 0 < y & y \<le> (q * j) div p})"
-  proof -
-    have "{y. y \<in> Q_set & y \<le> (q * j) div p} =
-        {y. 0 < y & y \<le> (q * j) div p}"
-      apply (auto simp add: Q_set_def)
-    proof -
-      fix x
-      assume "0 < x" and "x \<le> q * j div p"
-      with j_fact P_set_def  have "j \<le> (p - 1) div 2" by auto
-      with q_g_2 have "q * j \<le> q * ((p - 1) div 2)"
-        by (auto simp add: mult_le_cancel_left)
-      with p_g_2 have "q * j div p \<le> q * ((p - 1) div 2) div p"
-        by (auto simp add: zdiv_mono1)
-      also from prems P_set_def have "... \<le> (q - 1) div 2"
-        apply simp
-        apply (insert aux2)
-        apply (simp add: QRTEMP_def)
-        done
-      finally show "x \<le> (q - 1) div 2" using prems by auto
-    qed
-    then show ?thesis by auto
-  qed
-  also have "... = (q * j) div p"
-  proof -
-    from j_fact P_set_def have "0 \<le> j" by auto
-    with q_g_2 have "q * 0 \<le> q * j" by (auto simp only: mult_left_mono)
-    then have "0 \<le> q * j" by auto
-    then have "0 div p \<le> (q * j) div p"
-      apply (rule_tac a = 0 in zdiv_mono1)
-      apply (insert p_g_2, auto)
-      done
-    also have "0 div p = 0" by auto
-    finally show ?thesis by (auto simp add: card_bdd_int_set_l_le)
-  qed
-  finally show "int (card (f1 j)) = q * j div p" .
-qed
-
-lemma aux3b: "\<forall>j \<in> Q_set. int (card (f2 j)) = (p * j) div q"
-proof
-  fix j
-  assume j_fact: "j \<in> Q_set"
-  have "int (card (f2 j)) = int (card {y. y \<in> P_set & y \<le> (p * j) div q})"
-  proof -
-    have "finite (f2 j)"
-    proof -
-      have "(f2 j) \<subseteq> S" by (auto simp add: f2_def)
-      with S_finite show ?thesis by (auto simp add: finite_subset)
-    qed
-    moreover have "inj_on (%(x,y). x) (f2 j)"
-      by (auto simp add: f2_def inj_on_def)
-    ultimately have "card ((%(x,y). x) ` (f2 j)) = card  (f2 j)"
-      by (auto simp add: f2_def card_image)
-    moreover have "((%(x,y). x) ` (f2 j)) = {y. y \<in> P_set & y \<le> (p * j) div q}"
-      using prems by (auto simp add: f2_def S_def Q_set_def P_set_def image_def)
-    ultimately show ?thesis by (auto simp add: f2_def)
-  qed
-  also have "... = int (card {y. 0 < y & y \<le> (p * j) div q})"
-  proof -
-    have "{y. y \<in> P_set & y \<le> (p * j) div q} =
-        {y. 0 < y & y \<le> (p * j) div q}"
-      apply (auto simp add: P_set_def)
-    proof -
-      fix x
-      assume "0 < x" and "x \<le> p * j div q"
-      with j_fact Q_set_def  have "j \<le> (q - 1) div 2" by auto
-      with p_g_2 have "p * j \<le> p * ((q - 1) div 2)"
-        by (auto simp add: mult_le_cancel_left)
-      with q_g_2 have "p * j div q \<le> p * ((q - 1) div 2) div q"
-        by (auto simp add: zdiv_mono1)
-      also from prems have "... \<le> (p - 1) div 2"
-        by (auto simp add: aux2 QRTEMP_def)
-      finally show "x \<le> (p - 1) div 2" using prems by auto
-      qed
-    then show ?thesis by auto
-  qed
-  also have "... = (p * j) div q"
-  proof -
-    from j_fact Q_set_def have "0 \<le> j" by auto
-    with p_g_2 have "p * 0 \<le> p * j" by (auto simp only: mult_left_mono)
-    then have "0 \<le> p * j" by auto
-    then have "0 div q \<le> (p * j) div q"
-      apply (rule_tac a = 0 in zdiv_mono1)
-      apply (insert q_g_2, auto)
-      done
-    also have "0 div q = 0" by auto
-    finally show ?thesis by (auto simp add: card_bdd_int_set_l_le)
-  qed
-  finally show "int (card (f2 j)) = p * j div q" .
-qed
-
-lemma S1_card: "int (card(S1)) = setsum (%j. (q * j) div p) P_set"
-proof -
-  have "\<forall>x \<in> P_set. finite (f1 x)"
-  proof
-    fix x
-    have "f1 x \<subseteq> S" by (auto simp add: f1_def)
-    with S_finite show "finite (f1 x)" by (auto simp add: finite_subset)
-  qed
-  moreover have "(\<forall>x \<in> P_set. \<forall>y \<in> P_set. x \<noteq> y --> (f1 x) \<inter> (f1 y) = {})"
-    by (auto simp add: f1_def)
-  moreover note P_set_finite
-  ultimately have "int(card (UNION P_set f1)) =
-      setsum (%x. int(card (f1 x))) P_set"
-    by(simp add:card_UN_disjoint int_setsum o_def)
-  moreover have "S1 = UNION P_set f1"
-    by (auto simp add: f1_def S_def S1_def S2_def P_set_def Q_set_def aux1a)
-  ultimately have "int(card (S1)) = setsum (%j. int(card (f1 j))) P_set"
-    by auto
-  also have "... = setsum (%j. q * j div p) P_set"
-    using aux3a by(fastsimp intro: setsum_cong)
-  finally show ?thesis .
-qed
-
-lemma S2_card: "int (card(S2)) = setsum (%j. (p * j) div q) Q_set"
-proof -
-  have "\<forall>x \<in> Q_set. finite (f2 x)"
-  proof
-    fix x
-    have "f2 x \<subseteq> S" by (auto simp add: f2_def)
-    with S_finite show "finite (f2 x)" by (auto simp add: finite_subset)
-  qed
-  moreover have "(\<forall>x \<in> Q_set. \<forall>y \<in> Q_set. x \<noteq> y -->
-      (f2 x) \<inter> (f2 y) = {})"
-    by (auto simp add: f2_def)
-  moreover note Q_set_finite
-  ultimately have "int(card (UNION Q_set f2)) =
-      setsum (%x. int(card (f2 x))) Q_set"
-    by(simp add:card_UN_disjoint int_setsum o_def)
-  moreover have "S2 = UNION Q_set f2"
-    by (auto simp add: f2_def S_def S1_def S2_def P_set_def Q_set_def aux1b)
-  ultimately have "int(card (S2)) = setsum (%j. int(card (f2 j))) Q_set"
-    by auto
-  also have "... = setsum (%j. p * j div q) Q_set"
-    using aux3b by(fastsimp intro: setsum_cong)
-  finally show ?thesis .
-qed
-
-lemma S1_carda: "int (card(S1)) =
-    setsum (%j. (j * q) div p) P_set"
-  by (auto simp add: S1_card zmult_ac)
-
-lemma S2_carda: "int (card(S2)) =
-    setsum (%j. (j * p) div q) Q_set"
-  by (auto simp add: S2_card zmult_ac)
-
-lemma pq_sum_prop: "(setsum (%j. (j * p) div q) Q_set) +
-    (setsum (%j. (j * q) div p) P_set) = ((p - 1) div 2) * ((q - 1) div 2)"
-proof -
-  have "(setsum (%j. (j * p) div q) Q_set) +
-      (setsum (%j. (j * q) div p) P_set) = int (card S2) + int (card S1)"
-    by (auto simp add: S1_carda S2_carda)
-  also have "... = int (card S1) + int (card S2)"
-    by auto
-  also have "... = ((p - 1) div 2) * ((q - 1) div 2)"
-    by (auto simp add: card_sum_S1_S2)
-  finally show ?thesis .
-qed
-
-
-lemma (in -) pq_prime_neq: "[| zprime p; zprime q; p \<noteq> q |] ==> (~[p = 0] (mod q))"
-  apply (auto simp add: zcong_eq_zdvd_prop zprime_def)
-  apply (drule_tac x = q in allE)
-  apply (drule_tac x = p in allE)
-  apply auto
-  done
-
-
-lemma QR_short: "(Legendre p q) * (Legendre q p) =
-    (-1::int)^nat(((p - 1) div 2)*((q - 1) div 2))"
-proof -
-  from prems have "~([p = 0] (mod q))"
-    by (auto simp add: pq_prime_neq QRTEMP_def)
-  with prems Q_set_def have a1: "(Legendre p q) = (-1::int) ^
-      nat(setsum (%x. ((x * p) div q)) Q_set)"
-    apply (rule_tac p = q in  MainQRLemma)
-    apply (auto simp add: zprime_zOdd_eq_grt_2 QRTEMP_def)
-    done
-  from prems have "~([q = 0] (mod p))"
-    apply (rule_tac p = q and q = p in pq_prime_neq)
-    apply (simp add: QRTEMP_def)+
-    done
-  with prems P_set_def have a2: "(Legendre q p) =
-      (-1::int) ^ nat(setsum (%x. ((x * q) div p)) P_set)"
-    apply (rule_tac p = p in  MainQRLemma)
-    apply (auto simp add: zprime_zOdd_eq_grt_2 QRTEMP_def)
-    done
-  from a1 a2 have "(Legendre p q) * (Legendre q p) =
-      (-1::int) ^ nat(setsum (%x. ((x * p) div q)) Q_set) *
-        (-1::int) ^ nat(setsum (%x. ((x * q) div p)) P_set)"
-    by auto
-  also have "... = (-1::int) ^ (nat(setsum (%x. ((x * p) div q)) Q_set) +
-                   nat(setsum (%x. ((x * q) div p)) P_set))"
-    by (auto simp add: zpower_zadd_distrib)
-  also have "nat(setsum (%x. ((x * p) div q)) Q_set) +
-      nat(setsum (%x. ((x * q) div p)) P_set) =
-        nat((setsum (%x. ((x * p) div q)) Q_set) +
-          (setsum (%x. ((x * q) div p)) P_set))"
-    apply (rule_tac z = "setsum (%x. ((x * p) div q)) Q_set" in
-      nat_add_distrib [symmetric])
-    apply (auto simp add: S1_carda [symmetric] S2_carda [symmetric])
-    done
-  also have "... = nat(((p - 1) div 2) * ((q - 1) div 2))"
-    by (auto simp add: pq_sum_prop)
-  finally show ?thesis .
-qed
-
-end
-
-theorem Quadratic_Reciprocity:
-     "[| p \<in> zOdd; zprime p; q \<in> zOdd; zprime q;
-         p \<noteq> q |]
-      ==> (Legendre p q) * (Legendre q p) =
-          (-1::int)^nat(((p - 1) div 2)*((q - 1) div 2))"
-  by (auto simp add: QRTEMP.QR_short zprime_zOdd_eq_grt_2 [symmetric]
-                     QRTEMP_def)
-
-end
--- a/src/HOL/NumberTheory/ROOT.ML	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,5 +0,0 @@
-(* $Id$ *)
-
-no_document use_thys ["Infinite_Set", "Permutation", "Primes"];
-use_thys ["Fib", "Factorization", "Chinese", "WilsonRuss",
-  "WilsonBij", "Quadratic_Reciprocity"];
--- a/src/HOL/NumberTheory/Residues.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,172 +0,0 @@
-(*  Title:      HOL/Quadratic_Reciprocity/Residues.thy
-    ID:         $Id$
-    Authors:    Jeremy Avigad, David Gray, and Adam Kramer
-*)
-
-header {* Residue Sets *}
-
-theory Residues imports Int2 begin
-
-text {*
-  \medskip Define the residue of a set, the standard residue,
-  quadratic residues, and prove some basic properties. *}
-
-definition
-  ResSet      :: "int => int set => bool" where
-  "ResSet m X = (\<forall>y1 y2. (y1 \<in> X & y2 \<in> X & [y1 = y2] (mod m) --> y1 = y2))"
-
-definition
-  StandardRes :: "int => int => int" where
-  "StandardRes m x = x mod m"
-
-definition
-  QuadRes     :: "int => int => bool" where
-  "QuadRes m x = (\<exists>y. ([(y ^ 2) = x] (mod m)))"
-
-definition
-  Legendre    :: "int => int => int" where
-  "Legendre a p = (if ([a = 0] (mod p)) then 0
-                     else if (QuadRes p a) then 1
-                     else -1)"
-
-definition
-  SR          :: "int => int set" where
-  "SR p = {x. (0 \<le> x) & (x < p)}"
-
-definition
-  SRStar      :: "int => int set" where
-  "SRStar p = {x. (0 < x) & (x < p)}"
-
-
-subsection {* Some useful properties of StandardRes *}
-
-lemma StandardRes_prop1: "[x = StandardRes m x] (mod m)"
-  by (auto simp add: StandardRes_def zcong_zmod)
-
-lemma StandardRes_prop2: "0 < m ==> (StandardRes m x1 = StandardRes m x2)
-      = ([x1 = x2] (mod m))"
-  by (auto simp add: StandardRes_def zcong_zmod_eq)
-
-lemma StandardRes_prop3: "(~[x = 0] (mod p)) = (~(StandardRes p x = 0))"
-  by (auto simp add: StandardRes_def zcong_def dvd_eq_mod_eq_0)
-
-lemma StandardRes_prop4: "2 < m 
-     ==> [StandardRes m x * StandardRes m y = (x * y)] (mod m)"
-  by (auto simp add: StandardRes_def zcong_zmod_eq 
-                     mod_mult_eq [of x y m])
-
-lemma StandardRes_lbound: "0 < p ==> 0 \<le> StandardRes p x"
-  by (auto simp add: StandardRes_def pos_mod_sign)
-
-lemma StandardRes_ubound: "0 < p ==> StandardRes p x < p"
-  by (auto simp add: StandardRes_def pos_mod_bound)
-
-lemma StandardRes_eq_zcong: 
-   "(StandardRes m x = 0) = ([x = 0](mod m))"
-  by (auto simp add: StandardRes_def zcong_eq_zdvd_prop dvd_def) 
-
-
-subsection {* Relations between StandardRes, SRStar, and SR *}
-
-lemma SRStar_SR_prop: "x \<in> SRStar p ==> x \<in> SR p"
-  by (auto simp add: SRStar_def SR_def)
-
-lemma StandardRes_SR_prop: "x \<in> SR p ==> StandardRes p x = x"
-  by (auto simp add: SR_def StandardRes_def mod_pos_pos_trivial)
-
-lemma StandardRes_SRStar_prop1: "2 < p ==> (StandardRes p x \<in> SRStar p) 
-     = (~[x = 0] (mod p))"
-  apply (auto simp add: StandardRes_prop3 StandardRes_def
-                        SRStar_def pos_mod_bound)
-  apply (subgoal_tac "0 < p")
-  apply (drule_tac a = x in pos_mod_sign, arith, simp)
-  done
-
-lemma StandardRes_SRStar_prop1a: "x \<in> SRStar p ==> ~([x = 0] (mod p))"
-  by (auto simp add: SRStar_def zcong_def zdvd_not_zless)
-
-lemma StandardRes_SRStar_prop2: "[| 2 < p; zprime p; x \<in> SRStar p |] 
-     ==> StandardRes p (MultInv p x) \<in> SRStar p"
-  apply (frule_tac x = "(MultInv p x)" in StandardRes_SRStar_prop1, simp)
-  apply (rule MultInv_prop3)
-  apply (auto simp add: SRStar_def zcong_def zdvd_not_zless)
-  done
-
-lemma StandardRes_SRStar_prop3: "x \<in> SRStar p ==> StandardRes p x = x"
-  by (auto simp add: SRStar_SR_prop StandardRes_SR_prop)
-
-lemma StandardRes_SRStar_prop4: "[| zprime p; 2 < p; x \<in> SRStar p |] 
-     ==> StandardRes p x \<in> SRStar p"
-  by (frule StandardRes_SRStar_prop3, auto)
-
-lemma SRStar_mult_prop1: "[| zprime p; 2 < p; x \<in> SRStar p; y \<in> SRStar p|] 
-     ==> (StandardRes p (x * y)):SRStar p"
-  apply (frule_tac x = x in StandardRes_SRStar_prop4, auto)
-  apply (frule_tac x = y in StandardRes_SRStar_prop4, auto)
-  apply (auto simp add: StandardRes_SRStar_prop1 zcong_zmult_prop3)
-  done
-
-lemma SRStar_mult_prop2: "[| zprime p; 2 < p; ~([a = 0](mod p)); 
-     x \<in> SRStar p |] 
-     ==> StandardRes p (a * MultInv p x) \<in> SRStar p"
-  apply (frule_tac x = x in StandardRes_SRStar_prop2, auto)
-  apply (frule_tac x = "MultInv p x" in StandardRes_SRStar_prop1)
-  apply (auto simp add: StandardRes_SRStar_prop1 zcong_zmult_prop3)
-  done
-
-lemma SRStar_card: "2 < p ==> int(card(SRStar p)) = p - 1"
-  by (auto simp add: SRStar_def int_card_bdd_int_set_l_l)
-
-lemma SRStar_finite: "2 < p ==> finite( SRStar p)"
-  by (auto simp add: SRStar_def bdd_int_set_l_l_finite)
-
-
-subsection {* Properties relating ResSets with StandardRes *}
-
-lemma aux: "x mod m = y mod m ==> [x = y] (mod m)"
-  apply (subgoal_tac "x = y ==> [x = y](mod m)")
-  apply (subgoal_tac "[x mod m = y mod m] (mod m) ==> [x = y] (mod m)")
-  apply (auto simp add: zcong_zmod [of x y m])
-  done
-
-lemma StandardRes_inj_on_ResSet: "ResSet m X ==> (inj_on (StandardRes m) X)"
-  apply (auto simp add: ResSet_def StandardRes_def inj_on_def)
-  apply (drule_tac m = m in aux, auto)
-  done
-
-lemma StandardRes_Sum: "[| finite X; 0 < m |] 
-     ==> [setsum f X = setsum (StandardRes m o f) X](mod m)" 
-  apply (rule_tac F = X in finite_induct)
-  apply (auto intro!: zcong_zadd simp add: StandardRes_prop1)
-  done
-
-lemma SR_pos: "0 < m ==> (StandardRes m ` X) \<subseteq> {x. 0 \<le> x & x < m}"
-  by (auto simp add: StandardRes_ubound StandardRes_lbound)
-
-lemma ResSet_finite: "0 < m ==> ResSet m X ==> finite X"
-  apply (rule_tac f = "StandardRes m" in finite_imageD) 
-  apply (rule_tac B = "{x. (0 :: int) \<le> x & x < m}" in finite_subset)
-  apply (auto simp add: StandardRes_inj_on_ResSet bdd_int_set_l_finite SR_pos)
-  done
-
-lemma mod_mod_is_mod: "[x = x mod m](mod m)"
-  by (auto simp add: zcong_zmod)
-
-lemma StandardRes_prod: "[| finite X; 0 < m |] 
-     ==> [setprod f X = setprod (StandardRes m o f) X] (mod m)"
-  apply (rule_tac F = X in finite_induct)
-  apply (auto intro!: zcong_zmult simp add: StandardRes_prop1)
-  done
-
-lemma ResSet_image:
-  "[| 0 < m; ResSet m A; \<forall>x \<in> A. \<forall>y \<in> A. ([f x = f y](mod m) --> x = y) |] ==>
-    ResSet m (f ` A)"
-  by (auto simp add: ResSet_def)
-
-
-subsection {* Property for SRStar *}
-
-lemma ResSet_SRStar_prop: "ResSet p (SRStar p)"
-  by (auto simp add: SRStar_def ResSet_def zcong_zless_imp_eq)
-
-end
--- a/src/HOL/NumberTheory/WilsonBij.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,263 +0,0 @@
-(*  Title:      HOL/NumberTheory/WilsonBij.thy
-    ID:         $Id$
-    Author:     Thomas M. Rasmussen
-    Copyright   2000  University of Cambridge
-*)
-
-header {* Wilson's Theorem using a more abstract approach *}
-
-theory WilsonBij imports BijectionRel IntFact begin
-
-text {*
-  Wilson's Theorem using a more ``abstract'' approach based on
-  bijections between sets.  Does not use Fermat's Little Theorem
-  (unlike Russinoff).
-*}
-
-
-subsection {* Definitions and lemmas *}
-
-definition
-  reciR :: "int => int => int => bool" where
-  "reciR p = (\<lambda>a b. zcong (a * b) 1 p \<and> 1 < a \<and> a < p - 1 \<and> 1 < b \<and> b < p - 1)"
-
-definition
-  inv :: "int => int => int" where
-  "inv p a =
-    (if zprime p \<and> 0 < a \<and> a < p then
-      (SOME x. 0 \<le> x \<and> x < p \<and> zcong (a * x) 1 p)
-     else 0)"
-
-
-text {* \medskip Inverse *}
-
-lemma inv_correct:
-  "zprime p ==> 0 < a ==> a < p
-    ==> 0 \<le> inv p a \<and> inv p a < p \<and> [a * inv p a = 1] (mod p)"
-  apply (unfold inv_def)
-  apply (simp (no_asm_simp))
-  apply (rule zcong_lineq_unique [THEN ex1_implies_ex, THEN someI_ex])
-   apply (erule_tac [2] zless_zprime_imp_zrelprime)
-    apply (unfold zprime_def)
-    apply auto
-  done
-
-lemmas inv_ge = inv_correct [THEN conjunct1, standard]
-lemmas inv_less = inv_correct [THEN conjunct2, THEN conjunct1, standard]
-lemmas inv_is_inv = inv_correct [THEN conjunct2, THEN conjunct2, standard]
-
-lemma inv_not_0:
-  "zprime p ==> 1 < a ==> a < p - 1 ==> inv p a \<noteq> 0"
-  -- {* same as @{text WilsonRuss} *}
-  apply safe
-  apply (cut_tac a = a and p = p in inv_is_inv)
-     apply (unfold zcong_def)
-     apply auto
-  apply (subgoal_tac "\<not> p dvd 1")
-   apply (rule_tac [2] zdvd_not_zless)
-    apply (subgoal_tac "p dvd 1")
-     prefer 2
-     apply (subst dvd_minus_iff [symmetric])
-     apply auto
-  done
-
-lemma inv_not_1:
-  "zprime p ==> 1 < a ==> a < p - 1 ==> inv p a \<noteq> 1"
-  -- {* same as @{text WilsonRuss} *}
-  apply safe
-  apply (cut_tac a = a and p = p in inv_is_inv)
-     prefer 4
-     apply simp
-     apply (subgoal_tac "a = 1")
-      apply (rule_tac [2] zcong_zless_imp_eq)
-          apply auto
-  done
-
-lemma aux: "[a * (p - 1) = 1] (mod p) = [a = p - 1] (mod p)"
-  -- {* same as @{text WilsonRuss} *}
-  apply (unfold zcong_def)
-  apply (simp add: OrderedGroup.diff_diff_eq diff_diff_eq2 zdiff_zmult_distrib2)
-  apply (rule_tac s = "p dvd -((a + 1) + (p * -a))" in trans)
-   apply (simp add: mult_commute)
-  apply (subst dvd_minus_iff)
-  apply (subst zdvd_reduce)
-  apply (rule_tac s = "p dvd (a + 1) + (p * -1)" in trans)
-   apply (subst zdvd_reduce)
-   apply auto
-  done
-
-lemma inv_not_p_minus_1:
-  "zprime p ==> 1 < a ==> a < p - 1 ==> inv p a \<noteq> p - 1"
-  -- {* same as @{text WilsonRuss} *}
-  apply safe
-  apply (cut_tac a = a and p = p in inv_is_inv)
-     apply auto
-  apply (simp add: aux)
-  apply (subgoal_tac "a = p - 1")
-   apply (rule_tac [2] zcong_zless_imp_eq)
-       apply auto
-  done
-
-text {*
-  Below is slightly different as we don't expand @{term [source] inv}
-  but use ``@{text correct}'' theorems.
-*}
-
-lemma inv_g_1: "zprime p ==> 1 < a ==> a < p - 1 ==> 1 < inv p a"
-  apply (subgoal_tac "inv p a \<noteq> 1")
-   apply (subgoal_tac "inv p a \<noteq> 0")
-    apply (subst order_less_le)
-    apply (subst zle_add1_eq_le [symmetric])
-    apply (subst order_less_le)
-    apply (rule_tac [2] inv_not_0)
-      apply (rule_tac [5] inv_not_1)
-        apply auto
-  apply (rule inv_ge)
-    apply auto
-  done
-
-lemma inv_less_p_minus_1:
-  "zprime p ==> 1 < a ==> a < p - 1 ==> inv p a < p - 1"
-  -- {* ditto *}
-  apply (subst order_less_le)
-  apply (simp add: inv_not_p_minus_1 inv_less)
-  done
-
-
-text {* \medskip Bijection *}
-
-lemma aux1: "1 < x ==> 0 \<le> (x::int)"
-  apply auto
-  done
-
-lemma aux2: "1 < x ==> 0 < (x::int)"
-  apply auto
-  done
-
-lemma aux3: "x \<le> p - 2 ==> x < (p::int)"
-  apply auto
-  done
-
-lemma aux4: "x \<le> p - 2 ==> x < (p::int) - 1"
-  apply auto
-  done
-
-lemma inv_inj: "zprime p ==> inj_on (inv p) (d22set (p - 2))"
-  apply (unfold inj_on_def)
-  apply auto
-  apply (rule zcong_zless_imp_eq)
-      apply (tactic {* stac (thm "zcong_cancel" RS sym) 5 *})
-        apply (rule_tac [7] zcong_trans)
-         apply (tactic {* stac (thm "zcong_sym") 8 *})
-         apply (erule_tac [7] inv_is_inv)
-          apply (tactic "asm_simp_tac @{simpset} 9")
-          apply (erule_tac [9] inv_is_inv)
-           apply (rule_tac [6] zless_zprime_imp_zrelprime)
-             apply (rule_tac [8] inv_less)
-               apply (rule_tac [7] inv_g_1 [THEN aux2])
-                 apply (unfold zprime_def)
-                 apply (auto intro: d22set_g_1 d22set_le
-		   aux1 aux2 aux3 aux4)
-  done
-
-lemma inv_d22set_d22set:
-    "zprime p ==> inv p ` d22set (p - 2) = d22set (p - 2)"
-  apply (rule endo_inj_surj)
-    apply (rule d22set_fin)
-   apply (erule_tac [2] inv_inj)
-  apply auto
-  apply (rule d22set_mem)
-   apply (erule inv_g_1)
-    apply (subgoal_tac [3] "inv p xa < p - 1")
-     apply (erule_tac [4] inv_less_p_minus_1)
-      apply (auto intro: d22set_g_1 d22set_le aux4)
-  done
-
-lemma d22set_d22set_bij:
-    "zprime p ==> (d22set (p - 2), d22set (p - 2)) \<in> bijR (reciR p)"
-  apply (unfold reciR_def)
-  apply (rule_tac s = "(d22set (p - 2), inv p ` d22set (p - 2))" in subst)
-   apply (simp add: inv_d22set_d22set)
-  apply (rule inj_func_bijR)
-    apply (rule_tac [3] d22set_fin)
-   apply (erule_tac [2] inv_inj)
-  apply auto
-      apply (erule inv_is_inv)
-       apply (erule_tac [5] inv_g_1)
-        apply (erule_tac [7] inv_less_p_minus_1)
-         apply (auto intro: d22set_g_1 d22set_le aux2 aux3 aux4)
-  done
-
-lemma reciP_bijP: "zprime p ==> bijP (reciR p) (d22set (p - 2))"
-  apply (unfold reciR_def bijP_def)
-  apply auto
-  apply (rule d22set_mem)
-   apply auto
-  done
-
-lemma reciP_uniq: "zprime p ==> uniqP (reciR p)"
-  apply (unfold reciR_def uniqP_def)
-  apply auto
-   apply (rule zcong_zless_imp_eq)
-       apply (tactic {* stac (thm "zcong_cancel2" RS sym) 5 *})
-         apply (rule_tac [7] zcong_trans)
-          apply (tactic {* stac (thm "zcong_sym") 8 *})
-          apply (rule_tac [6] zless_zprime_imp_zrelprime)
-            apply auto
-  apply (rule zcong_zless_imp_eq)
-      apply (tactic {* stac (thm "zcong_cancel" RS sym) 5 *})
-        apply (rule_tac [7] zcong_trans)
-         apply (tactic {* stac (thm "zcong_sym") 8 *})
-         apply (rule_tac [6] zless_zprime_imp_zrelprime)
-           apply auto
-  done
-
-lemma reciP_sym: "zprime p ==> symP (reciR p)"
-  apply (unfold reciR_def symP_def)
-  apply (simp add: zmult_commute)
-  apply auto
-  done
-
-lemma bijER_d22set: "zprime p ==> d22set (p - 2) \<in> bijER (reciR p)"
-  apply (rule bijR_bijER)
-     apply (erule d22set_d22set_bij)
-    apply (erule reciP_bijP)
-   apply (erule reciP_uniq)
-  apply (erule reciP_sym)
-  done
-
-
-subsection {* Wilson *}
-
-lemma bijER_zcong_prod_1:
-    "zprime p ==> A \<in> bijER (reciR p) ==> [\<Prod>A = 1] (mod p)"
-  apply (unfold reciR_def)
-  apply (erule bijER.induct)
-    apply (subgoal_tac [2] "a = 1 \<or> a = p - 1")
-     apply (rule_tac [3] zcong_square_zless)
-        apply auto
-  apply (subst setprod_insert)
-    prefer 3
-    apply (subst setprod_insert)
-      apply (auto simp add: fin_bijER)
-  apply (subgoal_tac "zcong ((a * b) * \<Prod>A) (1 * 1) p")
-   apply (simp add: zmult_assoc)
-  apply (rule zcong_zmult)
-   apply auto
-  done
-
-theorem Wilson_Bij: "zprime p ==> [zfact (p - 1) = -1] (mod p)"
-  apply (subgoal_tac "zcong ((p - 1) * zfact (p - 2)) (-1 * 1) p")
-   apply (rule_tac [2] zcong_zmult)
-    apply (simp add: zprime_def)
-    apply (subst zfact.simps)
-    apply (rule_tac t = "p - 1 - 1" and s = "p - 2" in subst)
-     apply auto
-   apply (simp add: zcong_def)
-  apply (subst d22set_prod_zfact [symmetric])
-  apply (rule bijER_zcong_prod_1)
-   apply (rule_tac [2] bijER_d22set)
-   apply auto
-  done
-
-end
--- a/src/HOL/NumberTheory/WilsonRuss.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,329 +0,0 @@
-(*  Title:      HOL/NumberTheory/WilsonRuss.thy
-    ID:         $Id$
-    Author:     Thomas M. Rasmussen
-    Copyright   2000  University of Cambridge
-*)
-
-header {* Wilson's Theorem according to Russinoff *}
-
-theory WilsonRuss imports EulerFermat begin
-
-text {*
-  Wilson's Theorem following quite closely Russinoff's approach
-  using Boyer-Moore (using finite sets instead of lists, though).
-*}
-
-subsection {* Definitions and lemmas *}
-
-definition
-  inv :: "int => int => int" where
-  "inv p a = (a^(nat (p - 2))) mod p"
-
-consts
-  wset :: "int * int => int set"
-
-recdef wset
-  "measure ((\<lambda>(a, p). nat a) :: int * int => nat)"
-  "wset (a, p) =
-    (if 1 < a then
-      let ws = wset (a - 1, p)
-      in (if a \<in> ws then ws else insert a (insert (inv p a) ws)) else {})"
-
-
-text {* \medskip @{term [source] inv} *}
-
-lemma inv_is_inv_aux: "1 < m ==> Suc (nat (m - 2)) = nat (m - 1)"
-by (subst int_int_eq [symmetric], auto)
-
-lemma inv_is_inv:
-    "zprime p \<Longrightarrow> 0 < a \<Longrightarrow> a < p ==> [a * inv p a = 1] (mod p)"
-  apply (unfold inv_def)
-  apply (subst zcong_zmod)
-  apply (subst zmod_zmult1_eq [symmetric])
-  apply (subst zcong_zmod [symmetric])
-  apply (subst power_Suc [symmetric])
-  apply (subst inv_is_inv_aux)
-   apply (erule_tac [2] Little_Fermat)
-   apply (erule_tac [2] zdvd_not_zless)
-   apply (unfold zprime_def, auto)
-  done
-
-lemma inv_distinct:
-    "zprime p \<Longrightarrow> 1 < a \<Longrightarrow> a < p - 1 ==> a \<noteq> inv p a"
-  apply safe
-  apply (cut_tac a = a and p = p in zcong_square)
-     apply (cut_tac [3] a = a and p = p in inv_is_inv, auto)
-   apply (subgoal_tac "a = 1")
-    apply (rule_tac [2] m = p in zcong_zless_imp_eq)
-        apply (subgoal_tac [7] "a = p - 1")
-         apply (rule_tac [8] m = p in zcong_zless_imp_eq, auto)
-  done
-
-lemma inv_not_0:
-    "zprime p \<Longrightarrow> 1 < a \<Longrightarrow> a < p - 1 ==> inv p a \<noteq> 0"
-  apply safe
-  apply (cut_tac a = a and p = p in inv_is_inv)
-     apply (unfold zcong_def, auto)
-  apply (subgoal_tac "\<not> p dvd 1")
-   apply (rule_tac [2] zdvd_not_zless)
-    apply (subgoal_tac "p dvd 1")
-     prefer 2
-     apply (subst dvd_minus_iff [symmetric], auto)
-  done
-
-lemma inv_not_1:
-    "zprime p \<Longrightarrow> 1 < a \<Longrightarrow> a < p - 1 ==> inv p a \<noteq> 1"
-  apply safe
-  apply (cut_tac a = a and p = p in inv_is_inv)
-     prefer 4
-     apply simp
-     apply (subgoal_tac "a = 1")
-      apply (rule_tac [2] zcong_zless_imp_eq, auto)
-  done
-
-lemma inv_not_p_minus_1_aux:
-    "[a * (p - 1) = 1] (mod p) = [a = p - 1] (mod p)"
-  apply (unfold zcong_def)
-  apply (simp add: OrderedGroup.diff_diff_eq diff_diff_eq2 zdiff_zmult_distrib2)
-  apply (rule_tac s = "p dvd -((a + 1) + (p * -a))" in trans)
-   apply (simp add: mult_commute)
-  apply (subst dvd_minus_iff)
-  apply (subst zdvd_reduce)
-  apply (rule_tac s = "p dvd (a + 1) + (p * -1)" in trans)
-   apply (subst zdvd_reduce, auto)
-  done
-
-lemma inv_not_p_minus_1:
-    "zprime p \<Longrightarrow> 1 < a \<Longrightarrow> a < p - 1 ==> inv p a \<noteq> p - 1"
-  apply safe
-  apply (cut_tac a = a and p = p in inv_is_inv, auto)
-  apply (simp add: inv_not_p_minus_1_aux)
-  apply (subgoal_tac "a = p - 1")
-   apply (rule_tac [2] zcong_zless_imp_eq, auto)
-  done
-
-lemma inv_g_1:
-    "zprime p \<Longrightarrow> 1 < a \<Longrightarrow> a < p - 1 ==> 1 < inv p a"
-  apply (case_tac "0\<le> inv p a")
-   apply (subgoal_tac "inv p a \<noteq> 1")
-    apply (subgoal_tac "inv p a \<noteq> 0")
-     apply (subst order_less_le)
-     apply (subst zle_add1_eq_le [symmetric])
-     apply (subst order_less_le)
-     apply (rule_tac [2] inv_not_0)
-       apply (rule_tac [5] inv_not_1, auto)
-  apply (unfold inv_def zprime_def, simp)
-  done
-
-lemma inv_less_p_minus_1:
-    "zprime p \<Longrightarrow> 1 < a \<Longrightarrow> a < p - 1 ==> inv p a < p - 1"
-  apply (case_tac "inv p a < p")
-   apply (subst order_less_le)
-   apply (simp add: inv_not_p_minus_1, auto)
-  apply (unfold inv_def zprime_def, simp)
-  done
-
-lemma inv_inv_aux: "5 \<le> p ==>
-    nat (p - 2) * nat (p - 2) = Suc (nat (p - 1) * nat (p - 3))"
-  apply (subst int_int_eq [symmetric])
-  apply (simp add: zmult_int [symmetric])
-  apply (simp add: zdiff_zmult_distrib zdiff_zmult_distrib2)
-  done
-
-lemma zcong_zpower_zmult:
-    "[x^y = 1] (mod p) \<Longrightarrow> [x^(y * z) = 1] (mod p)"
-  apply (induct z)
-   apply (auto simp add: zpower_zadd_distrib)
-  apply (subgoal_tac "zcong (x^y * x^(y * z)) (1 * 1) p")
-   apply (rule_tac [2] zcong_zmult, simp_all)
-  done
-
-lemma inv_inv: "zprime p \<Longrightarrow>
-    5 \<le> p \<Longrightarrow> 0 < a \<Longrightarrow> a < p ==> inv p (inv p a) = a"
-  apply (unfold inv_def)
-  apply (subst zpower_zmod)
-  apply (subst zpower_zpower)
-  apply (rule zcong_zless_imp_eq)
-      prefer 5
-      apply (subst zcong_zmod)
-      apply (subst mod_mod_trivial)
-      apply (subst zcong_zmod [symmetric])
-      apply (subst inv_inv_aux)
-       apply (subgoal_tac [2]
-	 "zcong (a * a^(nat (p - 1) * nat (p - 3))) (a * 1) p")
-        apply (rule_tac [3] zcong_zmult)
-         apply (rule_tac [4] zcong_zpower_zmult)
-         apply (erule_tac [4] Little_Fermat)
-         apply (rule_tac [4] zdvd_not_zless, simp_all)
-  done
-
-
-text {* \medskip @{term wset} *}
-
-declare wset.simps [simp del]
-
-lemma wset_induct:
-  assumes "!!a p. P {} a p"
-    and "!!a p. 1 < (a::int) \<Longrightarrow>
-      P (wset (a - 1, p)) (a - 1) p ==> P (wset (a, p)) a p"
-  shows "P (wset (u, v)) u v"
-  apply (rule wset.induct, safe)
-   prefer 2
-   apply (case_tac "1 < a")
-    apply (rule prems)
-     apply simp_all
-   apply (simp_all add: wset.simps prems)
-  done
-
-lemma wset_mem_imp_or [rule_format]:
-  "1 < a \<Longrightarrow> b \<notin> wset (a - 1, p)
-    ==> b \<in> wset (a, p) --> b = a \<or> b = inv p a"
-  apply (subst wset.simps)
-  apply (unfold Let_def, simp)
-  done
-
-lemma wset_mem_mem [simp]: "1 < a ==> a \<in> wset (a, p)"
-  apply (subst wset.simps)
-  apply (unfold Let_def, simp)
-  done
-
-lemma wset_subset: "1 < a \<Longrightarrow> b \<in> wset (a - 1, p) ==> b \<in> wset (a, p)"
-  apply (subst wset.simps)
-  apply (unfold Let_def, auto)
-  done
-
-lemma wset_g_1 [rule_format]:
-    "zprime p --> a < p - 1 --> b \<in> wset (a, p) --> 1 < b"
-  apply (induct a p rule: wset_induct, auto)
-  apply (case_tac "b = a")
-   apply (case_tac [2] "b = inv p a")
-    apply (subgoal_tac [3] "b = a \<or> b = inv p a")
-     apply (rule_tac [4] wset_mem_imp_or)
-       prefer 2
-       apply simp
-       apply (rule inv_g_1, auto)
-  done
-
-lemma wset_less [rule_format]:
-    "zprime p --> a < p - 1 --> b \<in> wset (a, p) --> b < p - 1"
-  apply (induct a p rule: wset_induct, auto)
-  apply (case_tac "b = a")
-   apply (case_tac [2] "b = inv p a")
-    apply (subgoal_tac [3] "b = a \<or> b = inv p a")
-     apply (rule_tac [4] wset_mem_imp_or)
-       prefer 2
-       apply simp
-       apply (rule inv_less_p_minus_1, auto)
-  done
-
-lemma wset_mem [rule_format]:
-  "zprime p -->
-    a < p - 1 --> 1 < b --> b \<le> a --> b \<in> wset (a, p)"
-  apply (induct a p rule: wset.induct, auto)
-  apply (rule_tac wset_subset)
-  apply (simp (no_asm_simp))
-  apply auto
-  done
-
-lemma wset_mem_inv_mem [rule_format]:
-  "zprime p --> 5 \<le> p --> a < p - 1 --> b \<in> wset (a, p)
-    --> inv p b \<in> wset (a, p)"
-  apply (induct a p rule: wset_induct, auto)
-   apply (case_tac "b = a")
-    apply (subst wset.simps)
-    apply (unfold Let_def)
-    apply (rule_tac [3] wset_subset, auto)
-  apply (case_tac "b = inv p a")
-   apply (simp (no_asm_simp))
-   apply (subst inv_inv)
-       apply (subgoal_tac [6] "b = a \<or> b = inv p a")
-        apply (rule_tac [7] wset_mem_imp_or, auto)
-  done
-
-lemma wset_inv_mem_mem:
-  "zprime p \<Longrightarrow> 5 \<le> p \<Longrightarrow> a < p - 1 \<Longrightarrow> 1 < b \<Longrightarrow> b < p - 1
-    \<Longrightarrow> inv p b \<in> wset (a, p) \<Longrightarrow> b \<in> wset (a, p)"
-  apply (rule_tac s = "inv p (inv p b)" and t = b in subst)
-   apply (rule_tac [2] wset_mem_inv_mem)
-      apply (rule inv_inv, simp_all)
-  done
-
-lemma wset_fin: "finite (wset (a, p))"
-  apply (induct a p rule: wset_induct)
-   prefer 2
-   apply (subst wset.simps)
-   apply (unfold Let_def, auto)
-  done
-
-lemma wset_zcong_prod_1 [rule_format]:
-  "zprime p -->
-    5 \<le> p --> a < p - 1 --> [(\<Prod>x\<in>wset(a, p). x) = 1] (mod p)"
-  apply (induct a p rule: wset_induct)
-   prefer 2
-   apply (subst wset.simps)
-   apply (unfold Let_def, auto)
-  apply (subst setprod_insert)
-    apply (tactic {* stac (thm "setprod_insert") 3 *})
-      apply (subgoal_tac [5]
-	"zcong (a * inv p a * (\<Prod>x\<in> wset(a - 1, p). x)) (1 * 1) p")
-       prefer 5
-       apply (simp add: zmult_assoc)
-      apply (rule_tac [5] zcong_zmult)
-       apply (rule_tac [5] inv_is_inv)
-         apply (tactic "clarify_tac @{claset} 4")
-         apply (subgoal_tac [4] "a \<in> wset (a - 1, p)")
-          apply (rule_tac [5] wset_inv_mem_mem)
-               apply (simp_all add: wset_fin)
-  apply (rule inv_distinct, auto)
-  done
-
-lemma d22set_eq_wset: "zprime p ==> d22set (p - 2) = wset (p - 2, p)"
-  apply safe
-   apply (erule wset_mem)
-     apply (rule_tac [2] d22set_g_1)
-     apply (rule_tac [3] d22set_le)
-     apply (rule_tac [4] d22set_mem)
-      apply (erule_tac [4] wset_g_1)
-       prefer 6
-       apply (subst zle_add1_eq_le [symmetric])
-       apply (subgoal_tac "p - 2 + 1 = p - 1")
-        apply (simp (no_asm_simp))
-        apply (erule wset_less, auto)
-  done
-
-
-subsection {* Wilson *}
-
-lemma prime_g_5: "zprime p \<Longrightarrow> p \<noteq> 2 \<Longrightarrow> p \<noteq> 3 ==> 5 \<le> p"
-  apply (unfold zprime_def dvd_def)
-  apply (case_tac "p = 4", auto)
-   apply (rule notE)
-    prefer 2
-    apply assumption
-   apply (simp (no_asm))
-   apply (rule_tac x = 2 in exI)
-   apply (safe, arith)
-     apply (rule_tac x = 2 in exI, auto)
-  done
-
-theorem Wilson_Russ:
-    "zprime p ==> [zfact (p - 1) = -1] (mod p)"
-  apply (subgoal_tac "[(p - 1) * zfact (p - 2) = -1 * 1] (mod p)")
-   apply (rule_tac [2] zcong_zmult)
-    apply (simp only: zprime_def)
-    apply (subst zfact.simps)
-    apply (rule_tac t = "p - 1 - 1" and s = "p - 2" in subst, auto)
-   apply (simp only: zcong_def)
-   apply (simp (no_asm_simp))
-  apply (case_tac "p = 2")
-   apply (simp add: zfact.simps)
-  apply (case_tac "p = 3")
-   apply (simp add: zfact.simps)
-  apply (subgoal_tac "5 \<le> p")
-   apply (erule_tac [2] prime_g_5)
-    apply (subst d22set_prod_zfact [symmetric])
-    apply (subst d22set_eq_wset)
-     apply (rule_tac [2] wset_zcong_prod_1, auto)
-  done
-
-end
--- a/src/HOL/NumberTheory/document/root.tex	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,56 +0,0 @@
-
-\documentclass[11pt,a4paper]{article}
-\usepackage{graphicx}
-\usepackage{isabelle,isabellesym,pdfsetup}
-
-\urlstyle{rm}
-\isabellestyle{it}
-
-\begin{document}
-
-\title{Some results of number theory}
-\author{Jeremy Avigad\\
-    David Gray\\
-    Adam Kramer\\
-    Thomas M Rasmussen}
-
-\maketitle
-
-\begin{abstract}
-This is a collection of formalized proofs of many results of number theory.
-The proofs of the Chinese Remainder Theorem and Wilson's Theorem are due to
-Rasmussen.  The proof of Gauss's law of quadratic reciprocity is due to
-Avigad, Gray and Kramer.  Proofs can be found in most introductory number
-theory textbooks; Goldman's \emph{The Queen of Mathematics: a Historically
-Motivated Guide to Number Theory} provides some historical context.
-
-Avigad, Gray and Kramer have also provided library theories dealing with
-finite sets and finite sums, divisibility and congruences, parity and
-residues.  The authors are engaged in redesigning and polishing these theories
-for more serious use.  For the latest information in this respect, please see
-the web page \url{http://www.andrew.cmu.edu/~avigad/isabelle}.  Other theories
-contain proofs of Euler's criteria, Gauss' lemma, and the law of quadratic
-reciprocity.  The formalization follows Eisenstein's proof, which is the one
-most commonly found in introductory textbooks; in particular, it follows the
-presentation in Niven and Zuckerman, \emph{The Theory of Numbers}.
-
-To avoid having to count roots of polynomials, however, we relied on a trick
-previously used by David Russinoff in formalizing quadratic reciprocity for
-the Boyer-Moore theorem prover; see Russinoff, David, ``A mechanical proof
-of quadratic reciprocity,'' \emph{Journal of Automated Reasoning} 8:3-21,
-1992.  We are grateful to Larry Paulson for calling our attention to this
-reference.
-\end{abstract}
-
-\tableofcontents
-
-\begin{center}
-  \includegraphics[scale=0.5]{session_graph}  
-\end{center}
-
-\newpage
-
-\parindent 0pt\parskip 0.5ex
-\input{session}
-
-\end{document}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Number_Theory/Binomial.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,373 @@
+(*  Title:      Binomial.thy
+    Authors:    Lawrence C. Paulson, Jeremy Avigad, Tobias Nipkow
+
+
+Defines the "choose" function, and establishes basic properties.
+
+The original theory "Binomial" was by Lawrence C. Paulson, based on
+the work of Andy Gordon and Florian Kammueller. The approach here,
+which derives the definition of binomial coefficients in terms of the
+factorial function, is due to Jeremy Avigad. The binomial theorem was
+formalized by Tobias Nipkow.
+
+*)
+
+
+header {* Binomial *}
+
+theory Binomial
+imports Cong Fact
+begin
+
+
+subsection {* Main definitions *}
+
+class binomial =
+
+fixes 
+  binomial :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "choose" 65)
+
+(* definitions for the natural numbers *)
+
+instantiation nat :: binomial
+
+begin 
+
+fun
+  binomial_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+where
+  "binomial_nat n k =
+   (if k = 0 then 1 else
+    if n = 0 then 0 else
+      (binomial (n - 1) k) + (binomial (n - 1) (k - 1)))"
+
+instance proof qed
+
+end
+
+(* definitions for the integers *)
+
+instantiation int :: binomial
+
+begin 
+
+definition
+  binomial_int :: "int => int \<Rightarrow> int"
+where
+  "binomial_int n k = (if n \<ge> 0 \<and> k \<ge> 0 then int (binomial (nat n) (nat k))
+      else 0)"
+instance proof qed
+
+end
+
+
+subsection {* Set up Transfer *}
+
+lemma transfer_nat_int_binomial:
+  "(n::int) >= 0 \<Longrightarrow> k >= 0 \<Longrightarrow> binomial (nat n) (nat k) = 
+      nat (binomial n k)"
+  unfolding binomial_int_def 
+  by auto
+
+lemma transfer_nat_int_binomial_closure:
+  "n >= (0::int) \<Longrightarrow> k >= 0 \<Longrightarrow> binomial n k >= 0"
+  by (auto simp add: binomial_int_def)
+
+declare TransferMorphism_nat_int[transfer add return: 
+    transfer_nat_int_binomial transfer_nat_int_binomial_closure]
+
+lemma transfer_int_nat_binomial:
+  "binomial (int n) (int k) = int (binomial n k)"
+  unfolding fact_int_def binomial_int_def by auto
+
+lemma transfer_int_nat_binomial_closure:
+  "is_nat n \<Longrightarrow> is_nat k \<Longrightarrow> binomial n k >= 0"
+  by (auto simp add: binomial_int_def)
+
+declare TransferMorphism_int_nat[transfer add return: 
+    transfer_int_nat_binomial transfer_int_nat_binomial_closure]
+
+
+subsection {* Binomial coefficients *}
+
+lemma choose_zero_nat [simp]: "(n::nat) choose 0 = 1"
+  by simp
+
+lemma choose_zero_int [simp]: "n \<ge> 0 \<Longrightarrow> (n::int) choose 0 = 1"
+  by (simp add: binomial_int_def)
+
+lemma zero_choose_nat [rule_format,simp]: "ALL (k::nat) > n. n choose k = 0"
+  by (induct n rule: induct'_nat, auto)
+
+lemma zero_choose_int [rule_format,simp]: "(k::int) > n \<Longrightarrow> n choose k = 0"
+  unfolding binomial_int_def apply (case_tac "n < 0")
+  apply force
+  apply (simp del: binomial_nat.simps)
+done
+
+lemma choose_reduce_nat: "(n::nat) > 0 \<Longrightarrow> 0 < k \<Longrightarrow>
+    (n choose k) = ((n - 1) choose k) + ((n - 1) choose (k - 1))"
+  by simp
+
+lemma choose_reduce_int: "(n::int) > 0 \<Longrightarrow> 0 < k \<Longrightarrow>
+    (n choose k) = ((n - 1) choose k) + ((n - 1) choose (k - 1))"
+  unfolding binomial_int_def apply (subst choose_reduce_nat)
+    apply (auto simp del: binomial_nat.simps 
+      simp add: nat_diff_distrib)
+done
+
+lemma choose_plus_one_nat: "((n::nat) + 1) choose (k + 1) = 
+    (n choose (k + 1)) + (n choose k)"
+  by (simp add: choose_reduce_nat)
+
+lemma choose_Suc_nat: "(Suc n) choose (Suc k) = 
+    (n choose (Suc k)) + (n choose k)"
+  by (simp add: choose_reduce_nat One_nat_def)
+
+lemma choose_plus_one_int: "n \<ge> 0 \<Longrightarrow> k \<ge> 0 \<Longrightarrow> ((n::int) + 1) choose (k + 1) = 
+    (n choose (k + 1)) + (n choose k)"
+  by (simp add: binomial_int_def choose_plus_one_nat nat_add_distrib del: binomial_nat.simps)
+
+declare binomial_nat.simps [simp del]
+
+lemma choose_self_nat [simp]: "((n::nat) choose n) = 1"
+  by (induct n rule: induct'_nat, auto simp add: choose_plus_one_nat)
+
+lemma choose_self_int [simp]: "n \<ge> 0 \<Longrightarrow> ((n::int) choose n) = 1"
+  by (auto simp add: binomial_int_def)
+
+lemma choose_one_nat [simp]: "(n::nat) choose 1 = n"
+  by (induct n rule: induct'_nat, auto simp add: choose_reduce_nat)
+
+lemma choose_one_int [simp]: "n \<ge> 0 \<Longrightarrow> (n::int) choose 1 = n"
+  by (auto simp add: binomial_int_def)
+
+lemma plus_one_choose_self_nat [simp]: "(n::nat) + 1 choose n = n + 1"
+  apply (induct n rule: induct'_nat, force)
+  apply (case_tac "n = 0")
+  apply auto
+  apply (subst choose_reduce_nat)
+  apply (auto simp add: One_nat_def)  
+  (* natdiff_cancel_numerals introduces Suc *)
+done
+
+lemma Suc_choose_self_nat [simp]: "(Suc n) choose n = Suc n"
+  using plus_one_choose_self_nat by (simp add: One_nat_def)
+
+lemma plus_one_choose_self_int [rule_format, simp]: 
+    "(n::int) \<ge> 0 \<longrightarrow> n + 1 choose n = n + 1"
+   by (auto simp add: binomial_int_def nat_add_distrib)
+
+(* bounded quantification doesn't work with the unicode characters? *)
+lemma choose_pos_nat [rule_format]: "ALL k <= (n::nat). 
+    ((n::nat) choose k) > 0"
+  apply (induct n rule: induct'_nat) 
+  apply force
+  apply clarify
+  apply (case_tac "k = 0")
+  apply force
+  apply (subst choose_reduce_nat)
+  apply auto
+done
+
+lemma choose_pos_int: "n \<ge> 0 \<Longrightarrow> k >= 0 \<Longrightarrow> k \<le> n \<Longrightarrow>
+    ((n::int) choose k) > 0"
+  by (auto simp add: binomial_int_def choose_pos_nat)
+
+lemma binomial_induct [rule_format]: "(ALL (n::nat). P n n) \<longrightarrow> 
+    (ALL n. P (n + 1) 0) \<longrightarrow> (ALL n. (ALL k < n. P n k \<longrightarrow> P n (k + 1) \<longrightarrow>
+    P (n + 1) (k + 1))) \<longrightarrow> (ALL k <= n. P n k)"
+  apply (induct n rule: induct'_nat)
+  apply auto
+  apply (case_tac "k = 0")
+  apply auto
+  apply (case_tac "k = n + 1")
+  apply auto
+  apply (drule_tac x = n in spec) back back 
+  apply (drule_tac x = "k - 1" in spec) back back back
+  apply auto
+done
+
+lemma choose_altdef_aux_nat: "(k::nat) \<le> n \<Longrightarrow> 
+    fact k * fact (n - k) * (n choose k) = fact n"
+  apply (rule binomial_induct [of _ k n])
+  apply auto
+proof -
+  fix k :: nat and n
+  assume less: "k < n"
+  assume ih1: "fact k * fact (n - k) * (n choose k) = fact n"
+  hence one: "fact (k + 1) * fact (n - k) * (n choose k) = (k + 1) * fact n"
+    by (subst fact_plus_one_nat, auto)
+  assume ih2: "fact (k + 1) * fact (n - (k + 1)) * (n choose (k + 1)) = 
+      fact n"
+  with less have "fact (k + 1) * fact ((n - (k + 1)) + 1) * 
+      (n choose (k + 1)) = (n - k) * fact n"
+    by (subst (2) fact_plus_one_nat, auto)
+  with less have two: "fact (k + 1) * fact (n - k) * (n choose (k + 1)) = 
+      (n - k) * fact n" by simp
+  have "fact (k + 1) * fact (n - k) * (n + 1 choose (k + 1)) =
+      fact (k + 1) * fact (n - k) * (n choose (k + 1)) + 
+      fact (k + 1) * fact (n - k) * (n choose k)" 
+    by (subst choose_reduce_nat, auto simp add: ring_simps)
+  also note one
+  also note two
+  also with less have "(n - k) * fact n + (k + 1) * fact n= fact (n + 1)" 
+    apply (subst fact_plus_one_nat)
+    apply (subst left_distrib [symmetric])
+    apply simp
+    done
+  finally show "fact (k + 1) * fact (n - k) * (n + 1 choose (k + 1)) = 
+    fact (n + 1)" .
+qed
+
+lemma choose_altdef_nat: "(k::nat) \<le> n \<Longrightarrow> 
+    n choose k = fact n div (fact k * fact (n - k))"
+  apply (frule choose_altdef_aux_nat)
+  apply (erule subst)
+  apply (simp add: mult_ac)
+done
+
+
+lemma choose_altdef_int: 
+  assumes "(0::int) <= k" and "k <= n"
+  shows "n choose k = fact n div (fact k * fact (n - k))"
+  
+  apply (subst tsub_eq [symmetric], rule prems)
+  apply (rule choose_altdef_nat [transferred])
+  using prems apply auto
+done
+
+lemma choose_dvd_nat: "(k::nat) \<le> n \<Longrightarrow> fact k * fact (n - k) dvd fact n"
+  unfolding dvd_def apply (frule choose_altdef_aux_nat)
+  (* why don't blast and auto get this??? *)
+  apply (rule exI)
+  apply (erule sym)
+done
+
+lemma choose_dvd_int: 
+  assumes "(0::int) <= k" and "k <= n"
+  shows "fact k * fact (n - k) dvd fact n"
+ 
+  apply (subst tsub_eq [symmetric], rule prems)
+  apply (rule choose_dvd_nat [transferred])
+  using prems apply auto
+done
+
+(* generalizes Tobias Nipkow's proof to any commutative semiring *)
+theorem binomial: "(a+b::'a::{comm_ring_1,power})^n = 
+  (SUM k=0..n. (of_nat (n choose k)) * a^k * b^(n-k))" (is "?P n")
+proof (induct n rule: induct'_nat)
+  show "?P 0" by simp
+next
+  fix n
+  assume ih: "?P n"
+  have decomp: "{0..n+1} = {0} Un {n+1} Un {1..n}"
+    by auto
+  have decomp2: "{0..n} = {0} Un {1..n}"
+    by auto
+  have decomp3: "{1..n+1} = {n+1} Un {1..n}"
+    by auto
+  have "(a+b)^(n+1) = 
+      (a+b) * (SUM k=0..n. of_nat (n choose k) * a^k * b^(n-k))"
+    using ih by (simp add: power_plus_one)
+  also have "... =  a*(SUM k=0..n. of_nat (n choose k) * a^k * b^(n-k)) +
+                   b*(SUM k=0..n. of_nat (n choose k) * a^k * b^(n-k))"
+    by (rule distrib)
+  also have "... = (SUM k=0..n. of_nat (n choose k) * a^(k+1) * b^(n-k)) +
+                  (SUM k=0..n. of_nat (n choose k) * a^k * b^(n-k+1))"
+    by (subst (1 2) power_plus_one, simp add: setsum_right_distrib mult_ac)
+  also have "... = (SUM k=0..n. of_nat (n choose k) * a^k * b^(n+1-k)) +
+                  (SUM k=1..n+1. of_nat (n choose (k - 1)) * a^k * b^(n+1-k))"
+    by (simp add:setsum_shift_bounds_cl_Suc_ivl Suc_diff_le 
+             power_Suc ring_simps One_nat_def del:setsum_cl_ivl_Suc)
+  also have "... = a^(n+1) + b^(n+1) +
+                  (SUM k=1..n. of_nat (n choose (k - 1)) * a^k * b^(n+1-k)) +
+                  (SUM k=1..n. of_nat (n choose k) * a^k * b^(n+1-k))"
+    by (simp add: decomp2 decomp3)
+  also have
+      "... = a^(n+1) + b^(n+1) + 
+         (SUM k=1..n. of_nat(n+1 choose k) * a^k * b^(n+1-k))"
+    by (auto simp add: ring_simps setsum_addf [symmetric]
+      choose_reduce_nat)
+  also have "... = (SUM k=0..n+1. of_nat (n+1 choose k) * a^k * b^(n+1-k))"
+    using decomp by (simp add: ring_simps)
+  finally show "?P (n + 1)" by simp
+qed
+
+lemma set_explicit: "{S. S = T \<and> P S} = (if P T then {T} else {})"
+  by auto
+
+lemma card_subsets_nat [rule_format]:
+  fixes S :: "'a set"
+  assumes "finite S"
+  shows "ALL k. card {T. T \<le> S \<and> card T = k} = card S choose k" 
+      (is "?P S")
+using `finite S`
+proof (induct set: finite)
+  show "?P {}" by (auto simp add: set_explicit)
+  next fix x :: "'a" and F
+  assume iassms: "finite F" "x ~: F"
+  assume ih: "?P F"
+  show "?P (insert x F)" (is "ALL k. ?Q k")
+  proof
+    fix k
+    show "card {T. T \<subseteq> (insert x F) \<and> card T = k} = 
+        card (insert x F) choose k" (is "?Q k")
+    proof (induct k rule: induct'_nat)
+      from iassms have "{T. T \<le> (insert x F) \<and> card T = 0} = {{}}"
+        apply auto
+        apply (subst (asm) card_0_eq)
+        apply (auto elim: finite_subset)
+        done
+      thus "?Q 0" 
+        by auto
+      next fix k
+      show "?Q (k + 1)"
+      proof -
+        from iassms have fin: "finite (insert x F)" by auto
+        hence "{ T. T \<subseteq> insert x F \<and> card T = k + 1} =
+          {T. T \<le> F & card T = k + 1} Un 
+          {T. T \<le> insert x F & x : T & card T = k + 1}"
+          by (auto intro!: subsetI)
+        with iassms fin have "card ({T. T \<le> insert x F \<and> card T = k + 1}) = 
+          card ({T. T \<subseteq> F \<and> card T = k + 1}) + 
+          card ({T. T \<subseteq> insert x F \<and> x : T \<and> card T = k + 1})"
+          apply (subst card_Un_disjoint [symmetric])
+          apply auto
+          (* note: nice! Didn't have to say anything here *)
+          done
+        also from ih have "card ({T. T \<subseteq> F \<and> card T = k + 1}) = 
+          card F choose (k+1)" by auto
+        also have "card ({T. T \<subseteq> insert x F \<and> x : T \<and> card T = k + 1}) =
+          card ({T. T <= F & card T = k})"
+        proof -
+          let ?f = "%T. T Un {x}"
+          from iassms have "inj_on ?f {T. T <= F & card T = k}"
+            unfolding inj_on_def by (auto intro!: subsetI)
+          hence "card ({T. T <= F & card T = k}) = 
+            card(?f ` {T. T <= F & card T = k})"
+            by (rule card_image [symmetric])
+          also from iassms fin have "?f ` {T. T <= F & card T = k} = 
+            {T. T \<subseteq> insert x F \<and> x : T \<and> card T = k + 1}"
+            unfolding image_def 
+            (* I can't figure out why this next line takes so long *)
+            apply auto
+            apply (frule (1) finite_subset, force)
+            apply (rule_tac x = "xa - {x}" in exI)
+            apply (subst card_Diff_singleton)
+            apply (auto elim: finite_subset)
+            done
+          finally show ?thesis by (rule sym)
+        qed
+        also from ih have "card ({T. T <= F & card T = k}) = card F choose k"
+          by auto
+        finally have "card ({T. T \<le> insert x F \<and> card T = k + 1}) = 
+          card F choose (k + 1) + (card F choose k)".
+        with iassms choose_plus_one_nat show ?thesis
+          by auto
+      qed
+    qed
+  qed
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Number_Theory/Cong.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,1091 @@
+(*  Title:      HOL/Library/Cong.thy
+    ID:         
+    Authors:    Christophe Tabacznyj, Lawrence C. Paulson, Amine Chaieb,
+                Thomas M. Rasmussen, Jeremy Avigad
+
+
+Defines congruence (notation: [x = y] (mod z)) for natural numbers and
+integers.
+
+This file combines and revises a number of prior developments.
+
+The original theories "GCD" and "Primes" were by Christophe Tabacznyj
+and Lawrence C. Paulson, based on \cite{davenport92}. They introduced
+gcd, lcm, and prime for the natural numbers.
+
+The original theory "IntPrimes" was by Thomas M. Rasmussen, and
+extended gcd, lcm, primes to the integers. Amine Chaieb provided
+another extension of the notions to the integers, and added a number
+of results to "Primes" and "GCD". 
+
+The original theory, "IntPrimes", by Thomas M. Rasmussen, defined and
+developed the congruence relations on the integers. The notion was
+extended to the natural numbers by Chiaeb. Jeremy Avigad combined
+these, revised and tidied them, made the development uniform for the
+natural numbers and the integers, and added a number of new theorems.
+
+*)
+
+
+header {* Congruence *}
+
+theory Cong
+imports GCD Primes
+begin
+
+subsection {* Turn off One_nat_def *}
+
+lemma induct'_nat [case_names zero plus1, induct type: nat]: 
+    "\<lbrakk> P (0::nat); !!n. P n \<Longrightarrow> P (n + 1)\<rbrakk> \<Longrightarrow> P n"
+by (erule nat_induct) (simp add:One_nat_def)
+
+lemma cases_nat [case_names zero plus1, cases type: nat]: 
+    "P (0::nat) \<Longrightarrow> (!!n. P (n + 1)) \<Longrightarrow> P n"
+by(metis induct'_nat)
+
+lemma power_plus_one [simp]: "(x::'a::power)^(n + 1) = x * x^n"
+by (simp add: One_nat_def)
+
+lemma power_eq_one_eq_nat [simp]: 
+  "((x::nat)^m = 1) = (m = 0 | x = 1)"
+by (induct m, auto)
+
+lemma card_insert_if' [simp]: "finite A \<Longrightarrow>
+  card (insert x A) = (if x \<in> A then (card A) else (card A) + 1)"
+by (auto simp add: insert_absorb)
+
+(* why wasn't card_insert_if a simp rule? *)
+declare card_insert_disjoint [simp del]
+
+lemma nat_1' [simp]: "nat 1 = 1"
+by simp
+
+(* For those annoying moments where Suc reappears, use Suc_eq_plus1 *)
+
+declare nat_1 [simp del]
+declare add_2_eq_Suc [simp del] 
+declare add_2_eq_Suc' [simp del]
+
+
+declare mod_pos_pos_trivial [simp]
+
+
+subsection {* Main definitions *}
+
+class cong =
+
+fixes 
+  cong :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" ("(1[_ = _] '(mod _'))")
+
+begin
+
+abbreviation
+  notcong :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" ("(1[_ \<noteq> _] '(mod _'))")
+where
+  "notcong x y m == (~cong x y m)" 
+
+end
+
+(* definitions for the natural numbers *)
+
+instantiation nat :: cong
+
+begin 
+
+definition 
+  cong_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool"
+where 
+  "cong_nat x y m = ((x mod m) = (y mod m))"
+
+instance proof qed
+
+end
+
+
+(* definitions for the integers *)
+
+instantiation int :: cong
+
+begin 
+
+definition 
+  cong_int :: "int \<Rightarrow> int \<Rightarrow> int \<Rightarrow> bool"
+where 
+  "cong_int x y m = ((x mod m) = (y mod m))"
+
+instance proof qed
+
+end
+
+
+subsection {* Set up Transfer *}
+
+
+lemma transfer_nat_int_cong:
+  "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> m >= 0 \<Longrightarrow> 
+    ([(nat x) = (nat y)] (mod (nat m))) = ([x = y] (mod m))"
+  unfolding cong_int_def cong_nat_def 
+  apply (auto simp add: nat_mod_distrib [symmetric])
+  apply (subst (asm) eq_nat_nat_iff)
+  apply (case_tac "m = 0", force, rule pos_mod_sign, force)+
+  apply assumption
+done
+
+declare TransferMorphism_nat_int[transfer add return: 
+    transfer_nat_int_cong]
+
+lemma transfer_int_nat_cong:
+  "[(int x) = (int y)] (mod (int m)) = [x = y] (mod m)"
+  apply (auto simp add: cong_int_def cong_nat_def)
+  apply (auto simp add: zmod_int [symmetric])
+done
+
+declare TransferMorphism_int_nat[transfer add return: 
+    transfer_int_nat_cong]
+
+
+subsection {* Congruence *}
+
+(* was zcong_0, etc. *)
+lemma cong_0_nat [simp, presburger]: "([(a::nat) = b] (mod 0)) = (a = b)"
+  by (unfold cong_nat_def, auto)
+
+lemma cong_0_int [simp, presburger]: "([(a::int) = b] (mod 0)) = (a = b)"
+  by (unfold cong_int_def, auto)
+
+lemma cong_1_nat [simp, presburger]: "[(a::nat) = b] (mod 1)"
+  by (unfold cong_nat_def, auto)
+
+lemma cong_Suc_0_nat [simp, presburger]: "[(a::nat) = b] (mod Suc 0)"
+  by (unfold cong_nat_def, auto simp add: One_nat_def)
+
+lemma cong_1_int [simp, presburger]: "[(a::int) = b] (mod 1)"
+  by (unfold cong_int_def, auto)
+
+lemma cong_refl_nat [simp]: "[(k::nat) = k] (mod m)"
+  by (unfold cong_nat_def, auto)
+
+lemma cong_refl_int [simp]: "[(k::int) = k] (mod m)"
+  by (unfold cong_int_def, auto)
+
+lemma cong_sym_nat: "[(a::nat) = b] (mod m) \<Longrightarrow> [b = a] (mod m)"
+  by (unfold cong_nat_def, auto)
+
+lemma cong_sym_int: "[(a::int) = b] (mod m) \<Longrightarrow> [b = a] (mod m)"
+  by (unfold cong_int_def, auto)
+
+lemma cong_sym_eq_nat: "[(a::nat) = b] (mod m) = [b = a] (mod m)"
+  by (unfold cong_nat_def, auto)
+
+lemma cong_sym_eq_int: "[(a::int) = b] (mod m) = [b = a] (mod m)"
+  by (unfold cong_int_def, auto)
+
+lemma cong_trans_nat [trans]:
+    "[(a::nat) = b] (mod m) \<Longrightarrow> [b = c] (mod m) \<Longrightarrow> [a = c] (mod m)"
+  by (unfold cong_nat_def, auto)
+
+lemma cong_trans_int [trans]:
+    "[(a::int) = b] (mod m) \<Longrightarrow> [b = c] (mod m) \<Longrightarrow> [a = c] (mod m)"
+  by (unfold cong_int_def, auto)
+
+lemma cong_add_nat:
+    "[(a::nat) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a + c = b + d] (mod m)"
+  apply (unfold cong_nat_def)
+  apply (subst (1 2) mod_add_eq)
+  apply simp
+done
+
+lemma cong_add_int:
+    "[(a::int) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a + c = b + d] (mod m)"
+  apply (unfold cong_int_def)
+  apply (subst (1 2) mod_add_left_eq)
+  apply (subst (1 2) mod_add_right_eq)
+  apply simp
+done
+
+lemma cong_diff_int:
+    "[(a::int) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a - c = b - d] (mod m)"
+  apply (unfold cong_int_def)
+  apply (subst (1 2) mod_diff_eq)
+  apply simp
+done
+
+lemma cong_diff_aux_int:
+  "(a::int) >= c \<Longrightarrow> b >= d \<Longrightarrow> [(a::int) = b] (mod m) \<Longrightarrow> 
+      [c = d] (mod m) \<Longrightarrow> [tsub a c = tsub b d] (mod m)"
+  apply (subst (1 2) tsub_eq)
+  apply (auto intro: cong_diff_int)
+done;
+
+lemma cong_diff_nat:
+  assumes "(a::nat) >= c" and "b >= d" and "[a = b] (mod m)" and
+    "[c = d] (mod m)"
+  shows "[a - c = b - d] (mod m)"
+
+  using prems by (rule cong_diff_aux_int [transferred]);
+
+lemma cong_mult_nat:
+    "[(a::nat) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a * c = b * d] (mod m)"
+  apply (unfold cong_nat_def)
+  apply (subst (1 2) mod_mult_eq)
+  apply simp
+done
+
+lemma cong_mult_int:
+    "[(a::int) = b] (mod m) \<Longrightarrow> [c = d] (mod m) \<Longrightarrow> [a * c = b * d] (mod m)"
+  apply (unfold cong_int_def)
+  apply (subst (1 2) zmod_zmult1_eq)
+  apply (subst (1 2) mult_commute)
+  apply (subst (1 2) zmod_zmult1_eq)
+  apply simp
+done
+
+lemma cong_exp_nat: "[(x::nat) = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"
+  apply (induct k)
+  apply (auto simp add: cong_refl_nat cong_mult_nat)
+done
+
+lemma cong_exp_int: "[(x::int) = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"
+  apply (induct k)
+  apply (auto simp add: cong_refl_int cong_mult_int)
+done
+
+lemma cong_setsum_nat [rule_format]: 
+    "(ALL x: A. [((f x)::nat) = g x] (mod m)) \<longrightarrow> 
+      [(SUM x:A. f x) = (SUM x:A. g x)] (mod m)"
+  apply (case_tac "finite A")
+  apply (induct set: finite)
+  apply (auto intro: cong_add_nat)
+done
+
+lemma cong_setsum_int [rule_format]:
+    "(ALL x: A. [((f x)::int) = g x] (mod m)) \<longrightarrow> 
+      [(SUM x:A. f x) = (SUM x:A. g x)] (mod m)"
+  apply (case_tac "finite A")
+  apply (induct set: finite)
+  apply (auto intro: cong_add_int)
+done
+
+lemma cong_setprod_nat [rule_format]: 
+    "(ALL x: A. [((f x)::nat) = g x] (mod m)) \<longrightarrow> 
+      [(PROD x:A. f x) = (PROD x:A. g x)] (mod m)"
+  apply (case_tac "finite A")
+  apply (induct set: finite)
+  apply (auto intro: cong_mult_nat)
+done
+
+lemma cong_setprod_int [rule_format]: 
+    "(ALL x: A. [((f x)::int) = g x] (mod m)) \<longrightarrow> 
+      [(PROD x:A. f x) = (PROD x:A. g x)] (mod m)"
+  apply (case_tac "finite A")
+  apply (induct set: finite)
+  apply (auto intro: cong_mult_int)
+done
+
+lemma cong_scalar_nat: "[(a::nat)= b] (mod m) \<Longrightarrow> [a * k = b * k] (mod m)"
+  by (rule cong_mult_nat, simp_all)
+
+lemma cong_scalar_int: "[(a::int)= b] (mod m) \<Longrightarrow> [a * k = b * k] (mod m)"
+  by (rule cong_mult_int, simp_all)
+
+lemma cong_scalar2_nat: "[(a::nat)= b] (mod m) \<Longrightarrow> [k * a = k * b] (mod m)"
+  by (rule cong_mult_nat, simp_all)
+
+lemma cong_scalar2_int: "[(a::int)= b] (mod m) \<Longrightarrow> [k * a = k * b] (mod m)"
+  by (rule cong_mult_int, simp_all)
+
+lemma cong_mult_self_nat: "[(a::nat) * m = 0] (mod m)"
+  by (unfold cong_nat_def, auto)
+
+lemma cong_mult_self_int: "[(a::int) * m = 0] (mod m)"
+  by (unfold cong_int_def, auto)
+
+lemma cong_eq_diff_cong_0_int: "[(a::int) = b] (mod m) = [a - b = 0] (mod m)"
+  apply (rule iffI)
+  apply (erule cong_diff_int [of a b m b b, simplified])
+  apply (erule cong_add_int [of "a - b" 0 m b b, simplified])
+done
+
+lemma cong_eq_diff_cong_0_aux_int: "a >= b \<Longrightarrow>
+    [(a::int) = b] (mod m) = [tsub a b = 0] (mod m)"
+  by (subst tsub_eq, assumption, rule cong_eq_diff_cong_0_int)
+
+lemma cong_eq_diff_cong_0_nat:
+  assumes "(a::nat) >= b"
+  shows "[a = b] (mod m) = [a - b = 0] (mod m)"
+
+  using prems by (rule cong_eq_diff_cong_0_aux_int [transferred])
+
+lemma cong_diff_cong_0'_nat: 
+  "[(x::nat) = y] (mod n) \<longleftrightarrow> 
+    (if x <= y then [y - x = 0] (mod n) else [x - y = 0] (mod n))"
+  apply (case_tac "y <= x")
+  apply (frule cong_eq_diff_cong_0_nat [where m = n])
+  apply auto [1]
+  apply (subgoal_tac "x <= y")
+  apply (frule cong_eq_diff_cong_0_nat [where m = n])
+  apply (subst cong_sym_eq_nat)
+  apply auto
+done
+
+lemma cong_altdef_nat: "(a::nat) >= b \<Longrightarrow> [a = b] (mod m) = (m dvd (a - b))"
+  apply (subst cong_eq_diff_cong_0_nat, assumption)
+  apply (unfold cong_nat_def)
+  apply (simp add: dvd_eq_mod_eq_0 [symmetric])
+done
+
+lemma cong_altdef_int: "[(a::int) = b] (mod m) = (m dvd (a - b))"
+  apply (subst cong_eq_diff_cong_0_int)
+  apply (unfold cong_int_def)
+  apply (simp add: dvd_eq_mod_eq_0 [symmetric])
+done
+
+lemma cong_abs_int: "[(x::int) = y] (mod abs m) = [x = y] (mod m)"
+  by (simp add: cong_altdef_int)
+
+lemma cong_square_int:
+   "\<lbrakk> prime (p::int); 0 < a; [a * a = 1] (mod p) \<rbrakk>
+    \<Longrightarrow> [a = 1] (mod p) \<or> [a = - 1] (mod p)"
+  apply (simp only: cong_altdef_int)
+  apply (subst prime_dvd_mult_eq_int [symmetric], assumption)
+  (* any way around this? *)
+  apply (subgoal_tac "a * a - 1 = (a - 1) * (a - -1)")
+  apply (auto simp add: ring_simps)
+done
+
+lemma cong_mult_rcancel_int:
+  "coprime k (m::int) \<Longrightarrow> [a * k = b * k] (mod m) = [a = b] (mod m)"
+  apply (subst (1 2) cong_altdef_int)
+  apply (subst left_diff_distrib [symmetric])
+  apply (rule coprime_dvd_mult_iff_int)
+  apply (subst gcd_commute_int, assumption)
+done
+
+lemma cong_mult_rcancel_nat:
+  assumes  "coprime k (m::nat)"
+  shows "[a * k = b * k] (mod m) = [a = b] (mod m)"
+
+  apply (rule cong_mult_rcancel_int [transferred])
+  using prems apply auto
+done
+
+lemma cong_mult_lcancel_nat:
+  "coprime k (m::nat) \<Longrightarrow> [k * a = k * b ] (mod m) = [a = b] (mod m)"
+  by (simp add: mult_commute cong_mult_rcancel_nat)
+
+lemma cong_mult_lcancel_int:
+  "coprime k (m::int) \<Longrightarrow> [k * a = k * b] (mod m) = [a = b] (mod m)"
+  by (simp add: mult_commute cong_mult_rcancel_int)
+
+(* was zcong_zgcd_zmult_zmod *)
+lemma coprime_cong_mult_int:
+  "[(a::int) = b] (mod m) \<Longrightarrow> [a = b] (mod n) \<Longrightarrow> coprime m n
+    \<Longrightarrow> [a = b] (mod m * n)"
+  apply (simp only: cong_altdef_int)
+  apply (erule (2) divides_mult_int)
+done
+
+lemma coprime_cong_mult_nat:
+  assumes "[(a::nat) = b] (mod m)" and "[a = b] (mod n)" and "coprime m n"
+  shows "[a = b] (mod m * n)"
+
+  apply (rule coprime_cong_mult_int [transferred])
+  using prems apply auto
+done
+
+lemma cong_less_imp_eq_nat: "0 \<le> (a::nat) \<Longrightarrow>
+    a < m \<Longrightarrow> 0 \<le> b \<Longrightarrow> b < m \<Longrightarrow> [a = b] (mod m) \<Longrightarrow> a = b"
+  by (auto simp add: cong_nat_def mod_pos_pos_trivial)
+
+lemma cong_less_imp_eq_int: "0 \<le> (a::int) \<Longrightarrow>
+    a < m \<Longrightarrow> 0 \<le> b \<Longrightarrow> b < m \<Longrightarrow> [a = b] (mod m) \<Longrightarrow> a = b"
+  by (auto simp add: cong_int_def mod_pos_pos_trivial)
+
+lemma cong_less_unique_nat:
+    "0 < (m::nat) \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))"
+  apply auto
+  apply (rule_tac x = "a mod m" in exI)
+  apply (unfold cong_nat_def, auto)
+done
+
+lemma cong_less_unique_int:
+    "0 < (m::int) \<Longrightarrow> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))"
+  apply auto
+  apply (rule_tac x = "a mod m" in exI)
+  apply (unfold cong_int_def, auto simp add: mod_pos_pos_trivial)
+done
+
+lemma cong_iff_lin_int: "([(a::int) = b] (mod m)) = (\<exists>k. b = a + m * k)"
+  apply (auto simp add: cong_altdef_int dvd_def ring_simps)
+  apply (rule_tac [!] x = "-k" in exI, auto)
+done
+
+lemma cong_iff_lin_nat: "([(a::nat) = b] (mod m)) = 
+    (\<exists>k1 k2. b + k1 * m = a + k2 * m)"
+  apply (rule iffI)
+  apply (case_tac "b <= a")
+  apply (subst (asm) cong_altdef_nat, assumption)
+  apply (unfold dvd_def, auto)
+  apply (rule_tac x = k in exI)
+  apply (rule_tac x = 0 in exI)
+  apply (auto simp add: ring_simps)
+  apply (subst (asm) cong_sym_eq_nat)
+  apply (subst (asm) cong_altdef_nat)
+  apply force
+  apply (unfold dvd_def, auto)
+  apply (rule_tac x = 0 in exI)
+  apply (rule_tac x = k in exI)
+  apply (auto simp add: ring_simps)
+  apply (unfold cong_nat_def)
+  apply (subgoal_tac "a mod m = (a + k2 * m) mod m")
+  apply (erule ssubst)back
+  apply (erule subst)
+  apply auto
+done
+
+lemma cong_gcd_eq_int: "[(a::int) = b] (mod m) \<Longrightarrow> gcd a m = gcd b m"
+  apply (subst (asm) cong_iff_lin_int, auto)
+  apply (subst add_commute) 
+  apply (subst (2) gcd_commute_int)
+  apply (subst mult_commute)
+  apply (subst gcd_add_mult_int)
+  apply (rule gcd_commute_int)
+done
+
+lemma cong_gcd_eq_nat: 
+  assumes "[(a::nat) = b] (mod m)"
+  shows "gcd a m = gcd b m"
+
+  apply (rule cong_gcd_eq_int [transferred])
+  using prems apply auto
+done
+
+lemma cong_imp_coprime_nat: "[(a::nat) = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow> 
+    coprime b m"
+  by (auto simp add: cong_gcd_eq_nat)
+
+lemma cong_imp_coprime_int: "[(a::int) = b] (mod m) \<Longrightarrow> coprime a m \<Longrightarrow> 
+    coprime b m"
+  by (auto simp add: cong_gcd_eq_int)
+
+lemma cong_cong_mod_nat: "[(a::nat) = b] (mod m) = 
+    [a mod m = b mod m] (mod m)"
+  by (auto simp add: cong_nat_def)
+
+lemma cong_cong_mod_int: "[(a::int) = b] (mod m) = 
+    [a mod m = b mod m] (mod m)"
+  by (auto simp add: cong_int_def)
+
+lemma cong_minus_int [iff]: "[(a::int) = b] (mod -m) = [a = b] (mod m)"
+  by (subst (1 2) cong_altdef_int, auto)
+
+lemma cong_zero_nat [iff]: "[(a::nat) = b] (mod 0) = (a = b)"
+  by (auto simp add: cong_nat_def)
+
+lemma cong_zero_int [iff]: "[(a::int) = b] (mod 0) = (a = b)"
+  by (auto simp add: cong_int_def)
+
+(*
+lemma mod_dvd_mod_int:
+    "0 < (m::int) \<Longrightarrow> m dvd b \<Longrightarrow> (a mod b mod m) = (a mod m)"
+  apply (unfold dvd_def, auto)
+  apply (rule mod_mod_cancel)
+  apply auto
+done
+
+lemma mod_dvd_mod:
+  assumes "0 < (m::nat)" and "m dvd b"
+  shows "(a mod b mod m) = (a mod m)"
+
+  apply (rule mod_dvd_mod_int [transferred])
+  using prems apply auto
+done
+*)
+
+lemma cong_add_lcancel_nat: 
+    "[(a::nat) + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)" 
+  by (simp add: cong_iff_lin_nat)
+
+lemma cong_add_lcancel_int: 
+    "[(a::int) + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)" 
+  by (simp add: cong_iff_lin_int)
+
+lemma cong_add_rcancel_nat: "[(x::nat) + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
+  by (simp add: cong_iff_lin_nat)
+
+lemma cong_add_rcancel_int: "[(x::int) + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
+  by (simp add: cong_iff_lin_int)
+
+lemma cong_add_lcancel_0_nat: "[(a::nat) + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" 
+  by (simp add: cong_iff_lin_nat)
+
+lemma cong_add_lcancel_0_int: "[(a::int) + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" 
+  by (simp add: cong_iff_lin_int)
+
+lemma cong_add_rcancel_0_nat: "[x + (a::nat) = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" 
+  by (simp add: cong_iff_lin_nat)
+
+lemma cong_add_rcancel_0_int: "[x + (a::int) = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" 
+  by (simp add: cong_iff_lin_int)
+
+lemma cong_dvd_modulus_nat: "[(x::nat) = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow> 
+    [x = y] (mod n)"
+  apply (auto simp add: cong_iff_lin_nat dvd_def)
+  apply (rule_tac x="k1 * k" in exI)
+  apply (rule_tac x="k2 * k" in exI)
+  apply (simp add: ring_simps)
+done
+
+lemma cong_dvd_modulus_int: "[(x::int) = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow> 
+    [x = y] (mod n)"
+  by (auto simp add: cong_altdef_int dvd_def)
+
+lemma cong_dvd_eq_nat: "[(x::nat) = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"
+  by (unfold cong_nat_def, auto simp add: dvd_eq_mod_eq_0)
+
+lemma cong_dvd_eq_int: "[(x::int) = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"
+  by (unfold cong_int_def, auto simp add: dvd_eq_mod_eq_0)
+
+lemma cong_mod_nat: "(n::nat) ~= 0 \<Longrightarrow> [a mod n = a] (mod n)" 
+  by (simp add: cong_nat_def)
+
+lemma cong_mod_int: "(n::int) ~= 0 \<Longrightarrow> [a mod n = a] (mod n)" 
+  by (simp add: cong_int_def)
+
+lemma mod_mult_cong_nat: "(a::nat) ~= 0 \<Longrightarrow> b ~= 0 
+    \<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)"
+  by (simp add: cong_nat_def mod_mult2_eq  mod_add_left_eq)
+
+lemma neg_cong_int: "([(a::int) = b] (mod m)) = ([-a = -b] (mod m))"
+  apply (simp add: cong_altdef_int)
+  apply (subst dvd_minus_iff [symmetric])
+  apply (simp add: ring_simps)
+done
+
+lemma cong_modulus_neg_int: "([(a::int) = b] (mod m)) = ([a = b] (mod -m))"
+  by (auto simp add: cong_altdef_int)
+
+lemma mod_mult_cong_int: "(a::int) ~= 0 \<Longrightarrow> b ~= 0 
+    \<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)"
+  apply (case_tac "b > 0")
+  apply (simp add: cong_int_def mod_mod_cancel mod_add_left_eq)
+  apply (subst (1 2) cong_modulus_neg_int)
+  apply (unfold cong_int_def)
+  apply (subgoal_tac "a * b = (-a * -b)")
+  apply (erule ssubst)
+  apply (subst zmod_zmult2_eq)
+  apply (auto simp add: mod_add_left_eq) 
+done
+
+lemma cong_to_1_nat: "([(a::nat) = 1] (mod n)) \<Longrightarrow> (n dvd (a - 1))"
+  apply (case_tac "a = 0")
+  apply force
+  apply (subst (asm) cong_altdef_nat)
+  apply auto
+done
+
+lemma cong_0_1_nat: "[(0::nat) = 1] (mod n) = (n = 1)"
+  by (unfold cong_nat_def, auto)
+
+lemma cong_0_1_int: "[(0::int) = 1] (mod n) = ((n = 1) | (n = -1))"
+  by (unfold cong_int_def, auto simp add: zmult_eq_1_iff)
+
+lemma cong_to_1'_nat: "[(a::nat) = 1] (mod n) \<longleftrightarrow> 
+    a = 0 \<and> n = 1 \<or> (\<exists>m. a = 1 + m * n)"
+  apply (case_tac "n = 1")
+  apply auto [1]
+  apply (drule_tac x = "a - 1" in spec)
+  apply force
+  apply (case_tac "a = 0")
+  apply (auto simp add: cong_0_1_nat) [1]
+  apply (rule iffI)
+  apply (drule cong_to_1_nat)
+  apply (unfold dvd_def)
+  apply auto [1]
+  apply (rule_tac x = k in exI)
+  apply (auto simp add: ring_simps) [1]
+  apply (subst cong_altdef_nat)
+  apply (auto simp add: dvd_def)
+done
+
+lemma cong_le_nat: "(y::nat) <= x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)"
+  apply (subst cong_altdef_nat)
+  apply assumption
+  apply (unfold dvd_def, auto simp add: ring_simps)
+  apply (rule_tac x = k in exI)
+  apply auto
+done
+
+lemma cong_solve_nat: "(a::nat) \<noteq> 0 \<Longrightarrow> EX x. [a * x = gcd a n] (mod n)"
+  apply (case_tac "n = 0")
+  apply force
+  apply (frule bezout_nat [of a n], auto)
+  apply (rule exI, erule ssubst)
+  apply (rule cong_trans_nat)
+  apply (rule cong_add_nat)
+  apply (subst mult_commute)
+  apply (rule cong_mult_self_nat)
+  prefer 2
+  apply simp
+  apply (rule cong_refl_nat)
+  apply (rule cong_refl_nat)
+done
+
+lemma cong_solve_int: "(a::int) \<noteq> 0 \<Longrightarrow> EX x. [a * x = gcd a n] (mod n)"
+  apply (case_tac "n = 0")
+  apply (case_tac "a \<ge> 0")
+  apply auto
+  apply (rule_tac x = "-1" in exI)
+  apply auto
+  apply (insert bezout_int [of a n], auto)
+  apply (rule exI)
+  apply (erule subst)
+  apply (rule cong_trans_int)
+  prefer 2
+  apply (rule cong_add_int)
+  apply (rule cong_refl_int)
+  apply (rule cong_sym_int)
+  apply (rule cong_mult_self_int)
+  apply simp
+  apply (subst mult_commute)
+  apply (rule cong_refl_int)
+done
+  
+lemma cong_solve_dvd_nat: 
+  assumes a: "(a::nat) \<noteq> 0" and b: "gcd a n dvd d"
+  shows "EX x. [a * x = d] (mod n)"
+proof -
+  from cong_solve_nat [OF a] obtain x where 
+      "[a * x = gcd a n](mod n)"
+    by auto
+  hence "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)" 
+    by (elim cong_scalar2_nat)
+  also from b have "(d div gcd a n) * gcd a n = d"
+    by (rule dvd_div_mult_self)
+  also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)"
+    by auto
+  finally show ?thesis
+    by auto
+qed
+
+lemma cong_solve_dvd_int: 
+  assumes a: "(a::int) \<noteq> 0" and b: "gcd a n dvd d"
+  shows "EX x. [a * x = d] (mod n)"
+proof -
+  from cong_solve_int [OF a] obtain x where 
+      "[a * x = gcd a n](mod n)"
+    by auto
+  hence "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)" 
+    by (elim cong_scalar2_int)
+  also from b have "(d div gcd a n) * gcd a n = d"
+    by (rule dvd_div_mult_self)
+  also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)"
+    by auto
+  finally show ?thesis
+    by auto
+qed
+
+lemma cong_solve_coprime_nat: "coprime (a::nat) n \<Longrightarrow> 
+    EX x. [a * x = 1] (mod n)"
+  apply (case_tac "a = 0")
+  apply force
+  apply (frule cong_solve_nat [of a n])
+  apply auto
+done
+
+lemma cong_solve_coprime_int: "coprime (a::int) n \<Longrightarrow> 
+    EX x. [a * x = 1] (mod n)"
+  apply (case_tac "a = 0")
+  apply auto
+  apply (case_tac "n \<ge> 0")
+  apply auto
+  apply (subst cong_int_def, auto)
+  apply (frule cong_solve_int [of a n])
+  apply auto
+done
+
+lemma coprime_iff_invertible_nat: "m > (1::nat) \<Longrightarrow> coprime a m = 
+    (EX x. [a * x = 1] (mod m))"
+  apply (auto intro: cong_solve_coprime_nat)
+  apply (unfold cong_nat_def, auto intro: invertible_coprime_nat)
+done
+
+lemma coprime_iff_invertible_int: "m > (1::int) \<Longrightarrow> coprime a m = 
+    (EX x. [a * x = 1] (mod m))"
+  apply (auto intro: cong_solve_coprime_int)
+  apply (unfold cong_int_def)
+  apply (auto intro: invertible_coprime_int)
+done
+
+lemma coprime_iff_invertible'_int: "m > (1::int) \<Longrightarrow> coprime a m = 
+    (EX x. 0 <= x & x < m & [a * x = 1] (mod m))"
+  apply (subst coprime_iff_invertible_int)
+  apply auto
+  apply (auto simp add: cong_int_def)
+  apply (rule_tac x = "x mod m" in exI)
+  apply (auto simp add: mod_mult_right_eq [symmetric])
+done
+
+
+lemma cong_cong_lcm_nat: "[(x::nat) = y] (mod a) \<Longrightarrow>
+    [x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)"
+  apply (case_tac "y \<le> x")
+  apply (auto simp add: cong_altdef_nat lcm_least_nat) [1]
+  apply (rule cong_sym_nat)
+  apply (subst (asm) (1 2) cong_sym_eq_nat)
+  apply (auto simp add: cong_altdef_nat lcm_least_nat)
+done
+
+lemma cong_cong_lcm_int: "[(x::int) = y] (mod a) \<Longrightarrow>
+    [x = y] (mod b) \<Longrightarrow> [x = y] (mod lcm a b)"
+  by (auto simp add: cong_altdef_int lcm_least_int) [1]
+
+lemma cong_cong_coprime_nat: "coprime a b \<Longrightarrow> [(x::nat) = y] (mod a) \<Longrightarrow>
+    [x = y] (mod b) \<Longrightarrow> [x = y] (mod a * b)"
+  apply (frule (1) cong_cong_lcm_nat)back
+  apply (simp add: lcm_nat_def)
+done
+
+lemma cong_cong_coprime_int: "coprime a b \<Longrightarrow> [(x::int) = y] (mod a) \<Longrightarrow>
+    [x = y] (mod b) \<Longrightarrow> [x = y] (mod a * b)"
+  apply (frule (1) cong_cong_lcm_int)back
+  apply (simp add: lcm_altdef_int cong_abs_int abs_mult [symmetric])
+done
+
+lemma cong_cong_setprod_coprime_nat [rule_format]: "finite A \<Longrightarrow>
+    (ALL i:A. (ALL j:A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>
+    (ALL i:A. [(x::nat) = y] (mod m i)) \<longrightarrow>
+      [x = y] (mod (PROD i:A. m i))"
+  apply (induct set: finite)
+  apply auto
+  apply (rule cong_cong_coprime_nat)
+  apply (subst gcd_commute_nat)
+  apply (rule setprod_coprime_nat)
+  apply auto
+done
+
+lemma cong_cong_setprod_coprime_int [rule_format]: "finite A \<Longrightarrow>
+    (ALL i:A. (ALL j:A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>
+    (ALL i:A. [(x::int) = y] (mod m i)) \<longrightarrow>
+      [x = y] (mod (PROD i:A. m i))"
+  apply (induct set: finite)
+  apply auto
+  apply (rule cong_cong_coprime_int)
+  apply (subst gcd_commute_int)
+  apply (rule setprod_coprime_int)
+  apply auto
+done
+
+lemma binary_chinese_remainder_aux_nat: 
+  assumes a: "coprime (m1::nat) m2"
+  shows "EX b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and>
+    [b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)"
+proof -
+  from cong_solve_coprime_nat [OF a]
+      obtain x1 where one: "[m1 * x1 = 1] (mod m2)"
+    by auto
+  from a have b: "coprime m2 m1" 
+    by (subst gcd_commute_nat)
+  from cong_solve_coprime_nat [OF b]
+      obtain x2 where two: "[m2 * x2 = 1] (mod m1)"
+    by auto
+  have "[m1 * x1 = 0] (mod m1)"
+    by (subst mult_commute, rule cong_mult_self_nat)
+  moreover have "[m2 * x2 = 0] (mod m2)"
+    by (subst mult_commute, rule cong_mult_self_nat)
+  moreover note one two
+  ultimately show ?thesis by blast
+qed
+
+lemma binary_chinese_remainder_aux_int: 
+  assumes a: "coprime (m1::int) m2"
+  shows "EX b1 b2. [b1 = 1] (mod m1) \<and> [b1 = 0] (mod m2) \<and>
+    [b2 = 0] (mod m1) \<and> [b2 = 1] (mod m2)"
+proof -
+  from cong_solve_coprime_int [OF a]
+      obtain x1 where one: "[m1 * x1 = 1] (mod m2)"
+    by auto
+  from a have b: "coprime m2 m1" 
+    by (subst gcd_commute_int)
+  from cong_solve_coprime_int [OF b]
+      obtain x2 where two: "[m2 * x2 = 1] (mod m1)"
+    by auto
+  have "[m1 * x1 = 0] (mod m1)"
+    by (subst mult_commute, rule cong_mult_self_int)
+  moreover have "[m2 * x2 = 0] (mod m2)"
+    by (subst mult_commute, rule cong_mult_self_int)
+  moreover note one two
+  ultimately show ?thesis by blast
+qed
+
+lemma binary_chinese_remainder_nat:
+  assumes a: "coprime (m1::nat) m2"
+  shows "EX x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)"
+proof -
+  from binary_chinese_remainder_aux_nat [OF a] obtain b1 b2
+    where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" and
+          "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)"
+    by blast
+  let ?x = "u1 * b1 + u2 * b2"
+  have "[?x = u1 * 1 + u2 * 0] (mod m1)"
+    apply (rule cong_add_nat)
+    apply (rule cong_scalar2_nat)
+    apply (rule `[b1 = 1] (mod m1)`)
+    apply (rule cong_scalar2_nat)
+    apply (rule `[b2 = 0] (mod m1)`)
+    done
+  hence "[?x = u1] (mod m1)" by simp
+  have "[?x = u1 * 0 + u2 * 1] (mod m2)"
+    apply (rule cong_add_nat)
+    apply (rule cong_scalar2_nat)
+    apply (rule `[b1 = 0] (mod m2)`)
+    apply (rule cong_scalar2_nat)
+    apply (rule `[b2 = 1] (mod m2)`)
+    done
+  hence "[?x = u2] (mod m2)" by simp
+  with `[?x = u1] (mod m1)` show ?thesis by blast
+qed
+
+lemma binary_chinese_remainder_int:
+  assumes a: "coprime (m1::int) m2"
+  shows "EX x. [x = u1] (mod m1) \<and> [x = u2] (mod m2)"
+proof -
+  from binary_chinese_remainder_aux_int [OF a] obtain b1 b2
+    where "[b1 = 1] (mod m1)" and "[b1 = 0] (mod m2)" and
+          "[b2 = 0] (mod m1)" and "[b2 = 1] (mod m2)"
+    by blast
+  let ?x = "u1 * b1 + u2 * b2"
+  have "[?x = u1 * 1 + u2 * 0] (mod m1)"
+    apply (rule cong_add_int)
+    apply (rule cong_scalar2_int)
+    apply (rule `[b1 = 1] (mod m1)`)
+    apply (rule cong_scalar2_int)
+    apply (rule `[b2 = 0] (mod m1)`)
+    done
+  hence "[?x = u1] (mod m1)" by simp
+  have "[?x = u1 * 0 + u2 * 1] (mod m2)"
+    apply (rule cong_add_int)
+    apply (rule cong_scalar2_int)
+    apply (rule `[b1 = 0] (mod m2)`)
+    apply (rule cong_scalar2_int)
+    apply (rule `[b2 = 1] (mod m2)`)
+    done
+  hence "[?x = u2] (mod m2)" by simp
+  with `[?x = u1] (mod m1)` show ?thesis by blast
+qed
+
+lemma cong_modulus_mult_nat: "[(x::nat) = y] (mod m * n) \<Longrightarrow> 
+    [x = y] (mod m)"
+  apply (case_tac "y \<le> x")
+  apply (simp add: cong_altdef_nat)
+  apply (erule dvd_mult_left)
+  apply (rule cong_sym_nat)
+  apply (subst (asm) cong_sym_eq_nat)
+  apply (simp add: cong_altdef_nat) 
+  apply (erule dvd_mult_left)
+done
+
+lemma cong_modulus_mult_int: "[(x::int) = y] (mod m * n) \<Longrightarrow> 
+    [x = y] (mod m)"
+  apply (simp add: cong_altdef_int) 
+  apply (erule dvd_mult_left)
+done
+
+lemma cong_less_modulus_unique_nat: 
+    "[(x::nat) = y] (mod m) \<Longrightarrow> x < m \<Longrightarrow> y < m \<Longrightarrow> x = y"
+  by (simp add: cong_nat_def)
+
+lemma binary_chinese_remainder_unique_nat:
+  assumes a: "coprime (m1::nat) m2" and
+         nz: "m1 \<noteq> 0" "m2 \<noteq> 0"
+  shows "EX! x. x < m1 * m2 \<and> [x = u1] (mod m1) \<and> [x = u2] (mod m2)"
+proof -
+  from binary_chinese_remainder_nat [OF a] obtain y where 
+      "[y = u1] (mod m1)" and "[y = u2] (mod m2)"
+    by blast
+  let ?x = "y mod (m1 * m2)"
+  from nz have less: "?x < m1 * m2"
+    by auto   
+  have one: "[?x = u1] (mod m1)"
+    apply (rule cong_trans_nat)
+    prefer 2
+    apply (rule `[y = u1] (mod m1)`)
+    apply (rule cong_modulus_mult_nat)
+    apply (rule cong_mod_nat)
+    using nz apply auto
+    done
+  have two: "[?x = u2] (mod m2)"
+    apply (rule cong_trans_nat)
+    prefer 2
+    apply (rule `[y = u2] (mod m2)`)
+    apply (subst mult_commute)
+    apply (rule cong_modulus_mult_nat)
+    apply (rule cong_mod_nat)
+    using nz apply auto
+    done
+  have "ALL z. z < m1 * m2 \<and> [z = u1] (mod m1) \<and> [z = u2] (mod m2) \<longrightarrow>
+      z = ?x"
+  proof (clarify)
+    fix z
+    assume "z < m1 * m2"
+    assume "[z = u1] (mod m1)" and  "[z = u2] (mod m2)"
+    have "[?x = z] (mod m1)"
+      apply (rule cong_trans_nat)
+      apply (rule `[?x = u1] (mod m1)`)
+      apply (rule cong_sym_nat)
+      apply (rule `[z = u1] (mod m1)`)
+      done
+    moreover have "[?x = z] (mod m2)"
+      apply (rule cong_trans_nat)
+      apply (rule `[?x = u2] (mod m2)`)
+      apply (rule cong_sym_nat)
+      apply (rule `[z = u2] (mod m2)`)
+      done
+    ultimately have "[?x = z] (mod m1 * m2)"
+      by (auto intro: coprime_cong_mult_nat a)
+    with `z < m1 * m2` `?x < m1 * m2` show "z = ?x"
+      apply (intro cong_less_modulus_unique_nat)
+      apply (auto, erule cong_sym_nat)
+      done
+  qed  
+  with less one two show ?thesis
+    by auto
+ qed
+
+lemma chinese_remainder_aux_nat:
+  fixes A :: "'a set" and
+        m :: "'a \<Rightarrow> nat"
+  assumes fin: "finite A" and
+          cop: "ALL i : A. (ALL j : A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"
+  shows "EX b. (ALL i : A. 
+      [b i = 1] (mod m i) \<and> [b i = 0] (mod (PROD j : A - {i}. m j)))"
+proof (rule finite_set_choice, rule fin, rule ballI)
+  fix i
+  assume "i : A"
+  with cop have "coprime (PROD j : A - {i}. m j) (m i)"
+    by (intro setprod_coprime_nat, auto)
+  hence "EX x. [(PROD j : A - {i}. m j) * x = 1] (mod m i)"
+    by (elim cong_solve_coprime_nat)
+  then obtain x where "[(PROD j : A - {i}. m j) * x = 1] (mod m i)"
+    by auto
+  moreover have "[(PROD j : A - {i}. m j) * x = 0] 
+    (mod (PROD j : A - {i}. m j))"
+    by (subst mult_commute, rule cong_mult_self_nat)
+  ultimately show "\<exists>a. [a = 1] (mod m i) \<and> [a = 0] 
+      (mod setprod m (A - {i}))"
+    by blast
+qed
+
+lemma chinese_remainder_nat:
+  fixes A :: "'a set" and
+        m :: "'a \<Rightarrow> nat" and
+        u :: "'a \<Rightarrow> nat"
+  assumes 
+        fin: "finite A" and
+        cop: "ALL i:A. (ALL j : A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"
+  shows "EX x. (ALL i:A. [x = u i] (mod m i))"
+proof -
+  from chinese_remainder_aux_nat [OF fin cop] obtain b where
+    bprop: "ALL i:A. [b i = 1] (mod m i) \<and> 
+      [b i = 0] (mod (PROD j : A - {i}. m j))"
+    by blast
+  let ?x = "SUM i:A. (u i) * (b i)"
+  show "?thesis"
+  proof (rule exI, clarify)
+    fix i
+    assume a: "i : A"
+    show "[?x = u i] (mod m i)" 
+    proof -
+      from fin a have "?x = (SUM j:{i}. u j * b j) + 
+          (SUM j:A-{i}. u j * b j)"
+        by (subst setsum_Un_disjoint [symmetric], auto intro: setsum_cong)
+      hence "[?x = u i * b i + (SUM j:A-{i}. u j * b j)] (mod m i)"
+        by auto
+      also have "[u i * b i + (SUM j:A-{i}. u j * b j) =
+                  u i * 1 + (SUM j:A-{i}. u j * 0)] (mod m i)"
+        apply (rule cong_add_nat)
+        apply (rule cong_scalar2_nat)
+        using bprop a apply blast
+        apply (rule cong_setsum_nat)
+        apply (rule cong_scalar2_nat)
+        using bprop apply auto
+        apply (rule cong_dvd_modulus_nat)
+        apply (drule (1) bspec)
+        apply (erule conjE)
+        apply assumption
+        apply (rule dvd_setprod)
+        using fin a apply auto
+        done
+      finally show ?thesis
+        by simp
+    qed
+  qed
+qed
+
+lemma coprime_cong_prod_nat [rule_format]: "finite A \<Longrightarrow> 
+    (ALL i: A. (ALL j: A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))) \<longrightarrow>
+      (ALL i: A. [(x::nat) = y] (mod m i)) \<longrightarrow>
+         [x = y] (mod (PROD i:A. m i))" 
+  apply (induct set: finite)
+  apply auto
+  apply (erule (1) coprime_cong_mult_nat)
+  apply (subst gcd_commute_nat)
+  apply (rule setprod_coprime_nat)
+  apply auto
+done
+
+lemma chinese_remainder_unique_nat:
+  fixes A :: "'a set" and
+        m :: "'a \<Rightarrow> nat" and
+        u :: "'a \<Rightarrow> nat"
+  assumes 
+        fin: "finite A" and
+         nz: "ALL i:A. m i \<noteq> 0" and
+        cop: "ALL i:A. (ALL j : A. i \<noteq> j \<longrightarrow> coprime (m i) (m j))"
+  shows "EX! x. x < (PROD i:A. m i) \<and> (ALL i:A. [x = u i] (mod m i))"
+proof -
+  from chinese_remainder_nat [OF fin cop] obtain y where
+      one: "(ALL i:A. [y = u i] (mod m i))" 
+    by blast
+  let ?x = "y mod (PROD i:A. m i)"
+  from fin nz have prodnz: "(PROD i:A. m i) \<noteq> 0"
+    by auto
+  hence less: "?x < (PROD i:A. m i)"
+    by auto
+  have cong: "ALL i:A. [?x = u i] (mod m i)"
+    apply auto
+    apply (rule cong_trans_nat)
+    prefer 2
+    using one apply auto
+    apply (rule cong_dvd_modulus_nat)
+    apply (rule cong_mod_nat)
+    using prodnz apply auto
+    apply (rule dvd_setprod)
+    apply (rule fin)
+    apply assumption
+    done
+  have unique: "ALL z. z < (PROD i:A. m i) \<and> 
+      (ALL i:A. [z = u i] (mod m i)) \<longrightarrow> z = ?x"
+  proof (clarify)
+    fix z
+    assume zless: "z < (PROD i:A. m i)"
+    assume zcong: "(ALL i:A. [z = u i] (mod m i))"
+    have "ALL i:A. [?x = z] (mod m i)"
+      apply clarify     
+      apply (rule cong_trans_nat)
+      using cong apply (erule bspec)
+      apply (rule cong_sym_nat)
+      using zcong apply auto
+      done
+    with fin cop have "[?x = z] (mod (PROD i:A. m i))"
+      by (intro coprime_cong_prod_nat, auto)
+    with zless less show "z = ?x"
+      apply (intro cong_less_modulus_unique_nat)
+      apply (auto, erule cong_sym_nat)
+      done
+  qed 
+  from less cong unique show ?thesis
+    by blast  
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Number_Theory/Fib.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,319 @@
+(*  Title:      Fib.thy
+    Authors:    Lawrence C. Paulson, Jeremy Avigad
+
+
+Defines the fibonacci function.
+
+The original "Fib" is due to Lawrence C. Paulson, and was adapted by
+Jeremy Avigad.
+*)
+
+
+header {* Fib *}
+
+theory Fib
+imports Binomial
+begin
+
+
+subsection {* Main definitions *}
+
+class fib =
+
+fixes 
+  fib :: "'a \<Rightarrow> 'a"
+
+
+(* definition for the natural numbers *)
+
+instantiation nat :: fib
+
+begin 
+
+fun 
+  fib_nat :: "nat \<Rightarrow> nat"
+where
+  "fib_nat n =
+   (if n = 0 then 0 else
+   (if n = 1 then 1 else
+     fib (n - 1) + fib (n - 2)))"
+
+instance proof qed
+
+end
+
+(* definition for the integers *)
+
+instantiation int :: fib
+
+begin 
+
+definition
+  fib_int :: "int \<Rightarrow> int"
+where  
+  "fib_int n = (if n >= 0 then int (fib (nat n)) else 0)"
+
+instance proof qed
+
+end
+
+
+subsection {* Set up Transfer *}
+
+
+lemma transfer_nat_int_fib:
+  "(x::int) >= 0 \<Longrightarrow> fib (nat x) = nat (fib x)"
+  unfolding fib_int_def by auto
+
+lemma transfer_nat_int_fib_closure:
+  "n >= (0::int) \<Longrightarrow> fib n >= 0"
+  by (auto simp add: fib_int_def)
+
+declare TransferMorphism_nat_int[transfer add return: 
+    transfer_nat_int_fib transfer_nat_int_fib_closure]
+
+lemma transfer_int_nat_fib:
+  "fib (int n) = int (fib n)"
+  unfolding fib_int_def by auto
+
+lemma transfer_int_nat_fib_closure:
+  "is_nat n \<Longrightarrow> fib n >= 0"
+  unfolding fib_int_def by auto
+
+declare TransferMorphism_int_nat[transfer add return: 
+    transfer_int_nat_fib transfer_int_nat_fib_closure]
+
+
+subsection {* Fibonacci numbers *}
+
+lemma fib_0_nat [simp]: "fib (0::nat) = 0"
+  by simp
+
+lemma fib_0_int [simp]: "fib (0::int) = 0"
+  unfolding fib_int_def by simp
+
+lemma fib_1_nat [simp]: "fib (1::nat) = 1"
+  by simp
+
+lemma fib_Suc_0_nat [simp]: "fib (Suc 0) = Suc 0"
+  by simp
+
+lemma fib_1_int [simp]: "fib (1::int) = 1"
+  unfolding fib_int_def by simp
+
+lemma fib_reduce_nat: "(n::nat) >= 2 \<Longrightarrow> fib n = fib (n - 1) + fib (n - 2)"
+  by simp
+
+declare fib_nat.simps [simp del]
+
+lemma fib_reduce_int: "(n::int) >= 2 \<Longrightarrow> fib n = fib (n - 1) + fib (n - 2)"
+  unfolding fib_int_def
+  by (auto simp add: fib_reduce_nat nat_diff_distrib)
+
+lemma fib_neg_int [simp]: "(n::int) < 0 \<Longrightarrow> fib n = 0"
+  unfolding fib_int_def by auto
+
+lemma fib_2_nat [simp]: "fib (2::nat) = 1"
+  by (subst fib_reduce_nat, auto)
+
+lemma fib_2_int [simp]: "fib (2::int) = 1"
+  by (subst fib_reduce_int, auto)
+
+lemma fib_plus_2_nat: "fib ((n::nat) + 2) = fib (n + 1) + fib n"
+  by (subst fib_reduce_nat, auto simp add: One_nat_def)
+(* the need for One_nat_def is due to the natdiff_cancel_numerals
+   procedure *)
+
+lemma fib_induct_nat: "P (0::nat) \<Longrightarrow> P (1::nat) \<Longrightarrow> 
+    (!!n. P n \<Longrightarrow> P (n + 1) \<Longrightarrow> P (n + 2)) \<Longrightarrow> P n"
+  apply (atomize, induct n rule: nat_less_induct)
+  apply auto
+  apply (case_tac "n = 0", force)
+  apply (case_tac "n = 1", force)
+  apply (subgoal_tac "n >= 2")
+  apply (frule_tac x = "n - 1" in spec)
+  apply (drule_tac x = "n - 2" in spec)
+  apply (drule_tac x = "n - 2" in spec)
+  apply auto
+  apply (auto simp add: One_nat_def) (* again, natdiff_cancel *)
+done
+
+lemma fib_add_nat: "fib ((n::nat) + k + 1) = fib (k + 1) * fib (n + 1) + 
+    fib k * fib n"
+  apply (induct n rule: fib_induct_nat)
+  apply auto
+  apply (subst fib_reduce_nat)
+  apply (auto simp add: ring_simps)
+  apply (subst (1 3 5) fib_reduce_nat)
+  apply (auto simp add: ring_simps Suc_eq_plus1)
+(* hmmm. Why doesn't "n + (1 + (1 + k))" simplify to "n + k + 2"? *)
+  apply (subgoal_tac "n + (k + 2) = n + (1 + (1 + k))")
+  apply (erule ssubst) back back
+  apply (erule ssubst) back 
+  apply auto
+done
+
+lemma fib_add'_nat: "fib (n + Suc k) = fib (Suc k) * fib (Suc n) + 
+    fib k * fib n"
+  using fib_add_nat by (auto simp add: One_nat_def)
+
+
+(* transfer from nats to ints *)
+lemma fib_add_int [rule_format]: "(n::int) >= 0 \<Longrightarrow> k >= 0 \<Longrightarrow>
+    fib (n + k + 1) = fib (k + 1) * fib (n + 1) + 
+    fib k * fib n "
+
+  by (rule fib_add_nat [transferred])
+
+lemma fib_neq_0_nat: "(n::nat) > 0 \<Longrightarrow> fib n ~= 0"
+  apply (induct n rule: fib_induct_nat)
+  apply (auto simp add: fib_plus_2_nat)
+done
+
+lemma fib_gr_0_nat: "(n::nat) > 0 \<Longrightarrow> fib n > 0"
+  by (frule fib_neq_0_nat, simp)
+
+lemma fib_gr_0_int: "(n::int) > 0 \<Longrightarrow> fib n > 0"
+  unfolding fib_int_def by (simp add: fib_gr_0_nat)
+
+text {*
+  \medskip Concrete Mathematics, page 278: Cassini's identity.  The proof is
+  much easier using integers, not natural numbers!
+*}
+
+lemma fib_Cassini_aux_int: "fib (int n + 2) * fib (int n) - 
+    (fib (int n + 1))^2 = (-1)^(n + 1)"
+  apply (induct n)
+  apply (auto simp add: ring_simps power2_eq_square fib_reduce_int
+      power_add)
+done
+
+lemma fib_Cassini_int: "n >= 0 \<Longrightarrow> fib (n + 2) * fib n - 
+    (fib (n + 1))^2 = (-1)^(nat n + 1)"
+  by (insert fib_Cassini_aux_int [of "nat n"], auto)
+
+(*
+lemma fib_Cassini'_int: "n >= 0 \<Longrightarrow> fib (n + 2) * fib n = 
+    (fib (n + 1))^2 + (-1)^(nat n + 1)"
+  by (frule fib_Cassini_int, simp) 
+*)
+
+lemma fib_Cassini'_int: "n >= 0 \<Longrightarrow> fib ((n::int) + 2) * fib n =
+  (if even n then tsub ((fib (n + 1))^2) 1
+   else (fib (n + 1))^2 + 1)"
+  apply (frule fib_Cassini_int, auto simp add: pos_int_even_equiv_nat_even)
+  apply (subst tsub_eq)
+  apply (insert fib_gr_0_int [of "n + 1"], force)
+  apply auto
+done
+
+lemma fib_Cassini_nat: "fib ((n::nat) + 2) * fib n =
+  (if even n then (fib (n + 1))^2 - 1
+   else (fib (n + 1))^2 + 1)"
+
+  by (rule fib_Cassini'_int [transferred, of n], auto)
+
+
+text {* \medskip Toward Law 6.111 of Concrete Mathematics *}
+
+lemma coprime_fib_plus_1_nat: "coprime (fib (n::nat)) (fib (n + 1))"
+  apply (induct n rule: fib_induct_nat)
+  apply auto
+  apply (subst (2) fib_reduce_nat)
+  apply (auto simp add: Suc_eq_plus1) (* again, natdiff_cancel *)
+  apply (subst add_commute, auto)
+  apply (subst gcd_commute_nat, auto simp add: ring_simps)
+done
+
+lemma coprime_fib_Suc_nat: "coprime (fib n) (fib (Suc n))"
+  using coprime_fib_plus_1_nat by (simp add: One_nat_def)
+
+lemma coprime_fib_plus_1_int: 
+    "n >= 0 \<Longrightarrow> coprime (fib (n::int)) (fib (n + 1))"
+  by (erule coprime_fib_plus_1_nat [transferred])
+
+lemma gcd_fib_add_nat: "gcd (fib (m::nat)) (fib (n + m)) = gcd (fib m) (fib n)"
+  apply (simp add: gcd_commute_nat [of "fib m"])
+  apply (rule cases_nat [of _ m])
+  apply simp
+  apply (subst add_assoc [symmetric])
+  apply (simp add: fib_add_nat)
+  apply (subst gcd_commute_nat)
+  apply (subst mult_commute)
+  apply (subst gcd_add_mult_nat)
+  apply (subst gcd_commute_nat)
+  apply (rule gcd_mult_cancel_nat)
+  apply (rule coprime_fib_plus_1_nat)
+done
+
+lemma gcd_fib_add_int [rule_format]: "m >= 0 \<Longrightarrow> n >= 0 \<Longrightarrow> 
+    gcd (fib (m::int)) (fib (n + m)) = gcd (fib m) (fib n)"
+  by (erule gcd_fib_add_nat [transferred])
+
+lemma gcd_fib_diff_nat: "(m::nat) \<le> n \<Longrightarrow> 
+    gcd (fib m) (fib (n - m)) = gcd (fib m) (fib n)"
+  by (simp add: gcd_fib_add_nat [symmetric, of _ "n-m"])
+
+lemma gcd_fib_diff_int: "0 <= (m::int) \<Longrightarrow> m \<le> n \<Longrightarrow> 
+    gcd (fib m) (fib (n - m)) = gcd (fib m) (fib n)"
+  by (simp add: gcd_fib_add_int [symmetric, of _ "n-m"])
+
+lemma gcd_fib_mod_nat: "0 < (m::nat) \<Longrightarrow> 
+    gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
+proof (induct n rule: less_induct)
+  case (less n)
+  from less.prems have pos_m: "0 < m" .
+  show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
+  proof (cases "m < n")
+    case True note m_n = True
+    then have m_n': "m \<le> n" by auto
+    with pos_m have pos_n: "0 < n" by auto
+    with pos_m m_n have diff: "n - m < n" by auto
+    have "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib ((n - m) mod m))"
+    by (simp add: mod_if [of n]) (insert m_n, auto)
+    also have "\<dots> = gcd (fib m)  (fib (n - m))" 
+      by (simp add: less.hyps diff pos_m)
+    also have "\<dots> = gcd (fib m) (fib n)" by (simp add: gcd_fib_diff_nat m_n')
+    finally show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)" .
+  next
+    case False then show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
+    by (cases "m = n") auto
+  qed
+qed
+
+lemma gcd_fib_mod_int: 
+  assumes "0 < (m::int)" and "0 <= n"
+  shows "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
+
+  apply (rule gcd_fib_mod_nat [transferred])
+  using prems apply auto
+done
+
+lemma fib_gcd_nat: "fib (gcd (m::nat) n) = gcd (fib m) (fib n)"  
+    -- {* Law 6.111 *}
+  apply (induct m n rule: gcd_nat_induct)
+  apply (simp_all add: gcd_non_0_nat gcd_commute_nat gcd_fib_mod_nat)
+done
+
+lemma fib_gcd_int: "m >= 0 \<Longrightarrow> n >= 0 \<Longrightarrow>
+    fib (gcd (m::int) n) = gcd (fib m) (fib n)"
+  by (erule fib_gcd_nat [transferred])
+
+lemma atMost_plus_one_nat: "{..(k::nat) + 1} = insert (k + 1) {..k}" 
+  by auto
+
+theorem fib_mult_eq_setsum_nat:
+    "fib ((n::nat) + 1) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
+  apply (induct n)
+  apply (auto simp add: atMost_plus_one_nat fib_plus_2_nat ring_simps)
+done
+
+theorem fib_mult_eq_setsum'_nat:
+    "fib (Suc n) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
+  using fib_mult_eq_setsum_nat by (simp add: One_nat_def)
+
+theorem fib_mult_eq_setsum_int [rule_format]:
+    "n >= 0 \<Longrightarrow> fib ((n::int) + 1) * fib n = (\<Sum>k \<in> {0..n}. fib k * fib k)"
+  by (erule fib_mult_eq_setsum_nat [transferred])
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Number_Theory/MiscAlgebra.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,355 @@
+(*  Title:      MiscAlgebra.thy
+    Author:     Jeremy Avigad
+
+These are things that can be added to the Algebra library.
+*)
+
+theory MiscAlgebra
+imports
+  "~~/src/HOL/Algebra/Ring"
+  "~~/src/HOL/Algebra/FiniteProduct"
+begin;
+
+(* finiteness stuff *)
+
+lemma bounded_set1_int [intro]: "finite {(x::int). a < x & x < b & P x}" 
+  apply (subgoal_tac "{x. a < x & x < b & P x} <= {a<..<b}")
+  apply (erule finite_subset)
+  apply auto
+done
+
+
+(* The rest is for the algebra libraries *)
+
+(* These go in Group.thy. *)
+
+(*
+  Show that the units in any monoid give rise to a group.
+
+  The file Residues.thy provides some infrastructure to use
+  facts about the unit group within the ring locale.
+*)
+
+
+constdefs 
+  units_of :: "('a, 'b) monoid_scheme => 'a monoid"
+  "units_of G == (| carrier = Units G,
+     Group.monoid.mult = Group.monoid.mult G,
+     one  = one G |)";
+
+(*
+
+lemma (in monoid) Units_mult_closed [intro]:
+  "x : Units G ==> y : Units G ==> x \<otimes> y : Units G"
+  apply (unfold Units_def)
+  apply (clarsimp)
+  apply (rule_tac x = "xaa \<otimes> xa" in bexI)
+  apply auto
+  apply (subst m_assoc)
+  apply auto
+  apply (subst (2) m_assoc [symmetric])
+  apply auto
+  apply (subst m_assoc)
+  apply auto
+  apply (subst (2) m_assoc [symmetric])
+  apply auto
+done
+
+*)
+
+lemma (in monoid) units_group: "group(units_of G)"
+  apply (unfold units_of_def)
+  apply (rule groupI)
+  apply auto
+  apply (subst m_assoc)
+  apply auto
+  apply (rule_tac x = "inv x" in bexI)
+  apply auto
+done
+
+lemma (in comm_monoid) units_comm_group: "comm_group(units_of G)"
+  apply (rule group.group_comm_groupI)
+  apply (rule units_group)
+  apply (insert prems)
+  apply (unfold units_of_def Units_def comm_monoid_def comm_monoid_axioms_def)
+  apply auto;
+done;
+
+lemma units_of_carrier: "carrier (units_of G) = Units G"
+  by (unfold units_of_def, auto)
+
+lemma units_of_mult: "mult(units_of G) = mult G"
+  by (unfold units_of_def, auto)
+
+lemma units_of_one: "one(units_of G) = one G"
+  by (unfold units_of_def, auto)
+
+lemma (in monoid) units_of_inv: "x : Units G ==> 
+    m_inv (units_of G) x = m_inv G x"
+  apply (rule sym)
+  apply (subst m_inv_def)
+  apply (rule the1_equality)
+  apply (rule ex_ex1I)
+  apply (subst (asm) Units_def)
+  apply auto
+  apply (erule inv_unique)
+  apply auto
+  apply (rule Units_closed)
+  apply (simp_all only: units_of_carrier [symmetric])
+  apply (insert units_group)
+  apply auto
+  apply (subst units_of_mult [symmetric])
+  apply (subst units_of_one [symmetric])
+  apply (erule group.r_inv, assumption)
+  apply (subst units_of_mult [symmetric])
+  apply (subst units_of_one [symmetric])
+  apply (erule group.l_inv, assumption)
+done
+
+lemma (in group) inj_on_const_mult: "a: (carrier G) ==> 
+    inj_on (%x. a \<otimes> x) (carrier G)"
+  by (unfold inj_on_def, auto)
+
+lemma (in group) surj_const_mult: "a : (carrier G) ==>
+    (%x. a \<otimes> x) ` (carrier G) = (carrier G)" 
+  apply (auto simp add: image_def)
+  apply (rule_tac x = "(m_inv G a) \<otimes> x" in bexI)
+  apply auto
+(* auto should get this. I suppose we need "comm_monoid_simprules"
+   for mult_ac rewriting. *)
+  apply (subst m_assoc [symmetric])
+  apply auto
+done
+
+lemma (in group) l_cancel_one [simp]: "x : carrier G \<Longrightarrow> a : carrier G \<Longrightarrow>
+    (x \<otimes> a = x) = (a = one G)"
+  apply auto
+  apply (subst l_cancel [symmetric])
+  prefer 4
+  apply (erule ssubst)
+  apply auto
+done
+
+lemma (in group) r_cancel_one [simp]: "x : carrier G \<Longrightarrow> a : carrier G \<Longrightarrow>
+    (a \<otimes> x = x) = (a = one G)"
+  apply auto
+  apply (subst r_cancel [symmetric])
+  prefer 4
+  apply (erule ssubst)
+  apply auto
+done
+
+(* Is there a better way to do this? *)
+
+lemma (in group) l_cancel_one' [simp]: "x : carrier G \<Longrightarrow> a : carrier G \<Longrightarrow>
+    (x = x \<otimes> a) = (a = one G)"
+  by (subst eq_commute, simp)
+
+lemma (in group) r_cancel_one' [simp]: "x : carrier G \<Longrightarrow> a : carrier G \<Longrightarrow>
+    (x = a \<otimes> x) = (a = one G)"
+  by (subst eq_commute, simp)
+
+(* This should be generalized to arbitrary groups, not just commutative
+   ones, using Lagrange's theorem. *)
+
+lemma (in comm_group) power_order_eq_one:
+    assumes fin [simp]: "finite (carrier G)"
+        and a [simp]: "a : carrier G" 
+      shows "a (^) card(carrier G) = one G" 
+proof -
+  have "(\<Otimes>x:carrier G. x) = (\<Otimes>x:carrier G. a \<otimes> x)"
+    by (subst (2) finprod_reindex [symmetric], 
+      auto simp add: Pi_def inj_on_const_mult surj_const_mult)
+  also have "\<dots> = (\<Otimes>x:carrier G. a) \<otimes> (\<Otimes>x:carrier G. x)"
+    by (auto simp add: finprod_multf Pi_def)
+  also have "(\<Otimes>x:carrier G. a) = a (^) card(carrier G)"
+    by (auto simp add: finprod_const)
+  finally show ?thesis
+(* uses the preceeding lemma *)
+    by auto
+qed
+
+
+(* Miscellaneous *)
+
+lemma (in cring) field_intro2: "\<zero>\<^bsub>R\<^esub> ~= \<one>\<^bsub>R\<^esub> \<Longrightarrow> ALL x : carrier R - {\<zero>\<^bsub>R\<^esub>}.
+    x : Units R \<Longrightarrow> field R"
+  apply (unfold_locales)
+  apply (insert prems, auto)
+  apply (rule trans)
+  apply (subgoal_tac "a = (a \<otimes> b) \<otimes> inv b")
+  apply assumption
+  apply (subst m_assoc) 
+  apply (auto simp add: Units_r_inv)
+  apply (unfold Units_def)
+  apply auto
+done
+
+lemma (in monoid) inv_char: "x : carrier G \<Longrightarrow> y : carrier G \<Longrightarrow>
+  x \<otimes> y = \<one> \<Longrightarrow> y \<otimes> x = \<one> \<Longrightarrow> inv x = y"
+  apply (subgoal_tac "x : Units G")
+  apply (subgoal_tac "y = inv x \<otimes> \<one>")
+  apply simp
+  apply (erule subst)
+  apply (subst m_assoc [symmetric])
+  apply auto
+  apply (unfold Units_def)
+  apply auto
+done
+
+lemma (in comm_monoid) comm_inv_char: "x : carrier G \<Longrightarrow> y : carrier G \<Longrightarrow>
+  x \<otimes> y = \<one> \<Longrightarrow> inv x = y"
+  apply (rule inv_char)
+  apply auto
+  apply (subst m_comm, auto) 
+done
+
+lemma (in ring) inv_neg_one [simp]: "inv (\<ominus> \<one>) = \<ominus> \<one>"  
+  apply (rule inv_char)
+  apply (auto simp add: l_minus r_minus)
+done
+
+lemma (in monoid) inv_eq_imp_eq: "x : Units G \<Longrightarrow> y : Units G \<Longrightarrow> 
+    inv x = inv y \<Longrightarrow> x = y"
+  apply (subgoal_tac "inv(inv x) = inv(inv y)")
+  apply (subst (asm) Units_inv_inv)+
+  apply auto
+done
+
+lemma (in ring) Units_minus_one_closed [intro]: "\<ominus> \<one> : Units R"
+  apply (unfold Units_def)
+  apply auto
+  apply (rule_tac x = "\<ominus> \<one>" in bexI)
+  apply auto
+  apply (simp add: l_minus r_minus)
+done
+
+lemma (in monoid) inv_one [simp]: "inv \<one> = \<one>"
+  apply (rule inv_char)
+  apply auto
+done
+
+lemma (in ring) inv_eq_neg_one_eq: "x : Units R \<Longrightarrow> (inv x = \<ominus> \<one>) = (x = \<ominus> \<one>)"
+  apply auto
+  apply (subst Units_inv_inv [symmetric])
+  apply auto
+done
+
+lemma (in monoid) inv_eq_one_eq: "x : Units G \<Longrightarrow> (inv x = \<one>) = (x = \<one>)"
+  apply auto
+  apply (subst Units_inv_inv [symmetric])
+  apply auto
+done
+
+
+(* This goes in FiniteProduct *)
+
+lemma (in comm_monoid) finprod_UN_disjoint:
+  "finite I \<Longrightarrow> (ALL i:I. finite (A i)) \<longrightarrow> (ALL i:I. ALL j:I. i ~= j \<longrightarrow>
+     (A i) Int (A j) = {}) \<longrightarrow>
+      (ALL i:I. ALL x: (A i). g x : carrier G) \<longrightarrow>
+        finprod G g (UNION I A) = finprod G (%i. finprod G g (A i)) I"
+  apply (induct set: finite)
+  apply force
+  apply clarsimp
+  apply (subst finprod_Un_disjoint)
+  apply blast
+  apply (erule finite_UN_I)
+  apply blast
+  apply (fastsimp)
+  apply (auto intro!: funcsetI finprod_closed) 
+done
+
+lemma (in comm_monoid) finprod_Union_disjoint:
+  "[| finite C; (ALL A:C. finite A & (ALL x:A. f x : carrier G));
+      (ALL A:C. ALL B:C. A ~= B --> A Int B = {}) |] 
+   ==> finprod G f (Union C) = finprod G (finprod G f) C" 
+  apply (frule finprod_UN_disjoint [of C id f])
+  apply (unfold Union_def id_def, auto)
+done
+
+lemma (in comm_monoid) finprod_one [rule_format]: 
+  "finite A \<Longrightarrow> (ALL x:A. f x = \<one>) \<longrightarrow>
+     finprod G f A = \<one>"
+by (induct set: finite) auto
+
+
+(* need better simplification rules for rings *)
+(* the next one holds more generally for abelian groups *)
+
+lemma (in cring) sum_zero_eq_neg:
+  "x : carrier R \<Longrightarrow> y : carrier R \<Longrightarrow> x \<oplus> y = \<zero> \<Longrightarrow> x = \<ominus> y"
+  apply (subgoal_tac "\<ominus> y = \<zero> \<oplus> \<ominus> y") 
+  apply (erule ssubst)back
+  apply (erule subst)
+  apply (simp add: ring_simprules)+
+done
+
+(* there's a name conflict -- maybe "domain" should be
+   "integral_domain" *)
+
+lemma (in Ring.domain) square_eq_one: 
+  fixes x
+  assumes [simp]: "x : carrier R" and
+    "x \<otimes> x = \<one>"
+  shows "x = \<one> | x = \<ominus>\<one>"
+proof -
+  have "(x \<oplus> \<one>) \<otimes> (x \<oplus> \<ominus> \<one>) = x \<otimes> x \<oplus> \<ominus> \<one>"
+    by (simp add: ring_simprules)
+  also with `x \<otimes> x = \<one>` have "\<dots> = \<zero>"
+    by (simp add: ring_simprules)
+  finally have "(x \<oplus> \<one>) \<otimes> (x \<oplus> \<ominus> \<one>) = \<zero>" .
+  hence "(x \<oplus> \<one>) = \<zero> | (x \<oplus> \<ominus> \<one>) = \<zero>"
+    by (intro integral, auto)
+  thus ?thesis
+    apply auto
+    apply (erule notE)
+    apply (rule sum_zero_eq_neg)
+    apply auto
+    apply (subgoal_tac "x = \<ominus> (\<ominus> \<one>)")
+    apply (simp add: ring_simprules) 
+    apply (rule sum_zero_eq_neg)
+    apply auto
+    done
+qed
+
+lemma (in Ring.domain) inv_eq_self: "x : Units R \<Longrightarrow>
+    x = inv x \<Longrightarrow> x = \<one> | x = \<ominus> \<one>"
+  apply (rule square_eq_one)
+  apply auto
+  apply (erule ssubst)back
+  apply (erule Units_r_inv)
+done
+
+
+(*
+  The following translates theorems about groups to the facts about
+  the units of a ring. (The list should be expanded as more things are
+  needed.)
+*)
+
+lemma (in ring) finite_ring_finite_units [intro]: "finite (carrier R) \<Longrightarrow> 
+    finite (Units R)"
+  by (rule finite_subset, auto)
+
+(* this belongs with MiscAlgebra.thy *)
+lemma (in monoid) units_of_pow: 
+    "x : Units G \<Longrightarrow> x (^)\<^bsub>units_of G\<^esub> (n::nat) = x (^)\<^bsub>G\<^esub> n"
+  apply (induct n)
+  apply (auto simp add: units_group group.is_monoid  
+    monoid.nat_pow_0 monoid.nat_pow_Suc units_of_one units_of_mult
+    One_nat_def)
+done
+
+lemma (in cring) units_power_order_eq_one: "finite (Units R) \<Longrightarrow> a : Units R
+    \<Longrightarrow> a (^) card(Units R) = \<one>"
+  apply (subst units_of_carrier [symmetric])
+  apply (subst units_of_one [symmetric])
+  apply (subst units_of_pow [symmetric])
+  apply assumption
+  apply (rule comm_group.power_order_eq_one)
+  apply (rule units_comm_group)
+  apply (unfold units_of_def, auto)
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Number_Theory/Number_Theory.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,8 @@
+
+header {* Comprehensive number theory *}
+
+theory Number_Theory
+imports Fib Residues
+begin
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Number_Theory/Primes.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,423 @@
+(*  Authors:    Christophe Tabacznyj, Lawrence C. Paulson, Amine Chaieb,
+                Thomas M. Rasmussen, Jeremy Avigad, Tobias Nipkow
+
+
+This file deals with properties of primes. Definitions and lemmas are
+proved uniformly for the natural numbers and integers.
+
+This file combines and revises a number of prior developments.
+
+The original theories "GCD" and "Primes" were by Christophe Tabacznyj
+and Lawrence C. Paulson, based on \cite{davenport92}. They introduced
+gcd, lcm, and prime for the natural numbers.
+
+The original theory "IntPrimes" was by Thomas M. Rasmussen, and
+extended gcd, lcm, primes to the integers. Amine Chaieb provided
+another extension of the notions to the integers, and added a number
+of results to "Primes" and "GCD". IntPrimes also defined and developed
+the congruence relations on the integers. The notion was extended to
+the natural numbers by Chiaeb.
+
+Jeremy Avigad combined all of these, made everything uniform for the
+natural numbers and the integers, and added a number of new theorems.
+
+Tobias Nipkow cleaned up a lot.
+*)
+
+
+header {* Primes *}
+
+theory Primes
+imports GCD
+begin
+
+declare One_nat_def [simp del]
+
+class prime = one +
+
+fixes
+  prime :: "'a \<Rightarrow> bool"
+
+instantiation nat :: prime
+
+begin
+
+definition
+  prime_nat :: "nat \<Rightarrow> bool"
+where
+  [code del]: "prime_nat p = (1 < p \<and> (\<forall>m. m dvd p --> m = 1 \<or> m = p))"
+
+instance proof qed
+
+end
+
+instantiation int :: prime
+
+begin
+
+definition
+  prime_int :: "int \<Rightarrow> bool"
+where
+  [code del]: "prime_int p = prime (nat p)"
+
+instance proof qed
+
+end
+
+
+subsection {* Set up Transfer *}
+
+
+lemma transfer_nat_int_prime:
+  "(x::int) >= 0 \<Longrightarrow> prime (nat x) = prime x"
+  unfolding gcd_int_def lcm_int_def prime_int_def
+  by auto
+
+declare TransferMorphism_nat_int[transfer add return:
+    transfer_nat_int_prime]
+
+lemma transfer_int_nat_prime:
+  "prime (int x) = prime x"
+  by (unfold gcd_int_def lcm_int_def prime_int_def, auto)
+
+declare TransferMorphism_int_nat[transfer add return:
+    transfer_int_nat_prime]
+
+
+subsection {* Primes *}
+
+lemma prime_odd_nat: "prime (p::nat) \<Longrightarrow> p > 2 \<Longrightarrow> odd p"
+  unfolding prime_nat_def
+  apply (subst even_mult_two_ex)
+  apply clarify
+  apply (drule_tac x = 2 in spec)
+  apply auto
+done
+
+lemma prime_odd_int: "prime (p::int) \<Longrightarrow> p > 2 \<Longrightarrow> odd p"
+  unfolding prime_int_def
+  apply (frule prime_odd_nat)
+  apply (auto simp add: even_nat_def)
+done
+
+(* FIXME Is there a better way to handle these, rather than making them elim rules? *)
+
+lemma prime_ge_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= 0"
+  by (unfold prime_nat_def, auto)
+
+lemma prime_gt_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p > 0"
+  by (unfold prime_nat_def, auto)
+
+lemma prime_ge_1_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= 1"
+  by (unfold prime_nat_def, auto)
+
+lemma prime_gt_1_nat [elim]: "prime (p::nat) \<Longrightarrow> p > 1"
+  by (unfold prime_nat_def, auto)
+
+lemma prime_ge_Suc_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= Suc 0"
+  by (unfold prime_nat_def, auto)
+
+lemma prime_gt_Suc_0_nat [elim]: "prime (p::nat) \<Longrightarrow> p > Suc 0"
+  by (unfold prime_nat_def, auto)
+
+lemma prime_ge_2_nat [elim]: "prime (p::nat) \<Longrightarrow> p >= 2"
+  by (unfold prime_nat_def, auto)
+
+lemma prime_ge_0_int [elim]: "prime (p::int) \<Longrightarrow> p >= 0"
+  by (unfold prime_int_def prime_nat_def) auto
+
+lemma prime_gt_0_int [elim]: "prime (p::int) \<Longrightarrow> p > 0"
+  by (unfold prime_int_def prime_nat_def, auto)
+
+lemma prime_ge_1_int [elim]: "prime (p::int) \<Longrightarrow> p >= 1"
+  by (unfold prime_int_def prime_nat_def, auto)
+
+lemma prime_gt_1_int [elim]: "prime (p::int) \<Longrightarrow> p > 1"
+  by (unfold prime_int_def prime_nat_def, auto)
+
+lemma prime_ge_2_int [elim]: "prime (p::int) \<Longrightarrow> p >= 2"
+  by (unfold prime_int_def prime_nat_def, auto)
+
+
+lemma prime_int_altdef: "prime (p::int) = (1 < p \<and> (\<forall>m \<ge> 0. m dvd p \<longrightarrow>
+    m = 1 \<or> m = p))"
+  using prime_nat_def [transferred]
+    apply (case_tac "p >= 0")
+    by (blast, auto simp add: prime_ge_0_int)
+
+lemma prime_imp_coprime_nat: "prime (p::nat) \<Longrightarrow> \<not> p dvd n \<Longrightarrow> coprime p n"
+  apply (unfold prime_nat_def)
+  apply (metis gcd_dvd1_nat gcd_dvd2_nat)
+  done
+
+lemma prime_imp_coprime_int: "prime (p::int) \<Longrightarrow> \<not> p dvd n \<Longrightarrow> coprime p n"
+  apply (unfold prime_int_altdef)
+  apply (metis gcd_dvd1_int gcd_dvd2_int gcd_ge_0_int)
+  done
+
+lemma prime_dvd_mult_nat: "prime (p::nat) \<Longrightarrow> p dvd m * n \<Longrightarrow> p dvd m \<or> p dvd n"
+  by (blast intro: coprime_dvd_mult_nat prime_imp_coprime_nat)
+
+lemma prime_dvd_mult_int: "prime (p::int) \<Longrightarrow> p dvd m * n \<Longrightarrow> p dvd m \<or> p dvd n"
+  by (blast intro: coprime_dvd_mult_int prime_imp_coprime_int)
+
+lemma prime_dvd_mult_eq_nat [simp]: "prime (p::nat) \<Longrightarrow>
+    p dvd m * n = (p dvd m \<or> p dvd n)"
+  by (rule iffI, rule prime_dvd_mult_nat, auto)
+
+lemma prime_dvd_mult_eq_int [simp]: "prime (p::int) \<Longrightarrow>
+    p dvd m * n = (p dvd m \<or> p dvd n)"
+  by (rule iffI, rule prime_dvd_mult_int, auto)
+
+lemma not_prime_eq_prod_nat: "(n::nat) > 1 \<Longrightarrow> ~ prime n \<Longrightarrow>
+    EX m k. n = m * k & 1 < m & m < n & 1 < k & k < n"
+  unfolding prime_nat_def dvd_def apply auto
+  by(metis mult_commute linorder_neq_iff linorder_not_le mult_1 n_less_n_mult_m one_le_mult_iff less_imp_le_nat)
+
+lemma not_prime_eq_prod_int: "(n::int) > 1 \<Longrightarrow> ~ prime n \<Longrightarrow>
+    EX m k. n = m * k & 1 < m & m < n & 1 < k & k < n"
+  unfolding prime_int_altdef dvd_def
+  apply auto
+  by(metis div_mult_self1_is_id div_mult_self2_is_id int_div_less_self int_one_le_iff_zero_less zero_less_mult_pos zless_le)
+
+lemma prime_dvd_power_nat [rule_format]: "prime (p::nat) -->
+    n > 0 --> (p dvd x^n --> p dvd x)"
+  by (induct n rule: nat_induct, auto)
+
+lemma prime_dvd_power_int [rule_format]: "prime (p::int) -->
+    n > 0 --> (p dvd x^n --> p dvd x)"
+  apply (induct n rule: nat_induct, auto)
+  apply (frule prime_ge_0_int)
+  apply auto
+done
+
+subsubsection{* Make prime naively executable *}
+
+lemma zero_not_prime_nat [simp]: "~prime (0::nat)"
+  by (simp add: prime_nat_def)
+
+lemma zero_not_prime_int [simp]: "~prime (0::int)"
+  by (simp add: prime_int_def)
+
+lemma one_not_prime_nat [simp]: "~prime (1::nat)"
+  by (simp add: prime_nat_def)
+
+lemma Suc_0_not_prime_nat [simp]: "~prime (Suc 0)"
+  by (simp add: prime_nat_def One_nat_def)
+
+lemma one_not_prime_int [simp]: "~prime (1::int)"
+  by (simp add: prime_int_def)
+
+lemma prime_nat_code[code]:
+ "prime(p::nat) = (p > 1 & (ALL n : {1<..<p}. ~(n dvd p)))"
+apply(simp add: Ball_def)
+apply (metis less_not_refl prime_nat_def dvd_triv_right not_prime_eq_prod_nat)
+done
+
+lemma prime_nat_simp:
+ "prime(p::nat) = (p > 1 & (list_all (%n. ~ n dvd p) [2..<p]))"
+apply(simp only:prime_nat_code list_ball_code greaterThanLessThan_upt)
+apply(simp add:nat_number One_nat_def)
+done
+
+lemmas prime_nat_simp_number_of[simp] = prime_nat_simp[of "number_of m", standard]
+
+lemma prime_int_code[code]:
+  "prime(p::int) = (p > 1 & (ALL n : {1<..<p}. ~(n dvd p)))" (is "?L = ?R")
+proof
+  assume "?L" thus "?R"
+    by (clarsimp simp: prime_gt_1_int) (metis int_one_le_iff_zero_less prime_int_altdef zless_le)
+next
+    assume "?R" thus "?L" by (clarsimp simp:Ball_def) (metis dvdI not_prime_eq_prod_int)
+qed
+
+lemma prime_int_simp:
+  "prime(p::int) = (p > 1 & (list_all (%n. ~ n dvd p) [2..p - 1]))"
+apply(simp only:prime_int_code list_ball_code greaterThanLessThan_upto)
+apply simp
+done
+
+lemmas prime_int_simp_number_of[simp] = prime_int_simp[of "number_of m", standard]
+
+lemma two_is_prime_nat [simp]: "prime (2::nat)"
+by simp
+
+lemma two_is_prime_int [simp]: "prime (2::int)"
+by simp
+
+text{* A bit of regression testing: *}
+
+lemma "prime(97::nat)"
+by simp
+
+lemma "prime(97::int)"
+by simp
+
+lemma "prime(997::nat)"
+by eval
+
+lemma "prime(997::int)"
+by eval
+
+
+lemma prime_imp_power_coprime_nat: "prime (p::nat) \<Longrightarrow> ~ p dvd a \<Longrightarrow> coprime a (p^m)"
+  apply (rule coprime_exp_nat)
+  apply (subst gcd_commute_nat)
+  apply (erule (1) prime_imp_coprime_nat)
+done
+
+lemma prime_imp_power_coprime_int: "prime (p::int) \<Longrightarrow> ~ p dvd a \<Longrightarrow> coprime a (p^m)"
+  apply (rule coprime_exp_int)
+  apply (subst gcd_commute_int)
+  apply (erule (1) prime_imp_coprime_int)
+done
+
+lemma primes_coprime_nat: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"
+  apply (rule prime_imp_coprime_nat, assumption)
+  apply (unfold prime_nat_def, auto)
+done
+
+lemma primes_coprime_int: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"
+  apply (rule prime_imp_coprime_int, assumption)
+  apply (unfold prime_int_altdef, clarify)
+  apply (drule_tac x = q in spec)
+  apply (drule_tac x = p in spec)
+  apply auto
+done
+
+lemma primes_imp_powers_coprime_nat: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow> p ~= q \<Longrightarrow> coprime (p^m) (q^n)"
+  by (rule coprime_exp2_nat, rule primes_coprime_nat)
+
+lemma primes_imp_powers_coprime_int: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow> p ~= q \<Longrightarrow> coprime (p^m) (q^n)"
+  by (rule coprime_exp2_int, rule primes_coprime_int)
+
+lemma prime_factor_nat: "n \<noteq> (1::nat) \<Longrightarrow> \<exists> p. prime p \<and> p dvd n"
+  apply (induct n rule: nat_less_induct)
+  apply (case_tac "n = 0")
+  using two_is_prime_nat apply blast
+  apply (case_tac "prime n")
+  apply blast
+  apply (subgoal_tac "n > 1")
+  apply (frule (1) not_prime_eq_prod_nat)
+  apply (auto intro: dvd_mult dvd_mult2)
+done
+
+(* An Isar version:
+
+lemma prime_factor_b_nat:
+  fixes n :: nat
+  assumes "n \<noteq> 1"
+  shows "\<exists>p. prime p \<and> p dvd n"
+
+using `n ~= 1`
+proof (induct n rule: less_induct_nat)
+  fix n :: nat
+  assume "n ~= 1" and
+    ih: "\<forall>m<n. m \<noteq> 1 \<longrightarrow> (\<exists>p. prime p \<and> p dvd m)"
+  thus "\<exists>p. prime p \<and> p dvd n"
+  proof -
+  {
+    assume "n = 0"
+    moreover note two_is_prime_nat
+    ultimately have ?thesis
+      by (auto simp del: two_is_prime_nat)
+  }
+  moreover
+  {
+    assume "prime n"
+    hence ?thesis by auto
+  }
+  moreover
+  {
+    assume "n ~= 0" and "~ prime n"
+    with `n ~= 1` have "n > 1" by auto
+    with `~ prime n` and not_prime_eq_prod_nat obtain m k where
+      "n = m * k" and "1 < m" and "m < n" by blast
+    with ih obtain p where "prime p" and "p dvd m" by blast
+    with `n = m * k` have ?thesis by auto
+  }
+  ultimately show ?thesis by blast
+  qed
+qed
+
+*)
+
+text {* One property of coprimality is easier to prove via prime factors. *}
+
+lemma prime_divprod_pow_nat:
+  assumes p: "prime (p::nat)" and ab: "coprime a b" and pab: "p^n dvd a * b"
+  shows "p^n dvd a \<or> p^n dvd b"
+proof-
+  {assume "n = 0 \<or> a = 1 \<or> b = 1" with pab have ?thesis
+      apply (cases "n=0", simp_all)
+      apply (cases "a=1", simp_all) done}
+  moreover
+  {assume n: "n \<noteq> 0" and a: "a\<noteq>1" and b: "b\<noteq>1"
+    then obtain m where m: "n = Suc m" by (cases n, auto)
+    from n have "p dvd p^n" by (intro dvd_power, auto)
+    also note pab
+    finally have pab': "p dvd a * b".
+    from prime_dvd_mult_nat[OF p pab']
+    have "p dvd a \<or> p dvd b" .
+    moreover
+    {assume pa: "p dvd a"
+      have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)
+      from coprime_common_divisor_nat [OF ab, OF pa] p have "\<not> p dvd b" by auto
+      with p have "coprime b p"
+        by (subst gcd_commute_nat, intro prime_imp_coprime_nat)
+      hence pnb: "coprime (p^n) b"
+        by (subst gcd_commute_nat, rule coprime_exp_nat)
+      from coprime_divprod_nat[OF pnba pnb] have ?thesis by blast }
+    moreover
+    {assume pb: "p dvd b"
+      have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)
+      from coprime_common_divisor_nat [OF ab, of p] pb p have "\<not> p dvd a"
+        by auto
+      with p have "coprime a p"
+        by (subst gcd_commute_nat, intro prime_imp_coprime_nat)
+      hence pna: "coprime (p^n) a"
+        by (subst gcd_commute_nat, rule coprime_exp_nat)
+      from coprime_divprod_nat[OF pab pna] have ?thesis by blast }
+    ultimately have ?thesis by blast}
+  ultimately show ?thesis by blast
+qed
+
+subsection {* Infinitely many primes *}
+
+lemma next_prime_bound: "\<exists>(p::nat). prime p \<and> n < p \<and> p <= fact n + 1"
+proof-
+  have f1: "fact n + 1 \<noteq> 1" using fact_ge_one_nat [of n] by arith 
+  from prime_factor_nat [OF f1]
+      obtain p where "prime p" and "p dvd fact n + 1" by auto
+  hence "p \<le> fact n + 1" 
+    by (intro dvd_imp_le, auto)
+  {assume "p \<le> n"
+    from `prime p` have "p \<ge> 1" 
+      by (cases p, simp_all)
+    with `p <= n` have "p dvd fact n" 
+      by (intro dvd_fact_nat)
+    with `p dvd fact n + 1` have "p dvd fact n + 1 - fact n"
+      by (rule dvd_diff_nat)
+    hence "p dvd 1" by simp
+    hence "p <= 1" by auto
+    moreover from `prime p` have "p > 1" by auto
+    ultimately have False by auto}
+  hence "n < p" by arith
+  with `prime p` and `p <= fact n + 1` show ?thesis by auto
+qed
+
+lemma bigger_prime: "\<exists>p. prime p \<and> p > (n::nat)" 
+using next_prime_bound by auto
+
+lemma primes_infinite: "\<not> (finite {(p::nat). prime p})"
+proof
+  assume "finite {(p::nat). prime p}"
+  with Max_ge have "(EX b. (ALL x : {(p::nat). prime p}. x <= b))"
+    by auto
+  then obtain b where "ALL (x::nat). prime x \<longrightarrow> x <= b"
+    by auto
+  with bigger_prime [of b] show False by auto
+qed
+
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Number_Theory/ROOT.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,2 @@
+
+use_thy "Number_Theory";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Number_Theory/Residues.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,466 @@
+(*  Title:      HOL/Library/Residues.thy
+    ID:         
+    Author:     Jeremy Avigad
+
+    An algebraic treatment of residue rings, and resulting proofs of
+    Euler's theorem and Wilson's theorem. 
+*)
+
+header {* Residue rings *}
+
+theory Residues
+imports
+   UniqueFactorization
+   Binomial
+   MiscAlgebra
+begin
+
+
+(*
+ 
+  A locale for residue rings
+
+*)
+
+constdefs 
+  residue_ring :: "int => int ring"
+  "residue_ring m == (| 
+    carrier =       {0..m - 1}, 
+    mult =          (%x y. (x * y) mod m),
+    one =           1,
+    zero =          0,
+    add =           (%x y. (x + y) mod m) |)"
+
+locale residues =
+  fixes m :: int and R (structure)
+  assumes m_gt_one: "m > 1"
+  defines "R == residue_ring m"
+
+context residues begin
+
+lemma abelian_group: "abelian_group R"
+  apply (insert m_gt_one)
+  apply (rule abelian_groupI)
+  apply (unfold R_def residue_ring_def)
+  apply (auto simp add: mod_pos_pos_trivial mod_add_right_eq [symmetric]
+    add_ac)
+  apply (case_tac "x = 0")
+  apply force
+  apply (subgoal_tac "(x + (m - x)) mod m = 0")
+  apply (erule bexI)
+  apply auto
+done
+
+lemma comm_monoid: "comm_monoid R"
+  apply (insert m_gt_one)
+  apply (unfold R_def residue_ring_def)
+  apply (rule comm_monoidI)
+  apply auto
+  apply (subgoal_tac "x * y mod m * z mod m = z * (x * y mod m) mod m")
+  apply (erule ssubst)
+  apply (subst zmod_zmult1_eq [symmetric])+
+  apply (simp_all only: mult_ac)
+done
+
+lemma cring: "cring R"
+  apply (rule cringI)
+  apply (rule abelian_group)
+  apply (rule comm_monoid)
+  apply (unfold R_def residue_ring_def, auto)
+  apply (subst mod_add_eq [symmetric])
+  apply (subst mult_commute)
+  apply (subst zmod_zmult1_eq [symmetric])
+  apply (simp add: ring_simps)
+done
+
+end
+
+sublocale residues < cring
+  by (rule cring)
+
+
+context residues begin
+
+(* These lemmas translate back and forth between internal and 
+   external concepts *)
+
+lemma res_carrier_eq: "carrier R = {0..m - 1}"
+  by (unfold R_def residue_ring_def, auto)
+
+lemma res_add_eq: "x \<oplus> y = (x + y) mod m"
+  by (unfold R_def residue_ring_def, auto)
+
+lemma res_mult_eq: "x \<otimes> y = (x * y) mod m"
+  by (unfold R_def residue_ring_def, auto)
+
+lemma res_zero_eq: "\<zero> = 0"
+  by (unfold R_def residue_ring_def, auto)
+
+lemma res_one_eq: "\<one> = 1"
+  by (unfold R_def residue_ring_def units_of_def residue_ring_def, auto)
+
+lemma res_units_eq: "Units R = { x. 0 < x & x < m & coprime x m}"
+  apply (insert m_gt_one)
+  apply (unfold Units_def R_def residue_ring_def)
+  apply auto
+  apply (subgoal_tac "x ~= 0")
+  apply auto
+  apply (rule invertible_coprime_int)
+  apply (subgoal_tac "x ~= 0")
+  apply auto
+  apply (subst (asm) coprime_iff_invertible'_int)
+  apply (rule m_gt_one)
+  apply (auto simp add: cong_int_def mult_commute)
+done
+
+lemma res_neg_eq: "\<ominus> x = (- x) mod m"
+  apply (insert m_gt_one)
+  apply (unfold R_def a_inv_def m_inv_def residue_ring_def)
+  apply auto
+  apply (rule the_equality)
+  apply auto
+  apply (subst mod_add_right_eq [symmetric])
+  apply auto
+  apply (subst mod_add_left_eq [symmetric])
+  apply auto
+  apply (subgoal_tac "y mod m = - x mod m")
+  apply simp
+  apply (subst zmod_eq_dvd_iff)
+  apply auto
+done
+
+lemma finite [iff]: "finite(carrier R)"
+  by (subst res_carrier_eq, auto)
+
+lemma finite_Units [iff]: "finite(Units R)"
+  by (subst res_units_eq, auto)
+
+(* The function a -> a mod m maps the integers to the 
+   residue classes. The following lemmas show that this mapping 
+   respects addition and multiplication on the integers. *)
+
+lemma mod_in_carrier [iff]: "a mod m : carrier R"
+  apply (unfold res_carrier_eq)
+  apply (insert m_gt_one, auto)
+done
+
+lemma add_cong: "(x mod m) \<oplus> (y mod m) = (x + y) mod m"
+  by (unfold R_def residue_ring_def, auto, arith)
+
+lemma mult_cong: "(x mod m) \<otimes> (y mod m) = (x * y) mod m"
+  apply (unfold R_def residue_ring_def, auto)
+  apply (subst zmod_zmult1_eq [symmetric])
+  apply (subst mult_commute)
+  apply (subst zmod_zmult1_eq [symmetric])
+  apply (subst mult_commute)
+  apply auto
+done  
+
+lemma zero_cong: "\<zero> = 0"
+  apply (unfold R_def residue_ring_def, auto)
+done
+
+lemma one_cong: "\<one> = 1 mod m"
+  apply (insert m_gt_one)
+  apply (unfold R_def residue_ring_def, auto)
+done
+
+(* revise algebra library to use 1? *)
+lemma pow_cong: "(x mod m) (^) n = x^n mod m"
+  apply (insert m_gt_one)
+  apply (induct n)
+  apply (auto simp add: nat_pow_def one_cong One_nat_def)
+  apply (subst mult_commute)
+  apply (rule mult_cong)
+done
+
+lemma neg_cong: "\<ominus> (x mod m) = (- x) mod m"
+  apply (rule sym)
+  apply (rule sum_zero_eq_neg)
+  apply auto
+  apply (subst add_cong)
+  apply (subst zero_cong)
+  apply auto
+done
+
+lemma (in residues) prod_cong: 
+  "finite A \<Longrightarrow> (\<Otimes> i:A. (f i) mod m) = (PROD i:A. f i) mod m"
+  apply (induct set: finite)
+  apply (auto simp: one_cong mult_cong)
+done
+
+lemma (in residues) sum_cong:
+  "finite A \<Longrightarrow> (\<Oplus> i:A. (f i) mod m) = (SUM i: A. f i) mod m"
+  apply (induct set: finite)
+  apply (auto simp: zero_cong add_cong)
+done
+
+lemma mod_in_res_units [simp]: "1 < m \<Longrightarrow> coprime a m \<Longrightarrow> 
+    a mod m : Units R"
+  apply (subst res_units_eq, auto)
+  apply (insert pos_mod_sign [of m a])
+  apply (subgoal_tac "a mod m ~= 0")
+  apply arith
+  apply auto
+  apply (subst (asm) gcd_red_int)
+  apply (subst gcd_commute_int, assumption)
+done
+
+lemma res_eq_to_cong: "((a mod m) = (b mod m)) = [a = b] (mod (m::int))" 
+  unfolding cong_int_def by auto
+
+(* Simplifying with these will translate a ring equation in R to a 
+   congruence. *)
+
+lemmas res_to_cong_simps = add_cong mult_cong pow_cong one_cong
+    prod_cong sum_cong neg_cong res_eq_to_cong
+
+(* Other useful facts about the residue ring *)
+
+lemma one_eq_neg_one: "\<one> = \<ominus> \<one> \<Longrightarrow> m = 2"
+  apply (simp add: res_one_eq res_neg_eq)
+  apply (insert m_gt_one)
+  apply (subgoal_tac "~(m > 2)")
+  apply arith
+  apply (rule notI)
+  apply (subgoal_tac "-1 mod m = m - 1")
+  apply force
+  apply (subst mod_add_self2 [symmetric])
+  apply (subst mod_pos_pos_trivial)
+  apply auto
+done
+
+end
+
+
+(* prime residues *)
+
+locale residues_prime =
+  fixes p :: int and R (structure)
+  assumes p_prime [intro]: "prime p"
+  defines "R == residue_ring p"
+
+sublocale residues_prime < residues p
+  apply (unfold R_def residues_def)
+  using p_prime apply auto
+done
+
+context residues_prime begin
+
+lemma is_field: "field R"
+  apply (rule cring.field_intro2)
+  apply (rule cring)
+  apply (auto simp add: res_carrier_eq res_one_eq res_zero_eq
+    res_units_eq)
+  apply (rule classical)
+  apply (erule notE)
+  apply (subst gcd_commute_int)
+  apply (rule prime_imp_coprime_int)
+  apply (rule p_prime)
+  apply (rule notI)
+  apply (frule zdvd_imp_le)
+  apply auto
+done
+
+lemma res_prime_units_eq: "Units R = {1..p - 1}"
+  apply (subst res_units_eq)
+  apply auto
+  apply (subst gcd_commute_int)
+  apply (rule prime_imp_coprime_int)
+  apply (rule p_prime)
+  apply (rule zdvd_not_zless)
+  apply auto
+done
+
+end
+
+sublocale residues_prime < field
+  by (rule is_field)
+
+
+(*
+  Test cases: Euler's theorem and Wilson's theorem.
+*)
+
+
+subsection{* Euler's theorem *}
+
+(* the definition of the phi function *)
+
+constdefs
+  phi :: "int => nat"
+  "phi m == card({ x. 0 < x & x < m & gcd x m = 1})" 
+
+lemma phi_zero [simp]: "phi 0 = 0"
+  apply (subst phi_def)
+(* Auto hangs here. Once again, where is the simplification rule 
+   1 == Suc 0 coming from? *)
+  apply (auto simp add: card_eq_0_iff)
+(* Add card_eq_0_iff as a simp rule? delete card_empty_imp? *)
+done
+
+lemma phi_one [simp]: "phi 1 = 0"
+  apply (auto simp add: phi_def card_eq_0_iff)
+done
+
+lemma (in residues) phi_eq: "phi m = card(Units R)"
+  by (simp add: phi_def res_units_eq)
+
+lemma (in residues) euler_theorem1: 
+  assumes a: "gcd a m = 1"
+  shows "[a^phi m = 1] (mod m)"
+proof -
+  from a m_gt_one have [simp]: "a mod m : Units R"
+    by (intro mod_in_res_units)
+  from phi_eq have "(a mod m) (^) (phi m) = (a mod m) (^) (card (Units R))"
+    by simp
+  also have "\<dots> = \<one>" 
+    by (intro units_power_order_eq_one, auto)
+  finally show ?thesis
+    by (simp add: res_to_cong_simps)
+qed
+
+(* In fact, there is a two line proof!
+
+lemma (in residues) euler_theorem1: 
+  assumes a: "gcd a m = 1"
+  shows "[a^phi m = 1] (mod m)"
+proof -
+  have "(a mod m) (^) (phi m) = \<one>"
+    by (simp add: phi_eq units_power_order_eq_one a m_gt_one)
+  thus ?thesis
+    by (simp add: res_to_cong_simps)
+qed
+
+*)
+
+(* outside the locale, we can relax the restriction m > 1 *)
+
+lemma euler_theorem:
+  assumes "m >= 0" and "gcd a m = 1"
+  shows "[a^phi m = 1] (mod m)"
+proof (cases)
+  assume "m = 0 | m = 1"
+  thus ?thesis by auto
+next
+  assume "~(m = 0 | m = 1)"
+  with prems show ?thesis
+    by (intro residues.euler_theorem1, unfold residues_def, auto)
+qed
+
+lemma (in residues_prime) phi_prime: "phi p = (nat p - 1)"
+  apply (subst phi_eq)
+  apply (subst res_prime_units_eq)
+  apply auto
+done
+
+lemma phi_prime: "prime p \<Longrightarrow> phi p = (nat p - 1)"
+  apply (rule residues_prime.phi_prime)
+  apply (erule residues_prime.intro)
+done
+
+lemma fermat_theorem:
+  assumes "prime p" and "~ (p dvd a)"
+  shows "[a^(nat p - 1) = 1] (mod p)"
+proof -
+  from prems have "[a^phi p = 1] (mod p)"
+    apply (intro euler_theorem)
+    (* auto should get this next part. matching across
+       substitutions is needed. *)
+    apply (frule prime_gt_1_int, arith)
+    apply (subst gcd_commute_int, erule prime_imp_coprime_int, assumption)
+    done
+  also have "phi p = nat p - 1"
+    by (rule phi_prime, rule prems)
+  finally show ?thesis .
+qed
+
+
+subsection {* Wilson's theorem *}
+
+lemma (in field) inv_pair_lemma: "x : Units R \<Longrightarrow> y : Units R \<Longrightarrow> 
+  {x, inv x} ~= {y, inv y} \<Longrightarrow> {x, inv x} Int {y, inv y} = {}" 
+  apply auto
+  apply (erule notE)
+  apply (erule inv_eq_imp_eq)
+  apply auto
+  apply (erule notE)
+  apply (erule inv_eq_imp_eq)
+  apply auto
+done
+
+lemma (in residues_prime) wilson_theorem1:
+  assumes a: "p > 2"
+  shows "[fact (p - 1) = - 1] (mod p)"
+proof -
+  let ?InversePairs = "{ {x, inv x} | x. x : Units R - {\<one>, \<ominus> \<one>}}" 
+  have UR: "Units R = {\<one>, \<ominus> \<one>} Un (Union ?InversePairs)"
+    by auto
+  have "(\<Otimes>i: Units R. i) = 
+    (\<Otimes>i: {\<one>, \<ominus> \<one>}. i) \<otimes> (\<Otimes>i: Union ?InversePairs. i)"
+    apply (subst UR)
+    apply (subst finprod_Un_disjoint)
+    apply (auto intro:funcsetI)
+    apply (drule sym, subst (asm) inv_eq_one_eq)
+    apply auto
+    apply (drule sym, subst (asm) inv_eq_neg_one_eq)
+    apply auto
+    done
+  also have "(\<Otimes>i: {\<one>, \<ominus> \<one>}. i) = \<ominus> \<one>"
+    apply (subst finprod_insert)
+    apply auto
+    apply (frule one_eq_neg_one)
+    apply (insert a, force)
+    done
+  also have "(\<Otimes>i:(Union ?InversePairs). i) = 
+      (\<Otimes> A: ?InversePairs. (\<Otimes> y:A. y))"
+    apply (subst finprod_Union_disjoint)
+    apply force
+    apply force
+    apply clarify
+    apply (rule inv_pair_lemma)
+    apply auto
+    done
+  also have "\<dots> = \<one>"
+    apply (rule finprod_one)
+    apply auto
+    apply (subst finprod_insert)
+    apply auto
+    apply (frule inv_eq_self)
+    apply (auto)
+    done
+  finally have "(\<Otimes>i: Units R. i) = \<ominus> \<one>"
+    by simp
+  also have "(\<Otimes>i: Units R. i) = (\<Otimes>i: Units R. i mod p)"
+    apply (rule finprod_cong')
+    apply (auto)
+    apply (subst (asm) res_prime_units_eq)
+    apply auto
+    done
+  also have "\<dots> = (PROD i: Units R. i) mod p"
+    apply (rule prod_cong)
+    apply auto
+    done
+  also have "\<dots> = fact (p - 1) mod p"
+    apply (subst fact_altdef_int)
+    apply (insert prems, force)
+    apply (subst res_prime_units_eq, rule refl)
+    done
+  finally have "fact (p - 1) mod p = \<ominus> \<one>".
+  thus ?thesis
+    by (simp add: res_to_cong_simps)
+qed
+
+lemma wilson_theorem: "prime (p::int) \<Longrightarrow> [fact (p - 1) = - 1] (mod p)"
+  apply (frule prime_gt_1_int)
+  apply (case_tac "p = 2")
+  apply (subst fact_altdef_int, simp)
+  apply (subst cong_int_def)
+  apply simp
+  apply (rule residues_prime.wilson_theorem1)
+  apply (rule residues_prime.intro)
+  apply auto
+done
+
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Number_Theory/UniqueFactorization.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,967 @@
+(*  Title:      UniqueFactorization.thy
+    ID:         
+    Author:     Jeremy Avigad
+
+    
+    Unique factorization for the natural numbers and the integers.
+
+    Note: there were previous Isabelle formalizations of unique
+    factorization due to Thomas Marthedal Rasmussen, and, building on
+    that, by Jeremy Avigad and David Gray.  
+*)
+
+header {* UniqueFactorization *}
+
+theory UniqueFactorization
+imports Cong Multiset
+begin
+
+(* inherited from Multiset *)
+declare One_nat_def [simp del] 
+
+(* As a simp or intro rule,
+
+     prime p \<Longrightarrow> p > 0
+
+   wreaks havoc here. When the premise includes ALL x :# M. prime x, it 
+   leads to the backchaining
+
+     x > 0  
+     prime x 
+     x :# M   which is, unfortunately,
+     count M x > 0
+*)
+
+
+(* useful facts *)
+
+lemma setsum_Un2: "finite (A Un B) \<Longrightarrow> 
+    setsum f (A Un B) = setsum f (A - B) + setsum f (B - A) + 
+      setsum f (A Int B)"
+  apply (subgoal_tac "A Un B = (A - B) Un (B - A) Un (A Int B)")
+  apply (erule ssubst)
+  apply (subst setsum_Un_disjoint)
+  apply auto
+  apply (subst setsum_Un_disjoint)
+  apply auto
+done
+
+lemma setprod_Un2: "finite (A Un B) \<Longrightarrow> 
+    setprod f (A Un B) = setprod f (A - B) * setprod f (B - A) * 
+      setprod f (A Int B)"
+  apply (subgoal_tac "A Un B = (A - B) Un (B - A) Un (A Int B)")
+  apply (erule ssubst)
+  apply (subst setprod_Un_disjoint)
+  apply auto
+  apply (subst setprod_Un_disjoint)
+  apply auto
+done
+ 
+(* Should this go in Multiset.thy? *)
+(* TN: No longer an intro-rule; needed only once and might get in the way *)
+lemma multiset_eqI: "[| !!x. count M x = count N x |] ==> M = N"
+  by (subst multiset_eq_conv_count_eq, blast)
+
+(* Here is a version of set product for multisets. Is it worth moving
+   to multiset.thy? If so, one should similarly define msetsum for abelian 
+   semirings, using of_nat. Also, is it worth developing bounded quantifiers 
+   "ALL i :# M. P i"? 
+*)
+
+constdefs
+  msetprod :: "('a => ('b::{power,comm_monoid_mult})) => 'a multiset => 'b"
+  "msetprod f M == setprod (%x. (f x)^(count M x)) (set_of M)"
+
+syntax
+  "_msetprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" 
+      ("(3PROD _:#_. _)" [0, 51, 10] 10)
+
+translations
+  "PROD i :# A. b" == "msetprod (%i. b) A"
+
+lemma msetprod_Un: "msetprod f (A+B) = msetprod f A * msetprod f B" 
+  apply (simp add: msetprod_def power_add)
+  apply (subst setprod_Un2)
+  apply auto
+  apply (subgoal_tac 
+      "(PROD x:set_of A - set_of B. f x ^ count A x * f x ^ count B x) =
+       (PROD x:set_of A - set_of B. f x ^ count A x)")
+  apply (erule ssubst)
+  apply (subgoal_tac 
+      "(PROD x:set_of B - set_of A. f x ^ count A x * f x ^ count B x) =
+       (PROD x:set_of B - set_of A. f x ^ count B x)")
+  apply (erule ssubst)
+  apply (subgoal_tac "(PROD x:set_of A. f x ^ count A x) = 
+    (PROD x:set_of A - set_of B. f x ^ count A x) *
+    (PROD x:set_of A Int set_of B. f x ^ count A x)")
+  apply (erule ssubst)
+  apply (subgoal_tac "(PROD x:set_of B. f x ^ count B x) = 
+    (PROD x:set_of B - set_of A. f x ^ count B x) *
+    (PROD x:set_of A Int set_of B. f x ^ count B x)")
+  apply (erule ssubst)
+  apply (subst setprod_timesf)
+  apply (force simp add: mult_ac)
+  apply (subst setprod_Un_disjoint [symmetric])
+  apply (auto intro: setprod_cong)
+  apply (subst setprod_Un_disjoint [symmetric])
+  apply (auto intro: setprod_cong)
+done
+
+
+subsection {* unique factorization: multiset version *}
+
+lemma multiset_prime_factorization_exists [rule_format]: "n > 0 --> 
+    (EX M. (ALL (p::nat) : set_of M. prime p) & n = (PROD i :# M. i))"
+proof (rule nat_less_induct, clarify)
+  fix n :: nat
+  assume ih: "ALL m < n. 0 < m --> (EX M. (ALL p : set_of M. prime p) & m = 
+      (PROD i :# M. i))"
+  assume "(n::nat) > 0"
+  then have "n = 1 | (n > 1 & prime n) | (n > 1 & ~ prime n)"
+    by arith
+  moreover 
+  {
+    assume "n = 1"
+    then have "(ALL p : set_of {#}. prime p) & n = (PROD i :# {#}. i)"
+        by (auto simp add: msetprod_def)
+  } 
+  moreover 
+  {
+    assume "n > 1" and "prime n"
+    then have "(ALL p : set_of {# n #}. prime p) & n = (PROD i :# {# n #}. i)"
+      by (auto simp add: msetprod_def)
+  } 
+  moreover 
+  {
+    assume "n > 1" and "~ prime n"
+    from prems not_prime_eq_prod_nat
+      obtain m k where "n = m * k & 1 < m & m < n & 1 < k & k < n"
+        by blast
+    with ih obtain Q R where "(ALL p : set_of Q. prime p) & m = (PROD i:#Q. i)"
+        and "(ALL p: set_of R. prime p) & k = (PROD i:#R. i)"
+      by blast
+    hence "(ALL p: set_of (Q + R). prime p) & n = (PROD i :# Q + R. i)"
+      by (auto simp add: prems msetprod_Un set_of_union)
+    then have "EX M. (ALL p : set_of M. prime p) & n = (PROD i :# M. i)"..
+  }
+  ultimately show "EX M. (ALL p : set_of M. prime p) & n = (PROD i::nat:#M. i)"
+    by blast
+qed
+
+lemma multiset_prime_factorization_unique_aux:
+  fixes a :: nat
+  assumes "(ALL p : set_of M. prime p)" and
+    "(ALL p : set_of N. prime p)" and
+    "(PROD i :# M. i) dvd (PROD i:# N. i)"
+  shows
+    "count M a <= count N a"
+proof cases
+  assume "a : set_of M"
+  with prems have a: "prime a"
+    by auto
+  with prems have "a ^ count M a dvd (PROD i :# M. i)"
+    by (auto intro: dvd_setprod simp add: msetprod_def)
+  also have "... dvd (PROD i :# N. i)"
+    by (rule prems)
+  also have "... = (PROD i : (set_of N). i ^ (count N i))"
+    by (simp add: msetprod_def)
+  also have "... = 
+      a^(count N a) * (PROD i : (set_of N - {a}). i ^ (count N i))"
+    proof (cases)
+      assume "a : set_of N"
+      hence b: "set_of N = {a} Un (set_of N - {a})"
+        by auto
+      thus ?thesis
+        by (subst (1) b, subst setprod_Un_disjoint, auto)
+    next
+      assume "a ~: set_of N" 
+      thus ?thesis
+        by auto
+    qed
+  finally have "a ^ count M a dvd 
+      a^(count N a) * (PROD i : (set_of N - {a}). i ^ (count N i))".
+  moreover have "coprime (a ^ count M a)
+      (PROD i : (set_of N - {a}). i ^ (count N i))"
+    apply (subst gcd_commute_nat)
+    apply (rule setprod_coprime_nat)
+    apply (rule primes_imp_powers_coprime_nat)
+    apply (insert prems, auto) 
+    done
+  ultimately have "a ^ count M a dvd a^(count N a)"
+    by (elim coprime_dvd_mult_nat)
+  with a show ?thesis 
+    by (intro power_dvd_imp_le, auto)
+next
+  assume "a ~: set_of M"
+  thus ?thesis by auto
+qed
+
+lemma multiset_prime_factorization_unique:
+  assumes "(ALL (p::nat) : set_of M. prime p)" and
+    "(ALL p : set_of N. prime p)" and
+    "(PROD i :# M. i) = (PROD i:# N. i)"
+  shows
+    "M = N"
+proof -
+  {
+    fix a
+    from prems have "count M a <= count N a"
+      by (intro multiset_prime_factorization_unique_aux, auto) 
+    moreover from prems have "count N a <= count M a"
+      by (intro multiset_prime_factorization_unique_aux, auto) 
+    ultimately have "count M a = count N a"
+      by auto
+  }
+  thus ?thesis by (simp add:multiset_eq_conv_count_eq)
+qed
+
+constdefs
+  multiset_prime_factorization :: "nat => nat multiset"
+  "multiset_prime_factorization n ==
+     if n > 0 then (THE M. ((ALL p : set_of M. prime p) & 
+       n = (PROD i :# M. i)))
+     else {#}"
+
+lemma multiset_prime_factorization: "n > 0 ==>
+    (ALL p : set_of (multiset_prime_factorization n). prime p) &
+       n = (PROD i :# (multiset_prime_factorization n). i)"
+  apply (unfold multiset_prime_factorization_def)
+  apply clarsimp
+  apply (frule multiset_prime_factorization_exists)
+  apply clarify
+  apply (rule theI)
+  apply (insert multiset_prime_factorization_unique, blast)+
+done
+
+
+subsection {* Prime factors and multiplicity for nats and ints *}
+
+class unique_factorization =
+
+fixes
+  multiplicity :: "'a \<Rightarrow> 'a \<Rightarrow> nat" and
+  prime_factors :: "'a \<Rightarrow> 'a set"
+
+(* definitions for the natural numbers *)
+
+instantiation nat :: unique_factorization
+
+begin
+
+definition
+  multiplicity_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+where
+  "multiplicity_nat p n = count (multiset_prime_factorization n) p"
+
+definition
+  prime_factors_nat :: "nat \<Rightarrow> nat set"
+where
+  "prime_factors_nat n = set_of (multiset_prime_factorization n)"
+
+instance proof qed
+
+end
+
+(* definitions for the integers *)
+
+instantiation int :: unique_factorization
+
+begin
+
+definition
+  multiplicity_int :: "int \<Rightarrow> int \<Rightarrow> nat"
+where
+  "multiplicity_int p n = multiplicity (nat p) (nat n)"
+
+definition
+  prime_factors_int :: "int \<Rightarrow> int set"
+where
+  "prime_factors_int n = int ` (prime_factors (nat n))"
+
+instance proof qed
+
+end
+
+
+subsection {* Set up transfer *}
+
+lemma transfer_nat_int_prime_factors: 
+  "prime_factors (nat n) = nat ` prime_factors n"
+  unfolding prime_factors_int_def apply auto
+  by (subst transfer_int_nat_set_return_embed, assumption)
+
+lemma transfer_nat_int_prime_factors_closure: "n >= 0 \<Longrightarrow> 
+    nat_set (prime_factors n)"
+  by (auto simp add: nat_set_def prime_factors_int_def)
+
+lemma transfer_nat_int_multiplicity: "p >= 0 \<Longrightarrow> n >= 0 \<Longrightarrow>
+  multiplicity (nat p) (nat n) = multiplicity p n"
+  by (auto simp add: multiplicity_int_def)
+
+declare TransferMorphism_nat_int[transfer add return: 
+  transfer_nat_int_prime_factors transfer_nat_int_prime_factors_closure
+  transfer_nat_int_multiplicity]
+
+
+lemma transfer_int_nat_prime_factors:
+    "prime_factors (int n) = int ` prime_factors n"
+  unfolding prime_factors_int_def by auto
+
+lemma transfer_int_nat_prime_factors_closure: "is_nat n \<Longrightarrow> 
+    nat_set (prime_factors n)"
+  by (simp only: transfer_nat_int_prime_factors_closure is_nat_def)
+
+lemma transfer_int_nat_multiplicity: 
+    "multiplicity (int p) (int n) = multiplicity p n"
+  by (auto simp add: multiplicity_int_def)
+
+declare TransferMorphism_int_nat[transfer add return: 
+  transfer_int_nat_prime_factors transfer_int_nat_prime_factors_closure
+  transfer_int_nat_multiplicity]
+
+
+subsection {* Properties of prime factors and multiplicity for nats and ints *}
+
+lemma prime_factors_ge_0_int [elim]: "p : prime_factors (n::int) \<Longrightarrow> p >= 0"
+  by (unfold prime_factors_int_def, auto)
+
+lemma prime_factors_prime_nat [intro]: "p : prime_factors (n::nat) \<Longrightarrow> prime p"
+  apply (case_tac "n = 0")
+  apply (simp add: prime_factors_nat_def multiset_prime_factorization_def)
+  apply (auto simp add: prime_factors_nat_def multiset_prime_factorization)
+done
+
+lemma prime_factors_prime_int [intro]:
+  assumes "n >= 0" and "p : prime_factors (n::int)"
+  shows "prime p"
+
+  apply (rule prime_factors_prime_nat [transferred, of n p])
+  using prems apply auto
+done
+
+lemma prime_factors_gt_0_nat [elim]: "p : prime_factors x \<Longrightarrow> p > (0::nat)"
+  by (frule prime_factors_prime_nat, auto)
+
+lemma prime_factors_gt_0_int [elim]: "x >= 0 \<Longrightarrow> p : prime_factors x \<Longrightarrow> 
+    p > (0::int)"
+  by (frule (1) prime_factors_prime_int, auto)
+
+lemma prime_factors_finite_nat [iff]: "finite (prime_factors (n::nat))"
+  by (unfold prime_factors_nat_def, auto)
+
+lemma prime_factors_finite_int [iff]: "finite (prime_factors (n::int))"
+  by (unfold prime_factors_int_def, auto)
+
+lemma prime_factors_altdef_nat: "prime_factors (n::nat) = 
+    {p. multiplicity p n > 0}"
+  by (force simp add: prime_factors_nat_def multiplicity_nat_def)
+
+lemma prime_factors_altdef_int: "prime_factors (n::int) = 
+    {p. p >= 0 & multiplicity p n > 0}"
+  apply (unfold prime_factors_int_def multiplicity_int_def)
+  apply (subst prime_factors_altdef_nat)
+  apply (auto simp add: image_def)
+done
+
+lemma prime_factorization_nat: "(n::nat) > 0 \<Longrightarrow> 
+    n = (PROD p : prime_factors n. p^(multiplicity p n))"
+  by (frule multiset_prime_factorization, 
+    simp add: prime_factors_nat_def multiplicity_nat_def msetprod_def)
+
+thm prime_factorization_nat [transferred] 
+
+lemma prime_factorization_int: 
+  assumes "(n::int) > 0"
+  shows "n = (PROD p : prime_factors n. p^(multiplicity p n))"
+
+  apply (rule prime_factorization_nat [transferred, of n])
+  using prems apply auto
+done
+
+lemma neq_zero_eq_gt_zero_nat: "((x::nat) ~= 0) = (x > 0)"
+  by auto
+
+lemma prime_factorization_unique_nat: 
+    "S = { (p::nat) . f p > 0} \<Longrightarrow> finite S \<Longrightarrow> (ALL p : S. prime p) \<Longrightarrow>
+      n = (PROD p : S. p^(f p)) \<Longrightarrow>
+        S = prime_factors n & (ALL p. f p = multiplicity p n)"
+  apply (subgoal_tac "multiset_prime_factorization n = Abs_multiset
+      f")
+  apply (unfold prime_factors_nat_def multiplicity_nat_def)
+  apply (simp add: set_of_def count_def Abs_multiset_inverse multiset_def)
+  apply (unfold multiset_prime_factorization_def)
+  apply (subgoal_tac "n > 0")
+  prefer 2
+  apply force
+  apply (subst if_P, assumption)
+  apply (rule the1_equality)
+  apply (rule ex_ex1I)
+  apply (rule multiset_prime_factorization_exists, assumption)
+  apply (rule multiset_prime_factorization_unique)
+  apply force
+  apply force
+  apply force
+  unfolding set_of_def count_def msetprod_def
+  apply (subgoal_tac "f : multiset")
+  apply (auto simp only: Abs_multiset_inverse)
+  unfolding multiset_def apply force 
+done
+
+lemma prime_factors_characterization_nat: "S = {p. 0 < f (p::nat)} \<Longrightarrow> 
+    finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
+      prime_factors n = S"
+  by (rule prime_factorization_unique_nat [THEN conjunct1, symmetric],
+    assumption+)
+
+lemma prime_factors_characterization'_nat: 
+  "finite {p. 0 < f (p::nat)} \<Longrightarrow>
+    (ALL p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow>
+      prime_factors (PROD p | 0 < f p . p ^ f p) = {p. 0 < f p}"
+  apply (rule prime_factors_characterization_nat)
+  apply auto
+done
+
+(* A minor glitch:*)
+
+thm prime_factors_characterization'_nat 
+    [where f = "%x. f (int (x::nat))", 
+      transferred direction: nat "op <= (0::int)", rule_format]
+
+(*
+  Transfer isn't smart enough to know that the "0 < f p" should 
+  remain a comparison between nats. But the transfer still works. 
+*)
+
+lemma primes_characterization'_int [rule_format]: 
+    "finite {p. p >= 0 & 0 < f (p::int)} \<Longrightarrow>
+      (ALL p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow>
+        prime_factors (PROD p | p >=0 & 0 < f p . p ^ f p) = 
+          {p. p >= 0 & 0 < f p}"
+
+  apply (insert prime_factors_characterization'_nat 
+    [where f = "%x. f (int (x::nat))", 
+    transferred direction: nat "op <= (0::int)"])
+  apply auto
+done
+
+lemma prime_factors_characterization_int: "S = {p. 0 < f (p::int)} \<Longrightarrow> 
+    finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
+      prime_factors n = S"
+  apply simp
+  apply (subgoal_tac "{p. 0 < f p} = {p. 0 <= p & 0 < f p}")
+  apply (simp only:)
+  apply (subst primes_characterization'_int)
+  apply auto
+  apply (auto simp add: prime_ge_0_int)
+done
+
+lemma multiplicity_characterization_nat: "S = {p. 0 < f (p::nat)} \<Longrightarrow> 
+    finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
+      multiplicity p n = f p"
+  by (frule prime_factorization_unique_nat [THEN conjunct2, rule_format, 
+    symmetric], auto)
+
+lemma multiplicity_characterization'_nat: "finite {p. 0 < f (p::nat)} \<longrightarrow>
+    (ALL p. 0 < f p \<longrightarrow> prime p) \<longrightarrow>
+      multiplicity p (PROD p | 0 < f p . p ^ f p) = f p"
+  apply (rule impI)+
+  apply (rule multiplicity_characterization_nat)
+  apply auto
+done
+
+lemma multiplicity_characterization'_int [rule_format]: 
+  "finite {p. p >= 0 & 0 < f (p::int)} \<Longrightarrow>
+    (ALL p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow> p >= 0 \<Longrightarrow>
+      multiplicity p (PROD p | p >= 0 & 0 < f p . p ^ f p) = f p"
+
+  apply (insert multiplicity_characterization'_nat 
+    [where f = "%x. f (int (x::nat))", 
+      transferred direction: nat "op <= (0::int)", rule_format])
+  apply auto
+done
+
+lemma multiplicity_characterization_int: "S = {p. 0 < f (p::int)} \<Longrightarrow> 
+    finite S \<Longrightarrow> (ALL p:S. prime p) \<Longrightarrow> n = (PROD p:S. p ^ f p) \<Longrightarrow>
+      p >= 0 \<Longrightarrow> multiplicity p n = f p"
+  apply simp
+  apply (subgoal_tac "{p. 0 < f p} = {p. 0 <= p & 0 < f p}")
+  apply (simp only:)
+  apply (subst multiplicity_characterization'_int)
+  apply auto
+  apply (auto simp add: prime_ge_0_int)
+done
+
+lemma multiplicity_zero_nat [simp]: "multiplicity (p::nat) 0 = 0"
+  by (simp add: multiplicity_nat_def multiset_prime_factorization_def)
+
+lemma multiplicity_zero_int [simp]: "multiplicity (p::int) 0 = 0"
+  by (simp add: multiplicity_int_def) 
+
+lemma multiplicity_one_nat [simp]: "multiplicity p (1::nat) = 0"
+  by (subst multiplicity_characterization_nat [where f = "%x. 0"], auto)
+
+lemma multiplicity_one_int [simp]: "multiplicity p (1::int) = 0"
+  by (simp add: multiplicity_int_def)
+
+lemma multiplicity_prime_nat [simp]: "prime (p::nat) \<Longrightarrow> multiplicity p p = 1"
+  apply (subst multiplicity_characterization_nat
+      [where f = "(%q. if q = p then 1 else 0)"])
+  apply auto
+  apply (case_tac "x = p")
+  apply auto
+done
+
+lemma multiplicity_prime_int [simp]: "prime (p::int) \<Longrightarrow> multiplicity p p = 1"
+  unfolding prime_int_def multiplicity_int_def by auto
+
+lemma multiplicity_prime_power_nat [simp]: "prime (p::nat) \<Longrightarrow> 
+    multiplicity p (p^n) = n"
+  apply (case_tac "n = 0")
+  apply auto
+  apply (subst multiplicity_characterization_nat
+      [where f = "(%q. if q = p then n else 0)"])
+  apply auto
+  apply (case_tac "x = p")
+  apply auto
+done
+
+lemma multiplicity_prime_power_int [simp]: "prime (p::int) \<Longrightarrow> 
+    multiplicity p (p^n) = n"
+  apply (frule prime_ge_0_int)
+  apply (auto simp add: prime_int_def multiplicity_int_def nat_power_eq)
+done
+
+lemma multiplicity_nonprime_nat [simp]: "~ prime (p::nat) \<Longrightarrow> 
+    multiplicity p n = 0"
+  apply (case_tac "n = 0")
+  apply auto
+  apply (frule multiset_prime_factorization)
+  apply (auto simp add: set_of_def multiplicity_nat_def)
+done
+
+lemma multiplicity_nonprime_int [simp]: "~ prime (p::int) \<Longrightarrow> multiplicity p n = 0"
+  by (unfold multiplicity_int_def prime_int_def, auto)
+
+lemma multiplicity_not_factor_nat [simp]: 
+    "p ~: prime_factors (n::nat) \<Longrightarrow> multiplicity p n = 0"
+  by (subst (asm) prime_factors_altdef_nat, auto)
+
+lemma multiplicity_not_factor_int [simp]: 
+    "p >= 0 \<Longrightarrow> p ~: prime_factors (n::int) \<Longrightarrow> multiplicity p n = 0"
+  by (subst (asm) prime_factors_altdef_int, auto)
+
+lemma multiplicity_product_aux_nat: "(k::nat) > 0 \<Longrightarrow> l > 0 \<Longrightarrow>
+    (prime_factors k) Un (prime_factors l) = prime_factors (k * l) &
+    (ALL p. multiplicity p k + multiplicity p l = multiplicity p (k * l))"
+  apply (rule prime_factorization_unique_nat)
+  apply (simp only: prime_factors_altdef_nat)
+  apply auto
+  apply (subst power_add)
+  apply (subst setprod_timesf)
+  apply (rule arg_cong2)back back
+  apply (subgoal_tac "prime_factors k Un prime_factors l = prime_factors k Un 
+      (prime_factors l - prime_factors k)")
+  apply (erule ssubst)
+  apply (subst setprod_Un_disjoint)
+  apply auto
+  apply (subgoal_tac "(\<Prod>p\<in>prime_factors l - prime_factors k. p ^ multiplicity p k) = 
+      (\<Prod>p\<in>prime_factors l - prime_factors k. 1)")
+  apply (erule ssubst)
+  apply (simp add: setprod_1)
+  apply (erule prime_factorization_nat)
+  apply (rule setprod_cong, auto)
+  apply (subgoal_tac "prime_factors k Un prime_factors l = prime_factors l Un 
+      (prime_factors k - prime_factors l)")
+  apply (erule ssubst)
+  apply (subst setprod_Un_disjoint)
+  apply auto
+  apply (subgoal_tac "(\<Prod>p\<in>prime_factors k - prime_factors l. p ^ multiplicity p l) = 
+      (\<Prod>p\<in>prime_factors k - prime_factors l. 1)")
+  apply (erule ssubst)
+  apply (simp add: setprod_1)
+  apply (erule prime_factorization_nat)
+  apply (rule setprod_cong, auto)
+done
+
+(* transfer doesn't have the same problem here with the right 
+   choice of rules. *)
+
+lemma multiplicity_product_aux_int: 
+  assumes "(k::int) > 0" and "l > 0"
+  shows 
+    "(prime_factors k) Un (prime_factors l) = prime_factors (k * l) &
+    (ALL p >= 0. multiplicity p k + multiplicity p l = multiplicity p (k * l))"
+
+  apply (rule multiplicity_product_aux_nat [transferred, of l k])
+  using prems apply auto
+done
+
+lemma prime_factors_product_nat: "(k::nat) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> prime_factors (k * l) = 
+    prime_factors k Un prime_factors l"
+  by (rule multiplicity_product_aux_nat [THEN conjunct1, symmetric])
+
+lemma prime_factors_product_int: "(k::int) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> prime_factors (k * l) = 
+    prime_factors k Un prime_factors l"
+  by (rule multiplicity_product_aux_int [THEN conjunct1, symmetric])
+
+lemma multiplicity_product_nat: "(k::nat) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> multiplicity p (k * l) = 
+    multiplicity p k + multiplicity p l"
+  by (rule multiplicity_product_aux_nat [THEN conjunct2, rule_format, 
+      symmetric])
+
+lemma multiplicity_product_int: "(k::int) > 0 \<Longrightarrow> l > 0 \<Longrightarrow> p >= 0 \<Longrightarrow> 
+    multiplicity p (k * l) = multiplicity p k + multiplicity p l"
+  by (rule multiplicity_product_aux_int [THEN conjunct2, rule_format, 
+      symmetric])
+
+lemma multiplicity_setprod_nat: "finite S \<Longrightarrow> (ALL x : S. f x > 0) \<Longrightarrow> 
+    multiplicity (p::nat) (PROD x : S. f x) = 
+      (SUM x : S. multiplicity p (f x))"
+  apply (induct set: finite)
+  apply auto
+  apply (subst multiplicity_product_nat)
+  apply auto
+done
+
+(* Transfer is delicate here for two reasons: first, because there is
+   an implicit quantifier over functions (f), and, second, because the 
+   product over the multiplicity should not be translated to an integer 
+   product.
+
+   The way to handle the first is to use quantifier rules for functions.
+   The way to handle the second is to turn off the offending rule.
+*)
+
+lemma transfer_nat_int_sum_prod_closure3:
+  "(SUM x : A. int (f x)) >= 0"
+  "(PROD x : A. int (f x)) >= 0"
+  apply (rule setsum_nonneg, auto)
+  apply (rule setprod_nonneg, auto)
+done
+
+declare TransferMorphism_nat_int[transfer 
+  add return: transfer_nat_int_sum_prod_closure3
+  del: transfer_nat_int_sum_prod2 (1)]
+
+lemma multiplicity_setprod_int: "p >= 0 \<Longrightarrow> finite S \<Longrightarrow> 
+  (ALL x : S. f x > 0) \<Longrightarrow> 
+    multiplicity (p::int) (PROD x : S. f x) = 
+      (SUM x : S. multiplicity p (f x))"
+
+  apply (frule multiplicity_setprod_nat
+    [where f = "%x. nat(int(nat(f x)))", 
+      transferred direction: nat "op <= (0::int)"])
+  apply auto
+  apply (subst (asm) setprod_cong)
+  apply (rule refl)
+  apply (rule if_P)
+  apply auto
+  apply (rule setsum_cong)
+  apply auto
+done
+
+declare TransferMorphism_nat_int[transfer 
+  add return: transfer_nat_int_sum_prod2 (1)]
+
+lemma multiplicity_prod_prime_powers_nat:
+    "finite S \<Longrightarrow> (ALL p : S. prime (p::nat)) \<Longrightarrow>
+       multiplicity p (PROD p : S. p ^ f p) = (if p : S then f p else 0)"
+  apply (subgoal_tac "(PROD p : S. p ^ f p) = 
+      (PROD p : S. p ^ (%x. if x : S then f x else 0) p)")
+  apply (erule ssubst)
+  apply (subst multiplicity_characterization_nat)
+  prefer 5 apply (rule refl)
+  apply (rule refl)
+  apply auto
+  apply (subst setprod_mono_one_right)
+  apply assumption
+  prefer 3
+  apply (rule setprod_cong)
+  apply (rule refl)
+  apply auto
+done
+
+(* Here the issue with transfer is the implicit quantifier over S *)
+
+lemma multiplicity_prod_prime_powers_int:
+    "(p::int) >= 0 \<Longrightarrow> finite S \<Longrightarrow> (ALL p : S. prime p) \<Longrightarrow>
+       multiplicity p (PROD p : S. p ^ f p) = (if p : S then f p else 0)"
+
+  apply (subgoal_tac "int ` nat ` S = S")
+  apply (frule multiplicity_prod_prime_powers_nat [where f = "%x. f(int x)" 
+    and S = "nat ` S", transferred])
+  apply auto
+  apply (subst prime_int_def [symmetric])
+  apply auto
+  apply (subgoal_tac "xb >= 0")
+  apply force
+  apply (rule prime_ge_0_int)
+  apply force
+  apply (subst transfer_nat_int_set_return_embed)
+  apply (unfold nat_set_def, auto)
+done
+
+lemma multiplicity_distinct_prime_power_nat: "prime (p::nat) \<Longrightarrow> prime q \<Longrightarrow>
+    p ~= q \<Longrightarrow> multiplicity p (q^n) = 0"
+  apply (subgoal_tac "q^n = setprod (%x. x^n) {q}")
+  apply (erule ssubst)
+  apply (subst multiplicity_prod_prime_powers_nat)
+  apply auto
+done
+
+lemma multiplicity_distinct_prime_power_int: "prime (p::int) \<Longrightarrow> prime q \<Longrightarrow>
+    p ~= q \<Longrightarrow> multiplicity p (q^n) = 0"
+  apply (frule prime_ge_0_int [of q])
+  apply (frule multiplicity_distinct_prime_power_nat [transferred leaving: n]) 
+  prefer 4
+  apply assumption
+  apply auto
+done
+
+lemma dvd_multiplicity_nat: 
+    "(0::nat) < y \<Longrightarrow> x dvd y \<Longrightarrow> multiplicity p x <= multiplicity p y"
+  apply (case_tac "x = 0")
+  apply (auto simp add: dvd_def multiplicity_product_nat)
+done
+
+lemma dvd_multiplicity_int: 
+    "(0::int) < y \<Longrightarrow> 0 <= x \<Longrightarrow> x dvd y \<Longrightarrow> p >= 0 \<Longrightarrow> 
+      multiplicity p x <= multiplicity p y"
+  apply (case_tac "x = 0")
+  apply (auto simp add: dvd_def)
+  apply (subgoal_tac "0 < k")
+  apply (auto simp add: multiplicity_product_int)
+  apply (erule zero_less_mult_pos)
+  apply arith
+done
+
+lemma dvd_prime_factors_nat [intro]:
+    "0 < (y::nat) \<Longrightarrow> x dvd y \<Longrightarrow> prime_factors x <= prime_factors y"
+  apply (simp only: prime_factors_altdef_nat)
+  apply auto
+  apply (frule dvd_multiplicity_nat)
+  apply auto
+(* It is a shame that auto and arith don't get this. *)
+  apply (erule order_less_le_trans)back
+  apply assumption
+done
+
+lemma dvd_prime_factors_int [intro]:
+    "0 < (y::int) \<Longrightarrow> 0 <= x \<Longrightarrow> x dvd y \<Longrightarrow> prime_factors x <= prime_factors y"
+  apply (auto simp add: prime_factors_altdef_int)
+  apply (erule order_less_le_trans)
+  apply (rule dvd_multiplicity_int)
+  apply auto
+done
+
+lemma multiplicity_dvd_nat: "0 < (x::nat) \<Longrightarrow> 0 < y \<Longrightarrow> 
+    ALL p. multiplicity p x <= multiplicity p y \<Longrightarrow>
+      x dvd y"
+  apply (subst prime_factorization_nat [of x], assumption)
+  apply (subst prime_factorization_nat [of y], assumption)
+  apply (rule setprod_dvd_setprod_subset2)
+  apply force
+  apply (subst prime_factors_altdef_nat)+
+  apply auto
+(* Again, a shame that auto and arith don't get this. *)
+  apply (drule_tac x = xa in spec, auto)
+  apply (rule le_imp_power_dvd)
+  apply blast
+done
+
+lemma multiplicity_dvd_int: "0 < (x::int) \<Longrightarrow> 0 < y \<Longrightarrow> 
+    ALL p >= 0. multiplicity p x <= multiplicity p y \<Longrightarrow>
+      x dvd y"
+  apply (subst prime_factorization_int [of x], assumption)
+  apply (subst prime_factorization_int [of y], assumption)
+  apply (rule setprod_dvd_setprod_subset2)
+  apply force
+  apply (subst prime_factors_altdef_int)+
+  apply auto
+  apply (rule dvd_power_le)
+  apply auto
+  apply (drule_tac x = xa in spec)
+  apply (erule impE)
+  apply auto
+done
+
+lemma multiplicity_dvd'_nat: "(0::nat) < x \<Longrightarrow> 
+    \<forall>p. prime p \<longrightarrow> multiplicity p x \<le> multiplicity p y \<Longrightarrow> x dvd y"
+  apply (cases "y = 0")
+  apply auto
+  apply (rule multiplicity_dvd_nat, auto)
+  apply (case_tac "prime p")
+  apply auto
+done
+
+lemma multiplicity_dvd'_int: "(0::int) < x \<Longrightarrow> 0 <= y \<Longrightarrow>
+    \<forall>p. prime p \<longrightarrow> multiplicity p x \<le> multiplicity p y \<Longrightarrow> x dvd y"
+  apply (cases "y = 0")
+  apply auto
+  apply (rule multiplicity_dvd_int, auto)
+  apply (case_tac "prime p")
+  apply auto
+done
+
+lemma dvd_multiplicity_eq_nat: "0 < (x::nat) \<Longrightarrow> 0 < y \<Longrightarrow>
+    (x dvd y) = (ALL p. multiplicity p x <= multiplicity p y)"
+  by (auto intro: dvd_multiplicity_nat multiplicity_dvd_nat)
+
+lemma dvd_multiplicity_eq_int: "0 < (x::int) \<Longrightarrow> 0 < y \<Longrightarrow>
+    (x dvd y) = (ALL p >= 0. multiplicity p x <= multiplicity p y)"
+  by (auto intro: dvd_multiplicity_int multiplicity_dvd_int)
+
+lemma prime_factors_altdef2_nat: "(n::nat) > 0 \<Longrightarrow> 
+    (p : prime_factors n) = (prime p & p dvd n)"
+  apply (case_tac "prime p")
+  apply auto
+  apply (subst prime_factorization_nat [where n = n], assumption)
+  apply (rule dvd_trans) 
+  apply (rule dvd_power [where x = p and n = "multiplicity p n"])
+  apply (subst (asm) prime_factors_altdef_nat, force)
+  apply (rule dvd_setprod)
+  apply auto  
+  apply (subst prime_factors_altdef_nat)
+  apply (subst (asm) dvd_multiplicity_eq_nat)
+  apply auto
+  apply (drule spec [where x = p])
+  apply auto
+done
+
+lemma prime_factors_altdef2_int: 
+  assumes "(n::int) > 0" 
+  shows "(p : prime_factors n) = (prime p & p dvd n)"
+
+  apply (case_tac "p >= 0")
+  apply (rule prime_factors_altdef2_nat [transferred])
+  using prems apply auto
+  apply (auto simp add: prime_ge_0_int prime_factors_ge_0_int)
+done
+
+lemma multiplicity_eq_nat:
+  fixes x and y::nat 
+  assumes [arith]: "x > 0" "y > 0" and
+    mult_eq [simp]: "!!p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
+  shows "x = y"
+
+  apply (rule dvd_anti_sym)
+  apply (auto intro: multiplicity_dvd'_nat) 
+done
+
+lemma multiplicity_eq_int:
+  fixes x and y::int 
+  assumes [arith]: "x > 0" "y > 0" and
+    mult_eq [simp]: "!!p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
+  shows "x = y"
+
+  apply (rule dvd_anti_sym [transferred])
+  apply (auto intro: multiplicity_dvd'_int) 
+done
+
+
+subsection {* An application *}
+
+lemma gcd_eq_nat: 
+  assumes pos [arith]: "x > 0" "y > 0"
+  shows "gcd (x::nat) y = 
+    (PROD p: prime_factors x Un prime_factors y. 
+      p ^ (min (multiplicity p x) (multiplicity p y)))"
+proof -
+  def z == "(PROD p: prime_factors (x::nat) Un prime_factors y. 
+      p ^ (min (multiplicity p x) (multiplicity p y)))"
+  have [arith]: "z > 0"
+    unfolding z_def by (rule setprod_pos_nat, auto)
+  have aux: "!!p. prime p \<Longrightarrow> multiplicity p z = 
+      min (multiplicity p x) (multiplicity p y)"
+    unfolding z_def
+    apply (subst multiplicity_prod_prime_powers_nat)
+    apply (auto simp add: multiplicity_not_factor_nat)
+    done
+  have "z dvd x" 
+    by (intro multiplicity_dvd'_nat, auto simp add: aux)
+  moreover have "z dvd y" 
+    by (intro multiplicity_dvd'_nat, auto simp add: aux)
+  moreover have "ALL w. w dvd x & w dvd y \<longrightarrow> w dvd z"
+    apply auto
+    apply (case_tac "w = 0", auto)
+    apply (erule multiplicity_dvd'_nat)
+    apply (auto intro: dvd_multiplicity_nat simp add: aux)
+    done
+  ultimately have "z = gcd x y"
+    by (subst gcd_unique_nat [symmetric], blast)
+  thus ?thesis
+    unfolding z_def by auto
+qed
+
+lemma lcm_eq_nat: 
+  assumes pos [arith]: "x > 0" "y > 0"
+  shows "lcm (x::nat) y = 
+    (PROD p: prime_factors x Un prime_factors y. 
+      p ^ (max (multiplicity p x) (multiplicity p y)))"
+proof -
+  def z == "(PROD p: prime_factors (x::nat) Un prime_factors y. 
+      p ^ (max (multiplicity p x) (multiplicity p y)))"
+  have [arith]: "z > 0"
+    unfolding z_def by (rule setprod_pos_nat, auto)
+  have aux: "!!p. prime p \<Longrightarrow> multiplicity p z = 
+      max (multiplicity p x) (multiplicity p y)"
+    unfolding z_def
+    apply (subst multiplicity_prod_prime_powers_nat)
+    apply (auto simp add: multiplicity_not_factor_nat)
+    done
+  have "x dvd z" 
+    by (intro multiplicity_dvd'_nat, auto simp add: aux)
+  moreover have "y dvd z" 
+    by (intro multiplicity_dvd'_nat, auto simp add: aux)
+  moreover have "ALL w. x dvd w & y dvd w \<longrightarrow> z dvd w"
+    apply auto
+    apply (case_tac "w = 0", auto)
+    apply (rule multiplicity_dvd'_nat)
+    apply (auto intro: dvd_multiplicity_nat simp add: aux)
+    done
+  ultimately have "z = lcm x y"
+    by (subst lcm_unique_nat [symmetric], blast)
+  thus ?thesis
+    unfolding z_def by auto
+qed
+
+lemma multiplicity_gcd_nat: 
+  assumes [arith]: "x > 0" "y > 0"
+  shows "multiplicity (p::nat) (gcd x y) = 
+    min (multiplicity p x) (multiplicity p y)"
+
+  apply (subst gcd_eq_nat)
+  apply auto
+  apply (subst multiplicity_prod_prime_powers_nat)
+  apply auto
+done
+
+lemma multiplicity_lcm_nat: 
+  assumes [arith]: "x > 0" "y > 0"
+  shows "multiplicity (p::nat) (lcm x y) = 
+    max (multiplicity p x) (multiplicity p y)"
+
+  apply (subst lcm_eq_nat)
+  apply auto
+  apply (subst multiplicity_prod_prime_powers_nat)
+  apply auto
+done
+
+lemma gcd_lcm_distrib_nat: "gcd (x::nat) (lcm y z) = lcm (gcd x y) (gcd x z)"
+  apply (case_tac "x = 0 | y = 0 | z = 0") 
+  apply auto
+  apply (rule multiplicity_eq_nat)
+  apply (auto simp add: multiplicity_gcd_nat multiplicity_lcm_nat 
+      lcm_pos_nat)
+done
+
+lemma gcd_lcm_distrib_int: "gcd (x::int) (lcm y z) = lcm (gcd x y) (gcd x z)"
+  apply (subst (1 2 3) gcd_abs_int)
+  apply (subst lcm_abs_int)
+  apply (subst (2) abs_of_nonneg)
+  apply force
+  apply (rule gcd_lcm_distrib_nat [transferred])
+  apply auto
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/BijectionRel.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,229 @@
+(*  Author:     Thomas M. Rasmussen
+    Copyright   2000  University of Cambridge
+*)
+
+header {* Bijections between sets *}
+
+theory BijectionRel imports Main begin
+
+text {*
+  Inductive definitions of bijections between two different sets and
+  between the same set.  Theorem for relating the two definitions.
+
+  \bigskip
+*}
+
+inductive_set
+  bijR :: "('a => 'b => bool) => ('a set * 'b set) set"
+  for P :: "'a => 'b => bool"
+where
+  empty [simp]: "({}, {}) \<in> bijR P"
+| insert: "P a b ==> a \<notin> A ==> b \<notin> B ==> (A, B) \<in> bijR P
+    ==> (insert a A, insert b B) \<in> bijR P"
+
+text {*
+  Add extra condition to @{term insert}: @{term "\<forall>b \<in> B. \<not> P a b"}
+  (and similar for @{term A}).
+*}
+
+definition
+  bijP :: "('a => 'a => bool) => 'a set => bool" where
+  "bijP P F = (\<forall>a b. a \<in> F \<and> P a b --> b \<in> F)"
+
+definition
+  uniqP :: "('a => 'a => bool) => bool" where
+  "uniqP P = (\<forall>a b c d. P a b \<and> P c d --> (a = c) = (b = d))"
+
+definition
+  symP :: "('a => 'a => bool) => bool" where
+  "symP P = (\<forall>a b. P a b = P b a)"
+
+inductive_set
+  bijER :: "('a => 'a => bool) => 'a set set"
+  for P :: "'a => 'a => bool"
+where
+  empty [simp]: "{} \<in> bijER P"
+| insert1: "P a a ==> a \<notin> A ==> A \<in> bijER P ==> insert a A \<in> bijER P"
+| insert2: "P a b ==> a \<noteq> b ==> a \<notin> A ==> b \<notin> A ==> A \<in> bijER P
+    ==> insert a (insert b A) \<in> bijER P"
+
+
+text {* \medskip @{term bijR} *}
+
+lemma fin_bijRl: "(A, B) \<in> bijR P ==> finite A"
+  apply (erule bijR.induct)
+  apply auto
+  done
+
+lemma fin_bijRr: "(A, B) \<in> bijR P ==> finite B"
+  apply (erule bijR.induct)
+  apply auto
+  done
+
+lemma aux_induct:
+  assumes major: "finite F"
+    and subs: "F \<subseteq> A"
+    and cases: "P {}"
+      "!!F a. F \<subseteq> A ==> a \<in> A ==> a \<notin> F ==> P F ==> P (insert a F)"
+  shows "P F"
+  using major subs
+  apply (induct set: finite)
+   apply (blast intro: cases)+
+  done
+
+
+lemma inj_func_bijR_aux1:
+    "A \<subseteq> B ==> a \<notin> A ==> a \<in> B ==> inj_on f B ==> f a \<notin> f ` A"
+  apply (unfold inj_on_def)
+  apply auto
+  done
+
+lemma inj_func_bijR_aux2:
+  "\<forall>a. a \<in> A --> P a (f a) ==> inj_on f A ==> finite A ==> F <= A
+    ==> (F, f ` F) \<in> bijR P"
+  apply (rule_tac F = F and A = A in aux_induct)
+     apply (rule finite_subset)
+      apply auto
+  apply (rule bijR.insert)
+     apply (rule_tac [3] inj_func_bijR_aux1)
+        apply auto
+  done
+
+lemma inj_func_bijR:
+  "\<forall>a. a \<in> A --> P a (f a) ==> inj_on f A ==> finite A
+    ==> (A, f ` A) \<in> bijR P"
+  apply (rule inj_func_bijR_aux2)
+     apply auto
+  done
+
+
+text {* \medskip @{term bijER} *}
+
+lemma fin_bijER: "A \<in> bijER P ==> finite A"
+  apply (erule bijER.induct)
+    apply auto
+  done
+
+lemma aux1:
+  "a \<notin> A ==> a \<notin> B ==> F \<subseteq> insert a A ==> F \<subseteq> insert a B ==> a \<in> F
+    ==> \<exists>C. F = insert a C \<and> a \<notin> C \<and> C <= A \<and> C <= B"
+  apply (rule_tac x = "F - {a}" in exI)
+  apply auto
+  done
+
+lemma aux2: "a \<noteq> b ==> a \<notin> A ==> b \<notin> B ==> a \<in> F ==> b \<in> F
+    ==> F \<subseteq> insert a A ==> F \<subseteq> insert b B
+    ==> \<exists>C. F = insert a (insert b C) \<and> a \<notin> C \<and> b \<notin> C \<and> C \<subseteq> A \<and> C \<subseteq> B"
+  apply (rule_tac x = "F - {a, b}" in exI)
+  apply auto
+  done
+
+lemma aux_uniq: "uniqP P ==> P a b ==> P c d ==> (a = c) = (b = d)"
+  apply (unfold uniqP_def)
+  apply auto
+  done
+
+lemma aux_sym: "symP P ==> P a b = P b a"
+  apply (unfold symP_def)
+  apply auto
+  done
+
+lemma aux_in1:
+    "uniqP P ==> b \<notin> C ==> P b b ==> bijP P (insert b C) ==> bijP P C"
+  apply (unfold bijP_def)
+  apply auto
+  apply (subgoal_tac "b \<noteq> a")
+   prefer 2
+   apply clarify
+  apply (simp add: aux_uniq)
+  apply auto
+  done
+
+lemma aux_in2:
+  "symP P ==> uniqP P ==> a \<notin> C ==> b \<notin> C ==> a \<noteq> b ==> P a b
+    ==> bijP P (insert a (insert b C)) ==> bijP P C"
+  apply (unfold bijP_def)
+  apply auto
+  apply (subgoal_tac "aa \<noteq> a")
+   prefer 2
+   apply clarify
+  apply (subgoal_tac "aa \<noteq> b")
+   prefer 2
+   apply clarify
+  apply (simp add: aux_uniq)
+  apply (subgoal_tac "ba \<noteq> a")
+   apply auto
+  apply (subgoal_tac "P a aa")
+   prefer 2
+   apply (simp add: aux_sym)
+  apply (subgoal_tac "b = aa")
+   apply (rule_tac [2] iffD1)
+    apply (rule_tac [2] a = a and c = a and P = P in aux_uniq)
+      apply auto
+  done
+
+lemma aux_foo: "\<forall>a b. Q a \<and> P a b --> R b ==> P a b ==> Q a ==> R b"
+  apply auto
+  done
+
+lemma aux_bij: "bijP P F ==> symP P ==> P a b ==> (a \<in> F) = (b \<in> F)"
+  apply (unfold bijP_def)
+  apply (rule iffI)
+  apply (erule_tac [!] aux_foo)
+      apply simp_all
+  apply (rule iffD2)
+   apply (rule_tac P = P in aux_sym)
+   apply simp_all
+  done
+
+
+lemma aux_bijRER:
+  "(A, B) \<in> bijR P ==> uniqP P ==> symP P
+    ==> \<forall>F. bijP P F \<and> F \<subseteq> A \<and> F \<subseteq> B --> F \<in> bijER P"
+  apply (erule bijR.induct)
+   apply simp
+  apply (case_tac "a = b")
+   apply clarify
+   apply (case_tac "b \<in> F")
+    prefer 2
+    apply (simp add: subset_insert)
+   apply (cut_tac F = F and a = b and A = A and B = B in aux1)
+        prefer 6
+        apply clarify
+        apply (rule bijER.insert1)
+          apply simp_all
+   apply (subgoal_tac "bijP P C")
+    apply simp
+   apply (rule aux_in1)
+      apply simp_all
+  apply clarify
+  apply (case_tac "a \<in> F")
+   apply (case_tac [!] "b \<in> F")
+     apply (cut_tac F = F and a = a and b = b and A = A and B = B
+       in aux2)
+            apply (simp_all add: subset_insert)
+    apply clarify
+    apply (rule bijER.insert2)
+        apply simp_all
+    apply (subgoal_tac "bijP P C")
+     apply simp
+    apply (rule aux_in2)
+          apply simp_all
+   apply (subgoal_tac "b \<in> F")
+    apply (rule_tac [2] iffD1)
+     apply (rule_tac [2] a = a and F = F and P = P in aux_bij)
+       apply (simp_all (no_asm_simp))
+   apply (subgoal_tac [2] "a \<in> F")
+    apply (rule_tac [3] iffD2)
+     apply (rule_tac [3] b = b and F = F and P = P in aux_bij)
+       apply auto
+  done
+
+lemma bijR_bijER:
+  "(A, A) \<in> bijR P ==>
+    bijP P A ==> uniqP P ==> symP P ==> A \<in> bijER P"
+  apply (cut_tac A = A and B = A and P = P in aux_bijRER)
+     apply auto
+  done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/Chinese.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,257 @@
+(*  Author:     Thomas M. Rasmussen
+    Copyright   2000  University of Cambridge
+*)
+
+header {* The Chinese Remainder Theorem *}
+
+theory Chinese 
+imports IntPrimes
+begin
+
+text {*
+  The Chinese Remainder Theorem for an arbitrary finite number of
+  equations.  (The one-equation case is included in theory @{text
+  IntPrimes}.  Uses functions for indexing.\footnote{Maybe @{term
+  funprod} and @{term funsum} should be based on general @{term fold}
+  on indices?}
+*}
+
+
+subsection {* Definitions *}
+
+consts
+  funprod :: "(nat => int) => nat => nat => int"
+  funsum :: "(nat => int) => nat => nat => int"
+
+primrec
+  "funprod f i 0 = f i"
+  "funprod f i (Suc n) = f (Suc (i + n)) * funprod f i n"
+
+primrec
+  "funsum f i 0 = f i"
+  "funsum f i (Suc n) = f (Suc (i + n)) + funsum f i n"
+
+definition
+  m_cond :: "nat => (nat => int) => bool" where
+  "m_cond n mf =
+    ((\<forall>i. i \<le> n --> 0 < mf i) \<and>
+      (\<forall>i j. i \<le> n \<and> j \<le> n \<and> i \<noteq> j --> zgcd (mf i) (mf j) = 1))"
+
+definition
+  km_cond :: "nat => (nat => int) => (nat => int) => bool" where
+  "km_cond n kf mf = (\<forall>i. i \<le> n --> zgcd (kf i) (mf i) = 1)"
+
+definition
+  lincong_sol ::
+    "nat => (nat => int) => (nat => int) => (nat => int) => int => bool" where
+  "lincong_sol n kf bf mf x = (\<forall>i. i \<le> n --> zcong (kf i * x) (bf i) (mf i))"
+
+definition
+  mhf :: "(nat => int) => nat => nat => int" where
+  "mhf mf n i =
+    (if i = 0 then funprod mf (Suc 0) (n - Suc 0)
+     else if i = n then funprod mf 0 (n - Suc 0)
+     else funprod mf 0 (i - Suc 0) * funprod mf (Suc i) (n - Suc 0 - i))"
+
+definition
+  xilin_sol ::
+    "nat => nat => (nat => int) => (nat => int) => (nat => int) => int" where
+  "xilin_sol i n kf bf mf =
+    (if 0 < n \<and> i \<le> n \<and> m_cond n mf \<and> km_cond n kf mf then
+        (SOME x. 0 \<le> x \<and> x < mf i \<and> zcong (kf i * mhf mf n i * x) (bf i) (mf i))
+     else 0)"
+
+definition
+  x_sol :: "nat => (nat => int) => (nat => int) => (nat => int) => int" where
+  "x_sol n kf bf mf = funsum (\<lambda>i. xilin_sol i n kf bf mf * mhf mf n i) 0 n"
+
+
+text {* \medskip @{term funprod} and @{term funsum} *}
+
+lemma funprod_pos: "(\<forall>i. i \<le> n --> 0 < mf i) ==> 0 < funprod mf 0 n"
+  apply (induct n)
+   apply auto
+  apply (simp add: zero_less_mult_iff)
+  done
+
+lemma funprod_zgcd [rule_format (no_asm)]:
+  "(\<forall>i. k \<le> i \<and> i \<le> k + l --> zgcd (mf i) (mf m) = 1) -->
+    zgcd (funprod mf k l) (mf m) = 1"
+  apply (induct l)
+   apply simp_all
+  apply (rule impI)+
+  apply (subst zgcd_zmult_cancel)
+  apply auto
+  done
+
+lemma funprod_zdvd [rule_format]:
+    "k \<le> i --> i \<le> k + l --> mf i dvd funprod mf k l"
+  apply (induct l)
+   apply auto
+  apply (subgoal_tac "i = Suc (k + l)")
+   apply (simp_all (no_asm_simp))
+  done
+
+lemma funsum_mod:
+    "funsum f k l mod m = funsum (\<lambda>i. (f i) mod m) k l mod m"
+  apply (induct l)
+   apply auto
+  apply (rule trans)
+   apply (rule mod_add_eq)
+  apply simp
+  apply (rule mod_add_right_eq [symmetric])
+  done
+
+lemma funsum_zero [rule_format (no_asm)]:
+    "(\<forall>i. k \<le> i \<and> i \<le> k + l --> f i = 0) --> (funsum f k l) = 0"
+  apply (induct l)
+   apply auto
+  done
+
+lemma funsum_oneelem [rule_format (no_asm)]:
+  "k \<le> j --> j \<le> k + l -->
+    (\<forall>i. k \<le> i \<and> i \<le> k + l \<and> i \<noteq> j --> f i = 0) -->
+    funsum f k l = f j"
+  apply (induct l)
+   prefer 2
+   apply clarify
+   defer
+   apply clarify
+   apply (subgoal_tac "k = j")
+    apply (simp_all (no_asm_simp))
+  apply (case_tac "Suc (k + l) = j")
+   apply (subgoal_tac "funsum f k l = 0")
+    apply (rule_tac [2] funsum_zero)
+    apply (subgoal_tac [3] "f (Suc (k + l)) = 0")
+     apply (subgoal_tac [3] "j \<le> k + l")
+      prefer 4
+      apply arith
+     apply auto
+  done
+
+
+subsection {* Chinese: uniqueness *}
+
+lemma zcong_funprod_aux:
+  "m_cond n mf ==> km_cond n kf mf
+    ==> lincong_sol n kf bf mf x ==> lincong_sol n kf bf mf y
+    ==> [x = y] (mod mf n)"
+  apply (unfold m_cond_def km_cond_def lincong_sol_def)
+  apply (rule iffD1)
+   apply (rule_tac k = "kf n" in zcong_cancel2)
+    apply (rule_tac [3] b = "bf n" in zcong_trans)
+     prefer 4
+     apply (subst zcong_sym)
+     defer
+     apply (rule order_less_imp_le)
+     apply simp_all
+  done
+
+lemma zcong_funprod [rule_format]:
+  "m_cond n mf --> km_cond n kf mf -->
+    lincong_sol n kf bf mf x --> lincong_sol n kf bf mf y -->
+    [x = y] (mod funprod mf 0 n)"
+  apply (induct n)
+   apply (simp_all (no_asm))
+   apply (blast intro: zcong_funprod_aux)
+  apply (rule impI)+
+  apply (rule zcong_zgcd_zmult_zmod)
+    apply (blast intro: zcong_funprod_aux)
+    prefer 2
+    apply (subst zgcd_commute)
+    apply (rule funprod_zgcd)
+   apply (auto simp add: m_cond_def km_cond_def lincong_sol_def)
+  done
+
+
+subsection {* Chinese: existence *}
+
+lemma unique_xi_sol:
+  "0 < n ==> i \<le> n ==> m_cond n mf ==> km_cond n kf mf
+    ==> \<exists>!x. 0 \<le> x \<and> x < mf i \<and> [kf i * mhf mf n i * x = bf i] (mod mf i)"
+  apply (rule zcong_lineq_unique)
+   apply (tactic {* stac (thm "zgcd_zmult_cancel") 2 *})
+    apply (unfold m_cond_def km_cond_def mhf_def)
+    apply (simp_all (no_asm_simp))
+  apply safe
+    apply (tactic {* stac (thm "zgcd_zmult_cancel") 3 *})
+     apply (rule_tac [!] funprod_zgcd)
+     apply safe
+     apply simp_all
+   apply (subgoal_tac "i<n")
+    prefer 2
+    apply arith
+   apply (case_tac [2] i)
+    apply simp_all
+  done
+
+lemma x_sol_lin_aux:
+    "0 < n ==> i \<le> n ==> j \<le> n ==> j \<noteq> i ==> mf j dvd mhf mf n i"
+  apply (unfold mhf_def)
+  apply (case_tac "i = 0")
+   apply (case_tac [2] "i = n")
+    apply (simp_all (no_asm_simp))
+    apply (case_tac [3] "j < i")
+     apply (rule_tac [3] dvd_mult2)
+     apply (rule_tac [4] dvd_mult)
+     apply (rule_tac [!] funprod_zdvd)
+     apply arith
+     apply arith
+     apply arith
+     apply arith
+     apply arith
+     apply arith
+     apply arith
+     apply arith
+  done
+
+lemma x_sol_lin:
+  "0 < n ==> i \<le> n
+    ==> x_sol n kf bf mf mod mf i =
+      xilin_sol i n kf bf mf * mhf mf n i mod mf i"
+  apply (unfold x_sol_def)
+  apply (subst funsum_mod)
+  apply (subst funsum_oneelem)
+     apply auto
+  apply (subst dvd_eq_mod_eq_0 [symmetric])
+  apply (rule dvd_mult)
+  apply (rule x_sol_lin_aux)
+  apply auto
+  done
+
+
+subsection {* Chinese *}
+
+lemma chinese_remainder:
+  "0 < n ==> m_cond n mf ==> km_cond n kf mf
+    ==> \<exists>!x. 0 \<le> x \<and> x < funprod mf 0 n \<and> lincong_sol n kf bf mf x"
+  apply safe
+   apply (rule_tac [2] m = "funprod mf 0 n" in zcong_zless_imp_eq)
+       apply (rule_tac [6] zcong_funprod)
+          apply auto
+  apply (rule_tac x = "x_sol n kf bf mf mod funprod mf 0 n" in exI)
+  apply (unfold lincong_sol_def)
+  apply safe
+    apply (tactic {* stac (thm "zcong_zmod") 3 *})
+    apply (tactic {* stac (thm "mod_mult_eq") 3 *})
+    apply (tactic {* stac (thm "mod_mod_cancel") 3 *})
+      apply (tactic {* stac (thm "x_sol_lin") 4 *})
+        apply (tactic {* stac (thm "mod_mult_eq" RS sym) 6 *})
+        apply (tactic {* stac (thm "zcong_zmod" RS sym) 6 *})
+        apply (subgoal_tac [6]
+          "0 \<le> xilin_sol i n kf bf mf \<and> xilin_sol i n kf bf mf < mf i
+          \<and> [kf i * mhf mf n i * xilin_sol i n kf bf mf = bf i] (mod mf i)")
+         prefer 6
+         apply (simp add: zmult_ac)
+        apply (unfold xilin_sol_def)
+        apply (tactic {* asm_simp_tac @{simpset} 6 *})
+        apply (rule_tac [6] ex1_implies_ex [THEN someI_ex])
+        apply (rule_tac [6] unique_xi_sol)
+           apply (rule_tac [3] funprod_zdvd)
+            apply (unfold m_cond_def)
+            apply (rule funprod_pos [THEN pos_mod_sign])
+            apply (rule_tac [2] funprod_pos [THEN pos_mod_bound])
+            apply auto
+  done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/Euler.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,304 @@
+(*  Title:      HOL/Quadratic_Reciprocity/Euler.thy
+    ID:         $Id$
+    Authors:    Jeremy Avigad, David Gray, and Adam Kramer
+*)
+
+header {* Euler's criterion *}
+
+theory Euler imports Residues EvenOdd begin
+
+definition
+  MultInvPair :: "int => int => int => int set" where
+  "MultInvPair a p j = {StandardRes p j, StandardRes p (a * (MultInv p j))}"
+
+definition
+  SetS        :: "int => int => int set set" where
+  "SetS        a p   =  (MultInvPair a p ` SRStar p)"
+
+
+subsection {* Property for MultInvPair *}
+
+lemma MultInvPair_prop1a:
+  "[| zprime p; 2 < p; ~([a = 0](mod p));
+      X \<in> (SetS a p); Y \<in> (SetS a p);
+      ~((X \<inter> Y) = {}) |] ==> X = Y"
+  apply (auto simp add: SetS_def)
+  apply (drule StandardRes_SRStar_prop1a)+ defer 1
+  apply (drule StandardRes_SRStar_prop1a)+
+  apply (auto simp add: MultInvPair_def StandardRes_prop2 zcong_sym)
+  apply (drule notE, rule MultInv_zcong_prop1, auto)[]
+  apply (drule notE, rule MultInv_zcong_prop2, auto simp add: zcong_sym)[]
+  apply (drule MultInv_zcong_prop2, auto simp add: zcong_sym)[]
+  apply (drule MultInv_zcong_prop3, auto simp add: zcong_sym)[]
+  apply (drule MultInv_zcong_prop1, auto)[]
+  apply (drule MultInv_zcong_prop2, auto simp add: zcong_sym)[]
+  apply (drule MultInv_zcong_prop2, auto simp add: zcong_sym)[]
+  apply (drule MultInv_zcong_prop3, auto simp add: zcong_sym)[]
+  done
+
+lemma MultInvPair_prop1b:
+  "[| zprime p; 2 < p; ~([a = 0](mod p));
+      X \<in> (SetS a p); Y \<in> (SetS a p);
+      X \<noteq> Y |] ==> X \<inter> Y = {}"
+  apply (rule notnotD)
+  apply (rule notI)
+  apply (drule MultInvPair_prop1a, auto)
+  done
+
+lemma MultInvPair_prop1c: "[| zprime p; 2 < p; ~([a = 0](mod p)) |] ==>  
+    \<forall>X \<in> SetS a p. \<forall>Y \<in> SetS a p. X \<noteq> Y --> X\<inter>Y = {}"
+  by (auto simp add: MultInvPair_prop1b)
+
+lemma MultInvPair_prop2: "[| zprime p; 2 < p; ~([a = 0](mod p)) |] ==> 
+                          Union ( SetS a p) = SRStar p"
+  apply (auto simp add: SetS_def MultInvPair_def StandardRes_SRStar_prop4 
+    SRStar_mult_prop2)
+  apply (frule StandardRes_SRStar_prop3)
+  apply (rule bexI, auto)
+  done
+
+lemma MultInvPair_distinct: "[| zprime p; 2 < p; ~([a = 0] (mod p)); 
+                                ~([j = 0] (mod p)); 
+                                ~(QuadRes p a) |]  ==> 
+                             ~([j = a * MultInv p j] (mod p))"
+proof
+  assume "zprime p" and "2 < p" and "~([a = 0] (mod p))" and 
+    "~([j = 0] (mod p))" and "~(QuadRes p a)"
+  assume "[j = a * MultInv p j] (mod p)"
+  then have "[j * j = (a * MultInv p j) * j] (mod p)"
+    by (auto simp add: zcong_scalar)
+  then have a:"[j * j = a * (MultInv p j * j)] (mod p)"
+    by (auto simp add: zmult_ac)
+  have "[j * j = a] (mod p)"
+    proof -
+      from prems have b: "[MultInv p j * j = 1] (mod p)"
+        by (simp add: MultInv_prop2a)
+      from b a show ?thesis
+        by (auto simp add: zcong_zmult_prop2)
+    qed
+  then have "[j^2 = a] (mod p)"
+    by (metis  number_of_is_id power2_eq_square succ_bin_simps)
+  with prems show False
+    by (simp add: QuadRes_def)
+qed
+
+lemma MultInvPair_card_two: "[| zprime p; 2 < p; ~([a = 0] (mod p)); 
+                                ~(QuadRes p a); ~([j = 0] (mod p)) |]  ==> 
+                             card (MultInvPair a p j) = 2"
+  apply (auto simp add: MultInvPair_def)
+  apply (subgoal_tac "~ (StandardRes p j = StandardRes p (a * MultInv p j))")
+  apply auto
+  apply (metis MultInvPair_distinct Pls_def StandardRes_def aux number_of_is_id one_is_num_one)
+  done
+
+
+subsection {* Properties of SetS *}
+
+lemma SetS_finite: "2 < p ==> finite (SetS a p)"
+  by (auto simp add: SetS_def SRStar_finite [of p] finite_imageI)
+
+lemma SetS_elems_finite: "\<forall>X \<in> SetS a p. finite X"
+  by (auto simp add: SetS_def MultInvPair_def)
+
+lemma SetS_elems_card: "[| zprime p; 2 < p; ~([a = 0] (mod p)); 
+                        ~(QuadRes p a) |]  ==>
+                        \<forall>X \<in> SetS a p. card X = 2"
+  apply (auto simp add: SetS_def)
+  apply (frule StandardRes_SRStar_prop1a)
+  apply (rule MultInvPair_card_two, auto)
+  done
+
+lemma Union_SetS_finite: "2 < p ==> finite (Union (SetS a p))"
+  by (auto simp add: SetS_finite SetS_elems_finite finite_Union)
+
+lemma card_setsum_aux: "[| finite S; \<forall>X \<in> S. finite (X::int set); 
+    \<forall>X \<in> S. card X = n |] ==> setsum card S = setsum (%x. n) S"
+  by (induct set: finite) auto
+
+lemma SetS_card: "[| zprime p; 2 < p; ~([a = 0] (mod p)); ~(QuadRes p a) |] ==> 
+                  int(card(SetS a p)) = (p - 1) div 2"
+proof -
+  assume "zprime p" and "2 < p" and  "~([a = 0] (mod p))" and "~(QuadRes p a)"
+  then have "(p - 1) = 2 * int(card(SetS a p))"
+  proof -
+    have "p - 1 = int(card(Union (SetS a p)))"
+      by (auto simp add: prems MultInvPair_prop2 SRStar_card)
+    also have "... = int (setsum card (SetS a p))"
+      by (auto simp add: prems SetS_finite SetS_elems_finite
+                         MultInvPair_prop1c [of p a] card_Union_disjoint)
+    also have "... = int(setsum (%x.2) (SetS a p))"
+      using prems
+      by (auto simp add: SetS_elems_card SetS_finite SetS_elems_finite 
+        card_setsum_aux simp del: setsum_constant)
+    also have "... = 2 * int(card( SetS a p))"
+      by (auto simp add: prems SetS_finite setsum_const2)
+    finally show ?thesis .
+  qed
+  from this show ?thesis
+    by auto
+qed
+
+lemma SetS_setprod_prop: "[| zprime p; 2 < p; ~([a = 0] (mod p));
+                              ~(QuadRes p a); x \<in> (SetS a p) |] ==> 
+                          [\<Prod>x = a] (mod p)"
+  apply (auto simp add: SetS_def MultInvPair_def)
+  apply (frule StandardRes_SRStar_prop1a)
+  apply (subgoal_tac "StandardRes p x \<noteq> StandardRes p (a * MultInv p x)")
+  apply (auto simp add: StandardRes_prop2 MultInvPair_distinct)
+  apply (frule_tac m = p and x = x and y = "(a * MultInv p x)" in 
+    StandardRes_prop4)
+  apply (subgoal_tac "[x * (a * MultInv p x) = a * (x * MultInv p x)] (mod p)")
+  apply (drule_tac a = "StandardRes p x * StandardRes p (a * MultInv p x)" and
+                   b = "x * (a * MultInv p x)" and
+                   c = "a * (x * MultInv p x)" in  zcong_trans, force)
+  apply (frule_tac p = p and x = x in MultInv_prop2, auto)
+apply (metis StandardRes_SRStar_prop3 mult_1_right mult_commute zcong_sym zcong_zmult_prop1)
+  apply (auto simp add: zmult_ac)
+  done
+
+lemma aux1: "[| 0 < x; (x::int) < a; x \<noteq> (a - 1) |] ==> x < a - 1"
+  by arith
+
+lemma aux2: "[| (a::int) < c; b < c |] ==> (a \<le> b | b \<le> a)"
+  by auto
+
+lemma SRStar_d22set_prop: "2 < p \<Longrightarrow> (SRStar p) = {1} \<union> (d22set (p - 1))"
+  apply (induct p rule: d22set.induct)
+  apply auto
+  apply (simp add: SRStar_def d22set.simps)
+  apply (simp add: SRStar_def d22set.simps, clarify)
+  apply (frule aux1)
+  apply (frule aux2, auto)
+  apply (simp_all add: SRStar_def)
+  apply (simp add: d22set.simps)
+  apply (frule d22set_le)
+  apply (frule d22set_g_1, auto)
+  done
+
+lemma Union_SetS_setprod_prop1: "[| zprime p; 2 < p; ~([a = 0] (mod p)); ~(QuadRes p a) |] ==>
+                                 [\<Prod>(Union (SetS a p)) = a ^ nat ((p - 1) div 2)] (mod p)"
+proof -
+  assume "zprime p" and "2 < p" and  "~([a = 0] (mod p))" and "~(QuadRes p a)"
+  then have "[\<Prod>(Union (SetS a p)) = 
+      setprod (setprod (%x. x)) (SetS a p)] (mod p)"
+    by (auto simp add: SetS_finite SetS_elems_finite
+                       MultInvPair_prop1c setprod_Union_disjoint)
+  also have "[setprod (setprod (%x. x)) (SetS a p) = 
+      setprod (%x. a) (SetS a p)] (mod p)"
+    by (rule setprod_same_function_zcong)
+      (auto simp add: prems SetS_setprod_prop SetS_finite)
+  also (zcong_trans) have "[setprod (%x. a) (SetS a p) = 
+      a^(card (SetS a p))] (mod p)"
+    by (auto simp add: prems SetS_finite setprod_constant)
+  finally (zcong_trans) show ?thesis
+    apply (rule zcong_trans)
+    apply (subgoal_tac "card(SetS a p) = nat((p - 1) div 2)", auto)
+    apply (subgoal_tac "nat(int(card(SetS a p))) = nat((p - 1) div 2)", force)
+    apply (auto simp add: prems SetS_card)
+    done
+qed
+
+lemma Union_SetS_setprod_prop2: "[| zprime p; 2 < p; ~([a = 0](mod p)) |] ==> 
+                                    \<Prod>(Union (SetS a p)) = zfact (p - 1)"
+proof -
+  assume "zprime p" and "2 < p" and "~([a = 0](mod p))"
+  then have "\<Prod>(Union (SetS a p)) = \<Prod>(SRStar p)"
+    by (auto simp add: MultInvPair_prop2)
+  also have "... = \<Prod>({1} \<union> (d22set (p - 1)))"
+    by (auto simp add: prems SRStar_d22set_prop)
+  also have "... = zfact(p - 1)"
+  proof -
+    have "~(1 \<in> d22set (p - 1)) & finite( d22set (p - 1))"
+      by (metis d22set_fin d22set_g_1 linorder_neq_iff)
+    then have "\<Prod>({1} \<union> (d22set (p - 1))) = \<Prod>(d22set (p - 1))"
+      by auto
+    then show ?thesis
+      by (auto simp add: d22set_prod_zfact)
+  qed
+  finally show ?thesis .
+qed
+
+lemma zfact_prop: "[| zprime p; 2 < p; ~([a = 0] (mod p)); ~(QuadRes p a) |] ==>
+                   [zfact (p - 1) = a ^ nat ((p - 1) div 2)] (mod p)"
+  apply (frule Union_SetS_setprod_prop1) 
+  apply (auto simp add: Union_SetS_setprod_prop2)
+  done
+
+text {* \medskip Prove the first part of Euler's Criterion: *}
+
+lemma Euler_part1: "[| 2 < p; zprime p; ~([x = 0](mod p)); 
+    ~(QuadRes p x) |] ==> 
+      [x^(nat (((p) - 1) div 2)) = -1](mod p)"
+  by (metis Wilson_Russ number_of_is_id zcong_sym zcong_trans zfact_prop)
+
+text {* \medskip Prove another part of Euler Criterion: *}
+
+lemma aux_1: "0 < p ==> (a::int) ^ nat (p) = a * a ^ (nat (p) - 1)"
+proof -
+  assume "0 < p"
+  then have "a ^ (nat p) =  a ^ (1 + (nat p - 1))"
+    by (auto simp add: diff_add_assoc)
+  also have "... = (a ^ 1) * a ^ (nat(p) - 1)"
+    by (simp only: zpower_zadd_distrib)
+  also have "... = a * a ^ (nat(p) - 1)"
+    by auto
+  finally show ?thesis .
+qed
+
+lemma aux_2: "[| (2::int) < p; p \<in> zOdd |] ==> 0 < ((p - 1) div 2)"
+proof -
+  assume "2 < p" and "p \<in> zOdd"
+  then have "(p - 1):zEven"
+    by (auto simp add: zEven_def zOdd_def)
+  then have aux_1: "2 * ((p - 1) div 2) = (p - 1)"
+    by (auto simp add: even_div_2_prop2)
+  with `2 < p` have "1 < (p - 1)"
+    by auto
+  then have " 1 < (2 * ((p - 1) div 2))"
+    by (auto simp add: aux_1)
+  then have "0 < (2 * ((p - 1) div 2)) div 2"
+    by auto
+  then show ?thesis by auto
+qed
+
+lemma Euler_part2:
+    "[| 2 < p; zprime p; [a = 0] (mod p) |] ==> [0 = a ^ nat ((p - 1) div 2)] (mod p)"
+  apply (frule zprime_zOdd_eq_grt_2)
+  apply (frule aux_2, auto)
+  apply (frule_tac a = a in aux_1, auto)
+  apply (frule zcong_zmult_prop1, auto)
+  done
+
+text {* \medskip Prove the final part of Euler's Criterion: *}
+
+lemma aux__1: "[| ~([x = 0] (mod p)); [y ^ 2 = x] (mod p)|] ==> ~(p dvd y)"
+  by (metis dvdI power2_eq_square zcong_sym zcong_trans zcong_zero_equiv_div dvd_trans)
+
+lemma aux__2: "2 * nat((p - 1) div 2) =  nat (2 * ((p - 1) div 2))"
+  by (auto simp add: nat_mult_distrib)
+
+lemma Euler_part3: "[| 2 < p; zprime p; ~([x = 0](mod p)); QuadRes p x |] ==> 
+                      [x^(nat (((p) - 1) div 2)) = 1](mod p)"
+  apply (subgoal_tac "p \<in> zOdd")
+  apply (auto simp add: QuadRes_def)
+   prefer 2 
+   apply (metis number_of_is_id numeral_1_eq_1 zprime_zOdd_eq_grt_2)
+  apply (frule aux__1, auto)
+  apply (drule_tac z = "nat ((p - 1) div 2)" in zcong_zpower)
+  apply (auto simp add: zpower_zpower) 
+  apply (rule zcong_trans)
+  apply (auto simp add: zcong_sym [of "x ^ nat ((p - 1) div 2)"])
+  apply (metis Little_Fermat even_div_2_prop2 mult_Bit0 number_of_is_id odd_minus_one_even one_is_num_one zmult_1 aux__2)
+  done
+
+
+text {* \medskip Finally show Euler's Criterion: *}
+
+theorem Euler_Criterion: "[| 2 < p; zprime p |] ==> [(Legendre a p) =
+    a^(nat (((p) - 1) div 2))] (mod p)"
+  apply (auto simp add: Legendre_def Euler_part2)
+  apply (frule Euler_part3, auto simp add: zcong_sym)[]
+  apply (frule Euler_part1, auto simp add: zcong_sym)[]
+  done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/EulerFermat.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,346 @@
+(*  Author:     Thomas M. Rasmussen
+    Copyright   2000  University of Cambridge
+*)
+
+header {* Fermat's Little Theorem extended to Euler's Totient function *}
+
+theory EulerFermat
+imports BijectionRel IntFact
+begin
+
+text {*
+  Fermat's Little Theorem extended to Euler's Totient function. More
+  abstract approach than Boyer-Moore (which seems necessary to achieve
+  the extended version).
+*}
+
+
+subsection {* Definitions and lemmas *}
+
+inductive_set
+  RsetR :: "int => int set set"
+  for m :: int
+  where
+    empty [simp]: "{} \<in> RsetR m"
+  | insert: "A \<in> RsetR m ==> zgcd a m = 1 ==>
+      \<forall>a'. a' \<in> A --> \<not> zcong a a' m ==> insert a A \<in> RsetR m"
+
+consts
+  BnorRset :: "int * int => int set"
+
+recdef BnorRset
+  "measure ((\<lambda>(a, m). nat a) :: int * int => nat)"
+  "BnorRset (a, m) =
+   (if 0 < a then
+    let na = BnorRset (a - 1, m)
+    in (if zgcd a m = 1 then insert a na else na)
+    else {})"
+
+definition
+  norRRset :: "int => int set" where
+  "norRRset m = BnorRset (m - 1, m)"
+
+definition
+  noXRRset :: "int => int => int set" where
+  "noXRRset m x = (\<lambda>a. a * x) ` norRRset m"
+
+definition
+  phi :: "int => nat" where
+  "phi m = card (norRRset m)"
+
+definition
+  is_RRset :: "int set => int => bool" where
+  "is_RRset A m = (A \<in> RsetR m \<and> card A = phi m)"
+
+definition
+  RRset2norRR :: "int set => int => int => int" where
+  "RRset2norRR A m a =
+     (if 1 < m \<and> is_RRset A m \<and> a \<in> A then
+        SOME b. zcong a b m \<and> b \<in> norRRset m
+      else 0)"
+
+definition
+  zcongm :: "int => int => int => bool" where
+  "zcongm m = (\<lambda>a b. zcong a b m)"
+
+lemma abs_eq_1_iff [iff]: "(abs z = (1::int)) = (z = 1 \<or> z = -1)"
+  -- {* LCP: not sure why this lemma is needed now *}
+  by (auto simp add: abs_if)
+
+
+text {* \medskip @{text norRRset} *}
+
+declare BnorRset.simps [simp del]
+
+lemma BnorRset_induct:
+  assumes "!!a m. P {} a m"
+    and "!!a m. 0 < (a::int) ==> P (BnorRset (a - 1, m::int)) (a - 1) m
+      ==> P (BnorRset(a,m)) a m"
+  shows "P (BnorRset(u,v)) u v"
+  apply (rule BnorRset.induct)
+  apply safe
+   apply (case_tac [2] "0 < a")
+    apply (rule_tac [2] prems)
+     apply simp_all
+   apply (simp_all add: BnorRset.simps prems)
+  done
+
+lemma Bnor_mem_zle [rule_format]: "b \<in> BnorRset (a, m) \<longrightarrow> b \<le> a"
+  apply (induct a m rule: BnorRset_induct)
+   apply simp
+  apply (subst BnorRset.simps)
+   apply (unfold Let_def, auto)
+  done
+
+lemma Bnor_mem_zle_swap: "a < b ==> b \<notin> BnorRset (a, m)"
+  by (auto dest: Bnor_mem_zle)
+
+lemma Bnor_mem_zg [rule_format]: "b \<in> BnorRset (a, m) --> 0 < b"
+  apply (induct a m rule: BnorRset_induct)
+   prefer 2
+   apply (subst BnorRset.simps)
+   apply (unfold Let_def, auto)
+  done
+
+lemma Bnor_mem_if [rule_format]:
+    "zgcd b m = 1 --> 0 < b --> b \<le> a --> b \<in> BnorRset (a, m)"
+  apply (induct a m rule: BnorRset.induct, auto)
+   apply (subst BnorRset.simps)
+   defer
+   apply (subst BnorRset.simps)
+   apply (unfold Let_def, auto)
+  done
+
+lemma Bnor_in_RsetR [rule_format]: "a < m --> BnorRset (a, m) \<in> RsetR m"
+  apply (induct a m rule: BnorRset_induct, simp)
+  apply (subst BnorRset.simps)
+  apply (unfold Let_def, auto)
+  apply (rule RsetR.insert)
+    apply (rule_tac [3] allI)
+    apply (rule_tac [3] impI)
+    apply (rule_tac [3] zcong_not)
+       apply (subgoal_tac [6] "a' \<le> a - 1")
+        apply (rule_tac [7] Bnor_mem_zle)
+        apply (rule_tac [5] Bnor_mem_zg, auto)
+  done
+
+lemma Bnor_fin: "finite (BnorRset (a, m))"
+  apply (induct a m rule: BnorRset_induct)
+   prefer 2
+   apply (subst BnorRset.simps)
+   apply (unfold Let_def, auto)
+  done
+
+lemma norR_mem_unique_aux: "a \<le> b - 1 ==> a < (b::int)"
+  apply auto
+  done
+
+lemma norR_mem_unique:
+  "1 < m ==>
+    zgcd a m = 1 ==> \<exists>!b. [a = b] (mod m) \<and> b \<in> norRRset m"
+  apply (unfold norRRset_def)
+  apply (cut_tac a = a and m = m in zcong_zless_unique, auto)
+   apply (rule_tac [2] m = m in zcong_zless_imp_eq)
+       apply (auto intro: Bnor_mem_zle Bnor_mem_zg zcong_trans
+	 order_less_imp_le norR_mem_unique_aux simp add: zcong_sym)
+  apply (rule_tac x = b in exI, safe)
+  apply (rule Bnor_mem_if)
+    apply (case_tac [2] "b = 0")
+     apply (auto intro: order_less_le [THEN iffD2])
+   prefer 2
+   apply (simp only: zcong_def)
+   apply (subgoal_tac "zgcd a m = m")
+    prefer 2
+    apply (subst zdvd_iff_zgcd [symmetric])
+     apply (rule_tac [4] zgcd_zcong_zgcd)
+       apply (simp_all add: zcong_sym)
+  done
+
+
+text {* \medskip @{term noXRRset} *}
+
+lemma RRset_gcd [rule_format]:
+    "is_RRset A m ==> a \<in> A --> zgcd a m = 1"
+  apply (unfold is_RRset_def)
+  apply (rule RsetR.induct [where P="%A. a \<in> A --> zgcd a m = 1"], auto)
+  done
+
+lemma RsetR_zmult_mono:
+  "A \<in> RsetR m ==>
+    0 < m ==> zgcd x m = 1 ==> (\<lambda>a. a * x) ` A \<in> RsetR m"
+  apply (erule RsetR.induct, simp_all)
+  apply (rule RsetR.insert, auto)
+   apply (blast intro: zgcd_zgcd_zmult)
+  apply (simp add: zcong_cancel)
+  done
+
+lemma card_nor_eq_noX:
+  "0 < m ==>
+    zgcd x m = 1 ==> card (noXRRset m x) = card (norRRset m)"
+  apply (unfold norRRset_def noXRRset_def)
+  apply (rule card_image)
+   apply (auto simp add: inj_on_def Bnor_fin)
+  apply (simp add: BnorRset.simps)
+  done
+
+lemma noX_is_RRset:
+    "0 < m ==> zgcd x m = 1 ==> is_RRset (noXRRset m x) m"
+  apply (unfold is_RRset_def phi_def)
+  apply (auto simp add: card_nor_eq_noX)
+  apply (unfold noXRRset_def norRRset_def)
+  apply (rule RsetR_zmult_mono)
+    apply (rule Bnor_in_RsetR, simp_all)
+  done
+
+lemma aux_some:
+  "1 < m ==> is_RRset A m ==> a \<in> A
+    ==> zcong a (SOME b. [a = b] (mod m) \<and> b \<in> norRRset m) m \<and>
+      (SOME b. [a = b] (mod m) \<and> b \<in> norRRset m) \<in> norRRset m"
+  apply (rule norR_mem_unique [THEN ex1_implies_ex, THEN someI_ex])
+   apply (rule_tac [2] RRset_gcd, simp_all)
+  done
+
+lemma RRset2norRR_correct:
+  "1 < m ==> is_RRset A m ==> a \<in> A ==>
+    [a = RRset2norRR A m a] (mod m) \<and> RRset2norRR A m a \<in> norRRset m"
+  apply (unfold RRset2norRR_def, simp)
+  apply (rule aux_some, simp_all)
+  done
+
+lemmas RRset2norRR_correct1 =
+  RRset2norRR_correct [THEN conjunct1, standard]
+lemmas RRset2norRR_correct2 =
+  RRset2norRR_correct [THEN conjunct2, standard]
+
+lemma RsetR_fin: "A \<in> RsetR m ==> finite A"
+  by (induct set: RsetR) auto
+
+lemma RRset_zcong_eq [rule_format]:
+  "1 < m ==>
+    is_RRset A m ==> [a = b] (mod m) ==> a \<in> A --> b \<in> A --> a = b"
+  apply (unfold is_RRset_def)
+  apply (rule RsetR.induct [where P="%A. a \<in> A --> b \<in> A --> a = b"])
+    apply (auto simp add: zcong_sym)
+  done
+
+lemma aux:
+  "P (SOME a. P a) ==> Q (SOME a. Q a) ==>
+    (SOME a. P a) = (SOME a. Q a) ==> \<exists>a. P a \<and> Q a"
+  apply auto
+  done
+
+lemma RRset2norRR_inj:
+    "1 < m ==> is_RRset A m ==> inj_on (RRset2norRR A m) A"
+  apply (unfold RRset2norRR_def inj_on_def, auto)
+  apply (subgoal_tac "\<exists>b. ([x = b] (mod m) \<and> b \<in> norRRset m) \<and>
+      ([y = b] (mod m) \<and> b \<in> norRRset m)")
+   apply (rule_tac [2] aux)
+     apply (rule_tac [3] aux_some)
+       apply (rule_tac [2] aux_some)
+         apply (rule RRset_zcong_eq, auto)
+  apply (rule_tac b = b in zcong_trans)
+   apply (simp_all add: zcong_sym)
+  done
+
+lemma RRset2norRR_eq_norR:
+    "1 < m ==> is_RRset A m ==> RRset2norRR A m ` A = norRRset m"
+  apply (rule card_seteq)
+    prefer 3
+    apply (subst card_image)
+      apply (rule_tac RRset2norRR_inj, auto)
+     apply (rule_tac [3] RRset2norRR_correct2, auto)
+    apply (unfold is_RRset_def phi_def norRRset_def)
+    apply (auto simp add: Bnor_fin)
+  done
+
+
+lemma Bnor_prod_power_aux: "a \<notin> A ==> inj f ==> f a \<notin> f ` A"
+by (unfold inj_on_def, auto)
+
+lemma Bnor_prod_power [rule_format]:
+  "x \<noteq> 0 ==> a < m --> \<Prod>((\<lambda>a. a * x) ` BnorRset (a, m)) =
+      \<Prod>(BnorRset(a, m)) * x^card (BnorRset (a, m))"
+  apply (induct a m rule: BnorRset_induct)
+   prefer 2
+   apply (simplesubst BnorRset.simps)  --{*multiple redexes*}
+   apply (unfold Let_def, auto)
+  apply (simp add: Bnor_fin Bnor_mem_zle_swap)
+  apply (subst setprod_insert)
+    apply (rule_tac [2] Bnor_prod_power_aux)
+     apply (unfold inj_on_def)
+     apply (simp_all add: zmult_ac Bnor_fin finite_imageI
+       Bnor_mem_zle_swap)
+  done
+
+
+subsection {* Fermat *}
+
+lemma bijzcong_zcong_prod:
+    "(A, B) \<in> bijR (zcongm m) ==> [\<Prod>A = \<Prod>B] (mod m)"
+  apply (unfold zcongm_def)
+  apply (erule bijR.induct)
+   apply (subgoal_tac [2] "a \<notin> A \<and> b \<notin> B \<and> finite A \<and> finite B")
+    apply (auto intro: fin_bijRl fin_bijRr zcong_zmult)
+  done
+
+lemma Bnor_prod_zgcd [rule_format]:
+    "a < m --> zgcd (\<Prod>(BnorRset(a, m))) m = 1"
+  apply (induct a m rule: BnorRset_induct)
+   prefer 2
+   apply (subst BnorRset.simps)
+   apply (unfold Let_def, auto)
+  apply (simp add: Bnor_fin Bnor_mem_zle_swap)
+  apply (blast intro: zgcd_zgcd_zmult)
+  done
+
+theorem Euler_Fermat:
+    "0 < m ==> zgcd x m = 1 ==> [x^(phi m) = 1] (mod m)"
+  apply (unfold norRRset_def phi_def)
+  apply (case_tac "x = 0")
+   apply (case_tac [2] "m = 1")
+    apply (rule_tac [3] iffD1)
+     apply (rule_tac [3] k = "\<Prod>(BnorRset(m - 1, m))"
+       in zcong_cancel2)
+      prefer 5
+      apply (subst Bnor_prod_power [symmetric])
+        apply (rule_tac [7] Bnor_prod_zgcd, simp_all)
+  apply (rule bijzcong_zcong_prod)
+  apply (fold norRRset_def noXRRset_def)
+  apply (subst RRset2norRR_eq_norR [symmetric])
+    apply (rule_tac [3] inj_func_bijR, auto)
+     apply (unfold zcongm_def)
+     apply (rule_tac [2] RRset2norRR_correct1)
+       apply (rule_tac [5] RRset2norRR_inj)
+        apply (auto intro: order_less_le [THEN iffD2]
+	   simp add: noX_is_RRset)
+  apply (unfold noXRRset_def norRRset_def)
+  apply (rule finite_imageI)
+  apply (rule Bnor_fin)
+  done
+
+lemma Bnor_prime:
+  "\<lbrakk> zprime p; a < p \<rbrakk> \<Longrightarrow> card (BnorRset (a, p)) = nat a"
+  apply (induct a p rule: BnorRset.induct)
+  apply (subst BnorRset.simps)
+  apply (unfold Let_def, auto simp add:zless_zprime_imp_zrelprime)
+  apply (subgoal_tac "finite (BnorRset (a - 1,m))")
+   apply (subgoal_tac "a ~: BnorRset (a - 1,m)")
+    apply (auto simp add: card_insert_disjoint Suc_nat_eq_nat_zadd1)
+   apply (frule Bnor_mem_zle, arith)
+  apply (frule Bnor_fin)
+  done
+
+lemma phi_prime: "zprime p ==> phi p = nat (p - 1)"
+  apply (unfold phi_def norRRset_def)
+  apply (rule Bnor_prime, auto)
+  done
+
+theorem Little_Fermat:
+    "zprime p ==> \<not> p dvd x ==> [x^(nat (p - 1)) = 1] (mod p)"
+  apply (subst phi_prime [symmetric])
+   apply (rule_tac [2] Euler_Fermat)
+    apply (erule_tac [3] zprime_imp_zrelprime)
+    apply (unfold zprime_def, auto)
+  done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/EvenOdd.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,256 @@
+(*  Title:      HOL/Quadratic_Reciprocity/EvenOdd.thy
+    Authors:    Jeremy Avigad, David Gray, and Adam Kramer
+*)
+
+header {*Parity: Even and Odd Integers*}
+
+theory EvenOdd
+imports Int2
+begin
+
+definition
+  zOdd    :: "int set" where
+  "zOdd = {x. \<exists>k. x = 2 * k + 1}"
+
+definition
+  zEven   :: "int set" where
+  "zEven = {x. \<exists>k. x = 2 * k}"
+
+subsection {* Some useful properties about even and odd *}
+
+lemma zOddI [intro?]: "x = 2 * k + 1 \<Longrightarrow> x \<in> zOdd"
+  and zOddE [elim?]: "x \<in> zOdd \<Longrightarrow> (!!k. x = 2 * k + 1 \<Longrightarrow> C) \<Longrightarrow> C"
+  by (auto simp add: zOdd_def)
+
+lemma zEvenI [intro?]: "x = 2 * k \<Longrightarrow> x \<in> zEven"
+  and zEvenE [elim?]: "x \<in> zEven \<Longrightarrow> (!!k. x = 2 * k \<Longrightarrow> C) \<Longrightarrow> C"
+  by (auto simp add: zEven_def)
+
+lemma one_not_even: "~(1 \<in> zEven)"
+proof
+  assume "1 \<in> zEven"
+  then obtain k :: int where "1 = 2 * k" ..
+  then show False by arith
+qed
+
+lemma even_odd_conj: "~(x \<in> zOdd & x \<in> zEven)"
+proof -
+  {
+    fix a b
+    assume "2 * (a::int) = 2 * (b::int) + 1"
+    then have "2 * (a::int) - 2 * (b :: int) = 1"
+      by arith
+    then have "2 * (a - b) = 1"
+      by (auto simp add: zdiff_zmult_distrib)
+    moreover have "(2 * (a - b)):zEven"
+      by (auto simp only: zEven_def)
+    ultimately have False
+      by (auto simp add: one_not_even)
+  }
+  then show ?thesis
+    by (auto simp add: zOdd_def zEven_def)
+qed
+
+lemma even_odd_disj: "(x \<in> zOdd | x \<in> zEven)"
+  by (simp add: zOdd_def zEven_def) arith
+
+lemma not_odd_impl_even: "~(x \<in> zOdd) ==> x \<in> zEven"
+  using even_odd_disj by auto
+
+lemma odd_mult_odd_prop: "(x*y):zOdd ==> x \<in> zOdd"
+proof (rule classical)
+  assume "\<not> ?thesis"
+  then have "x \<in> zEven" by (rule not_odd_impl_even)
+  then obtain a where a: "x = 2 * a" ..
+  assume "x * y : zOdd"
+  then obtain b where "x * y = 2 * b + 1" ..
+  with a have "2 * a * y = 2 * b + 1" by simp
+  then have "2 * a * y - 2 * b = 1"
+    by arith
+  then have "2 * (a * y - b) = 1"
+    by (auto simp add: zdiff_zmult_distrib)
+  moreover have "(2 * (a * y - b)):zEven"
+    by (auto simp only: zEven_def)
+  ultimately have False
+    by (auto simp add: one_not_even)
+  then show ?thesis ..
+qed
+
+lemma odd_minus_one_even: "x \<in> zOdd ==> (x - 1):zEven"
+  by (auto simp add: zOdd_def zEven_def)
+
+lemma even_div_2_prop1: "x \<in> zEven ==> (x mod 2) = 0"
+  by (auto simp add: zEven_def)
+
+lemma even_div_2_prop2: "x \<in> zEven ==> (2 * (x div 2)) = x"
+  by (auto simp add: zEven_def)
+
+lemma even_plus_even: "[| x \<in> zEven; y \<in> zEven |] ==> x + y \<in> zEven"
+  apply (auto simp add: zEven_def)
+  apply (auto simp only: zadd_zmult_distrib2 [symmetric])
+  done
+
+lemma even_times_either: "x \<in> zEven ==> x * y \<in> zEven"
+  by (auto simp add: zEven_def)
+
+lemma even_minus_even: "[| x \<in> zEven; y \<in> zEven |] ==> x - y \<in> zEven"
+  apply (auto simp add: zEven_def)
+  apply (auto simp only: zdiff_zmult_distrib2 [symmetric])
+  done
+
+lemma odd_minus_odd: "[| x \<in> zOdd; y \<in> zOdd |] ==> x - y \<in> zEven"
+  apply (auto simp add: zOdd_def zEven_def)
+  apply (auto simp only: zdiff_zmult_distrib2 [symmetric])
+  done
+
+lemma even_minus_odd: "[| x \<in> zEven; y \<in> zOdd |] ==> x - y \<in> zOdd"
+  apply (auto simp add: zOdd_def zEven_def)
+  apply (rule_tac x = "k - ka - 1" in exI)
+  apply auto
+  done
+
+lemma odd_minus_even: "[| x \<in> zOdd; y \<in> zEven |] ==> x - y \<in> zOdd"
+  apply (auto simp add: zOdd_def zEven_def)
+  apply (auto simp only: zdiff_zmult_distrib2 [symmetric])
+  done
+
+lemma odd_times_odd: "[| x \<in> zOdd;  y \<in> zOdd |] ==> x * y \<in> zOdd"
+  apply (auto simp add: zOdd_def zadd_zmult_distrib zadd_zmult_distrib2)
+  apply (rule_tac x = "2 * ka * k + ka + k" in exI)
+  apply (auto simp add: zadd_zmult_distrib)
+  done
+
+lemma odd_iff_not_even: "(x \<in> zOdd) = (~ (x \<in> zEven))"
+  using even_odd_conj even_odd_disj by auto
+
+lemma even_product: "x * y \<in> zEven ==> x \<in> zEven | y \<in> zEven"
+  using odd_iff_not_even odd_times_odd by auto
+
+lemma even_diff: "x - y \<in> zEven = ((x \<in> zEven) = (y \<in> zEven))"
+proof
+  assume xy: "x - y \<in> zEven"
+  {
+    assume x: "x \<in> zEven"
+    have "y \<in> zEven"
+    proof (rule classical)
+      assume "\<not> ?thesis"
+      then have "y \<in> zOdd"
+        by (simp add: odd_iff_not_even)
+      with x have "x - y \<in> zOdd"
+        by (simp add: even_minus_odd)
+      with xy have False
+        by (auto simp add: odd_iff_not_even)
+      then show ?thesis ..
+    qed
+  } moreover {
+    assume y: "y \<in> zEven"
+    have "x \<in> zEven"
+    proof (rule classical)
+      assume "\<not> ?thesis"
+      then have "x \<in> zOdd"
+        by (auto simp add: odd_iff_not_even)
+      with y have "x - y \<in> zOdd"
+        by (simp add: odd_minus_even)
+      with xy have False
+        by (auto simp add: odd_iff_not_even)
+      then show ?thesis ..
+    qed
+  }
+  ultimately show "(x \<in> zEven) = (y \<in> zEven)"
+    by (auto simp add: odd_iff_not_even even_minus_even odd_minus_odd
+      even_minus_odd odd_minus_even)
+next
+  assume "(x \<in> zEven) = (y \<in> zEven)"
+  then show "x - y \<in> zEven"
+    by (auto simp add: odd_iff_not_even even_minus_even odd_minus_odd
+      even_minus_odd odd_minus_even)
+qed
+
+lemma neg_one_even_power: "[| x \<in> zEven; 0 \<le> x |] ==> (-1::int)^(nat x) = 1"
+proof -
+  assume "x \<in> zEven" and "0 \<le> x"
+  from `x \<in> zEven` obtain a where "x = 2 * a" ..
+  with `0 \<le> x` have "0 \<le> a" by simp
+  from `0 \<le> x` and `x = 2 * a` have "nat x = nat (2 * a)"
+    by simp
+  also from `x = 2 * a` have "nat (2 * a) = 2 * nat a"
+    by (simp add: nat_mult_distrib)
+  finally have "(-1::int)^nat x = (-1)^(2 * nat a)"
+    by simp
+  also have "... = ((-1::int)^2)^ (nat a)"
+    by (simp add: zpower_zpower [symmetric])
+  also have "(-1::int)^2 = 1"
+    by simp
+  finally show ?thesis
+    by simp
+qed
+
+lemma neg_one_odd_power: "[| x \<in> zOdd; 0 \<le> x |] ==> (-1::int)^(nat x) = -1"
+proof -
+  assume "x \<in> zOdd" and "0 \<le> x"
+  from `x \<in> zOdd` obtain a where "x = 2 * a + 1" ..
+  with `0 \<le> x` have a: "0 \<le> a" by simp
+  with `0 \<le> x` and `x = 2 * a + 1` have "nat x = nat (2 * a + 1)"
+    by simp
+  also from a have "nat (2 * a + 1) = 2 * nat a + 1"
+    by (auto simp add: nat_mult_distrib nat_add_distrib)
+  finally have "(-1::int)^nat x = (-1)^(2 * nat a + 1)"
+    by simp
+  also have "... = ((-1::int)^2)^ (nat a) * (-1)^1"
+    by (auto simp add: zpower_zpower [symmetric] zpower_zadd_distrib)
+  also have "(-1::int)^2 = 1"
+    by simp
+  finally show ?thesis
+    by simp
+qed
+
+lemma neg_one_power_parity: "[| 0 \<le> x; 0 \<le> y; (x \<in> zEven) = (y \<in> zEven) |] ==>
+    (-1::int)^(nat x) = (-1::int)^(nat y)"
+  using even_odd_disj [of x] even_odd_disj [of y]
+  by (auto simp add: neg_one_even_power neg_one_odd_power)
+
+
+lemma one_not_neg_one_mod_m: "2 < m ==> ~([1 = -1] (mod m))"
+  by (auto simp add: zcong_def zdvd_not_zless)
+
+lemma even_div_2_l: "[| y \<in> zEven; x < y |] ==> x div 2 < y div 2"
+proof -
+  assume "y \<in> zEven" and "x < y"
+  from `y \<in> zEven` obtain k where k: "y = 2 * k" ..
+  with `x < y` have "x < 2 * k" by simp
+  then have "x div 2 < k" by (auto simp add: div_prop1)
+  also have "k = (2 * k) div 2" by simp
+  finally have "x div 2 < 2 * k div 2" by simp
+  with k show ?thesis by simp
+qed
+
+lemma even_sum_div_2: "[| x \<in> zEven; y \<in> zEven |] ==> (x + y) div 2 = x div 2 + y div 2"
+  by (auto simp add: zEven_def)
+
+lemma even_prod_div_2: "[| x \<in> zEven |] ==> (x * y) div 2 = (x div 2) * y"
+  by (auto simp add: zEven_def)
+
+(* An odd prime is greater than 2 *)
+
+lemma zprime_zOdd_eq_grt_2: "zprime p ==> (p \<in> zOdd) = (2 < p)"
+  apply (auto simp add: zOdd_def zprime_def)
+  apply (drule_tac x = 2 in allE)
+  using odd_iff_not_even [of p]
+  apply (auto simp add: zOdd_def zEven_def)
+  done
+
+(* Powers of -1 and parity *)
+
+lemma neg_one_special: "finite A ==>
+    ((-1 :: int) ^ card A) * (-1 ^ card A) = 1"
+  by (induct set: finite) auto
+
+lemma neg_one_power: "(-1::int)^n = 1 | (-1::int)^n = -1"
+  by (induct n) auto
+
+lemma neg_one_power_eq_mod_m: "[| 2 < m; [(-1::int)^j = (-1::int)^k] (mod m) |]
+    ==> ((-1::int)^j = (-1::int)^k)"
+  using neg_one_power [of j] and ListMem.insert neg_one_power [of k]
+  by (auto simp add: one_not_neg_one_mod_m zcong_sym)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/Factorization.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,339 @@
+(*  Author:     Thomas Marthedal Rasmussen
+    Copyright   2000  University of Cambridge
+*)
+
+header {* Fundamental Theorem of Arithmetic (unique factorization into primes) *}
+
+theory Factorization
+imports Main "~~/src/HOL/Old_Number_Theory/Primes" Permutation
+begin
+
+
+subsection {* Definitions *}
+
+definition
+  primel :: "nat list => bool" where
+  "primel xs = (\<forall>p \<in> set xs. prime p)"
+
+consts
+  nondec :: "nat list => bool "
+  prod :: "nat list => nat"
+  oinsert :: "nat => nat list => nat list"
+  sort :: "nat list => nat list"
+
+primrec
+  "nondec [] = True"
+  "nondec (x # xs) = (case xs of [] => True | y # ys => x \<le> y \<and> nondec xs)"
+
+primrec
+  "prod [] = Suc 0"
+  "prod (x # xs) = x * prod xs"
+
+primrec
+  "oinsert x [] = [x]"
+  "oinsert x (y # ys) = (if x \<le> y then x # y # ys else y # oinsert x ys)"
+
+primrec
+  "sort [] = []"
+  "sort (x # xs) = oinsert x (sort xs)"
+
+
+subsection {* Arithmetic *}
+
+lemma one_less_m: "(m::nat) \<noteq> m * k ==> m \<noteq> Suc 0 ==> Suc 0 < m"
+  apply (cases m)
+   apply auto
+  done
+
+lemma one_less_k: "(m::nat) \<noteq> m * k ==> Suc 0 < m * k ==> Suc 0 < k"
+  apply (cases k)
+   apply auto
+  done
+
+lemma mult_left_cancel: "(0::nat) < k ==> k * n = k * m ==> n = m"
+  apply auto
+  done
+
+lemma mn_eq_m_one: "(0::nat) < m ==> m * n = m ==> n = Suc 0"
+  apply (cases n)
+   apply auto
+  done
+
+lemma prod_mn_less_k:
+    "(0::nat) < n ==> 0 < k ==> Suc 0 < m ==> m * n = k ==> n < k"
+  apply (induct m)
+   apply auto
+  done
+
+
+subsection {* Prime list and product *}
+
+lemma prod_append: "prod (xs @ ys) = prod xs * prod ys"
+  apply (induct xs)
+   apply (simp_all add: mult_assoc)
+  done
+
+lemma prod_xy_prod:
+    "prod (x # xs) = prod (y # ys) ==> x * prod xs = y * prod ys"
+  apply auto
+  done
+
+lemma primel_append: "primel (xs @ ys) = (primel xs \<and> primel ys)"
+  apply (unfold primel_def)
+  apply auto
+  done
+
+lemma prime_primel: "prime n ==> primel [n] \<and> prod [n] = n"
+  apply (unfold primel_def)
+  apply auto
+  done
+
+lemma prime_nd_one: "prime p ==> \<not> p dvd Suc 0"
+  apply (unfold prime_def dvd_def)
+  apply auto
+  done
+
+lemma hd_dvd_prod: "prod (x # xs) = prod ys ==> x dvd (prod ys)" 
+  by (metis dvd_mult_left dvd_refl prod.simps(2))
+
+lemma primel_tl: "primel (x # xs) ==> primel xs"
+  apply (unfold primel_def)
+  apply auto
+  done
+
+lemma primel_hd_tl: "(primel (x # xs)) = (prime x \<and> primel xs)"
+  apply (unfold primel_def)
+  apply auto
+  done
+
+lemma primes_eq: "prime p ==> prime q ==> p dvd q ==> p = q"
+  apply (unfold prime_def)
+  apply auto
+  done
+
+lemma primel_one_empty: "primel xs ==> prod xs = Suc 0 ==> xs = []"
+  apply (cases xs)
+   apply (simp_all add: primel_def prime_def)
+  done
+
+lemma prime_g_one: "prime p ==> Suc 0 < p"
+  apply (unfold prime_def)
+  apply auto
+  done
+
+lemma prime_g_zero: "prime p ==> 0 < p"
+  apply (unfold prime_def)
+  apply auto
+  done
+
+lemma primel_nempty_g_one:
+    "primel xs \<Longrightarrow> xs \<noteq> [] \<Longrightarrow> Suc 0 < prod xs"
+  apply (induct xs)
+   apply simp
+  apply (fastsimp simp: primel_def prime_def elim: one_less_mult)
+  done
+
+lemma primel_prod_gz: "primel xs ==> 0 < prod xs"
+  apply (induct xs)
+   apply (auto simp: primel_def prime_def)
+  done
+
+
+subsection {* Sorting *}
+
+lemma nondec_oinsert: "nondec xs \<Longrightarrow> nondec (oinsert x xs)"
+  apply (induct xs)
+   apply simp
+   apply (case_tac xs)
+    apply (simp_all cong del: list.weak_case_cong)
+  done
+
+lemma nondec_sort: "nondec (sort xs)"
+  apply (induct xs)
+   apply simp_all
+  apply (erule nondec_oinsert)
+  done
+
+lemma x_less_y_oinsert: "x \<le> y ==> l = y # ys ==> x # l = oinsert x l"
+  apply simp_all
+  done
+
+lemma nondec_sort_eq [rule_format]: "nondec xs \<longrightarrow> xs = sort xs"
+  apply (induct xs)
+   apply safe
+    apply simp_all
+   apply (case_tac xs)
+    apply simp_all
+  apply (case_tac xs)
+   apply simp
+  apply (rule_tac y = aa and ys = list in x_less_y_oinsert)
+   apply simp_all
+  done
+
+lemma oinsert_x_y: "oinsert x (oinsert y l) = oinsert y (oinsert x l)"
+  apply (induct l)
+  apply auto
+  done
+
+
+subsection {* Permutation *}
+
+lemma perm_primel [rule_format]: "xs <~~> ys ==> primel xs --> primel ys"
+  apply (unfold primel_def)
+  apply (induct set: perm)
+     apply simp
+    apply simp
+   apply (simp (no_asm))
+   apply blast
+  apply blast
+  done
+
+lemma perm_prod: "xs <~~> ys ==> prod xs = prod ys"
+  apply (induct set: perm)
+     apply (simp_all add: mult_ac)
+  done
+
+lemma perm_subst_oinsert: "xs <~~> ys ==> oinsert a xs <~~> oinsert a ys"
+  apply (induct set: perm)
+     apply auto
+  done
+
+lemma perm_oinsert: "x # xs <~~> oinsert x xs"
+  apply (induct xs)
+   apply auto
+  done
+
+lemma perm_sort: "xs <~~> sort xs"
+  apply (induct xs)
+  apply (auto intro: perm_oinsert elim: perm_subst_oinsert)
+  done
+
+lemma perm_sort_eq: "xs <~~> ys ==> sort xs = sort ys"
+  apply (induct set: perm)
+     apply (simp_all add: oinsert_x_y)
+  done
+
+
+subsection {* Existence *}
+
+lemma ex_nondec_lemma:
+    "primel xs ==> \<exists>ys. primel ys \<and> nondec ys \<and> prod ys = prod xs"
+  apply (blast intro: nondec_sort perm_prod perm_primel perm_sort perm_sym)
+  done
+
+lemma not_prime_ex_mk:
+  "Suc 0 < n \<and> \<not> prime n ==>
+    \<exists>m k. Suc 0 < m \<and> Suc 0 < k \<and> m < n \<and> k < n \<and> n = m * k"
+  apply (unfold prime_def dvd_def)
+  apply (auto intro: n_less_m_mult_n n_less_n_mult_m one_less_m one_less_k)
+  done
+
+lemma split_primel:
+  "primel xs \<Longrightarrow> primel ys \<Longrightarrow> \<exists>l. primel l \<and> prod l = prod xs * prod ys"
+  apply (rule exI)
+  apply safe
+   apply (rule_tac [2] prod_append)
+  apply (simp add: primel_append)
+  done
+
+lemma factor_exists [rule_format]: "Suc 0 < n --> (\<exists>l. primel l \<and> prod l = n)"
+  apply (induct n rule: nat_less_induct)
+  apply (rule impI)
+  apply (case_tac "prime n")
+   apply (rule exI)
+   apply (erule prime_primel)
+  apply (cut_tac n = n in not_prime_ex_mk)
+   apply (auto intro!: split_primel)
+  done
+
+lemma nondec_factor_exists: "Suc 0 < n ==> \<exists>l. primel l \<and> nondec l \<and> prod l = n"
+  apply (erule factor_exists [THEN exE])
+  apply (blast intro!: ex_nondec_lemma)
+  done
+
+
+subsection {* Uniqueness *}
+
+lemma prime_dvd_mult_list [rule_format]:
+    "prime p ==> p dvd (prod xs) --> (\<exists>m. m:set xs \<and> p dvd m)"
+  apply (induct xs)
+   apply (force simp add: prime_def)
+   apply (force dest: prime_dvd_mult)
+  done
+
+lemma hd_xs_dvd_prod:
+  "primel (x # xs) ==> primel ys ==> prod (x # xs) = prod ys
+    ==> \<exists>m. m \<in> set ys \<and> x dvd m"
+  apply (rule prime_dvd_mult_list)
+   apply (simp add: primel_hd_tl)
+  apply (erule hd_dvd_prod)
+  done
+
+lemma prime_dvd_eq: "primel (x # xs) ==> primel ys ==> m \<in> set ys ==> x dvd m ==> x = m"
+  apply (rule primes_eq)
+    apply (auto simp add: primel_def primel_hd_tl)
+  done
+
+lemma hd_xs_eq_prod:
+  "primel (x # xs) ==>
+    primel ys ==> prod (x # xs) = prod ys ==> x \<in> set ys"
+  apply (frule hd_xs_dvd_prod)
+    apply auto
+  apply (drule prime_dvd_eq)
+     apply auto
+  done
+
+lemma perm_primel_ex:
+  "primel (x # xs) ==>
+    primel ys ==> prod (x # xs) = prod ys ==> \<exists>l. ys <~~> (x # l)"
+  apply (rule exI)
+  apply (rule perm_remove)
+  apply (erule hd_xs_eq_prod)
+   apply simp_all
+  done
+
+lemma primel_prod_less:
+  "primel (x # xs) ==>
+    primel ys ==> prod (x # xs) = prod ys ==> prod xs < prod ys"
+  by (metis less_asym linorder_neqE_nat mult_less_cancel2 nat_0_less_mult_iff
+    nat_less_le nat_mult_1 prime_def primel_hd_tl primel_prod_gz prod.simps(2))
+
+lemma prod_one_empty:
+    "primel xs ==> p * prod xs = p ==> prime p ==> xs = []"
+  apply (auto intro: primel_one_empty simp add: prime_def)
+  done
+
+lemma uniq_ex_aux:
+  "\<forall>m. m < prod ys --> (\<forall>xs ys. primel xs \<and> primel ys \<and>
+      prod xs = prod ys \<and> prod xs = m --> xs <~~> ys) ==>
+    primel list ==> primel x ==> prod list = prod x ==> prod x < prod ys
+    ==> x <~~> list"
+  apply simp
+  done
+
+lemma factor_unique [rule_format]:
+  "\<forall>xs ys. primel xs \<and> primel ys \<and> prod xs = prod ys \<and> prod xs = n
+    --> xs <~~> ys"
+  apply (induct n rule: nat_less_induct)
+  apply safe
+  apply (case_tac xs)
+   apply (force intro: primel_one_empty)
+  apply (rule perm_primel_ex [THEN exE])
+     apply simp_all
+  apply (rule perm.trans [THEN perm_sym])
+  apply assumption
+  apply (rule perm.Cons)
+  apply (case_tac "x = []")
+   apply (metis perm_prod perm_refl prime_primel primel_hd_tl primel_tl prod_one_empty)
+  apply (metis nat_0_less_mult_iff nat_mult_eq_cancel1 perm_primel perm_prod primel_prod_gz primel_prod_less primel_tl prod.simps(2))
+  done
+
+lemma perm_nondec_unique:
+    "xs <~~> ys ==> nondec xs ==> nondec ys ==> xs = ys"
+  by (metis nondec_sort_eq perm_sort_eq)
+
+theorem unique_prime_factorization [rule_format]:
+    "\<forall>n. Suc 0 < n --> (\<exists>!l. primel l \<and> nondec l \<and> prod l = n)"
+  by (metis factor_unique nondec_factor_exists perm_nondec_unique)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/Fib.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,150 @@
+(*  ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1997  University of Cambridge
+*)
+
+header {* The Fibonacci function *}
+
+theory Fib
+imports Primes
+begin
+
+text {*
+  Fibonacci numbers: proofs of laws taken from:
+  R. L. Graham, D. E. Knuth, O. Patashnik.  Concrete Mathematics.
+  (Addison-Wesley, 1989)
+
+  \bigskip
+*}
+
+fun fib :: "nat \<Rightarrow> nat"
+where
+         "fib 0 = 0"
+|        "fib (Suc 0) = 1"
+| fib_2: "fib (Suc (Suc n)) = fib n + fib (Suc n)"
+
+text {*
+  \medskip The difficulty in these proofs is to ensure that the
+  induction hypotheses are applied before the definition of @{term
+  fib}.  Towards this end, the @{term fib} equations are not declared
+  to the Simplifier and are applied very selectively at first.
+*}
+
+text{*We disable @{text fib.fib_2fib_2} for simplification ...*}
+declare fib_2 [simp del]
+
+text{*...then prove a version that has a more restrictive pattern.*}
+lemma fib_Suc3: "fib (Suc (Suc (Suc n))) = fib (Suc n) + fib (Suc (Suc n))"
+  by (rule fib_2)
+
+text {* \medskip Concrete Mathematics, page 280 *}
+
+lemma fib_add: "fib (Suc (n + k)) = fib (Suc k) * fib (Suc n) + fib k * fib n"
+proof (induct n rule: fib.induct)
+  case 1 show ?case by simp
+next
+  case 2 show ?case  by (simp add: fib_2)
+next
+  case 3 thus ?case by (simp add: fib_2 add_mult_distrib2)
+qed
+
+lemma fib_Suc_neq_0: "fib (Suc n) \<noteq> 0"
+  apply (induct n rule: fib.induct)
+    apply (simp_all add: fib_2)
+  done
+
+lemma fib_Suc_gr_0: "0 < fib (Suc n)"
+  by (insert fib_Suc_neq_0 [of n], simp)  
+
+lemma fib_gr_0: "0 < n ==> 0 < fib n"
+  by (case_tac n, auto simp add: fib_Suc_gr_0) 
+
+
+text {*
+  \medskip Concrete Mathematics, page 278: Cassini's identity.  The proof is
+  much easier using integers, not natural numbers!
+*}
+
+lemma fib_Cassini_int:
+ "int (fib (Suc (Suc n)) * fib n) =
+  (if n mod 2 = 0 then int (fib (Suc n) * fib (Suc n)) - 1
+   else int (fib (Suc n) * fib (Suc n)) + 1)"
+proof(induct n rule: fib.induct)
+  case 1 thus ?case by (simp add: fib_2)
+next
+  case 2 thus ?case by (simp add: fib_2 mod_Suc)
+next 
+  case (3 x) 
+  have "Suc 0 \<noteq> x mod 2 \<longrightarrow> x mod 2 = 0" by presburger
+  with "3.hyps" show ?case by (simp add: fib.simps add_mult_distrib add_mult_distrib2)
+qed
+
+text{*We now obtain a version for the natural numbers via the coercion 
+   function @{term int}.*}
+theorem fib_Cassini:
+ "fib (Suc (Suc n)) * fib n =
+  (if n mod 2 = 0 then fib (Suc n) * fib (Suc n) - 1
+   else fib (Suc n) * fib (Suc n) + 1)"
+  apply (rule int_int_eq [THEN iffD1]) 
+  apply (simp add: fib_Cassini_int)
+  apply (subst zdiff_int [symmetric]) 
+   apply (insert fib_Suc_gr_0 [of n], simp_all)
+  done
+
+
+text {* \medskip Toward Law 6.111 of Concrete Mathematics *}
+
+lemma gcd_fib_Suc_eq_1: "gcd (fib n) (fib (Suc n)) = Suc 0"
+  apply (induct n rule: fib.induct)
+    prefer 3
+    apply (simp add: gcd_commute fib_Suc3)
+   apply (simp_all add: fib_2)
+  done
+
+lemma gcd_fib_add: "gcd (fib m) (fib (n + m)) = gcd (fib m) (fib n)"
+  apply (simp add: gcd_commute [of "fib m"])
+  apply (case_tac m)
+   apply simp 
+  apply (simp add: fib_add)
+  apply (simp add: add_commute gcd_non_0 [OF fib_Suc_gr_0])
+  apply (simp add: gcd_non_0 [OF fib_Suc_gr_0, symmetric])
+  apply (simp add: gcd_fib_Suc_eq_1 gcd_mult_cancel)
+  done
+
+lemma gcd_fib_diff: "m \<le> n ==> gcd (fib m) (fib (n - m)) = gcd (fib m) (fib n)"
+  by (simp add: gcd_fib_add [symmetric, of _ "n-m"]) 
+
+lemma gcd_fib_mod: "0 < m ==> gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
+proof (induct n rule: less_induct)
+  case (less n)
+  from less.prems have pos_m: "0 < m" .
+  show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
+  proof (cases "m < n")
+    case True note m_n = True
+    then have m_n': "m \<le> n" by auto
+    with pos_m have pos_n: "0 < n" by auto
+    with pos_m m_n have diff: "n - m < n" by auto
+    have "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib ((n - m) mod m))"
+    by (simp add: mod_if [of n]) (insert m_n, auto)
+    also have "\<dots> = gcd (fib m) (fib (n - m))" by (simp add: less.hyps diff pos_m)
+    also have "\<dots> = gcd (fib m) (fib n)" by (simp add: gcd_fib_diff m_n')
+    finally show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)" .
+  next
+    case False then show "gcd (fib m) (fib (n mod m)) = gcd (fib m) (fib n)"
+    by (cases "m = n") auto
+  qed
+qed
+
+lemma fib_gcd: "fib (gcd m n) = gcd (fib m) (fib n)"  -- {* Law 6.111 *}
+  apply (induct m n rule: gcd_induct)
+  apply (simp_all add: gcd_non_0 gcd_commute gcd_fib_mod)
+  done
+
+theorem fib_mult_eq_setsum:
+    "fib (Suc n) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
+  apply (induct n rule: fib.induct)
+    apply (auto simp add: atMost_Suc fib_2)
+  apply (simp add: add_mult_distrib add_mult_distrib2)
+  done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/Finite2.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,223 @@
+(*  Title:      HOL/Quadratic_Reciprocity/Finite2.thy
+    ID:         $Id$
+    Authors:    Jeremy Avigad, David Gray, and Adam Kramer
+*)
+
+header {*Finite Sets and Finite Sums*}
+
+theory Finite2
+imports Main IntFact Infinite_Set
+begin
+
+text{*
+  These are useful for combinatorial and number-theoretic counting
+  arguments.
+*}
+
+
+subsection {* Useful properties of sums and products *}
+
+lemma setsum_same_function_zcong:
+  assumes a: "\<forall>x \<in> S. [f x = g x](mod m)"
+  shows "[setsum f S = setsum g S] (mod m)"
+proof cases
+  assume "finite S"
+  thus ?thesis using a by induct (simp_all add: zcong_zadd)
+next
+  assume "infinite S" thus ?thesis by(simp add:setsum_def)
+qed
+
+lemma setprod_same_function_zcong:
+  assumes a: "\<forall>x \<in> S. [f x = g x](mod m)"
+  shows "[setprod f S = setprod g S] (mod m)"
+proof cases
+  assume "finite S"
+  thus ?thesis using a by induct (simp_all add: zcong_zmult)
+next
+  assume "infinite S" thus ?thesis by(simp add:setprod_def)
+qed
+
+lemma setsum_const: "finite X ==> setsum (%x. (c :: int)) X = c * int(card X)"
+  apply (induct set: finite)
+  apply (auto simp add: left_distrib right_distrib int_eq_of_nat)
+  done
+
+lemma setsum_const2: "finite X ==> int (setsum (%x. (c :: nat)) X) =
+    int(c) * int(card X)"
+  apply (induct set: finite)
+  apply (auto simp add: zadd_zmult_distrib2)
+  done
+
+lemma setsum_const_mult: "finite A ==> setsum (%x. c * ((f x)::int)) A =
+    c * setsum f A"
+  by (induct set: finite) (auto simp add: zadd_zmult_distrib2)
+
+
+subsection {* Cardinality of explicit finite sets *}
+
+lemma finite_surjI: "[| B \<subseteq> f ` A; finite A |] ==> finite B"
+  by (simp add: finite_subset finite_imageI)
+
+lemma bdd_nat_set_l_finite: "finite {y::nat . y < x}"
+  by (rule bounded_nat_set_is_finite) blast
+
+lemma bdd_nat_set_le_finite: "finite {y::nat . y \<le> x}"
+proof -
+  have "{y::nat . y \<le> x} = {y::nat . y < Suc x}" by auto
+  then show ?thesis by (auto simp add: bdd_nat_set_l_finite)
+qed
+
+lemma  bdd_int_set_l_finite: "finite {x::int. 0 \<le> x & x < n}"
+  apply (subgoal_tac " {(x :: int). 0 \<le> x & x < n} \<subseteq>
+      int ` {(x :: nat). x < nat n}")
+   apply (erule finite_surjI)
+   apply (auto simp add: bdd_nat_set_l_finite image_def)
+  apply (rule_tac x = "nat x" in exI, simp)
+  done
+
+lemma bdd_int_set_le_finite: "finite {x::int. 0 \<le> x & x \<le> n}"
+  apply (subgoal_tac "{x. 0 \<le> x & x \<le> n} = {x. 0 \<le> x & x < n + 1}")
+   apply (erule ssubst)
+   apply (rule bdd_int_set_l_finite)
+  apply auto
+  done
+
+lemma bdd_int_set_l_l_finite: "finite {x::int. 0 < x & x < n}"
+proof -
+  have "{x::int. 0 < x & x < n} \<subseteq> {x::int. 0 \<le> x & x < n}"
+    by auto
+  then show ?thesis by (auto simp add: bdd_int_set_l_finite finite_subset)
+qed
+
+lemma bdd_int_set_l_le_finite: "finite {x::int. 0 < x & x \<le> n}"
+proof -
+  have "{x::int. 0 < x & x \<le> n} \<subseteq> {x::int. 0 \<le> x & x \<le> n}"
+    by auto
+  then show ?thesis by (auto simp add: bdd_int_set_le_finite finite_subset)
+qed
+
+lemma card_bdd_nat_set_l: "card {y::nat . y < x} = x"
+proof (induct x)
+  case 0
+  show "card {y::nat . y < 0} = 0" by simp
+next
+  case (Suc n)
+  have "{y. y < Suc n} = insert n {y. y < n}"
+    by auto
+  then have "card {y. y < Suc n} = card (insert n {y. y < n})"
+    by auto
+  also have "... = Suc (card {y. y < n})"
+    by (rule card_insert_disjoint) (auto simp add: bdd_nat_set_l_finite)
+  finally show "card {y. y < Suc n} = Suc n"
+    using `card {y. y < n} = n` by simp
+qed
+
+lemma card_bdd_nat_set_le: "card { y::nat. y \<le> x} = Suc x"
+proof -
+  have "{y::nat. y \<le> x} = { y::nat. y < Suc x}"
+    by auto
+  then show ?thesis by (auto simp add: card_bdd_nat_set_l)
+qed
+
+lemma card_bdd_int_set_l: "0 \<le> (n::int) ==> card {y. 0 \<le> y & y < n} = nat n"
+proof -
+  assume "0 \<le> n"
+  have "inj_on (%y. int y) {y. y < nat n}"
+    by (auto simp add: inj_on_def)
+  hence "card (int ` {y. y < nat n}) = card {y. y < nat n}"
+    by (rule card_image)
+  also from `0 \<le> n` have "int ` {y. y < nat n} = {y. 0 \<le> y & y < n}"
+    apply (auto simp add: zless_nat_eq_int_zless image_def)
+    apply (rule_tac x = "nat x" in exI)
+    apply (auto simp add: nat_0_le)
+    done
+  also have "card {y. y < nat n} = nat n"
+    by (rule card_bdd_nat_set_l)
+  finally show "card {y. 0 \<le> y & y < n} = nat n" .
+qed
+
+lemma card_bdd_int_set_le: "0 \<le> (n::int) ==> card {y. 0 \<le> y & y \<le> n} =
+  nat n + 1"
+proof -
+  assume "0 \<le> n"
+  moreover have "{y. 0 \<le> y & y \<le> n} = {y. 0 \<le> y & y < n+1}" by auto
+  ultimately show ?thesis
+    using card_bdd_int_set_l [of "n + 1"]
+    by (auto simp add: nat_add_distrib)
+qed
+
+lemma card_bdd_int_set_l_le: "0 \<le> (n::int) ==>
+    card {x. 0 < x & x \<le> n} = nat n"
+proof -
+  assume "0 \<le> n"
+  have "inj_on (%x. x+1) {x. 0 \<le> x & x < n}"
+    by (auto simp add: inj_on_def)
+  hence "card ((%x. x+1) ` {x. 0 \<le> x & x < n}) =
+     card {x. 0 \<le> x & x < n}"
+    by (rule card_image)
+  also from `0 \<le> n` have "... = nat n"
+    by (rule card_bdd_int_set_l)
+  also have "(%x. x + 1) ` {x. 0 \<le> x & x < n} = {x. 0 < x & x<= n}"
+    apply (auto simp add: image_def)
+    apply (rule_tac x = "x - 1" in exI)
+    apply arith
+    done
+  finally show "card {x. 0 < x & x \<le> n} = nat n" .
+qed
+
+lemma card_bdd_int_set_l_l: "0 < (n::int) ==>
+  card {x. 0 < x & x < n} = nat n - 1"
+proof -
+  assume "0 < n"
+  moreover have "{x. 0 < x & x < n} = {x. 0 < x & x \<le> n - 1}"
+    by simp
+  ultimately show ?thesis
+    using insert card_bdd_int_set_l_le [of "n - 1"]
+    by (auto simp add: nat_diff_distrib)
+qed
+
+lemma int_card_bdd_int_set_l_l: "0 < n ==>
+    int(card {x. 0 < x & x < n}) = n - 1"
+  apply (auto simp add: card_bdd_int_set_l_l)
+  done
+
+lemma int_card_bdd_int_set_l_le: "0 \<le> n ==>
+    int(card {x. 0 < x & x \<le> n}) = n"
+  by (auto simp add: card_bdd_int_set_l_le)
+
+
+subsection {* Cardinality of finite cartesian products *}
+
+(* FIXME could be useful in general but not needed here
+lemma insert_Sigma [simp]: "(insert x A) <*> B = ({ x } <*> B) \<union> (A <*> B)"
+  by blast
+ *)
+
+text {* Lemmas for counting arguments. *}
+
+lemma setsum_bij_eq: "[| finite A; finite B; f ` A \<subseteq> B; inj_on f A;
+    g ` B \<subseteq> A; inj_on g B |] ==> setsum g B = setsum (g \<circ> f) A"
+  apply (frule_tac h = g and f = f in setsum_reindex)
+  apply (subgoal_tac "setsum g B = setsum g (f ` A)")
+   apply (simp add: inj_on_def)
+  apply (subgoal_tac "card A = card B")
+   apply (drule_tac A = "f ` A" and B = B in card_seteq)
+     apply (auto simp add: card_image)
+  apply (frule_tac A = A and B = B and f = f in card_inj_on_le, auto)
+  apply (frule_tac A = B and B = A and f = g in card_inj_on_le)
+    apply auto
+  done
+
+lemma setprod_bij_eq: "[| finite A; finite B; f ` A \<subseteq> B; inj_on f A;
+    g ` B \<subseteq> A; inj_on g B |] ==> setprod g B = setprod (g \<circ> f) A"
+  apply (frule_tac h = g and f = f in setprod_reindex)
+  apply (subgoal_tac "setprod g B = setprod g (f ` A)")
+   apply (simp add: inj_on_def)
+  apply (subgoal_tac "card A = card B")
+   apply (drule_tac A = "f ` A" and B = B in card_seteq)
+     apply (auto simp add: card_image)
+  apply (frule_tac A = A and B = B and f = f in card_inj_on_le, auto)
+  apply (frule_tac A = B and B = A and f = g in card_inj_on_le, auto)
+  done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/Gauss.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,535 @@
+(*  Title:      HOL/Quadratic_Reciprocity/Gauss.thy
+    ID:         $Id$
+    Authors:    Jeremy Avigad, David Gray, and Adam Kramer)
+*)
+
+header {* Gauss' Lemma *}
+
+theory Gauss
+imports Euler
+begin
+
+locale GAUSS =
+  fixes p :: "int"
+  fixes a :: "int"
+
+  assumes p_prime: "zprime p"
+  assumes p_g_2: "2 < p"
+  assumes p_a_relprime: "~[a = 0](mod p)"
+  assumes a_nonzero:    "0 < a"
+begin
+
+definition
+  A :: "int set" where
+  "A = {(x::int). 0 < x & x \<le> ((p - 1) div 2)}"
+
+definition
+  B :: "int set" where
+  "B = (%x. x * a) ` A"
+
+definition
+  C :: "int set" where
+  "C = StandardRes p ` B"
+
+definition
+  D :: "int set" where
+  "D = C \<inter> {x. x \<le> ((p - 1) div 2)}"
+
+definition
+  E :: "int set" where
+  "E = C \<inter> {x. ((p - 1) div 2) < x}"
+
+definition
+  F :: "int set" where
+  "F = (%x. (p - x)) ` E"
+
+
+subsection {* Basic properties of p *}
+
+lemma p_odd: "p \<in> zOdd"
+  by (auto simp add: p_prime p_g_2 zprime_zOdd_eq_grt_2)
+
+lemma p_g_0: "0 < p"
+  using p_g_2 by auto
+
+lemma int_nat: "int (nat ((p - 1) div 2)) = (p - 1) div 2"
+  using ListMem.insert p_g_2 by (auto simp add: pos_imp_zdiv_nonneg_iff)
+
+lemma p_minus_one_l: "(p - 1) div 2 < p"
+proof -
+  have "(p - 1) div 2 \<le> (p - 1) div 1"
+    by (rule zdiv_mono2) (auto simp add: p_g_0)
+  also have "\<dots> = p - 1" by simp
+  finally show ?thesis by simp
+qed
+
+lemma p_eq: "p = (2 * (p - 1) div 2) + 1"
+  using div_mult_self1_is_id [of 2 "p - 1"] by auto
+
+
+lemma (in -) zodd_imp_zdiv_eq: "x \<in> zOdd ==> 2 * (x - 1) div 2 = 2 * ((x - 1) div 2)"
+  apply (frule odd_minus_one_even)
+  apply (simp add: zEven_def)
+  apply (subgoal_tac "2 \<noteq> 0")
+  apply (frule_tac b = "2 :: int" and a = "x - 1" in div_mult_self1_is_id)
+  apply (auto simp add: even_div_2_prop2)
+  done
+
+
+lemma p_eq2: "p = (2 * ((p - 1) div 2)) + 1"
+  apply (insert p_eq p_prime p_g_2 zprime_zOdd_eq_grt_2 [of p], auto)
+  apply (frule zodd_imp_zdiv_eq, auto)
+  done
+
+
+subsection {* Basic Properties of the Gauss Sets *}
+
+lemma finite_A: "finite (A)"
+  apply (auto simp add: A_def)
+  apply (subgoal_tac "{x. 0 < x & x \<le> (p - 1) div 2} \<subseteq> {x. 0 \<le> x & x < 1 + (p - 1) div 2}")
+  apply (auto simp add: bdd_int_set_l_finite finite_subset)
+  done
+
+lemma finite_B: "finite (B)"
+  by (auto simp add: B_def finite_A finite_imageI)
+
+lemma finite_C: "finite (C)"
+  by (auto simp add: C_def finite_B finite_imageI)
+
+lemma finite_D: "finite (D)"
+  by (auto simp add: D_def finite_Int finite_C)
+
+lemma finite_E: "finite (E)"
+  by (auto simp add: E_def finite_Int finite_C)
+
+lemma finite_F: "finite (F)"
+  by (auto simp add: F_def finite_E finite_imageI)
+
+lemma C_eq: "C = D \<union> E"
+  by (auto simp add: C_def D_def E_def)
+
+lemma A_card_eq: "card A = nat ((p - 1) div 2)"
+  apply (auto simp add: A_def)
+  apply (insert int_nat)
+  apply (erule subst)
+  apply (auto simp add: card_bdd_int_set_l_le)
+  done
+
+lemma inj_on_xa_A: "inj_on (%x. x * a) A"
+  using a_nonzero by (simp add: A_def inj_on_def)
+
+lemma A_res: "ResSet p A"
+  apply (auto simp add: A_def ResSet_def)
+  apply (rule_tac m = p in zcong_less_eq)
+  apply (insert p_g_2, auto)
+  done
+
+lemma B_res: "ResSet p B"
+  apply (insert p_g_2 p_a_relprime p_minus_one_l)
+  apply (auto simp add: B_def)
+  apply (rule ResSet_image)
+  apply (auto simp add: A_res)
+  apply (auto simp add: A_def)
+proof -
+  fix x fix y
+  assume a: "[x * a = y * a] (mod p)"
+  assume b: "0 < x"
+  assume c: "x \<le> (p - 1) div 2"
+  assume d: "0 < y"
+  assume e: "y \<le> (p - 1) div 2"
+  from a p_a_relprime p_prime a_nonzero zcong_cancel [of p a x y]
+  have "[x = y](mod p)"
+    by (simp add: zprime_imp_zrelprime zcong_def p_g_0 order_le_less)
+  with zcong_less_eq [of x y p] p_minus_one_l
+      order_le_less_trans [of x "(p - 1) div 2" p]
+      order_le_less_trans [of y "(p - 1) div 2" p] show "x = y"
+    by (simp add: prems p_minus_one_l p_g_0)
+qed
+
+lemma SR_B_inj: "inj_on (StandardRes p) B"
+  apply (auto simp add: B_def StandardRes_def inj_on_def A_def prems)
+proof -
+  fix x fix y
+  assume a: "x * a mod p = y * a mod p"
+  assume b: "0 < x"
+  assume c: "x \<le> (p - 1) div 2"
+  assume d: "0 < y"
+  assume e: "y \<le> (p - 1) div 2"
+  assume f: "x \<noteq> y"
+  from a have "[x * a = y * a](mod p)"
+    by (simp add: zcong_zmod_eq p_g_0)
+  with p_a_relprime p_prime a_nonzero zcong_cancel [of p a x y]
+  have "[x = y](mod p)"
+    by (simp add: zprime_imp_zrelprime zcong_def p_g_0 order_le_less)
+  with zcong_less_eq [of x y p] p_minus_one_l
+    order_le_less_trans [of x "(p - 1) div 2" p]
+    order_le_less_trans [of y "(p - 1) div 2" p] have "x = y"
+    by (simp add: prems p_minus_one_l p_g_0)
+  then have False
+    by (simp add: f)
+  then show "a = 0"
+    by simp
+qed
+
+lemma inj_on_pminusx_E: "inj_on (%x. p - x) E"
+  apply (auto simp add: E_def C_def B_def A_def)
+  apply (rule_tac g = "%x. -1 * (x - p)" in inj_on_inverseI)
+  apply auto
+  done
+
+lemma A_ncong_p: "x \<in> A ==> ~[x = 0](mod p)"
+  apply (auto simp add: A_def)
+  apply (frule_tac m = p in zcong_not_zero)
+  apply (insert p_minus_one_l)
+  apply auto
+  done
+
+lemma A_greater_zero: "x \<in> A ==> 0 < x"
+  by (auto simp add: A_def)
+
+lemma B_ncong_p: "x \<in> B ==> ~[x = 0](mod p)"
+  apply (auto simp add: B_def)
+  apply (frule A_ncong_p)
+  apply (insert p_a_relprime p_prime a_nonzero)
+  apply (frule_tac a = x and b = a in zcong_zprime_prod_zero_contra)
+  apply (auto simp add: A_greater_zero)
+  done
+
+lemma B_greater_zero: "x \<in> B ==> 0 < x"
+  using a_nonzero by (auto simp add: B_def mult_pos_pos A_greater_zero)
+
+lemma C_ncong_p: "x \<in> C ==>  ~[x = 0](mod p)"
+  apply (auto simp add: C_def)
+  apply (frule B_ncong_p)
+  apply (subgoal_tac "[x = StandardRes p x](mod p)")
+  defer apply (simp add: StandardRes_prop1)
+  apply (frule_tac a = x and b = "StandardRes p x" and c = 0 in zcong_trans)
+  apply auto
+  done
+
+lemma C_greater_zero: "y \<in> C ==> 0 < y"
+  apply (auto simp add: C_def)
+proof -
+  fix x
+  assume a: "x \<in> B"
+  from p_g_0 have "0 \<le> StandardRes p x"
+    by (simp add: StandardRes_lbound)
+  moreover have "~[x = 0] (mod p)"
+    by (simp add: a B_ncong_p)
+  then have "StandardRes p x \<noteq> 0"
+    by (simp add: StandardRes_prop3)
+  ultimately show "0 < StandardRes p x"
+    by (simp add: order_le_less)
+qed
+
+lemma D_ncong_p: "x \<in> D ==> ~[x = 0](mod p)"
+  by (auto simp add: D_def C_ncong_p)
+
+lemma E_ncong_p: "x \<in> E ==> ~[x = 0](mod p)"
+  by (auto simp add: E_def C_ncong_p)
+
+lemma F_ncong_p: "x \<in> F ==> ~[x = 0](mod p)"
+  apply (auto simp add: F_def)
+proof -
+  fix x assume a: "x \<in> E" assume b: "[p - x = 0] (mod p)"
+  from E_ncong_p have "~[x = 0] (mod p)"
+    by (simp add: a)
+  moreover from a have "0 < x"
+    by (simp add: a E_def C_greater_zero)
+  moreover from a have "x < p"
+    by (auto simp add: E_def C_def p_g_0 StandardRes_ubound)
+  ultimately have "~[p - x = 0] (mod p)"
+    by (simp add: zcong_not_zero)
+  from this show False by (simp add: b)
+qed
+
+lemma F_subset: "F \<subseteq> {x. 0 < x & x \<le> ((p - 1) div 2)}"
+  apply (auto simp add: F_def E_def)
+  apply (insert p_g_0)
+  apply (frule_tac x = xa in StandardRes_ubound)
+  apply (frule_tac x = x in StandardRes_ubound)
+  apply (subgoal_tac "xa = StandardRes p xa")
+  apply (auto simp add: C_def StandardRes_prop2 StandardRes_prop1)
+proof -
+  from zodd_imp_zdiv_eq p_prime p_g_2 zprime_zOdd_eq_grt_2 have
+    "2 * (p - 1) div 2 = 2 * ((p - 1) div 2)"
+    by simp
+  with p_eq2 show " !!x. [| (p - 1) div 2 < StandardRes p x; x \<in> B |]
+      ==> p - StandardRes p x \<le> (p - 1) div 2"
+    by simp
+qed
+
+lemma D_subset: "D \<subseteq> {x. 0 < x & x \<le> ((p - 1) div 2)}"
+  by (auto simp add: D_def C_greater_zero)
+
+lemma F_eq: "F = {x. \<exists>y \<in> A. ( x = p - (StandardRes p (y*a)) & (p - 1) div 2 < StandardRes p (y*a))}"
+  by (auto simp add: F_def E_def D_def C_def B_def A_def)
+
+lemma D_eq: "D = {x. \<exists>y \<in> A. ( x = StandardRes p (y*a) & StandardRes p (y*a) \<le> (p - 1) div 2)}"
+  by (auto simp add: D_def C_def B_def A_def)
+
+lemma D_leq: "x \<in> D ==> x \<le> (p - 1) div 2"
+  by (auto simp add: D_eq)
+
+lemma F_ge: "x \<in> F ==> x \<le> (p - 1) div 2"
+  apply (auto simp add: F_eq A_def)
+proof -
+  fix y
+  assume "(p - 1) div 2 < StandardRes p (y * a)"
+  then have "p - StandardRes p (y * a) < p - ((p - 1) div 2)"
+    by arith
+  also from p_eq2 have "... = 2 * ((p - 1) div 2) + 1 - ((p - 1) div 2)"
+    by auto
+  also have "2 * ((p - 1) div 2) + 1 - (p - 1) div 2 = (p - 1) div 2 + 1"
+    by arith
+  finally show "p - StandardRes p (y * a) \<le> (p - 1) div 2"
+    using zless_add1_eq [of "p - StandardRes p (y * a)" "(p - 1) div 2"] by auto
+qed
+
+lemma all_A_relprime: "\<forall>x \<in> A. zgcd x p = 1"
+  using p_prime p_minus_one_l by (auto simp add: A_def zless_zprime_imp_zrelprime)
+
+lemma A_prod_relprime: "zgcd (setprod id A) p = 1"
+by(rule all_relprime_prod_relprime[OF finite_A all_A_relprime])
+
+
+subsection {* Relationships Between Gauss Sets *}
+
+lemma B_card_eq_A: "card B = card A"
+  using finite_A by (simp add: finite_A B_def inj_on_xa_A card_image)
+
+lemma B_card_eq: "card B = nat ((p - 1) div 2)"
+  by (simp add: B_card_eq_A A_card_eq)
+
+lemma F_card_eq_E: "card F = card E"
+  using finite_E by (simp add: F_def inj_on_pminusx_E card_image)
+
+lemma C_card_eq_B: "card C = card B"
+  apply (insert finite_B)
+  apply (subgoal_tac "inj_on (StandardRes p) B")
+  apply (simp add: B_def C_def card_image)
+  apply (rule StandardRes_inj_on_ResSet)
+  apply (simp add: B_res)
+  done
+
+lemma D_E_disj: "D \<inter> E = {}"
+  by (auto simp add: D_def E_def)
+
+lemma C_card_eq_D_plus_E: "card C = card D + card E"
+  by (auto simp add: C_eq card_Un_disjoint D_E_disj finite_D finite_E)
+
+lemma C_prod_eq_D_times_E: "setprod id E * setprod id D = setprod id C"
+  apply (insert D_E_disj finite_D finite_E C_eq)
+  apply (frule setprod_Un_disjoint [of D E id])
+  apply auto
+  done
+
+lemma C_B_zcong_prod: "[setprod id C = setprod id B] (mod p)"
+  apply (auto simp add: C_def)
+  apply (insert finite_B SR_B_inj)
+  apply (frule_tac f = "StandardRes p" in setprod_reindex_id [symmetric], auto)
+  apply (rule setprod_same_function_zcong)
+  apply (auto simp add: StandardRes_prop1 zcong_sym p_g_0)
+  done
+
+lemma F_Un_D_subset: "(F \<union> D) \<subseteq> A"
+  apply (rule Un_least)
+  apply (auto simp add: A_def F_subset D_subset)
+  done
+
+lemma F_D_disj: "(F \<inter> D) = {}"
+  apply (simp add: F_eq D_eq)
+  apply (auto simp add: F_eq D_eq)
+proof -
+  fix y fix ya
+  assume "p - StandardRes p (y * a) = StandardRes p (ya * a)"
+  then have "p = StandardRes p (y * a) + StandardRes p (ya * a)"
+    by arith
+  moreover have "p dvd p"
+    by auto
+  ultimately have "p dvd (StandardRes p (y * a) + StandardRes p (ya * a))"
+    by auto
+  then have a: "[StandardRes p (y * a) + StandardRes p (ya * a) = 0] (mod p)"
+    by (auto simp add: zcong_def)
+  have "[y * a = StandardRes p (y * a)] (mod p)"
+    by (simp only: zcong_sym StandardRes_prop1)
+  moreover have "[ya * a = StandardRes p (ya * a)] (mod p)"
+    by (simp only: zcong_sym StandardRes_prop1)
+  ultimately have "[y * a + ya * a =
+    StandardRes p (y * a) + StandardRes p (ya * a)] (mod p)"
+    by (rule zcong_zadd)
+  with a have "[y * a + ya * a = 0] (mod p)"
+    apply (elim zcong_trans)
+    by (simp only: zcong_refl)
+  also have "y * a + ya * a = a * (y + ya)"
+    by (simp add: zadd_zmult_distrib2 zmult_commute)
+  finally have "[a * (y + ya) = 0] (mod p)" .
+  with p_prime a_nonzero zcong_zprime_prod_zero [of p a "y + ya"]
+    p_a_relprime
+  have a: "[y + ya = 0] (mod p)"
+    by auto
+  assume b: "y \<in> A" and c: "ya: A"
+  with A_def have "0 < y + ya"
+    by auto
+  moreover from b c A_def have "y + ya \<le> (p - 1) div 2 + (p - 1) div 2"
+    by auto
+  moreover from b c p_eq2 A_def have "y + ya < p"
+    by auto
+  ultimately show False
+    apply simp
+    apply (frule_tac m = p in zcong_not_zero)
+    apply (auto simp add: a)
+    done
+qed
+
+lemma F_Un_D_card: "card (F \<union> D) = nat ((p - 1) div 2)"
+proof -
+  have "card (F \<union> D) = card E + card D"
+    by (auto simp add: finite_F finite_D F_D_disj
+      card_Un_disjoint F_card_eq_E)
+  then have "card (F \<union> D) = card C"
+    by (simp add: C_card_eq_D_plus_E)
+  from this show "card (F \<union> D) = nat ((p - 1) div 2)"
+    by (simp add: C_card_eq_B B_card_eq)
+qed
+
+lemma F_Un_D_eq_A: "F \<union> D = A"
+  using finite_A F_Un_D_subset A_card_eq F_Un_D_card by (auto simp add: card_seteq)
+
+lemma prod_D_F_eq_prod_A:
+    "(setprod id D) * (setprod id F) = setprod id A"
+  apply (insert F_D_disj finite_D finite_F)
+  apply (frule setprod_Un_disjoint [of F D id])
+  apply (auto simp add: F_Un_D_eq_A)
+  done
+
+lemma prod_F_zcong:
+  "[setprod id F = ((-1) ^ (card E)) * (setprod id E)] (mod p)"
+proof -
+  have "setprod id F = setprod id (op - p ` E)"
+    by (auto simp add: F_def)
+  then have "setprod id F = setprod (op - p) E"
+    apply simp
+    apply (insert finite_E inj_on_pminusx_E)
+    apply (frule_tac f = "op - p" in setprod_reindex_id, auto)
+    done
+  then have one:
+    "[setprod id F = setprod (StandardRes p o (op - p)) E] (mod p)"
+    apply simp
+    apply (insert p_g_0 finite_E StandardRes_prod)
+    by (auto)
+  moreover have a: "\<forall>x \<in> E. [p - x = 0 - x] (mod p)"
+    apply clarify
+    apply (insert zcong_id [of p])
+    apply (rule_tac a = p and m = p and c = x and d = x in zcong_zdiff, auto)
+    done
+  moreover have b: "\<forall>x \<in> E. [StandardRes p (p - x) = p - x](mod p)"
+    apply clarify
+    apply (simp add: StandardRes_prop1 zcong_sym)
+    done
+  moreover have "\<forall>x \<in> E. [StandardRes p (p - x) = - x](mod p)"
+    apply clarify
+    apply (insert a b)
+    apply (rule_tac b = "p - x" in zcong_trans, auto)
+    done
+  ultimately have c:
+    "[setprod (StandardRes p o (op - p)) E = setprod (uminus) E](mod p)"
+    apply simp
+    using finite_E p_g_0
+      setprod_same_function_zcong [of E "StandardRes p o (op - p)" uminus p]
+    by auto
+  then have two: "[setprod id F = setprod (uminus) E](mod p)"
+    apply (insert one c)
+    apply (rule zcong_trans [of "setprod id F"
+                               "setprod (StandardRes p o op - p) E" p
+                               "setprod uminus E"], auto)
+    done
+  also have "setprod uminus E = (setprod id E) * (-1)^(card E)"
+    using finite_E by (induct set: finite) auto
+  then have "setprod uminus E = (-1) ^ (card E) * (setprod id E)"
+    by (simp add: zmult_commute)
+  with two show ?thesis
+    by simp
+qed
+
+
+subsection {* Gauss' Lemma *}
+
+lemma aux: "setprod id A * -1 ^ card E * a ^ card A * -1 ^ card E = setprod id A * a ^ card A"
+  by (auto simp add: finite_E neg_one_special)
+
+theorem pre_gauss_lemma:
+  "[a ^ nat((p - 1) div 2) = (-1) ^ (card E)] (mod p)"
+proof -
+  have "[setprod id A = setprod id F * setprod id D](mod p)"
+    by (auto simp add: prod_D_F_eq_prod_A zmult_commute cong del:setprod_cong)
+  then have "[setprod id A = ((-1)^(card E) * setprod id E) *
+      setprod id D] (mod p)"
+    apply (rule zcong_trans)
+    apply (auto simp add: prod_F_zcong zcong_scalar cong del: setprod_cong)
+    done
+  then have "[setprod id A = ((-1)^(card E) * setprod id C)] (mod p)"
+    apply (rule zcong_trans)
+    apply (insert C_prod_eq_D_times_E, erule subst)
+    apply (subst zmult_assoc, auto)
+    done
+  then have "[setprod id A = ((-1)^(card E) * setprod id B)] (mod p)"
+    apply (rule zcong_trans)
+    apply (simp add: C_B_zcong_prod zcong_scalar2 cong del:setprod_cong)
+    done
+  then have "[setprod id A = ((-1)^(card E) *
+    (setprod id ((%x. x * a) ` A)))] (mod p)"
+    by (simp add: B_def)
+  then have "[setprod id A = ((-1)^(card E) * (setprod (%x. x * a) A))]
+    (mod p)"
+    by (simp add:finite_A inj_on_xa_A setprod_reindex_id[symmetric] cong del:setprod_cong)
+  moreover have "setprod (%x. x * a) A =
+    setprod (%x. a) A * setprod id A"
+    using finite_A by (induct set: finite) auto
+  ultimately have "[setprod id A = ((-1)^(card E) * (setprod (%x. a) A *
+    setprod id A))] (mod p)"
+    by simp
+  then have "[setprod id A = ((-1)^(card E) * a^(card A) *
+      setprod id A)](mod p)"
+    apply (rule zcong_trans)
+    apply (simp add: zcong_scalar2 zcong_scalar finite_A setprod_constant zmult_assoc)
+    done
+  then have a: "[setprod id A * (-1)^(card E) =
+      ((-1)^(card E) * a^(card A) * setprod id A * (-1)^(card E))](mod p)"
+    by (rule zcong_scalar)
+  then have "[setprod id A * (-1)^(card E) = setprod id A *
+      (-1)^(card E) * a^(card A) * (-1)^(card E)](mod p)"
+    apply (rule zcong_trans)
+    apply (simp add: a mult_commute mult_left_commute)
+    done
+  then have "[setprod id A * (-1)^(card E) = setprod id A *
+      a^(card A)](mod p)"
+    apply (rule zcong_trans)
+    apply (simp add: aux cong del:setprod_cong)
+    done
+  with this zcong_cancel2 [of p "setprod id A" "-1 ^ card E" "a ^ card A"]
+      p_g_0 A_prod_relprime have "[-1 ^ card E = a ^ card A](mod p)"
+    by (simp add: order_less_imp_le)
+  from this show ?thesis
+    by (simp add: A_card_eq zcong_sym)
+qed
+
+theorem gauss_lemma: "(Legendre a p) = (-1) ^ (card E)"
+proof -
+  from Euler_Criterion p_prime p_g_2 have
+      "[(Legendre a p) = a^(nat (((p) - 1) div 2))] (mod p)"
+    by auto
+  moreover note pre_gauss_lemma
+  ultimately have "[(Legendre a p) = (-1) ^ (card E)] (mod p)"
+    by (rule zcong_trans)
+  moreover from p_a_relprime have "(Legendre a p) = 1 | (Legendre a p) = (-1)"
+    by (auto simp add: Legendre_def)
+  moreover have "(-1::int) ^ (card E) = 1 | (-1::int) ^ (card E) = -1"
+    by (rule neg_one_power)
+  ultimately show ?thesis
+    by (auto simp add: p_g_2 one_not_neg_one_mod_m zcong_sym)
+qed
+
+end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/Int2.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,299 @@
+(*  Title:      HOL/Quadratic_Reciprocity/Gauss.thy
+    ID:         $Id$
+    Authors:    Jeremy Avigad, David Gray, and Adam Kramer
+*)
+
+header {*Integers: Divisibility and Congruences*}
+
+theory Int2
+imports Finite2 WilsonRuss
+begin
+
+definition
+  MultInv :: "int => int => int" where
+  "MultInv p x = x ^ nat (p - 2)"
+
+
+subsection {* Useful lemmas about dvd and powers *}
+
+lemma zpower_zdvd_prop1:
+  "0 < n \<Longrightarrow> p dvd y \<Longrightarrow> p dvd ((y::int) ^ n)"
+  by (induct n) (auto simp add: dvd_mult2 [of p y])
+
+lemma zdvd_bounds: "n dvd m ==> m \<le> (0::int) | n \<le> m"
+proof -
+  assume "n dvd m"
+  then have "~(0 < m & m < n)"
+    using zdvd_not_zless [of m n] by auto
+  then show ?thesis by auto
+qed
+
+lemma zprime_zdvd_zmult_better: "[| zprime p;  p dvd (m * n) |] ==>
+    (p dvd m) | (p dvd n)"
+  apply (cases "0 \<le> m")
+  apply (simp add: zprime_zdvd_zmult)
+  apply (insert zprime_zdvd_zmult [of "-m" p n])
+  apply auto
+  done
+
+lemma zpower_zdvd_prop2:
+    "zprime p \<Longrightarrow> p dvd ((y::int) ^ n) \<Longrightarrow> 0 < n \<Longrightarrow> p dvd y"
+  apply (induct n)
+   apply simp
+  apply (frule zprime_zdvd_zmult_better)
+   apply simp
+  apply (force simp del:dvd_mult)
+  done
+
+lemma div_prop1: "[| 0 < z; (x::int) < y * z |] ==> x div z < y"
+proof -
+  assume "0 < z" then have modth: "x mod z \<ge> 0" by simp
+  have "(x div z) * z \<le> (x div z) * z" by simp
+  then have "(x div z) * z \<le> (x div z) * z + x mod z" using modth by arith 
+  also have "\<dots> = x"
+    by (auto simp add: zmod_zdiv_equality [symmetric] zmult_ac)
+  also assume  "x < y * z"
+  finally show ?thesis
+    by (auto simp add: prems mult_less_cancel_right, insert prems, arith)
+qed
+
+lemma div_prop2: "[| 0 < z; (x::int) < (y * z) + z |] ==> x div z \<le> y"
+proof -
+  assume "0 < z" and "x < (y * z) + z"
+  then have "x < (y + 1) * z" by (auto simp add: int_distrib)
+  then have "x div z < y + 1"
+    apply -
+    apply (rule_tac y = "y + 1" in div_prop1)
+    apply (auto simp add: prems)
+    done
+  then show ?thesis by auto
+qed
+
+lemma zdiv_leq_prop: "[| 0 < y |] ==> y * (x div y) \<le> (x::int)"
+proof-
+  assume "0 < y"
+  from zmod_zdiv_equality have "x = y * (x div y) + x mod y" by auto
+  moreover have "0 \<le> x mod y"
+    by (auto simp add: prems pos_mod_sign)
+  ultimately show ?thesis
+    by arith
+qed
+
+
+subsection {* Useful properties of congruences *}
+
+lemma zcong_eq_zdvd_prop: "[x = 0](mod p) = (p dvd x)"
+  by (auto simp add: zcong_def)
+
+lemma zcong_id: "[m = 0] (mod m)"
+  by (auto simp add: zcong_def)
+
+lemma zcong_shift: "[a = b] (mod m) ==> [a + c = b + c] (mod m)"
+  by (auto simp add: zcong_refl zcong_zadd)
+
+lemma zcong_zpower: "[x = y](mod m) ==> [x^z = y^z](mod m)"
+  by (induct z) (auto simp add: zcong_zmult)
+
+lemma zcong_eq_trans: "[| [a = b](mod m); b = c; [c = d](mod m) |] ==>
+    [a = d](mod m)"
+  apply (erule zcong_trans)
+  apply simp
+  done
+
+lemma aux1: "a - b = (c::int) ==> a = c + b"
+  by auto
+
+lemma zcong_zmult_prop1: "[a = b](mod m) ==> ([c = a * d](mod m) =
+    [c = b * d] (mod m))"
+  apply (auto simp add: zcong_def dvd_def)
+  apply (rule_tac x = "ka + k * d" in exI)
+  apply (drule aux1)+
+  apply (auto simp add: int_distrib)
+  apply (rule_tac x = "ka - k * d" in exI)
+  apply (drule aux1)+
+  apply (auto simp add: int_distrib)
+  done
+
+lemma zcong_zmult_prop2: "[a = b](mod m) ==>
+    ([c = d * a](mod m) = [c = d * b] (mod m))"
+  by (auto simp add: zmult_ac zcong_zmult_prop1)
+
+lemma zcong_zmult_prop3: "[| zprime p; ~[x = 0] (mod p);
+    ~[y = 0] (mod p) |] ==> ~[x * y = 0] (mod p)"
+  apply (auto simp add: zcong_def)
+  apply (drule zprime_zdvd_zmult_better, auto)
+  done
+
+lemma zcong_less_eq: "[| 0 < x; 0 < y; 0 < m; [x = y] (mod m);
+    x < m; y < m |] ==> x = y"
+  by (metis zcong_not zcong_sym zless_linear)
+
+lemma zcong_neg_1_impl_ne_1: "[| 2 < p; [x = -1] (mod p) |] ==>
+    ~([x = 1] (mod p))"
+proof
+  assume "2 < p" and "[x = 1] (mod p)" and "[x = -1] (mod p)"
+  then have "[1 = -1] (mod p)"
+    apply (auto simp add: zcong_sym)
+    apply (drule zcong_trans, auto)
+    done
+  then have "[1 + 1 = -1 + 1] (mod p)"
+    by (simp only: zcong_shift)
+  then have "[2 = 0] (mod p)"
+    by auto
+  then have "p dvd 2"
+    by (auto simp add: dvd_def zcong_def)
+  with prems show False
+    by (auto simp add: zdvd_not_zless)
+qed
+
+lemma zcong_zero_equiv_div: "[a = 0] (mod m) = (m dvd a)"
+  by (auto simp add: zcong_def)
+
+lemma zcong_zprime_prod_zero: "[| zprime p; 0 < a |] ==>
+    [a * b = 0] (mod p) ==> [a = 0] (mod p) | [b = 0] (mod p)"
+  by (auto simp add: zcong_zero_equiv_div zprime_zdvd_zmult)
+
+lemma zcong_zprime_prod_zero_contra: "[| zprime p; 0 < a |] ==>
+  ~[a = 0](mod p) & ~[b = 0](mod p) ==> ~[a * b = 0] (mod p)"
+  apply auto
+  apply (frule_tac a = a and b = b and p = p in zcong_zprime_prod_zero)
+  apply auto
+  done
+
+lemma zcong_not_zero: "[| 0 < x; x < m |] ==> ~[x = 0] (mod m)"
+  by (auto simp add: zcong_zero_equiv_div zdvd_not_zless)
+
+lemma zcong_zero: "[| 0 \<le> x; x < m; [x = 0](mod m) |] ==> x = 0"
+  apply (drule order_le_imp_less_or_eq, auto)
+  apply (frule_tac m = m in zcong_not_zero)
+  apply auto
+  done
+
+lemma all_relprime_prod_relprime: "[| finite A; \<forall>x \<in> A. zgcd x y = 1 |]
+    ==> zgcd (setprod id A) y = 1"
+  by (induct set: finite) (auto simp add: zgcd_zgcd_zmult)
+
+
+subsection {* Some properties of MultInv *}
+
+lemma MultInv_prop1: "[| 2 < p; [x = y] (mod p) |] ==>
+    [(MultInv p x) = (MultInv p y)] (mod p)"
+  by (auto simp add: MultInv_def zcong_zpower)
+
+lemma MultInv_prop2: "[| 2 < p; zprime p; ~([x = 0](mod p)) |] ==>
+  [(x * (MultInv p x)) = 1] (mod p)"
+proof (simp add: MultInv_def zcong_eq_zdvd_prop)
+  assume "2 < p" and "zprime p" and "~ p dvd x"
+  have "x * x ^ nat (p - 2) = x ^ (nat (p - 2) + 1)"
+    by auto
+  also from prems have "nat (p - 2) + 1 = nat (p - 2 + 1)"
+    by (simp only: nat_add_distrib)
+  also have "p - 2 + 1 = p - 1" by arith
+  finally have "[x * x ^ nat (p - 2) = x ^ nat (p - 1)] (mod p)"
+    by (rule ssubst, auto)
+  also from prems have "[x ^ nat (p - 1) = 1] (mod p)"
+    by (auto simp add: Little_Fermat)
+  finally (zcong_trans) show "[x * x ^ nat (p - 2) = 1] (mod p)" .
+qed
+
+lemma MultInv_prop2a: "[| 2 < p; zprime p; ~([x = 0](mod p)) |] ==>
+    [(MultInv p x) * x = 1] (mod p)"
+  by (auto simp add: MultInv_prop2 zmult_ac)
+
+lemma aux_1: "2 < p ==> ((nat p) - 2) = (nat (p - 2))"
+  by (simp add: nat_diff_distrib)
+
+lemma aux_2: "2 < p ==> 0 < nat (p - 2)"
+  by auto
+
+lemma MultInv_prop3: "[| 2 < p; zprime p; ~([x = 0](mod p)) |] ==>
+    ~([MultInv p x = 0](mod p))"
+  apply (auto simp add: MultInv_def zcong_eq_zdvd_prop aux_1)
+  apply (drule aux_2)
+  apply (drule zpower_zdvd_prop2, auto)
+  done
+
+lemma aux__1: "[| 2 < p; zprime p; ~([x = 0](mod p))|] ==>
+    [(MultInv p (MultInv p x)) = (x * (MultInv p x) *
+      (MultInv p (MultInv p x)))] (mod p)"
+  apply (drule MultInv_prop2, auto)
+  apply (drule_tac k = "MultInv p (MultInv p x)" in zcong_scalar, auto)
+  apply (auto simp add: zcong_sym)
+  done
+
+lemma aux__2: "[| 2 < p; zprime p; ~([x = 0](mod p))|] ==>
+    [(x * (MultInv p x) * (MultInv p (MultInv p x))) = x] (mod p)"
+  apply (frule MultInv_prop3, auto)
+  apply (insert MultInv_prop2 [of p "MultInv p x"], auto)
+  apply (drule MultInv_prop2, auto)
+  apply (drule_tac k = x in zcong_scalar2, auto)
+  apply (auto simp add: zmult_ac)
+  done
+
+lemma MultInv_prop4: "[| 2 < p; zprime p; ~([x = 0](mod p)) |] ==>
+    [(MultInv p (MultInv p x)) = x] (mod p)"
+  apply (frule aux__1, auto)
+  apply (drule aux__2, auto)
+  apply (drule zcong_trans, auto)
+  done
+
+lemma MultInv_prop5: "[| 2 < p; zprime p; ~([x = 0](mod p));
+    ~([y = 0](mod p)); [(MultInv p x) = (MultInv p y)] (mod p) |] ==>
+    [x = y] (mod p)"
+  apply (drule_tac a = "MultInv p x" and b = "MultInv p y" and
+    m = p and k = x in zcong_scalar)
+  apply (insert MultInv_prop2 [of p x], simp)
+  apply (auto simp only: zcong_sym [of "MultInv p x * x"])
+  apply (auto simp add:  zmult_ac)
+  apply (drule zcong_trans, auto)
+  apply (drule_tac a = "x * MultInv p y" and k = y in zcong_scalar, auto)
+  apply (insert MultInv_prop2a [of p y], auto simp add: zmult_ac)
+  apply (insert zcong_zmult_prop2 [of "y * MultInv p y" 1 p y x])
+  apply (auto simp add: zcong_sym)
+  done
+
+lemma MultInv_zcong_prop1: "[| 2 < p; [j = k] (mod p) |] ==>
+    [a * MultInv p j = a * MultInv p k] (mod p)"
+  by (drule MultInv_prop1, auto simp add: zcong_scalar2)
+
+lemma aux___1: "[j = a * MultInv p k] (mod p) ==>
+    [j * k = a * MultInv p k * k] (mod p)"
+  by (auto simp add: zcong_scalar)
+
+lemma aux___2: "[|2 < p; zprime p; ~([k = 0](mod p));
+    [j * k = a * MultInv p k * k] (mod p) |] ==> [j * k = a] (mod p)"
+  apply (insert MultInv_prop2a [of p k] zcong_zmult_prop2
+    [of "MultInv p k * k" 1 p "j * k" a])
+  apply (auto simp add: zmult_ac)
+  done
+
+lemma aux___3: "[j * k = a] (mod p) ==> [(MultInv p j) * j * k =
+     (MultInv p j) * a] (mod p)"
+  by (auto simp add: zmult_assoc zcong_scalar2)
+
+lemma aux___4: "[|2 < p; zprime p; ~([j = 0](mod p));
+    [(MultInv p j) * j * k = (MultInv p j) * a] (mod p) |]
+       ==> [k = a * (MultInv p j)] (mod p)"
+  apply (insert MultInv_prop2a [of p j] zcong_zmult_prop1
+    [of "MultInv p j * j" 1 p "MultInv p j * a" k])
+  apply (auto simp add: zmult_ac zcong_sym)
+  done
+
+lemma MultInv_zcong_prop2: "[| 2 < p; zprime p; ~([k = 0](mod p));
+    ~([j = 0](mod p)); [j = a * MultInv p k] (mod p) |] ==>
+    [k = a * MultInv p j] (mod p)"
+  apply (drule aux___1)
+  apply (frule aux___2, auto)
+  by (drule aux___3, drule aux___4, auto)
+
+lemma MultInv_zcong_prop3: "[| 2 < p; zprime p; ~([a = 0](mod p));
+    ~([k = 0](mod p)); ~([j = 0](mod p));
+    [a * MultInv p j = a * MultInv p k] (mod p) |] ==>
+      [j = k] (mod p)"
+  apply (auto simp add: zcong_eq_zdvd_prop [of a p])
+  apply (frule zprime_imp_zrelprime, auto)
+  apply (insert zcong_cancel2 [of p a "MultInv p j" "MultInv p k"], auto)
+  apply (drule MultInv_prop5, auto)
+  done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/IntFact.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,94 @@
+(*  Author:     Thomas M. Rasmussen
+    Copyright   2000  University of Cambridge
+*)
+
+header {* Factorial on integers *}
+
+theory IntFact imports IntPrimes begin
+
+text {*
+  Factorial on integers and recursively defined set including all
+  Integers from @{text 2} up to @{text a}.  Plus definition of product
+  of finite set.
+
+  \bigskip
+*}
+
+consts
+  zfact :: "int => int"
+  d22set :: "int => int set"
+
+recdef zfact  "measure ((\<lambda>n. nat n) :: int => nat)"
+  "zfact n = (if n \<le> 0 then 1 else n * zfact (n - 1))"
+
+recdef d22set  "measure ((\<lambda>a. nat a) :: int => nat)"
+  "d22set a = (if 1 < a then insert a (d22set (a - 1)) else {})"
+
+
+text {*
+  \medskip @{term d22set} --- recursively defined set including all
+  integers from @{text 2} up to @{text a}
+*}
+
+declare d22set.simps [simp del]
+
+
+lemma d22set_induct:
+  assumes "!!a. P {} a"
+    and "!!a. 1 < (a::int) ==> P (d22set (a - 1)) (a - 1) ==> P (d22set a) a"
+  shows "P (d22set u) u"
+  apply (rule d22set.induct)
+  apply safe
+   prefer 2
+   apply (case_tac "1 < a")
+    apply (rule_tac prems)
+     apply (simp_all (no_asm_simp))
+   apply (simp_all (no_asm_simp) add: d22set.simps prems)
+  done
+
+lemma d22set_g_1 [rule_format]: "b \<in> d22set a --> 1 < b"
+  apply (induct a rule: d22set_induct)
+   apply simp
+  apply (subst d22set.simps)
+  apply auto
+  done
+
+lemma d22set_le [rule_format]: "b \<in> d22set a --> b \<le> a"
+  apply (induct a rule: d22set_induct)
+  apply simp
+   apply (subst d22set.simps)
+   apply auto
+  done
+
+lemma d22set_le_swap: "a < b ==> b \<notin> d22set a"
+  by (auto dest: d22set_le)
+
+lemma d22set_mem: "1 < b \<Longrightarrow> b \<le> a \<Longrightarrow> b \<in> d22set a"
+  apply (induct a rule: d22set.induct)
+  apply auto
+   apply (simp_all add: d22set.simps)
+  done
+
+lemma d22set_fin: "finite (d22set a)"
+  apply (induct a rule: d22set_induct)
+   prefer 2
+   apply (subst d22set.simps)
+   apply auto
+  done
+
+
+declare zfact.simps [simp del]
+
+lemma d22set_prod_zfact: "\<Prod>(d22set a) = zfact a"
+  apply (induct a rule: d22set.induct)
+  apply safe
+   apply (simp add: d22set.simps zfact.simps)
+  apply (subst d22set.simps)
+  apply (subst zfact.simps)
+  apply (case_tac "1 < a")
+   prefer 2
+   apply (simp add: d22set.simps zfact.simps)
+  apply (simp add: d22set_fin d22set_le_swap)
+  done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/IntPrimes.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,421 @@
+(*  Author:     Thomas M. Rasmussen
+    Copyright   2000  University of Cambridge
+*)
+
+header {* Divisibility and prime numbers (on integers) *}
+
+theory IntPrimes
+imports Main Primes
+begin
+
+text {*
+  The @{text dvd} relation, GCD, Euclid's extended algorithm, primes,
+  congruences (all on the Integers).  Comparable to theory @{text
+  Primes}, but @{text dvd} is included here as it is not present in
+  main HOL.  Also includes extended GCD and congruences not present in
+  @{text Primes}.
+*}
+
+
+subsection {* Definitions *}
+
+consts
+  xzgcda :: "int * int * int * int * int * int * int * int => int * int * int"
+
+recdef xzgcda
+  "measure ((\<lambda>(m, n, r', r, s', s, t', t). nat r)
+    :: int * int * int * int *int * int * int * int => nat)"
+  "xzgcda (m, n, r', r, s', s, t', t) =
+	(if r \<le> 0 then (r', s', t')
+	 else xzgcda (m, n, r, r' mod r, 
+		      s, s' - (r' div r) * s, 
+		      t, t' - (r' div r) * t))"
+
+definition
+  zprime :: "int \<Rightarrow> bool" where
+  "zprime p = (1 < p \<and> (\<forall>m. 0 <= m & m dvd p --> m = 1 \<or> m = p))"
+
+definition
+  xzgcd :: "int => int => int * int * int" where
+  "xzgcd m n = xzgcda (m, n, m, n, 1, 0, 0, 1)"
+
+definition
+  zcong :: "int => int => int => bool"  ("(1[_ = _] '(mod _'))") where
+  "[a = b] (mod m) = (m dvd (a - b))"
+
+subsection {* Euclid's Algorithm and GCD *}
+
+
+lemma zrelprime_zdvd_zmult_aux:
+     "zgcd n k = 1 ==> k dvd m * n ==> 0 \<le> m ==> k dvd m"
+    by (metis abs_of_nonneg dvd_triv_right zgcd_greatest_iff zgcd_zmult_distrib2_abs zmult_1_right)
+
+lemma zrelprime_zdvd_zmult: "zgcd n k = 1 ==> k dvd m * n ==> k dvd m"
+  apply (case_tac "0 \<le> m")
+   apply (blast intro: zrelprime_zdvd_zmult_aux)
+  apply (subgoal_tac "k dvd -m")
+   apply (rule_tac [2] zrelprime_zdvd_zmult_aux, auto)
+  done
+
+lemma zgcd_geq_zero: "0 <= zgcd x y"
+  by (auto simp add: zgcd_def)
+
+text{*This is merely a sanity check on zprime, since the previous version
+      denoted the empty set.*}
+lemma "zprime 2"
+  apply (auto simp add: zprime_def) 
+  apply (frule zdvd_imp_le, simp) 
+  apply (auto simp add: order_le_less dvd_def) 
+  done
+
+lemma zprime_imp_zrelprime:
+    "zprime p ==> \<not> p dvd n ==> zgcd n p = 1"
+  apply (auto simp add: zprime_def)
+  apply (metis zgcd_geq_zero zgcd_zdvd1 zgcd_zdvd2)
+  done
+
+lemma zless_zprime_imp_zrelprime:
+    "zprime p ==> 0 < n ==> n < p ==> zgcd n p = 1"
+  apply (erule zprime_imp_zrelprime)
+  apply (erule zdvd_not_zless, assumption)
+  done
+
+lemma zprime_zdvd_zmult:
+    "0 \<le> (m::int) ==> zprime p ==> p dvd m * n ==> p dvd m \<or> p dvd n"
+  by (metis zgcd_zdvd1 zgcd_zdvd2 zgcd_pos zprime_def zrelprime_dvd_mult)
+
+lemma zgcd_zadd_zmult [simp]: "zgcd (m + n * k) n = zgcd m n"
+  apply (rule zgcd_eq [THEN trans])
+  apply (simp add: mod_add_eq)
+  apply (rule zgcd_eq [symmetric])
+  done
+
+lemma zgcd_zdvd_zgcd_zmult: "zgcd m n dvd zgcd (k * m) n"
+by (simp add: zgcd_greatest_iff)
+
+lemma zgcd_zmult_zdvd_zgcd:
+    "zgcd k n = 1 ==> zgcd (k * m) n dvd zgcd m n"
+  apply (simp add: zgcd_greatest_iff)
+  apply (rule_tac n = k in zrelprime_zdvd_zmult)
+   prefer 2
+   apply (simp add: zmult_commute)
+  apply (metis zgcd_1 zgcd_commute zgcd_left_commute)
+  done
+
+lemma zgcd_zmult_cancel: "zgcd k n = 1 ==> zgcd (k * m) n = zgcd m n"
+  by (simp add: zgcd_def nat_abs_mult_distrib gcd_mult_cancel)
+
+lemma zgcd_zgcd_zmult:
+    "zgcd k m = 1 ==> zgcd n m = 1 ==> zgcd (k * n) m = 1"
+  by (simp add: zgcd_zmult_cancel)
+
+lemma zdvd_iff_zgcd: "0 < m ==> m dvd n \<longleftrightarrow> zgcd n m = m"
+  by (metis abs_of_pos zdvd_mult_div_cancel zgcd_0 zgcd_commute zgcd_geq_zero zgcd_zdvd2 zgcd_zmult_eq_self)
+
+
+
+subsection {* Congruences *}
+
+lemma zcong_1 [simp]: "[a = b] (mod 1)"
+  by (unfold zcong_def, auto)
+
+lemma zcong_refl [simp]: "[k = k] (mod m)"
+  by (unfold zcong_def, auto)
+
+lemma zcong_sym: "[a = b] (mod m) = [b = a] (mod m)"
+  unfolding zcong_def minus_diff_eq [of a, symmetric] dvd_minus_iff ..
+
+lemma zcong_zadd:
+    "[a = b] (mod m) ==> [c = d] (mod m) ==> [a + c = b + d] (mod m)"
+  apply (unfold zcong_def)
+  apply (rule_tac s = "(a - b) + (c - d)" in subst)
+   apply (rule_tac [2] dvd_add, auto)
+  done
+
+lemma zcong_zdiff:
+    "[a = b] (mod m) ==> [c = d] (mod m) ==> [a - c = b - d] (mod m)"
+  apply (unfold zcong_def)
+  apply (rule_tac s = "(a - b) - (c - d)" in subst)
+   apply (rule_tac [2] dvd_diff, auto)
+  done
+
+lemma zcong_trans:
+  "[a = b] (mod m) ==> [b = c] (mod m) ==> [a = c] (mod m)"
+unfolding zcong_def by (auto elim!: dvdE simp add: algebra_simps)
+
+lemma zcong_zmult:
+    "[a = b] (mod m) ==> [c = d] (mod m) ==> [a * c = b * d] (mod m)"
+  apply (rule_tac b = "b * c" in zcong_trans)
+   apply (unfold zcong_def)
+  apply (metis zdiff_zmult_distrib2 dvd_mult zmult_commute)
+  apply (metis zdiff_zmult_distrib2 dvd_mult)
+  done
+
+lemma zcong_scalar: "[a = b] (mod m) ==> [a * k = b * k] (mod m)"
+  by (rule zcong_zmult, simp_all)
+
+lemma zcong_scalar2: "[a = b] (mod m) ==> [k * a = k * b] (mod m)"
+  by (rule zcong_zmult, simp_all)
+
+lemma zcong_zmult_self: "[a * m = b * m] (mod m)"
+  apply (unfold zcong_def)
+  apply (rule dvd_diff, simp_all)
+  done
+
+lemma zcong_square:
+   "[| zprime p;  0 < a;  [a * a = 1] (mod p)|]
+    ==> [a = 1] (mod p) \<or> [a = p - 1] (mod p)"
+  apply (unfold zcong_def)
+  apply (rule zprime_zdvd_zmult)
+    apply (rule_tac [3] s = "a * a - 1 + p * (1 - a)" in subst)
+     prefer 4
+     apply (simp add: zdvd_reduce)
+    apply (simp_all add: zdiff_zmult_distrib zmult_commute zdiff_zmult_distrib2)
+  done
+
+lemma zcong_cancel:
+  "0 \<le> m ==>
+    zgcd k m = 1 ==> [a * k = b * k] (mod m) = [a = b] (mod m)"
+  apply safe
+   prefer 2
+   apply (blast intro: zcong_scalar)
+  apply (case_tac "b < a")
+   prefer 2
+   apply (subst zcong_sym)
+   apply (unfold zcong_def)
+   apply (rule_tac [!] zrelprime_zdvd_zmult)
+     apply (simp_all add: zdiff_zmult_distrib)
+  apply (subgoal_tac "m dvd (-(a * k - b * k))")
+   apply simp
+  apply (subst dvd_minus_iff, assumption)
+  done
+
+lemma zcong_cancel2:
+  "0 \<le> m ==>
+    zgcd k m = 1 ==> [k * a = k * b] (mod m) = [a = b] (mod m)"
+  by (simp add: zmult_commute zcong_cancel)
+
+lemma zcong_zgcd_zmult_zmod:
+  "[a = b] (mod m) ==> [a = b] (mod n) ==> zgcd m n = 1
+    ==> [a = b] (mod m * n)"
+  apply (auto simp add: zcong_def dvd_def)
+  apply (subgoal_tac "m dvd n * ka")
+   apply (subgoal_tac "m dvd ka")
+    apply (case_tac [2] "0 \<le> ka")
+  apply (metis zdvd_mult_div_cancel dvd_refl dvd_mult_left zmult_commute zrelprime_zdvd_zmult)
+  apply (metis abs_dvd_iff abs_of_nonneg zadd_0 zgcd_0_left zgcd_commute zgcd_zadd_zmult zgcd_zdvd_zgcd_zmult zgcd_zmult_distrib2_abs zmult_1_right zmult_commute)
+  apply (metis mult_le_0_iff  zdvd_mono zdvd_mult_cancel dvd_triv_left zero_le_mult_iff zle_anti_sym zle_linear zle_refl zmult_commute zrelprime_zdvd_zmult)
+  apply (metis dvd_triv_left)
+  done
+
+lemma zcong_zless_imp_eq:
+  "0 \<le> a ==>
+    a < m ==> 0 \<le> b ==> b < m ==> [a = b] (mod m) ==> a = b"
+  apply (unfold zcong_def dvd_def, auto)
+  apply (drule_tac f = "\<lambda>z. z mod m" in arg_cong)
+  apply (metis diff_add_cancel mod_pos_pos_trivial zadd_0 zadd_commute zmod_eq_0_iff mod_add_right_eq)
+  done
+
+lemma zcong_square_zless:
+  "zprime p ==> 0 < a ==> a < p ==>
+    [a * a = 1] (mod p) ==> a = 1 \<or> a = p - 1"
+  apply (cut_tac p = p and a = a in zcong_square)
+     apply (simp add: zprime_def)
+    apply (auto intro: zcong_zless_imp_eq)
+  done
+
+lemma zcong_not:
+    "0 < a ==> a < m ==> 0 < b ==> b < a ==> \<not> [a = b] (mod m)"
+  apply (unfold zcong_def)
+  apply (rule zdvd_not_zless, auto)
+  done
+
+lemma zcong_zless_0:
+    "0 \<le> a ==> a < m ==> [a = 0] (mod m) ==> a = 0"
+  apply (unfold zcong_def dvd_def, auto)
+  apply (metis div_pos_pos_trivial linorder_not_less div_mult_self1_is_id)
+  done
+
+lemma zcong_zless_unique:
+    "0 < m ==> (\<exists>!b. 0 \<le> b \<and> b < m \<and> [a = b] (mod m))"
+  apply auto
+   prefer 2 apply (metis zcong_sym zcong_trans zcong_zless_imp_eq)
+  apply (unfold zcong_def dvd_def)
+  apply (rule_tac x = "a mod m" in exI, auto)
+  apply (metis zmult_div_cancel)
+  done
+
+lemma zcong_iff_lin: "([a = b] (mod m)) = (\<exists>k. b = a + m * k)"
+  unfolding zcong_def
+  apply (auto elim!: dvdE simp add: algebra_simps)
+  apply (rule_tac x = "-k" in exI) apply simp
+  done
+
+lemma zgcd_zcong_zgcd:
+  "0 < m ==>
+    zgcd a m = 1 ==> [a = b] (mod m) ==> zgcd b m = 1"
+  by (auto simp add: zcong_iff_lin)
+
+lemma zcong_zmod_aux:
+     "a - b = (m::int) * (a div m - b div m) + (a mod m - b mod m)"
+  by(simp add: zdiff_zmult_distrib2 add_diff_eq eq_diff_eq add_ac)
+
+lemma zcong_zmod: "[a = b] (mod m) = [a mod m = b mod m] (mod m)"
+  apply (unfold zcong_def)
+  apply (rule_tac t = "a - b" in ssubst)
+  apply (rule_tac m = m in zcong_zmod_aux)
+  apply (rule trans)
+   apply (rule_tac [2] k = m and m = "a div m - b div m" in zdvd_reduce)
+  apply (simp add: zadd_commute)
+  done
+
+lemma zcong_zmod_eq: "0 < m ==> [a = b] (mod m) = (a mod m = b mod m)"
+  apply auto
+  apply (metis pos_mod_conj zcong_zless_imp_eq zcong_zmod)
+  apply (metis zcong_refl zcong_zmod)
+  done
+
+lemma zcong_zminus [iff]: "[a = b] (mod -m) = [a = b] (mod m)"
+  by (auto simp add: zcong_def)
+
+lemma zcong_zero [iff]: "[a = b] (mod 0) = (a = b)"
+  by (auto simp add: zcong_def)
+
+lemma "[a = b] (mod m) = (a mod m = b mod m)"
+  apply (case_tac "m = 0", simp add: DIVISION_BY_ZERO)
+  apply (simp add: linorder_neq_iff)
+  apply (erule disjE)  
+   prefer 2 apply (simp add: zcong_zmod_eq)
+  txt{*Remainding case: @{term "m<0"}*}
+  apply (rule_tac t = m in zminus_zminus [THEN subst])
+  apply (subst zcong_zminus)
+  apply (subst zcong_zmod_eq, arith)
+  apply (frule neg_mod_bound [of _ a], frule neg_mod_bound [of _ b]) 
+  apply (simp add: zmod_zminus2_eq_if del: neg_mod_bound)
+  done
+
+subsection {* Modulo *}
+
+lemma zmod_zdvd_zmod:
+    "0 < (m::int) ==> m dvd b ==> (a mod b mod m) = (a mod m)"
+  by (rule mod_mod_cancel) 
+
+
+subsection {* Extended GCD *}
+
+declare xzgcda.simps [simp del]
+
+lemma xzgcd_correct_aux1:
+  "zgcd r' r = k --> 0 < r -->
+    (\<exists>sn tn. xzgcda (m, n, r', r, s', s, t', t) = (k, sn, tn))"
+  apply (rule_tac u = m and v = n and w = r' and x = r and y = s' and
+    z = s and aa = t' and ab = t in xzgcda.induct)
+  apply (subst zgcd_eq)
+  apply (subst xzgcda.simps, auto)
+  apply (case_tac "r' mod r = 0")
+   prefer 2
+   apply (frule_tac a = "r'" in pos_mod_sign, auto)
+  apply (rule exI)
+  apply (rule exI)
+  apply (subst xzgcda.simps, auto)
+  done
+
+lemma xzgcd_correct_aux2:
+  "(\<exists>sn tn. xzgcda (m, n, r', r, s', s, t', t) = (k, sn, tn)) --> 0 < r -->
+    zgcd r' r = k"
+  apply (rule_tac u = m and v = n and w = r' and x = r and y = s' and
+    z = s and aa = t' and ab = t in xzgcda.induct)
+  apply (subst zgcd_eq)
+  apply (subst xzgcda.simps)
+  apply (auto simp add: linorder_not_le)
+  apply (case_tac "r' mod r = 0")
+   prefer 2
+   apply (frule_tac a = "r'" in pos_mod_sign, auto)
+  apply (metis Pair_eq simps zle_refl)
+  done
+
+lemma xzgcd_correct:
+    "0 < n ==> (zgcd m n = k) = (\<exists>s t. xzgcd m n = (k, s, t))"
+  apply (unfold xzgcd_def)
+  apply (rule iffI)
+   apply (rule_tac [2] xzgcd_correct_aux2 [THEN mp, THEN mp])
+    apply (rule xzgcd_correct_aux1 [THEN mp, THEN mp], auto)
+  done
+
+
+text {* \medskip @{term xzgcd} linear *}
+
+lemma xzgcda_linear_aux1:
+  "(a - r * b) * m + (c - r * d) * (n::int) =
+   (a * m + c * n) - r * (b * m + d * n)"
+  by (simp add: zdiff_zmult_distrib zadd_zmult_distrib2 zmult_assoc)
+
+lemma xzgcda_linear_aux2:
+  "r' = s' * m + t' * n ==> r = s * m + t * n
+    ==> (r' mod r) = (s' - (r' div r) * s) * m + (t' - (r' div r) * t) * (n::int)"
+  apply (rule trans)
+   apply (rule_tac [2] xzgcda_linear_aux1 [symmetric])
+  apply (simp add: eq_diff_eq mult_commute)
+  done
+
+lemma order_le_neq_implies_less: "(x::'a::order) \<le> y ==> x \<noteq> y ==> x < y"
+  by (rule iffD2 [OF order_less_le conjI])
+
+lemma xzgcda_linear [rule_format]:
+  "0 < r --> xzgcda (m, n, r', r, s', s, t', t) = (rn, sn, tn) -->
+    r' = s' * m + t' * n -->  r = s * m + t * n --> rn = sn * m + tn * n"
+  apply (rule_tac u = m and v = n and w = r' and x = r and y = s' and
+    z = s and aa = t' and ab = t in xzgcda.induct)
+  apply (subst xzgcda.simps)
+  apply (simp (no_asm))
+  apply (rule impI)+
+  apply (case_tac "r' mod r = 0")
+   apply (simp add: xzgcda.simps, clarify)
+  apply (subgoal_tac "0 < r' mod r")
+   apply (rule_tac [2] order_le_neq_implies_less)
+   apply (rule_tac [2] pos_mod_sign)
+    apply (cut_tac m = m and n = n and r' = r' and r = r and s' = s' and
+      s = s and t' = t' and t = t in xzgcda_linear_aux2, auto)
+  done
+
+lemma xzgcd_linear:
+    "0 < n ==> xzgcd m n = (r, s, t) ==> r = s * m + t * n"
+  apply (unfold xzgcd_def)
+  apply (erule xzgcda_linear, assumption, auto)
+  done
+
+lemma zgcd_ex_linear:
+    "0 < n ==> zgcd m n = k ==> (\<exists>s t. k = s * m + t * n)"
+  apply (simp add: xzgcd_correct, safe)
+  apply (rule exI)+
+  apply (erule xzgcd_linear, auto)
+  done
+
+lemma zcong_lineq_ex:
+    "0 < n ==> zgcd a n = 1 ==> \<exists>x. [a * x = 1] (mod n)"
+  apply (cut_tac m = a and n = n and k = 1 in zgcd_ex_linear, safe)
+  apply (rule_tac x = s in exI)
+  apply (rule_tac b = "s * a + t * n" in zcong_trans)
+   prefer 2
+   apply simp
+  apply (unfold zcong_def)
+  apply (simp (no_asm) add: zmult_commute)
+  done
+
+lemma zcong_lineq_unique:
+  "0 < n ==>
+    zgcd a n = 1 ==> \<exists>!x. 0 \<le> x \<and> x < n \<and> [a * x = b] (mod n)"
+  apply auto
+   apply (rule_tac [2] zcong_zless_imp_eq)
+       apply (tactic {* stac (thm "zcong_cancel2" RS sym) 6 *})
+         apply (rule_tac [8] zcong_trans)
+          apply (simp_all (no_asm_simp))
+   prefer 2
+   apply (simp add: zcong_sym)
+  apply (cut_tac a = a and n = n in zcong_lineq_ex, auto)
+  apply (rule_tac x = "x * b mod n" in exI, safe)
+    apply (simp_all (no_asm_simp))
+  apply (metis zcong_scalar zcong_zmod zmod_zmult1_eq zmult_1 zmult_assoc)
+  done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/Legacy_GCD.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,787 @@
+(*  Title:      HOL/GCD.thy
+    Author:     Christophe Tabacznyj and Lawrence C Paulson
+    Copyright   1996  University of Cambridge
+*)
+
+header {* The Greatest Common Divisor *}
+
+theory Legacy_GCD
+imports Main
+begin
+
+text {*
+  See \cite{davenport92}. \bigskip
+*}
+
+subsection {* Specification of GCD on nats *}
+
+definition
+  is_gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where -- {* @{term gcd} as a relation *}
+  [code del]: "is_gcd m n p \<longleftrightarrow> p dvd m \<and> p dvd n \<and>
+    (\<forall>d. d dvd m \<longrightarrow> d dvd n \<longrightarrow> d dvd p)"
+
+text {* Uniqueness *}
+
+lemma is_gcd_unique: "is_gcd a b m \<Longrightarrow> is_gcd a b n \<Longrightarrow> m = n"
+  by (simp add: is_gcd_def) (blast intro: dvd_anti_sym)
+
+text {* Connection to divides relation *}
+
+lemma is_gcd_dvd: "is_gcd a b m \<Longrightarrow> k dvd a \<Longrightarrow> k dvd b \<Longrightarrow> k dvd m"
+  by (auto simp add: is_gcd_def)
+
+text {* Commutativity *}
+
+lemma is_gcd_commute: "is_gcd m n k = is_gcd n m k"
+  by (auto simp add: is_gcd_def)
+
+
+subsection {* GCD on nat by Euclid's algorithm *}
+
+fun
+  gcd  :: "nat => nat => nat"
+where
+  "gcd m n = (if n = 0 then m else gcd n (m mod n))"
+lemma gcd_induct [case_names "0" rec]:
+  fixes m n :: nat
+  assumes "\<And>m. P m 0"
+    and "\<And>m n. 0 < n \<Longrightarrow> P n (m mod n) \<Longrightarrow> P m n"
+  shows "P m n"
+proof (induct m n rule: gcd.induct)
+  case (1 m n) with assms show ?case by (cases "n = 0") simp_all
+qed
+
+lemma gcd_0 [simp, algebra]: "gcd m 0 = m"
+  by simp
+
+lemma gcd_0_left [simp,algebra]: "gcd 0 m = m"
+  by simp
+
+lemma gcd_non_0: "n > 0 \<Longrightarrow> gcd m n = gcd n (m mod n)"
+  by simp
+
+lemma gcd_1 [simp, algebra]: "gcd m (Suc 0) = Suc 0"
+  by simp
+
+lemma nat_gcd_1_right [simp, algebra]: "gcd m 1 = 1"
+  unfolding One_nat_def by (rule gcd_1)
+
+declare gcd.simps [simp del]
+
+text {*
+  \medskip @{term "gcd m n"} divides @{text m} and @{text n}.  The
+  conjunctions don't seem provable separately.
+*}
+
+lemma gcd_dvd1 [iff, algebra]: "gcd m n dvd m"
+  and gcd_dvd2 [iff, algebra]: "gcd m n dvd n"
+  apply (induct m n rule: gcd_induct)
+     apply (simp_all add: gcd_non_0)
+  apply (blast dest: dvd_mod_imp_dvd)
+  done
+
+text {*
+  \medskip Maximality: for all @{term m}, @{term n}, @{term k}
+  naturals, if @{term k} divides @{term m} and @{term k} divides
+  @{term n} then @{term k} divides @{term "gcd m n"}.
+*}
+
+lemma gcd_greatest: "k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd gcd m n"
+  by (induct m n rule: gcd_induct) (simp_all add: gcd_non_0 dvd_mod)
+
+text {*
+  \medskip Function gcd yields the Greatest Common Divisor.
+*}
+
+lemma is_gcd: "is_gcd m n (gcd m n) "
+  by (simp add: is_gcd_def gcd_greatest)
+
+
+subsection {* Derived laws for GCD *}
+
+lemma gcd_greatest_iff [iff, algebra]: "k dvd gcd m n \<longleftrightarrow> k dvd m \<and> k dvd n"
+  by (blast intro!: gcd_greatest intro: dvd_trans)
+
+lemma gcd_zero[algebra]: "gcd m n = 0 \<longleftrightarrow> m = 0 \<and> n = 0"
+  by (simp only: dvd_0_left_iff [symmetric] gcd_greatest_iff)
+
+lemma gcd_commute: "gcd m n = gcd n m"
+  apply (rule is_gcd_unique)
+   apply (rule is_gcd)
+  apply (subst is_gcd_commute)
+  apply (simp add: is_gcd)
+  done
+
+lemma gcd_assoc: "gcd (gcd k m) n = gcd k (gcd m n)"
+  apply (rule is_gcd_unique)
+   apply (rule is_gcd)
+  apply (simp add: is_gcd_def)
+  apply (blast intro: dvd_trans)
+  done
+
+lemma gcd_1_left [simp, algebra]: "gcd (Suc 0) m = Suc 0"
+  by (simp add: gcd_commute)
+
+lemma nat_gcd_1_left [simp, algebra]: "gcd 1 m = 1"
+  unfolding One_nat_def by (rule gcd_1_left)
+
+text {*
+  \medskip Multiplication laws
+*}
+
+lemma gcd_mult_distrib2: "k * gcd m n = gcd (k * m) (k * n)"
+    -- {* \cite[page 27]{davenport92} *}
+  apply (induct m n rule: gcd_induct)
+   apply simp
+  apply (case_tac "k = 0")
+   apply (simp_all add: mod_geq gcd_non_0 mod_mult_distrib2)
+  done
+
+lemma gcd_mult [simp, algebra]: "gcd k (k * n) = k"
+  apply (rule gcd_mult_distrib2 [of k 1 n, simplified, symmetric])
+  done
+
+lemma gcd_self [simp, algebra]: "gcd k k = k"
+  apply (rule gcd_mult [of k 1, simplified])
+  done
+
+lemma relprime_dvd_mult: "gcd k n = 1 ==> k dvd m * n ==> k dvd m"
+  apply (insert gcd_mult_distrib2 [of m k n])
+  apply simp
+  apply (erule_tac t = m in ssubst)
+  apply simp
+  done
+
+lemma relprime_dvd_mult_iff: "gcd k n = 1 ==> (k dvd m * n) = (k dvd m)"
+  by (auto intro: relprime_dvd_mult dvd_mult2)
+
+lemma gcd_mult_cancel: "gcd k n = 1 ==> gcd (k * m) n = gcd m n"
+  apply (rule dvd_anti_sym)
+   apply (rule gcd_greatest)
+    apply (rule_tac n = k in relprime_dvd_mult)
+     apply (simp add: gcd_assoc)
+     apply (simp add: gcd_commute)
+    apply (simp_all add: mult_commute)
+  done
+
+
+text {* \medskip Addition laws *}
+
+lemma gcd_add1 [simp, algebra]: "gcd (m + n) n = gcd m n"
+  by (cases "n = 0") (auto simp add: gcd_non_0)
+
+lemma gcd_add2 [simp, algebra]: "gcd m (m + n) = gcd m n"
+proof -
+  have "gcd m (m + n) = gcd (m + n) m" by (rule gcd_commute)
+  also have "... = gcd (n + m) m" by (simp add: add_commute)
+  also have "... = gcd n m" by simp
+  also have  "... = gcd m n" by (rule gcd_commute)
+  finally show ?thesis .
+qed
+
+lemma gcd_add2' [simp, algebra]: "gcd m (n + m) = gcd m n"
+  apply (subst add_commute)
+  apply (rule gcd_add2)
+  done
+
+lemma gcd_add_mult[algebra]: "gcd m (k * m + n) = gcd m n"
+  by (induct k) (simp_all add: add_assoc)
+
+lemma gcd_dvd_prod: "gcd m n dvd m * n" 
+  using mult_dvd_mono [of 1] by auto
+
+text {*
+  \medskip Division by gcd yields rrelatively primes.
+*}
+
+lemma div_gcd_relprime:
+  assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
+  shows "gcd (a div gcd a b) (b div gcd a b) = 1"
+proof -
+  let ?g = "gcd a b"
+  let ?a' = "a div ?g"
+  let ?b' = "b div ?g"
+  let ?g' = "gcd ?a' ?b'"
+  have dvdg: "?g dvd a" "?g dvd b" by simp_all
+  have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by simp_all
+  from dvdg dvdg' obtain ka kb ka' kb' where
+      kab: "a = ?g * ka" "b = ?g * kb" "?a' = ?g' * ka'" "?b' = ?g' * kb'"
+    unfolding dvd_def by blast
+  then have "?g * ?a' = (?g * ?g') * ka'" "?g * ?b' = (?g * ?g') * kb'" by simp_all
+  then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
+    by (auto simp add: dvd_mult_div_cancel [OF dvdg(1)]
+      dvd_mult_div_cancel [OF dvdg(2)] dvd_def)
+  have "?g \<noteq> 0" using nz by (simp add: gcd_zero)
+  then have gp: "?g > 0" by simp
+  from gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
+  with dvd_mult_cancel1 [OF gp] show "?g' = 1" by simp
+qed
+
+
+lemma gcd_unique: "d dvd a\<and>d dvd b \<and> (\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"
+proof(auto)
+  assume H: "d dvd a" "d dvd b" "\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d"
+  from H(3)[rule_format] gcd_dvd1[of a b] gcd_dvd2[of a b] 
+  have th: "gcd a b dvd d" by blast
+  from dvd_anti_sym[OF th gcd_greatest[OF H(1,2)]]  show "d = gcd a b" by blast 
+qed
+
+lemma gcd_eq: assumes H: "\<forall>d. d dvd x \<and> d dvd y \<longleftrightarrow> d dvd u \<and> d dvd v"
+  shows "gcd x y = gcd u v"
+proof-
+  from H have "\<forall>d. d dvd x \<and> d dvd y \<longleftrightarrow> d dvd gcd u v" by simp
+  with gcd_unique[of "gcd u v" x y]  show ?thesis by auto
+qed
+
+lemma ind_euclid: 
+  assumes c: " \<forall>a b. P (a::nat) b \<longleftrightarrow> P b a" and z: "\<forall>a. P a 0" 
+  and add: "\<forall>a b. P a b \<longrightarrow> P a (a + b)" 
+  shows "P a b"
+proof(induct n\<equiv>"a+b" arbitrary: a b rule: nat_less_induct)
+  fix n a b
+  assume H: "\<forall>m < n. \<forall>a b. m = a + b \<longrightarrow> P a b" "n = a + b"
+  have "a = b \<or> a < b \<or> b < a" by arith
+  moreover {assume eq: "a= b"
+    from add[rule_format, OF z[rule_format, of a]] have "P a b" using eq by simp}
+  moreover
+  {assume lt: "a < b"
+    hence "a + b - a < n \<or> a = 0"  using H(2) by arith
+    moreover
+    {assume "a =0" with z c have "P a b" by blast }
+    moreover
+    {assume ab: "a + b - a < n"
+      have th0: "a + b - a = a + (b - a)" using lt by arith
+      from add[rule_format, OF H(1)[rule_format, OF ab th0]]
+      have "P a b" by (simp add: th0[symmetric])}
+    ultimately have "P a b" by blast}
+  moreover
+  {assume lt: "a > b"
+    hence "b + a - b < n \<or> b = 0"  using H(2) by arith
+    moreover
+    {assume "b =0" with z c have "P a b" by blast }
+    moreover
+    {assume ab: "b + a - b < n"
+      have th0: "b + a - b = b + (a - b)" using lt by arith
+      from add[rule_format, OF H(1)[rule_format, OF ab th0]]
+      have "P b a" by (simp add: th0[symmetric])
+      hence "P a b" using c by blast }
+    ultimately have "P a b" by blast}
+ultimately  show "P a b" by blast
+qed
+
+lemma bezout_lemma: 
+  assumes ex: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x = b * y + d \<or> b * x = a * y + d)"
+  shows "\<exists>d x y. d dvd a \<and> d dvd a + b \<and> (a * x = (a + b) * y + d \<or> (a + b) * x = a * y + d)"
+using ex
+apply clarsimp
+apply (rule_tac x="d" in exI, simp add: dvd_add)
+apply (case_tac "a * x = b * y + d" , simp_all)
+apply (rule_tac x="x + y" in exI)
+apply (rule_tac x="y" in exI)
+apply algebra
+apply (rule_tac x="x" in exI)
+apply (rule_tac x="x + y" in exI)
+apply algebra
+done
+
+lemma bezout_add: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x = b * y + d \<or> b * x = a * y + d)"
+apply(induct a b rule: ind_euclid)
+apply blast
+apply clarify
+apply (rule_tac x="a" in exI, simp add: dvd_add)
+apply clarsimp
+apply (rule_tac x="d" in exI)
+apply (case_tac "a * x = b * y + d", simp_all add: dvd_add)
+apply (rule_tac x="x+y" in exI)
+apply (rule_tac x="y" in exI)
+apply algebra
+apply (rule_tac x="x" in exI)
+apply (rule_tac x="x+y" in exI)
+apply algebra
+done
+
+lemma bezout: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x - b * y = d \<or> b * x - a * y = d)"
+using bezout_add[of a b]
+apply clarsimp
+apply (rule_tac x="d" in exI, simp)
+apply (rule_tac x="x" in exI)
+apply (rule_tac x="y" in exI)
+apply auto
+done
+
+
+text {* We can get a stronger version with a nonzeroness assumption. *}
+lemma divides_le: "m dvd n ==> m <= n \<or> n = (0::nat)" by (auto simp add: dvd_def)
+
+lemma bezout_add_strong: assumes nz: "a \<noteq> (0::nat)"
+  shows "\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d"
+proof-
+  from nz have ap: "a > 0" by simp
+ from bezout_add[of a b] 
+ have "(\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d) \<or> (\<exists>d x y. d dvd a \<and> d dvd b \<and> b * x = a * y + d)" by blast
+ moreover
+ {fix d x y assume H: "d dvd a" "d dvd b" "a * x = b * y + d"
+   from H have ?thesis by blast }
+ moreover
+ {fix d x y assume H: "d dvd a" "d dvd b" "b * x = a * y + d"
+   {assume b0: "b = 0" with H  have ?thesis by simp}
+   moreover 
+   {assume b: "b \<noteq> 0" hence bp: "b > 0" by simp
+     from divides_le[OF H(2)] b have "d < b \<or> d = b" using le_less by blast
+     moreover
+     {assume db: "d=b"
+       from prems have ?thesis apply simp
+	 apply (rule exI[where x = b], simp)
+	 apply (rule exI[where x = b])
+	by (rule exI[where x = "a - 1"], simp add: diff_mult_distrib2)}
+    moreover
+    {assume db: "d < b" 
+	{assume "x=0" hence ?thesis  using prems by simp }
+	moreover
+	{assume x0: "x \<noteq> 0" hence xp: "x > 0" by simp
+	  
+	  from db have "d \<le> b - 1" by simp
+	  hence "d*b \<le> b*(b - 1)" by simp
+	  with xp mult_mono[of "1" "x" "d*b" "b*(b - 1)"]
+	  have dble: "d*b \<le> x*b*(b - 1)" using bp by simp
+	  from H (3) have "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x*b*(b - 1)" by algebra
+	  hence "a * ((b - 1) * y) = d + x*b*(b - 1) - d*b" using bp by simp
+	  hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)" 
+	    by (simp only: diff_add_assoc[OF dble, of d, symmetric])
+	  hence "a * ((b - 1) * y) = b*(x*(b - 1) - d) + d"
+	    by (simp only: diff_mult_distrib2 add_commute mult_ac)
+	  hence ?thesis using H(1,2)
+	    apply -
+	    apply (rule exI[where x=d], simp)
+	    apply (rule exI[where x="(b - 1) * y"])
+	    by (rule exI[where x="x*(b - 1) - d"], simp)}
+	ultimately have ?thesis by blast}
+    ultimately have ?thesis by blast}
+  ultimately have ?thesis by blast}
+ ultimately show ?thesis by blast
+qed
+
+
+lemma bezout_gcd: "\<exists>x y. a * x - b * y = gcd a b \<or> b * x - a * y = gcd a b"
+proof-
+  let ?g = "gcd a b"
+  from bezout[of a b] obtain d x y where d: "d dvd a" "d dvd b" "a * x - b * y = d \<or> b * x - a * y = d" by blast
+  from d(1,2) have "d dvd ?g" by simp
+  then obtain k where k: "?g = d*k" unfolding dvd_def by blast
+  from d(3) have "(a * x - b * y)*k = d*k \<or> (b * x - a * y)*k = d*k" by blast 
+  hence "a * x * k - b * y*k = d*k \<or> b * x * k - a * y*k = d*k" 
+    by (algebra add: diff_mult_distrib)
+  hence "a * (x * k) - b * (y*k) = ?g \<or> b * (x * k) - a * (y*k) = ?g" 
+    by (simp add: k mult_assoc)
+  thus ?thesis by blast
+qed
+
+lemma bezout_gcd_strong: assumes a: "a \<noteq> 0" 
+  shows "\<exists>x y. a * x = b * y + gcd a b"
+proof-
+  let ?g = "gcd a b"
+  from bezout_add_strong[OF a, of b]
+  obtain d x y where d: "d dvd a" "d dvd b" "a * x = b * y + d" by blast
+  from d(1,2) have "d dvd ?g" by simp
+  then obtain k where k: "?g = d*k" unfolding dvd_def by blast
+  from d(3) have "a * x * k = (b * y + d) *k " by algebra
+  hence "a * (x * k) = b * (y*k) + ?g" by (algebra add: k)
+  thus ?thesis by blast
+qed
+
+lemma gcd_mult_distrib: "gcd(a * c) (b * c) = c * gcd a b"
+by(simp add: gcd_mult_distrib2 mult_commute)
+
+lemma gcd_bezout: "(\<exists>x y. a * x - b * y = d \<or> b * x - a * y = d) \<longleftrightarrow> gcd a b dvd d"
+  (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  let ?g = "gcd a b"
+  {assume H: ?rhs then obtain k where k: "d = ?g*k" unfolding dvd_def by blast
+    from bezout_gcd[of a b] obtain x y where xy: "a * x - b * y = ?g \<or> b * x - a * y = ?g"
+      by blast
+    hence "(a * x - b * y)*k = ?g*k \<or> (b * x - a * y)*k = ?g*k" by auto
+    hence "a * x*k - b * y*k = ?g*k \<or> b * x * k - a * y*k = ?g*k" 
+      by (simp only: diff_mult_distrib)
+    hence "a * (x*k) - b * (y*k) = d \<or> b * (x * k) - a * (y*k) = d"
+      by (simp add: k[symmetric] mult_assoc)
+    hence ?lhs by blast}
+  moreover
+  {fix x y assume H: "a * x - b * y = d \<or> b * x - a * y = d"
+    have dv: "?g dvd a*x" "?g dvd b * y" "?g dvd b*x" "?g dvd a * y"
+      using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all
+    from dvd_diff_nat[OF dv(1,2)] dvd_diff_nat[OF dv(3,4)] H
+    have ?rhs by auto}
+  ultimately show ?thesis by blast
+qed
+
+lemma gcd_bezout_sum: assumes H:"a * x + b * y = d" shows "gcd a b dvd d"
+proof-
+  let ?g = "gcd a b"
+    have dv: "?g dvd a*x" "?g dvd b * y" 
+      using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all
+    from dvd_add[OF dv] H
+    show ?thesis by auto
+qed
+
+lemma gcd_mult': "gcd b (a * b) = b"
+by (simp add: gcd_mult mult_commute[of a b]) 
+
+lemma gcd_add: "gcd(a + b) b = gcd a b" 
+  "gcd(b + a) b = gcd a b" "gcd a (a + b) = gcd a b" "gcd a (b + a) = gcd a b"
+apply (simp_all add: gcd_add1)
+by (simp add: gcd_commute gcd_add1)
+
+lemma gcd_sub: "b <= a ==> gcd(a - b) b = gcd a b" "a <= b ==> gcd a (b - a) = gcd a b"
+proof-
+  {fix a b assume H: "b \<le> (a::nat)"
+    hence th: "a - b + b = a" by arith
+    from gcd_add(1)[of "a - b" b] th  have "gcd(a - b) b = gcd a b" by simp}
+  note th = this
+{
+  assume ab: "b \<le> a"
+  from th[OF ab] show "gcd (a - b)  b = gcd a b" by blast
+next
+  assume ab: "a \<le> b"
+  from th[OF ab] show "gcd a (b - a) = gcd a b" 
+    by (simp add: gcd_commute)}
+qed
+
+
+subsection {* LCM defined by GCD *}
+
+
+definition
+  lcm :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+where
+  lcm_def: "lcm m n = m * n div gcd m n"
+
+lemma prod_gcd_lcm:
+  "m * n = gcd m n * lcm m n"
+  unfolding lcm_def by (simp add: dvd_mult_div_cancel [OF gcd_dvd_prod])
+
+lemma lcm_0 [simp]: "lcm m 0 = 0"
+  unfolding lcm_def by simp
+
+lemma lcm_1 [simp]: "lcm m 1 = m"
+  unfolding lcm_def by simp
+
+lemma lcm_0_left [simp]: "lcm 0 n = 0"
+  unfolding lcm_def by simp
+
+lemma lcm_1_left [simp]: "lcm 1 m = m"
+  unfolding lcm_def by simp
+
+lemma dvd_pos:
+  fixes n m :: nat
+  assumes "n > 0" and "m dvd n"
+  shows "m > 0"
+using assms by (cases m) auto
+
+lemma lcm_least:
+  assumes "m dvd k" and "n dvd k"
+  shows "lcm m n dvd k"
+proof (cases k)
+  case 0 then show ?thesis by auto
+next
+  case (Suc _) then have pos_k: "k > 0" by auto
+  from assms dvd_pos [OF this] have pos_mn: "m > 0" "n > 0" by auto
+  with gcd_zero [of m n] have pos_gcd: "gcd m n > 0" by simp
+  from assms obtain p where k_m: "k = m * p" using dvd_def by blast
+  from assms obtain q where k_n: "k = n * q" using dvd_def by blast
+  from pos_k k_m have pos_p: "p > 0" by auto
+  from pos_k k_n have pos_q: "q > 0" by auto
+  have "k * k * gcd q p = k * gcd (k * q) (k * p)"
+    by (simp add: mult_ac gcd_mult_distrib2)
+  also have "\<dots> = k * gcd (m * p * q) (n * q * p)"
+    by (simp add: k_m [symmetric] k_n [symmetric])
+  also have "\<dots> = k * p * q * gcd m n"
+    by (simp add: mult_ac gcd_mult_distrib2)
+  finally have "(m * p) * (n * q) * gcd q p = k * p * q * gcd m n"
+    by (simp only: k_m [symmetric] k_n [symmetric])
+  then have "p * q * m * n * gcd q p = p * q * k * gcd m n"
+    by (simp add: mult_ac)
+  with pos_p pos_q have "m * n * gcd q p = k * gcd m n"
+    by simp
+  with prod_gcd_lcm [of m n]
+  have "lcm m n * gcd q p * gcd m n = k * gcd m n"
+    by (simp add: mult_ac)
+  with pos_gcd have "lcm m n * gcd q p = k" by simp
+  then show ?thesis using dvd_def by auto
+qed
+
+lemma lcm_dvd1 [iff]:
+  "m dvd lcm m n"
+proof (cases m)
+  case 0 then show ?thesis by simp
+next
+  case (Suc _)
+  then have mpos: "m > 0" by simp
+  show ?thesis
+  proof (cases n)
+    case 0 then show ?thesis by simp
+  next
+    case (Suc _)
+    then have npos: "n > 0" by simp
+    have "gcd m n dvd n" by simp
+    then obtain k where "n = gcd m n * k" using dvd_def by auto
+    then have "m * n div gcd m n = m * (gcd m n * k) div gcd m n" by (simp add: mult_ac)
+    also have "\<dots> = m * k" using mpos npos gcd_zero by simp
+    finally show ?thesis by (simp add: lcm_def)
+  qed
+qed
+
+lemma lcm_dvd2 [iff]: 
+  "n dvd lcm m n"
+proof (cases n)
+  case 0 then show ?thesis by simp
+next
+  case (Suc _)
+  then have npos: "n > 0" by simp
+  show ?thesis
+  proof (cases m)
+    case 0 then show ?thesis by simp
+  next
+    case (Suc _)
+    then have mpos: "m > 0" by simp
+    have "gcd m n dvd m" by simp
+    then obtain k where "m = gcd m n * k" using dvd_def by auto
+    then have "m * n div gcd m n = (gcd m n * k) * n div gcd m n" by (simp add: mult_ac)
+    also have "\<dots> = n * k" using mpos npos gcd_zero by simp
+    finally show ?thesis by (simp add: lcm_def)
+  qed
+qed
+
+lemma gcd_add1_eq: "gcd (m + k) k = gcd (m + k) m"
+  by (simp add: gcd_commute)
+
+lemma gcd_diff2: "m \<le> n ==> gcd n (n - m) = gcd n m"
+  apply (subgoal_tac "n = m + (n - m)")
+  apply (erule ssubst, rule gcd_add1_eq, simp)  
+  done
+
+
+subsection {* GCD and LCM on integers *}
+
+definition
+  zgcd :: "int \<Rightarrow> int \<Rightarrow> int" where
+  "zgcd i j = int (gcd (nat (abs i)) (nat (abs j)))"
+
+lemma zgcd_zdvd1 [iff,simp, algebra]: "zgcd i j dvd i"
+by (simp add: zgcd_def int_dvd_iff)
+
+lemma zgcd_zdvd2 [iff,simp, algebra]: "zgcd i j dvd j"
+by (simp add: zgcd_def int_dvd_iff)
+
+lemma zgcd_pos: "zgcd i j \<ge> 0"
+by (simp add: zgcd_def)
+
+lemma zgcd0 [simp,algebra]: "(zgcd i j = 0) = (i = 0 \<and> j = 0)"
+by (simp add: zgcd_def gcd_zero)
+
+lemma zgcd_commute: "zgcd i j = zgcd j i"
+unfolding zgcd_def by (simp add: gcd_commute)
+
+lemma zgcd_zminus [simp, algebra]: "zgcd (- i) j = zgcd i j"
+unfolding zgcd_def by simp
+
+lemma zgcd_zminus2 [simp, algebra]: "zgcd i (- j) = zgcd i j"
+unfolding zgcd_def by simp
+
+  (* should be solved by algebra*)
+lemma zrelprime_dvd_mult: "zgcd i j = 1 \<Longrightarrow> i dvd k * j \<Longrightarrow> i dvd k"
+  unfolding zgcd_def
+proof -
+  assume "int (gcd (nat \<bar>i\<bar>) (nat \<bar>j\<bar>)) = 1" "i dvd k * j"
+  then have g: "gcd (nat \<bar>i\<bar>) (nat \<bar>j\<bar>) = 1" by simp
+  from `i dvd k * j` obtain h where h: "k*j = i*h" unfolding dvd_def by blast
+  have th: "nat \<bar>i\<bar> dvd nat \<bar>k\<bar> * nat \<bar>j\<bar>"
+    unfolding dvd_def
+    by (rule_tac x= "nat \<bar>h\<bar>" in exI, simp add: h nat_abs_mult_distrib [symmetric])
+  from relprime_dvd_mult [OF g th] obtain h' where h': "nat \<bar>k\<bar> = nat \<bar>i\<bar> * h'"
+    unfolding dvd_def by blast
+  from h' have "int (nat \<bar>k\<bar>) = int (nat \<bar>i\<bar> * h')" by simp
+  then have "\<bar>k\<bar> = \<bar>i\<bar> * int h'" by (simp add: int_mult)
+  then show ?thesis
+    apply (subst abs_dvd_iff [symmetric])
+    apply (subst dvd_abs_iff [symmetric])
+    apply (unfold dvd_def)
+    apply (rule_tac x = "int h'" in exI, simp)
+    done
+qed
+
+lemma int_nat_abs: "int (nat (abs x)) = abs x" by arith
+
+lemma zgcd_greatest:
+  assumes "k dvd m" and "k dvd n"
+  shows "k dvd zgcd m n"
+proof -
+  let ?k' = "nat \<bar>k\<bar>"
+  let ?m' = "nat \<bar>m\<bar>"
+  let ?n' = "nat \<bar>n\<bar>"
+  from `k dvd m` and `k dvd n` have dvd': "?k' dvd ?m'" "?k' dvd ?n'"
+    unfolding zdvd_int by (simp_all only: int_nat_abs abs_dvd_iff dvd_abs_iff)
+  from gcd_greatest [OF dvd'] have "int (nat \<bar>k\<bar>) dvd zgcd m n"
+    unfolding zgcd_def by (simp only: zdvd_int)
+  then have "\<bar>k\<bar> dvd zgcd m n" by (simp only: int_nat_abs)
+  then show "k dvd zgcd m n" by simp
+qed
+
+lemma div_zgcd_relprime:
+  assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
+  shows "zgcd (a div (zgcd a b)) (b div (zgcd a b)) = 1"
+proof -
+  from nz have nz': "nat \<bar>a\<bar> \<noteq> 0 \<or> nat \<bar>b\<bar> \<noteq> 0" by arith 
+  let ?g = "zgcd a b"
+  let ?a' = "a div ?g"
+  let ?b' = "b div ?g"
+  let ?g' = "zgcd ?a' ?b'"
+  have dvdg: "?g dvd a" "?g dvd b" by (simp_all add: zgcd_zdvd1 zgcd_zdvd2)
+  have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by (simp_all add: zgcd_zdvd1 zgcd_zdvd2)
+  from dvdg dvdg' obtain ka kb ka' kb' where
+   kab: "a = ?g*ka" "b = ?g*kb" "?a' = ?g'*ka'" "?b' = ?g' * kb'"
+    unfolding dvd_def by blast
+  then have "?g* ?a' = (?g * ?g') * ka'" "?g* ?b' = (?g * ?g') * kb'" by simp_all
+  then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b"
+    by (auto simp add: zdvd_mult_div_cancel [OF dvdg(1)]
+      zdvd_mult_div_cancel [OF dvdg(2)] dvd_def)
+  have "?g \<noteq> 0" using nz by simp
+  then have gp: "?g \<noteq> 0" using zgcd_pos[where i="a" and j="b"] by arith
+  from zgcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" .
+  with zdvd_mult_cancel1 [OF gp] have "\<bar>?g'\<bar> = 1" by simp
+  with zgcd_pos show "?g' = 1" by simp
+qed
+
+lemma zgcd_0 [simp, algebra]: "zgcd m 0 = abs m"
+  by (simp add: zgcd_def abs_if)
+
+lemma zgcd_0_left [simp, algebra]: "zgcd 0 m = abs m"
+  by (simp add: zgcd_def abs_if)
+
+lemma zgcd_non_0: "0 < n ==> zgcd m n = zgcd n (m mod n)"
+  apply (frule_tac b = n and a = m in pos_mod_sign)
+  apply (simp del: pos_mod_sign add: zgcd_def abs_if nat_mod_distrib)
+  apply (auto simp add: gcd_non_0 nat_mod_distrib [symmetric] zmod_zminus1_eq_if)
+  apply (frule_tac a = m in pos_mod_bound)
+  apply (simp del: pos_mod_bound add: nat_diff_distrib gcd_diff2 nat_le_eq_zle)
+  done
+
+lemma zgcd_eq: "zgcd m n = zgcd n (m mod n)"
+  apply (case_tac "n = 0", simp add: DIVISION_BY_ZERO)
+  apply (auto simp add: linorder_neq_iff zgcd_non_0)
+  apply (cut_tac m = "-m" and n = "-n" in zgcd_non_0, auto)
+  done
+
+lemma zgcd_1 [simp, algebra]: "zgcd m 1 = 1"
+  by (simp add: zgcd_def abs_if)
+
+lemma zgcd_0_1_iff [simp, algebra]: "zgcd 0 m = 1 \<longleftrightarrow> \<bar>m\<bar> = 1"
+  by (simp add: zgcd_def abs_if)
+
+lemma zgcd_greatest_iff[algebra]: "k dvd zgcd m n = (k dvd m \<and> k dvd n)"
+  by (simp add: zgcd_def abs_if int_dvd_iff dvd_int_iff nat_dvd_iff)
+
+lemma zgcd_1_left [simp, algebra]: "zgcd 1 m = 1"
+  by (simp add: zgcd_def gcd_1_left)
+
+lemma zgcd_assoc: "zgcd (zgcd k m) n = zgcd k (zgcd m n)"
+  by (simp add: zgcd_def gcd_assoc)
+
+lemma zgcd_left_commute: "zgcd k (zgcd m n) = zgcd m (zgcd k n)"
+  apply (rule zgcd_commute [THEN trans])
+  apply (rule zgcd_assoc [THEN trans])
+  apply (rule zgcd_commute [THEN arg_cong])
+  done
+
+lemmas zgcd_ac = zgcd_assoc zgcd_commute zgcd_left_commute
+  -- {* addition is an AC-operator *}
+
+lemma zgcd_zmult_distrib2: "0 \<le> k ==> k * zgcd m n = zgcd (k * m) (k * n)"
+  by (simp del: minus_mult_right [symmetric]
+      add: minus_mult_right nat_mult_distrib zgcd_def abs_if
+          mult_less_0_iff gcd_mult_distrib2 [symmetric] zmult_int [symmetric])
+
+lemma zgcd_zmult_distrib2_abs: "zgcd (k * m) (k * n) = abs k * zgcd m n"
+  by (simp add: abs_if zgcd_zmult_distrib2)
+
+lemma zgcd_self [simp]: "0 \<le> m ==> zgcd m m = m"
+  by (cut_tac k = m and m = 1 and n = 1 in zgcd_zmult_distrib2, simp_all)
+
+lemma zgcd_zmult_eq_self [simp]: "0 \<le> k ==> zgcd k (k * n) = k"
+  by (cut_tac k = k and m = 1 and n = n in zgcd_zmult_distrib2, simp_all)
+
+lemma zgcd_zmult_eq_self2 [simp]: "0 \<le> k ==> zgcd (k * n) k = k"
+  by (cut_tac k = k and m = n and n = 1 in zgcd_zmult_distrib2, simp_all)
+
+
+definition "zlcm i j = int (lcm(nat(abs i)) (nat(abs j)))"
+
+lemma dvd_zlcm_self1[simp, algebra]: "i dvd zlcm i j"
+by(simp add:zlcm_def dvd_int_iff)
+
+lemma dvd_zlcm_self2[simp, algebra]: "j dvd zlcm i j"
+by(simp add:zlcm_def dvd_int_iff)
+
+
+lemma dvd_imp_dvd_zlcm1:
+  assumes "k dvd i" shows "k dvd (zlcm i j)"
+proof -
+  have "nat(abs k) dvd nat(abs i)" using `k dvd i`
+    by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric])
+  thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans)
+qed
+
+lemma dvd_imp_dvd_zlcm2:
+  assumes "k dvd j" shows "k dvd (zlcm i j)"
+proof -
+  have "nat(abs k) dvd nat(abs j)" using `k dvd j`
+    by(simp add:int_dvd_iff[symmetric] dvd_int_iff[symmetric])
+  thus ?thesis by(simp add:zlcm_def dvd_int_iff)(blast intro: dvd_trans)
+qed
+
+
+lemma zdvd_self_abs1: "(d::int) dvd (abs d)"
+by (case_tac "d <0", simp_all)
+
+lemma zdvd_self_abs2: "(abs (d::int)) dvd d"
+by (case_tac "d<0", simp_all)
+
+(* lcm a b is positive for positive a and b *)
+
+lemma lcm_pos: 
+  assumes mpos: "m > 0"
+  and npos: "n>0"
+  shows "lcm m n > 0"
+proof(rule ccontr, simp add: lcm_def gcd_zero)
+assume h:"m*n div gcd m n = 0"
+from mpos npos have "gcd m n \<noteq> 0" using gcd_zero by simp
+hence gcdp: "gcd m n > 0" by simp
+with h
+have "m*n < gcd m n"
+  by (cases "m * n < gcd m n") (auto simp add: div_if[OF gcdp, where m="m*n"])
+moreover 
+have "gcd m n dvd m" by simp
+ with mpos dvd_imp_le have t1:"gcd m n \<le> m" by simp
+ with npos have t1:"gcd m n *n \<le> m*n" by simp
+ have "gcd m n \<le> gcd m n*n" using npos by simp
+ with t1 have "gcd m n \<le> m*n" by arith
+ultimately show "False" by simp
+qed
+
+lemma zlcm_pos: 
+  assumes anz: "a \<noteq> 0"
+  and bnz: "b \<noteq> 0" 
+  shows "0 < zlcm a b"
+proof-
+  let ?na = "nat (abs a)"
+  let ?nb = "nat (abs b)"
+  have nap: "?na >0" using anz by simp
+  have nbp: "?nb >0" using bnz by simp
+  have "0 < lcm ?na ?nb" by (rule lcm_pos[OF nap nbp])
+  thus ?thesis by (simp add: zlcm_def)
+qed
+
+lemma zgcd_code [code]:
+  "zgcd k l = \<bar>if l = 0 then k else zgcd l (\<bar>k\<bar> mod \<bar>l\<bar>)\<bar>"
+  by (simp add: zgcd_def gcd.simps [of "nat \<bar>k\<bar>"] nat_mod_distrib)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/Pocklington.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,1263 @@
+(*  Title:      HOL/Library/Pocklington.thy
+    Author:     Amine Chaieb
+*)
+
+header {* Pocklington's Theorem for Primes *}
+
+theory Pocklington
+imports Main Primes
+begin
+
+definition modeq:: "nat => nat => nat => bool"    ("(1[_ = _] '(mod _'))")
+  where "[a = b] (mod p) == ((a mod p) = (b mod p))"
+
+definition modneq:: "nat => nat => nat => bool"    ("(1[_ \<noteq> _] '(mod _'))")
+  where "[a \<noteq> b] (mod p) == ((a mod p) \<noteq> (b mod p))"
+
+lemma modeq_trans:
+  "\<lbrakk> [a = b] (mod p); [b = c] (mod p) \<rbrakk> \<Longrightarrow> [a = c] (mod p)"
+  by (simp add:modeq_def)
+
+
+lemma nat_mod_lemma: assumes xyn: "[x = y] (mod n)" and xy:"y \<le> x"
+  shows "\<exists>q. x = y + n * q"
+using xyn xy unfolding modeq_def using nat_mod_eq_lemma by blast
+
+lemma nat_mod[algebra]: "[x = y] (mod n) \<longleftrightarrow> (\<exists>q1 q2. x + n * q1 = y + n * q2)"
+unfolding modeq_def nat_mod_eq_iff ..
+
+(* Lemmas about previously defined terms.                                    *)
+
+lemma prime: "prime p \<longleftrightarrow> p \<noteq> 0 \<and> p\<noteq>1 \<and> (\<forall>m. 0 < m \<and> m < p \<longrightarrow> coprime p m)"
+  (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  {assume "p=0 \<or> p=1" hence ?thesis using prime_0 prime_1 by (cases "p=0", simp_all)}
+  moreover
+  {assume p0: "p\<noteq>0" "p\<noteq>1"
+    {assume H: "?lhs"
+      {fix m assume m: "m > 0" "m < p"
+	{assume "m=1" hence "coprime p m" by simp}
+	moreover
+	{assume "p dvd m" hence "p \<le> m" using dvd_imp_le m by blast with m(2)
+	  have "coprime p m" by simp}
+	ultimately have "coprime p m" using prime_coprime[OF H, of m] by blast}
+      hence ?rhs using p0 by auto}
+    moreover
+    { assume H: "\<forall>m. 0 < m \<and> m < p \<longrightarrow> coprime p m"
+      from prime_factor[OF p0(2)] obtain q where q: "prime q" "q dvd p" by blast
+      from prime_ge_2[OF q(1)] have q0: "q > 0" by arith
+      from dvd_imp_le[OF q(2)] p0 have qp: "q \<le> p" by arith
+      {assume "q = p" hence ?lhs using q(1) by blast}
+      moreover
+      {assume "q\<noteq>p" with qp have qplt: "q < p" by arith
+	from H[rule_format, of q] qplt q0 have "coprime p q" by arith
+	with coprime_prime[of p q q] q have False by simp hence ?lhs by blast}
+      ultimately have ?lhs by blast}
+    ultimately have ?thesis by blast}
+  ultimately show ?thesis  by (cases"p=0 \<or> p=1", auto)
+qed
+
+lemma finite_number_segment: "card { m. 0 < m \<and> m < n } = n - 1"
+proof-
+  have "{ m. 0 < m \<and> m < n } = {1..<n}" by auto
+  thus ?thesis by simp
+qed
+
+lemma coprime_mod: assumes n: "n \<noteq> 0" shows "coprime (a mod n) n \<longleftrightarrow> coprime a n"
+  using n dvd_mod_iff[of _ n a] by (auto simp add: coprime)
+
+(* Congruences.                                                              *)
+
+lemma cong_mod_01[simp,presburger]:
+  "[x = y] (mod 0) \<longleftrightarrow> x = y" "[x = y] (mod 1)" "[x = 0] (mod n) \<longleftrightarrow> n dvd x"
+  by (simp_all add: modeq_def, presburger)
+
+lemma cong_sub_cases:
+  "[x = y] (mod n) \<longleftrightarrow> (if x <= y then [y - x = 0] (mod n) else [x - y = 0] (mod n))"
+apply (auto simp add: nat_mod)
+apply (rule_tac x="q2" in exI)
+apply (rule_tac x="q1" in exI, simp)
+apply (rule_tac x="q2" in exI)
+apply (rule_tac x="q1" in exI, simp)
+apply (rule_tac x="q1" in exI)
+apply (rule_tac x="q2" in exI, simp)
+apply (rule_tac x="q1" in exI)
+apply (rule_tac x="q2" in exI, simp)
+done
+
+lemma cong_mult_lcancel: assumes an: "coprime a n" and axy:"[a * x = a * y] (mod n)"
+  shows "[x = y] (mod n)"
+proof-
+  {assume "a = 0" with an axy coprime_0'[of n] have ?thesis by (simp add: modeq_def) }
+  moreover
+  {assume az: "a\<noteq>0"
+    {assume xy: "x \<le> y" hence axy': "a*x \<le> a*y" by simp
+      with axy cong_sub_cases[of "a*x" "a*y" n]  have "[a*(y - x) = 0] (mod n)"
+	by (simp only: if_True diff_mult_distrib2)
+      hence th: "n dvd a*(y -x)" by simp
+      from coprime_divprod[OF th] an have "n dvd y - x"
+	by (simp add: coprime_commute)
+      hence ?thesis using xy cong_sub_cases[of x y n] by simp}
+    moreover
+    {assume H: "\<not>x \<le> y" hence xy: "y \<le> x"  by arith
+      from H az have axy': "\<not> a*x \<le> a*y" by auto
+      with axy H cong_sub_cases[of "a*x" "a*y" n]  have "[a*(x - y) = 0] (mod n)"
+	by (simp only: if_False diff_mult_distrib2)
+      hence th: "n dvd a*(x - y)" by simp
+      from coprime_divprod[OF th] an have "n dvd x - y"
+	by (simp add: coprime_commute)
+      hence ?thesis using xy cong_sub_cases[of x y n] by simp}
+    ultimately have ?thesis by blast}
+  ultimately show ?thesis by blast
+qed
+
+lemma cong_mult_rcancel: assumes an: "coprime a n" and axy:"[x*a = y*a] (mod n)"
+  shows "[x = y] (mod n)"
+  using cong_mult_lcancel[OF an axy[unfolded mult_commute[of _a]]] .
+
+lemma cong_refl: "[x = x] (mod n)" by (simp add: modeq_def)
+
+lemma eq_imp_cong: "a = b \<Longrightarrow> [a = b] (mod n)" by (simp add: cong_refl)
+
+lemma cong_commute: "[x = y] (mod n) \<longleftrightarrow> [y = x] (mod n)"
+  by (auto simp add: modeq_def)
+
+lemma cong_trans[trans]: "[x = y] (mod n) \<Longrightarrow> [y = z] (mod n) \<Longrightarrow> [x = z] (mod n)"
+  by (simp add: modeq_def)
+
+lemma cong_add: assumes xx': "[x = x'] (mod n)" and yy':"[y = y'] (mod n)"
+  shows "[x + y = x' + y'] (mod n)"
+proof-
+  have "(x + y) mod n = (x mod n + y mod n) mod n"
+    by (simp add: mod_add_left_eq[of x y n] mod_add_right_eq[of "x mod n" y n])
+  also have "\<dots> = (x' mod n + y' mod n) mod n" using xx' yy' modeq_def by simp
+  also have "\<dots> = (x' + y') mod n"
+    by (simp add: mod_add_left_eq[of x' y' n] mod_add_right_eq[of "x' mod n" y' n])
+  finally show ?thesis unfolding modeq_def .
+qed
+
+lemma cong_mult: assumes xx': "[x = x'] (mod n)" and yy':"[y = y'] (mod n)"
+  shows "[x * y = x' * y'] (mod n)"
+proof-
+  have "(x * y) mod n = (x mod n) * (y mod n) mod n"
+    by (simp add: mod_mult_left_eq[of x y n] mod_mult_right_eq[of "x mod n" y n])
+  also have "\<dots> = (x' mod n) * (y' mod n) mod n" using xx'[unfolded modeq_def] yy'[unfolded modeq_def] by simp
+  also have "\<dots> = (x' * y') mod n"
+    by (simp add: mod_mult_left_eq[of x' y' n] mod_mult_right_eq[of "x' mod n" y' n])
+  finally show ?thesis unfolding modeq_def .
+qed
+
+lemma cong_exp: "[x = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"
+  by (induct k, auto simp add: cong_refl cong_mult)
+lemma cong_sub: assumes xx': "[x = x'] (mod n)" and yy': "[y = y'] (mod n)"
+  and yx: "y <= x" and yx': "y' <= x'"
+  shows "[x - y = x' - y'] (mod n)"
+proof-
+  { fix x a x' a' y b y' b'
+    have "(x::nat) + a = x' + a' \<Longrightarrow> y + b = y' + b' \<Longrightarrow> y <= x \<Longrightarrow> y' <= x'
+      \<Longrightarrow> (x - y) + (a + b') = (x' - y') + (a' + b)" by arith}
+  note th = this
+  from xx' yy' obtain q1 q2 q1' q2' where q12: "x + n*q1 = x'+n*q2"
+    and q12': "y + n*q1' = y'+n*q2'" unfolding nat_mod by blast+
+  from th[OF q12 q12' yx yx']
+  have "(x - y) + n*(q1 + q2') = (x' - y') + n*(q2 + q1')"
+    by (simp add: right_distrib)
+  thus ?thesis unfolding nat_mod by blast
+qed
+
+lemma cong_mult_lcancel_eq: assumes an: "coprime a n"
+  shows "[a * x = a * y] (mod n) \<longleftrightarrow> [x = y] (mod n)" (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume H: "?rhs" from cong_mult[OF cong_refl[of a n] H] show ?lhs .
+next
+  assume H: "?lhs" hence H': "[x*a = y*a] (mod n)" by (simp add: mult_commute)
+  from cong_mult_rcancel[OF an H'] show ?rhs  .
+qed
+
+lemma cong_mult_rcancel_eq: assumes an: "coprime a n"
+  shows "[x * a = y * a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
+using cong_mult_lcancel_eq[OF an, of x y] by (simp add: mult_commute)
+
+lemma cong_add_lcancel_eq: "[a + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)"
+  by (simp add: nat_mod)
+
+lemma cong_add_rcancel_eq: "[x + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
+  by (simp add: nat_mod)
+
+lemma cong_add_rcancel: "[x + a = y + a] (mod n) \<Longrightarrow> [x = y] (mod n)"
+  by (simp add: nat_mod)
+
+lemma cong_add_lcancel: "[a + x = a + y] (mod n) \<Longrightarrow> [x = y] (mod n)"
+  by (simp add: nat_mod)
+
+lemma cong_add_lcancel_eq_0: "[a + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
+  by (simp add: nat_mod)
+
+lemma cong_add_rcancel_eq_0: "[x + a = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
+  by (simp add: nat_mod)
+
+lemma cong_imp_eq: assumes xn: "x < n" and yn: "y < n" and xy: "[x = y] (mod n)"
+  shows "x = y"
+  using xy[unfolded modeq_def mod_less[OF xn] mod_less[OF yn]] .
+
+lemma cong_divides_modulus: "[x = y] (mod m) \<Longrightarrow> n dvd m ==> [x = y] (mod n)"
+  apply (auto simp add: nat_mod dvd_def)
+  apply (rule_tac x="k*q1" in exI)
+  apply (rule_tac x="k*q2" in exI)
+  by simp
+
+lemma cong_0_divides: "[x = 0] (mod n) \<longleftrightarrow> n dvd x" by simp
+
+lemma cong_1_divides:"[x = 1] (mod n) ==> n dvd x - 1"
+  apply (cases "x\<le>1", simp_all)
+  using cong_sub_cases[of x 1 n] by auto
+
+lemma cong_divides: "[x = y] (mod n) \<Longrightarrow> n dvd x \<longleftrightarrow> n dvd y"
+apply (auto simp add: nat_mod dvd_def)
+apply (rule_tac x="k + q1 - q2" in exI, simp add: add_mult_distrib2 diff_mult_distrib2)
+apply (rule_tac x="k + q2 - q1" in exI, simp add: add_mult_distrib2 diff_mult_distrib2)
+done
+
+lemma cong_coprime: assumes xy: "[x = y] (mod n)"
+  shows "coprime n x \<longleftrightarrow> coprime n y"
+proof-
+  {assume "n=0" hence ?thesis using xy by simp}
+  moreover
+  {assume nz: "n \<noteq> 0"
+  have "coprime n x \<longleftrightarrow> coprime (x mod n) n"
+    by (simp add: coprime_mod[OF nz, of x] coprime_commute[of n x])
+  also have "\<dots> \<longleftrightarrow> coprime (y mod n) n" using xy[unfolded modeq_def] by simp
+  also have "\<dots> \<longleftrightarrow> coprime y n" by (simp add: coprime_mod[OF nz, of y])
+  finally have ?thesis by (simp add: coprime_commute) }
+ultimately show ?thesis by blast
+qed
+
+lemma cong_mod: "~(n = 0) \<Longrightarrow> [a mod n = a] (mod n)" by (simp add: modeq_def)
+
+lemma mod_mult_cong: "~(a = 0) \<Longrightarrow> ~(b = 0)
+  \<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)"
+  by (simp add: modeq_def mod_mult2_eq mod_add_left_eq)
+
+lemma cong_mod_mult: "[x = y] (mod n) \<Longrightarrow> m dvd n \<Longrightarrow> [x = y] (mod m)"
+  apply (auto simp add: nat_mod dvd_def)
+  apply (rule_tac x="k*q1" in exI)
+  apply (rule_tac x="k*q2" in exI, simp)
+  done
+
+(* Some things when we know more about the order.                            *)
+
+lemma cong_le: "y <= x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)"
+  using nat_mod_lemma[of x y n]
+  apply auto
+  apply (simp add: nat_mod)
+  apply (rule_tac x="q" in exI)
+  apply (rule_tac x="q + q" in exI)
+  by (auto simp: algebra_simps)
+
+lemma cong_to_1: "[a = 1] (mod n) \<longleftrightarrow> a = 0 \<and> n = 1 \<or> (\<exists>m. a = 1 + m * n)"
+proof-
+  {assume "n = 0 \<or> n = 1\<or> a = 0 \<or> a = 1" hence ?thesis
+      apply (cases "n=0", simp_all add: cong_commute)
+      apply (cases "n=1", simp_all add: cong_commute modeq_def)
+      apply arith
+      by (cases "a=1", simp_all add: modeq_def cong_commute)}
+  moreover
+  {assume n: "n\<noteq>0" "n\<noteq>1" and a:"a\<noteq>0" "a \<noteq> 1" hence a': "a \<ge> 1" by simp
+    hence ?thesis using cong_le[OF a', of n] by auto }
+  ultimately show ?thesis by auto
+qed
+
+(* Some basic theorems about solving congruences.                            *)
+
+
+lemma cong_solve: assumes an: "coprime a n" shows "\<exists>x. [a * x = b] (mod n)"
+proof-
+  {assume "a=0" hence ?thesis using an by (simp add: modeq_def)}
+  moreover
+  {assume az: "a\<noteq>0"
+  from bezout_add_strong[OF az, of n]
+  obtain d x y where dxy: "d dvd a" "d dvd n" "a*x = n*y + d" by blast
+  from an[unfolded coprime, rule_format, of d] dxy(1,2) have d1: "d = 1" by blast
+  hence "a*x*b = (n*y + 1)*b" using dxy(3) by simp
+  hence "a*(x*b) = n*(y*b) + b" by algebra
+  hence "a*(x*b) mod n = (n*(y*b) + b) mod n" by simp
+  hence "a*(x*b) mod n = b mod n" by (simp add: mod_add_left_eq)
+  hence "[a*(x*b) = b] (mod n)" unfolding modeq_def .
+  hence ?thesis by blast}
+ultimately  show ?thesis by blast
+qed
+
+lemma cong_solve_unique: assumes an: "coprime a n" and nz: "n \<noteq> 0"
+  shows "\<exists>!x. x < n \<and> [a * x = b] (mod n)"
+proof-
+  let ?P = "\<lambda>x. x < n \<and> [a * x = b] (mod n)"
+  from cong_solve[OF an] obtain x where x: "[a*x = b] (mod n)" by blast
+  let ?x = "x mod n"
+  from x have th: "[a * ?x = b] (mod n)"
+    by (simp add: modeq_def mod_mult_right_eq[of a x n])
+  from mod_less_divisor[ of n x] nz th have Px: "?P ?x" by simp
+  {fix y assume Py: "y < n" "[a * y = b] (mod n)"
+    from Py(2) th have "[a * y = a*?x] (mod n)" by (simp add: modeq_def)
+    hence "[y = ?x] (mod n)" by (simp add: cong_mult_lcancel_eq[OF an])
+    with mod_less[OF Py(1)] mod_less_divisor[ of n x] nz
+    have "y = ?x" by (simp add: modeq_def)}
+  with Px show ?thesis by blast
+qed
+
+lemma cong_solve_unique_nontrivial:
+  assumes p: "prime p" and pa: "coprime p a" and x0: "0 < x" and xp: "x < p"
+  shows "\<exists>!y. 0 < y \<and> y < p \<and> [x * y = a] (mod p)"
+proof-
+  from p have p1: "p > 1" using prime_ge_2[OF p] by arith
+  hence p01: "p \<noteq> 0" "p \<noteq> 1" by arith+
+  from pa have ap: "coprime a p" by (simp add: coprime_commute)
+  from prime_coprime[OF p, of x] dvd_imp_le[of p x] x0 xp have px:"coprime x p"
+    by (auto simp add: coprime_commute)
+  from cong_solve_unique[OF px p01(1)]
+  obtain y where y: "y < p" "[x * y = a] (mod p)" "\<forall>z. z < p \<and> [x * z = a] (mod p) \<longrightarrow> z = y" by blast
+  {assume y0: "y = 0"
+    with y(2) have th: "p dvd a" by (simp add: cong_commute[of 0 a p])
+    with p coprime_prime[OF pa, of p] have False by simp}
+  with y show ?thesis unfolding Ex1_def using neq0_conv by blast
+qed
+lemma cong_unique_inverse_prime:
+  assumes p: "prime p" and x0: "0 < x" and xp: "x < p"
+  shows "\<exists>!y. 0 < y \<and> y < p \<and> [x * y = 1] (mod p)"
+  using cong_solve_unique_nontrivial[OF p coprime_1[of p] x0 xp] .
+
+(* Forms of the Chinese remainder theorem.                                   *)
+
+lemma cong_chinese:
+  assumes ab: "coprime a b" and  xya: "[x = y] (mod a)"
+  and xyb: "[x = y] (mod b)"
+  shows "[x = y] (mod a*b)"
+  using ab xya xyb
+  by (simp add: cong_sub_cases[of x y a] cong_sub_cases[of x y b]
+    cong_sub_cases[of x y "a*b"])
+(cases "x \<le> y", simp_all add: divides_mul[of a _ b])
+
+lemma chinese_remainder_unique:
+  assumes ab: "coprime a b" and az: "a \<noteq> 0" and bz: "b\<noteq>0"
+  shows "\<exists>!x. x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"
+proof-
+  from az bz have abpos: "a*b > 0" by simp
+  from chinese_remainder[OF ab az bz] obtain x q1 q2 where
+    xq12: "x = m + q1 * a" "x = n + q2 * b" by blast
+  let ?w = "x mod (a*b)"
+  have wab: "?w < a*b" by (simp add: mod_less_divisor[OF abpos])
+  from xq12(1) have "?w mod a = ((m + q1 * a) mod (a*b)) mod a" by simp
+  also have "\<dots> = m mod a" apply (simp add: mod_mult2_eq)
+    apply (subst mod_add_left_eq)
+    by simp
+  finally have th1: "[?w = m] (mod a)" by (simp add: modeq_def)
+  from xq12(2) have "?w mod b = ((n + q2 * b) mod (a*b)) mod b" by simp
+  also have "\<dots> = ((n + q2 * b) mod (b*a)) mod b" by (simp add: mult_commute)
+  also have "\<dots> = n mod b" apply (simp add: mod_mult2_eq)
+    apply (subst mod_add_left_eq)
+    by simp
+  finally have th2: "[?w = n] (mod b)" by (simp add: modeq_def)
+  {fix y assume H: "y < a*b" "[y = m] (mod a)" "[y = n] (mod b)"
+    with th1 th2 have H': "[y = ?w] (mod a)" "[y = ?w] (mod b)"
+      by (simp_all add: modeq_def)
+    from cong_chinese[OF ab H'] mod_less[OF H(1)] mod_less[OF wab]
+    have "y = ?w" by (simp add: modeq_def)}
+  with th1 th2 wab show ?thesis by blast
+qed
+
+lemma chinese_remainder_coprime_unique:
+  assumes ab: "coprime a b" and az: "a \<noteq> 0" and bz: "b \<noteq> 0"
+  and ma: "coprime m a" and nb: "coprime n b"
+  shows "\<exists>!x. coprime x (a * b) \<and> x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"
+proof-
+  let ?P = "\<lambda>x. x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"
+  from chinese_remainder_unique[OF ab az bz]
+  obtain x where x: "x < a * b" "[x = m] (mod a)" "[x = n] (mod b)"
+    "\<forall>y. ?P y \<longrightarrow> y = x" by blast
+  from ma nb cong_coprime[OF x(2)] cong_coprime[OF x(3)]
+  have "coprime x a" "coprime x b" by (simp_all add: coprime_commute)
+  with coprime_mul[of x a b] have "coprime x (a*b)" by simp
+  with x show ?thesis by blast
+qed
+
+(* Euler totient function.                                                   *)
+
+definition phi_def: "\<phi> n = card { m. 0 < m \<and> m <= n \<and> coprime m n }"
+
+lemma phi_0[simp]: "\<phi> 0 = 0"
+  unfolding phi_def by auto
+
+lemma phi_finite[simp]: "finite ({ m. 0 < m \<and> m <= n \<and> coprime m n })"
+proof-
+  have "{ m. 0 < m \<and> m <= n \<and> coprime m n } \<subseteq> {0..n}" by auto
+  thus ?thesis by (auto intro: finite_subset)
+qed
+
+declare coprime_1[presburger]
+lemma phi_1[simp]: "\<phi> 1 = 1"
+proof-
+  {fix m
+    have "0 < m \<and> m <= 1 \<and> coprime m 1 \<longleftrightarrow> m = 1" by presburger }
+  thus ?thesis by (simp add: phi_def)
+qed
+
+lemma [simp]: "\<phi> (Suc 0) = Suc 0" using phi_1 by simp
+
+lemma phi_alt: "\<phi>(n) = card { m. coprime m n \<and> m < n}"
+proof-
+  {assume "n=0 \<or> n=1" hence ?thesis by (cases "n=0", simp_all)}
+  moreover
+  {assume n: "n\<noteq>0" "n\<noteq>1"
+    {fix m
+      from n have "0 < m \<and> m <= n \<and> coprime m n \<longleftrightarrow> coprime m n \<and> m < n"
+	apply (cases "m = 0", simp_all)
+	apply (cases "m = 1", simp_all)
+	apply (cases "m = n", auto)
+	done }
+    hence ?thesis unfolding phi_def by simp}
+  ultimately show ?thesis by auto
+qed
+
+lemma phi_finite_lemma[simp]: "finite {m. coprime m n \<and>  m < n}" (is "finite ?S")
+  by (rule finite_subset[of "?S" "{0..n}"], auto)
+
+lemma phi_another: assumes n: "n\<noteq>1"
+  shows "\<phi> n = card {m. 0 < m \<and> m < n \<and> coprime m n }"
+proof-
+  {fix m
+    from n have "0 < m \<and> m < n \<and> coprime m n \<longleftrightarrow> coprime m n \<and> m < n"
+      by (cases "m=0", auto)}
+  thus ?thesis unfolding phi_alt by auto
+qed
+
+lemma phi_limit: "\<phi> n \<le> n"
+proof-
+  have "{ m. coprime m n \<and> m < n} \<subseteq> {0 ..<n}" by auto
+  with card_mono[of "{0 ..<n}" "{ m. coprime m n \<and> m < n}"]
+  show ?thesis unfolding phi_alt by auto
+qed
+
+lemma stupid[simp]: "{m. (0::nat) < m \<and> m < n} = {1..<n}"
+  by auto
+
+lemma phi_limit_strong: assumes n: "n\<noteq>1"
+  shows "\<phi>(n) \<le> n - 1"
+proof-
+  show ?thesis
+    unfolding phi_another[OF n] finite_number_segment[of n, symmetric]
+    by (rule card_mono[of "{m. 0 < m \<and> m < n}" "{m. 0 < m \<and> m < n \<and> coprime m n}"], auto)
+qed
+
+lemma phi_lowerbound_1_strong: assumes n: "n \<ge> 1"
+  shows "\<phi>(n) \<ge> 1"
+proof-
+  let ?S = "{ m. 0 < m \<and> m <= n \<and> coprime m n }"
+  from card_0_eq[of ?S] n have "\<phi> n \<noteq> 0" unfolding phi_alt
+    apply auto
+    apply (cases "n=1", simp_all)
+    apply (rule exI[where x=1], simp)
+    done
+  thus ?thesis by arith
+qed
+
+lemma phi_lowerbound_1: "2 <= n ==> 1 <= \<phi>(n)"
+  using phi_lowerbound_1_strong[of n] by auto
+
+lemma phi_lowerbound_2: assumes n: "3 <= n" shows "2 <= \<phi> (n)"
+proof-
+  let ?S = "{ m. 0 < m \<and> m <= n \<and> coprime m n }"
+  have inS: "{1, n - 1} \<subseteq> ?S" using n coprime_plus1[of "n - 1"]
+    by (auto simp add: coprime_commute)
+  from n have c2: "card {1, n - 1} = 2" by (auto simp add: card_insert_if)
+  from card_mono[of ?S "{1, n - 1}", simplified inS c2] show ?thesis
+    unfolding phi_def by auto
+qed
+
+lemma phi_prime: "\<phi> n = n - 1 \<and> n\<noteq>0 \<and> n\<noteq>1 \<longleftrightarrow> prime n"
+proof-
+  {assume "n=0 \<or> n=1" hence ?thesis by (cases "n=1", simp_all)}
+  moreover
+  {assume n: "n\<noteq>0" "n\<noteq>1"
+    let ?S = "{m. 0 < m \<and> m < n}"
+    have fS: "finite ?S" by simp
+    let ?S' = "{m. 0 < m \<and> m < n \<and> coprime m n}"
+    have fS':"finite ?S'" apply (rule finite_subset[of ?S' ?S]) by auto
+    {assume H: "\<phi> n = n - 1 \<and> n\<noteq>0 \<and> n\<noteq>1"
+      hence ceq: "card ?S' = card ?S"
+      using n finite_number_segment[of n] phi_another[OF n(2)] by simp
+      {fix m assume m: "0 < m" "m < n" "\<not> coprime m n"
+	hence mS': "m \<notin> ?S'" by auto
+	have "insert m ?S' \<le> ?S" using m by auto
+	from m have "card (insert m ?S') \<le> card ?S"
+	  by - (rule card_mono[of ?S "insert m ?S'"], auto)
+	hence False
+	  unfolding card_insert_disjoint[of "?S'" m, OF fS' mS'] ceq
+	  by simp }
+      hence "\<forall>m. 0 <m \<and> m < n \<longrightarrow> coprime m n" by blast
+      hence "prime n" unfolding prime using n by (simp add: coprime_commute)}
+    moreover
+    {assume H: "prime n"
+      hence "?S = ?S'" unfolding prime using n
+	by (auto simp add: coprime_commute)
+      hence "card ?S = card ?S'" by simp
+      hence "\<phi> n = n - 1" unfolding phi_another[OF n(2)] by simp}
+    ultimately have ?thesis using n by blast}
+  ultimately show ?thesis by (cases "n=0") blast+
+qed
+
+(* Multiplicativity property.                                                *)
+
+lemma phi_multiplicative: assumes ab: "coprime a b"
+  shows "\<phi> (a * b) = \<phi> a * \<phi> b"
+proof-
+  {assume "a = 0 \<or> b = 0 \<or> a = 1 \<or> b = 1"
+    hence ?thesis
+      by (cases "a=0", simp, cases "b=0", simp, cases"a=1", simp_all) }
+  moreover
+  {assume a: "a\<noteq>0" "a\<noteq>1" and b: "b\<noteq>0" "b\<noteq>1"
+    hence ab0: "a*b \<noteq> 0" by simp
+    let ?S = "\<lambda>k. {m. coprime m k \<and> m < k}"
+    let ?f = "\<lambda>x. (x mod a, x mod b)"
+    have eq: "?f ` (?S (a*b)) = (?S a \<times> ?S b)"
+    proof-
+      {fix x assume x:"x \<in> ?S (a*b)"
+	hence x': "coprime x (a*b)" "x < a*b" by simp_all
+	hence xab: "coprime x a" "coprime x b" by (simp_all add: coprime_mul_eq)
+	from mod_less_divisor a b have xab':"x mod a < a" "x mod b < b" by auto
+	from xab xab' have "?f x \<in> (?S a \<times> ?S b)"
+	  by (simp add: coprime_mod[OF a(1)] coprime_mod[OF b(1)])}
+      moreover
+      {fix x y assume x: "x \<in> ?S a" and y: "y \<in> ?S b"
+	hence x': "coprime x a" "x < a" and y': "coprime y b" "y < b" by simp_all
+	from chinese_remainder_coprime_unique[OF ab a(1) b(1) x'(1) y'(1)]
+	obtain z where z: "coprime z (a * b)" "z < a * b" "[z = x] (mod a)"
+	  "[z = y] (mod b)" by blast
+	hence "(x,y) \<in> ?f ` (?S (a*b))"
+	  using y'(2) mod_less_divisor[of b y] x'(2) mod_less_divisor[of a x]
+	  by (auto simp add: image_iff modeq_def)}
+      ultimately show ?thesis by auto
+    qed
+    have finj: "inj_on ?f (?S (a*b))"
+      unfolding inj_on_def
+    proof(clarify)
+      fix x y assume H: "coprime x (a * b)" "x < a * b" "coprime y (a * b)"
+	"y < a * b" "x mod a = y mod a" "x mod b = y mod b"
+      hence cp: "coprime x a" "coprime x b" "coprime y a" "coprime y b"
+	by (simp_all add: coprime_mul_eq)
+      from chinese_remainder_coprime_unique[OF ab a(1) b(1) cp(3,4)] H
+      show "x = y" unfolding modeq_def by blast
+    qed
+    from card_image[OF finj, unfolded eq] have ?thesis
+      unfolding phi_alt by simp }
+  ultimately show ?thesis by auto
+qed
+
+(* Fermat's Little theorem / Fermat-Euler theorem.                           *)
+
+
+lemma nproduct_mod:
+  assumes fS: "finite S" and n0: "n \<noteq> 0"
+  shows "[setprod (\<lambda>m. a(m) mod n) S = setprod a S] (mod n)"
+proof-
+  have th1:"[1 = 1] (mod n)" by (simp add: modeq_def)
+  from cong_mult
+  have th3:"\<forall>x1 y1 x2 y2.
+    [x1 = x2] (mod n) \<and> [y1 = y2] (mod n) \<longrightarrow> [x1 * y1 = x2 * y2] (mod n)"
+    by blast
+  have th4:"\<forall>x\<in>S. [a x mod n = a x] (mod n)" by (simp add: modeq_def)
+  from fold_image_related[where h="(\<lambda>m. a(m) mod n)" and g=a, OF th1 th3 fS, OF th4] show ?thesis unfolding setprod_def by (simp add: fS)
+qed
+
+lemma nproduct_cmul:
+  assumes fS:"finite S"
+  shows "setprod (\<lambda>m. (c::'a::{comm_monoid_mult})* a(m)) S = c ^ (card S) * setprod a S"
+unfolding setprod_timesf setprod_constant[OF fS, of c] ..
+
+lemma coprime_nproduct:
+  assumes fS: "finite S" and Sn: "\<forall>x\<in>S. coprime n (a x)"
+  shows "coprime n (setprod a S)"
+  using fS unfolding setprod_def by (rule finite_subset_induct)
+    (insert Sn, auto simp add: coprime_mul)
+
+lemma fermat_little: assumes an: "coprime a n"
+  shows "[a ^ (\<phi> n) = 1] (mod n)"
+proof-
+  {assume "n=0" hence ?thesis by simp}
+  moreover
+  {assume "n=1" hence ?thesis by (simp add: modeq_def)}
+  moreover
+  {assume nz: "n \<noteq> 0" and n1: "n \<noteq> 1"
+    let ?S = "{m. coprime m n \<and> m < n}"
+    let ?P = "\<Prod> ?S"
+    have fS: "finite ?S" by simp
+    have cardfS: "\<phi> n = card ?S" unfolding phi_alt ..
+    {fix m assume m: "m \<in> ?S"
+      hence "coprime m n" by simp
+      with coprime_mul[of n a m] an have "coprime (a*m) n"
+	by (simp add: coprime_commute)}
+    hence Sn: "\<forall>m\<in> ?S. coprime (a*m) n " by blast
+    from coprime_nproduct[OF fS, of n "\<lambda>m. m"] have nP:"coprime ?P n"
+      by (simp add: coprime_commute)
+    have Paphi: "[?P*a^ (\<phi> n) = ?P*1] (mod n)"
+    proof-
+      let ?h = "\<lambda>m. m mod n"
+      {fix m assume mS: "m\<in> ?S"
+	hence "?h m \<in> ?S" by simp}
+      hence hS: "?h ` ?S = ?S"by (auto simp add: image_iff)
+      have "a\<noteq>0" using an n1 nz apply- apply (rule ccontr) by simp
+      hence inj: "inj_on (op * a) ?S" unfolding inj_on_def by simp
+
+      have eq0: "fold_image op * (?h \<circ> op * a) 1 {m. coprime m n \<and> m < n} =
+     fold_image op * (\<lambda>m. m) 1 {m. coprime m n \<and> m < n}"
+      proof (rule fold_image_eq_general[where h="?h o (op * a)"])
+	show "finite ?S" using fS .
+      next
+	{fix y assume yS: "y \<in> ?S" hence y: "coprime y n" "y < n" by simp_all
+	  from cong_solve_unique[OF an nz, of y]
+	  obtain x where x:"x < n" "[a * x = y] (mod n)" "\<forall>z. z < n \<and> [a * z = y] (mod n) \<longrightarrow> z=x" by blast
+	  from cong_coprime[OF x(2)] y(1)
+	  have xm: "coprime x n" by (simp add: coprime_mul_eq coprime_commute)
+	  {fix z assume "z \<in> ?S" "(?h \<circ> op * a) z = y"
+	    hence z: "coprime z n" "z < n" "(?h \<circ> op * a) z = y" by simp_all
+	    from x(3)[rule_format, of z] z(2,3) have "z=x"
+	      unfolding modeq_def mod_less[OF y(2)] by simp}
+	  with xm x(1,2) have "\<exists>!x. x \<in> ?S \<and> (?h \<circ> op * a) x = y"
+	    unfolding modeq_def mod_less[OF y(2)] by auto }
+	thus "\<forall>y\<in>{m. coprime m n \<and> m < n}.
+       \<exists>!x. x \<in> {m. coprime m n \<and> m < n} \<and> ((\<lambda>m. m mod n) \<circ> op * a) x = y" by blast
+      next
+	{fix x assume xS: "x\<in> ?S"
+	  hence x: "coprime x n" "x < n" by simp_all
+	  with an have "coprime (a*x) n"
+	    by (simp add: coprime_mul_eq[of n a x] coprime_commute)
+	  hence "?h (a*x) \<in> ?S" using nz
+	    by (simp add: coprime_mod[OF nz] mod_less_divisor)}
+	thus " \<forall>x\<in>{m. coprime m n \<and> m < n}.
+       ((\<lambda>m. m mod n) \<circ> op * a) x \<in> {m. coprime m n \<and> m < n} \<and>
+       ((\<lambda>m. m mod n) \<circ> op * a) x = ((\<lambda>m. m mod n) \<circ> op * a) x" by simp
+      qed
+      from nproduct_mod[OF fS nz, of "op * a"]
+      have "[(setprod (op *a) ?S) = (setprod (?h o (op * a)) ?S)] (mod n)"
+	unfolding o_def
+	by (simp add: cong_commute)
+      also have "[setprod (?h o (op * a)) ?S = ?P ] (mod n)"
+	using eq0 fS an by (simp add: setprod_def modeq_def o_def)
+      finally show "[?P*a^ (\<phi> n) = ?P*1] (mod n)"
+	unfolding cardfS mult_commute[of ?P "a^ (card ?S)"]
+	  nproduct_cmul[OF fS, symmetric] mult_1_right by simp
+    qed
+    from cong_mult_lcancel[OF nP Paphi] have ?thesis . }
+  ultimately show ?thesis by blast
+qed
+
+lemma fermat_little_prime: assumes p: "prime p" and ap: "coprime a p"
+  shows "[a^ (p - 1) = 1] (mod p)"
+  using fermat_little[OF ap] p[unfolded phi_prime[symmetric]]
+by simp
+
+
+(* Lucas's theorem.                                                          *)
+
+lemma lucas_coprime_lemma:
+  assumes m: "m\<noteq>0" and am: "[a^m = 1] (mod n)"
+  shows "coprime a n"
+proof-
+  {assume "n=1" hence ?thesis by simp}
+  moreover
+  {assume "n = 0" hence ?thesis using am m exp_eq_1[of a m] by simp}
+  moreover
+  {assume n: "n\<noteq>0" "n\<noteq>1"
+    from m obtain m' where m': "m = Suc m'" by (cases m, blast+)
+    {fix d
+      assume d: "d dvd a" "d dvd n"
+      from n have n1: "1 < n" by arith
+      from am mod_less[OF n1] have am1: "a^m mod n = 1" unfolding modeq_def by simp
+      from dvd_mult2[OF d(1), of "a^m'"] have dam:"d dvd a^m" by (simp add: m')
+      from dvd_mod_iff[OF d(2), of "a^m"] dam am1
+      have "d = 1" by simp }
+    hence ?thesis unfolding coprime by auto
+  }
+  ultimately show ?thesis by blast
+qed
+
+lemma lucas_weak:
+  assumes n: "n \<ge> 2" and an:"[a^(n - 1) = 1] (mod n)"
+  and nm: "\<forall>m. 0 <m \<and> m < n - 1 \<longrightarrow> \<not> [a^m = 1] (mod n)"
+  shows "prime n"
+proof-
+  from n have n1: "n \<noteq> 1" "n\<noteq>0" "n - 1 \<noteq> 0" "n - 1 > 0" "n - 1 < n" by arith+
+  from lucas_coprime_lemma[OF n1(3) an] have can: "coprime a n" .
+  from fermat_little[OF can] have afn: "[a ^ \<phi> n = 1] (mod n)" .
+  {assume "\<phi> n \<noteq> n - 1"
+    with phi_limit_strong[OF n1(1)] phi_lowerbound_1[OF n]
+    have c:"\<phi> n > 0 \<and> \<phi> n < n - 1" by arith
+    from nm[rule_format, OF c] afn have False ..}
+  hence "\<phi> n = n - 1" by blast
+  with phi_prime[of n] n1(1,2) show ?thesis by simp
+qed
+
+lemma nat_exists_least_iff: "(\<exists>(n::nat). P n) \<longleftrightarrow> (\<exists>n. P n \<and> (\<forall>m < n. \<not> P m))"
+  (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume ?rhs thus ?lhs by blast
+next
+  assume H: ?lhs then obtain n where n: "P n" by blast
+  let ?x = "Least P"
+  {fix m assume m: "m < ?x"
+    from not_less_Least[OF m] have "\<not> P m" .}
+  with LeastI_ex[OF H] show ?rhs by blast
+qed
+
+lemma nat_exists_least_iff': "(\<exists>(n::nat). P n) \<longleftrightarrow> (P (Least P) \<and> (\<forall>m < (Least P). \<not> P m))"
+  (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  {assume ?rhs hence ?lhs by blast}
+  moreover
+  { assume H: ?lhs then obtain n where n: "P n" by blast
+    let ?x = "Least P"
+    {fix m assume m: "m < ?x"
+      from not_less_Least[OF m] have "\<not> P m" .}
+    with LeastI_ex[OF H] have ?rhs by blast}
+  ultimately show ?thesis by blast
+qed
+
+lemma power_mod: "((x::nat) mod m)^n mod m = x^n mod m"
+proof(induct n)
+  case 0 thus ?case by simp
+next
+  case (Suc n)
+  have "(x mod m)^(Suc n) mod m = ((x mod m) * (((x mod m) ^ n) mod m)) mod m"
+    by (simp add: mod_mult_right_eq[symmetric])
+  also have "\<dots> = ((x mod m) * (x^n mod m)) mod m" using Suc.hyps by simp
+  also have "\<dots> = x^(Suc n) mod m"
+    by (simp add: mod_mult_left_eq[symmetric] mod_mult_right_eq[symmetric])
+  finally show ?case .
+qed
+
+lemma lucas:
+  assumes n2: "n \<ge> 2" and an1: "[a^(n - 1) = 1] (mod n)"
+  and pn: "\<forall>p. prime p \<and> p dvd n - 1 \<longrightarrow> \<not> [a^((n - 1) div p) = 1] (mod n)"
+  shows "prime n"
+proof-
+  from n2 have n01: "n\<noteq>0" "n\<noteq>1" "n - 1 \<noteq> 0" by arith+
+  from mod_less_divisor[of n 1] n01 have onen: "1 mod n = 1" by simp
+  from lucas_coprime_lemma[OF n01(3) an1] cong_coprime[OF an1]
+  have an: "coprime a n" "coprime (a^(n - 1)) n" by (simp_all add: coprime_commute)
+  {assume H0: "\<exists>m. 0 < m \<and> m < n - 1 \<and> [a ^ m = 1] (mod n)" (is "EX m. ?P m")
+    from H0[unfolded nat_exists_least_iff[of ?P]] obtain m where
+      m: "0 < m" "m < n - 1" "[a ^ m = 1] (mod n)" "\<forall>k <m. \<not>?P k" by blast
+    {assume nm1: "(n - 1) mod m > 0"
+      from mod_less_divisor[OF m(1)] have th0:"(n - 1) mod m < m" by blast
+      let ?y = "a^ ((n - 1) div m * m)"
+      note mdeq = mod_div_equality[of "(n - 1)" m]
+      from coprime_exp[OF an(1)[unfolded coprime_commute[of a n]],
+	of "(n - 1) div m * m"]
+      have yn: "coprime ?y n" by (simp add: coprime_commute)
+      have "?y mod n = (a^m)^((n - 1) div m) mod n"
+	by (simp add: algebra_simps power_mult)
+      also have "\<dots> = (a^m mod n)^((n - 1) div m) mod n"
+	using power_mod[of "a^m" n "(n - 1) div m"] by simp
+      also have "\<dots> = 1" using m(3)[unfolded modeq_def onen] onen
+	by (simp add: power_Suc0)
+      finally have th3: "?y mod n = 1"  .
+      have th2: "[?y * a ^ ((n - 1) mod m) = ?y* 1] (mod n)"
+	using an1[unfolded modeq_def onen] onen
+	  mod_div_equality[of "(n - 1)" m, symmetric]
+	by (simp add:power_add[symmetric] modeq_def th3 del: One_nat_def)
+      from cong_mult_lcancel[of ?y n "a^((n - 1) mod m)" 1, OF yn th2]
+      have th1: "[a ^ ((n - 1) mod m) = 1] (mod n)"  .
+      from m(4)[rule_format, OF th0] nm1
+	less_trans[OF mod_less_divisor[OF m(1), of "n - 1"] m(2)] th1
+      have False by blast }
+    hence "(n - 1) mod m = 0" by auto
+    then have mn: "m dvd n - 1" by presburger
+    then obtain r where r: "n - 1 = m*r" unfolding dvd_def by blast
+    from n01 r m(2) have r01: "r\<noteq>0" "r\<noteq>1" by - (rule ccontr, simp)+
+    from prime_factor[OF r01(2)] obtain p where p: "prime p" "p dvd r" by blast
+    hence th: "prime p \<and> p dvd n - 1" unfolding r by (auto intro: dvd_mult)
+    have "(a ^ ((n - 1) div p)) mod n = (a^(m*r div p)) mod n" using r
+      by (simp add: power_mult)
+    also have "\<dots> = (a^(m*(r div p))) mod n" using div_mult1_eq[of m r p] p(2)[unfolded dvd_eq_mod_eq_0] by simp
+    also have "\<dots> = ((a^m)^(r div p)) mod n" by (simp add: power_mult)
+    also have "\<dots> = ((a^m mod n)^(r div p)) mod n" using power_mod[of "a^m" "n" "r div p" ] ..
+    also have "\<dots> = 1" using m(3) onen by (simp add: modeq_def power_Suc0)
+    finally have "[(a ^ ((n - 1) div p))= 1] (mod n)"
+      using onen by (simp add: modeq_def)
+    with pn[rule_format, OF th] have False by blast}
+  hence th: "\<forall>m. 0 < m \<and> m < n - 1 \<longrightarrow> \<not> [a ^ m = 1] (mod n)" by blast
+  from lucas_weak[OF n2 an1 th] show ?thesis .
+qed
+
+(* Definition of the order of a number mod n (0 in non-coprime case).        *)
+
+definition "ord n a = (if coprime n a then Least (\<lambda>d. d > 0 \<and> [a ^d = 1] (mod n)) else 0)"
+
+(* This has the expected properties.                                         *)
+
+lemma coprime_ord:
+  assumes na: "coprime n a"
+  shows "ord n a > 0 \<and> [a ^(ord n a) = 1] (mod n) \<and> (\<forall>m. 0 < m \<and> m < ord n a \<longrightarrow> \<not> [a^ m = 1] (mod n))"
+proof-
+  let ?P = "\<lambda>d. 0 < d \<and> [a ^ d = 1] (mod n)"
+  from euclid[of a] obtain p where p: "prime p" "a < p" by blast
+  from na have o: "ord n a = Least ?P" by (simp add: ord_def)
+  {assume "n=0 \<or> n=1" with na have "\<exists>m>0. ?P m" apply auto apply (rule exI[where x=1]) by (simp  add: modeq_def)}
+  moreover
+  {assume "n\<noteq>0 \<and> n\<noteq>1" hence n2:"n \<ge> 2" by arith
+    from na have na': "coprime a n" by (simp add: coprime_commute)
+    from phi_lowerbound_1[OF n2] fermat_little[OF na']
+    have ex: "\<exists>m>0. ?P m" by - (rule exI[where x="\<phi> n"], auto) }
+  ultimately have ex: "\<exists>m>0. ?P m" by blast
+  from nat_exists_least_iff'[of ?P] ex na show ?thesis
+    unfolding o[symmetric] by auto
+qed
+(* With the special value 0 for non-coprime case, it's more convenient.      *)
+lemma ord_works:
+ "[a ^ (ord n a) = 1] (mod n) \<and> (\<forall>m. 0 < m \<and> m < ord n a \<longrightarrow> ~[a^ m = 1] (mod n))"
+apply (cases "coprime n a")
+using coprime_ord[of n a]
+by (blast, simp add: ord_def modeq_def)
+
+lemma ord: "[a^(ord n a) = 1] (mod n)" using ord_works by blast
+lemma ord_minimal: "0 < m \<Longrightarrow> m < ord n a \<Longrightarrow> ~[a^m = 1] (mod n)"
+  using ord_works by blast
+lemma ord_eq_0: "ord n a = 0 \<longleftrightarrow> ~coprime n a"
+by (cases "coprime n a", simp add: neq0_conv coprime_ord, simp add: neq0_conv ord_def)
+
+lemma ord_divides:
+ "[a ^ d = 1] (mod n) \<longleftrightarrow> ord n a dvd d" (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume rh: ?rhs
+  then obtain k where "d = ord n a * k" unfolding dvd_def by blast
+  hence "[a ^ d = (a ^ (ord n a) mod n)^k] (mod n)"
+    by (simp add : modeq_def power_mult power_mod)
+  also have "[(a ^ (ord n a) mod n)^k = 1] (mod n)"
+    using ord[of a n, unfolded modeq_def]
+    by (simp add: modeq_def power_mod power_Suc0)
+  finally  show ?lhs .
+next
+  assume lh: ?lhs
+  { assume H: "\<not> coprime n a"
+    hence o: "ord n a = 0" by (simp add: ord_def)
+    {assume d: "d=0" with o H have ?rhs by (simp add: modeq_def)}
+    moreover
+    {assume d0: "d\<noteq>0" then obtain d' where d': "d = Suc d'" by (cases d, auto)
+      from H[unfolded coprime]
+      obtain p where p: "p dvd n" "p dvd a" "p \<noteq> 1" by auto
+      from lh[unfolded nat_mod]
+      obtain q1 q2 where q12:"a ^ d + n * q1 = 1 + n * q2" by blast
+      hence "a ^ d + n * q1 - n * q2 = 1" by simp
+      with dvd_diff_nat [OF dvd_add [OF divides_rexp[OF p(2), of d'] dvd_mult2[OF p(1), of q1]] dvd_mult2[OF p(1), of q2]] d' have "p dvd 1" by simp
+      with p(3) have False by simp
+      hence ?rhs ..}
+    ultimately have ?rhs by blast}
+  moreover
+  {assume H: "coprime n a"
+    let ?o = "ord n a"
+    let ?q = "d div ord n a"
+    let ?r = "d mod ord n a"
+    from cong_exp[OF ord[of a n], of ?q]
+    have eqo: "[(a^?o)^?q = 1] (mod n)"  by (simp add: modeq_def power_Suc0)
+    from H have onz: "?o \<noteq> 0" by (simp add: ord_eq_0)
+    hence op: "?o > 0" by simp
+    from mod_div_equality[of d "ord n a"] lh
+    have "[a^(?o*?q + ?r) = 1] (mod n)" by (simp add: modeq_def mult_commute)
+    hence "[(a^?o)^?q * (a^?r) = 1] (mod n)"
+      by (simp add: modeq_def power_mult[symmetric] power_add[symmetric])
+    hence th: "[a^?r = 1] (mod n)"
+      using eqo mod_mult_left_eq[of "(a^?o)^?q" "a^?r" n]
+      apply (simp add: modeq_def del: One_nat_def)
+      by (simp add: mod_mult_left_eq[symmetric])
+    {assume r: "?r = 0" hence ?rhs by (simp add: dvd_eq_mod_eq_0)}
+    moreover
+    {assume r: "?r \<noteq> 0"
+      with mod_less_divisor[OF op, of d] have r0o:"?r >0 \<and> ?r < ?o" by simp
+      from conjunct2[OF ord_works[of a n], rule_format, OF r0o] th
+      have ?rhs by blast}
+    ultimately have ?rhs by blast}
+  ultimately  show ?rhs by blast
+qed
+
+lemma order_divides_phi: "coprime n a \<Longrightarrow> ord n a dvd \<phi> n"
+using ord_divides fermat_little coprime_commute by simp
+lemma order_divides_expdiff:
+  assumes na: "coprime n a"
+  shows "[a^d = a^e] (mod n) \<longleftrightarrow> [d = e] (mod (ord n a))"
+proof-
+  {fix n a d e
+    assume na: "coprime n a" and ed: "(e::nat) \<le> d"
+    hence "\<exists>c. d = e + c" by arith
+    then obtain c where c: "d = e + c" by arith
+    from na have an: "coprime a n" by (simp add: coprime_commute)
+    from coprime_exp[OF na, of e]
+    have aen: "coprime (a^e) n" by (simp add: coprime_commute)
+    from coprime_exp[OF na, of c]
+    have acn: "coprime (a^c) n" by (simp add: coprime_commute)
+    have "[a^d = a^e] (mod n) \<longleftrightarrow> [a^(e + c) = a^(e + 0)] (mod n)"
+      using c by simp
+    also have "\<dots> \<longleftrightarrow> [a^e* a^c = a^e *a^0] (mod n)" by (simp add: power_add)
+    also have  "\<dots> \<longleftrightarrow> [a ^ c = 1] (mod n)"
+      using cong_mult_lcancel_eq[OF aen, of "a^c" "a^0"] by simp
+    also  have "\<dots> \<longleftrightarrow> ord n a dvd c" by (simp only: ord_divides)
+    also have "\<dots> \<longleftrightarrow> [e + c = e + 0] (mod ord n a)"
+      using cong_add_lcancel_eq[of e c 0 "ord n a", simplified cong_0_divides]
+      by simp
+    finally have "[a^d = a^e] (mod n) \<longleftrightarrow> [d = e] (mod (ord n a))"
+      using c by simp }
+  note th = this
+  have "e \<le> d \<or> d \<le> e" by arith
+  moreover
+  {assume ed: "e \<le> d" from th[OF na ed] have ?thesis .}
+  moreover
+  {assume de: "d \<le> e"
+    from th[OF na de] have ?thesis by (simp add: cong_commute) }
+  ultimately show ?thesis by blast
+qed
+
+(* Another trivial primality characterization.                               *)
+
+lemma prime_prime_factor:
+  "prime n \<longleftrightarrow> n \<noteq> 1\<and> (\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n)"
+proof-
+  {assume n: "n=0 \<or> n=1" hence ?thesis using prime_0 two_is_prime by auto}
+  moreover
+  {assume n: "n\<noteq>0" "n\<noteq>1"
+    {assume pn: "prime n"
+
+      from pn[unfolded prime_def] have "\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n"
+	using n
+	apply (cases "n = 0 \<or> n=1",simp)
+	by (clarsimp, erule_tac x="p" in allE, auto)}
+    moreover
+    {assume H: "\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n"
+      from n have n1: "n > 1" by arith
+      {fix m assume m: "m dvd n" "m\<noteq>1"
+	from prime_factor[OF m(2)] obtain p where
+	  p: "prime p" "p dvd m" by blast
+	from dvd_trans[OF p(2) m(1)] p(1) H have "p = n" by blast
+	with p(2) have "n dvd m"  by simp
+	hence "m=n"  using dvd_anti_sym[OF m(1)] by simp }
+      with n1 have "prime n"  unfolding prime_def by auto }
+    ultimately have ?thesis using n by blast}
+  ultimately       show ?thesis by auto
+qed
+
+lemma prime_divisor_sqrt:
+  "prime n \<longleftrightarrow> n \<noteq> 1 \<and> (\<forall>d. d dvd n \<and> d^2 \<le> n \<longrightarrow> d = 1)"
+proof-
+  {assume "n=0 \<or> n=1" hence ?thesis using prime_0 prime_1
+    by (auto simp add: nat_power_eq_0_iff)}
+  moreover
+  {assume n: "n\<noteq>0" "n\<noteq>1"
+    hence np: "n > 1" by arith
+    {fix d assume d: "d dvd n" "d^2 \<le> n" and H: "\<forall>m. m dvd n \<longrightarrow> m=1 \<or> m=n"
+      from H d have d1n: "d = 1 \<or> d=n" by blast
+      {assume dn: "d=n"
+	have "n^2 > n*1" using n
+	  by (simp add: power2_eq_square mult_less_cancel1)
+	with dn d(2) have "d=1" by simp}
+      with d1n have "d = 1" by blast  }
+    moreover
+    {fix d assume d: "d dvd n" and H: "\<forall>d'. d' dvd n \<and> d'^2 \<le> n \<longrightarrow> d' = 1"
+      from d n have "d \<noteq> 0" apply - apply (rule ccontr) by simp
+      hence dp: "d > 0" by simp
+      from d[unfolded dvd_def] obtain e where e: "n= d*e" by blast
+      from n dp e have ep:"e > 0" by simp
+      have "d^2 \<le> n \<or> e^2 \<le> n" using dp ep
+	by (auto simp add: e power2_eq_square mult_le_cancel_left)
+      moreover
+      {assume h: "d^2 \<le> n"
+	from H[rule_format, of d] h d have "d = 1" by blast}
+      moreover
+      {assume h: "e^2 \<le> n"
+	from e have "e dvd n" unfolding dvd_def by (simp add: mult_commute)
+	with H[rule_format, of e] h have "e=1" by simp
+	with e have "d = n" by simp}
+      ultimately have "d=1 \<or> d=n"  by blast}
+    ultimately have ?thesis unfolding prime_def using np n(2) by blast}
+  ultimately show ?thesis by auto
+qed
+lemma prime_prime_factor_sqrt:
+  "prime n \<longleftrightarrow> n \<noteq> 0 \<and> n \<noteq> 1 \<and> \<not> (\<exists>p. prime p \<and> p dvd n \<and> p^2 \<le> n)"
+  (is "?lhs \<longleftrightarrow>?rhs")
+proof-
+  {assume "n=0 \<or> n=1" hence ?thesis using prime_0 prime_1 by auto}
+  moreover
+  {assume n: "n\<noteq>0" "n\<noteq>1"
+    {assume H: ?lhs
+      from H[unfolded prime_divisor_sqrt] n
+      have ?rhs  apply clarsimp by (erule_tac x="p" in allE, simp add: prime_1)
+    }
+    moreover
+    {assume H: ?rhs
+      {fix d assume d: "d dvd n" "d^2 \<le> n" "d\<noteq>1"
+	from prime_factor[OF d(3)]
+	obtain p where p: "prime p" "p dvd d" by blast
+	from n have np: "n > 0" by arith
+	from d(1) n have "d \<noteq> 0" by - (rule ccontr, auto)
+	hence dp: "d > 0" by arith
+	from mult_mono[OF dvd_imp_le[OF p(2) dp] dvd_imp_le[OF p(2) dp]] d(2)
+	have "p^2 \<le> n" unfolding power2_eq_square by arith
+	with H n p(1) dvd_trans[OF p(2) d(1)] have False  by blast}
+      with n prime_divisor_sqrt  have ?lhs by auto}
+    ultimately have ?thesis by blast }
+  ultimately show ?thesis by (cases "n=0 \<or> n=1", auto)
+qed
+(* Pocklington theorem. *)
+
+lemma pocklington_lemma:
+  assumes n: "n \<ge> 2" and nqr: "n - 1 = q*r" and an: "[a^ (n - 1) = 1] (mod n)"
+  and aq:"\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a^ ((n - 1) div p) - 1) n"
+  and pp: "prime p" and pn: "p dvd n"
+  shows "[p = 1] (mod q)"
+proof-
+  from pp prime_0 prime_1 have p01: "p \<noteq> 0" "p \<noteq> 1" by - (rule ccontr, simp)+
+  from cong_1_divides[OF an, unfolded nqr, unfolded dvd_def]
+  obtain k where k: "a ^ (q * r) - 1 = n*k" by blast
+  from pn[unfolded dvd_def] obtain l where l: "n = p*l" by blast
+  {assume a0: "a = 0"
+    hence "a^ (n - 1) = 0" using n by (simp add: power_0_left)
+    with n an mod_less[of 1 n]  have False by (simp add: power_0_left modeq_def)}
+  hence a0: "a\<noteq>0" ..
+  from n nqr have aqr0: "a ^ (q * r) \<noteq> 0" using a0 by (simp add: neq0_conv)
+  hence "(a ^ (q * r) - 1) + 1  = a ^ (q * r)" by simp
+  with k l have "a ^ (q * r) = p*l*k + 1" by simp
+  hence "a ^ (r * q) + p * 0 = 1 + p * (l*k)" by (simp add: mult_ac)
+  hence odq: "ord p (a^r) dvd q"
+    unfolding ord_divides[symmetric] power_mult[symmetric] nat_mod  by blast
+  from odq[unfolded dvd_def] obtain d where d: "q = ord p (a^r) * d" by blast
+  {assume d1: "d \<noteq> 1"
+    from prime_factor[OF d1] obtain P where P: "prime P" "P dvd d" by blast
+    from d dvd_mult[OF P(2), of "ord p (a^r)"] have Pq: "P dvd q" by simp
+    from aq P(1) Pq have caP:"coprime (a^ ((n - 1) div P) - 1) n" by blast
+    from Pq obtain s where s: "q = P*s" unfolding dvd_def by blast
+    have P0: "P \<noteq> 0" using P(1) prime_0 by - (rule ccontr, simp)
+    from P(2) obtain t where t: "d = P*t" unfolding dvd_def by blast
+    from d s t P0  have s': "ord p (a^r) * t = s" by algebra
+    have "ord p (a^r) * t*r = r * ord p (a^r) * t" by algebra
+    hence exps: "a^(ord p (a^r) * t*r) = ((a ^ r) ^ ord p (a^r)) ^ t"
+      by (simp only: power_mult)
+    have "[((a ^ r) ^ ord p (a^r)) ^ t= 1^t] (mod p)"
+      by (rule cong_exp, rule ord)
+    then have th: "[((a ^ r) ^ ord p (a^r)) ^ t= 1] (mod p)"
+      by (simp add: power_Suc0)
+    from cong_1_divides[OF th] exps have pd0: "p dvd a^(ord p (a^r) * t*r) - 1" by simp
+    from nqr s s' have "(n - 1) div P = ord p (a^r) * t*r" using P0 by simp
+    with caP have "coprime (a^(ord p (a^r) * t*r) - 1) n" by simp
+    with p01 pn pd0 have False unfolding coprime by auto}
+  hence d1: "d = 1" by blast
+  hence o: "ord p (a^r) = q" using d by simp
+  from pp phi_prime[of p] have phip: " \<phi> p = p - 1" by simp
+  {fix d assume d: "d dvd p" "d dvd a" "d \<noteq> 1"
+    from pp[unfolded prime_def] d have dp: "d = p" by blast
+    from n have n12:"Suc (n - 2) = n - 1" by arith
+    with divides_rexp[OF d(2)[unfolded dp], of "n - 2"]
+    have th0: "p dvd a ^ (n - 1)" by simp
+    from n have n0: "n \<noteq> 0" by simp
+    from d(2) an n12[symmetric] have a0: "a \<noteq> 0"
+      by - (rule ccontr, simp add: modeq_def)
+    have th1: "a^ (n - 1) \<noteq> 0" using n d(2) dp a0 by (auto simp add: neq0_conv)
+    from coprime_minus1[OF th1, unfolded coprime]
+      dvd_trans[OF pn cong_1_divides[OF an]] th0 d(3) dp
+    have False by auto}
+  hence cpa: "coprime p a" using coprime by auto
+  from coprime_exp[OF cpa, of r] coprime_commute
+  have arp: "coprime (a^r) p" by blast
+  from fermat_little[OF arp, simplified ord_divides] o phip
+  have "q dvd (p - 1)" by simp
+  then obtain d where d:"p - 1 = q * d" unfolding dvd_def by blast
+  from prime_0 pp have p0:"p \<noteq> 0" by -  (rule ccontr, auto)
+  from p0 d have "p + q * 0 = 1 + q * d" by simp
+  with nat_mod[of p 1 q, symmetric]
+  show ?thesis by blast
+qed
+
+lemma pocklington:
+  assumes n: "n \<ge> 2" and nqr: "n - 1 = q*r" and sqr: "n \<le> q^2"
+  and an: "[a^ (n - 1) = 1] (mod n)"
+  and aq:"\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a^ ((n - 1) div p) - 1) n"
+  shows "prime n"
+unfolding prime_prime_factor_sqrt[of n]
+proof-
+  let ?ths = "n \<noteq> 0 \<and> n \<noteq> 1 \<and> \<not> (\<exists>p. prime p \<and> p dvd n \<and> p\<twosuperior> \<le> n)"
+  from n have n01: "n\<noteq>0" "n\<noteq>1" by arith+
+  {fix p assume p: "prime p" "p dvd n" "p^2 \<le> n"
+    from p(3) sqr have "p^(Suc 1) \<le> q^(Suc 1)" by (simp add: power2_eq_square)
+    hence pq: "p \<le> q" unfolding exp_mono_le .
+    from pocklington_lemma[OF n nqr an aq p(1,2)]  cong_1_divides
+    have th: "q dvd p - 1" by blast
+    have "p - 1 \<noteq> 0"using prime_ge_2[OF p(1)] by arith
+    with divides_ge[OF th] pq have False by arith }
+  with n01 show ?ths by blast
+qed
+
+(* Variant for application, to separate the exponentiation.                  *)
+lemma pocklington_alt:
+  assumes n: "n \<ge> 2" and nqr: "n - 1 = q*r" and sqr: "n \<le> q^2"
+  and an: "[a^ (n - 1) = 1] (mod n)"
+  and aq:"\<forall>p. prime p \<and> p dvd q \<longrightarrow> (\<exists>b. [a^((n - 1) div p) = b] (mod n) \<and> coprime (b - 1) n)"
+  shows "prime n"
+proof-
+  {fix p assume p: "prime p" "p dvd q"
+    from aq[rule_format] p obtain b where
+      b: "[a^((n - 1) div p) = b] (mod n)" "coprime (b - 1) n" by blast
+    {assume a0: "a=0"
+      from n an have "[0 = 1] (mod n)" unfolding a0 power_0_left by auto
+      hence False using n by (simp add: modeq_def dvd_eq_mod_eq_0[symmetric])}
+    hence a0: "a\<noteq> 0" ..
+    hence a1: "a \<ge> 1" by arith
+    from one_le_power[OF a1] have ath: "1 \<le> a ^ ((n - 1) div p)" .
+    {assume b0: "b = 0"
+      from p(2) nqr have "(n - 1) mod p = 0"
+	apply (simp only: dvd_eq_mod_eq_0[symmetric]) by (rule dvd_mult2, simp)
+      with mod_div_equality[of "n - 1" p]
+      have "(n - 1) div p * p= n - 1" by auto
+      hence eq: "(a^((n - 1) div p))^p = a^(n - 1)"
+	by (simp only: power_mult[symmetric])
+      from prime_ge_2[OF p(1)] have pS: "Suc (p - 1) = p" by arith
+      from b(1) have d: "n dvd a^((n - 1) div p)" unfolding b0 cong_0_divides .
+      from divides_rexp[OF d, of "p - 1"] pS eq cong_divides[OF an] n
+      have False by simp}
+    then have b0: "b \<noteq> 0" ..
+    hence b1: "b \<ge> 1" by arith
+    from cong_coprime[OF cong_sub[OF b(1) cong_refl[of 1] ath b1]] b(2) nqr
+    have "coprime (a ^ ((n - 1) div p) - 1) n" by (simp add: coprime_commute)}
+  hence th: "\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a ^ ((n - 1) div p) - 1) n "
+    by blast
+  from pocklington[OF n nqr sqr an th] show ?thesis .
+qed
+
+(* Prime factorizations.                                                     *)
+
+definition "primefact ps n = (foldr op * ps  1 = n \<and> (\<forall>p\<in> set ps. prime p))"
+
+lemma primefact: assumes n: "n \<noteq> 0"
+  shows "\<exists>ps. primefact ps n"
+using n
+proof(induct n rule: nat_less_induct)
+  fix n assume H: "\<forall>m<n. m \<noteq> 0 \<longrightarrow> (\<exists>ps. primefact ps m)" and n: "n\<noteq>0"
+  let ?ths = "\<exists>ps. primefact ps n"
+  {assume "n = 1"
+    hence "primefact [] n" by (simp add: primefact_def)
+    hence ?ths by blast }
+  moreover
+  {assume n1: "n \<noteq> 1"
+    with n have n2: "n \<ge> 2" by arith
+    from prime_factor[OF n1] obtain p where p: "prime p" "p dvd n" by blast
+    from p(2) obtain m where m: "n = p*m" unfolding dvd_def by blast
+    from n m have m0: "m > 0" "m\<noteq>0" by auto
+    from prime_ge_2[OF p(1)] have "1 < p" by arith
+    with m0 m have mn: "m < n" by auto
+    from H[rule_format, OF mn m0(2)] obtain ps where ps: "primefact ps m" ..
+    from ps m p(1) have "primefact (p#ps) n" by (simp add: primefact_def)
+    hence ?ths by blast}
+  ultimately show ?ths by blast
+qed
+
+lemma primefact_contains:
+  assumes pf: "primefact ps n" and p: "prime p" and pn: "p dvd n"
+  shows "p \<in> set ps"
+  using pf p pn
+proof(induct ps arbitrary: p n)
+  case Nil thus ?case by (auto simp add: primefact_def)
+next
+  case (Cons q qs p n)
+  from Cons.prems[unfolded primefact_def]
+  have q: "prime q" "q * foldr op * qs 1 = n" "\<forall>p \<in>set qs. prime p"  and p: "prime p" "p dvd q * foldr op * qs 1" by simp_all
+  {assume "p dvd q"
+    with p(1) q(1) have "p = q" unfolding prime_def by auto
+    hence ?case by simp}
+  moreover
+  { assume h: "p dvd foldr op * qs 1"
+    from q(3) have pqs: "primefact qs (foldr op * qs 1)"
+      by (simp add: primefact_def)
+    from Cons.hyps[OF pqs p(1) h] have ?case by simp}
+  ultimately show ?case using prime_divprod[OF p] by blast
+qed
+
+lemma primefact_variant: "primefact ps n \<longleftrightarrow> foldr op * ps 1 = n \<and> list_all prime ps" by (auto simp add: primefact_def list_all_iff)
+
+(* Variant of Lucas theorem.                                                 *)
+
+lemma lucas_primefact:
+  assumes n: "n \<ge> 2" and an: "[a^(n - 1) = 1] (mod n)"
+  and psn: "foldr op * ps 1 = n - 1"
+  and psp: "list_all (\<lambda>p. prime p \<and> \<not> [a^((n - 1) div p) = 1] (mod n)) ps"
+  shows "prime n"
+proof-
+  {fix p assume p: "prime p" "p dvd n - 1" "[a ^ ((n - 1) div p) = 1] (mod n)"
+    from psn psp have psn1: "primefact ps (n - 1)"
+      by (auto simp add: list_all_iff primefact_variant)
+    from p(3) primefact_contains[OF psn1 p(1,2)] psp
+    have False by (induct ps, auto)}
+  with lucas[OF n an] show ?thesis by blast
+qed
+
+(* Variant of Pocklington theorem.                                           *)
+
+lemma mod_le: assumes n: "n \<noteq> (0::nat)" shows "m mod n \<le> m"
+proof-
+    from mod_div_equality[of m n]
+    have "\<exists>x. x + m mod n = m" by blast
+    then show ?thesis by auto
+qed
+
+
+lemma pocklington_primefact:
+  assumes n: "n \<ge> 2" and qrn: "q*r = n - 1" and nq2: "n \<le> q^2"
+  and arnb: "(a^r) mod n = b" and psq: "foldr op * ps 1 = q"
+  and bqn: "(b^q) mod n = 1"
+  and psp: "list_all (\<lambda>p. prime p \<and> coprime ((b^(q div p)) mod n - 1) n) ps"
+  shows "prime n"
+proof-
+  from bqn psp qrn
+  have bqn: "a ^ (n - 1) mod n = 1"
+    and psp: "list_all (\<lambda>p. prime p \<and> coprime (a^(r *(q div p)) mod n - 1) n) ps"  unfolding arnb[symmetric] power_mod
+    by (simp_all add: power_mult[symmetric] algebra_simps)
+  from n  have n0: "n > 0" by arith
+  from mod_div_equality[of "a^(n - 1)" n]
+    mod_less_divisor[OF n0, of "a^(n - 1)"]
+  have an1: "[a ^ (n - 1) = 1] (mod n)"
+    unfolding nat_mod bqn
+    apply -
+    apply (rule exI[where x="0"])
+    apply (rule exI[where x="a^(n - 1) div n"])
+    by (simp add: algebra_simps)
+  {fix p assume p: "prime p" "p dvd q"
+    from psp psq have pfpsq: "primefact ps q"
+      by (auto simp add: primefact_variant list_all_iff)
+    from psp primefact_contains[OF pfpsq p]
+    have p': "coprime (a ^ (r * (q div p)) mod n - 1) n"
+      by (simp add: list_all_iff)
+    from prime_ge_2[OF p(1)] have p01: "p \<noteq> 0" "p \<noteq> 1" "p =Suc(p - 1)" by arith+
+    from div_mult1_eq[of r q p] p(2)
+    have eq1: "r* (q div p) = (n - 1) div p"
+      unfolding qrn[symmetric] dvd_eq_mod_eq_0 by (simp add: mult_commute)
+    have ath: "\<And>a (b::nat). a <= b \<Longrightarrow> a \<noteq> 0 ==> 1 <= a \<and> 1 <= b" by arith
+    from n0 have n00: "n \<noteq> 0" by arith
+    from mod_le[OF n00]
+    have th10: "a ^ ((n - 1) div p) mod n \<le> a ^ ((n - 1) div p)" .
+    {assume "a ^ ((n - 1) div p) mod n = 0"
+      then obtain s where s: "a ^ ((n - 1) div p) = n*s"
+	unfolding mod_eq_0_iff by blast
+      hence eq0: "(a^((n - 1) div p))^p = (n*s)^p" by simp
+      from qrn[symmetric] have qn1: "q dvd n - 1" unfolding dvd_def by auto
+      from dvd_trans[OF p(2) qn1] div_mod_equality'[of "n - 1" p]
+      have npp: "(n - 1) div p * p = n - 1" by (simp add: dvd_eq_mod_eq_0)
+      with eq0 have "a^ (n - 1) = (n*s)^p"
+	by (simp add: power_mult[symmetric])
+      hence "1 = (n*s)^(Suc (p - 1)) mod n" using bqn p01 by simp
+      also have "\<dots> = 0" by (simp add: mult_assoc)
+      finally have False by simp }
+      then have th11: "a ^ ((n - 1) div p) mod n \<noteq> 0" by auto
+    have th1: "[a ^ ((n - 1) div p) mod n = a ^ ((n - 1) div p)] (mod n)"
+      unfolding modeq_def by simp
+    from cong_sub[OF th1 cong_refl[of 1]]  ath[OF th10 th11]
+    have th: "[a ^ ((n - 1) div p) mod n - 1 = a ^ ((n - 1) div p) - 1] (mod n)"
+      by blast
+    from cong_coprime[OF th] p'[unfolded eq1]
+    have "coprime (a ^ ((n - 1) div p) - 1) n" by (simp add: coprime_commute) }
+  with pocklington[OF n qrn[symmetric] nq2 an1]
+  show ?thesis by blast
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/Primes.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,826 @@
+(*  Title:      HOL/Library/Primes.thy
+    Author:     Amine Chaieb, Christophe Tabacznyj and Lawrence C Paulson
+    Copyright   1996  University of Cambridge
+*)
+
+header {* Primality on nat *}
+
+theory Primes
+imports Complex_Main Legacy_GCD
+begin
+
+definition
+  coprime :: "nat => nat => bool" where
+  "coprime m n \<longleftrightarrow> gcd m n = 1"
+
+definition
+  prime :: "nat \<Rightarrow> bool" where
+  [code del]: "prime p \<longleftrightarrow> (1 < p \<and> (\<forall>m. m dvd p --> m = 1 \<or> m = p))"
+
+
+lemma two_is_prime: "prime 2"
+  apply (auto simp add: prime_def)
+  apply (case_tac m)
+   apply (auto dest!: dvd_imp_le)
+  done
+
+lemma prime_imp_relprime: "prime p ==> \<not> p dvd n ==> gcd p n = 1"
+  apply (auto simp add: prime_def)
+  apply (metis One_nat_def gcd_dvd1 gcd_dvd2)
+  done
+
+text {*
+  This theorem leads immediately to a proof of the uniqueness of
+  factorization.  If @{term p} divides a product of primes then it is
+  one of those primes.
+*}
+
+lemma prime_dvd_mult: "prime p ==> p dvd m * n ==> p dvd m \<or> p dvd n"
+  by (blast intro: relprime_dvd_mult prime_imp_relprime)
+
+lemma prime_dvd_square: "prime p ==> p dvd m^Suc (Suc 0) ==> p dvd m"
+  by (auto dest: prime_dvd_mult)
+
+lemma prime_dvd_power_two: "prime p ==> p dvd m\<twosuperior> ==> p dvd m"
+  by (rule prime_dvd_square) (simp_all add: power2_eq_square)
+
+
+lemma exp_eq_1:"(x::nat)^n = 1 \<longleftrightarrow> x = 1 \<or> n = 0"
+by (induct n, auto)
+
+lemma exp_mono_lt: "(x::nat) ^ (Suc n) < y ^ (Suc n) \<longleftrightarrow> x < y"
+by(metis linorder_not_less not_less0 power_le_imp_le_base power_less_imp_less_base)
+
+lemma exp_mono_le: "(x::nat) ^ (Suc n) \<le> y ^ (Suc n) \<longleftrightarrow> x \<le> y"
+by (simp only: linorder_not_less[symmetric] exp_mono_lt)
+
+lemma exp_mono_eq: "(x::nat) ^ Suc n = y ^ Suc n \<longleftrightarrow> x = y"
+using power_inject_base[of x n y] by auto
+
+
+lemma even_square: assumes e: "even (n::nat)" shows "\<exists>x. n ^ 2 = 4*x"
+proof-
+  from e have "2 dvd n" by presburger
+  then obtain k where k: "n = 2*k" using dvd_def by auto
+  hence "n^2 = 4* (k^2)" by (simp add: power2_eq_square)
+  thus ?thesis by blast
+qed
+
+lemma odd_square: assumes e: "odd (n::nat)" shows "\<exists>x. n ^ 2 = 4*x + 1"
+proof-
+  from e have np: "n > 0" by presburger
+  from e have "2 dvd (n - 1)" by presburger
+  then obtain k where "n - 1 = 2*k" using dvd_def by auto
+  hence k: "n = 2*k + 1"  using e by presburger 
+  hence "n^2 = 4* (k^2 + k) + 1" by algebra   
+  thus ?thesis by blast
+qed
+
+lemma diff_square: "(x::nat)^2 - y^2 = (x+y)*(x - y)" 
+proof-
+  have "x \<le> y \<or> y \<le> x" by (rule nat_le_linear)
+  moreover
+  {assume le: "x \<le> y"
+    hence "x ^2 \<le> y^2" by (simp only: numeral_2_eq_2 exp_mono_le Let_def)
+    with le have ?thesis by simp }
+  moreover
+  {assume le: "y \<le> x"
+    hence le2: "y ^2 \<le> x^2" by (simp only: numeral_2_eq_2 exp_mono_le Let_def)
+    from le have "\<exists>z. y + z = x" by presburger
+    then obtain z where z: "x = y + z" by blast 
+    from le2 have "\<exists>z. x^2 = y^2 + z" by presburger
+    then obtain z2 where z2: "x^2 = y^2 + z2"  by blast
+    from z z2 have ?thesis apply simp by algebra }
+  ultimately show ?thesis by blast  
+qed
+
+text {* Elementary theory of divisibility *}
+lemma divides_ge: "(a::nat) dvd b \<Longrightarrow> b = 0 \<or> a \<le> b" unfolding dvd_def by auto
+lemma divides_antisym: "(x::nat) dvd y \<and> y dvd x \<longleftrightarrow> x = y"
+  using dvd_anti_sym[of x y] by auto
+
+lemma divides_add_revr: assumes da: "(d::nat) dvd a" and dab:"d dvd (a + b)"
+  shows "d dvd b"
+proof-
+  from da obtain k where k:"a = d*k" by (auto simp add: dvd_def)
+  from dab obtain k' where k': "a + b = d*k'" by (auto simp add: dvd_def)
+  from k k' have "b = d *(k' - k)" by (simp add : diff_mult_distrib2)
+  thus ?thesis unfolding dvd_def by blast
+qed
+
+declare nat_mult_dvd_cancel_disj[presburger]
+lemma nat_mult_dvd_cancel_disj'[presburger]: 
+  "(m\<Colon>nat)*k dvd n*k \<longleftrightarrow> k = 0 \<or> m dvd n" unfolding mult_commute[of m k] mult_commute[of n k] by presburger 
+
+lemma divides_mul_l: "(a::nat) dvd b ==> (c * a) dvd (c * b)"
+  by presburger
+
+lemma divides_mul_r: "(a::nat) dvd b ==> (a * c) dvd (b * c)" by presburger
+lemma divides_cases: "(n::nat) dvd m ==> m = 0 \<or> m = n \<or> 2 * n <= m" 
+  by (auto simp add: dvd_def)
+
+lemma divides_div_not: "(x::nat) = (q * n) + r \<Longrightarrow> 0 < r \<Longrightarrow> r < n ==> ~(n dvd x)"
+proof(auto simp add: dvd_def)
+  fix k assume H: "0 < r" "r < n" "q * n + r = n * k"
+  from H(3) have r: "r = n* (k -q)" by(simp add: diff_mult_distrib2 mult_commute)
+  {assume "k - q = 0" with r H(1) have False by simp}
+  moreover
+  {assume "k - q \<noteq> 0" with r have "r \<ge> n" by auto
+    with H(2) have False by simp}
+  ultimately show False by blast
+qed
+lemma divides_exp: "(x::nat) dvd y ==> x ^ n dvd y ^ n"
+  by (auto simp add: power_mult_distrib dvd_def)
+
+lemma divides_exp2: "n \<noteq> 0 \<Longrightarrow> (x::nat) ^ n dvd y \<Longrightarrow> x dvd y" 
+  by (induct n ,auto simp add: dvd_def)
+
+fun fact :: "nat \<Rightarrow> nat" where
+  "fact 0 = 1"
+| "fact (Suc n) = Suc n * fact n"	
+
+lemma fact_lt: "0 < fact n" by(induct n, simp_all)
+lemma fact_le: "fact n \<ge> 1" using fact_lt[of n] by simp 
+lemma fact_mono: assumes le: "m \<le> n" shows "fact m \<le> fact n"
+proof-
+  from le have "\<exists>i. n = m+i" by presburger
+  then obtain i where i: "n = m+i" by blast 
+  have "fact m \<le> fact (m + i)"
+  proof(induct m)
+    case 0 thus ?case using fact_le[of i] by simp
+  next
+    case (Suc m)
+    have "fact (Suc m) = Suc m * fact m" by simp
+    have th1: "Suc m \<le> Suc (m + i)" by simp
+    from mult_le_mono[of "Suc m" "Suc (m+i)" "fact m" "fact (m+i)", OF th1 Suc.hyps]
+    show ?case by simp
+  qed
+  thus ?thesis using i by simp
+qed
+
+lemma divides_fact: "1 <= p \<Longrightarrow> p <= n ==> p dvd fact n"
+proof(induct n arbitrary: p)
+  case 0 thus ?case by simp
+next
+  case (Suc n p)
+  from Suc.prems have "p = Suc n \<or> p \<le> n" by presburger 
+  moreover
+  {assume "p = Suc n" hence ?case  by (simp only: fact.simps dvd_triv_left)}
+  moreover
+  {assume "p \<le> n"
+    with Suc.prems(1) Suc.hyps have th: "p dvd fact n" by simp
+    from dvd_mult[OF th] have ?case by (simp only: fact.simps) }
+  ultimately show ?case by blast
+qed
+
+declare dvd_triv_left[presburger]
+declare dvd_triv_right[presburger]
+lemma divides_rexp: 
+  "x dvd y \<Longrightarrow> (x::nat) dvd (y^(Suc n))" by (simp add: dvd_mult2[of x y])
+
+text {* Coprimality *}
+
+lemma coprime: "coprime a b \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> d = 1)"
+using gcd_unique[of 1 a b, simplified] by (auto simp add: coprime_def)
+lemma coprime_commute: "coprime a b \<longleftrightarrow> coprime b a" by (simp add: coprime_def gcd_commute)
+
+lemma coprime_bezout: "coprime a b \<longleftrightarrow> (\<exists>x y. a * x - b * y = 1 \<or> b * x - a * y = 1)"
+using coprime_def gcd_bezout by auto
+
+lemma coprime_divprod: "d dvd a * b  \<Longrightarrow> coprime d a \<Longrightarrow> d dvd b"
+  using relprime_dvd_mult_iff[of d a b] by (auto simp add: coprime_def mult_commute)
+
+lemma coprime_1[simp]: "coprime a 1" by (simp add: coprime_def)
+lemma coprime_1'[simp]: "coprime 1 a" by (simp add: coprime_def)
+lemma coprime_Suc0[simp]: "coprime a (Suc 0)" by (simp add: coprime_def)
+lemma coprime_Suc0'[simp]: "coprime (Suc 0) a" by (simp add: coprime_def)
+
+lemma gcd_coprime: 
+  assumes z: "gcd a b \<noteq> 0" and a: "a = a' * gcd a b" and b: "b = b' * gcd a b" 
+  shows    "coprime a' b'"
+proof-
+  let ?g = "gcd a b"
+  {assume bz: "a = 0" from b bz z a have ?thesis by (simp add: gcd_zero coprime_def)}
+  moreover 
+  {assume az: "a\<noteq> 0" 
+    from z have z': "?g > 0" by simp
+    from bezout_gcd_strong[OF az, of b] 
+    obtain x y where xy: "a*x = b*y + ?g" by blast
+    from xy a b have "?g * a'*x = ?g * (b'*y + 1)" by (simp add: algebra_simps)
+    hence "?g * (a'*x) = ?g * (b'*y + 1)" by (simp add: mult_assoc)
+    hence "a'*x = (b'*y + 1)"
+      by (simp only: nat_mult_eq_cancel1[OF z']) 
+    hence "a'*x - b'*y = 1" by simp
+    with coprime_bezout[of a' b'] have ?thesis by auto}
+  ultimately show ?thesis by blast
+qed
+lemma coprime_0: "coprime d 0 \<longleftrightarrow> d = 1" by (simp add: coprime_def)
+lemma coprime_mul: assumes da: "coprime d a" and db: "coprime d b"
+  shows "coprime d (a * b)"
+proof-
+  from da have th: "gcd a d = 1" by (simp add: coprime_def gcd_commute)
+  from gcd_mult_cancel[of a d b, OF th] db[unfolded coprime_def] have "gcd d (a*b) = 1"
+    by (simp add: gcd_commute)
+  thus ?thesis unfolding coprime_def .
+qed
+lemma coprime_lmul2: assumes dab: "coprime d (a * b)" shows "coprime d b"
+using prems unfolding coprime_bezout
+apply clarsimp
+apply (case_tac "d * x - a * b * y = Suc 0 ", simp_all)
+apply (rule_tac x="x" in exI)
+apply (rule_tac x="a*y" in exI)
+apply (simp add: mult_ac)
+apply (rule_tac x="a*x" in exI)
+apply (rule_tac x="y" in exI)
+apply (simp add: mult_ac)
+done
+
+lemma coprime_rmul2: "coprime d (a * b) \<Longrightarrow> coprime d a"
+unfolding coprime_bezout
+apply clarsimp
+apply (case_tac "d * x - a * b * y = Suc 0 ", simp_all)
+apply (rule_tac x="x" in exI)
+apply (rule_tac x="b*y" in exI)
+apply (simp add: mult_ac)
+apply (rule_tac x="b*x" in exI)
+apply (rule_tac x="y" in exI)
+apply (simp add: mult_ac)
+done
+lemma coprime_mul_eq: "coprime d (a * b) \<longleftrightarrow> coprime d a \<and>  coprime d b"
+  using coprime_rmul2[of d a b] coprime_lmul2[of d a b] coprime_mul[of d a b] 
+  by blast
+
+lemma gcd_coprime_exists:
+  assumes nz: "gcd a b \<noteq> 0" 
+  shows "\<exists>a' b'. a = a' * gcd a b \<and> b = b' * gcd a b \<and> coprime a' b'"
+proof-
+  let ?g = "gcd a b"
+  from gcd_dvd1[of a b] gcd_dvd2[of a b] 
+  obtain a' b' where "a = ?g*a'"  "b = ?g*b'" unfolding dvd_def by blast
+  hence ab': "a = a'*?g" "b = b'*?g" by algebra+
+  from ab' gcd_coprime[OF nz ab'] show ?thesis by blast
+qed
+
+lemma coprime_exp: "coprime d a ==> coprime d (a^n)" 
+  by(induct n, simp_all add: coprime_mul)
+
+lemma coprime_exp_imp: "coprime a b ==> coprime (a ^n) (b ^n)"
+  by (induct n, simp_all add: coprime_mul_eq coprime_commute coprime_exp)
+lemma coprime_refl[simp]: "coprime n n \<longleftrightarrow> n = 1" by (simp add: coprime_def)
+lemma coprime_plus1[simp]: "coprime (n + 1) n"
+  apply (simp add: coprime_bezout)
+  apply (rule exI[where x=1])
+  apply (rule exI[where x=1])
+  apply simp
+  done
+lemma coprime_minus1: "n \<noteq> 0 ==> coprime (n - 1) n"
+  using coprime_plus1[of "n - 1"] coprime_commute[of "n - 1" n] by auto
+
+lemma bezout_gcd_pow: "\<exists>x y. a ^n * x - b ^ n * y = gcd a b ^ n \<or> b ^ n * x - a ^ n * y = gcd a b ^ n"
+proof-
+  let ?g = "gcd a b"
+  {assume z: "?g = 0" hence ?thesis 
+      apply (cases n, simp)
+      apply arith
+      apply (simp only: z power_0_Suc)
+      apply (rule exI[where x=0])
+      apply (rule exI[where x=0])
+      by simp}
+  moreover
+  {assume z: "?g \<noteq> 0"
+    from gcd_dvd1[of a b] gcd_dvd2[of a b] obtain a' b' where
+      ab': "a = a'*?g" "b = b'*?g" unfolding dvd_def by (auto simp add: mult_ac)
+    hence ab'': "?g*a' = a" "?g * b' = b" by algebra+
+    from coprime_exp_imp[OF gcd_coprime[OF z ab'], unfolded coprime_bezout, of n]
+    obtain x y where "a'^n * x - b'^n * y = 1 \<or> b'^n * x - a'^n * y = 1"  by blast
+    hence "?g^n * (a'^n * x - b'^n * y) = ?g^n \<or> ?g^n*(b'^n * x - a'^n * y) = ?g^n"
+      using z by auto 
+    then have "a^n * x - b^n * y = ?g^n \<or> b^n * x - a^n * y = ?g^n"
+      using z ab'' by (simp only: power_mult_distrib[symmetric] 
+	diff_mult_distrib2 mult_assoc[symmetric])
+    hence  ?thesis by blast }
+  ultimately show ?thesis by blast
+qed
+
+lemma gcd_exp: "gcd (a^n) (b^n) = gcd a b^n"
+proof-
+  let ?g = "gcd (a^n) (b^n)"
+  let ?gn = "gcd a b^n"
+  {fix e assume H: "e dvd a^n" "e dvd b^n"
+    from bezout_gcd_pow[of a n b] obtain x y 
+      where xy: "a ^ n * x - b ^ n * y = ?gn \<or> b ^ n * x - a ^ n * y = ?gn" by blast
+    from dvd_diff_nat [OF dvd_mult2[OF H(1), of x] dvd_mult2[OF H(2), of y]]
+      dvd_diff_nat [OF dvd_mult2[OF H(2), of x] dvd_mult2[OF H(1), of y]] xy
+    have "e dvd ?gn" by (cases "a ^ n * x - b ^ n * y = gcd a b ^ n", simp_all)}
+  hence th:  "\<forall>e. e dvd a^n \<and> e dvd b^n \<longrightarrow> e dvd ?gn" by blast
+  from divides_exp[OF gcd_dvd1[of a b], of n] divides_exp[OF gcd_dvd2[of a b], of n] th
+    gcd_unique have "?gn = ?g" by blast thus ?thesis by simp 
+qed
+
+lemma coprime_exp2:  "coprime (a ^ Suc n) (b^ Suc n) \<longleftrightarrow> coprime a b"
+by (simp only: coprime_def gcd_exp exp_eq_1) simp
+
+lemma division_decomp: assumes dc: "(a::nat) dvd b * c"
+  shows "\<exists>b' c'. a = b' * c' \<and> b' dvd b \<and> c' dvd c"
+proof-
+  let ?g = "gcd a b"
+  {assume "?g = 0" with dc have ?thesis apply (simp add: gcd_zero)
+      apply (rule exI[where x="0"])
+      by (rule exI[where x="c"], simp)}
+  moreover
+  {assume z: "?g \<noteq> 0"
+    from gcd_coprime_exists[OF z]
+    obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'" by blast
+    from gcd_dvd2[of a b] have thb: "?g dvd b" .
+    from ab'(1) have "a' dvd a"  unfolding dvd_def by blast  
+    with dc have th0: "a' dvd b*c" using dvd_trans[of a' a "b*c"] by simp
+    from dc ab'(1,2) have "a'*?g dvd (b'*?g) *c" by auto
+    hence "?g*a' dvd ?g * (b' * c)" by (simp add: mult_assoc)
+    with z have th_1: "a' dvd b'*c" by simp
+    from coprime_divprod[OF th_1 ab'(3)] have thc: "a' dvd c" . 
+    from ab' have "a = ?g*a'" by algebra
+    with thb thc have ?thesis by blast }
+  ultimately show ?thesis by blast
+qed
+
+lemma nat_power_eq_0_iff: "(m::nat) ^ n = 0 \<longleftrightarrow> n \<noteq> 0 \<and> m = 0" by (induct n, auto)
+
+lemma divides_rev: assumes ab: "(a::nat) ^ n dvd b ^n" and n:"n \<noteq> 0" shows "a dvd b"
+proof-
+  let ?g = "gcd a b"
+  from n obtain m where m: "n = Suc m" by (cases n, simp_all)
+  {assume "?g = 0" with ab n have ?thesis by (simp add: gcd_zero)}
+  moreover
+  {assume z: "?g \<noteq> 0"
+    hence zn: "?g ^ n \<noteq> 0" using n by (simp add: neq0_conv)
+    from gcd_coprime_exists[OF z] 
+    obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'" by blast
+    from ab have "(a' * ?g) ^ n dvd (b' * ?g)^n" by (simp add: ab'(1,2)[symmetric])
+    hence "?g^n*a'^n dvd ?g^n *b'^n" by (simp only: power_mult_distrib mult_commute)
+    with zn z n have th0:"a'^n dvd b'^n" by (auto simp add: nat_power_eq_0_iff)
+    have "a' dvd a'^n" by (simp add: m)
+    with th0 have "a' dvd b'^n" using dvd_trans[of a' "a'^n" "b'^n"] by simp
+    hence th1: "a' dvd b'^m * b'" by (simp add: m mult_commute)
+    from coprime_divprod[OF th1 coprime_exp[OF ab'(3), of m]]
+    have "a' dvd b'" .
+    hence "a'*?g dvd b'*?g" by simp
+    with ab'(1,2)  have ?thesis by simp }
+  ultimately show ?thesis by blast
+qed
+
+lemma divides_mul: assumes mr: "m dvd r" and nr: "n dvd r" and mn:"coprime m n" 
+  shows "m * n dvd r"
+proof-
+  from mr nr obtain m' n' where m': "r = m*m'" and n': "r = n*n'"
+    unfolding dvd_def by blast
+  from mr n' have "m dvd n'*n" by (simp add: mult_commute)
+  hence "m dvd n'" using relprime_dvd_mult_iff[OF mn[unfolded coprime_def]] by simp
+  then obtain k where k: "n' = m*k" unfolding dvd_def by blast
+  from n' k show ?thesis unfolding dvd_def by auto
+qed
+
+
+text {* A binary form of the Chinese Remainder Theorem. *}
+
+lemma chinese_remainder: assumes ab: "coprime a b" and a:"a \<noteq> 0" and b:"b \<noteq> 0"
+  shows "\<exists>x q1 q2. x = u + q1 * a \<and> x = v + q2 * b"
+proof-
+  from bezout_add_strong[OF a, of b] bezout_add_strong[OF b, of a]
+  obtain d1 x1 y1 d2 x2 y2 where dxy1: "d1 dvd a" "d1 dvd b" "a * x1 = b * y1 + d1" 
+    and dxy2: "d2 dvd b" "d2 dvd a" "b * x2 = a * y2 + d2" by blast
+  from gcd_unique[of 1 a b, simplified ab[unfolded coprime_def], simplified] 
+    dxy1(1,2) dxy2(1,2) have d12: "d1 = 1" "d2 =1" by auto
+  let ?x = "v * a * x1 + u * b * x2"
+  let ?q1 = "v * x1 + u * y2"
+  let ?q2 = "v * y1 + u * x2"
+  from dxy2(3)[simplified d12] dxy1(3)[simplified d12] 
+  have "?x = u + ?q1 * a" "?x = v + ?q2 * b" by algebra+ 
+  thus ?thesis by blast
+qed
+
+text {* Primality *}
+
+text {* A few useful theorems about primes *}
+
+lemma prime_0[simp]: "~prime 0" by (simp add: prime_def)
+lemma prime_1[simp]: "~ prime 1"  by (simp add: prime_def)
+lemma prime_Suc0[simp]: "~ prime (Suc 0)"  by (simp add: prime_def)
+
+lemma prime_ge_2: "prime p ==> p \<ge> 2" by (simp add: prime_def)
+lemma prime_factor: assumes n: "n \<noteq> 1" shows "\<exists> p. prime p \<and> p dvd n"
+using n
+proof(induct n rule: nat_less_induct)
+  fix n
+  assume H: "\<forall>m<n. m \<noteq> 1 \<longrightarrow> (\<exists>p. prime p \<and> p dvd m)" "n \<noteq> 1"
+  let ?ths = "\<exists>p. prime p \<and> p dvd n"
+  {assume "n=0" hence ?ths using two_is_prime by auto}
+  moreover
+  {assume nz: "n\<noteq>0" 
+    {assume "prime n" hence ?ths by - (rule exI[where x="n"], simp)}
+    moreover
+    {assume n: "\<not> prime n"
+      with nz H(2) 
+      obtain k where k:"k dvd n" "k \<noteq> 1" "k \<noteq> n" by (auto simp add: prime_def) 
+      from dvd_imp_le[OF k(1)] nz k(3) have kn: "k < n" by simp
+      from H(1)[rule_format, OF kn k(2)] obtain p where p: "prime p" "p dvd k" by blast
+      from dvd_trans[OF p(2) k(1)] p(1) have ?ths by blast}
+    ultimately have ?ths by blast}
+  ultimately show ?ths by blast
+qed
+
+lemma prime_factor_lt: assumes p: "prime p" and n: "n \<noteq> 0" and npm:"n = p * m"
+  shows "m < n"
+proof-
+  {assume "m=0" with n have ?thesis by simp}
+  moreover
+  {assume m: "m \<noteq> 0"
+    from npm have mn: "m dvd n" unfolding dvd_def by auto
+    from npm m have "n \<noteq> m" using p by auto
+    with dvd_imp_le[OF mn] n have ?thesis by simp}
+  ultimately show ?thesis by blast
+qed
+
+lemma euclid_bound: "\<exists>p. prime p \<and> n < p \<and>  p <= Suc (fact n)"
+proof-
+  have f1: "fact n + 1 \<noteq> 1" using fact_le[of n] by arith 
+  from prime_factor[OF f1] obtain p where p: "prime p" "p dvd fact n + 1" by blast
+  from dvd_imp_le[OF p(2)] have pfn: "p \<le> fact n + 1" by simp
+  {assume np: "p \<le> n"
+    from p(1) have p1: "p \<ge> 1" by (cases p, simp_all)
+    from divides_fact[OF p1 np] have pfn': "p dvd fact n" .
+    from divides_add_revr[OF pfn' p(2)] p(1) have False by simp}
+  hence "n < p" by arith
+  with p(1) pfn show ?thesis by auto
+qed
+
+lemma euclid: "\<exists>p. prime p \<and> p > n" using euclid_bound by auto
+
+lemma primes_infinite: "\<not> (finite {p. prime p})"
+apply(simp add: finite_nat_set_iff_bounded_le)
+apply (metis euclid linorder_not_le)
+done
+
+lemma coprime_prime: assumes ab: "coprime a b"
+  shows "~(prime p \<and> p dvd a \<and> p dvd b)"
+proof
+  assume "prime p \<and> p dvd a \<and> p dvd b"
+  thus False using ab gcd_greatest[of p a b] by (simp add: coprime_def)
+qed
+lemma coprime_prime_eq: "coprime a b \<longleftrightarrow> (\<forall>p. ~(prime p \<and> p dvd a \<and> p dvd b))" 
+  (is "?lhs = ?rhs")
+proof-
+  {assume "?lhs" with coprime_prime  have ?rhs by blast}
+  moreover
+  {assume r: "?rhs" and c: "\<not> ?lhs"
+    then obtain g where g: "g\<noteq>1" "g dvd a" "g dvd b" unfolding coprime_def by blast
+    from prime_factor[OF g(1)] obtain p where p: "prime p" "p dvd g" by blast
+    from dvd_trans [OF p(2) g(2)] dvd_trans [OF p(2) g(3)] 
+    have "p dvd a" "p dvd b" . with p(1) r have False by blast}
+  ultimately show ?thesis by blast
+qed
+
+lemma prime_coprime: assumes p: "prime p" 
+  shows "n = 1 \<or> p dvd n \<or> coprime p n"
+using p prime_imp_relprime[of p n] by (auto simp add: coprime_def)
+
+lemma prime_coprime_strong: "prime p \<Longrightarrow> p dvd n \<or> coprime p n"
+  using prime_coprime[of p n] by auto
+
+declare  coprime_0[simp]
+
+lemma coprime_0'[simp]: "coprime 0 d \<longleftrightarrow> d = 1" by (simp add: coprime_commute[of 0 d])
+lemma coprime_bezout_strong: assumes ab: "coprime a b" and b: "b \<noteq> 1"
+  shows "\<exists>x y. a * x = b * y + 1"
+proof-
+  from ab b have az: "a \<noteq> 0" by - (rule ccontr, auto)
+  from bezout_gcd_strong[OF az, of b] ab[unfolded coprime_def]
+  show ?thesis by auto
+qed
+
+lemma bezout_prime: assumes p: "prime p"  and pa: "\<not> p dvd a"
+  shows "\<exists>x y. a*x = p*y + 1"
+proof-
+  from p have p1: "p \<noteq> 1" using prime_1 by blast 
+  from prime_coprime[OF p, of a] p1 pa have ap: "coprime a p" 
+    by (auto simp add: coprime_commute)
+  from coprime_bezout_strong[OF ap p1] show ?thesis . 
+qed
+lemma prime_divprod: assumes p: "prime p" and pab: "p dvd a*b"
+  shows "p dvd a \<or> p dvd b"
+proof-
+  {assume "a=1" hence ?thesis using pab by simp }
+  moreover
+  {assume "p dvd a" hence ?thesis by blast}
+  moreover
+  {assume pa: "coprime p a" from coprime_divprod[OF pab pa]  have ?thesis .. }
+  ultimately show ?thesis using prime_coprime[OF p, of a] by blast
+qed
+
+lemma prime_divprod_eq: assumes p: "prime p"
+  shows "p dvd a*b \<longleftrightarrow> p dvd a \<or> p dvd b"
+using p prime_divprod dvd_mult dvd_mult2 by auto
+
+lemma prime_divexp: assumes p:"prime p" and px: "p dvd x^n"
+  shows "p dvd x"
+using px
+proof(induct n)
+  case 0 thus ?case by simp
+next
+  case (Suc n) 
+  hence th: "p dvd x*x^n" by simp
+  {assume H: "p dvd x^n"
+    from Suc.hyps[OF H] have ?case .}
+  with prime_divprod[OF p th] show ?case by blast
+qed
+
+lemma prime_divexp_n: "prime p \<Longrightarrow> p dvd x^n \<Longrightarrow> p^n dvd x^n"
+  using prime_divexp[of p x n] divides_exp[of p x n] by blast
+
+lemma coprime_prime_dvd_ex: assumes xy: "\<not>coprime x y"
+  shows "\<exists>p. prime p \<and> p dvd x \<and> p dvd y"
+proof-
+  from xy[unfolded coprime_def] obtain g where g: "g \<noteq> 1" "g dvd x" "g dvd y" 
+    by blast
+  from prime_factor[OF g(1)] obtain p where p: "prime p" "p dvd g" by blast
+  from g(2,3) dvd_trans[OF p(2)] p(1) show ?thesis by auto
+qed
+lemma coprime_sos: assumes xy: "coprime x y" 
+  shows "coprime (x * y) (x^2 + y^2)"
+proof-
+  {assume c: "\<not> coprime (x * y) (x^2 + y^2)"
+    from coprime_prime_dvd_ex[OF c] obtain p 
+      where p: "prime p" "p dvd x*y" "p dvd x^2 + y^2" by blast
+    {assume px: "p dvd x"
+      from dvd_mult[OF px, of x] p(3) 
+        obtain r s where "x * x = p * r" and "x^2 + y^2 = p * s"
+          by (auto elim!: dvdE)
+        then have "y^2 = p * (s - r)" 
+          by (auto simp add: power2_eq_square diff_mult_distrib2)
+        then have "p dvd y^2" ..
+      with prime_divexp[OF p(1), of y 2] have py: "p dvd y" .
+      from p(1) px py xy[unfolded coprime, rule_format, of p] prime_1  
+      have False by simp }
+    moreover
+    {assume py: "p dvd y"
+      from dvd_mult[OF py, of y] p(3)
+        obtain r s where "y * y = p * r" and "x^2 + y^2 = p * s"
+          by (auto elim!: dvdE)
+        then have "x^2 = p * (s - r)" 
+          by (auto simp add: power2_eq_square diff_mult_distrib2)
+        then have "p dvd x^2" ..
+      with prime_divexp[OF p(1), of x 2] have px: "p dvd x" .
+      from p(1) px py xy[unfolded coprime, rule_format, of p] prime_1  
+      have False by simp }
+    ultimately have False using prime_divprod[OF p(1,2)] by blast}
+  thus ?thesis by blast
+qed
+
+lemma distinct_prime_coprime: "prime p \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"
+  unfolding prime_def coprime_prime_eq by blast
+
+lemma prime_coprime_lt: assumes p: "prime p" and x: "0 < x" and xp: "x < p"
+  shows "coprime x p"
+proof-
+  {assume c: "\<not> coprime x p"
+    then obtain g where g: "g \<noteq> 1" "g dvd x" "g dvd p" unfolding coprime_def by blast
+  from dvd_imp_le[OF g(2)] x xp have gp: "g < p" by arith
+  from g(2) x have "g \<noteq> 0" by - (rule ccontr, simp)
+  with g gp p[unfolded prime_def] have False by blast}
+thus ?thesis by blast
+qed
+
+lemma even_dvd[simp]: "even (n::nat) \<longleftrightarrow> 2 dvd n" by presburger
+lemma prime_odd: "prime p \<Longrightarrow> p = 2 \<or> odd p" unfolding prime_def by auto
+
+
+text {* One property of coprimality is easier to prove via prime factors. *}
+
+lemma prime_divprod_pow: 
+  assumes p: "prime p" and ab: "coprime a b" and pab: "p^n dvd a * b"
+  shows "p^n dvd a \<or> p^n dvd b"
+proof-
+  {assume "n = 0 \<or> a = 1 \<or> b = 1" with pab have ?thesis 
+      apply (cases "n=0", simp_all)
+      apply (cases "a=1", simp_all) done}
+  moreover
+  {assume n: "n \<noteq> 0" and a: "a\<noteq>1" and b: "b\<noteq>1" 
+    then obtain m where m: "n = Suc m" by (cases n, auto)
+    from divides_exp2[OF n pab] have pab': "p dvd a*b" .
+    from prime_divprod[OF p pab'] 
+    have "p dvd a \<or> p dvd b" .
+    moreover
+    {assume pa: "p dvd a"
+      have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)
+      from coprime_prime[OF ab, of p] p pa have "\<not> p dvd b" by blast
+      with prime_coprime[OF p, of b] b 
+      have cpb: "coprime b p" using coprime_commute by blast 
+      from coprime_exp[OF cpb] have pnb: "coprime (p^n) b" 
+	by (simp add: coprime_commute)
+      from coprime_divprod[OF pnba pnb] have ?thesis by blast }
+    moreover
+    {assume pb: "p dvd b"
+      have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)
+      from coprime_prime[OF ab, of p] p pb have "\<not> p dvd a" by blast
+      with prime_coprime[OF p, of a] a
+      have cpb: "coprime a p" using coprime_commute by blast 
+      from coprime_exp[OF cpb] have pnb: "coprime (p^n) a" 
+	by (simp add: coprime_commute)
+      from coprime_divprod[OF pab pnb] have ?thesis by blast }
+    ultimately have ?thesis by blast}
+  ultimately show ?thesis by blast
+qed
+
+lemma nat_mult_eq_one: "(n::nat) * m = 1 \<longleftrightarrow> n = 1 \<and> m = 1" (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume H: "?lhs"
+  hence "n dvd 1" "m dvd 1" unfolding dvd_def by (auto simp add: mult_commute)
+  thus ?rhs by auto
+next
+  assume ?rhs then show ?lhs by auto
+qed
+  
+lemma power_Suc0[simp]: "Suc 0 ^ n = Suc 0" 
+  unfolding One_nat_def[symmetric] power_one ..
+lemma coprime_pow: assumes ab: "coprime a b" and abcn: "a * b = c ^n"
+  shows "\<exists>r s. a = r^n  \<and> b = s ^n"
+  using ab abcn
+proof(induct c arbitrary: a b rule: nat_less_induct)
+  fix c a b
+  assume H: "\<forall>m<c. \<forall>a b. coprime a b \<longrightarrow> a * b = m ^ n \<longrightarrow> (\<exists>r s. a = r ^ n \<and> b = s ^ n)" "coprime a b" "a * b = c ^ n" 
+  let ?ths = "\<exists>r s. a = r^n  \<and> b = s ^n"
+  {assume n: "n = 0"
+    with H(3) power_one have "a*b = 1" by simp
+    hence "a = 1 \<and> b = 1" by simp
+    hence ?ths 
+      apply -
+      apply (rule exI[where x=1])
+      apply (rule exI[where x=1])
+      using power_one[of  n]
+      by simp}
+  moreover
+  {assume n: "n \<noteq> 0" then obtain m where m: "n = Suc m" by (cases n, auto)
+    {assume c: "c = 0"
+      with H(3) m H(2) have ?ths apply simp 
+	apply (cases "a=0", simp_all) 
+	apply (rule exI[where x="0"], simp)
+	apply (rule exI[where x="0"], simp)
+	done}
+    moreover
+    {assume "c=1" with H(3) power_one have "a*b = 1" by simp 
+	hence "a = 1 \<and> b = 1" by simp
+	hence ?ths 
+      apply -
+      apply (rule exI[where x=1])
+      apply (rule exI[where x=1])
+      using power_one[of  n]
+      by simp}
+  moreover
+  {assume c: "c\<noteq>1" "c \<noteq> 0"
+    from prime_factor[OF c(1)] obtain p where p: "prime p" "p dvd c" by blast
+    from prime_divprod_pow[OF p(1) H(2), unfolded H(3), OF divides_exp[OF p(2), of n]] 
+    have pnab: "p ^ n dvd a \<or> p^n dvd b" . 
+    from p(2) obtain l where l: "c = p*l" unfolding dvd_def by blast
+    have pn0: "p^n \<noteq> 0" using n prime_ge_2 [OF p(1)] by (simp add: neq0_conv)
+    {assume pa: "p^n dvd a"
+      then obtain k where k: "a = p^n * k" unfolding dvd_def by blast
+      from l have "l dvd c" by auto
+      with dvd_imp_le[of l c] c have "l \<le> c" by auto
+      moreover {assume "l = c" with l c  have "p = 1" by simp with p have False by simp}
+      ultimately have lc: "l < c" by arith
+      from coprime_lmul2 [OF H(2)[unfolded k coprime_commute[of "p^n*k" b]]]
+      have kb: "coprime k b" by (simp add: coprime_commute) 
+      from H(3) l k pn0 have kbln: "k * b = l ^ n" 
+	by (auto simp add: power_mult_distrib)
+      from H(1)[rule_format, OF lc kb kbln]
+      obtain r s where rs: "k = r ^n" "b = s^n" by blast
+      from k rs(1) have "a = (p*r)^n" by (simp add: power_mult_distrib)
+      with rs(2) have ?ths by blast }
+    moreover
+    {assume pb: "p^n dvd b"
+      then obtain k where k: "b = p^n * k" unfolding dvd_def by blast
+      from l have "l dvd c" by auto
+      with dvd_imp_le[of l c] c have "l \<le> c" by auto
+      moreover {assume "l = c" with l c  have "p = 1" by simp with p have False by simp}
+      ultimately have lc: "l < c" by arith
+      from coprime_lmul2 [OF H(2)[unfolded k coprime_commute[of "p^n*k" a]]]
+      have kb: "coprime k a" by (simp add: coprime_commute) 
+      from H(3) l k pn0 n have kbln: "k * a = l ^ n" 
+	by (simp add: power_mult_distrib mult_commute)
+      from H(1)[rule_format, OF lc kb kbln]
+      obtain r s where rs: "k = r ^n" "a = s^n" by blast
+      from k rs(1) have "b = (p*r)^n" by (simp add: power_mult_distrib)
+      with rs(2) have ?ths by blast }
+    ultimately have ?ths using pnab by blast}
+  ultimately have ?ths by blast}
+ultimately show ?ths by blast
+qed
+
+text {* More useful lemmas. *}
+lemma prime_product: 
+  assumes "prime (p * q)"
+  shows "p = 1 \<or> q = 1"
+proof -
+  from assms have 
+    "1 < p * q" and P: "\<And>m. m dvd p * q \<Longrightarrow> m = 1 \<or> m = p * q"
+    unfolding prime_def by auto
+  from `1 < p * q` have "p \<noteq> 0" by (cases p) auto
+  then have Q: "p = p * q \<longleftrightarrow> q = 1" by auto
+  have "p dvd p * q" by simp
+  then have "p = 1 \<or> p = p * q" by (rule P)
+  then show ?thesis by (simp add: Q)
+qed
+
+lemma prime_exp: "prime (p^n) \<longleftrightarrow> prime p \<and> n = 1"
+proof(induct n)
+  case 0 thus ?case by simp
+next
+  case (Suc n)
+  {assume "p = 0" hence ?case by simp}
+  moreover
+  {assume "p=1" hence ?case by simp}
+  moreover
+  {assume p: "p \<noteq> 0" "p\<noteq>1"
+    {assume pp: "prime (p^Suc n)"
+      hence "p = 1 \<or> p^n = 1" using prime_product[of p "p^n"] by simp
+      with p have n: "n = 0" 
+	by (simp only: exp_eq_1 ) simp
+      with pp have "prime p \<and> Suc n = 1" by simp}
+    moreover
+    {assume n: "prime p \<and> Suc n = 1" hence "prime (p^Suc n)" by simp}
+    ultimately have ?case by blast}
+  ultimately show ?case by blast
+qed
+
+lemma prime_power_mult: 
+  assumes p: "prime p" and xy: "x * y = p ^ k"
+  shows "\<exists>i j. x = p ^i \<and> y = p^ j"
+  using xy
+proof(induct k arbitrary: x y)
+  case 0 thus ?case apply simp by (rule exI[where x="0"], simp)
+next
+  case (Suc k x y)
+  from Suc.prems have pxy: "p dvd x*y" by auto
+  from prime_divprod[OF p pxy] have pxyc: "p dvd x \<or> p dvd y" .
+  from p have p0: "p \<noteq> 0" by - (rule ccontr, simp) 
+  {assume px: "p dvd x"
+    then obtain d where d: "x = p*d" unfolding dvd_def by blast
+    from Suc.prems d  have "p*d*y = p^Suc k" by simp
+    hence th: "d*y = p^k" using p0 by simp
+    from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "y = p^j" by blast
+    with d have "x = p^Suc i" by simp 
+    with ij(2) have ?case by blast}
+  moreover 
+  {assume px: "p dvd y"
+    then obtain d where d: "y = p*d" unfolding dvd_def by blast
+    from Suc.prems d  have "p*d*x = p^Suc k" by (simp add: mult_commute)
+    hence th: "d*x = p^k" using p0 by simp
+    from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "x = p^j" by blast
+    with d have "y = p^Suc i" by simp 
+    with ij(2) have ?case by blast}
+  ultimately show ?case  using pxyc by blast
+qed
+
+lemma prime_power_exp: assumes p: "prime p" and n:"n \<noteq> 0" 
+  and xn: "x^n = p^k" shows "\<exists>i. x = p^i"
+  using n xn
+proof(induct n arbitrary: k)
+  case 0 thus ?case by simp
+next
+  case (Suc n k) hence th: "x*x^n = p^k" by simp
+  {assume "n = 0" with prems have ?case apply simp 
+      by (rule exI[where x="k"],simp)}
+  moreover
+  {assume n: "n \<noteq> 0"
+    from prime_power_mult[OF p th] 
+    obtain i j where ij: "x = p^i" "x^n = p^j"by blast
+    from Suc.hyps[OF n ij(2)] have ?case .}
+  ultimately show ?case by blast
+qed
+
+lemma divides_primepow: assumes p: "prime p" 
+  shows "d dvd p^k \<longleftrightarrow> (\<exists> i. i \<le> k \<and> d = p ^i)"
+proof
+  assume H: "d dvd p^k" then obtain e where e: "d*e = p^k" 
+    unfolding dvd_def  apply (auto simp add: mult_commute) by blast
+  from prime_power_mult[OF p e] obtain i j where ij: "d = p^i" "e=p^j" by blast
+  from prime_ge_2[OF p] have p1: "p > 1" by arith
+  from e ij have "p^(i + j) = p^k" by (simp add: power_add)
+  hence "i + j = k" using power_inject_exp[of p "i+j" k, OF p1] by simp 
+  hence "i \<le> k" by arith
+  with ij(1) show "\<exists>i\<le>k. d = p ^ i" by blast
+next
+  {fix i assume H: "i \<le> k" "d = p^i"
+    hence "\<exists>j. k = i + j" by arith
+    then obtain j where j: "k = i + j" by blast
+    hence "p^k = p^j*d" using H(2) by (simp add: power_add)
+    hence "d dvd p^k" unfolding dvd_def by auto}
+  thus "\<exists>i\<le>k. d = p ^ i \<Longrightarrow> d dvd p ^ k" by blast
+qed
+
+lemma coprime_divisors: "d dvd a \<Longrightarrow> e dvd b \<Longrightarrow> coprime a b \<Longrightarrow> coprime d e"
+  by (auto simp add: dvd_def coprime)
+
+declare power_Suc0[simp del]
+declare even_dvd[simp del]
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/Quadratic_Reciprocity.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,642 @@
+(*  Authors:    Jeremy Avigad, David Gray, and Adam Kramer
+*)
+
+header {* The law of Quadratic reciprocity *}
+
+theory Quadratic_Reciprocity
+imports Gauss
+begin
+
+text {*
+  Lemmas leading up to the proof of theorem 3.3 in Niven and
+  Zuckerman's presentation.
+*}
+
+context GAUSS
+begin
+
+lemma QRLemma1: "a * setsum id A =
+  p * setsum (%x. ((x * a) div p)) A + setsum id D + setsum id E"
+proof -
+  from finite_A have "a * setsum id A = setsum (%x. a * x) A"
+    by (auto simp add: setsum_const_mult id_def)
+  also have "setsum (%x. a * x) = setsum (%x. x * a)"
+    by (auto simp add: zmult_commute)
+  also have "setsum (%x. x * a) A = setsum id B"
+    by (simp add: B_def setsum_reindex_id[OF inj_on_xa_A])
+  also have "... = setsum (%x. p * (x div p) + StandardRes p x) B"
+    by (auto simp add: StandardRes_def zmod_zdiv_equality)
+  also have "... = setsum (%x. p * (x div p)) B + setsum (StandardRes p) B"
+    by (rule setsum_addf)
+  also have "setsum (StandardRes p) B = setsum id C"
+    by (auto simp add: C_def setsum_reindex_id[OF SR_B_inj])
+  also from C_eq have "... = setsum id (D \<union> E)"
+    by auto
+  also from finite_D finite_E have "... = setsum id D + setsum id E"
+    by (rule setsum_Un_disjoint) (auto simp add: D_def E_def)
+  also have "setsum (%x. p * (x div p)) B =
+      setsum ((%x. p * (x div p)) o (%x. (x * a))) A"
+    by (auto simp add: B_def setsum_reindex inj_on_xa_A)
+  also have "... = setsum (%x. p * ((x * a) div p)) A"
+    by (auto simp add: o_def)
+  also from finite_A have "setsum (%x. p * ((x * a) div p)) A =
+    p * setsum (%x. ((x * a) div p)) A"
+    by (auto simp add: setsum_const_mult)
+  finally show ?thesis by arith
+qed
+
+lemma QRLemma2: "setsum id A = p * int (card E) - setsum id E +
+  setsum id D"
+proof -
+  from F_Un_D_eq_A have "setsum id A = setsum id (D \<union> F)"
+    by (simp add: Un_commute)
+  also from F_D_disj finite_D finite_F
+  have "... = setsum id D + setsum id F"
+    by (auto simp add: Int_commute intro: setsum_Un_disjoint)
+  also from F_def have "F = (%x. (p - x)) ` E"
+    by auto
+  also from finite_E inj_on_pminusx_E have "setsum id ((%x. (p - x)) ` E) =
+      setsum (%x. (p - x)) E"
+    by (auto simp add: setsum_reindex)
+  also from finite_E have "setsum (op - p) E = setsum (%x. p) E - setsum id E"
+    by (auto simp add: setsum_subtractf id_def)
+  also from finite_E have "setsum (%x. p) E = p * int(card E)"
+    by (intro setsum_const)
+  finally show ?thesis
+    by arith
+qed
+
+lemma QRLemma3: "(a - 1) * setsum id A =
+    p * (setsum (%x. ((x * a) div p)) A - int(card E)) + 2 * setsum id E"
+proof -
+  have "(a - 1) * setsum id A = a * setsum id A - setsum id A"
+    by (auto simp add: zdiff_zmult_distrib)
+  also note QRLemma1
+  also from QRLemma2 have "p * (\<Sum>x \<in> A. x * a div p) + setsum id D +
+     setsum id E - setsum id A =
+      p * (\<Sum>x \<in> A. x * a div p) + setsum id D +
+      setsum id E - (p * int (card E) - setsum id E + setsum id D)"
+    by auto
+  also have "... = p * (\<Sum>x \<in> A. x * a div p) -
+      p * int (card E) + 2 * setsum id E"
+    by arith
+  finally show ?thesis
+    by (auto simp only: zdiff_zmult_distrib2)
+qed
+
+lemma QRLemma4: "a \<in> zOdd ==>
+    (setsum (%x. ((x * a) div p)) A \<in> zEven) = (int(card E): zEven)"
+proof -
+  assume a_odd: "a \<in> zOdd"
+  from QRLemma3 have a: "p * (setsum (%x. ((x * a) div p)) A - int(card E)) =
+      (a - 1) * setsum id A - 2 * setsum id E"
+    by arith
+  from a_odd have "a - 1 \<in> zEven"
+    by (rule odd_minus_one_even)
+  hence "(a - 1) * setsum id A \<in> zEven"
+    by (rule even_times_either)
+  moreover have "2 * setsum id E \<in> zEven"
+    by (auto simp add: zEven_def)
+  ultimately have "(a - 1) * setsum id A - 2 * setsum id E \<in> zEven"
+    by (rule even_minus_even)
+  with a have "p * (setsum (%x. ((x * a) div p)) A - int(card E)): zEven"
+    by simp
+  hence "p \<in> zEven | (setsum (%x. ((x * a) div p)) A - int(card E)): zEven"
+    by (rule EvenOdd.even_product)
+  with p_odd have "(setsum (%x. ((x * a) div p)) A - int(card E)): zEven"
+    by (auto simp add: odd_iff_not_even)
+  thus ?thesis
+    by (auto simp only: even_diff [symmetric])
+qed
+
+lemma QRLemma5: "a \<in> zOdd ==>
+   (-1::int)^(card E) = (-1::int)^(nat(setsum (%x. ((x * a) div p)) A))"
+proof -
+  assume "a \<in> zOdd"
+  from QRLemma4 [OF this] have
+    "(int(card E): zEven) = (setsum (%x. ((x * a) div p)) A \<in> zEven)" ..
+  moreover have "0 \<le> int(card E)"
+    by auto
+  moreover have "0 \<le> setsum (%x. ((x * a) div p)) A"
+    proof (intro setsum_nonneg)
+      show "\<forall>x \<in> A. 0 \<le> x * a div p"
+      proof
+        fix x
+        assume "x \<in> A"
+        then have "0 \<le> x"
+          by (auto simp add: A_def)
+        with a_nonzero have "0 \<le> x * a"
+          by (auto simp add: zero_le_mult_iff)
+        with p_g_2 show "0 \<le> x * a div p"
+          by (auto simp add: pos_imp_zdiv_nonneg_iff)
+      qed
+    qed
+  ultimately have "(-1::int)^nat((int (card E))) =
+      (-1)^nat(((\<Sum>x \<in> A. x * a div p)))"
+    by (intro neg_one_power_parity, auto)
+  also have "nat (int(card E)) = card E"
+    by auto
+  finally show ?thesis .
+qed
+
+end
+
+lemma MainQRLemma: "[| a \<in> zOdd; 0 < a; ~([a = 0] (mod p)); zprime p; 2 < p;
+  A = {x. 0 < x & x \<le> (p - 1) div 2} |] ==>
+  (Legendre a p) = (-1::int)^(nat(setsum (%x. ((x * a) div p)) A))"
+  apply (subst GAUSS.gauss_lemma)
+  apply (auto simp add: GAUSS_def)
+  apply (subst GAUSS.QRLemma5)
+  apply (auto simp add: GAUSS_def)
+  apply (simp add: GAUSS.A_def [OF GAUSS.intro] GAUSS_def)
+  done
+
+
+subsection {* Stuff about S, S1 and S2 *}
+
+locale QRTEMP =
+  fixes p     :: "int"
+  fixes q     :: "int"
+
+  assumes p_prime: "zprime p"
+  assumes p_g_2: "2 < p"
+  assumes q_prime: "zprime q"
+  assumes q_g_2: "2 < q"
+  assumes p_neq_q:      "p \<noteq> q"
+begin
+
+definition
+  P_set :: "int set" where
+  "P_set = {x. 0 < x & x \<le> ((p - 1) div 2) }"
+
+definition
+  Q_set :: "int set" where
+  "Q_set = {x. 0 < x & x \<le> ((q - 1) div 2) }"
+  
+definition
+  S :: "(int * int) set" where
+  "S = P_set <*> Q_set"
+
+definition
+  S1 :: "(int * int) set" where
+  "S1 = { (x, y). (x, y):S & ((p * y) < (q * x)) }"
+
+definition
+  S2 :: "(int * int) set" where
+  "S2 = { (x, y). (x, y):S & ((q * x) < (p * y)) }"
+
+definition
+  f1 :: "int => (int * int) set" where
+  "f1 j = { (j1, y). (j1, y):S & j1 = j & (y \<le> (q * j) div p) }"
+
+definition
+  f2 :: "int => (int * int) set" where
+  "f2 j = { (x, j1). (x, j1):S & j1 = j & (x \<le> (p * j) div q) }"
+
+lemma p_fact: "0 < (p - 1) div 2"
+proof -
+  from p_g_2 have "2 \<le> p - 1" by arith
+  then have "2 div 2 \<le> (p - 1) div 2" by (rule zdiv_mono1, auto)
+  then show ?thesis by auto
+qed
+
+lemma q_fact: "0 < (q - 1) div 2"
+proof -
+  from q_g_2 have "2 \<le> q - 1" by arith
+  then have "2 div 2 \<le> (q - 1) div 2" by (rule zdiv_mono1, auto)
+  then show ?thesis by auto
+qed
+
+lemma pb_neq_qa: "[|1 \<le> b; b \<le> (q - 1) div 2 |] ==>
+    (p * b \<noteq> q * a)"
+proof
+  assume "p * b = q * a" and "1 \<le> b" and "b \<le> (q - 1) div 2"
+  then have "q dvd (p * b)" by (auto simp add: dvd_def)
+  with q_prime p_g_2 have "q dvd p | q dvd b"
+    by (auto simp add: zprime_zdvd_zmult)
+  moreover have "~ (q dvd p)"
+  proof
+    assume "q dvd p"
+    with p_prime have "q = 1 | q = p"
+      apply (auto simp add: zprime_def QRTEMP_def)
+      apply (drule_tac x = q and R = False in allE)
+      apply (simp add: QRTEMP_def)
+      apply (subgoal_tac "0 \<le> q", simp add: QRTEMP_def)
+      apply (insert prems)
+      apply (auto simp add: QRTEMP_def)
+      done
+    with q_g_2 p_neq_q show False by auto
+  qed
+  ultimately have "q dvd b" by auto
+  then have "q \<le> b"
+  proof -
+    assume "q dvd b"
+    moreover from prems have "0 < b" by auto
+    ultimately show ?thesis using zdvd_bounds [of q b] by auto
+  qed
+  with prems have "q \<le> (q - 1) div 2" by auto
+  then have "2 * q \<le> 2 * ((q - 1) div 2)" by arith
+  then have "2 * q \<le> q - 1"
+  proof -
+    assume "2 * q \<le> 2 * ((q - 1) div 2)"
+    with prems have "q \<in> zOdd" by (auto simp add: QRTEMP_def zprime_zOdd_eq_grt_2)
+    with odd_minus_one_even have "(q - 1):zEven" by auto
+    with even_div_2_prop2 have "(q - 1) = 2 * ((q - 1) div 2)" by auto
+    with prems show ?thesis by auto
+  qed
+  then have p1: "q \<le> -1" by arith
+  with q_g_2 show False by auto
+qed
+
+lemma P_set_finite: "finite (P_set)"
+  using p_fact by (auto simp add: P_set_def bdd_int_set_l_le_finite)
+
+lemma Q_set_finite: "finite (Q_set)"
+  using q_fact by (auto simp add: Q_set_def bdd_int_set_l_le_finite)
+
+lemma S_finite: "finite S"
+  by (auto simp add: S_def  P_set_finite Q_set_finite finite_cartesian_product)
+
+lemma S1_finite: "finite S1"
+proof -
+  have "finite S" by (auto simp add: S_finite)
+  moreover have "S1 \<subseteq> S" by (auto simp add: S1_def S_def)
+  ultimately show ?thesis by (auto simp add: finite_subset)
+qed
+
+lemma S2_finite: "finite S2"
+proof -
+  have "finite S" by (auto simp add: S_finite)
+  moreover have "S2 \<subseteq> S" by (auto simp add: S2_def S_def)
+  ultimately show ?thesis by (auto simp add: finite_subset)
+qed
+
+lemma P_set_card: "(p - 1) div 2 = int (card (P_set))"
+  using p_fact by (auto simp add: P_set_def card_bdd_int_set_l_le)
+
+lemma Q_set_card: "(q - 1) div 2 = int (card (Q_set))"
+  using q_fact by (auto simp add: Q_set_def card_bdd_int_set_l_le)
+
+lemma S_card: "((p - 1) div 2) * ((q - 1) div 2) = int (card(S))"
+  using P_set_card Q_set_card P_set_finite Q_set_finite
+  by (auto simp add: S_def zmult_int setsum_constant)
+
+lemma S1_Int_S2_prop: "S1 \<inter> S2 = {}"
+  by (auto simp add: S1_def S2_def)
+
+lemma S1_Union_S2_prop: "S = S1 \<union> S2"
+  apply (auto simp add: S_def P_set_def Q_set_def S1_def S2_def)
+proof -
+  fix a and b
+  assume "~ q * a < p * b" and b1: "0 < b" and b2: "b \<le> (q - 1) div 2"
+  with zless_linear have "(p * b < q * a) | (p * b = q * a)" by auto
+  moreover from pb_neq_qa b1 b2 have "(p * b \<noteq> q * a)" by auto
+  ultimately show "p * b < q * a" by auto
+qed
+
+lemma card_sum_S1_S2: "((p - 1) div 2) * ((q - 1) div 2) =
+    int(card(S1)) + int(card(S2))"
+proof -
+  have "((p - 1) div 2) * ((q - 1) div 2) = int (card(S))"
+    by (auto simp add: S_card)
+  also have "... = int( card(S1) + card(S2))"
+    apply (insert S1_finite S2_finite S1_Int_S2_prop S1_Union_S2_prop)
+    apply (drule card_Un_disjoint, auto)
+    done
+  also have "... = int(card(S1)) + int(card(S2))" by auto
+  finally show ?thesis .
+qed
+
+lemma aux1a: "[| 0 < a; a \<le> (p - 1) div 2;
+                             0 < b; b \<le> (q - 1) div 2 |] ==>
+                          (p * b < q * a) = (b \<le> q * a div p)"
+proof -
+  assume "0 < a" and "a \<le> (p - 1) div 2" and "0 < b" and "b \<le> (q - 1) div 2"
+  have "p * b < q * a ==> b \<le> q * a div p"
+  proof -
+    assume "p * b < q * a"
+    then have "p * b \<le> q * a" by auto
+    then have "(p * b) div p \<le> (q * a) div p"
+      by (rule zdiv_mono1) (insert p_g_2, auto)
+    then show "b \<le> (q * a) div p"
+      apply (subgoal_tac "p \<noteq> 0")
+      apply (frule div_mult_self1_is_id, force)
+      apply (insert p_g_2, auto)
+      done
+  qed
+  moreover have "b \<le> q * a div p ==> p * b < q * a"
+  proof -
+    assume "b \<le> q * a div p"
+    then have "p * b \<le> p * ((q * a) div p)"
+      using p_g_2 by (auto simp add: mult_le_cancel_left)
+    also have "... \<le> q * a"
+      by (rule zdiv_leq_prop) (insert p_g_2, auto)
+    finally have "p * b \<le> q * a" .
+    then have "p * b < q * a | p * b = q * a"
+      by (simp only: order_le_imp_less_or_eq)
+    moreover have "p * b \<noteq> q * a"
+      by (rule  pb_neq_qa) (insert prems, auto)
+    ultimately show ?thesis by auto
+  qed
+  ultimately show ?thesis ..
+qed
+
+lemma aux1b: "[| 0 < a; a \<le> (p - 1) div 2;
+                             0 < b; b \<le> (q - 1) div 2 |] ==>
+                          (q * a < p * b) = (a \<le> p * b div q)"
+proof -
+  assume "0 < a" and "a \<le> (p - 1) div 2" and "0 < b" and "b \<le> (q - 1) div 2"
+  have "q * a < p * b ==> a \<le> p * b div q"
+  proof -
+    assume "q * a < p * b"
+    then have "q * a \<le> p * b" by auto
+    then have "(q * a) div q \<le> (p * b) div q"
+      by (rule zdiv_mono1) (insert q_g_2, auto)
+    then show "a \<le> (p * b) div q"
+      apply (subgoal_tac "q \<noteq> 0")
+      apply (frule div_mult_self1_is_id, force)
+      apply (insert q_g_2, auto)
+      done
+  qed
+  moreover have "a \<le> p * b div q ==> q * a < p * b"
+  proof -
+    assume "a \<le> p * b div q"
+    then have "q * a \<le> q * ((p * b) div q)"
+      using q_g_2 by (auto simp add: mult_le_cancel_left)
+    also have "... \<le> p * b"
+      by (rule zdiv_leq_prop) (insert q_g_2, auto)
+    finally have "q * a \<le> p * b" .
+    then have "q * a < p * b | q * a = p * b"
+      by (simp only: order_le_imp_less_or_eq)
+    moreover have "p * b \<noteq> q * a"
+      by (rule  pb_neq_qa) (insert prems, auto)
+    ultimately show ?thesis by auto
+  qed
+  ultimately show ?thesis ..
+qed
+
+lemma (in -) aux2: "[| zprime p; zprime q; 2 < p; 2 < q |] ==>
+             (q * ((p - 1) div 2)) div p \<le> (q - 1) div 2"
+proof-
+  assume "zprime p" and "zprime q" and "2 < p" and "2 < q"
+  (* Set up what's even and odd *)
+  then have "p \<in> zOdd & q \<in> zOdd"
+    by (auto simp add:  zprime_zOdd_eq_grt_2)
+  then have even1: "(p - 1):zEven & (q - 1):zEven"
+    by (auto simp add: odd_minus_one_even)
+  then have even2: "(2 * p):zEven & ((q - 1) * p):zEven"
+    by (auto simp add: zEven_def)
+  then have even3: "(((q - 1) * p) + (2 * p)):zEven"
+    by (auto simp: EvenOdd.even_plus_even)
+  (* using these prove it *)
+  from prems have "q * (p - 1) < ((q - 1) * p) + (2 * p)"
+    by (auto simp add: int_distrib)
+  then have "((p - 1) * q) div 2 < (((q - 1) * p) + (2 * p)) div 2"
+    apply (rule_tac x = "((p - 1) * q)" in even_div_2_l)
+    by (auto simp add: even3, auto simp add: zmult_ac)
+  also have "((p - 1) * q) div 2 = q * ((p - 1) div 2)"
+    by (auto simp add: even1 even_prod_div_2)
+  also have "(((q - 1) * p) + (2 * p)) div 2 = (((q - 1) div 2) * p) + p"
+    by (auto simp add: even1 even2 even_prod_div_2 even_sum_div_2)
+  finally show ?thesis
+    apply (rule_tac x = " q * ((p - 1) div 2)" and
+                    y = "(q - 1) div 2" in div_prop2)
+    using prems by auto
+qed
+
+lemma aux3a: "\<forall>j \<in> P_set. int (card (f1 j)) = (q * j) div p"
+proof
+  fix j
+  assume j_fact: "j \<in> P_set"
+  have "int (card (f1 j)) = int (card {y. y \<in> Q_set & y \<le> (q * j) div p})"
+  proof -
+    have "finite (f1 j)"
+    proof -
+      have "(f1 j) \<subseteq> S" by (auto simp add: f1_def)
+      with S_finite show ?thesis by (auto simp add: finite_subset)
+    qed
+    moreover have "inj_on (%(x,y). y) (f1 j)"
+      by (auto simp add: f1_def inj_on_def)
+    ultimately have "card ((%(x,y). y) ` (f1 j)) = card  (f1 j)"
+      by (auto simp add: f1_def card_image)
+    moreover have "((%(x,y). y) ` (f1 j)) = {y. y \<in> Q_set & y \<le> (q * j) div p}"
+      using prems by (auto simp add: f1_def S_def Q_set_def P_set_def image_def)
+    ultimately show ?thesis by (auto simp add: f1_def)
+  qed
+  also have "... = int (card {y. 0 < y & y \<le> (q * j) div p})"
+  proof -
+    have "{y. y \<in> Q_set & y \<le> (q * j) div p} =
+        {y. 0 < y & y \<le> (q * j) div p}"
+      apply (auto simp add: Q_set_def)
+    proof -
+      fix x
+      assume "0 < x" and "x \<le> q * j div p"
+      with j_fact P_set_def  have "j \<le> (p - 1) div 2" by auto
+      with q_g_2 have "q * j \<le> q * ((p - 1) div 2)"
+        by (auto simp add: mult_le_cancel_left)
+      with p_g_2 have "q * j div p \<le> q * ((p - 1) div 2) div p"
+        by (auto simp add: zdiv_mono1)
+      also from prems P_set_def have "... \<le> (q - 1) div 2"
+        apply simp
+        apply (insert aux2)
+        apply (simp add: QRTEMP_def)
+        done
+      finally show "x \<le> (q - 1) div 2" using prems by auto
+    qed
+    then show ?thesis by auto
+  qed
+  also have "... = (q * j) div p"
+  proof -
+    from j_fact P_set_def have "0 \<le> j" by auto
+    with q_g_2 have "q * 0 \<le> q * j" by (auto simp only: mult_left_mono)
+    then have "0 \<le> q * j" by auto
+    then have "0 div p \<le> (q * j) div p"
+      apply (rule_tac a = 0 in zdiv_mono1)
+      apply (insert p_g_2, auto)
+      done
+    also have "0 div p = 0" by auto
+    finally show ?thesis by (auto simp add: card_bdd_int_set_l_le)
+  qed
+  finally show "int (card (f1 j)) = q * j div p" .
+qed
+
+lemma aux3b: "\<forall>j \<in> Q_set. int (card (f2 j)) = (p * j) div q"
+proof
+  fix j
+  assume j_fact: "j \<in> Q_set"
+  have "int (card (f2 j)) = int (card {y. y \<in> P_set & y \<le> (p * j) div q})"
+  proof -
+    have "finite (f2 j)"
+    proof -
+      have "(f2 j) \<subseteq> S" by (auto simp add: f2_def)
+      with S_finite show ?thesis by (auto simp add: finite_subset)
+    qed
+    moreover have "inj_on (%(x,y). x) (f2 j)"
+      by (auto simp add: f2_def inj_on_def)
+    ultimately have "card ((%(x,y). x) ` (f2 j)) = card  (f2 j)"
+      by (auto simp add: f2_def card_image)
+    moreover have "((%(x,y). x) ` (f2 j)) = {y. y \<in> P_set & y \<le> (p * j) div q}"
+      using prems by (auto simp add: f2_def S_def Q_set_def P_set_def image_def)
+    ultimately show ?thesis by (auto simp add: f2_def)
+  qed
+  also have "... = int (card {y. 0 < y & y \<le> (p * j) div q})"
+  proof -
+    have "{y. y \<in> P_set & y \<le> (p * j) div q} =
+        {y. 0 < y & y \<le> (p * j) div q}"
+      apply (auto simp add: P_set_def)
+    proof -
+      fix x
+      assume "0 < x" and "x \<le> p * j div q"
+      with j_fact Q_set_def  have "j \<le> (q - 1) div 2" by auto
+      with p_g_2 have "p * j \<le> p * ((q - 1) div 2)"
+        by (auto simp add: mult_le_cancel_left)
+      with q_g_2 have "p * j div q \<le> p * ((q - 1) div 2) div q"
+        by (auto simp add: zdiv_mono1)
+      also from prems have "... \<le> (p - 1) div 2"
+        by (auto simp add: aux2 QRTEMP_def)
+      finally show "x \<le> (p - 1) div 2" using prems by auto
+      qed
+    then show ?thesis by auto
+  qed
+  also have "... = (p * j) div q"
+  proof -
+    from j_fact Q_set_def have "0 \<le> j" by auto
+    with p_g_2 have "p * 0 \<le> p * j" by (auto simp only: mult_left_mono)
+    then have "0 \<le> p * j" by auto
+    then have "0 div q \<le> (p * j) div q"
+      apply (rule_tac a = 0 in zdiv_mono1)
+      apply (insert q_g_2, auto)
+      done
+    also have "0 div q = 0" by auto
+    finally show ?thesis by (auto simp add: card_bdd_int_set_l_le)
+  qed
+  finally show "int (card (f2 j)) = p * j div q" .
+qed
+
+lemma S1_card: "int (card(S1)) = setsum (%j. (q * j) div p) P_set"
+proof -
+  have "\<forall>x \<in> P_set. finite (f1 x)"
+  proof
+    fix x
+    have "f1 x \<subseteq> S" by (auto simp add: f1_def)
+    with S_finite show "finite (f1 x)" by (auto simp add: finite_subset)
+  qed
+  moreover have "(\<forall>x \<in> P_set. \<forall>y \<in> P_set. x \<noteq> y --> (f1 x) \<inter> (f1 y) = {})"
+    by (auto simp add: f1_def)
+  moreover note P_set_finite
+  ultimately have "int(card (UNION P_set f1)) =
+      setsum (%x. int(card (f1 x))) P_set"
+    by(simp add:card_UN_disjoint int_setsum o_def)
+  moreover have "S1 = UNION P_set f1"
+    by (auto simp add: f1_def S_def S1_def S2_def P_set_def Q_set_def aux1a)
+  ultimately have "int(card (S1)) = setsum (%j. int(card (f1 j))) P_set"
+    by auto
+  also have "... = setsum (%j. q * j div p) P_set"
+    using aux3a by(fastsimp intro: setsum_cong)
+  finally show ?thesis .
+qed
+
+lemma S2_card: "int (card(S2)) = setsum (%j. (p * j) div q) Q_set"
+proof -
+  have "\<forall>x \<in> Q_set. finite (f2 x)"
+  proof
+    fix x
+    have "f2 x \<subseteq> S" by (auto simp add: f2_def)
+    with S_finite show "finite (f2 x)" by (auto simp add: finite_subset)
+  qed
+  moreover have "(\<forall>x \<in> Q_set. \<forall>y \<in> Q_set. x \<noteq> y -->
+      (f2 x) \<inter> (f2 y) = {})"
+    by (auto simp add: f2_def)
+  moreover note Q_set_finite
+  ultimately have "int(card (UNION Q_set f2)) =
+      setsum (%x. int(card (f2 x))) Q_set"
+    by(simp add:card_UN_disjoint int_setsum o_def)
+  moreover have "S2 = UNION Q_set f2"
+    by (auto simp add: f2_def S_def S1_def S2_def P_set_def Q_set_def aux1b)
+  ultimately have "int(card (S2)) = setsum (%j. int(card (f2 j))) Q_set"
+    by auto
+  also have "... = setsum (%j. p * j div q) Q_set"
+    using aux3b by(fastsimp intro: setsum_cong)
+  finally show ?thesis .
+qed
+
+lemma S1_carda: "int (card(S1)) =
+    setsum (%j. (j * q) div p) P_set"
+  by (auto simp add: S1_card zmult_ac)
+
+lemma S2_carda: "int (card(S2)) =
+    setsum (%j. (j * p) div q) Q_set"
+  by (auto simp add: S2_card zmult_ac)
+
+lemma pq_sum_prop: "(setsum (%j. (j * p) div q) Q_set) +
+    (setsum (%j. (j * q) div p) P_set) = ((p - 1) div 2) * ((q - 1) div 2)"
+proof -
+  have "(setsum (%j. (j * p) div q) Q_set) +
+      (setsum (%j. (j * q) div p) P_set) = int (card S2) + int (card S1)"
+    by (auto simp add: S1_carda S2_carda)
+  also have "... = int (card S1) + int (card S2)"
+    by auto
+  also have "... = ((p - 1) div 2) * ((q - 1) div 2)"
+    by (auto simp add: card_sum_S1_S2)
+  finally show ?thesis .
+qed
+
+
+lemma (in -) pq_prime_neq: "[| zprime p; zprime q; p \<noteq> q |] ==> (~[p = 0] (mod q))"
+  apply (auto simp add: zcong_eq_zdvd_prop zprime_def)
+  apply (drule_tac x = q in allE)
+  apply (drule_tac x = p in allE)
+  apply auto
+  done
+
+
+lemma QR_short: "(Legendre p q) * (Legendre q p) =
+    (-1::int)^nat(((p - 1) div 2)*((q - 1) div 2))"
+proof -
+  from prems have "~([p = 0] (mod q))"
+    by (auto simp add: pq_prime_neq QRTEMP_def)
+  with prems Q_set_def have a1: "(Legendre p q) = (-1::int) ^
+      nat(setsum (%x. ((x * p) div q)) Q_set)"
+    apply (rule_tac p = q in  MainQRLemma)
+    apply (auto simp add: zprime_zOdd_eq_grt_2 QRTEMP_def)
+    done
+  from prems have "~([q = 0] (mod p))"
+    apply (rule_tac p = q and q = p in pq_prime_neq)
+    apply (simp add: QRTEMP_def)+
+    done
+  with prems P_set_def have a2: "(Legendre q p) =
+      (-1::int) ^ nat(setsum (%x. ((x * q) div p)) P_set)"
+    apply (rule_tac p = p in  MainQRLemma)
+    apply (auto simp add: zprime_zOdd_eq_grt_2 QRTEMP_def)
+    done
+  from a1 a2 have "(Legendre p q) * (Legendre q p) =
+      (-1::int) ^ nat(setsum (%x. ((x * p) div q)) Q_set) *
+        (-1::int) ^ nat(setsum (%x. ((x * q) div p)) P_set)"
+    by auto
+  also have "... = (-1::int) ^ (nat(setsum (%x. ((x * p) div q)) Q_set) +
+                   nat(setsum (%x. ((x * q) div p)) P_set))"
+    by (auto simp add: zpower_zadd_distrib)
+  also have "nat(setsum (%x. ((x * p) div q)) Q_set) +
+      nat(setsum (%x. ((x * q) div p)) P_set) =
+        nat((setsum (%x. ((x * p) div q)) Q_set) +
+          (setsum (%x. ((x * q) div p)) P_set))"
+    apply (rule_tac z = "setsum (%x. ((x * p) div q)) Q_set" in
+      nat_add_distrib [symmetric])
+    apply (auto simp add: S1_carda [symmetric] S2_carda [symmetric])
+    done
+  also have "... = nat(((p - 1) div 2) * ((q - 1) div 2))"
+    by (auto simp add: pq_sum_prop)
+  finally show ?thesis .
+qed
+
+end
+
+theorem Quadratic_Reciprocity:
+     "[| p \<in> zOdd; zprime p; q \<in> zOdd; zprime q;
+         p \<noteq> q |]
+      ==> (Legendre p q) * (Legendre q p) =
+          (-1::int)^nat(((p - 1) div 2)*((q - 1) div 2))"
+  by (auto simp add: QRTEMP.QR_short zprime_zOdd_eq_grt_2 [symmetric]
+                     QRTEMP_def)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/ROOT.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,4 @@
+
+no_document use_thys ["Infinite_Set", "Permutation"];
+use_thys ["Fib", "Factorization", "Chinese", "WilsonRuss",
+  "WilsonBij", "Quadratic_Reciprocity", "Primes", "Pocklington"];
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/Residues.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,172 @@
+(*  Title:      HOL/Quadratic_Reciprocity/Residues.thy
+    ID:         $Id$
+    Authors:    Jeremy Avigad, David Gray, and Adam Kramer
+*)
+
+header {* Residue Sets *}
+
+theory Residues imports Int2 begin
+
+text {*
+  \medskip Define the residue of a set, the standard residue,
+  quadratic residues, and prove some basic properties. *}
+
+definition
+  ResSet      :: "int => int set => bool" where
+  "ResSet m X = (\<forall>y1 y2. (y1 \<in> X & y2 \<in> X & [y1 = y2] (mod m) --> y1 = y2))"
+
+definition
+  StandardRes :: "int => int => int" where
+  "StandardRes m x = x mod m"
+
+definition
+  QuadRes     :: "int => int => bool" where
+  "QuadRes m x = (\<exists>y. ([(y ^ 2) = x] (mod m)))"
+
+definition
+  Legendre    :: "int => int => int" where
+  "Legendre a p = (if ([a = 0] (mod p)) then 0
+                     else if (QuadRes p a) then 1
+                     else -1)"
+
+definition
+  SR          :: "int => int set" where
+  "SR p = {x. (0 \<le> x) & (x < p)}"
+
+definition
+  SRStar      :: "int => int set" where
+  "SRStar p = {x. (0 < x) & (x < p)}"
+
+
+subsection {* Some useful properties of StandardRes *}
+
+lemma StandardRes_prop1: "[x = StandardRes m x] (mod m)"
+  by (auto simp add: StandardRes_def zcong_zmod)
+
+lemma StandardRes_prop2: "0 < m ==> (StandardRes m x1 = StandardRes m x2)
+      = ([x1 = x2] (mod m))"
+  by (auto simp add: StandardRes_def zcong_zmod_eq)
+
+lemma StandardRes_prop3: "(~[x = 0] (mod p)) = (~(StandardRes p x = 0))"
+  by (auto simp add: StandardRes_def zcong_def dvd_eq_mod_eq_0)
+
+lemma StandardRes_prop4: "2 < m 
+     ==> [StandardRes m x * StandardRes m y = (x * y)] (mod m)"
+  by (auto simp add: StandardRes_def zcong_zmod_eq 
+                     mod_mult_eq [of x y m])
+
+lemma StandardRes_lbound: "0 < p ==> 0 \<le> StandardRes p x"
+  by (auto simp add: StandardRes_def pos_mod_sign)
+
+lemma StandardRes_ubound: "0 < p ==> StandardRes p x < p"
+  by (auto simp add: StandardRes_def pos_mod_bound)
+
+lemma StandardRes_eq_zcong: 
+   "(StandardRes m x = 0) = ([x = 0](mod m))"
+  by (auto simp add: StandardRes_def zcong_eq_zdvd_prop dvd_def) 
+
+
+subsection {* Relations between StandardRes, SRStar, and SR *}
+
+lemma SRStar_SR_prop: "x \<in> SRStar p ==> x \<in> SR p"
+  by (auto simp add: SRStar_def SR_def)
+
+lemma StandardRes_SR_prop: "x \<in> SR p ==> StandardRes p x = x"
+  by (auto simp add: SR_def StandardRes_def mod_pos_pos_trivial)
+
+lemma StandardRes_SRStar_prop1: "2 < p ==> (StandardRes p x \<in> SRStar p) 
+     = (~[x = 0] (mod p))"
+  apply (auto simp add: StandardRes_prop3 StandardRes_def
+                        SRStar_def pos_mod_bound)
+  apply (subgoal_tac "0 < p")
+  apply (drule_tac a = x in pos_mod_sign, arith, simp)
+  done
+
+lemma StandardRes_SRStar_prop1a: "x \<in> SRStar p ==> ~([x = 0] (mod p))"
+  by (auto simp add: SRStar_def zcong_def zdvd_not_zless)
+
+lemma StandardRes_SRStar_prop2: "[| 2 < p; zprime p; x \<in> SRStar p |] 
+     ==> StandardRes p (MultInv p x) \<in> SRStar p"
+  apply (frule_tac x = "(MultInv p x)" in StandardRes_SRStar_prop1, simp)
+  apply (rule MultInv_prop3)
+  apply (auto simp add: SRStar_def zcong_def zdvd_not_zless)
+  done
+
+lemma StandardRes_SRStar_prop3: "x \<in> SRStar p ==> StandardRes p x = x"
+  by (auto simp add: SRStar_SR_prop StandardRes_SR_prop)
+
+lemma StandardRes_SRStar_prop4: "[| zprime p; 2 < p; x \<in> SRStar p |] 
+     ==> StandardRes p x \<in> SRStar p"
+  by (frule StandardRes_SRStar_prop3, auto)
+
+lemma SRStar_mult_prop1: "[| zprime p; 2 < p; x \<in> SRStar p; y \<in> SRStar p|] 
+     ==> (StandardRes p (x * y)):SRStar p"
+  apply (frule_tac x = x in StandardRes_SRStar_prop4, auto)
+  apply (frule_tac x = y in StandardRes_SRStar_prop4, auto)
+  apply (auto simp add: StandardRes_SRStar_prop1 zcong_zmult_prop3)
+  done
+
+lemma SRStar_mult_prop2: "[| zprime p; 2 < p; ~([a = 0](mod p)); 
+     x \<in> SRStar p |] 
+     ==> StandardRes p (a * MultInv p x) \<in> SRStar p"
+  apply (frule_tac x = x in StandardRes_SRStar_prop2, auto)
+  apply (frule_tac x = "MultInv p x" in StandardRes_SRStar_prop1)
+  apply (auto simp add: StandardRes_SRStar_prop1 zcong_zmult_prop3)
+  done
+
+lemma SRStar_card: "2 < p ==> int(card(SRStar p)) = p - 1"
+  by (auto simp add: SRStar_def int_card_bdd_int_set_l_l)
+
+lemma SRStar_finite: "2 < p ==> finite( SRStar p)"
+  by (auto simp add: SRStar_def bdd_int_set_l_l_finite)
+
+
+subsection {* Properties relating ResSets with StandardRes *}
+
+lemma aux: "x mod m = y mod m ==> [x = y] (mod m)"
+  apply (subgoal_tac "x = y ==> [x = y](mod m)")
+  apply (subgoal_tac "[x mod m = y mod m] (mod m) ==> [x = y] (mod m)")
+  apply (auto simp add: zcong_zmod [of x y m])
+  done
+
+lemma StandardRes_inj_on_ResSet: "ResSet m X ==> (inj_on (StandardRes m) X)"
+  apply (auto simp add: ResSet_def StandardRes_def inj_on_def)
+  apply (drule_tac m = m in aux, auto)
+  done
+
+lemma StandardRes_Sum: "[| finite X; 0 < m |] 
+     ==> [setsum f X = setsum (StandardRes m o f) X](mod m)" 
+  apply (rule_tac F = X in finite_induct)
+  apply (auto intro!: zcong_zadd simp add: StandardRes_prop1)
+  done
+
+lemma SR_pos: "0 < m ==> (StandardRes m ` X) \<subseteq> {x. 0 \<le> x & x < m}"
+  by (auto simp add: StandardRes_ubound StandardRes_lbound)
+
+lemma ResSet_finite: "0 < m ==> ResSet m X ==> finite X"
+  apply (rule_tac f = "StandardRes m" in finite_imageD) 
+  apply (rule_tac B = "{x. (0 :: int) \<le> x & x < m}" in finite_subset)
+  apply (auto simp add: StandardRes_inj_on_ResSet bdd_int_set_l_finite SR_pos)
+  done
+
+lemma mod_mod_is_mod: "[x = x mod m](mod m)"
+  by (auto simp add: zcong_zmod)
+
+lemma StandardRes_prod: "[| finite X; 0 < m |] 
+     ==> [setprod f X = setprod (StandardRes m o f) X] (mod m)"
+  apply (rule_tac F = X in finite_induct)
+  apply (auto intro!: zcong_zmult simp add: StandardRes_prop1)
+  done
+
+lemma ResSet_image:
+  "[| 0 < m; ResSet m A; \<forall>x \<in> A. \<forall>y \<in> A. ([f x = f y](mod m) --> x = y) |] ==>
+    ResSet m (f ` A)"
+  by (auto simp add: ResSet_def)
+
+
+subsection {* Property for SRStar *}
+
+lemma ResSet_SRStar_prop: "ResSet p (SRStar p)"
+  by (auto simp add: SRStar_def ResSet_def zcong_zless_imp_eq)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/WilsonBij.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,261 @@
+(*  Author:     Thomas M. Rasmussen
+    Copyright   2000  University of Cambridge
+*)
+
+header {* Wilson's Theorem using a more abstract approach *}
+
+theory WilsonBij imports BijectionRel IntFact begin
+
+text {*
+  Wilson's Theorem using a more ``abstract'' approach based on
+  bijections between sets.  Does not use Fermat's Little Theorem
+  (unlike Russinoff).
+*}
+
+
+subsection {* Definitions and lemmas *}
+
+definition
+  reciR :: "int => int => int => bool" where
+  "reciR p = (\<lambda>a b. zcong (a * b) 1 p \<and> 1 < a \<and> a < p - 1 \<and> 1 < b \<and> b < p - 1)"
+
+definition
+  inv :: "int => int => int" where
+  "inv p a =
+    (if zprime p \<and> 0 < a \<and> a < p then
+      (SOME x. 0 \<le> x \<and> x < p \<and> zcong (a * x) 1 p)
+     else 0)"
+
+
+text {* \medskip Inverse *}
+
+lemma inv_correct:
+  "zprime p ==> 0 < a ==> a < p
+    ==> 0 \<le> inv p a \<and> inv p a < p \<and> [a * inv p a = 1] (mod p)"
+  apply (unfold inv_def)
+  apply (simp (no_asm_simp))
+  apply (rule zcong_lineq_unique [THEN ex1_implies_ex, THEN someI_ex])
+   apply (erule_tac [2] zless_zprime_imp_zrelprime)
+    apply (unfold zprime_def)
+    apply auto
+  done
+
+lemmas inv_ge = inv_correct [THEN conjunct1, standard]
+lemmas inv_less = inv_correct [THEN conjunct2, THEN conjunct1, standard]
+lemmas inv_is_inv = inv_correct [THEN conjunct2, THEN conjunct2, standard]
+
+lemma inv_not_0:
+  "zprime p ==> 1 < a ==> a < p - 1 ==> inv p a \<noteq> 0"
+  -- {* same as @{text WilsonRuss} *}
+  apply safe
+  apply (cut_tac a = a and p = p in inv_is_inv)
+     apply (unfold zcong_def)
+     apply auto
+  apply (subgoal_tac "\<not> p dvd 1")
+   apply (rule_tac [2] zdvd_not_zless)
+    apply (subgoal_tac "p dvd 1")
+     prefer 2
+     apply (subst dvd_minus_iff [symmetric])
+     apply auto
+  done
+
+lemma inv_not_1:
+  "zprime p ==> 1 < a ==> a < p - 1 ==> inv p a \<noteq> 1"
+  -- {* same as @{text WilsonRuss} *}
+  apply safe
+  apply (cut_tac a = a and p = p in inv_is_inv)
+     prefer 4
+     apply simp
+     apply (subgoal_tac "a = 1")
+      apply (rule_tac [2] zcong_zless_imp_eq)
+          apply auto
+  done
+
+lemma aux: "[a * (p - 1) = 1] (mod p) = [a = p - 1] (mod p)"
+  -- {* same as @{text WilsonRuss} *}
+  apply (unfold zcong_def)
+  apply (simp add: OrderedGroup.diff_diff_eq diff_diff_eq2 zdiff_zmult_distrib2)
+  apply (rule_tac s = "p dvd -((a + 1) + (p * -a))" in trans)
+   apply (simp add: mult_commute)
+  apply (subst dvd_minus_iff)
+  apply (subst zdvd_reduce)
+  apply (rule_tac s = "p dvd (a + 1) + (p * -1)" in trans)
+   apply (subst zdvd_reduce)
+   apply auto
+  done
+
+lemma inv_not_p_minus_1:
+  "zprime p ==> 1 < a ==> a < p - 1 ==> inv p a \<noteq> p - 1"
+  -- {* same as @{text WilsonRuss} *}
+  apply safe
+  apply (cut_tac a = a and p = p in inv_is_inv)
+     apply auto
+  apply (simp add: aux)
+  apply (subgoal_tac "a = p - 1")
+   apply (rule_tac [2] zcong_zless_imp_eq)
+       apply auto
+  done
+
+text {*
+  Below is slightly different as we don't expand @{term [source] inv}
+  but use ``@{text correct}'' theorems.
+*}
+
+lemma inv_g_1: "zprime p ==> 1 < a ==> a < p - 1 ==> 1 < inv p a"
+  apply (subgoal_tac "inv p a \<noteq> 1")
+   apply (subgoal_tac "inv p a \<noteq> 0")
+    apply (subst order_less_le)
+    apply (subst zle_add1_eq_le [symmetric])
+    apply (subst order_less_le)
+    apply (rule_tac [2] inv_not_0)
+      apply (rule_tac [5] inv_not_1)
+        apply auto
+  apply (rule inv_ge)
+    apply auto
+  done
+
+lemma inv_less_p_minus_1:
+  "zprime p ==> 1 < a ==> a < p - 1 ==> inv p a < p - 1"
+  -- {* ditto *}
+  apply (subst order_less_le)
+  apply (simp add: inv_not_p_minus_1 inv_less)
+  done
+
+
+text {* \medskip Bijection *}
+
+lemma aux1: "1 < x ==> 0 \<le> (x::int)"
+  apply auto
+  done
+
+lemma aux2: "1 < x ==> 0 < (x::int)"
+  apply auto
+  done
+
+lemma aux3: "x \<le> p - 2 ==> x < (p::int)"
+  apply auto
+  done
+
+lemma aux4: "x \<le> p - 2 ==> x < (p::int) - 1"
+  apply auto
+  done
+
+lemma inv_inj: "zprime p ==> inj_on (inv p) (d22set (p - 2))"
+  apply (unfold inj_on_def)
+  apply auto
+  apply (rule zcong_zless_imp_eq)
+      apply (tactic {* stac (thm "zcong_cancel" RS sym) 5 *})
+        apply (rule_tac [7] zcong_trans)
+         apply (tactic {* stac (thm "zcong_sym") 8 *})
+         apply (erule_tac [7] inv_is_inv)
+          apply (tactic "asm_simp_tac @{simpset} 9")
+          apply (erule_tac [9] inv_is_inv)
+           apply (rule_tac [6] zless_zprime_imp_zrelprime)
+             apply (rule_tac [8] inv_less)
+               apply (rule_tac [7] inv_g_1 [THEN aux2])
+                 apply (unfold zprime_def)
+                 apply (auto intro: d22set_g_1 d22set_le
+		   aux1 aux2 aux3 aux4)
+  done
+
+lemma inv_d22set_d22set:
+    "zprime p ==> inv p ` d22set (p - 2) = d22set (p - 2)"
+  apply (rule endo_inj_surj)
+    apply (rule d22set_fin)
+   apply (erule_tac [2] inv_inj)
+  apply auto
+  apply (rule d22set_mem)
+   apply (erule inv_g_1)
+    apply (subgoal_tac [3] "inv p xa < p - 1")
+     apply (erule_tac [4] inv_less_p_minus_1)
+      apply (auto intro: d22set_g_1 d22set_le aux4)
+  done
+
+lemma d22set_d22set_bij:
+    "zprime p ==> (d22set (p - 2), d22set (p - 2)) \<in> bijR (reciR p)"
+  apply (unfold reciR_def)
+  apply (rule_tac s = "(d22set (p - 2), inv p ` d22set (p - 2))" in subst)
+   apply (simp add: inv_d22set_d22set)
+  apply (rule inj_func_bijR)
+    apply (rule_tac [3] d22set_fin)
+   apply (erule_tac [2] inv_inj)
+  apply auto
+      apply (erule inv_is_inv)
+       apply (erule_tac [5] inv_g_1)
+        apply (erule_tac [7] inv_less_p_minus_1)
+         apply (auto intro: d22set_g_1 d22set_le aux2 aux3 aux4)
+  done
+
+lemma reciP_bijP: "zprime p ==> bijP (reciR p) (d22set (p - 2))"
+  apply (unfold reciR_def bijP_def)
+  apply auto
+  apply (rule d22set_mem)
+   apply auto
+  done
+
+lemma reciP_uniq: "zprime p ==> uniqP (reciR p)"
+  apply (unfold reciR_def uniqP_def)
+  apply auto
+   apply (rule zcong_zless_imp_eq)
+       apply (tactic {* stac (thm "zcong_cancel2" RS sym) 5 *})
+         apply (rule_tac [7] zcong_trans)
+          apply (tactic {* stac (thm "zcong_sym") 8 *})
+          apply (rule_tac [6] zless_zprime_imp_zrelprime)
+            apply auto
+  apply (rule zcong_zless_imp_eq)
+      apply (tactic {* stac (thm "zcong_cancel" RS sym) 5 *})
+        apply (rule_tac [7] zcong_trans)
+         apply (tactic {* stac (thm "zcong_sym") 8 *})
+         apply (rule_tac [6] zless_zprime_imp_zrelprime)
+           apply auto
+  done
+
+lemma reciP_sym: "zprime p ==> symP (reciR p)"
+  apply (unfold reciR_def symP_def)
+  apply (simp add: zmult_commute)
+  apply auto
+  done
+
+lemma bijER_d22set: "zprime p ==> d22set (p - 2) \<in> bijER (reciR p)"
+  apply (rule bijR_bijER)
+     apply (erule d22set_d22set_bij)
+    apply (erule reciP_bijP)
+   apply (erule reciP_uniq)
+  apply (erule reciP_sym)
+  done
+
+
+subsection {* Wilson *}
+
+lemma bijER_zcong_prod_1:
+    "zprime p ==> A \<in> bijER (reciR p) ==> [\<Prod>A = 1] (mod p)"
+  apply (unfold reciR_def)
+  apply (erule bijER.induct)
+    apply (subgoal_tac [2] "a = 1 \<or> a = p - 1")
+     apply (rule_tac [3] zcong_square_zless)
+        apply auto
+  apply (subst setprod_insert)
+    prefer 3
+    apply (subst setprod_insert)
+      apply (auto simp add: fin_bijER)
+  apply (subgoal_tac "zcong ((a * b) * \<Prod>A) (1 * 1) p")
+   apply (simp add: zmult_assoc)
+  apply (rule zcong_zmult)
+   apply auto
+  done
+
+theorem Wilson_Bij: "zprime p ==> [zfact (p - 1) = -1] (mod p)"
+  apply (subgoal_tac "zcong ((p - 1) * zfact (p - 2)) (-1 * 1) p")
+   apply (rule_tac [2] zcong_zmult)
+    apply (simp add: zprime_def)
+    apply (subst zfact.simps)
+    apply (rule_tac t = "p - 1 - 1" and s = "p - 2" in subst)
+     apply auto
+   apply (simp add: zcong_def)
+  apply (subst d22set_prod_zfact [symmetric])
+  apply (rule bijER_zcong_prod_1)
+   apply (rule_tac [2] bijER_d22set)
+   apply auto
+  done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/WilsonRuss.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,327 @@
+(*  Author:     Thomas M. Rasmussen
+    Copyright   2000  University of Cambridge
+*)
+
+header {* Wilson's Theorem according to Russinoff *}
+
+theory WilsonRuss imports EulerFermat begin
+
+text {*
+  Wilson's Theorem following quite closely Russinoff's approach
+  using Boyer-Moore (using finite sets instead of lists, though).
+*}
+
+subsection {* Definitions and lemmas *}
+
+definition
+  inv :: "int => int => int" where
+  "inv p a = (a^(nat (p - 2))) mod p"
+
+consts
+  wset :: "int * int => int set"
+
+recdef wset
+  "measure ((\<lambda>(a, p). nat a) :: int * int => nat)"
+  "wset (a, p) =
+    (if 1 < a then
+      let ws = wset (a - 1, p)
+      in (if a \<in> ws then ws else insert a (insert (inv p a) ws)) else {})"
+
+
+text {* \medskip @{term [source] inv} *}
+
+lemma inv_is_inv_aux: "1 < m ==> Suc (nat (m - 2)) = nat (m - 1)"
+by (subst int_int_eq [symmetric], auto)
+
+lemma inv_is_inv:
+    "zprime p \<Longrightarrow> 0 < a \<Longrightarrow> a < p ==> [a * inv p a = 1] (mod p)"
+  apply (unfold inv_def)
+  apply (subst zcong_zmod)
+  apply (subst zmod_zmult1_eq [symmetric])
+  apply (subst zcong_zmod [symmetric])
+  apply (subst power_Suc [symmetric])
+  apply (subst inv_is_inv_aux)
+   apply (erule_tac [2] Little_Fermat)
+   apply (erule_tac [2] zdvd_not_zless)
+   apply (unfold zprime_def, auto)
+  done
+
+lemma inv_distinct:
+    "zprime p \<Longrightarrow> 1 < a \<Longrightarrow> a < p - 1 ==> a \<noteq> inv p a"
+  apply safe
+  apply (cut_tac a = a and p = p in zcong_square)
+     apply (cut_tac [3] a = a and p = p in inv_is_inv, auto)
+   apply (subgoal_tac "a = 1")
+    apply (rule_tac [2] m = p in zcong_zless_imp_eq)
+        apply (subgoal_tac [7] "a = p - 1")
+         apply (rule_tac [8] m = p in zcong_zless_imp_eq, auto)
+  done
+
+lemma inv_not_0:
+    "zprime p \<Longrightarrow> 1 < a \<Longrightarrow> a < p - 1 ==> inv p a \<noteq> 0"
+  apply safe
+  apply (cut_tac a = a and p = p in inv_is_inv)
+     apply (unfold zcong_def, auto)
+  apply (subgoal_tac "\<not> p dvd 1")
+   apply (rule_tac [2] zdvd_not_zless)
+    apply (subgoal_tac "p dvd 1")
+     prefer 2
+     apply (subst dvd_minus_iff [symmetric], auto)
+  done
+
+lemma inv_not_1:
+    "zprime p \<Longrightarrow> 1 < a \<Longrightarrow> a < p - 1 ==> inv p a \<noteq> 1"
+  apply safe
+  apply (cut_tac a = a and p = p in inv_is_inv)
+     prefer 4
+     apply simp
+     apply (subgoal_tac "a = 1")
+      apply (rule_tac [2] zcong_zless_imp_eq, auto)
+  done
+
+lemma inv_not_p_minus_1_aux:
+    "[a * (p - 1) = 1] (mod p) = [a = p - 1] (mod p)"
+  apply (unfold zcong_def)
+  apply (simp add: OrderedGroup.diff_diff_eq diff_diff_eq2 zdiff_zmult_distrib2)
+  apply (rule_tac s = "p dvd -((a + 1) + (p * -a))" in trans)
+   apply (simp add: mult_commute)
+  apply (subst dvd_minus_iff)
+  apply (subst zdvd_reduce)
+  apply (rule_tac s = "p dvd (a + 1) + (p * -1)" in trans)
+   apply (subst zdvd_reduce, auto)
+  done
+
+lemma inv_not_p_minus_1:
+    "zprime p \<Longrightarrow> 1 < a \<Longrightarrow> a < p - 1 ==> inv p a \<noteq> p - 1"
+  apply safe
+  apply (cut_tac a = a and p = p in inv_is_inv, auto)
+  apply (simp add: inv_not_p_minus_1_aux)
+  apply (subgoal_tac "a = p - 1")
+   apply (rule_tac [2] zcong_zless_imp_eq, auto)
+  done
+
+lemma inv_g_1:
+    "zprime p \<Longrightarrow> 1 < a \<Longrightarrow> a < p - 1 ==> 1 < inv p a"
+  apply (case_tac "0\<le> inv p a")
+   apply (subgoal_tac "inv p a \<noteq> 1")
+    apply (subgoal_tac "inv p a \<noteq> 0")
+     apply (subst order_less_le)
+     apply (subst zle_add1_eq_le [symmetric])
+     apply (subst order_less_le)
+     apply (rule_tac [2] inv_not_0)
+       apply (rule_tac [5] inv_not_1, auto)
+  apply (unfold inv_def zprime_def, simp)
+  done
+
+lemma inv_less_p_minus_1:
+    "zprime p \<Longrightarrow> 1 < a \<Longrightarrow> a < p - 1 ==> inv p a < p - 1"
+  apply (case_tac "inv p a < p")
+   apply (subst order_less_le)
+   apply (simp add: inv_not_p_minus_1, auto)
+  apply (unfold inv_def zprime_def, simp)
+  done
+
+lemma inv_inv_aux: "5 \<le> p ==>
+    nat (p - 2) * nat (p - 2) = Suc (nat (p - 1) * nat (p - 3))"
+  apply (subst int_int_eq [symmetric])
+  apply (simp add: zmult_int [symmetric])
+  apply (simp add: zdiff_zmult_distrib zdiff_zmult_distrib2)
+  done
+
+lemma zcong_zpower_zmult:
+    "[x^y = 1] (mod p) \<Longrightarrow> [x^(y * z) = 1] (mod p)"
+  apply (induct z)
+   apply (auto simp add: zpower_zadd_distrib)
+  apply (subgoal_tac "zcong (x^y * x^(y * z)) (1 * 1) p")
+   apply (rule_tac [2] zcong_zmult, simp_all)
+  done
+
+lemma inv_inv: "zprime p \<Longrightarrow>
+    5 \<le> p \<Longrightarrow> 0 < a \<Longrightarrow> a < p ==> inv p (inv p a) = a"
+  apply (unfold inv_def)
+  apply (subst zpower_zmod)
+  apply (subst zpower_zpower)
+  apply (rule zcong_zless_imp_eq)
+      prefer 5
+      apply (subst zcong_zmod)
+      apply (subst mod_mod_trivial)
+      apply (subst zcong_zmod [symmetric])
+      apply (subst inv_inv_aux)
+       apply (subgoal_tac [2]
+	 "zcong (a * a^(nat (p - 1) * nat (p - 3))) (a * 1) p")
+        apply (rule_tac [3] zcong_zmult)
+         apply (rule_tac [4] zcong_zpower_zmult)
+         apply (erule_tac [4] Little_Fermat)
+         apply (rule_tac [4] zdvd_not_zless, simp_all)
+  done
+
+
+text {* \medskip @{term wset} *}
+
+declare wset.simps [simp del]
+
+lemma wset_induct:
+  assumes "!!a p. P {} a p"
+    and "!!a p. 1 < (a::int) \<Longrightarrow>
+      P (wset (a - 1, p)) (a - 1) p ==> P (wset (a, p)) a p"
+  shows "P (wset (u, v)) u v"
+  apply (rule wset.induct, safe)
+   prefer 2
+   apply (case_tac "1 < a")
+    apply (rule prems)
+     apply simp_all
+   apply (simp_all add: wset.simps prems)
+  done
+
+lemma wset_mem_imp_or [rule_format]:
+  "1 < a \<Longrightarrow> b \<notin> wset (a - 1, p)
+    ==> b \<in> wset (a, p) --> b = a \<or> b = inv p a"
+  apply (subst wset.simps)
+  apply (unfold Let_def, simp)
+  done
+
+lemma wset_mem_mem [simp]: "1 < a ==> a \<in> wset (a, p)"
+  apply (subst wset.simps)
+  apply (unfold Let_def, simp)
+  done
+
+lemma wset_subset: "1 < a \<Longrightarrow> b \<in> wset (a - 1, p) ==> b \<in> wset (a, p)"
+  apply (subst wset.simps)
+  apply (unfold Let_def, auto)
+  done
+
+lemma wset_g_1 [rule_format]:
+    "zprime p --> a < p - 1 --> b \<in> wset (a, p) --> 1 < b"
+  apply (induct a p rule: wset_induct, auto)
+  apply (case_tac "b = a")
+   apply (case_tac [2] "b = inv p a")
+    apply (subgoal_tac [3] "b = a \<or> b = inv p a")
+     apply (rule_tac [4] wset_mem_imp_or)
+       prefer 2
+       apply simp
+       apply (rule inv_g_1, auto)
+  done
+
+lemma wset_less [rule_format]:
+    "zprime p --> a < p - 1 --> b \<in> wset (a, p) --> b < p - 1"
+  apply (induct a p rule: wset_induct, auto)
+  apply (case_tac "b = a")
+   apply (case_tac [2] "b = inv p a")
+    apply (subgoal_tac [3] "b = a \<or> b = inv p a")
+     apply (rule_tac [4] wset_mem_imp_or)
+       prefer 2
+       apply simp
+       apply (rule inv_less_p_minus_1, auto)
+  done
+
+lemma wset_mem [rule_format]:
+  "zprime p -->
+    a < p - 1 --> 1 < b --> b \<le> a --> b \<in> wset (a, p)"
+  apply (induct a p rule: wset.induct, auto)
+  apply (rule_tac wset_subset)
+  apply (simp (no_asm_simp))
+  apply auto
+  done
+
+lemma wset_mem_inv_mem [rule_format]:
+  "zprime p --> 5 \<le> p --> a < p - 1 --> b \<in> wset (a, p)
+    --> inv p b \<in> wset (a, p)"
+  apply (induct a p rule: wset_induct, auto)
+   apply (case_tac "b = a")
+    apply (subst wset.simps)
+    apply (unfold Let_def)
+    apply (rule_tac [3] wset_subset, auto)
+  apply (case_tac "b = inv p a")
+   apply (simp (no_asm_simp))
+   apply (subst inv_inv)
+       apply (subgoal_tac [6] "b = a \<or> b = inv p a")
+        apply (rule_tac [7] wset_mem_imp_or, auto)
+  done
+
+lemma wset_inv_mem_mem:
+  "zprime p \<Longrightarrow> 5 \<le> p \<Longrightarrow> a < p - 1 \<Longrightarrow> 1 < b \<Longrightarrow> b < p - 1
+    \<Longrightarrow> inv p b \<in> wset (a, p) \<Longrightarrow> b \<in> wset (a, p)"
+  apply (rule_tac s = "inv p (inv p b)" and t = b in subst)
+   apply (rule_tac [2] wset_mem_inv_mem)
+      apply (rule inv_inv, simp_all)
+  done
+
+lemma wset_fin: "finite (wset (a, p))"
+  apply (induct a p rule: wset_induct)
+   prefer 2
+   apply (subst wset.simps)
+   apply (unfold Let_def, auto)
+  done
+
+lemma wset_zcong_prod_1 [rule_format]:
+  "zprime p -->
+    5 \<le> p --> a < p - 1 --> [(\<Prod>x\<in>wset(a, p). x) = 1] (mod p)"
+  apply (induct a p rule: wset_induct)
+   prefer 2
+   apply (subst wset.simps)
+   apply (unfold Let_def, auto)
+  apply (subst setprod_insert)
+    apply (tactic {* stac (thm "setprod_insert") 3 *})
+      apply (subgoal_tac [5]
+	"zcong (a * inv p a * (\<Prod>x\<in> wset(a - 1, p). x)) (1 * 1) p")
+       prefer 5
+       apply (simp add: zmult_assoc)
+      apply (rule_tac [5] zcong_zmult)
+       apply (rule_tac [5] inv_is_inv)
+         apply (tactic "clarify_tac @{claset} 4")
+         apply (subgoal_tac [4] "a \<in> wset (a - 1, p)")
+          apply (rule_tac [5] wset_inv_mem_mem)
+               apply (simp_all add: wset_fin)
+  apply (rule inv_distinct, auto)
+  done
+
+lemma d22set_eq_wset: "zprime p ==> d22set (p - 2) = wset (p - 2, p)"
+  apply safe
+   apply (erule wset_mem)
+     apply (rule_tac [2] d22set_g_1)
+     apply (rule_tac [3] d22set_le)
+     apply (rule_tac [4] d22set_mem)
+      apply (erule_tac [4] wset_g_1)
+       prefer 6
+       apply (subst zle_add1_eq_le [symmetric])
+       apply (subgoal_tac "p - 2 + 1 = p - 1")
+        apply (simp (no_asm_simp))
+        apply (erule wset_less, auto)
+  done
+
+
+subsection {* Wilson *}
+
+lemma prime_g_5: "zprime p \<Longrightarrow> p \<noteq> 2 \<Longrightarrow> p \<noteq> 3 ==> 5 \<le> p"
+  apply (unfold zprime_def dvd_def)
+  apply (case_tac "p = 4", auto)
+   apply (rule notE)
+    prefer 2
+    apply assumption
+   apply (simp (no_asm))
+   apply (rule_tac x = 2 in exI)
+   apply (safe, arith)
+     apply (rule_tac x = 2 in exI, auto)
+  done
+
+theorem Wilson_Russ:
+    "zprime p ==> [zfact (p - 1) = -1] (mod p)"
+  apply (subgoal_tac "[(p - 1) * zfact (p - 2) = -1 * 1] (mod p)")
+   apply (rule_tac [2] zcong_zmult)
+    apply (simp only: zprime_def)
+    apply (subst zfact.simps)
+    apply (rule_tac t = "p - 1 - 1" and s = "p - 2" in subst, auto)
+   apply (simp only: zcong_def)
+   apply (simp (no_asm_simp))
+  apply (case_tac "p = 2")
+   apply (simp add: zfact.simps)
+  apply (case_tac "p = 3")
+   apply (simp add: zfact.simps)
+  apply (subgoal_tac "5 \<le> p")
+   apply (erule_tac [2] prime_g_5)
+    apply (subst d22set_prod_zfact [symmetric])
+    apply (subst d22set_eq_wset)
+     apply (rule_tac [2] wset_zcong_prod_1, auto)
+  done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Old_Number_Theory/document/root.tex	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,57 @@
+
+\documentclass[11pt,a4paper]{article}
+\usepackage{graphicx}
+\usepackage{isabelle,isabellesym,pdfsetup}
+\usepackage[latin1]{inputenc}
+
+\urlstyle{rm}
+\isabellestyle{it}
+
+\begin{document}
+
+\title{Some results of number theory}
+\author{Jeremy Avigad\\
+    David Gray\\
+    Adam Kramer\\
+    Thomas M Rasmussen}
+
+\maketitle
+
+\begin{abstract}
+This is a collection of formalized proofs of many results of number theory.
+The proofs of the Chinese Remainder Theorem and Wilson's Theorem are due to
+Rasmussen.  The proof of Gauss's law of quadratic reciprocity is due to
+Avigad, Gray and Kramer.  Proofs can be found in most introductory number
+theory textbooks; Goldman's \emph{The Queen of Mathematics: a Historically
+Motivated Guide to Number Theory} provides some historical context.
+
+Avigad, Gray and Kramer have also provided library theories dealing with
+finite sets and finite sums, divisibility and congruences, parity and
+residues.  The authors are engaged in redesigning and polishing these theories
+for more serious use.  For the latest information in this respect, please see
+the web page \url{http://www.andrew.cmu.edu/~avigad/isabelle}.  Other theories
+contain proofs of Euler's criteria, Gauss' lemma, and the law of quadratic
+reciprocity.  The formalization follows Eisenstein's proof, which is the one
+most commonly found in introductory textbooks; in particular, it follows the
+presentation in Niven and Zuckerman, \emph{The Theory of Numbers}.
+
+To avoid having to count roots of polynomials, however, we relied on a trick
+previously used by David Russinoff in formalizing quadratic reciprocity for
+the Boyer-Moore theorem prover; see Russinoff, David, ``A mechanical proof
+of quadratic reciprocity,'' \emph{Journal of Automated Reasoning} 8:3-21,
+1992.  We are grateful to Larry Paulson for calling our attention to this
+reference.
+\end{abstract}
+
+\tableofcontents
+
+\begin{center}
+  \includegraphics[scale=0.5]{session_graph}  
+\end{center}
+
+\newpage
+
+\parindent 0pt\parskip 0.5ex
+\input{session}
+
+\end{document}
--- a/src/HOL/OrderedGroup.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/OrderedGroup.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -1075,16 +1075,17 @@
 lemma nprt_0[simp]: "nprt 0 = 0" by (simp add: nprt_def)
 
 lemma pprt_eq_id [simp, noatp]: "0 \<le> x \<Longrightarrow> pprt x = x"
-by (simp add: pprt_def le_iff_sup sup_aci)
+by (simp add: pprt_def sup_aci)
+
 
 lemma nprt_eq_id [simp, noatp]: "x \<le> 0 \<Longrightarrow> nprt x = x"
-by (simp add: nprt_def le_iff_inf inf_aci)
+by (simp add: nprt_def inf_aci)
 
 lemma pprt_eq_0 [simp, noatp]: "x \<le> 0 \<Longrightarrow> pprt x = 0"
-by (simp add: pprt_def le_iff_sup sup_aci)
+by (simp add: pprt_def sup_aci)
 
 lemma nprt_eq_0 [simp, noatp]: "0 \<le> x \<Longrightarrow> nprt x = 0"
-by (simp add: nprt_def le_iff_inf inf_aci)
+by (simp add: nprt_def inf_aci)
 
 lemma sup_0_imp_0: "sup a (- a) = 0 \<Longrightarrow> a = 0"
 proof -
@@ -1118,13 +1119,13 @@
   "0 \<le> a + a \<longleftrightarrow> 0 \<le> a"
 proof
   assume "0 <= a + a"
-  hence a:"inf (a+a) 0 = 0" by (simp add: le_iff_inf inf_commute)
+  hence a:"inf (a+a) 0 = 0" by (simp add: inf_commute)
   have "(inf a 0)+(inf a 0) = inf (inf (a+a) 0) a" (is "?l=_")
     by (simp add: add_sup_inf_distribs inf_aci)
   hence "?l = 0 + inf a 0" by (simp add: a, simp add: inf_commute)
   hence "inf a 0 = 0" by (simp only: add_right_cancel)
-  then show "0 <= a" by (simp add: le_iff_inf inf_commute)    
-next  
+  then show "0 <= a" unfolding le_iff_inf by (simp add: inf_commute)
+next
   assume a: "0 <= a"
   show "0 <= a + a" by (simp add: add_mono[OF a a, simplified])
 qed
@@ -1194,22 +1195,22 @@
 qed
 
 lemma zero_le_iff_zero_nprt: "0 \<le> a \<longleftrightarrow> nprt a = 0"
-by (simp add: le_iff_inf nprt_def inf_commute)
+unfolding le_iff_inf by (simp add: nprt_def inf_commute)
 
 lemma le_zero_iff_zero_pprt: "a \<le> 0 \<longleftrightarrow> pprt a = 0"
-by (simp add: le_iff_sup pprt_def sup_commute)
+unfolding le_iff_sup by (simp add: pprt_def sup_commute)
 
 lemma le_zero_iff_pprt_id: "0 \<le> a \<longleftrightarrow> pprt a = a"
-by (simp add: le_iff_sup pprt_def sup_commute)
+unfolding le_iff_sup by (simp add: pprt_def sup_commute)
 
 lemma zero_le_iff_nprt_id: "a \<le> 0 \<longleftrightarrow> nprt a = a"
-by (simp add: le_iff_inf nprt_def inf_commute)
+unfolding le_iff_inf by (simp add: nprt_def inf_commute)
 
 lemma pprt_mono [simp, noatp]: "a \<le> b \<Longrightarrow> pprt a \<le> pprt b"
-by (simp add: le_iff_sup pprt_def sup_aci sup_assoc [symmetric, of a])
+unfolding le_iff_sup by (simp add: pprt_def sup_aci sup_assoc [symmetric, of a])
 
 lemma nprt_mono [simp, noatp]: "a \<le> b \<Longrightarrow> nprt a \<le> nprt b"
-by (simp add: le_iff_inf nprt_def inf_aci inf_assoc [symmetric, of a])
+unfolding le_iff_inf by (simp add: nprt_def inf_aci inf_assoc [symmetric, of a])
 
 end
 
@@ -1274,7 +1275,7 @@
 proof -
   note add_le_cancel_right [of a a "- a", symmetric, simplified]
   moreover note add_le_cancel_right [of "-a" a a, symmetric, simplified]
-  then show ?thesis by (auto simp: sup_max max_def)
+  then show ?thesis by (auto simp: sup_max)
 qed
 
 lemma abs_if_lattice:
--- a/src/HOL/Recdef.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Recdef.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -19,6 +19,65 @@
   ("Tools/recdef.ML")
 begin
 
+inductive
+  wfrec_rel :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b => bool"
+  for R :: "('a * 'a) set"
+  and F :: "('a => 'b) => 'a => 'b"
+where
+  wfrecI: "ALL z. (z, x) : R --> wfrec_rel R F z (g z) ==>
+            wfrec_rel R F x (F g x)"
+
+constdefs
+  cut        :: "('a => 'b) => ('a * 'a)set => 'a => 'a => 'b"
+  "cut f r x == (%y. if (y,x):r then f y else undefined)"
+
+  adm_wf :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => bool"
+  "adm_wf R F == ALL f g x.
+     (ALL z. (z, x) : R --> f z = g z) --> F f x = F g x"
+
+  wfrec :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b"
+  [code del]: "wfrec R F == %x. THE y. wfrec_rel R (%f x. F (cut f R x) x) x y"
+
+subsection{*Well-Founded Recursion*}
+
+text{*cut*}
+
+lemma cuts_eq: "(cut f r x = cut g r x) = (ALL y. (y,x):r --> f(y)=g(y))"
+by (simp add: expand_fun_eq cut_def)
+
+lemma cut_apply: "(x,a):r ==> (cut f r a)(x) = f(x)"
+by (simp add: cut_def)
+
+text{*Inductive characterization of wfrec combinator; for details see:
+John Harrison, "Inductive definitions: automation and application"*}
+
+lemma wfrec_unique: "[| adm_wf R F; wf R |] ==> EX! y. wfrec_rel R F x y"
+apply (simp add: adm_wf_def)
+apply (erule_tac a=x in wf_induct)
+apply (rule ex1I)
+apply (rule_tac g = "%x. THE y. wfrec_rel R F x y" in wfrec_rel.wfrecI)
+apply (fast dest!: theI')
+apply (erule wfrec_rel.cases, simp)
+apply (erule allE, erule allE, erule allE, erule mp)
+apply (fast intro: the_equality [symmetric])
+done
+
+lemma adm_lemma: "adm_wf R (%f x. F (cut f R x) x)"
+apply (simp add: adm_wf_def)
+apply (intro strip)
+apply (rule cuts_eq [THEN iffD2, THEN subst], assumption)
+apply (rule refl)
+done
+
+lemma wfrec: "wf(r) ==> wfrec r H a = H (cut (wfrec r H) r a) a"
+apply (simp add: wfrec_def)
+apply (rule adm_lemma [THEN wfrec_unique, THEN the1_equality], assumption)
+apply (rule wfrec_rel.wfrecI)
+apply (intro strip)
+apply (erule adm_lemma [THEN wfrec_unique, THEN theI'])
+done
+
+
 text{** This form avoids giant explosions in proofs.  NOTE USE OF ==*}
 lemma def_wfrec: "[| f==wfrec r H;  wf(r) |] ==> f(a) = H (cut f r a) a"
 apply auto
--- a/src/HOL/Relation.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Relation.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -599,6 +599,9 @@
   apply blast
   done
 
+lemma in_inv_image[simp]: "((x,y) : inv_image r f) = ((f x, f y) : r)"
+  by (auto simp:inv_image_def)
+
 
 subsection {* Finiteness *}
 
--- a/src/HOL/SEQ.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/SEQ.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -582,7 +582,7 @@
       ultimately
       have "a (max no n) < a n" by auto
       with monotone[where m=n and n="max no n"]
-      show False by auto
+      show False by (auto simp:max_def split:split_if_asm)
     qed
   } note top_down = this
   { fix x n m fix a :: "nat \<Rightarrow> real"
--- a/src/HOL/SET-Protocol/Cardholder_Registration.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/SET-Protocol/Cardholder_Registration.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -715,6 +715,7 @@
       ==> P --> (Nonce N \<in> analz (Key`KK Un H)) = (Nonce N \<in> analz H)"
 by (blast intro: analz_mono [THEN [2] rev_subsetD])
 
+
 text{*The @{text "(no_asm)"} attribute is essential, since it retains
   the quantifier and allows the simprule's condition to itself be simplified.*}
 lemma Nonce_compromise [rule_format (no_asm)]:
@@ -741,12 +742,11 @@
 apply blast  --{*3*}
 apply blast  --{*5*}
 txt{*Message 6*}
-apply (force del: allE ballE impCE simp add: symKey_compromise)
+apply (metis symKey_compromise)
   --{*cardSK compromised*}
 txt{*Simplify again--necessary because the previous simplification introduces
-  some logical connectives*}
-apply (force del: allE ballE impCE
-          simp del: image_insert image_Un imp_disjL
+  some logical connectives*} 
+apply (force simp del: image_insert image_Un imp_disjL
           simp add: analz_image_keys_simps symKey_compromise)
 done
 
--- a/src/HOL/SET-Protocol/Purchase.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/SET-Protocol/Purchase.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -1040,9 +1040,8 @@
 apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
 apply simp_all
 apply blast
-apply (force dest!: signed_Hash_imp_used)
-apply (clarify) --{*speeds next step*}
-apply (blast dest: unique_LID_M)
+apply (metis subsetD insert_subset parts.Fst parts_increasing signed_Hash_imp_used)
+apply (metis unique_LID_M)
 apply (blast dest!: Notes_Cardholder_self_False)
 done
 
--- a/src/HOL/Set.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Set.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -1268,10 +1268,26 @@
     "(insert a B) Int C = (if a \<in> C then insert a (B \<inter> C) else B \<inter> C)"
   by auto
 
+lemma Int_insert_left_if0[simp]:
+    "a \<notin> C \<Longrightarrow> (insert a B) Int C = B \<inter> C"
+  by auto
+
+lemma Int_insert_left_if1[simp]:
+    "a \<in> C \<Longrightarrow> (insert a B) Int C = insert a (B Int C)"
+  by auto
+
 lemma Int_insert_right:
     "A \<inter> (insert a B) = (if a \<in> A then insert a (A \<inter> B) else A \<inter> B)"
   by auto
 
+lemma Int_insert_right_if0[simp]:
+    "a \<notin> A \<Longrightarrow> A Int (insert a B) = A Int B"
+  by auto
+
+lemma Int_insert_right_if1[simp]:
+    "a \<in> A \<Longrightarrow> A Int (insert a B) = insert a (A Int B)"
+  by auto
+
 lemma Un_Int_distrib: "A \<union> (B \<inter> C) = (A \<union> B) \<inter> (A \<union> C)"
   by blast
 
--- a/src/HOL/SetInterval.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/SetInterval.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -181,33 +181,108 @@
   "(i : {l..u}) = (l <= i & i <= u)"
 by (simp add: atLeastAtMost_def)
 
-text {* The above four lemmas could be declared as iffs.
-  If we do so, a call to blast in Hyperreal/Star.ML, lemma @{text STAR_Int}
-  seems to take forever (more than one hour). *}
+text {* The above four lemmas could be declared as iffs. Unfortunately this
+breaks many proofs. Since it only helps blast, it is better to leave well
+alone *}
+
 end
 
-subsubsection{* Emptyness and singletons *}
+subsubsection{* Emptyness, singletons, subset *}
 
 context order
 begin
 
-lemma atLeastAtMost_empty [simp]: "n < m ==> {m..n} = {}";
-by (auto simp add: atLeastAtMost_def atMost_def atLeast_def)
+lemma atLeastatMost_empty[simp]:
+  "b < a \<Longrightarrow> {a..b} = {}"
+by(auto simp: atLeastAtMost_def atLeast_def atMost_def)
+
+lemma atLeastatMost_empty_iff[simp]:
+  "{a..b} = {} \<longleftrightarrow> (~ a <= b)"
+by auto (blast intro: order_trans)
+
+lemma atLeastatMost_empty_iff2[simp]:
+  "{} = {a..b} \<longleftrightarrow> (~ a <= b)"
+by auto (blast intro: order_trans)
+
+lemma atLeastLessThan_empty[simp]:
+  "b <= a \<Longrightarrow> {a..<b} = {}"
+by(auto simp: atLeastLessThan_def)
 
-lemma atLeastLessThan_empty[simp]: "n \<le> m ==> {m..<n} = {}"
-by (auto simp add: atLeastLessThan_def)
+lemma atLeastLessThan_empty_iff[simp]:
+  "{a..<b} = {} \<longleftrightarrow> (~ a < b)"
+by auto (blast intro: le_less_trans)
+
+lemma atLeastLessThan_empty_iff2[simp]:
+  "{} = {a..<b} \<longleftrightarrow> (~ a < b)"
+by auto (blast intro: le_less_trans)
 
-lemma greaterThanAtMost_empty[simp]:"l \<le> k ==> {k<..l} = {}"
+lemma greaterThanAtMost_empty[simp]: "l \<le> k ==> {k<..l} = {}"
 by(auto simp:greaterThanAtMost_def greaterThan_def atMost_def)
 
+lemma greaterThanAtMost_empty_iff[simp]: "{k<..l} = {} \<longleftrightarrow> ~ k < l"
+by auto (blast intro: less_le_trans)
+
+lemma greaterThanAtMost_empty_iff2[simp]: "{} = {k<..l} \<longleftrightarrow> ~ k < l"
+by auto (blast intro: less_le_trans)
+
 lemma greaterThanLessThan_empty[simp]:"l \<le> k ==> {k<..<l} = {}"
 by(auto simp:greaterThanLessThan_def greaterThan_def lessThan_def)
 
 lemma atLeastAtMost_singleton [simp]: "{a..a} = {a}"
 by (auto simp add: atLeastAtMost_def atMost_def atLeast_def)
 
+lemma atLeastatMost_subset_iff[simp]:
+  "{a..b} <= {c..d} \<longleftrightarrow> (~ a <= b) | c <= a & b <= d"
+unfolding atLeastAtMost_def atLeast_def atMost_def
+by (blast intro: order_trans)
+
+lemma atLeastatMost_psubset_iff:
+  "{a..b} < {c..d} \<longleftrightarrow>
+   ((~ a <= b) | c <= a & b <= d & (c < a | b < d))  &  c <= d"
+by(simp add: psubset_eq expand_set_eq less_le_not_le)(blast intro: order_trans)
+
 end
 
+lemma (in linorder) atLeastLessThan_subset_iff:
+  "{a..<b} <= {c..<d} \<Longrightarrow> b <= a | c<=a & b<=d"
+apply (auto simp:subset_eq Ball_def)
+apply(frule_tac x=a in spec)
+apply(erule_tac x=d in allE)
+apply (simp add: less_imp_le)
+done
+
+subsubsection {* Intersection *}
+
+context linorder
+begin
+
+lemma Int_atLeastAtMost[simp]: "{a..b} Int {c..d} = {max a c .. min b d}"
+by auto
+
+lemma Int_atLeastAtMostR1[simp]: "{..b} Int {c..d} = {c .. min b d}"
+by auto
+
+lemma Int_atLeastAtMostR2[simp]: "{a..} Int {c..d} = {max a c .. d}"
+by auto
+
+lemma Int_atLeastAtMostL1[simp]: "{a..b} Int {..d} = {a .. min b d}"
+by auto
+
+lemma Int_atLeastAtMostL2[simp]: "{a..b} Int {c..} = {max a c .. b}"
+by auto
+
+lemma Int_atLeastLessThan[simp]: "{a..<b} Int {c..<d} = {max a c ..< min b d}"
+by auto
+
+lemma Int_greaterThanAtMost[simp]: "{a<..b} Int {c<..d} = {max a c <.. min b d}"
+by auto
+
+lemma Int_greaterThanLessThan[simp]: "{a<..<b} Int {c<..<d} = {max a c <..< min b d}"
+by auto
+
+end
+
+
 subsection {* Intervals of natural numbers *}
 
 subsubsection {* The Constant @{term lessThan} *}
@@ -662,17 +737,6 @@
 
 subsubsection {* Disjoint Intersections *}
 
-text {* Singletons and open intervals *}
-
-lemma ivl_disj_int_singleton:
-  "{l::'a::order} Int {l<..} = {}"
-  "{..<u} Int {u} = {}"
-  "{l} Int {l<..<u} = {}"
-  "{l<..<u} Int {u} = {}"
-  "{l} Int {l<..u} = {}"
-  "{l..<u} Int {u} = {}"
-  by simp+
-
 text {* One- and two-sided intervals *}
 
 lemma ivl_disj_int_one:
@@ -699,7 +763,7 @@
   "{l..m} Int {m<..u} = {}"
   by auto
 
-lemmas ivl_disj_int = ivl_disj_int_singleton ivl_disj_int_one ivl_disj_int_two
+lemmas ivl_disj_int = ivl_disj_int_one ivl_disj_int_two
 
 subsubsection {* Some Differences *}
 
--- a/src/HOL/Statespace/state_space.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Statespace/state_space.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -567,8 +567,8 @@
           (case distinct (snd_eq) (filter (curry fst_eq (n,T)) raw_parent_comps) of
              []  => []
            | [_] => []
-           | rs  => ["Different types for component " ^ n ^": " ^ commas
-                       (map (Pretty.string_of o Display.pretty_ctyp o ctyp_of thy o snd) rs)])
+           | rs  => ["Different types for component " ^ n ^": " ^
+                commas (map (Syntax.string_of_typ ctxt o snd) rs)])
 
     val err_dup_types = List.concat (map check_type (duplicates fst_eq raw_parent_comps))
 
--- a/src/HOL/Tools/ATP_Manager/SystemOnTPTP	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Tools/ATP_Manager/SystemOnTPTP	Fri Sep 11 15:57:16 2009 +1000
@@ -19,6 +19,7 @@
     "QuietFlag" => "-q01",
     "SubmitButton" => "RunSelectedSystems",
     "ProblemSource" => "UPLOAD",
+    "ForceSystem" => "-force",
     );
 
 #----Get format and transform options if specified
--- a/src/HOL/Tools/ATP_Manager/atp_manager.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Tools/ATP_Manager/atp_manager.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -24,7 +24,7 @@
   type prover = int -> (thm * (string * int)) list option ->
     (thm * (string * int)) list option -> string -> int ->
     Proof.context * (thm list * thm) ->
-    bool * string * string * string vector * (thm * (string * int)) list
+    bool * (string * string list) * int * string * string vector * (thm * (string * int)) list
   val add_prover: string -> prover -> theory -> theory
   val print_provers: theory -> unit
   val get_prover: string -> theory -> prover option
@@ -305,7 +305,7 @@
 type prover = int -> (thm * (string * int)) list option ->
   (thm * (string * int)) list option -> string -> int ->
   Proof.context * (thm list * thm) ->
-  bool * string * string * string vector * (thm * (string * int)) list
+  bool * (string * string list) * int * string * string vector * (thm * (string * int)) list
 
 fun err_dup_prover name = error ("Duplicate prover: " ^ quote name);
 
@@ -345,7 +345,7 @@
           let
             val _ = register birthtime deadtime (Thread.self (), desc)
             val result =
-              let val (success, message, _, _, _) =
+              let val (success, (message, _), _, _, _, _) =
                 prover (get_timeout ()) NONE NONE name i (Proof.get_goal proof_state)
               in (success, message) end
               handle ResHolClause.TOO_TRIVIAL
--- a/src/HOL/Tools/ATP_Manager/atp_minimal.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Tools/ATP_Manager/atp_minimal.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -4,7 +4,13 @@
 Minimalization of theorem list for metis by using an external automated theorem prover
 *)
 
-structure AtpMinimal: sig end =
+signature ATP_MINIMAL =
+sig
+  val minimalize: AtpManager.prover -> string -> int -> Proof.state ->
+    (string * thm list) list -> (string * thm list) list option * string
+end
+
+structure AtpMinimal: ATP_MINIMAL =
 struct
 
 (* output control *)
@@ -62,7 +68,10 @@
    @post v subset s & p(v) &
          forall e in v. ~p(v \ e)
    *)
-  fun minimal p s = min p [] s
+  fun minimal p s =
+    case min p [] s of
+      [x] => if p [] then [] else [x]
+    | m => m
 end
 
 
@@ -83,7 +92,7 @@
    ("# Cannot determine problem status within resource limit", Timeout),
    ("Error", Error)]
 
-fun produce_answer (success, message, result_string, thm_name_vec, filtered) =
+fun produce_answer (success, message, _, result_string, thm_name_vec, filtered) =
   if success then
     (Success (Vector.foldr op:: [] thm_name_vec, filtered), result_string)
   else
@@ -103,8 +112,7 @@
 fun sh_test_thms prover prover_name time_limit subgoalno state filtered name_thms_pairs =
   let
     val _ = println ("Testing " ^ (length_string name_thms_pairs) ^ " theorems... ")
-    val name_thm_pairs =
-      flat (map (fn (n, ths) => map_index (fn (i, th) => (n, th)) ths) name_thms_pairs)
+    val name_thm_pairs = maps (fn (n, ths) => map (pair n) ths) name_thms_pairs
     val _ = debug_fn (fn () => print_names name_thm_pairs)
     val axclauses = ResAxioms.cnf_rules_pairs (Proof.theory_of state) name_thm_pairs
     val (result, proof) =
@@ -130,6 +138,7 @@
     val test_thms_fun = sh_test_thms prover prover_name time_limit 1 state
     fun test_thms filtered thms =
       case test_thms_fun filtered thms of (Success _, _) => true | _ => false
+    val answer' = pair and answer'' = pair NONE
   in
     (* try prove first to check result and get used theorems *)
     (case test_thms_fun NONE name_thms_pairs of
@@ -141,25 +150,26 @@
               filter (fn (name1, _) => List.exists (equal name1) ordered_used) name_thms_pairs
             else
               name_thms_pairs
-          val min_thms = (minimal (test_thms (SOME filtered)) to_use)
+          val min_thms = if null to_use then []
+            else minimal (test_thms (SOME filtered)) to_use
           val min_names = order_unique (map fst min_thms)
           val _ = println ("Minimal " ^ (length_string min_thms) ^ " theorems")
           val _ = debug_fn (fn () => print_names min_thms)
         in
-          answer ("Try this command: " ^
+          answer' (SOME min_thms) ("Try this command: " ^
             Markup.markup Markup.sendback ("apply (metis " ^ space_implode " " min_names ^ ")"))
         end
     | (Timeout, _) =>
-        answer ("Timeout: You may need to increase the time limit of " ^
+        answer'' ("Timeout: You may need to increase the time limit of " ^
           Int.toString time_limit ^ " seconds. Call atp_minimize [time=...] ")
     | (Error, msg) =>
-        answer ("Error in prover: " ^ msg)
+        answer'' ("Error in prover: " ^ msg)
     | (Failure, _) =>
-        answer "Failure: No proof with the theorems supplied")
+        answer'' "Failure: No proof with the theorems supplied")
     handle ResHolClause.TOO_TRIVIAL =>
-        answer ("Trivial: Try this command: " ^ Markup.markup Markup.sendback "apply metis")
+        answer' (SOME []) ("Trivial: Try this command: " ^ Markup.markup Markup.sendback "apply metis")
     | ERROR msg =>
-        answer ("Error: " ^ msg)
+        answer'' ("Error: " ^ msg)
   end
 
 
@@ -204,8 +214,9 @@
         SOME prover => prover
       | NONE => error ("Unknown prover: " ^ quote prover_name)
     val name_thms_pairs = get_thms (Proof.context_of state) thm_names
+    fun print_answer (_, msg) = answer msg
   in
-    minimalize prover prover_name time_limit state name_thms_pairs
+    print_answer (minimalize prover prover_name time_limit state name_thms_pairs)
   end
 
 val parse_args = Scan.optional (Args.bracks (P.list (P.xname --| P.$$$ "=" -- P.xname))) []
--- a/src/HOL/Tools/ATP_Manager/atp_wrapper.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Tools/ATP_Manager/atp_wrapper.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -41,6 +41,12 @@
 
 (* basic template *)
 
+fun with_path cleanup after f path =
+  Exn.capture f path
+  |> tap (fn _ => cleanup path)
+  |> Exn.release
+  |> tap (after path)
+
 fun external_prover relevance_filter preparer writer (cmd, args) find_failure produce_answer
   timeout axiom_clauses filtered_clauses name subgoalno goal =
   let
@@ -73,28 +79,41 @@
       preparer goal_cls chain_ths the_axiom_clauses the_filtered_clauses thy
 
     (* write out problem file and call prover *)
-    val probfile = prob_pathname subgoalno
-    val conj_pos = writer probfile clauses
-    val (proof, rc) = system_out (
-      if File.exists cmd then
-        space_implode " " ["exec", File.shell_path cmd, args, File.platform_path probfile]
-      else error ("Bad executable: " ^ Path.implode cmd))
+    fun cmd_line probfile = "TIMEFORMAT='%3U'; ((time " ^ space_implode " "
+      [File.shell_path cmd, args, File.platform_path probfile] ^ ") 2>&1)"
+    fun split_time s =
+      let
+        val split = String.tokens (fn c => str c = "\n")
+        val (proof, t) = s |> split |> split_last |> apfst cat_lines
+        val num = Scan.many1 Symbol.is_ascii_digit >> (fst o read_int)
+        val time = num --| Scan.$$ "." -- num >> (fn (a, b) => a * 1000 + b)
+        val as_time = the_default 0 o Scan.read Symbol.stopper time o explode
+      in (proof, as_time t) end
+    fun run_on probfile =
+      if File.exists cmd
+      then
+        writer probfile clauses
+        |> pair (apfst split_time (system_out (cmd_line probfile)))
+      else error ("Bad executable: " ^ Path.implode cmd)
 
     (* if problemfile has not been exported, delete problemfile; otherwise export proof, too *)
-    val _ =
-      if destdir' = "" then File.rm probfile
+    fun cleanup probfile = if destdir' = "" then File.rm probfile else ()
+    fun export probfile (((proof, _), _), _) = if destdir' = "" then ()
       else File.write (Path.explode (Path.implode probfile ^ "_proof")) proof
 
+    val (((proof, time), rc), conj_pos) = with_path cleanup export run_on
+      (prob_pathname subgoalno)
+
     (* check for success and print out some information on failure *)
     val failure = find_failure proof
     val success = rc = 0 andalso is_none failure
     val message =
-      if is_some failure then "External prover failed."
-      else if rc <> 0 then "External prover failed: " ^ proof
-      else "Try this command: " ^
-        produce_answer name (proof, thm_names, conj_pos, ctxt, th, subgoalno)
+      if is_some failure then ("External prover failed.", [])
+      else if rc <> 0 then ("External prover failed: " ^ proof, [])
+      else apfst (fn s => "Try this command: " ^ s)
+        (produce_answer name (proof, thm_names, conj_pos, ctxt, th, subgoalno))
     val _ = Output.debug (fn () => "Sledgehammer response (rc = " ^ string_of_int rc ^ "):\n" ^ proof)
-  in (success, message, proof, thm_names, the_filtered_clauses) end;
+  in (success, message, time, proof, thm_names, the_filtered_clauses) end;
 
 
 
--- a/src/HOL/Tools/ComputeFloat.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,568 +0,0 @@
-(*  Title:  HOL/Tools/ComputeFloat.thy
-    Author: Steven Obua
-*)
-
-header {* Floating Point Representation of the Reals *}
-
-theory ComputeFloat
-imports Complex_Main
-uses "~~/src/Tools/float.ML" ("~~/src/HOL/Tools/float_arith.ML")
-begin
-
-definition
-  pow2 :: "int \<Rightarrow> real" where
-  "pow2 a = (if (0 <= a) then (2^(nat a)) else (inverse (2^(nat (-a)))))"
-
-definition
-  float :: "int * int \<Rightarrow> real" where
-  "float x = real (fst x) * pow2 (snd x)"
-
-lemma pow2_0[simp]: "pow2 0 = 1"
-by (simp add: pow2_def)
-
-lemma pow2_1[simp]: "pow2 1 = 2"
-by (simp add: pow2_def)
-
-lemma pow2_neg: "pow2 x = inverse (pow2 (-x))"
-by (simp add: pow2_def)
-
-lemma pow2_add1: "pow2 (1 + a) = 2 * (pow2 a)"
-proof -
-  have h: "! n. nat (2 + int n) - Suc 0 = nat (1 + int n)" by arith
-  have g: "! a b. a - -1 = a + (1::int)" by arith
-  have pos: "! n. pow2 (int n + 1) = 2 * pow2 (int n)"
-    apply (auto, induct_tac n)
-    apply (simp_all add: pow2_def)
-    apply (rule_tac m1="2" and n1="nat (2 + int na)" in ssubst[OF realpow_num_eq_if])
-    by (auto simp add: h)
-  show ?thesis
-  proof (induct a)
-    case (1 n)
-    from pos show ?case by (simp add: algebra_simps)
-  next
-    case (2 n)
-    show ?case
-      apply (auto)
-      apply (subst pow2_neg[of "- int n"])
-      apply (subst pow2_neg[of "-1 - int n"])
-      apply (auto simp add: g pos)
-      done
-  qed
-qed
-
-lemma pow2_add: "pow2 (a+b) = (pow2 a) * (pow2 b)"
-proof (induct b)
-  case (1 n)
-  show ?case
-  proof (induct n)
-    case 0
-    show ?case by simp
-  next
-    case (Suc m)
-    show ?case by (auto simp add: algebra_simps pow2_add1 prems)
-  qed
-next
-  case (2 n)
-  show ?case
-  proof (induct n)
-    case 0
-    show ?case
-      apply (auto)
-      apply (subst pow2_neg[of "a + -1"])
-      apply (subst pow2_neg[of "-1"])
-      apply (simp)
-      apply (insert pow2_add1[of "-a"])
-      apply (simp add: algebra_simps)
-      apply (subst pow2_neg[of "-a"])
-      apply (simp)
-      done
-    case (Suc m)
-    have a: "int m - (a + -2) =  1 + (int m - a + 1)" by arith
-    have b: "int m - -2 = 1 + (int m + 1)" by arith
-    show ?case
-      apply (auto)
-      apply (subst pow2_neg[of "a + (-2 - int m)"])
-      apply (subst pow2_neg[of "-2 - int m"])
-      apply (auto simp add: algebra_simps)
-      apply (subst a)
-      apply (subst b)
-      apply (simp only: pow2_add1)
-      apply (subst pow2_neg[of "int m - a + 1"])
-      apply (subst pow2_neg[of "int m + 1"])
-      apply auto
-      apply (insert prems)
-      apply (auto simp add: algebra_simps)
-      done
-  qed
-qed
-
-lemma "float (a, e) + float (b, e) = float (a + b, e)"
-by (simp add: float_def algebra_simps)
-
-definition
-  int_of_real :: "real \<Rightarrow> int" where
-  "int_of_real x = (SOME y. real y = x)"
-
-definition
-  real_is_int :: "real \<Rightarrow> bool" where
-  "real_is_int x = (EX (u::int). x = real u)"
-
-lemma real_is_int_def2: "real_is_int x = (x = real (int_of_real x))"
-by (auto simp add: real_is_int_def int_of_real_def)
-
-lemma float_transfer: "real_is_int ((real a)*(pow2 c)) \<Longrightarrow> float (a, b) = float (int_of_real ((real a)*(pow2 c)), b - c)"
-by (simp add: float_def real_is_int_def2 pow2_add[symmetric])
-
-lemma pow2_int: "pow2 (int c) = 2^c"
-by (simp add: pow2_def)
-
-lemma float_transfer_nat: "float (a, b) = float (a * 2^c, b - int c)"
-by (simp add: float_def pow2_int[symmetric] pow2_add[symmetric])
-
-lemma real_is_int_real[simp]: "real_is_int (real (x::int))"
-by (auto simp add: real_is_int_def int_of_real_def)
-
-lemma int_of_real_real[simp]: "int_of_real (real x) = x"
-by (simp add: int_of_real_def)
-
-lemma real_int_of_real[simp]: "real_is_int x \<Longrightarrow> real (int_of_real x) = x"
-by (auto simp add: int_of_real_def real_is_int_def)
-
-lemma real_is_int_add_int_of_real: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> (int_of_real (a+b)) = (int_of_real a) + (int_of_real b)"
-by (auto simp add: int_of_real_def real_is_int_def)
-
-lemma real_is_int_add[simp]: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> real_is_int (a+b)"
-apply (subst real_is_int_def2)
-apply (simp add: real_is_int_add_int_of_real real_int_of_real)
-done
-
-lemma int_of_real_sub: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> (int_of_real (a-b)) = (int_of_real a) - (int_of_real b)"
-by (auto simp add: int_of_real_def real_is_int_def)
-
-lemma real_is_int_sub[simp]: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> real_is_int (a-b)"
-apply (subst real_is_int_def2)
-apply (simp add: int_of_real_sub real_int_of_real)
-done
-
-lemma real_is_int_rep: "real_is_int x \<Longrightarrow> ?! (a::int). real a = x"
-by (auto simp add: real_is_int_def)
-
-lemma int_of_real_mult:
-  assumes "real_is_int a" "real_is_int b"
-  shows "(int_of_real (a*b)) = (int_of_real a) * (int_of_real b)"
-proof -
-  from prems have a: "?! (a'::int). real a' = a" by (rule_tac real_is_int_rep, auto)
-  from prems have b: "?! (b'::int). real b' = b" by (rule_tac real_is_int_rep, auto)
-  from a obtain a'::int where a':"a = real a'" by auto
-  from b obtain b'::int where b':"b = real b'" by auto
-  have r: "real a' * real b' = real (a' * b')" by auto
-  show ?thesis
-    apply (simp add: a' b')
-    apply (subst r)
-    apply (simp only: int_of_real_real)
-    done
-qed
-
-lemma real_is_int_mult[simp]: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> real_is_int (a*b)"
-apply (subst real_is_int_def2)
-apply (simp add: int_of_real_mult)
-done
-
-lemma real_is_int_0[simp]: "real_is_int (0::real)"
-by (simp add: real_is_int_def int_of_real_def)
-
-lemma real_is_int_1[simp]: "real_is_int (1::real)"
-proof -
-  have "real_is_int (1::real) = real_is_int(real (1::int))" by auto
-  also have "\<dots> = True" by (simp only: real_is_int_real)
-  ultimately show ?thesis by auto
-qed
-
-lemma real_is_int_n1: "real_is_int (-1::real)"
-proof -
-  have "real_is_int (-1::real) = real_is_int(real (-1::int))" by auto
-  also have "\<dots> = True" by (simp only: real_is_int_real)
-  ultimately show ?thesis by auto
-qed
-
-lemma real_is_int_number_of[simp]: "real_is_int ((number_of \<Colon> int \<Rightarrow> real) x)"
-proof -
-  have neg1: "real_is_int (-1::real)"
-  proof -
-    have "real_is_int (-1::real) = real_is_int(real (-1::int))" by auto
-    also have "\<dots> = True" by (simp only: real_is_int_real)
-    ultimately show ?thesis by auto
-  qed
-
-  {
-    fix x :: int
-    have "real_is_int ((number_of \<Colon> int \<Rightarrow> real) x)"
-      unfolding number_of_eq
-      apply (induct x)
-      apply (induct_tac n)
-      apply (simp)
-      apply (simp)
-      apply (induct_tac n)
-      apply (simp add: neg1)
-    proof -
-      fix n :: nat
-      assume rn: "(real_is_int (of_int (- (int (Suc n)))))"
-      have s: "-(int (Suc (Suc n))) = -1 + - (int (Suc n))" by simp
-      show "real_is_int (of_int (- (int (Suc (Suc n)))))"
-        apply (simp only: s of_int_add)
-        apply (rule real_is_int_add)
-        apply (simp add: neg1)
-        apply (simp only: rn)
-        done
-    qed
-  }
-  note Abs_Bin = this
-  {
-    fix x :: int
-    have "? u. x = u"
-      apply (rule exI[where x = "x"])
-      apply (simp)
-      done
-  }
-  then obtain u::int where "x = u" by auto
-  with Abs_Bin show ?thesis by auto
-qed
-
-lemma int_of_real_0[simp]: "int_of_real (0::real) = (0::int)"
-by (simp add: int_of_real_def)
-
-lemma int_of_real_1[simp]: "int_of_real (1::real) = (1::int)"
-proof -
-  have 1: "(1::real) = real (1::int)" by auto
-  show ?thesis by (simp only: 1 int_of_real_real)
-qed
-
-lemma int_of_real_number_of[simp]: "int_of_real (number_of b) = number_of b"
-proof -
-  have "real_is_int (number_of b)" by simp
-  then have uu: "?! u::int. number_of b = real u" by (auto simp add: real_is_int_rep)
-  then obtain u::int where u:"number_of b = real u" by auto
-  have "number_of b = real ((number_of b)::int)"
-    by (simp add: number_of_eq real_of_int_def)
-  have ub: "number_of b = real ((number_of b)::int)"
-    by (simp add: number_of_eq real_of_int_def)
-  from uu u ub have unb: "u = number_of b"
-    by blast
-  have "int_of_real (number_of b) = u" by (simp add: u)
-  with unb show ?thesis by simp
-qed
-
-lemma float_transfer_even: "even a \<Longrightarrow> float (a, b) = float (a div 2, b+1)"
-  apply (subst float_transfer[where a="a" and b="b" and c="-1", simplified])
-  apply (simp_all add: pow2_def even_def real_is_int_def algebra_simps)
-  apply (auto)
-proof -
-  fix q::int
-  have a:"b - (-1\<Colon>int) = (1\<Colon>int) + b" by arith
-  show "(float (q, (b - (-1\<Colon>int)))) = (float (q, ((1\<Colon>int) + b)))"
-    by (simp add: a)
-qed
-
-lemma int_div_zdiv: "int (a div b) = (int a) div (int b)"
-by (rule zdiv_int)
-
-lemma int_mod_zmod: "int (a mod b) = (int a) mod (int b)"
-by (rule zmod_int)
-
-lemma abs_div_2_less: "a \<noteq> 0 \<Longrightarrow> a \<noteq> -1 \<Longrightarrow> abs((a::int) div 2) < abs a"
-by arith
-
-function norm_float :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
-  "norm_float a b = (if a \<noteq> 0 \<and> even a then norm_float (a div 2) (b + 1)
-    else if a = 0 then (0, 0) else (a, b))"
-by auto
-
-termination by (relation "measure (nat o abs o fst)")
-  (auto intro: abs_div_2_less)
-
-lemma norm_float: "float x = float (split norm_float x)"
-proof -
-  {
-    fix a b :: int
-    have norm_float_pair: "float (a, b) = float (norm_float a b)"
-    proof (induct a b rule: norm_float.induct)
-      case (1 u v)
-      show ?case
-      proof cases
-        assume u: "u \<noteq> 0 \<and> even u"
-        with prems have ind: "float (u div 2, v + 1) = float (norm_float (u div 2) (v + 1))" by auto
-        with u have "float (u,v) = float (u div 2, v+1)" by (simp add: float_transfer_even)
-        then show ?thesis
-          apply (subst norm_float.simps)
-          apply (simp add: ind)
-          done
-      next
-        assume "~(u \<noteq> 0 \<and> even u)"
-        then show ?thesis
-          by (simp add: prems float_def)
-      qed
-    qed
-  }
-  note helper = this
-  have "? a b. x = (a,b)" by auto
-  then obtain a b where "x = (a, b)" by blast
-  then show ?thesis by (simp add: helper)
-qed
-
-lemma float_add_l0: "float (0, e) + x = x"
-  by (simp add: float_def)
-
-lemma float_add_r0: "x + float (0, e) = x"
-  by (simp add: float_def)
-
-lemma float_add:
-  "float (a1, e1) + float (a2, e2) =
-  (if e1<=e2 then float (a1+a2*2^(nat(e2-e1)), e1)
-  else float (a1*2^(nat (e1-e2))+a2, e2))"
-  apply (simp add: float_def algebra_simps)
-  apply (auto simp add: pow2_int[symmetric] pow2_add[symmetric])
-  done
-
-lemma float_add_assoc1:
-  "(x + float (y1, e1)) + float (y2, e2) = (float (y1, e1) + float (y2, e2)) + x"
-  by simp
-
-lemma float_add_assoc2:
-  "(float (y1, e1) + x) + float (y2, e2) = (float (y1, e1) + float (y2, e2)) + x"
-  by simp
-
-lemma float_add_assoc3:
-  "float (y1, e1) + (x + float (y2, e2)) = (float (y1, e1) + float (y2, e2)) + x"
-  by simp
-
-lemma float_add_assoc4:
-  "float (y1, e1) + (float (y2, e2) + x) = (float (y1, e1) + float (y2, e2)) + x"
-  by simp
-
-lemma float_mult_l0: "float (0, e) * x = float (0, 0)"
-  by (simp add: float_def)
-
-lemma float_mult_r0: "x * float (0, e) = float (0, 0)"
-  by (simp add: float_def)
-
-definition 
-  lbound :: "real \<Rightarrow> real"
-where
-  "lbound x = min 0 x"
-
-definition
-  ubound :: "real \<Rightarrow> real"
-where
-  "ubound x = max 0 x"
-
-lemma lbound: "lbound x \<le> x"   
-  by (simp add: lbound_def)
-
-lemma ubound: "x \<le> ubound x"
-  by (simp add: ubound_def)
-
-lemma float_mult:
-  "float (a1, e1) * float (a2, e2) =
-  (float (a1 * a2, e1 + e2))"
-  by (simp add: float_def pow2_add)
-
-lemma float_minus:
-  "- (float (a,b)) = float (-a, b)"
-  by (simp add: float_def)
-
-lemma zero_less_pow2:
-  "0 < pow2 x"
-proof -
-  {
-    fix y
-    have "0 <= y \<Longrightarrow> 0 < pow2 y"
-      by (induct y, induct_tac n, simp_all add: pow2_add)
-  }
-  note helper=this
-  show ?thesis
-    apply (case_tac "0 <= x")
-    apply (simp add: helper)
-    apply (subst pow2_neg)
-    apply (simp add: helper)
-    done
-qed
-
-lemma zero_le_float:
-  "(0 <= float (a,b)) = (0 <= a)"
-  apply (auto simp add: float_def)
-  apply (auto simp add: zero_le_mult_iff zero_less_pow2)
-  apply (insert zero_less_pow2[of b])
-  apply (simp_all)
-  done
-
-lemma float_le_zero:
-  "(float (a,b) <= 0) = (a <= 0)"
-  apply (auto simp add: float_def)
-  apply (auto simp add: mult_le_0_iff)
-  apply (insert zero_less_pow2[of b])
-  apply auto
-  done
-
-lemma float_abs:
-  "abs (float (a,b)) = (if 0 <= a then (float (a,b)) else (float (-a,b)))"
-  apply (auto simp add: abs_if)
-  apply (simp_all add: zero_le_float[symmetric, of a b] float_minus)
-  done
-
-lemma float_zero:
-  "float (0, b) = 0"
-  by (simp add: float_def)
-
-lemma float_pprt:
-  "pprt (float (a, b)) = (if 0 <= a then (float (a,b)) else (float (0, b)))"
-  by (auto simp add: zero_le_float float_le_zero float_zero)
-
-lemma pprt_lbound: "pprt (lbound x) = float (0, 0)"
-  apply (simp add: float_def)
-  apply (rule pprt_eq_0)
-  apply (simp add: lbound_def)
-  done
-
-lemma nprt_ubound: "nprt (ubound x) = float (0, 0)"
-  apply (simp add: float_def)
-  apply (rule nprt_eq_0)
-  apply (simp add: ubound_def)
-  done
-
-lemma float_nprt:
-  "nprt (float (a, b)) = (if 0 <= a then (float (0,b)) else (float (a, b)))"
-  by (auto simp add: zero_le_float float_le_zero float_zero)
-
-lemma norm_0_1: "(0::_::number_ring) = Numeral0 & (1::_::number_ring) = Numeral1"
-  by auto
-
-lemma add_left_zero: "0 + a = (a::'a::comm_monoid_add)"
-  by simp
-
-lemma add_right_zero: "a + 0 = (a::'a::comm_monoid_add)"
-  by simp
-
-lemma mult_left_one: "1 * a = (a::'a::semiring_1)"
-  by simp
-
-lemma mult_right_one: "a * 1 = (a::'a::semiring_1)"
-  by simp
-
-lemma int_pow_0: "(a::int)^(Numeral0) = 1"
-  by simp
-
-lemma int_pow_1: "(a::int)^(Numeral1) = a"
-  by simp
-
-lemma zero_eq_Numeral0_nring: "(0::'a::number_ring) = Numeral0"
-  by simp
-
-lemma one_eq_Numeral1_nring: "(1::'a::number_ring) = Numeral1"
-  by simp
-
-lemma zero_eq_Numeral0_nat: "(0::nat) = Numeral0"
-  by simp
-
-lemma one_eq_Numeral1_nat: "(1::nat) = Numeral1"
-  by simp
-
-lemma zpower_Pls: "(z::int)^Numeral0 = Numeral1"
-  by simp
-
-lemma zpower_Min: "(z::int)^((-1)::nat) = Numeral1"
-proof -
-  have 1:"((-1)::nat) = 0"
-    by simp
-  show ?thesis by (simp add: 1)
-qed
-
-lemma fst_cong: "a=a' \<Longrightarrow> fst (a,b) = fst (a',b)"
-  by simp
-
-lemma snd_cong: "b=b' \<Longrightarrow> snd (a,b) = snd (a,b')"
-  by simp
-
-lemma lift_bool: "x \<Longrightarrow> x=True"
-  by simp
-
-lemma nlift_bool: "~x \<Longrightarrow> x=False"
-  by simp
-
-lemma not_false_eq_true: "(~ False) = True" by simp
-
-lemma not_true_eq_false: "(~ True) = False" by simp
-
-lemmas binarith =
-  normalize_bin_simps
-  pred_bin_simps succ_bin_simps
-  add_bin_simps minus_bin_simps mult_bin_simps
-
-lemma int_eq_number_of_eq:
-  "(((number_of v)::int)=(number_of w)) = iszero ((number_of (v + uminus w))::int)"
-  by (rule eq_number_of_eq)
-
-lemma int_iszero_number_of_Pls: "iszero (Numeral0::int)"
-  by (simp only: iszero_number_of_Pls)
-
-lemma int_nonzero_number_of_Min: "~(iszero ((-1)::int))"
-  by simp
-
-lemma int_iszero_number_of_Bit0: "iszero ((number_of (Int.Bit0 w))::int) = iszero ((number_of w)::int)"
-  by simp
-
-lemma int_iszero_number_of_Bit1: "\<not> iszero ((number_of (Int.Bit1 w))::int)"
-  by simp
-
-lemma int_less_number_of_eq_neg: "(((number_of x)::int) < number_of y) = neg ((number_of (x + (uminus y)))::int)"
-  unfolding neg_def number_of_is_id by simp
-
-lemma int_not_neg_number_of_Pls: "\<not> (neg (Numeral0::int))"
-  by simp
-
-lemma int_neg_number_of_Min: "neg (-1::int)"
-  by simp
-
-lemma int_neg_number_of_Bit0: "neg ((number_of (Int.Bit0 w))::int) = neg ((number_of w)::int)"
-  by simp
-
-lemma int_neg_number_of_Bit1: "neg ((number_of (Int.Bit1 w))::int) = neg ((number_of w)::int)"
-  by simp
-
-lemma int_le_number_of_eq: "(((number_of x)::int) \<le> number_of y) = (\<not> neg ((number_of (y + (uminus x)))::int))"
-  unfolding neg_def number_of_is_id by (simp add: not_less)
-
-lemmas intarithrel =
-  int_eq_number_of_eq
-  lift_bool[OF int_iszero_number_of_Pls] nlift_bool[OF int_nonzero_number_of_Min] int_iszero_number_of_Bit0
-  lift_bool[OF int_iszero_number_of_Bit1] int_less_number_of_eq_neg nlift_bool[OF int_not_neg_number_of_Pls] lift_bool[OF int_neg_number_of_Min]
-  int_neg_number_of_Bit0 int_neg_number_of_Bit1 int_le_number_of_eq
-
-lemma int_number_of_add_sym: "((number_of v)::int) + number_of w = number_of (v + w)"
-  by simp
-
-lemma int_number_of_diff_sym: "((number_of v)::int) - number_of w = number_of (v + (uminus w))"
-  by simp
-
-lemma int_number_of_mult_sym: "((number_of v)::int) * number_of w = number_of (v * w)"
-  by simp
-
-lemma int_number_of_minus_sym: "- ((number_of v)::int) = number_of (uminus v)"
-  by simp
-
-lemmas intarith = int_number_of_add_sym int_number_of_minus_sym int_number_of_diff_sym int_number_of_mult_sym
-
-lemmas natarith = add_nat_number_of diff_nat_number_of mult_nat_number_of eq_nat_number_of less_nat_number_of
-
-lemmas powerarith = nat_number_of zpower_number_of_even
-  zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring]
-  zpower_Pls zpower_Min
-
-lemmas floatarith[simplified norm_0_1] = float_add float_add_l0 float_add_r0 float_mult float_mult_l0 float_mult_r0 
-          float_minus float_abs zero_le_float float_pprt float_nprt pprt_lbound nprt_ubound
-
-(* for use with the compute oracle *)
-lemmas arith = binarith intarith intarithrel natarith powerarith floatarith not_false_eq_true not_true_eq_false
-
-use "~~/src/HOL/Tools/float_arith.ML"
-
-end
--- a/src/HOL/Tools/ComputeHOL.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,191 +0,0 @@
-theory ComputeHOL
-imports Complex_Main "~~/src/Tools/Compute_Oracle/Compute_Oracle"
-begin
-
-lemma Trueprop_eq_eq: "Trueprop X == (X == True)" by (simp add: atomize_eq)
-lemma meta_eq_trivial: "x == y \<Longrightarrow> x == y" by simp
-lemma meta_eq_imp_eq: "x == y \<Longrightarrow> x = y" by auto
-lemma eq_trivial: "x = y \<Longrightarrow> x = y" by auto
-lemma bool_to_true: "x :: bool \<Longrightarrow> x == True"  by simp
-lemma transmeta_1: "x = y \<Longrightarrow> y == z \<Longrightarrow> x = z" by simp
-lemma transmeta_2: "x == y \<Longrightarrow> y = z \<Longrightarrow> x = z" by simp
-lemma transmeta_3: "x == y \<Longrightarrow> y == z \<Longrightarrow> x = z" by simp
-
-
-(**** compute_if ****)
-
-lemma If_True: "If True = (\<lambda> x y. x)" by ((rule ext)+,auto)
-lemma If_False: "If False = (\<lambda> x y. y)" by ((rule ext)+, auto)
-
-lemmas compute_if = If_True If_False
-
-(**** compute_bool ****)
-
-lemma bool1: "(\<not> True) = False"  by blast
-lemma bool2: "(\<not> False) = True"  by blast
-lemma bool3: "(P \<and> True) = P" by blast
-lemma bool4: "(True \<and> P) = P" by blast
-lemma bool5: "(P \<and> False) = False" by blast
-lemma bool6: "(False \<and> P) = False" by blast
-lemma bool7: "(P \<or> True) = True" by blast
-lemma bool8: "(True \<or> P) = True" by blast
-lemma bool9: "(P \<or> False) = P" by blast
-lemma bool10: "(False \<or> P) = P" by blast
-lemma bool11: "(True \<longrightarrow> P) = P" by blast
-lemma bool12: "(P \<longrightarrow> True) = True" by blast
-lemma bool13: "(True \<longrightarrow> P) = P" by blast
-lemma bool14: "(P \<longrightarrow> False) = (\<not> P)" by blast
-lemma bool15: "(False \<longrightarrow> P) = True" by blast
-lemma bool16: "(False = False) = True" by blast
-lemma bool17: "(True = True) = True" by blast
-lemma bool18: "(False = True) = False" by blast
-lemma bool19: "(True = False) = False" by blast
-
-lemmas compute_bool = bool1 bool2 bool3 bool4 bool5 bool6 bool7 bool8 bool9 bool10 bool11 bool12 bool13 bool14 bool15 bool16 bool17 bool18 bool19
-
-
-(*** compute_pair ***)
-
-lemma compute_fst: "fst (x,y) = x" by simp
-lemma compute_snd: "snd (x,y) = y" by simp
-lemma compute_pair_eq: "((a, b) = (c, d)) = (a = c \<and> b = d)" by auto
-
-lemma prod_case_simp: "prod_case f (x,y) = f x y" by simp
-
-lemmas compute_pair = compute_fst compute_snd compute_pair_eq prod_case_simp
-
-(*** compute_option ***)
-
-lemma compute_the: "the (Some x) = x" by simp
-lemma compute_None_Some_eq: "(None = Some x) = False" by auto
-lemma compute_Some_None_eq: "(Some x = None) = False" by auto
-lemma compute_None_None_eq: "(None = None) = True" by auto
-lemma compute_Some_Some_eq: "(Some x = Some y) = (x = y)" by auto
-
-definition
-   option_case_compute :: "'b option \<Rightarrow> 'a \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a"
-where
-   "option_case_compute opt a f = option_case a f opt"
-
-lemma option_case_compute: "option_case = (\<lambda> a f opt. option_case_compute opt a f)"
-  by (simp add: option_case_compute_def)
-
-lemma option_case_compute_None: "option_case_compute None = (\<lambda> a f. a)"
-  apply (rule ext)+
-  apply (simp add: option_case_compute_def)
-  done
-
-lemma option_case_compute_Some: "option_case_compute (Some x) = (\<lambda> a f. f x)"
-  apply (rule ext)+
-  apply (simp add: option_case_compute_def)
-  done
-
-lemmas compute_option_case = option_case_compute option_case_compute_None option_case_compute_Some
-
-lemmas compute_option = compute_the compute_None_Some_eq compute_Some_None_eq compute_None_None_eq compute_Some_Some_eq compute_option_case
-
-(**** compute_list_length ****)
-
-lemma length_cons:"length (x#xs) = 1 + (length xs)"
-  by simp
-
-lemma length_nil: "length [] = 0"
-  by simp
-
-lemmas compute_list_length = length_nil length_cons
-
-(*** compute_list_case ***)
-
-definition
-  list_case_compute :: "'b list \<Rightarrow> 'a \<Rightarrow> ('b \<Rightarrow> 'b list \<Rightarrow> 'a) \<Rightarrow> 'a"
-where
-  "list_case_compute l a f = list_case a f l"
-
-lemma list_case_compute: "list_case = (\<lambda> (a::'a) f (l::'b list). list_case_compute l a f)"
-  apply (rule ext)+
-  apply (simp add: list_case_compute_def)
-  done
-
-lemma list_case_compute_empty: "list_case_compute ([]::'b list) = (\<lambda> (a::'a) f. a)"
-  apply (rule ext)+
-  apply (simp add: list_case_compute_def)
-  done
-
-lemma list_case_compute_cons: "list_case_compute (u#v) = (\<lambda> (a::'a) f. (f (u::'b) v))"
-  apply (rule ext)+
-  apply (simp add: list_case_compute_def)
-  done
-
-lemmas compute_list_case = list_case_compute list_case_compute_empty list_case_compute_cons
-
-(*** compute_list_nth ***)
-(* Of course, you will need computation with nats for this to work \<dots> *)
-
-lemma compute_list_nth: "((x#xs) ! n) = (if n = 0 then x else (xs ! (n - 1)))"
-  by (cases n, auto)
-  
-(*** compute_list ***)
-
-lemmas compute_list = compute_list_case compute_list_length compute_list_nth
-
-(*** compute_let ***)
-
-lemmas compute_let = Let_def
-
-(***********************)
-(* Everything together *)
-(***********************)
-
-lemmas compute_hol = compute_if compute_bool compute_pair compute_option compute_list compute_let
-
-ML {*
-signature ComputeHOL =
-sig
-  val prep_thms : thm list -> thm list
-  val to_meta_eq : thm -> thm
-  val to_hol_eq : thm -> thm
-  val symmetric : thm -> thm 
-  val trans : thm -> thm -> thm
-end
-
-structure ComputeHOL : ComputeHOL =
-struct
-
-local
-fun lhs_of eq = fst (Thm.dest_equals (cprop_of eq));
-in
-fun rewrite_conv [] ct = raise CTERM ("rewrite_conv", [])
-  | rewrite_conv (eq :: eqs) ct =
-      Thm.instantiate (Thm.match (lhs_of eq, ct)) eq
-      handle Pattern.MATCH => rewrite_conv eqs ct;
-end
-
-val convert_conditions = Conv.fconv_rule (Conv.prems_conv ~1 (Conv.try_conv (rewrite_conv [@{thm "Trueprop_eq_eq"}])))
-
-val eq_th = @{thm "HOL.eq_reflection"}
-val meta_eq_trivial = @{thm "ComputeHOL.meta_eq_trivial"}
-val bool_to_true = @{thm "ComputeHOL.bool_to_true"}
-
-fun to_meta_eq th = eq_th OF [th] handle THM _ => meta_eq_trivial OF [th] handle THM _ => bool_to_true OF [th]
-
-fun to_hol_eq th = @{thm "meta_eq_imp_eq"} OF [th] handle THM _ => @{thm "eq_trivial"} OF [th] 
-
-fun prep_thms ths = map (convert_conditions o to_meta_eq) ths
-
-fun symmetric th = @{thm "HOL.sym"} OF [th] handle THM _ => @{thm "Pure.symmetric"} OF [th]
-
-local
-    val trans_HOL = @{thm "HOL.trans"}
-    val trans_HOL_1 = @{thm "ComputeHOL.transmeta_1"}
-    val trans_HOL_2 = @{thm "ComputeHOL.transmeta_2"}
-    val trans_HOL_3 = @{thm "ComputeHOL.transmeta_3"}
-    fun tr [] th1 th2 = trans_HOL OF [th1, th2]
-      | tr (t::ts) th1 th2 = (t OF [th1, th2] handle THM _ => tr ts th1 th2) 
-in
-  fun trans th1 th2 = tr [trans_HOL, trans_HOL_1, trans_HOL_2, trans_HOL_3] th1 th2
-end
-
-end
-*}
-
-end
--- a/src/HOL/Tools/ComputeNumeral.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,195 +0,0 @@
-theory ComputeNumeral
-imports ComputeHOL ComputeFloat
-begin
-
-(* normalization of bit strings *)
-lemmas bitnorm = normalize_bin_simps
-
-(* neg for bit strings *)
-lemma neg1: "neg Int.Pls = False" by (simp add: Int.Pls_def)
-lemma neg2: "neg Int.Min = True" apply (subst Int.Min_def) by auto
-lemma neg3: "neg (Int.Bit0 x) = neg x" apply (simp add: neg_def) apply (subst Bit0_def) by auto
-lemma neg4: "neg (Int.Bit1 x) = neg x" apply (simp add: neg_def) apply (subst Bit1_def) by auto  
-lemmas bitneg = neg1 neg2 neg3 neg4
-
-(* iszero for bit strings *)
-lemma iszero1: "iszero Int.Pls = True" by (simp add: Int.Pls_def iszero_def)
-lemma iszero2: "iszero Int.Min = False" apply (subst Int.Min_def) apply (subst iszero_def) by simp
-lemma iszero3: "iszero (Int.Bit0 x) = iszero x" apply (subst Int.Bit0_def) apply (subst iszero_def)+ by auto
-lemma iszero4: "iszero (Int.Bit1 x) = False" apply (subst Int.Bit1_def) apply (subst iszero_def)+  apply simp by arith
-lemmas bitiszero = iszero1 iszero2 iszero3 iszero4
-
-(* lezero for bit strings *)
-constdefs
-  "lezero x == (x \<le> 0)"
-lemma lezero1: "lezero Int.Pls = True" unfolding Int.Pls_def lezero_def by auto
-lemma lezero2: "lezero Int.Min = True" unfolding Int.Min_def lezero_def by auto
-lemma lezero3: "lezero (Int.Bit0 x) = lezero x" unfolding Int.Bit0_def lezero_def by auto
-lemma lezero4: "lezero (Int.Bit1 x) = neg x" unfolding Int.Bit1_def lezero_def neg_def by auto
-lemmas bitlezero = lezero1 lezero2 lezero3 lezero4
-
-(* equality for bit strings *)
-lemmas biteq = eq_bin_simps
-
-(* x < y for bit strings *)
-lemmas bitless = less_bin_simps
-
-(* x \<le> y for bit strings *)
-lemmas bitle = le_bin_simps
-
-(* succ for bit strings *)
-lemmas bitsucc = succ_bin_simps
-
-(* pred for bit strings *)
-lemmas bitpred = pred_bin_simps
-
-(* unary minus for bit strings *)
-lemmas bituminus = minus_bin_simps
-
-(* addition for bit strings *)
-lemmas bitadd = add_bin_simps
-
-(* multiplication for bit strings *) 
-lemma mult_Pls_right: "x * Int.Pls = Int.Pls" by (simp add: Pls_def)
-lemma mult_Min_right: "x * Int.Min = - x" by (subst mult_commute, simp add: mult_Min)
-lemma multb0x: "(Int.Bit0 x) * y = Int.Bit0 (x * y)" by (rule mult_Bit0)
-lemma multxb0: "x * (Int.Bit0 y) = Int.Bit0 (x * y)" unfolding Bit0_def by simp
-lemma multb1: "(Int.Bit1 x) * (Int.Bit1 y) = Int.Bit1 (Int.Bit0 (x * y) + x + y)"
-  unfolding Bit0_def Bit1_def by (simp add: algebra_simps)
-lemmas bitmul = mult_Pls mult_Min mult_Pls_right mult_Min_right multb0x multxb0 multb1
-
-lemmas bitarith = bitnorm bitiszero bitneg bitlezero biteq bitless bitle bitsucc bitpred bituminus bitadd bitmul 
-
-constdefs 
-  "nat_norm_number_of (x::nat) == x"
-
-lemma nat_norm_number_of: "nat_norm_number_of (number_of w) = (if lezero w then 0 else number_of w)"
-  apply (simp add: nat_norm_number_of_def)
-  unfolding lezero_def iszero_def neg_def
-  apply (simp add: numeral_simps)
-  done
-
-(* Normalization of nat literals *)
-lemma natnorm0: "(0::nat) = number_of (Int.Pls)" by auto
-lemma natnorm1: "(1 :: nat) = number_of (Int.Bit1 Int.Pls)"  by auto 
-lemmas natnorm = natnorm0 natnorm1 nat_norm_number_of
-
-(* Suc *)
-lemma natsuc: "Suc (number_of x) = (if neg x then 1 else number_of (Int.succ x))" by (auto simp add: number_of_is_id)
-
-(* Addition for nat *)
-lemma natadd: "number_of x + ((number_of y)::nat) = (if neg x then (number_of y) else (if neg y then number_of x else (number_of (x + y))))"
-  unfolding nat_number_of_def number_of_is_id neg_def
-  by auto
-
-(* Subtraction for nat *)
-lemma natsub: "(number_of x) - ((number_of y)::nat) = 
-  (if neg x then 0 else (if neg y then number_of x else (nat_norm_number_of (number_of (x + (- y))))))"
-  unfolding nat_norm_number_of
-  by (auto simp add: number_of_is_id neg_def lezero_def iszero_def Let_def nat_number_of_def)
-
-(* Multiplication for nat *)
-lemma natmul: "(number_of x) * ((number_of y)::nat) = 
-  (if neg x then 0 else (if neg y then 0 else number_of (x * y)))"
-  unfolding nat_number_of_def number_of_is_id neg_def
-  by (simp add: nat_mult_distrib)
-
-lemma nateq: "(((number_of x)::nat) = (number_of y)) = ((lezero x \<and> lezero y) \<or> (x = y))"
-  by (auto simp add: iszero_def lezero_def neg_def number_of_is_id)
-
-lemma natless: "(((number_of x)::nat) < (number_of y)) = ((x < y) \<and> (\<not> (lezero y)))"
-  by (simp add: lezero_def numeral_simps not_le)
-
-lemma natle: "(((number_of x)::nat) \<le> (number_of y)) = (y < x \<longrightarrow> lezero x)"
-  by (auto simp add: number_of_is_id lezero_def nat_number_of_def)
-
-fun natfac :: "nat \<Rightarrow> nat"
-where
-   "natfac n = (if n = 0 then 1 else n * (natfac (n - 1)))"
-
-lemmas compute_natarith = bitarith natnorm natsuc natadd natsub natmul nateq natless natle natfac.simps
-
-lemma number_eq: "(((number_of x)::'a::{number_ring, ordered_idom}) = (number_of y)) = (x = y)"
-  unfolding number_of_eq
-  apply simp
-  done
-
-lemma number_le: "(((number_of x)::'a::{number_ring, ordered_idom}) \<le>  (number_of y)) = (x \<le> y)"
-  unfolding number_of_eq
-  apply simp
-  done
-
-lemma number_less: "(((number_of x)::'a::{number_ring, ordered_idom}) <  (number_of y)) = (x < y)"
-  unfolding number_of_eq 
-  apply simp
-  done
-
-lemma number_diff: "((number_of x)::'a::{number_ring, ordered_idom}) - number_of y = number_of (x + (- y))"
-  apply (subst diff_number_of_eq)
-  apply simp
-  done
-
-lemmas number_norm = number_of_Pls[symmetric] numeral_1_eq_1[symmetric]
-
-lemmas compute_numberarith = number_of_minus[symmetric] number_of_add[symmetric] number_diff number_of_mult[symmetric] number_norm number_eq number_le number_less
-
-lemma compute_real_of_nat_number_of: "real ((number_of v)::nat) = (if neg v then 0 else number_of v)"
-  by (simp only: real_of_nat_number_of number_of_is_id)
-
-lemma compute_nat_of_int_number_of: "nat ((number_of v)::int) = (number_of v)"
-  by simp
-
-lemmas compute_num_conversions = compute_real_of_nat_number_of compute_nat_of_int_number_of real_number_of
-
-lemmas zpowerarith = zpower_number_of_even
-  zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring]
-  zpower_Pls zpower_Min
-
-(* div, mod *)
-
-lemma adjust: "adjust b (q, r) = (if 0 \<le> r - b then (2 * q + 1, r - b) else (2 * q, r))"
-  by (auto simp only: adjust_def)
-
-lemma negateSnd: "negateSnd (q, r) = (q, -r)" 
-  by (simp add: negateSnd_def)
-
-lemma divmod: "IntDiv.divmod a b = (if 0\<le>a then
-                  if 0\<le>b then posDivAlg a b
-                  else if a=0 then (0, 0)
-                       else negateSnd (negDivAlg (-a) (-b))
-               else 
-                  if 0<b then negDivAlg a b
-                  else negateSnd (posDivAlg (-a) (-b)))"
-  by (auto simp only: IntDiv.divmod_def)
-
-lemmas compute_div_mod = div_def mod_def divmod adjust negateSnd posDivAlg.simps negDivAlg.simps
-
-
-
-(* collecting all the theorems *)
-
-lemma even_Pls: "even (Int.Pls) = True"
-  apply (unfold Pls_def even_def)
-  by simp
-
-lemma even_Min: "even (Int.Min) = False"
-  apply (unfold Min_def even_def)
-  by simp
-
-lemma even_B0: "even (Int.Bit0 x) = True"
-  apply (unfold Bit0_def)
-  by simp
-
-lemma even_B1: "even (Int.Bit1 x) = False"
-  apply (unfold Bit1_def)
-  by simp
-
-lemma even_number_of: "even ((number_of w)::int) = even w"
-  by (simp only: number_of_is_id)
-
-lemmas compute_even = even_Pls even_Min even_B0 even_B1 even_number_of
-
-lemmas compute_numeral = compute_if compute_let compute_pair compute_bool 
-                         compute_natarith compute_numberarith max_def min_def compute_num_conversions zpowerarith compute_div_mod compute_even
-
-end
--- a/src/HOL/Tools/Function/fundef_core.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Tools/Function/fundef_core.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -577,8 +577,6 @@
 val acc_subset_induct = @{thm Orderings.predicate1I} RS @{thm accp_subset_induct}
 
 
-fun binder_conv cv ctxt = Conv.arg_conv (Conv.abs_conv (K cv) ctxt);
-
 fun mk_partial_induct_rule thy globals R complete_thm clauses =
     let
       val Globals {domT, x, z, a, P, D, ...} = globals
@@ -614,7 +612,7 @@
     val case_hyp_conv = K (case_hyp RS eq_reflection)
     local open Conv in
     val lhs_D = fconv_rule (arg_conv (arg_conv (case_hyp_conv))) x_D
-    val sih = fconv_rule (binder_conv (arg1_conv (arg_conv (arg_conv case_hyp_conv))) ctxt) aihyp
+    val sih = fconv_rule (More_Conv.binder_conv (K (arg1_conv (arg_conv (arg_conv case_hyp_conv)))) ctxt) aihyp
     end
 
     fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) =
--- a/src/HOL/Tools/Qelim/cooper.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Tools/Qelim/cooper.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -29,30 +29,30 @@
 val bT = HOLogic.boolT;
 val dest_numeral = HOLogic.dest_number #> snd;
 
-val [miconj, midisj, mieq, mineq, milt, mile, migt, mige, midvd, mindvd, miP] = 
+val [miconj, midisj, mieq, mineq, milt, mile, migt, mige, midvd, mindvd, miP] =
     map(instantiate' [SOME @{ctyp "int"}] []) @{thms "minf"};
 
-val [infDconj, infDdisj, infDdvd,infDndvd,infDP] = 
+val [infDconj, infDdisj, infDdvd,infDndvd,infDP] =
     map(instantiate' [SOME @{ctyp "int"}] []) @{thms "inf_period"};
 
-val [piconj, pidisj, pieq,pineq,pilt,pile,pigt,pige,pidvd,pindvd,piP] = 
+val [piconj, pidisj, pieq,pineq,pilt,pile,pigt,pige,pidvd,pindvd,piP] =
     map (instantiate' [SOME @{ctyp "int"}] []) @{thms "pinf"};
 
 val [miP, piP] = map (instantiate' [SOME @{ctyp "bool"}] []) [miP, piP];
 
 val infDP = instantiate' (map SOME [@{ctyp "int"}, @{ctyp "bool"}]) [] infDP;
 
-val [[asetconj, asetdisj, aseteq, asetneq, asetlt, asetle, 
+val [[asetconj, asetdisj, aseteq, asetneq, asetlt, asetle,
       asetgt, asetge, asetdvd, asetndvd,asetP],
-     [bsetconj, bsetdisj, bseteq, bsetneq, bsetlt, bsetle, 
+     [bsetconj, bsetdisj, bseteq, bsetneq, bsetlt, bsetle,
       bsetgt, bsetge, bsetdvd, bsetndvd,bsetP]]  = [@{thms "aset"}, @{thms "bset"}];
 
-val [miex, cpmi, piex, cppi] = [@{thm "minusinfinity"}, @{thm "cpmi"}, 
+val [miex, cpmi, piex, cppi] = [@{thm "minusinfinity"}, @{thm "cpmi"},
                                 @{thm "plusinfinity"}, @{thm "cppi"}];
 
 val unity_coeff_ex = instantiate' [SOME @{ctyp "int"}] [] @{thm "unity_coeff_ex"};
 
-val [zdvd_mono,simp_from_to,all_not_ex] = 
+val [zdvd_mono,simp_from_to,all_not_ex] =
      [@{thm "zdvd_mono"}, @{thm "simp_from_to"}, @{thm "all_not_ex"}];
 
 val [dvd_uminus, dvd_uminus'] = @{thms "uminus_dvd_conv"};
@@ -62,49 +62,49 @@
 
 (* recognising cterm without moving to terms *)
 
-datatype fm = And of cterm*cterm| Or of cterm*cterm| Eq of cterm | NEq of cterm 
+datatype fm = And of cterm*cterm| Or of cterm*cterm| Eq of cterm | NEq of cterm
             | Lt of cterm | Le of cterm | Gt of cterm | Ge of cterm
             | Dvd of cterm*cterm | NDvd of cterm*cterm | Nox
 
-fun whatis x ct = 
-( case (term_of ct) of 
+fun whatis x ct =
+( case (term_of ct) of
   Const("op &",_)$_$_ => And (Thm.dest_binop ct)
 | Const ("op |",_)$_$_ => Or (Thm.dest_binop ct)
 | Const ("op =",ty)$y$_ => if term_of x aconv y then Eq (Thm.dest_arg ct) else Nox
-| Const (@{const_name Not},_) $ (Const ("op =",_)$y$_) => 
+| Const (@{const_name Not},_) $ (Const ("op =",_)$y$_) =>
   if term_of x aconv y then NEq (funpow 2 Thm.dest_arg ct) else Nox
 | Const (@{const_name HOL.less}, _) $ y$ z =>
-   if term_of x aconv y then Lt (Thm.dest_arg ct) 
+   if term_of x aconv y then Lt (Thm.dest_arg ct)
    else if term_of x aconv z then Gt (Thm.dest_arg1 ct) else Nox
-| Const (@{const_name HOL.less_eq}, _) $ y $ z => 
-   if term_of x aconv y then Le (Thm.dest_arg ct) 
+| Const (@{const_name HOL.less_eq}, _) $ y $ z =>
+   if term_of x aconv y then Le (Thm.dest_arg ct)
    else if term_of x aconv z then Ge (Thm.dest_arg1 ct) else Nox
 | Const (@{const_name Ring_and_Field.dvd},_)$_$(Const(@{const_name HOL.plus},_)$y$_) =>
-   if term_of x aconv y then Dvd (Thm.dest_binop ct ||> Thm.dest_arg) else Nox 
+   if term_of x aconv y then Dvd (Thm.dest_binop ct ||> Thm.dest_arg) else Nox
 | Const (@{const_name Not},_) $ (Const (@{const_name Ring_and_Field.dvd},_)$_$(Const(@{const_name "HOL.plus"},_)$y$_)) =>
-   if term_of x aconv y then 
-   NDvd (Thm.dest_binop (Thm.dest_arg ct) ||> Thm.dest_arg) else Nox 
+   if term_of x aconv y then
+   NDvd (Thm.dest_binop (Thm.dest_arg ct) ||> Thm.dest_arg) else Nox
 | _ => Nox)
-  handle CTERM _ => Nox; 
+  handle CTERM _ => Nox;
 
-fun get_pmi_term t = 
-  let val (x,eq) = 
+fun get_pmi_term t =
+  let val (x,eq) =
      (Thm.dest_abs NONE o Thm.dest_arg o snd o Thm.dest_abs NONE o Thm.dest_arg)
         (Thm.dest_arg t)
 in (Thm.cabs x o Thm.dest_arg o Thm.dest_arg) eq end;
 
 val get_pmi = get_pmi_term o cprop_of;
 
-val p_v' = @{cpat "?P' :: int => bool"}; 
+val p_v' = @{cpat "?P' :: int => bool"};
 val q_v' = @{cpat "?Q' :: int => bool"};
 val p_v = @{cpat "?P:: int => bool"};
 val q_v = @{cpat "?Q:: int => bool"};
 
-fun myfwd (th1, th2, th3) p q 
-      [(th_1,th_2,th_3), (th_1',th_2',th_3')] = 
-  let  
+fun myfwd (th1, th2, th3) p q
+      [(th_1,th_2,th_3), (th_1',th_2',th_3')] =
+  let
    val (mp', mq') = (get_pmi th_1, get_pmi th_1')
-   val mi_th = FWD (instantiate ([],[(p_v,p),(q_v,q), (p_v',mp'),(q_v',mq')]) th1) 
+   val mi_th = FWD (instantiate ([],[(p_v,p),(q_v,q), (p_v',mp'),(q_v',mq')]) th1)
                    [th_1, th_1']
    val infD_th = FWD (instantiate ([],[(p_v,mp'), (q_v, mq')]) th3) [th_3,th_3']
    val set_th = FWD (instantiate ([],[(p_v,p), (q_v,q)]) th2) [th_2, th_2']
@@ -123,15 +123,15 @@
 val [addC, mulC, subC, negC] = map term_of [cadd, cmulC, cminus, cneg]
 val [zero, one] = [@{term "0 :: int"}, @{term "1 :: int"}];
 
-val is_numeral = can dest_numeral; 
+val is_numeral = can dest_numeral;
 
-fun numeral1 f n = HOLogic.mk_number iT (f (dest_numeral n)); 
+fun numeral1 f n = HOLogic.mk_number iT (f (dest_numeral n));
 fun numeral2 f m n = HOLogic.mk_number iT (f (dest_numeral m) (dest_numeral n));
 
-val [minus1,plus1] = 
+val [minus1,plus1] =
     map (fn c => fn t => Thm.capply (Thm.capply c t) cone) [cminus,cadd];
 
-fun decomp_pinf x dvd inS [aseteq, asetneq, asetlt, asetle, 
+fun decomp_pinf x dvd inS [aseteq, asetneq, asetlt, asetle,
                            asetgt, asetge,asetdvd,asetndvd,asetP,
                            infDdvd, infDndvd, asetconj,
                            asetdisj, infDconj, infDdisj] cp =
@@ -144,11 +144,11 @@
 | Le t => ([], K (inst' [t] pile, FWD (inst' [t] asetle) [inS (plus1 t)], infDFalse))
 | Gt t => ([], K (inst' [t] pigt, (inst' [t] asetgt), infDTrue))
 | Ge t => ([], K (inst' [t] pige, (inst' [t] asetge), infDTrue))
-| Dvd (d,s) => 
+| Dvd (d,s) =>
    ([],let val dd = dvd d
-	     in K (inst' [d,s] pidvd, FWD (inst' [d,s] asetdvd) [dd],FWD (inst' [d,s] infDdvd) [dd]) end)
+       in K (inst' [d,s] pidvd, FWD (inst' [d,s] asetdvd) [dd],FWD (inst' [d,s] infDdvd) [dd]) end)
 | NDvd(d,s) => ([],let val dd = dvd d
-	      in K (inst' [d,s] pindvd, FWD (inst' [d,s] asetndvd) [dd], FWD (inst' [d,s] infDndvd) [dd]) end)
+        in K (inst' [d,s] pindvd, FWD (inst' [d,s] asetndvd) [dd], FWD (inst' [d,s] infDndvd) [dd]) end)
 | _ => ([], K (inst' [cp] piP, inst' [cp] asetP, inst' [cp] infDP));
 
 fun decomp_minf x dvd inS [bseteq,bsetneq,bsetlt, bsetle, bsetgt,
@@ -165,110 +165,112 @@
 | Gt t => ([], K (inst' [t] migt, FWD (inst' [t] bsetgt) [inS t], infDFalse))
 | Ge t => ([], K (inst' [t] mige,FWD (inst' [t] bsetge) [inS (minus1 t)], infDFalse))
 | Dvd (d,s) => ([],let val dd = dvd d
-	      in K (inst' [d,s] midvd, FWD (inst' [d,s] bsetdvd) [dd] , FWD (inst' [d,s] infDdvd) [dd]) end)
+        in K (inst' [d,s] midvd, FWD (inst' [d,s] bsetdvd) [dd] , FWD (inst' [d,s] infDdvd) [dd]) end)
 | NDvd (d,s) => ([],let val dd = dvd d
-	      in K (inst' [d,s] mindvd, FWD (inst' [d,s] bsetndvd) [dd], FWD (inst' [d,s] infDndvd) [dd]) end)
+        in K (inst' [d,s] mindvd, FWD (inst' [d,s] bsetndvd) [dd], FWD (inst' [d,s] infDndvd) [dd]) end)
 | _ => ([], K (inst' [cp] miP, inst' [cp] bsetP, inst' [cp] infDP))
 
     (* Canonical linear form for terms, formulae etc.. *)
-fun provelin ctxt t = Goal.prove ctxt [] [] t 
+fun provelin ctxt t = Goal.prove ctxt [] [] t
   (fn _ => EVERY [simp_tac lin_ss 1, TRY (Lin_Arith.tac ctxt 1)]);
-fun linear_cmul 0 tm = zero 
-  | linear_cmul n tm = case tm of  
+fun linear_cmul 0 tm = zero
+  | linear_cmul n tm = case tm of
       Const (@{const_name HOL.plus}, _) $ a $ b => addC $ linear_cmul n a $ linear_cmul n b
     | Const (@{const_name HOL.times}, _) $ c $ x => mulC $ numeral1 (fn m => n * m) c $ x
     | Const (@{const_name HOL.minus}, _) $ a $ b => subC $ linear_cmul n a $ linear_cmul n b
     | (m as Const (@{const_name HOL.uminus}, _)) $ a => m $ linear_cmul n a
     | _ => numeral1 (fn m => n * m) tm;
-fun earlier [] x y = false 
-	| earlier (h::t) x y = 
-    if h aconv y then false else if h aconv x then true else earlier t x y; 
+fun earlier [] x y = false
+  | earlier (h::t) x y =
+    if h aconv y then false else if h aconv x then true else earlier t x y;
 
-fun linear_add vars tm1 tm2 = case (tm1, tm2) of 
+fun linear_add vars tm1 tm2 = case (tm1, tm2) of
     (Const (@{const_name HOL.plus}, _) $ (Const (@{const_name HOL.times}, _) $ c1 $ x1) $ r1,
     Const (@{const_name HOL.plus}, _) $ (Const (@{const_name HOL.times}, _) $ c2 $ x2) $ r2) =>
-   if x1 = x2 then 
+   if x1 = x2 then
      let val c = numeral2 (curry op +) c1 c2
       in if c = zero then linear_add vars r1 r2
          else addC$(mulC$c$x1)$(linear_add vars r1 r2)
-     end 
+     end
      else if earlier vars x1 x2 then addC $ (mulC $ c1 $ x1) $ linear_add vars r1 tm2
    else addC $ (mulC $ c2 $ x2) $ linear_add vars tm1 r2
  | (Const (@{const_name HOL.plus}, _) $ (Const (@{const_name HOL.times}, _) $ c1 $ x1) $ r1, _) =>
       addC $ (mulC $ c1 $ x1) $ linear_add vars r1 tm2
- | (_, Const (@{const_name HOL.plus}, _) $ (Const (@{const_name HOL.times}, _) $ c2 $ x2) $ r2) => 
+ | (_, Const (@{const_name HOL.plus}, _) $ (Const (@{const_name HOL.times}, _) $ c2 $ x2) $ r2) =>
       addC $ (mulC $ c2 $ x2) $ linear_add vars tm1 r2
  | (_, _) => numeral2 (curry op +) tm1 tm2;
- 
-fun linear_neg tm = linear_cmul ~1 tm; 
-fun linear_sub vars tm1 tm2 = linear_add vars tm1 (linear_neg tm2); 
+
+fun linear_neg tm = linear_cmul ~1 tm;
+fun linear_sub vars tm1 tm2 = linear_add vars tm1 (linear_neg tm2);
 
 
-fun lint vars tm =  if is_numeral tm then tm  else case tm of 
+fun lint vars tm =  if is_numeral tm then tm  else case tm of
   Const (@{const_name HOL.uminus}, _) $ t => linear_neg (lint vars t)
 | Const (@{const_name HOL.plus}, _) $ s $ t => linear_add vars (lint vars s) (lint vars t)
 | Const (@{const_name HOL.minus}, _) $ s $ t => linear_sub vars (lint vars s) (lint vars t)
 | Const (@{const_name HOL.times}, _) $ s $ t =>
-  let val s' = lint vars s  
-      val t' = lint vars t  
-  in if is_numeral s' then (linear_cmul (dest_numeral s') t') 
-     else if is_numeral t' then (linear_cmul (dest_numeral t') s') 
+  let val s' = lint vars s
+      val t' = lint vars t
+  in if is_numeral s' then (linear_cmul (dest_numeral s') t')
+     else if is_numeral t' then (linear_cmul (dest_numeral t') s')
      else raise COOPER ("Cooper Failed", TERM ("lint: not linear",[tm]))
-  end 
+  end
  | _ => addC $ (mulC $ one $ tm) $ zero;
 
-fun lin (vs as x::_) (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.less}, T) $ s $ t)) = 
+fun lin (vs as x::_) (Const (@{const_name Not}, _) $ (Const (@{const_name HOL.less}, T) $ s $ t)) =
     lin vs (Const (@{const_name HOL.less_eq}, T) $ t $ s)
-  | lin (vs as x::_) (Const (@{const_name Not},_) $ (Const(@{const_name HOL.less_eq}, T) $ s $ t)) = 
+  | lin (vs as x::_) (Const (@{const_name Not},_) $ (Const(@{const_name HOL.less_eq}, T) $ s $ t)) =
     lin vs (Const (@{const_name HOL.less}, T) $ t $ s)
   | lin vs (Const (@{const_name Not},T)$t) = Const (@{const_name Not},T)$ (lin vs t)
-  | lin (vs as x::_) (Const(@{const_name Ring_and_Field.dvd},_)$d$t) = 
+  | lin (vs as x::_) (Const(@{const_name Ring_and_Field.dvd},_)$d$t) =
     HOLogic.mk_binrel @{const_name Ring_and_Field.dvd} (numeral1 abs d, lint vs t)
-  | lin (vs as x::_) ((b as Const("op =",_))$s$t) = 
-     (case lint vs (subC$t$s) of 
-      (t as a$(m$c$y)$r) => 
+  | lin (vs as x::_) ((b as Const("op =",_))$s$t) =
+     (case lint vs (subC$t$s) of
+      (t as a$(m$c$y)$r) =>
         if x <> y then b$zero$t
         else if dest_numeral c < 0 then b$(m$(numeral1 ~ c)$y)$r
         else b$(m$c$y)$(linear_neg r)
       | t => b$zero$t)
-  | lin (vs as x::_) (b$s$t) = 
-     (case lint vs (subC$t$s) of 
-      (t as a$(m$c$y)$r) => 
+  | lin (vs as x::_) (b$s$t) =
+     (case lint vs (subC$t$s) of
+      (t as a$(m$c$y)$r) =>
         if x <> y then b$zero$t
         else if dest_numeral c < 0 then b$(m$(numeral1 ~ c)$y)$r
         else b$(linear_neg r)$(m$c$y)
       | t => b$zero$t)
   | lin vs fm = fm;
 
-fun lint_conv ctxt vs ct = 
+fun lint_conv ctxt vs ct =
 let val t = term_of ct
 in (provelin ctxt ((HOLogic.eq_const iT)$t$(lint vs t) |> HOLogic.mk_Trueprop))
              RS eq_reflection
 end;
 
-fun is_intrel (b$_$_) = domain_type (fastype_of b) = HOLogic.intT
-  | is_intrel (@{term "Not"}$(b$_$_)) = domain_type (fastype_of b) = HOLogic.intT
+fun is_intrel_type T = T = @{typ "int => int => bool"};
+
+fun is_intrel (b$_$_) = is_intrel_type (fastype_of b)
+  | is_intrel (@{term "Not"}$(b$_$_)) = is_intrel_type (fastype_of b)
   | is_intrel _ = false;
- 
+
 fun linearize_conv ctxt vs ct = case term_of ct of
-  Const(@{const_name Ring_and_Field.dvd},_)$d$t => 
-  let 
+  Const(@{const_name Ring_and_Field.dvd},_)$d$t =>
+  let
     val th = binop_conv (lint_conv ctxt vs) ct
     val (d',t') = Thm.dest_binop (Thm.rhs_of th)
     val (dt',tt') = (term_of d', term_of t')
-  in if is_numeral dt' andalso is_numeral tt' 
+  in if is_numeral dt' andalso is_numeral tt'
      then Conv.fconv_rule (arg_conv (Simplifier.rewrite presburger_ss)) th
-     else 
-     let 
-      val dth = 
-      ((if dest_numeral (term_of d') < 0 then 
+     else
+     let
+      val dth =
+      ((if dest_numeral (term_of d') < 0 then
           Conv.fconv_rule (arg_conv (arg1_conv (lint_conv ctxt vs)))
                            (Thm.transitive th (inst' [d',t'] dvd_uminus))
         else th) handle TERM _ => th)
       val d'' = Thm.rhs_of dth |> Thm.dest_arg1
      in
-      case tt' of 
-        Const(@{const_name HOL.plus},_)$(Const(@{const_name HOL.times},_)$c$_)$_ => 
+      case tt' of
+        Const(@{const_name HOL.plus},_)$(Const(@{const_name HOL.times},_)$c$_)$_ =>
         let val x = dest_numeral c
         in if x < 0 then Conv.fconv_rule (arg_conv (arg_conv (lint_conv ctxt vs)))
                                        (Thm.transitive dth (inst' [d'',t'] dvd_uminus'))
@@ -277,29 +279,29 @@
      end
   end
 | Const (@{const_name Not},_)$(Const(@{const_name Ring_and_Field.dvd},_)$_$_) => arg_conv (linearize_conv ctxt vs) ct
-| t => if is_intrel t 
+| t => if is_intrel t
       then (provelin ctxt ((HOLogic.eq_const bT)$t$(lin vs t) |> HOLogic.mk_Trueprop))
        RS eq_reflection
       else reflexive ct;
 
 val dvdc = @{cterm "op dvd :: int => _"};
 
-fun unify ctxt q = 
+fun unify ctxt q =
  let
   val (e,(cx,p)) = q |> Thm.dest_comb ||> Thm.dest_abs NONE
-  val x = term_of cx 
+  val x = term_of cx
   val ins = insert (op = : int * int -> bool)
-  fun h (acc,dacc) t = 
+  fun h (acc,dacc) t =
    case (term_of t) of
-    Const(s,_)$(Const(@{const_name HOL.times},_)$c$y)$ _ => 
+    Const(s,_)$(Const(@{const_name HOL.times},_)$c$y)$ _ =>
     if x aconv y andalso member (op =)
       ["op =", @{const_name HOL.less}, @{const_name HOL.less_eq}] s
     then (ins (dest_numeral c) acc,dacc) else (acc,dacc)
-  | Const(s,_)$_$(Const(@{const_name HOL.times},_)$c$y) => 
+  | Const(s,_)$_$(Const(@{const_name HOL.times},_)$c$y) =>
     if x aconv y andalso member (op =)
-       [@{const_name HOL.less}, @{const_name HOL.less_eq}] s 
+       [@{const_name HOL.less}, @{const_name HOL.less_eq}] s
     then (ins (dest_numeral c) acc, dacc) else (acc,dacc)
-  | Const(@{const_name Ring_and_Field.dvd},_)$_$(Const(@{const_name HOL.plus},_)$(Const(@{const_name HOL.times},_)$c$y)$_) => 
+  | Const(@{const_name Ring_and_Field.dvd},_)$_$(Const(@{const_name HOL.plus},_)$(Const(@{const_name HOL.times},_)$c$y)$_) =>
     if x aconv y then (acc,ins (dest_numeral c) dacc) else (acc,dacc)
   | Const("op &",_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
   | Const("op |",_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t)
@@ -307,51 +309,53 @@
   | _ => (acc, dacc)
   val (cs,ds) = h ([],[]) p
   val l = Integer.lcms (cs union ds)
-  fun cv k ct = 
-    let val (tm as b$s$t) = term_of ct 
+  fun cv k ct =
+    let val (tm as b$s$t) = term_of ct
     in ((HOLogic.eq_const bT)$tm$(b$(linear_cmul k s)$(linear_cmul k t))
          |> HOLogic.mk_Trueprop |> provelin ctxt) RS eq_reflection end
-  fun nzprop x = 
-   let 
-    val th = 
-     Simplifier.rewrite lin_ss 
-      (Thm.capply @{cterm Trueprop} (Thm.capply @{cterm "Not"} 
-           (Thm.capply (Thm.capply @{cterm "op = :: int => _"} (Numeral.mk_cnumber @{ctyp "int"} x)) 
+  fun nzprop x =
+   let
+    val th =
+     Simplifier.rewrite lin_ss
+      (Thm.capply @{cterm Trueprop} (Thm.capply @{cterm "Not"}
+           (Thm.capply (Thm.capply @{cterm "op = :: int => _"} (Numeral.mk_cnumber @{ctyp "int"} x))
            @{cterm "0::int"})))
    in equal_elim (Thm.symmetric th) TrueI end;
-  val notz = let val tab = fold Inttab.update 
-                               (ds ~~ (map (fn x => nzprop (l div x)) ds)) Inttab.empty 
-            in 
-             (fn ct => (valOf (Inttab.lookup tab (ct |> term_of |> dest_numeral)) 
-                handle Option => (writeln "noz: Theorems-Table contains no entry for"; 
-                                    Display.print_cterm ct ; raise Option)))
-           end
-  fun unit_conv t = 
+  val notz =
+    let val tab = fold Inttab.update
+          (ds ~~ (map (fn x => nzprop (l div x)) ds)) Inttab.empty
+    in
+      fn ct => valOf (Inttab.lookup tab (ct |> term_of |> dest_numeral))
+        handle Option =>
+          (writeln ("noz: Theorems-Table contains no entry for " ^
+              Syntax.string_of_term ctxt (Thm.term_of ct)); raise Option)
+    end
+  fun unit_conv t =
    case (term_of t) of
    Const("op &",_)$_$_ => binop_conv unit_conv t
   | Const("op |",_)$_$_ => binop_conv unit_conv t
   | Const (@{const_name Not},_)$_ => arg_conv unit_conv t
-  | Const(s,_)$(Const(@{const_name HOL.times},_)$c$y)$ _ => 
+  | Const(s,_)$(Const(@{const_name HOL.times},_)$c$y)$ _ =>
     if x=y andalso member (op =)
       ["op =", @{const_name HOL.less}, @{const_name HOL.less_eq}] s
     then cv (l div dest_numeral c) t else Thm.reflexive t
-  | Const(s,_)$_$(Const(@{const_name HOL.times},_)$c$y) => 
+  | Const(s,_)$_$(Const(@{const_name HOL.times},_)$c$y) =>
     if x=y andalso member (op =)
       [@{const_name HOL.less}, @{const_name HOL.less_eq}] s
     then cv (l div dest_numeral c) t else Thm.reflexive t
-  | Const(@{const_name Ring_and_Field.dvd},_)$d$(r as (Const(@{const_name HOL.plus},_)$(Const(@{const_name HOL.times},_)$c$y)$_)) => 
-    if x=y then 
-      let 
+  | Const(@{const_name Ring_and_Field.dvd},_)$d$(r as (Const(@{const_name HOL.plus},_)$(Const(@{const_name HOL.times},_)$c$y)$_)) =>
+    if x=y then
+      let
        val k = l div dest_numeral c
        val kt = HOLogic.mk_number iT k
-       val th1 = inst' [Thm.dest_arg1 t, Thm.dest_arg t] 
+       val th1 = inst' [Thm.dest_arg1 t, Thm.dest_arg t]
              ((Thm.dest_arg t |> funpow 2 Thm.dest_arg1 |> notz) RS zdvd_mono)
        val (d',t') = (mulC$kt$d, mulC$kt$r)
        val thc = (provelin ctxt ((HOLogic.eq_const iT)$d'$(lint [] d') |> HOLogic.mk_Trueprop))
                    RS eq_reflection
        val tht = (provelin ctxt ((HOLogic.eq_const iT)$t'$(linear_cmul k r) |> HOLogic.mk_Trueprop))
                  RS eq_reflection
-      in Thm.transitive th1 (Thm.combination (Drule.arg_cong_rule dvdc thc) tht) end                 
+      in Thm.transitive th1 (Thm.combination (Drule.arg_cong_rule dvdc thc) tht) end
     else Thm.reflexive t
   | _ => Thm.reflexive t
   val uth = unit_conv p
@@ -359,7 +363,7 @@
   val ltx = Thm.capply (Thm.capply cmulC clt) cx
   val th = Drule.arg_cong_rule e (Thm.abstract_rule (fst (dest_Free x )) cx uth)
   val th' = inst' [Thm.cabs ltx (Thm.rhs_of uth), clt] unity_coeff_ex
-  val thf = transitive th 
+  val thf = transitive th
       (transitive (symmetric (beta_conversion true (cprop_of th' |> Thm.dest_arg1))) th')
   val (lth,rth) = Thm.dest_comb (cprop_of thf) |>> Thm.dest_arg |>> Thm.beta_conversion true
                   ||> beta_conversion true |>> Thm.symmetric
@@ -372,25 +376,25 @@
 fun mkISet cts = fold_rev (Thm.capply insert_tm #> Thm.capply) cts emptyIS;
 val cTrp = @{cterm "Trueprop"};
 val eqelem_imp_imp = (thm"eqelem_imp_iff") RS iffD1;
-val [A_tm,B_tm] = map (fn th => cprop_of th |> funpow 2 Thm.dest_arg |> Thm.dest_abs NONE |> snd |> Thm.dest_arg1 |> Thm.dest_arg 
+val [A_tm,B_tm] = map (fn th => cprop_of th |> funpow 2 Thm.dest_arg |> Thm.dest_abs NONE |> snd |> Thm.dest_arg1 |> Thm.dest_arg
                                       |> Thm.dest_abs NONE |> snd |> Thm.dest_fun |> Thm.dest_arg)
                       [asetP,bsetP];
 
 val D_tm = @{cpat "?D::int"};
 
-fun cooperex_conv ctxt vs q = 
-let 
+fun cooperex_conv ctxt vs q =
+let
 
  val uth = unify ctxt q
  val (x,p) = Thm.dest_abs NONE (Thm.dest_arg (Thm.rhs_of uth))
  val ins = insert (op aconvc)
- fun h t (bacc,aacc,dacc) = 
+ fun h t (bacc,aacc,dacc) =
   case (whatis x t) of
     And (p,q) => h q (h p (bacc,aacc,dacc))
   | Or (p,q) => h q  (h p (bacc,aacc,dacc))
-  | Eq t => (ins (minus1 t) bacc, 
+  | Eq t => (ins (minus1 t) bacc,
              ins (plus1 t) aacc,dacc)
-  | NEq t => (ins t bacc, 
+  | NEq t => (ins t bacc,
               ins t aacc, dacc)
   | Lt t => (bacc, ins t aacc, dacc)
   | Le t => (bacc, ins (plus1 t) aacc,dacc)
@@ -403,89 +407,92 @@
  val d = Integer.lcms ds
  val cd = Numeral.mk_cnumber @{ctyp "int"} d
  val dt = term_of cd
- fun divprop x = 
-   let 
-    val th = 
-     Simplifier.rewrite lin_ss 
-      (Thm.capply @{cterm Trueprop} 
+ fun divprop x =
+   let
+    val th =
+     Simplifier.rewrite lin_ss
+      (Thm.capply @{cterm Trueprop}
            (Thm.capply (Thm.capply dvdc (Numeral.mk_cnumber @{ctyp "int"} x)) cd))
    in equal_elim (Thm.symmetric th) TrueI end;
- val dvd = let val tab = fold Inttab.update
-                               (ds ~~ (map divprop ds)) Inttab.empty in 
-           (fn ct => (valOf (Inttab.lookup tab (term_of ct |> dest_numeral)) 
-                    handle Option => (writeln "dvd: Theorems-Table contains no entry for"; 
-                                      Display.print_cterm ct ; raise Option)))
-           end
- val dp = 
-   let val th = Simplifier.rewrite lin_ss 
-      (Thm.capply @{cterm Trueprop} 
+ val dvd =
+   let val tab = fold Inttab.update (ds ~~ (map divprop ds)) Inttab.empty in
+     fn ct => valOf (Inttab.lookup tab (term_of ct |> dest_numeral))
+       handle Option =>
+        (writeln ("dvd: Theorems-Table contains no entry for" ^
+            Syntax.string_of_term ctxt (Thm.term_of ct)); raise Option)
+   end
+ val dp =
+   let val th = Simplifier.rewrite lin_ss
+      (Thm.capply @{cterm Trueprop}
            (Thm.capply (Thm.capply @{cterm "op < :: int => _"} @{cterm "0::int"}) cd))
    in equal_elim (Thm.symmetric th) TrueI end;
     (* A and B set *)
-   local 
+   local
      val insI1 = instantiate' [SOME @{ctyp "int"}] [] @{thm "insertI1"}
      val insI2 = instantiate' [SOME @{ctyp "int"}] [] @{thm "insertI2"}
    in
-    fun provein x S = 
+    fun provein x S =
      case term_of S of
         Const(@{const_name Orderings.bot}, _) => error "Unexpected error in Cooper, please email Amine Chaieb"
-      | Const(@{const_name insert}, _) $ y $ _ => 
+      | Const(@{const_name insert}, _) $ y $ _ =>
          let val (cy,S') = Thm.dest_binop S
          in if term_of x aconv y then instantiate' [] [SOME x, SOME S'] insI1
-         else implies_elim (instantiate' [] [SOME x, SOME S', SOME cy] insI2) 
+         else implies_elim (instantiate' [] [SOME x, SOME S', SOME cy] insI2)
                            (provein x S')
          end
    end
- 
+
  val al = map (lint vs o term_of) a0
  val bl = map (lint vs o term_of) b0
- val (sl,s0,f,abths,cpth) = 
-   if length (distinct (op aconv) bl) <= length (distinct (op aconv) al) 
-   then  
+ val (sl,s0,f,abths,cpth) =
+   if length (distinct (op aconv) bl) <= length (distinct (op aconv) al)
+   then
     (bl,b0,decomp_minf,
-     fn B => (map (fn th => implies_elim (Thm.instantiate ([],[(B_tm,B), (D_tm,cd)]) th) dp) 
+     fn B => (map (fn th => implies_elim (Thm.instantiate ([],[(B_tm,B), (D_tm,cd)]) th) dp)
                      [bseteq,bsetneq,bsetlt, bsetle, bsetgt,bsetge])@
-                   (map (Thm.instantiate ([],[(B_tm,B), (D_tm,cd)])) 
+                   (map (Thm.instantiate ([],[(B_tm,B), (D_tm,cd)]))
                         [bsetdvd,bsetndvd,bsetP,infDdvd, infDndvd,bsetconj,
                          bsetdisj,infDconj, infDdisj]),
-                       cpmi) 
-     else (al,a0,decomp_pinf,fn A => 
+                       cpmi)
+     else (al,a0,decomp_pinf,fn A =>
           (map (fn th => implies_elim (Thm.instantiate ([],[(A_tm,A), (D_tm,cd)]) th) dp)
                    [aseteq,asetneq,asetlt, asetle, asetgt,asetge])@
-                   (map (Thm.instantiate ([],[(A_tm,A), (D_tm,cd)])) 
+                   (map (Thm.instantiate ([],[(A_tm,A), (D_tm,cd)]))
                    [asetdvd,asetndvd, asetP, infDdvd, infDndvd,asetconj,
                          asetdisj,infDconj, infDdisj]),cppi)
- val cpth = 
+ val cpth =
   let
-   val sths = map (fn (tl,t0) => 
-                      if tl = term_of t0 
+   val sths = map (fn (tl,t0) =>
+                      if tl = term_of t0
                       then instantiate' [SOME @{ctyp "int"}] [SOME t0] refl
-                      else provelin ctxt ((HOLogic.eq_const iT)$tl$(term_of t0) 
-                                 |> HOLogic.mk_Trueprop)) 
+                      else provelin ctxt ((HOLogic.eq_const iT)$tl$(term_of t0)
+                                 |> HOLogic.mk_Trueprop))
                    (sl ~~ s0)
    val csl = distinct (op aconvc) (map (cprop_of #> Thm.dest_arg #> Thm.dest_arg1) sths)
    val S = mkISet csl
-   val inStab = fold (fn ct => fn tab => Termtab.update (term_of ct, provein ct S) tab) 
+   val inStab = fold (fn ct => fn tab => Termtab.update (term_of ct, provein ct S) tab)
                     csl Termtab.empty
    val eqelem_th = instantiate' [SOME @{ctyp "int"}] [NONE,NONE, SOME S] eqelem_imp_imp
-   val inS = 
-     let 
-      fun transmem th0 th1 = 
-       Thm.equal_elim 
-        (Drule.arg_cong_rule cTrp (Drule.fun_cong_rule (Drule.arg_cong_rule 
+   val inS =
+     let
+      fun transmem th0 th1 =
+       Thm.equal_elim
+        (Drule.arg_cong_rule cTrp (Drule.fun_cong_rule (Drule.arg_cong_rule
                ((Thm.dest_fun o Thm.dest_fun o Thm.dest_arg o cprop_of) th1) th0) S)) th1
       val tab = fold Termtab.update
-        (map (fn eq => 
-                let val (s,t) = cprop_of eq |> Thm.dest_arg |> Thm.dest_binop 
-                    val th = if term_of s = term_of t 
+        (map (fn eq =>
+                let val (s,t) = cprop_of eq |> Thm.dest_arg |> Thm.dest_binop
+                    val th = if term_of s = term_of t
                              then valOf(Termtab.lookup inStab (term_of s))
-                             else FWD (instantiate' [] [SOME s, SOME t] eqelem_th) 
+                             else FWD (instantiate' [] [SOME s, SOME t] eqelem_th)
                                 [eq, valOf(Termtab.lookup inStab (term_of s))]
                  in (term_of t, th) end)
                   sths) Termtab.empty
-        in fn ct => 
-          (valOf (Termtab.lookup tab (term_of ct))
-           handle Option => (writeln "inS: No theorem for " ; Display.print_cterm ct ; raise Option))
+        in
+          fn ct => valOf (Termtab.lookup tab (term_of ct))
+            handle Option =>
+              (writeln ("inS: No theorem for " ^ Syntax.string_of_term ctxt (Thm.term_of ct));
+                raise Option)
         end
        val (inf, nb, pd) = divide_and_conquer (f x dvd inS (abths S)) p
    in [dp, inf, nb, pd] MRS cpth
@@ -494,9 +501,9 @@
 in Thm.transitive cpth' ((simp_thms_conv ctxt then_conv eval_conv) (Thm.rhs_of cpth'))
 end;
 
-fun literals_conv bops uops env cv = 
+fun literals_conv bops uops env cv =
  let fun h t =
-  case (term_of t) of 
+  case (term_of t) of
    b$_$_ => if member (op aconv) bops b then binop_conv h t else cv env t
  | u$_ => if member (op aconv) uops u then arg_conv h t else cv env t
  | _ => cv env t
@@ -506,21 +513,21 @@
  nnf_conv then_conv literals_conv [HOLogic.conj, HOLogic.disj] [] env (linearize_conv ctxt);
 
 local
- val pcv = Simplifier.rewrite 
-     (HOL_basic_ss addsimps (simp_thms @ (List.take(ex_simps,4)) 
+ val pcv = Simplifier.rewrite
+     (HOL_basic_ss addsimps (simp_thms @ (List.take(ex_simps,4))
                       @ [not_all,all_not_ex, ex_disj_distrib]))
  val postcv = Simplifier.rewrite presburger_ss
- fun conv ctxt p = 
+ fun conv ctxt p =
   let val _ = ()
   in
-   Qelim.gen_qelim_conv pcv postcv pcv (cons o term_of) 
-      (OldTerm.term_frees (term_of p)) (linearize_conv ctxt) (integer_nnf_conv ctxt) 
-      (cooperex_conv ctxt) p 
+   Qelim.gen_qelim_conv pcv postcv pcv (cons o term_of)
+      (OldTerm.term_frees (term_of p)) (linearize_conv ctxt) (integer_nnf_conv ctxt)
+      (cooperex_conv ctxt) p
   end
   handle  CTERM s => raise COOPER ("Cooper Failed", CTERM s)
-        | THM s => raise COOPER ("Cooper Failed", THM s) 
-        | TYPE s => raise COOPER ("Cooper Failed", TYPE s) 
-in val cooper_conv = conv 
+        | THM s => raise COOPER ("Cooper Failed", THM s)
+        | TYPE s => raise COOPER ("Cooper Failed", TYPE s)
+in val cooper_conv = conv
 end;
 end;
 
@@ -542,12 +549,12 @@
   | Const(@{const_name HOL.uminus},_)$t' => Neg (i_of_term vs t')
   | Const(@{const_name HOL.plus},_)$t1$t2 => Add (i_of_term vs t1,i_of_term vs t2)
   | Const(@{const_name HOL.minus},_)$t1$t2 => Sub (i_of_term vs t1,i_of_term vs t2)
-  | Const(@{const_name HOL.times},_)$t1$t2 => 
+  | Const(@{const_name HOL.times},_)$t1$t2 =>
      (Mul (HOLogic.dest_number t1 |> snd, i_of_term vs t2)
-    handle TERM _ => 
+    handle TERM _ =>
        (Mul (HOLogic.dest_number t2 |> snd, i_of_term vs t1)
         handle TERM _ => cooper "Reification: Unsupported kind of multiplication"))
-  | _ => (C (HOLogic.dest_number t |> snd) 
+  | _ => (C (HOLogic.dest_number t |> snd)
            handle TERM _ => cooper "Reification: unknown term");
 
 fun qf_of_term ps vs t =  case t
@@ -555,7 +562,7 @@
   | Const("False",_) => F
   | Const(@{const_name HOL.less},_)$t1$t2 => Lt (Sub (i_of_term vs t1,i_of_term vs t2))
   | Const(@{const_name HOL.less_eq},_)$t1$t2 => Le (Sub(i_of_term vs t1,i_of_term vs t2))
-  | Const(@{const_name Ring_and_Field.dvd},_)$t1$t2 => 
+  | Const(@{const_name Ring_and_Field.dvd},_)$t1$t2 =>
       (Dvd(HOLogic.dest_number t1 |> snd, i_of_term vs t2) handle _ => cooper "Reification: unsupported dvd")  (* FIXME avoid handle _ *)
   | @{term "op = :: int => _"}$t1$t2 => Eq (Sub (i_of_term vs t1,i_of_term vs t2))
   | @{term "op = :: bool => _ "}$t1$t2 => Iff(qf_of_term ps vs t1,qf_of_term ps vs t2)
@@ -563,42 +570,42 @@
   | Const("op |",_)$t1$t2 => Or(qf_of_term ps vs t1,qf_of_term ps vs t2)
   | Const("op -->",_)$t1$t2 => Imp(qf_of_term ps vs t1,qf_of_term ps vs t2)
   | Const (@{const_name Not},_)$t' => Not(qf_of_term ps vs t')
-  | Const("Ex",_)$Abs(xn,xT,p) => 
+  | Const("Ex",_)$Abs(xn,xT,p) =>
      let val (xn',p') = variant_abs (xn,xT,p)
          val vs' = (Free (xn',xT), 0) :: (map (fn(v,n) => (v,1+ n)) vs)
      in E (qf_of_term ps vs' p')
      end
-  | Const("All",_)$Abs(xn,xT,p) => 
+  | Const("All",_)$Abs(xn,xT,p) =>
      let val (xn',p') = variant_abs (xn,xT,p)
          val vs' = (Free (xn',xT), 0) :: (map (fn(v,n) => (v,1+ n)) vs)
      in A (qf_of_term ps vs' p')
      end
-  | _ =>(case AList.lookup (op aconv) ps t of 
+  | _ =>(case AList.lookup (op aconv) ps t of
            NONE => cooper "Reification: unknown term!"
          | SOME n => Closed n);
 
 local
  val ops = [@{term "op &"}, @{term "op |"}, @{term "op -->"}, @{term "op = :: bool => _"},
-             @{term "op = :: int => _"}, @{term "op < :: int => _"}, 
-             @{term "op <= :: int => _"}, @{term "Not"}, @{term "All:: (int => _) => _"}, 
+             @{term "op = :: int => _"}, @{term "op < :: int => _"},
+             @{term "op <= :: int => _"}, @{term "Not"}, @{term "All:: (int => _) => _"},
              @{term "Ex:: (int => _) => _"}, @{term "True"}, @{term "False"}]
 fun ty t = Bool.not (fastype_of t = HOLogic.boolT)
 in
 fun term_bools acc t =
-case t of 
-    (l as f $ a) $ b => if ty t orelse f mem ops then term_bools (term_bools acc l)b 
+case t of
+    (l as f $ a) $ b => if ty t orelse f mem ops then term_bools (term_bools acc l)b
             else insert (op aconv) t acc
-  | f $ a => if ty t orelse f mem ops then term_bools (term_bools acc f) a  
+  | f $ a => if ty t orelse f mem ops then term_bools (term_bools acc f) a
             else insert (op aconv) t acc
   | Abs p => term_bools acc (snd (variant_abs p))
   | _ => if ty t orelse t mem ops then acc else insert (op aconv) t acc
 end;
- 
+
 fun myassoc2 l v =
     case l of
-	[] => NONE
+  [] => NONE
       | (x,v')::xs => if v = v' then SOME x
-		      else myassoc2 xs v;
+          else myassoc2 xs v;
 
 fun term_of_i vs t = case t
  of C i => HOLogic.mk_number HOLogic.intT i
@@ -610,9 +617,9 @@
       HOLogic.mk_number HOLogic.intT i $ term_of_i vs t2
   | Cn (n, i, t') => term_of_i vs (Add (Mul (i, Bound n), t'));
 
-fun term_of_qf ps vs t = 
- case t of 
-   T => HOLogic.true_const 
+fun term_of_qf ps vs t =
+ case t of
+   T => HOLogic.true_const
  | F => HOLogic.false_const
  | Lt t' => @{term "op < :: int => _ "}$ term_of_i vs t'$ @{term "0::int"}
  | Le t' => @{term "op <= :: int => _ "}$ term_of_i vs t' $ @{term "0::int"}
@@ -620,7 +627,7 @@
  | Ge t' => @{term "op <= :: int => _ "}$ @{term "0::int"}$ term_of_i vs t'
  | Eq t' => @{term "op = :: int => _ "}$ term_of_i vs t'$ @{term "0::int"}
  | NEq t' => term_of_qf ps vs (Not (Eq t'))
- | Dvd(i,t') => @{term "op dvd :: int => _ "} $ 
+ | Dvd(i,t') => @{term "op dvd :: int => _ "} $
     HOLogic.mk_number HOLogic.intT i $ term_of_i vs t'
  | NDvd(i,t')=> term_of_qf ps vs (Not(Dvd(i,t')))
  | Not t' => HOLogic.Not$(term_of_qf ps vs t')
--- a/src/HOL/Tools/TFL/rules.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Tools/TFL/rules.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -456,7 +456,7 @@
 fun is_cong thm =
   case (Thm.prop_of thm)
      of (Const("==>",_)$(Const("Trueprop",_)$ _) $
-         (Const("==",_) $ (Const (@{const_name "Wellfounded.cut"},_) $ f $ R $ a $ x) $ _)) => false
+         (Const("==",_) $ (Const (@{const_name "Recdef.cut"},_) $ f $ R $ a $ x) $ _)) => false
       | _ => true;
 
 
@@ -544,7 +544,7 @@
 (*---------------------------------------------------------------------------
  * Trace information for the rewriter
  *---------------------------------------------------------------------------*)
-val term_ref = ref[] : term list ref
+val term_ref = ref [] : term list ref
 val ss_ref = ref [] : simpset list ref;
 val thm_ref = ref [] : thm list ref;
 val tracing = ref false;
@@ -554,8 +554,8 @@
 fun print_thms s L =
   say (cat_lines (s :: map Display.string_of_thm_without_context L));
 
-fun print_cterms s L =
-  say (cat_lines (s :: map Display.string_of_cterm L));
+fun print_cterm s ct =
+  say (cat_lines [s, Syntax.string_of_term_global (Thm.theory_of_cterm ct) (Thm.term_of ct)]);
 
 
 (*---------------------------------------------------------------------------
@@ -659,7 +659,7 @@
   end;
 
 fun restricted t = isSome (S.find_term
-                            (fn (Const(@{const_name "Wellfounded.cut"},_)) =>true | _ => false)
+                            (fn (Const(@{const_name "Recdef.cut"},_)) =>true | _ => false)
                             t)
 
 fun CONTEXT_REWRITE_RULE (func, G, cut_lemma, congs) th =
@@ -683,7 +683,7 @@
              (* Unquantified eliminate *)
              fun uq_eliminate (thm,imp,thy) =
                  let val tych = cterm_of thy
-                     val dummy = print_cterms "To eliminate:" [tych imp]
+                     val dummy = print_cterm "To eliminate:" (tych imp)
                      val ants = map tych (Logic.strip_imp_prems imp)
                      val eq = Logic.strip_imp_concl imp
                      val lhs = tych(get_lhs eq)
@@ -781,9 +781,8 @@
               val antl = case rcontext of [] => []
                          | _   => [S.list_mk_conj(map cncl rcontext)]
               val TC = genl(S.list_mk_imp(antl, A))
-              val dummy = print_cterms "func:" [cterm_of thy func]
-              val dummy = print_cterms "TC:"
-                              [cterm_of thy (HOLogic.mk_Trueprop TC)]
+              val dummy = print_cterm "func:" (cterm_of thy func)
+              val dummy = print_cterm "TC:" (cterm_of thy (HOLogic.mk_Trueprop TC))
               val dummy = tc_list := (TC :: !tc_list)
               val nestedp = isSome (S.find_term is_func TC)
               val dummy = if nestedp then say "nested" else say "not_nested"
--- a/src/HOL/Tools/TFL/tfl.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Tools/TFL/tfl.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -8,7 +8,7 @@
 sig
   val trace: bool ref
   val trace_thms: string -> thm list -> unit
-  val trace_cterms: string -> cterm list -> unit
+  val trace_cterm: string -> cterm -> unit
   type pattern
   val mk_functional: theory -> term list -> {functional: term, pats: pattern list}
   val wfrec_definition0: theory -> string -> term -> term -> theory * thm
@@ -296,7 +296,7 @@
             raise TFL_ERR "no_repeat_vars"
                           (quote (#1 (dest_Free v)) ^
                           " occurs repeatedly in the pattern " ^
-                          quote (Display.string_of_cterm (Thry.typecheck thy pat)))
+                          quote (Syntax.string_of_term_global thy pat))
          else check rst
  in check (FV_multiset pat)
  end;
@@ -912,9 +912,10 @@
   if !trace then writeln (cat_lines (s :: map Display.string_of_thm_without_context L))
   else ();
 
-fun trace_cterms s L =
-  if !trace then writeln (cat_lines (s :: map Display.string_of_cterm L))
-  else ();;
+fun trace_cterm s ct =
+  if !trace then
+    writeln (cat_lines [s, Syntax.string_of_term_global (Thm.theory_of_cterm ct) (Thm.term_of ct)])
+  else ();
 
 
 fun postprocess strict {wf_tac, terminator, simplifier} theory {rules,induction,TCs} =
@@ -942,7 +943,7 @@
 
    fun simplify_tc tc (r,ind) =
        let val tc1 = tych tc
-           val _ = trace_cterms "TC before simplification: " [tc1]
+           val _ = trace_cterm "TC before simplification: " tc1
            val tc_eq = simplifier tc1
            val _ = trace_thms "result: " [tc_eq]
        in
--- a/src/HOL/Tools/hologic.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Tools/hologic.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -586,7 +586,7 @@
 
 (* string *)
 
-val stringT = Type ("String.string", []);
+val stringT = listT charT;
 
 val mk_string = mk_list charT o map (mk_char o ord) o explode;
 val dest_string = implode o map (chr o dest_char) o dest_list;
--- a/src/HOL/Tools/metis_tools.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Tools/metis_tools.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -10,7 +10,7 @@
   val type_lits: bool Config.T
   val metis_tac: Proof.context -> thm list -> int -> tactic
   val metisF_tac: Proof.context -> thm list -> int -> tactic
-  val metisH_tac: Proof.context -> thm list -> int -> tactic
+  val metisFT_tac: Proof.context -> thm list -> int -> tactic
   val setup: theory -> theory
 end
 
@@ -21,6 +21,8 @@
 
   val (type_lits, type_lits_setup) = Attrib.config_bool "metis_type_lits" true;
 
+  datatype mode = FO | HO | FT  (*first-order, higher-order, fully-typed*)
+
   (* ------------------------------------------------------------------------- *)
   (* Useful Theorems                                                           *)
   (* ------------------------------------------------------------------------- *)
@@ -81,20 +83,36 @@
     | hol_term_to_fol_HO (ResHolClause.CombApp(tm1,tm2)) =
          Metis.Term.Fn(".", map hol_term_to_fol_HO [tm1,tm2]);
 
-  fun hol_literal_to_fol true (ResHolClause.Literal (pol, tm)) =  (*first-order*)
+  (*The fully-typed translation, to avoid type errors*)
+  fun wrap_type (tm, ty) = Metis.Term.Fn("ti", [tm, hol_type_to_fol ty]);
+  
+  fun hol_term_to_fol_FT (ResHolClause.CombVar(a, ty)) = 
+        wrap_type (Metis.Term.Var a, ty)
+    | hol_term_to_fol_FT (ResHolClause.CombConst(a, ty, _)) =
+        wrap_type (Metis.Term.Fn(fn_isa_to_met a, []), ty)
+    | hol_term_to_fol_FT (tm as ResHolClause.CombApp(tm1,tm2)) =
+         wrap_type (Metis.Term.Fn(".", map hol_term_to_fol_FT [tm1,tm2]),    
+                    ResHolClause.type_of_combterm tm);
+
+  fun hol_literal_to_fol FO (ResHolClause.Literal (pol, tm)) =  
         let val (ResHolClause.CombConst(p,_,tys), tms) = ResHolClause.strip_comb tm
             val tylits = if p = "equal" then [] else map hol_type_to_fol tys
             val lits = map hol_term_to_fol_FO tms
         in metis_lit pol (fn_isa_to_met p) (tylits @ lits) end
-    | hol_literal_to_fol false (ResHolClause.Literal (pol, tm)) =    (*higher-order*)
-        case ResHolClause.strip_comb tm of
+    | hol_literal_to_fol HO (ResHolClause.Literal (pol, tm)) =
+       (case ResHolClause.strip_comb tm of
             (ResHolClause.CombConst("equal",_,_), tms) =>
               metis_lit pol "=" (map hol_term_to_fol_HO tms)
-          | _ => metis_lit pol "{}" [hol_term_to_fol_HO tm];
+          | _ => metis_lit pol "{}" [hol_term_to_fol_HO tm])   (*hBOOL*)
+    | hol_literal_to_fol FT (ResHolClause.Literal (pol, tm)) = 
+       (case ResHolClause.strip_comb tm of
+            (ResHolClause.CombConst("equal",_,_), tms) =>
+              metis_lit pol "=" (map hol_term_to_fol_FT tms)
+          | _ => metis_lit pol "{}" [hol_term_to_fol_FT tm])   (*hBOOL*);
 
-  fun literals_of_hol_thm thy isFO t =
+  fun literals_of_hol_thm thy mode t =
         let val (lits, types_sorts) = ResHolClause.literals_of_term thy t
-        in  (map (hol_literal_to_fol isFO) lits, types_sorts) end;
+        in  (map (hol_literal_to_fol mode) lits, types_sorts) end;
 
   (*Sign should be "true" for conjecture type constraints, "false" for type lits in clauses.*)
   fun metis_of_typeLit pos (ResClause.LTVar (s,x))  = metis_lit pos s [Metis.Term.Var x]
@@ -106,10 +124,10 @@
   fun metis_of_tfree tf =
     Metis.Thm.axiom (Metis.LiteralSet.singleton (metis_of_typeLit true tf));
 
-  fun hol_thm_to_fol is_conjecture ctxt isFO th =
+  fun hol_thm_to_fol is_conjecture ctxt mode th =
     let val thy = ProofContext.theory_of ctxt
         val (mlits, types_sorts) =
-               (literals_of_hol_thm thy isFO o HOLogic.dest_Trueprop o prop_of) th
+               (literals_of_hol_thm thy mode o HOLogic.dest_Trueprop o prop_of) th
     in
         if is_conjecture then
             (Metis.Thm.axiom (Metis.LiteralSet.fromList mlits), ResClause.add_typs types_sorts)
@@ -171,17 +189,13 @@
             " received " ^ commas (map (Syntax.string_of_term_global Pure.thy) trands))
     end;
 
-(*Instantiate constant c with the supplied types, but if they don't match its declared
-  sort constraints, replace by a general type.*)
-fun const_of ctxt (c,Ts) =  Const (c, dummyT)
-
 fun infer_types ctxt =
   Syntax.check_terms (ProofContext.set_mode ProofContext.mode_pattern ctxt);
 
   (*We use 1 rather than 0 because variable references in clauses may otherwise conflict
     with variable constraints in the goal...at least, type inference often fails otherwise.
     SEE ALSO axiom_inf below.*)
-  fun mk_var w = Term.Var((w,1), HOLogic.typeT);
+  fun mk_var (w,T) = Term.Var((w,1), T);
 
   (*include the default sort, if available*)
   fun mk_tfree ctxt w =
@@ -192,6 +206,18 @@
   fun strip_happ args (Metis.Term.Fn(".",[t,u])) = strip_happ (u::args) t
     | strip_happ args x = (x, args);
 
+  fun fol_type_to_isa ctxt (Metis.Term.Var v) = 
+       (case Recon.strip_prefix ResClause.tvar_prefix v of
+	    SOME w => Recon.make_tvar w
+	  | NONE   => Recon.make_tvar v)
+    | fol_type_to_isa ctxt (Metis.Term.Fn(x, tys)) =
+       (case Recon.strip_prefix ResClause.tconst_prefix x of
+	    SOME tc => Term.Type (Recon.invert_type_const tc, map (fol_type_to_isa ctxt) tys)
+	  | NONE    => 
+        case Recon.strip_prefix ResClause.tfree_prefix x of
+	    SOME tf => mk_tfree ctxt tf
+	  | NONE    => error ("fol_type_to_isa: " ^ x));
+
   (*Maps metis terms to isabelle terms*)
   fun fol_term_to_hol_RAW ctxt fol_tm =
     let val thy = ProofContext.theory_of ctxt
@@ -201,8 +227,8 @@
                     SOME w => Type (Recon.make_tvar w)
                   | NONE =>
                 case Recon.strip_prefix ResClause.schematic_var_prefix v of
-                    SOME w => Term (mk_var w)
-                  | NONE   => Term (mk_var v) )
+                    SOME w => Term (mk_var (w, HOLogic.typeT))
+                  | NONE   => Term (mk_var (v, HOLogic.typeT)) )
                       (*Var from Metis with a name like _nnn; possibly a type variable*)
           | tm_to_tt (Metis.Term.Fn ("{}", [arg])) = tm_to_tt arg   (*hBOOL*)
           | tm_to_tt (t as Metis.Term.Fn (".",_)) =
@@ -226,7 +252,7 @@
                         val tys = types_of (List.take(tts,ntypes))
                         val ntyargs = Recon.num_typargs thy c
                     in if length tys = ntyargs then
-                           apply_list (const_of ctxt (c, tys)) nterms (List.drop(tts,ntypes))
+                           apply_list (Const (c, dummyT)) nterms (List.drop(tts,ntypes))
                        else error ("Constant " ^ c ^ " expects " ^ Int.toString ntyargs ^
                                    " but gets " ^ Int.toString (length tys) ^
                                    " type arguments\n" ^
@@ -248,8 +274,45 @@
                 | NONE => error ("unexpected metis function: " ^ a)
     in  case tm_to_tt fol_tm of Term t => t | _ => error "fol_tm_to_tt: Term expected"  end;
 
-  fun fol_terms_to_hol ctxt fol_tms =
-    let val ts = map (fol_term_to_hol_RAW ctxt) fol_tms
+  (*Maps fully-typed metis terms to isabelle terms*)
+  fun fol_term_to_hol_FT ctxt fol_tm =
+    let val _ = Output.debug (fn () => "fol_term_to_hol_FT: " ^ Metis.Term.toString fol_tm)
+        fun cvt (Metis.Term.Fn ("ti", [Metis.Term.Var v, ty])) =
+               (case Recon.strip_prefix ResClause.schematic_var_prefix v of
+                    SOME w =>  mk_var(w, dummyT)
+                  | NONE   => mk_var(v, dummyT))
+          | cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn ("=",[]), ty])) = 
+              Const ("op =", HOLogic.typeT)
+          | cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn (x,[]), ty])) =
+             (case Recon.strip_prefix ResClause.const_prefix x of
+                  SOME c => Const (Recon.invert_const c, dummyT)
+                | NONE => (*Not a constant. Is it a fixed variable??*)
+              case Recon.strip_prefix ResClause.fixed_var_prefix x of
+                  SOME v => Free (v, fol_type_to_isa ctxt ty)
+                | NONE => error ("fol_term_to_hol_FT bad constant: " ^ x))
+          | cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn (".",[tm1,tm2]), _])) =
+              cvt tm1 $ cvt tm2
+          | cvt (Metis.Term.Fn (".",[tm1,tm2])) = (*untyped application*)
+              cvt tm1 $ cvt tm2
+          | cvt (Metis.Term.Fn ("{}", [arg])) = cvt arg   (*hBOOL*)
+          | cvt (Metis.Term.Fn ("=", [tm1,tm2])) = 
+              list_comb(Const ("op =", HOLogic.typeT), map cvt [tm1,tm2])
+          | cvt (t as Metis.Term.Fn (x, [])) = 
+             (case Recon.strip_prefix ResClause.const_prefix x of
+                  SOME c => Const (Recon.invert_const c, dummyT)
+                | NONE => (*Not a constant. Is it a fixed variable??*)
+              case Recon.strip_prefix ResClause.fixed_var_prefix x of
+                  SOME v => Free (v, dummyT)
+                | NONE =>  (Output.debug (fn () => "fol_term_to_hol_FT bad const: " ^ x); fol_term_to_hol_RAW ctxt t))
+          | cvt t = (Output.debug (fn () => "fol_term_to_hol_FT bad term: " ^ Metis.Term.toString t); fol_term_to_hol_RAW ctxt t)
+    in  cvt fol_tm   end;
+
+  fun fol_term_to_hol ctxt FO = fol_term_to_hol_RAW ctxt
+    | fol_term_to_hol ctxt HO = fol_term_to_hol_RAW ctxt
+    | fol_term_to_hol ctxt FT = fol_term_to_hol_FT ctxt;
+
+  fun fol_terms_to_hol ctxt mode fol_tms =
+    let val ts = map (fol_term_to_hol ctxt mode) fol_tms
         val _ = Output.debug (fn () => "  calling type inference:")
         val _ = app (fn t => Output.debug (fn () => Syntax.string_of_term ctxt t)) ts
         val ts' = infer_types ctxt ts;
@@ -262,6 +325,8 @@
   fun mk_not (Const ("Not", _) $ b) = b
     | mk_not b = HOLogic.mk_not b;
 
+  val metis_eq = Metis.Term.Fn ("=", []);
+
   (* ------------------------------------------------------------------------- *)
   (* FOL step Inference Rules                                                  *)
   (* ------------------------------------------------------------------------- *)
@@ -291,22 +356,22 @@
       (*This causes variables to have an index of 1 by default. SEE ALSO mk_var above.*)
 
   (* INFERENCE RULE: ASSUME *)
-  fun assume_inf ctxt atm =
+  fun assume_inf ctxt mode atm =
     inst_excluded_middle
       (ProofContext.theory_of ctxt)
-      (singleton (fol_terms_to_hol ctxt) (Metis.Term.Fn atm));
+      (singleton (fol_terms_to_hol ctxt mode) (Metis.Term.Fn atm));
 
   (* INFERENCE RULE: INSTANTIATE (Subst). Type instantiations are ignored. Trying to reconstruct
      them admits new possibilities of errors, e.g. concerning sorts. Instead we try to arrange
      that new TVars are distinct and that types can be inferred from terms.*)
-  fun inst_inf ctxt thpairs fsubst th =    
+  fun inst_inf ctxt mode thpairs fsubst th =    
     let val thy = ProofContext.theory_of ctxt
         val i_th   = lookth thpairs th
         val i_th_vars = Term.add_vars (prop_of i_th) []
         fun find_var x = valOf (List.find (fn ((a,_),_) => a=x) i_th_vars)
         fun subst_translation (x,y) =
               let val v = find_var x
-                  val t = fol_term_to_hol_RAW ctxt y (*we call infer_types below*)
+                  val t = fol_term_to_hol ctxt mode y (*we call infer_types below*)
               in  SOME (cterm_of thy (Var v), t)  end
               handle Option =>
                   (Output.debug (fn() => "List.find failed for the variable " ^ x ^
@@ -324,10 +389,11 @@
         val tms = infer_types ctxt rawtms;
         val ctm_of = cterm_incr_types thy (1 + Thm.maxidx_of i_th)
         val substs' = ListPair.zip (vars, map ctm_of tms)
-        val _ = Output.debug (fn() => "subst_translations:")
-        val _ = app (fn (x,y) => Output.debug (fn () => Display.string_of_cterm x ^ " |-> " ^
-                                                        Display.string_of_cterm y))
-                  substs'
+        val _ = Output.debug (fn () =>
+          cat_lines ("subst_translations:" ::
+            (substs' |> map (fn (x, y) =>
+              Syntax.string_of_term ctxt (term_of x) ^ " |-> " ^
+              Syntax.string_of_term ctxt (term_of y)))));
     in  cterm_instantiate substs' i_th  
         handle THM (msg, _, _) => error ("metis error (inst_inf): " ^ msg)
     end;
@@ -346,7 +412,7 @@
 	| _ => raise THM ("resolve_inc_tyvars: unique result expected", i, [tha,thb])
     end;
 
-  fun resolve_inf ctxt thpairs atm th1 th2 =
+  fun resolve_inf ctxt mode thpairs atm th1 th2 =
     let
       val thy = ProofContext.theory_of ctxt
       val i_th1 = lookth thpairs th1 and i_th2 = lookth thpairs th2
@@ -357,7 +423,7 @@
       else if is_TrueI i_th2 then i_th1
       else
         let
-          val i_atm = singleton (fol_terms_to_hol ctxt) (Metis.Term.Fn atm)
+          val i_atm = singleton (fol_terms_to_hol ctxt mode) (Metis.Term.Fn atm)
           val _ = Output.debug (fn () => "  atom: " ^ Syntax.string_of_term ctxt i_atm)
           val prems_th1 = prems_of i_th1
           val prems_th2 = prems_of i_th2
@@ -374,9 +440,9 @@
   val refl_x = cterm_of @{theory} (Var (hd (Term.add_vars (prop_of REFL_THM) [])));
   val refl_idx = 1 + Thm.maxidx_of REFL_THM;
 
-  fun refl_inf ctxt t =
+  fun refl_inf ctxt mode t =
     let val thy = ProofContext.theory_of ctxt
-        val i_t = singleton (fol_terms_to_hol ctxt) t
+        val i_t = singleton (fol_terms_to_hol ctxt mode) t
         val _ = Output.debug (fn () => "  term: " ^ Syntax.string_of_term ctxt i_t)
         val c_t = cterm_incr_types thy refl_idx i_t
     in  cterm_instantiate [(refl_x, c_t)] REFL_THM  end;
@@ -386,41 +452,64 @@
     | get_ty_arg_size thy _      = 0;
 
   (* INFERENCE RULE: EQUALITY *)
-  fun equality_inf ctxt isFO thpairs (pos,atm) fp fr =
+  fun equality_inf ctxt mode thpairs (pos,atm) fp fr =
     let val thy = ProofContext.theory_of ctxt
-        val [i_atm,i_tm] = fol_terms_to_hol ctxt [Metis.Term.Fn atm, fr]
+        val m_tm = Metis.Term.Fn atm
+        val [i_atm,i_tm] = fol_terms_to_hol ctxt mode [m_tm, fr]
         val _ = Output.debug (fn () => "sign of the literal: " ^ Bool.toString pos)
         fun replace_item_list lx 0 (l::ls) = lx::ls
           | replace_item_list lx i (l::ls) = l :: replace_item_list lx (i-1) ls
-        fun path_finder_FO tm (p::ps) =
+        fun path_finder_FO tm [] = (tm, Term.Bound 0)
+          | path_finder_FO tm (p::ps) =
               let val (tm1,args) = Term.strip_comb tm
                   val adjustment = get_ty_arg_size thy tm1
                   val p' = if adjustment > p then p else p-adjustment
                   val tm_p = List.nth(args,p')
                     handle Subscript => error ("equality_inf: " ^ Int.toString p ^ " adj " ^
                       Int.toString adjustment  ^ " term " ^  Syntax.string_of_term ctxt tm)
+                  val _ = Output.debug (fn () => "path_finder: " ^ Int.toString p ^
+                                        "  " ^ Syntax.string_of_term ctxt tm_p)
+		  val (r,t) = path_finder_FO tm_p ps
               in
-                  Output.debug (fn () => "path_finder: " ^ Int.toString p ^
-                                        "  " ^ Syntax.string_of_term ctxt tm_p);
-                  if null ps   (*FIXME: why not use pattern-matching and avoid repetition*)
-                  then (tm_p, list_comb (tm1, replace_item_list (Term.Bound 0) p' args))
-                  else let val (r,t) = path_finder_FO tm_p ps
-                       in (r, list_comb (tm1, replace_item_list t p' args)) end
+                  (r, list_comb (tm1, replace_item_list t p' args)) 
               end
         fun path_finder_HO tm [] = (tm, Term.Bound 0)
           | path_finder_HO (t$u) (0::ps) = (fn(x,y) => (x, y$u)) (path_finder_HO t ps)
           | path_finder_HO (t$u) (p::ps) = (fn(x,y) => (x, t$y)) (path_finder_HO u ps)
-        fun path_finder true tm ps = path_finder_FO tm ps
-          | path_finder false (tm as Const("op =",_) $ _ $ _) (p::ps) =
+        fun path_finder_FT tm [] _ = (tm, Term.Bound 0)
+          | path_finder_FT tm (0::ps) (Metis.Term.Fn ("ti", [t1,t2])) = 
+              path_finder_FT tm ps t1
+          | path_finder_FT (t$u) (0::ps) (Metis.Term.Fn (".", [t1,t2])) = 
+              (fn(x,y) => (x, y$u)) (path_finder_FT t ps t1)
+          | path_finder_FT (t$u) (1::ps) (Metis.Term.Fn (".", [t1,t2])) = 
+              (fn(x,y) => (x, t$y)) (path_finder_FT u ps t2)
+          | path_finder_FT tm ps t = error ("equality_inf, path_finder_FT: path = " ^
+                                          space_implode " " (map Int.toString ps) ^ 
+                                          " isa-term: " ^  Syntax.string_of_term ctxt tm ^
+                                          " fol-term: " ^ Metis.Term.toString t)
+        fun path_finder FO tm ps _ = path_finder_FO tm ps
+          | path_finder HO (tm as Const("op =",_) $ _ $ _) (p::ps) _ =
                (*equality: not curried, as other predicates are*)
                if p=0 then path_finder_HO tm (0::1::ps)  (*select first operand*)
                else path_finder_HO tm (p::ps)        (*1 selects second operand*)
-          | path_finder false tm (p::ps) =
+          | path_finder HO tm (p::ps) (Metis.Term.Fn ("{}", [t1])) =
                path_finder_HO tm ps      (*if not equality, ignore head to skip hBOOL*)
+          | path_finder FT (tm as Const("op =",_) $ _ $ _) (p::ps) 
+                              (Metis.Term.Fn ("=", [t1,t2])) =
+               (*equality: not curried, as other predicates are*)
+               if p=0 then path_finder_FT tm (0::1::ps) 
+                            (Metis.Term.Fn (".", [Metis.Term.Fn (".", [metis_eq,t1]), t2])) 
+                            (*select first operand*)
+               else path_finder_FT tm (p::ps) 
+                     (Metis.Term.Fn (".", [metis_eq,t2])) 
+                     (*1 selects second operand*)
+          | path_finder FT tm (p::ps) (Metis.Term.Fn ("{}", [t1])) = path_finder_FT tm ps t1     
+               (*if not equality, ignore head to skip the hBOOL predicate*)
+          | path_finder FT tm ps t = path_finder_FT tm ps t  (*really an error case!*)
         fun path_finder_lit ((nt as Term.Const ("Not", _)) $ tm_a) idx =
-              let val (tm, tm_rslt) = path_finder isFO tm_a idx
+              let val (tm, tm_rslt) = path_finder mode tm_a idx m_tm
               in (tm, nt $ tm_rslt) end
-          | path_finder_lit tm_a idx = path_finder isFO tm_a idx
+          | path_finder_lit tm_a idx = path_finder mode tm_a idx m_tm
         val (tm_subst, body) = path_finder_lit i_atm fp
         val tm_abs = Term.Abs("x", Term.type_of tm_subst, body)
         val _ = Output.debug (fn () => "abstraction: " ^ Syntax.string_of_term ctxt tm_abs)
@@ -435,34 +524,34 @@
 
   val factor = Seq.hd o distinct_subgoals_tac;
 
-  fun step ctxt isFO thpairs (fol_th, Metis.Proof.Axiom _)                        =
+  fun step ctxt mode thpairs (fol_th, Metis.Proof.Axiom _)                        =
         factor (axiom_inf ctxt thpairs fol_th)
-    | step ctxt isFO thpairs (_, Metis.Proof.Assume f_atm)                        =
-        assume_inf ctxt f_atm
-    | step ctxt isFO thpairs (_, Metis.Proof.Subst(f_subst, f_th1))               =
-        factor (inst_inf ctxt thpairs f_subst f_th1)
-    | step ctxt isFO thpairs (_, Metis.Proof.Resolve(f_atm, f_th1, f_th2))        =
-        factor (resolve_inf ctxt thpairs f_atm f_th1 f_th2)
-    | step ctxt isFO thpairs (_, Metis.Proof.Refl f_tm)                           =
-        refl_inf ctxt f_tm
-    | step ctxt isFO thpairs (_, Metis.Proof.Equality(f_lit, f_p, f_r)) =
-        equality_inf ctxt isFO thpairs f_lit f_p f_r;
+    | step ctxt mode thpairs (_, Metis.Proof.Assume f_atm)                        =
+        assume_inf ctxt mode f_atm
+    | step ctxt mode thpairs (_, Metis.Proof.Subst(f_subst, f_th1))               =
+        factor (inst_inf ctxt mode thpairs f_subst f_th1)
+    | step ctxt mode thpairs (_, Metis.Proof.Resolve(f_atm, f_th1, f_th2))        =
+        factor (resolve_inf ctxt mode thpairs f_atm f_th1 f_th2)
+    | step ctxt mode thpairs (_, Metis.Proof.Refl f_tm)                           =
+        refl_inf ctxt mode f_tm
+    | step ctxt mode thpairs (_, Metis.Proof.Equality(f_lit, f_p, f_r)) =
+        equality_inf ctxt mode thpairs f_lit f_p f_r;
 
   fun real_literal (b, (c, _)) = not (String.isPrefix ResClause.class_prefix c);
 
-  fun translate isFO _    thpairs [] = thpairs
-    | translate isFO ctxt thpairs ((fol_th, inf) :: infpairs) =
+  fun translate mode _    thpairs [] = thpairs
+    | translate mode ctxt thpairs ((fol_th, inf) :: infpairs) =
         let val _ = Output.debug (fn () => "=============================================")
             val _ = Output.debug (fn () => "METIS THM: " ^ Metis.Thm.toString fol_th)
             val _ = Output.debug (fn () => "INFERENCE: " ^ Metis.Proof.inferenceToString inf)
-            val th = Meson.flexflex_first_order (step ctxt isFO thpairs (fol_th, inf))
+            val th = Meson.flexflex_first_order (step ctxt mode thpairs (fol_th, inf))
             val _ = Output.debug (fn () => "ISABELLE THM: " ^ Display.string_of_thm ctxt th)
             val _ = Output.debug (fn () => "=============================================")
             val n_metis_lits = length (filter real_literal (Metis.LiteralSet.toList (Metis.Thm.clause fol_th)))
         in
             if nprems_of th = n_metis_lits then ()
             else error "Metis: proof reconstruction has gone wrong";
-            translate isFO ctxt ((fol_th, th) :: thpairs) infpairs
+            translate mode ctxt ((fol_th, th) :: thpairs) infpairs
         end;
 
   (*Determining which axiom clauses are actually used*)
@@ -499,8 +588,7 @@
   (* ------------------------------------------------------------------------- *)
 
   type logic_map =
-    {isFO   : bool,
-     axioms : (Metis.Thm.thm * Thm.thm) list,
+    {axioms : (Metis.Thm.thm * Thm.thm) list,
      tfrees : ResClause.type_literal list};
 
   fun const_in_metis c (pol,(pred,tm_list)) =
@@ -515,37 +603,39 @@
     let fun add ((a,i),s) Ts = if i = ~1 then TFree(a,s) :: Ts else Ts
     in  ResClause.add_typs (Vartab.fold add (#2 (Variable.constraints_of ctxt)) []) end;
 
-  (*transform isabelle clause to metis clause *)
-  fun add_thm is_conjecture ctxt (ith, {isFO, axioms, tfrees}) : logic_map =
-    let val (mth, tfree_lits) = hol_thm_to_fol is_conjecture ctxt isFO ith
-    in
-       {isFO = isFO,
-        axioms = (mth, Meson.make_meta_clause ith) :: axioms,
-        tfrees = tfree_lits union tfrees}
-    end;
-
   (*transform isabelle type / arity clause to metis clause *)
   fun add_type_thm [] lmap = lmap
-    | add_type_thm ((ith, mth) :: cls) {isFO, axioms, tfrees} =
-        add_type_thm cls {isFO = isFO,
-                          axioms = (mth, ith) :: axioms,
+    | add_type_thm ((ith, mth) :: cls) {axioms, tfrees} =
+        add_type_thm cls {axioms = (mth, ith) :: axioms,
                           tfrees = tfrees}
 
   (*Insert non-logical axioms corresponding to all accumulated TFrees*)
-  fun add_tfrees {isFO, axioms, tfrees} : logic_map =
-       {isFO = isFO,
-        axioms = (map (fn tf => (metis_of_tfree tf, TrueI)) (distinct op= tfrees)) @ axioms,
+  fun add_tfrees {axioms, tfrees} : logic_map =
+       {axioms = (map (fn tf => (metis_of_tfree tf, TrueI)) (distinct op= tfrees)) @ axioms,
         tfrees = tfrees};
 
+  fun string_of_mode FO = "FO"
+    | string_of_mode HO = "HO"
+    | string_of_mode FT = "FT"
+
   (* Function to generate metis clauses, including comb and type clauses *)
-  fun build_map mode ctxt cls ths =
+  fun build_map mode0 ctxt cls ths =
     let val thy = ProofContext.theory_of ctxt
-        val all_thms_FO = forall (Meson.is_fol_term thy o prop_of)
-        val isFO = (mode = ResAtp.Fol) orelse
-                   (mode <> ResAtp.Hol andalso all_thms_FO (cls @ ths))
-        val lmap0 = List.foldl (add_thm true ctxt)
-                          {isFO = isFO, axioms = [], tfrees = init_tfrees ctxt} cls
-        val lmap = List.foldl (add_thm false ctxt) (add_tfrees lmap0) ths
+        (*The modes FO and FT are sticky. HO can be downgraded to FO.*)
+	fun set_mode FO = FO
+	  | set_mode HO = if forall (Meson.is_fol_term thy o prop_of) (cls@ths) then FO else HO
+	  | set_mode FT = FT
+        val mode = set_mode mode0 
+	(*transform isabelle clause to metis clause *)
+	fun add_thm is_conjecture (ith, {axioms, tfrees}) : logic_map =
+	  let val (mth, tfree_lits) = hol_thm_to_fol is_conjecture ctxt mode ith
+	  in
+	     {axioms = (mth, Meson.make_meta_clause ith) :: axioms,
+	      tfrees = tfree_lits union tfrees}
+	  end;
+        val lmap0 = List.foldl (add_thm true)
+                          {axioms = [], tfrees = init_tfrees ctxt} cls
+        val lmap = List.foldl (add_thm false) (add_tfrees lmap0) ths
         val clause_lists = map (Metis.Thm.clause o #1) (#axioms lmap)
         fun used c = exists (Metis.LiteralSet.exists (const_in_metis c)) clause_lists
         (*Now check for the existence of certain combinators*)
@@ -555,10 +645,10 @@
         val thC   = if used "c_COMBC" then [comb_C] else []
         val thS   = if used "c_COMBS" then [comb_S] else []
         val thEQ  = if used "c_fequal" then [fequal_imp_equal', equal_imp_fequal'] else []
-        val lmap' = if isFO then lmap
-                    else List.foldl (add_thm false ctxt) lmap (thEQ @ thS @ thC @ thB @ thK @ thI)
+        val lmap' = if mode=FO then lmap
+                    else List.foldl (add_thm false) lmap (thEQ @ thS @ thC @ thB @ thK @ thI)
     in
-        add_type_thm (type_ext thy (map prop_of (cls @ ths))) lmap'
+        (mode, add_type_thm (type_ext thy (map prop_of (cls @ ths))) lmap')
     end;
 
   fun refute cls =
@@ -579,16 +669,14 @@
         val _ = app (fn th => Output.debug (fn () => Display.string_of_thm ctxt th)) cls
         val _ = Output.debug (fn () => "THEOREM CLAUSES")
         val _ = app (fn th => Output.debug (fn () => Display.string_of_thm ctxt th)) ths
-        val {isFO,axioms,tfrees} = build_map mode ctxt cls ths
+        val (mode, {axioms,tfrees}) = build_map mode ctxt cls ths
         val _ = if null tfrees then ()
                 else (Output.debug (fn () => "TFREE CLAUSES");
                       app (fn tf => Output.debug (fn _ => ResClause.tptp_of_typeLit true tf)) tfrees)
         val _ = Output.debug (fn () => "CLAUSES GIVEN TO METIS")
         val thms = map #1 axioms
         val _ = app (fn th => Output.debug (fn () => Metis.Thm.toString th)) thms
-        val _ = if isFO
-                then Output.debug (fn () => "goal is first-order")
-                else Output.debug (fn () => "goal is higher-order")
+        val _ = Output.debug (fn () => "mode = " ^ string_of_mode mode)
         val _ = Output.debug (fn () => "START METIS PROVE PROCESS")
     in
         case List.filter (is_false o prop_of) cls of
@@ -601,7 +689,7 @@
                   val ctxt' = fold Variable.declare_constraints (map prop_of cls) ctxt
                                (*add constraints arising from converting goal to clause form*)
                   val proof = Metis.Proof.proof mth
-                  val result = translate isFO ctxt' axioms proof
+                  val result = translate mode ctxt' axioms proof
                   and used = List.mapPartial (used_axioms axioms) proof
                   val _ = Output.debug (fn () => "METIS COMPLETED...clauses actually used:")
 	          val _ = app (fn th => Output.debug (fn () => Display.string_of_thm ctxt th)) used
@@ -634,18 +722,18 @@
     end
     handle METIS s => (warning ("Metis: " ^ s); Seq.empty);
 
-  val metis_tac = metis_general_tac ResAtp.Auto;
-  val metisF_tac = metis_general_tac ResAtp.Fol;
-  val metisH_tac = metis_general_tac ResAtp.Hol;
+  val metis_tac = metis_general_tac HO;
+  val metisF_tac = metis_general_tac FO;
+  val metisFT_tac = metis_general_tac FT;
 
   fun method name mode comment = Method.setup name (Attrib.thms >> (fn ths => fn ctxt =>
     SIMPLE_METHOD' (CHANGED_PROP o metis_general_tac mode ctxt ths))) comment;
 
   val setup =
     type_lits_setup #>
-    method @{binding metis} ResAtp.Auto "METIS for FOL & HOL problems" #>
-    method @{binding metisF} ResAtp.Fol "METIS for FOL problems" #>
-    method @{binding metisH} ResAtp.Hol "METIS for HOL problems" #>
+    method @{binding metis} HO "METIS for FOL & HOL problems" #>
+    method @{binding metisF} FO "METIS for FOL problems" #>
+    method @{binding metisFT} FT "METIS With-fully typed translation" #>
     Method.setup @{binding finish_clausify}
       (Scan.succeed (K (SIMPLE_METHOD (ResAxioms.expand_defs_tac refl))))
       "cleanup after conversion to clauses";
--- a/src/HOL/Tools/res_atp.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Tools/res_atp.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -520,12 +520,17 @@
 fun get_relevant max_new theory_const (ctxt, (chain_ths, th)) goal_cls =
   let
     val thy = ProofContext.theory_of ctxt
+    val isFO = isFO thy goal_cls
     val included_thms = get_clasimp_atp_lemmas ctxt
     val included_cls = included_thms |> ResAxioms.cnf_rules_pairs thy |> make_unique
-                                     |> restrict_to_logic thy (isFO thy goal_cls)
+                                     |> restrict_to_logic thy isFO
                                      |> remove_unwanted_clauses
+    val axcls = relevance_filter max_new theory_const thy included_cls (map prop_of goal_cls)
+    val white_thms = filter check_named (map ResAxioms.pairname
+      (whitelist_fo @ (if isFO then [] else whitelist_ho) @ chain_ths))
+    val white_cls = ResAxioms.cnf_rules_pairs thy white_thms
   in
-    relevance_filter max_new theory_const thy included_cls (map prop_of goal_cls)
+    white_cls @ axcls 
   end;
 
 (* prepare for passing to writer,
@@ -533,11 +538,6 @@
 fun prepare_clauses dfg goal_cls chain_ths axcls extra_cls thy =
   let
     val isFO = isFO thy goal_cls
-    val white_thms = filter check_named (map ResAxioms.pairname
-      (whitelist_fo @ (if isFO then [] else whitelist_ho) @ chain_ths))
-    val white_cls = ResAxioms.cnf_rules_pairs thy white_thms
-    val extra_cls = white_cls @ extra_cls
-    val axcls = white_cls @ axcls
     val ccls = subtract_cls goal_cls extra_cls
     val _ = app (fn th => Output.debug (fn _ => Display.string_of_thm_global thy th)) ccls
     val ccltms = map prop_of ccls
--- a/src/HOL/Tools/res_hol_clause.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Tools/res_hol_clause.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -23,6 +23,7 @@
   datatype literal = Literal of polarity * combterm
   datatype clause = Clause of {clause_id: clause_id, axiom_name: axiom_name, th: thm,
                     kind: ResClause.kind,literals: literal list, ctypes_sorts: typ list}
+  val type_of_combterm: combterm -> ResClause.fol_type
   val strip_comb: combterm -> combterm * combterm list
   val literals_of_term: theory -> term -> literal list * typ list
   exception TOO_TRIVIAL
--- a/src/HOL/Tools/res_reconstruct.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Tools/res_reconstruct.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -17,10 +17,10 @@
   (* extracting lemma list*)
   val find_failure: string -> string option
   val lemma_list: bool -> string ->
-    string * string vector * (int * int) * Proof.context * Thm.thm * int -> string
+    string * string vector * (int * int) * Proof.context * Thm.thm * int -> string * string list
   (* structured proofs *)
   val structured_proof: string ->
-    string * string vector * (int * int) * Proof.context * Thm.thm * int -> string
+    string * string vector * (int * int) * Proof.context * Thm.thm * int -> string * string list
 end;
 
 structure ResReconstruct : RES_RECONSTRUCT =
@@ -554,9 +554,10 @@
 
   fun lemma_list dfg name result =
     let val (lemmas, used_conj) = extract_lemmas (get_step_nums dfg) result
-    in sendback_metis_nochained lemmas ^ "\n" ^ minimize_line name lemmas ^
+    in (sendback_metis_nochained lemmas ^ "\n" ^ minimize_line name lemmas ^
       (if used_conj then ""
-       else "\nWarning: Goal is provable because context is inconsistent.")
+       else "\nWarning: Goal is provable because context is inconsistent."),
+       nochained lemmas)
     end;
 
   (* === Extracting structured Isar-proof === *)
@@ -567,11 +568,11 @@
     val nospaces = String.translate (fn c => if Char.isSpace c then "" else str c)
     val proofextract = get_proof_extract proof
     val cnfs = filter (String.isPrefix "cnf(") (map nospaces (lines proofextract))
-    val one_line_proof = lemma_list false name result
+    val (one_line_proof, lemma_names) = lemma_list false name result
     val structured = if chained_hint mem_string (String.tokens (fn c => c = #" ") one_line_proof) then ""
                 else decode_tstp_file cnfs ctxt goal subgoalno thm_names
     in
-    one_line_proof ^ "\n\n" ^ (Markup.markup Markup.sendback) structured
+    (one_line_proof ^ "\n\n" ^ (Markup.markup Markup.sendback) structured, lemma_names)
   end
 
 end;
--- a/src/HOL/Tools/sat_funcs.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Tools/sat_funcs.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -51,7 +51,7 @@
 	val trace_sat: bool ref    (* input: print trace messages *)
 	val solver: string ref  (* input: name of SAT solver to be used *)
 	val counter: int ref     (* output: number of resolution steps during last proof replay *)
-	val rawsat_thm: cterm list -> thm
+	val rawsat_thm: Proof.context -> cterm list -> thm
 	val rawsat_tac: Proof.context -> int -> tactic
 	val sat_tac: Proof.context -> int -> tactic
 	val satx_tac: Proof.context -> int -> tactic
@@ -295,9 +295,7 @@
 (*      hyps).                                                               *)
 (* ------------------------------------------------------------------------- *)
 
-(* Thm.cterm list -> Thm.thm *)
-
-fun rawsat_thm clauses =
+fun rawsat_thm ctxt clauses =
 let
 	(* remove premises that equal "True" *)
 	val clauses' = filter (fn clause =>
@@ -310,7 +308,7 @@
 		((cnf.is_clause o HOLogic.dest_Trueprop o Thm.term_of) clause
 			handle TERM ("dest_Trueprop", _) => false)
 		orelse (
-			warning ("Ignoring non-clausal premise " ^ Display.string_of_cterm clause);
+			warning ("Ignoring non-clausal premise " ^ Syntax.string_of_term ctxt (Thm.term_of clause));
 			false)) clauses'
 	(* remove trivial clauses -- this is necessary because zChaff removes *)
 	(* trivial clauses during preprocessing, and otherwise our clause     *)
@@ -323,7 +321,8 @@
 	(* sorted in ascending order                                          *)
 	val sorted_clauses = sort (TermOrd.fast_term_ord o pairself Thm.term_of) nontrivial_clauses
 	val _ = if !trace_sat then
-			tracing ("Sorted non-trivial clauses:\n" ^ cat_lines (map Display.string_of_cterm sorted_clauses))
+			tracing ("Sorted non-trivial clauses:\n" ^
+			  cat_lines (map (Syntax.string_of_term ctxt o Thm.term_of) sorted_clauses))
 		else ()
 	(* translate clauses from HOL terms to PropLogic.prop_formula *)
 	val (fms, atom_table) = fold_map (PropLogic.prop_formula_of_term o HOLogic.dest_Trueprop o Thm.term_of) sorted_clauses Termtab.empty
@@ -411,7 +410,8 @@
 (* ------------------------------------------------------------------------- *)
 
 fun rawsat_tac ctxt i =
-  Subgoal.FOCUS (fn {prems, ...} => rtac (rawsat_thm (map cprop_of prems)) 1) ctxt i;
+  Subgoal.FOCUS (fn {context = ctxt', prems, ...} =>
+    rtac (rawsat_thm ctxt' (map cprop_of prems)) 1) ctxt i;
 
 (* ------------------------------------------------------------------------- *)
 (* pre_cnf_tac: converts the i-th subgoal                                    *)
--- a/src/HOL/Tools/transfer_data.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Tools/transfer_data.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -223,7 +223,8 @@
     transf_add
     >> (fn (((((g, inj), embed), ret), cg), hints) => add (inj, embed, ret, cg, g, hints)))
 
-val transferred_att_syntax = (optional names -- Scan.option (keywordC directionN |-- (Args.term -- Args.term)) -- optional (keywordC leavingN |-- names) >> (fn ((hints, aD),leave) => transferred_attribute hints aD leave));
+val transferred_att_syntax = (optional names -- Scan.option (keywordC directionN |-- (Args.term -- Args.term))
+  -- optional (keywordC leavingN |-- names) >> (fn ((hints, aD),leave) => transferred_attribute hints aD leave));
 
 end;
 
--- a/src/HOL/UNITY/Simple/Common.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/UNITY/Simple/Common.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -65,7 +65,7 @@
 lemma "mk_total_program (UNIV, {range(%t.(t, max (ftime t) (gtime t)))}, UNIV)
        \<in> {m} co (maxfg m)"
 apply (simp add: mk_total_program_def) 
-apply (simp add: constrains_def maxfg_def max_def gasc)
+apply (simp add: constrains_def maxfg_def gasc)
 done
 
 (*This one is  t := t+1 if t <max (ftime t) (gtime t) *)
@@ -73,7 +73,7 @@
           (UNIV, { {(t, Suc t) | t. t < max (ftime t) (gtime t)} }, UNIV)   
        \<in> {m} co (maxfg m)"
 apply (simp add: mk_total_program_def) 
-apply (simp add: constrains_def maxfg_def max_def gasc)
+apply (simp add: constrains_def maxfg_def gasc)
 done
 
 
@@ -92,7 +92,7 @@
       ==> F \<in> (atMost n) LeadsTo common"
 apply (rule LeadsTo_weaken_R)
 apply (rule_tac f = id and l = n in GreaterThan_bounded_induct)
-apply (simp_all (no_asm_simp))
+apply (simp_all (no_asm_simp) del: Int_insert_right_if1)
 apply (rule_tac [2] subset_refl)
 apply (blast dest: PSP_Stable2 intro: common_stable LeadsTo_weaken_R)
 done
--- a/src/HOL/Wellfounded.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Wellfounded.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -13,14 +13,6 @@
 
 subsection {* Basic Definitions *}
 
-inductive
-  wfrec_rel :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b => bool"
-  for R :: "('a * 'a) set"
-  and F :: "('a => 'b) => 'a => 'b"
-where
-  wfrecI: "ALL z. (z, x) : R --> wfrec_rel R F z (g z) ==>
-            wfrec_rel R F x (F g x)"
-
 constdefs
   wf         :: "('a * 'a)set => bool"
   "wf(r) == (!P. (!x. (!y. (y,x):r --> P(y)) --> P(x)) --> (!x. P(x)))"
@@ -31,16 +23,6 @@
   acyclic :: "('a*'a)set => bool"
   "acyclic r == !x. (x,x) ~: r^+"
 
-  cut        :: "('a => 'b) => ('a * 'a)set => 'a => 'a => 'b"
-  "cut f r x == (%y. if (y,x):r then f y else undefined)"
-
-  adm_wf :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => bool"
-  "adm_wf R F == ALL f g x.
-     (ALL z. (z, x) : R --> f z = g z) --> F f x = F g x"
-
-  wfrec :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b"
-  [code del]: "wfrec R F == %x. THE y. wfrec_rel R (%f x. F (cut f R x) x) x y"
-
 abbreviation acyclicP :: "('a => 'a => bool) => bool" where
   "acyclicP r == acyclic {(x, y). r x y}"
 
@@ -425,54 +407,6 @@
 by (blast intro: finite_acyclic_wf wf_acyclic)
 
 
-subsection{*Well-Founded Recursion*}
-
-text{*cut*}
-
-lemma cuts_eq: "(cut f r x = cut g r x) = (ALL y. (y,x):r --> f(y)=g(y))"
-by (simp add: expand_fun_eq cut_def)
-
-lemma cut_apply: "(x,a):r ==> (cut f r a)(x) = f(x)"
-by (simp add: cut_def)
-
-text{*Inductive characterization of wfrec combinator; for details see:  
-John Harrison, "Inductive definitions: automation and application"*}
-
-lemma wfrec_unique: "[| adm_wf R F; wf R |] ==> EX! y. wfrec_rel R F x y"
-apply (simp add: adm_wf_def)
-apply (erule_tac a=x in wf_induct) 
-apply (rule ex1I)
-apply (rule_tac g = "%x. THE y. wfrec_rel R F x y" in wfrec_rel.wfrecI)
-apply (fast dest!: theI')
-apply (erule wfrec_rel.cases, simp)
-apply (erule allE, erule allE, erule allE, erule mp)
-apply (fast intro: the_equality [symmetric])
-done
-
-lemma adm_lemma: "adm_wf R (%f x. F (cut f R x) x)"
-apply (simp add: adm_wf_def)
-apply (intro strip)
-apply (rule cuts_eq [THEN iffD2, THEN subst], assumption)
-apply (rule refl)
-done
-
-lemma wfrec: "wf(r) ==> wfrec r H a = H (cut (wfrec r H) r a) a"
-apply (simp add: wfrec_def)
-apply (rule adm_lemma [THEN wfrec_unique, THEN the1_equality], assumption)
-apply (rule wfrec_rel.wfrecI)
-apply (intro strip)
-apply (erule adm_lemma [THEN wfrec_unique, THEN theI'])
-done
-
-subsection {* Code generator setup *}
-
-consts_code
-  "wfrec"   ("\<module>wfrec?")
-attach {*
-fun wfrec f x = f (wfrec f) x;
-*}
-
-
 subsection {* @{typ nat} is well-founded *}
 
 lemma less_nat_rel: "op < = (\<lambda>m n. n = Suc m)^++"
@@ -696,9 +630,6 @@
 apply blast
 done
 
-lemma in_inv_image[simp]: "((x,y) : inv_image r f) = ((f x, f y) : r)"
-  by (auto simp:inv_image_def)
-
 text {* Measure Datatypes into @{typ nat} *}
 
 definition measure :: "('a => nat) => ('a * 'a)set"
--- a/src/HOL/Word/BinBoolList.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Word/BinBoolList.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -918,8 +918,8 @@
   apply (frule asm_rl)
   apply (drule spec)
   apply (erule trans)
-  apply (drule_tac x = "bin_cat y n a" in spec) 
-  apply (simp add : bin_cat_assoc_sym min_def)
+  apply (drule_tac x = "bin_cat y n a" in spec)
+  apply (simp add : bin_cat_assoc_sym)
   done
 
 lemma bin_rcat_bl:
--- a/src/HOL/Word/BinGeneral.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Word/BinGeneral.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -493,7 +493,7 @@
 
 lemma sbintrunc_sbintrunc_l:
   "n <= m ==> (sbintrunc m (sbintrunc n w) = sbintrunc n w)"
-  by (rule bin_eqI) (auto simp: nth_sbintr min_def)
+  by (rule bin_eqI) (auto simp: nth_sbintr)
 
 lemma bintrunc_bintrunc_ge:
   "n <= m ==> (bintrunc n (bintrunc m w) = bintrunc n w)"
@@ -501,14 +501,12 @@
 
 lemma bintrunc_bintrunc_min [simp]:
   "bintrunc m (bintrunc n w) = bintrunc (min m n) w"
-  apply (unfold min_def)
   apply (rule bin_eqI)
   apply (auto simp: nth_bintr)
   done
 
 lemma sbintrunc_sbintrunc_min [simp]:
   "sbintrunc m (sbintrunc n w) = sbintrunc (min m n) w"
-  apply (unfold min_def)
   apply (rule bin_eqI)
   apply (auto simp: nth_sbintr)
   done
--- a/src/HOL/Word/WordDefinition.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Word/WordDefinition.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -380,7 +380,7 @@
   "n >= size w ==> bintrunc n (uint w) = uint w"
   apply (unfold word_size)
   apply (subst word_ubin.norm_Rep [symmetric]) 
-  apply (simp only: bintrunc_bintrunc_min word_size min_def)
+  apply (simp only: bintrunc_bintrunc_min word_size)
   apply simp
   done
 
@@ -388,7 +388,7 @@
   "wb = word_of_int bin ==> n >= size wb ==> 
     word_of_int (bintrunc n bin) = wb"
   unfolding word_size
-  by (clarsimp simp add : word_ubin.norm_eq_iff [symmetric] min_def)
+  by (clarsimp simp add : word_ubin.norm_eq_iff [symmetric])
 
 lemmas bintr_uint = bintr_uint' [unfolded word_size]
 lemmas wi_bintr = wi_bintr' [unfolded word_size]
--- a/src/HOL/Word/WordShift.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/Word/WordShift.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -319,7 +319,7 @@
 
 lemma bl_shiftl:
   "to_bl (w << n) = drop n (to_bl w) @ replicate (min (size w) n) False"
-  by (simp add: shiftl_bl word_rep_drop word_size min_def)
+  by (simp add: shiftl_bl word_rep_drop word_size)
 
 lemma shiftl_zero_size: 
   fixes x :: "'a::len0 word"
@@ -1018,8 +1018,7 @@
   word_of_int (bin_cat w (size b) (uint b))"
   apply (unfold word_cat_def word_size) 
   apply (clarsimp simp add : word_ubin.norm_eq_iff [symmetric]
-      word_ubin.eq_norm bintr_cat min_def)
-  apply arith
+      word_ubin.eq_norm bintr_cat)
   done
 
 lemma word_cat_split_alt:
--- a/src/HOL/ex/Codegenerator_Candidates.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/ex/Codegenerator_Candidates.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -16,7 +16,7 @@
   Nested_Environment
   Option_ord
   Permutation
-  Primes
+  "~~/src/HOL/Number_Theory/Primes"
   Product_ord
   SetsAndFunctions
   Tree
--- a/src/HOL/ex/Mirabelle.thy	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,14 +0,0 @@
-(* Title: Mirabelle.thy
-   Author: Jasmin Blanchette and Sascha Boehme
-*)
-
-theory Mirabelle
-imports Main
-uses "mirabelle.ML"
-begin
-
-(* FIXME: use a logfile for each theory file *)
-
-setup Mirabelle.setup
-
-end
--- a/src/HOL/ex/NormalForm.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/ex/NormalForm.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -120,4 +120,12 @@
 normal_form "(%m n f x. m (n f) x) (%f x. f(f(f(x)))) (%f x. f(f(f(x))))"
 normal_form "(%m n. n m) (%f x. f(f(f(x)))) (%f x. f(f(f(x))))"
 
+(* handling of type classes in connection with equality *)
+
+lemma "map f [x, y] = [f x, f y]" by normalization
+lemma "(map f [x, y], w) = ([f x, f y], w)" by normalization
+lemma "map f [x, y] = [f x \<Colon> 'a\<Colon>semigroup_add, f y]" by normalization
+lemma "map f [x \<Colon> 'a\<Colon>semigroup_add, y] = [f x, f y]" by normalization
+lemma "(map f [x \<Colon> 'a\<Colon>semigroup_add, y], w \<Colon> 'b\<Colon>finite) = ([f x, f y], w)" by normalization
+
 end
--- a/src/HOL/ex/Predicate_Compile_ex.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/ex/Predicate_Compile_ex.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -99,4 +99,62 @@
 values 20 "{(n, m). tranclp succ n m}"
 *)
 
+subsection{* IMP *}
+
+types
+  var = nat
+  state = "int list"
+
+datatype com =
+  Skip |
+  Ass var "state => int" |
+  Seq com com |
+  IF "state => bool" com com |
+  While "state => bool" com
+
+inductive exec :: "com => state => state => bool" where
+"exec Skip s s" |
+"exec (Ass x e) s (s[x := e(s)])" |
+"exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" |
+"b s ==> exec c1 s t ==> exec (IF b c1 c2) s t" |
+"~b s ==> exec c2 s t ==> exec (IF b c1 c2) s t" |
+"~b s ==> exec (While b c) s s" |
+"b s1 ==> exec c s1 s2 ==> exec (While b c) s2 s3 ==> exec (While b c) s1 s3"
+
+code_pred exec .
+
+values "{t. exec
+ (While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1))))
+ [3,5] t}"
+
+
+subsection{* CCS *}
+
+text{* This example formalizes finite CCS processes without communication or
+recursion. For simplicity, labels are natural numbers. *}
+
+datatype proc = nil | pre nat proc | or proc proc | par proc proc
+
+inductive step :: "proc \<Rightarrow> nat \<Rightarrow> proc \<Rightarrow> bool" where
+"step (pre n p) n p" |
+"step p1 a q \<Longrightarrow> step (or p1 p2) a q" |
+"step p2 a q \<Longrightarrow> step (or p1 p2) a q" |
+"step p1 a q \<Longrightarrow> step (par p1 p2) a (par q p2)" |
+"step p2 a q \<Longrightarrow> step (par p1 p2) a (par p1 q)"
+
+code_pred step .
+
+inductive steps where
+"steps p [] p" |
+"step p a q \<Longrightarrow> steps q as r \<Longrightarrow> steps p (a#as) r"
+
+code_pred steps .
+
+values 5
+ "{as . steps (par (or (pre 0 nil) (pre 1 nil)) (pre 2 nil)) as (par nil nil)}"
+
+(* FIXME
+values 3 "{(a,q). step (par nil nil) a q}"
+*)
+
 end
\ No newline at end of file
--- a/src/HOL/ex/ROOT.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/ex/ROOT.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -12,7 +12,6 @@
   "Codegenerator_Test",
   "Codegenerator_Pretty_Test",
   "NormalForm",
-  "../NumberTheory/Factorization",
   "Predicate_Compile",
   "Predicate_Compile_ex"
 ];
@@ -72,25 +71,16 @@
 (setmp_noncritical proofs 2 (setmp_noncritical Multithreading.max_threads 1 use_thy))
   "Hilbert_Classical";
 
-
 use_thy "SVC_Oracle";
-
-fun svc_enabled () = getenv "SVC_HOME" <> "";
-fun if_svc_enabled f x = if svc_enabled () then f x else ();
-
-if_svc_enabled use_thy "svc_test";
-
+if getenv "SVC_HOME" = "" then ()
+else use_thy "svc_test";
 
-(* requires a proof-generating SAT solver (zChaff or MiniSAT) to be *)
-(* installed:                                                       *)
+(*requires a proof-generating SAT solver (zChaff or MiniSAT)*)
 try use_thy "SAT_Examples";
-Future.shutdown ();
 
-(* requires zChaff (or some other reasonably fast SAT solver) to be *)
-(* installed:                                                       *)
-if getenv "ZCHAFF_HOME" <> "" then
-  use_thy "Sudoku"
-else ();
+(*requires zChaff (or some other reasonably fast SAT solver)*)
+if getenv "ZCHAFF_HOME" = "" then ()
+else use_thy "Sudoku";
 
 HTML.with_charset "utf-8" (no_document use_thys)
   ["Hebrew", "Chinese", "Serbian"];
--- a/src/HOL/ex/Sqrt.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/ex/Sqrt.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -5,7 +5,7 @@
 header {*  Square roots of primes are irrational *}
 
 theory Sqrt
-imports Complex_Main
+imports Complex_Main "~~/src/HOL/Number_Theory/Primes"
 begin
 
 text {*
--- a/src/HOL/ex/Sqrt_Script.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/ex/Sqrt_Script.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -6,7 +6,7 @@
 header {* Square roots of primes are irrational (script version) *}
 
 theory Sqrt_Script
-imports Complex_Main Primes
+imports Complex_Main "~~/src/HOL/Number_Theory/Primes"
 begin
 
 text {*
@@ -16,30 +16,30 @@
 
 subsection {* Preliminaries *}
 
-lemma prime_nonzero:  "prime p \<Longrightarrow> p \<noteq> 0"
-  by (force simp add: prime_def)
+lemma prime_nonzero:  "prime (p::nat) \<Longrightarrow> p \<noteq> 0"
+  by (force simp add: prime_nat_def)
 
 lemma prime_dvd_other_side:
-    "n * n = p * (k * k) \<Longrightarrow> prime p \<Longrightarrow> p dvd n"
-  apply (subgoal_tac "p dvd n * n", blast dest: prime_dvd_mult)
+    "(n::nat) * n = p * (k * k) \<Longrightarrow> prime p \<Longrightarrow> p dvd n"
+  apply (subgoal_tac "p dvd n * n", blast dest: prime_dvd_mult_nat)
   apply auto
   done
 
-lemma reduction: "prime p \<Longrightarrow>
+lemma reduction: "prime (p::nat) \<Longrightarrow>
     0 < k \<Longrightarrow> k * k = p * (j * j) \<Longrightarrow> k < p * j \<and> 0 < j"
   apply (rule ccontr)
   apply (simp add: linorder_not_less)
   apply (erule disjE)
    apply (frule mult_le_mono, assumption)
    apply auto
-  apply (force simp add: prime_def)
+  apply (force simp add: prime_nat_def)
   done
 
 lemma rearrange: "(j::nat) * (p * j) = k * k \<Longrightarrow> k * k = p * (j * j)"
   by (simp add: mult_ac)
 
 lemma prime_not_square:
-    "prime p \<Longrightarrow> (\<And>k. 0 < k \<Longrightarrow> m * m \<noteq> p * (k * k))"
+    "prime (p::nat) \<Longrightarrow> (\<And>k. 0 < k \<Longrightarrow> m * m \<noteq> p * (k * k))"
   apply (induct m rule: nat_less_induct)
   apply clarify
   apply (frule prime_dvd_other_side, assumption)
@@ -57,7 +57,7 @@
 *}
 
 theorem prime_sqrt_irrational:
-    "prime p \<Longrightarrow> x * x = real p \<Longrightarrow> 0 \<le> x \<Longrightarrow> x \<notin> \<rat>"
+    "prime (p::nat) \<Longrightarrow> x * x = real p \<Longrightarrow> 0 \<le> x \<Longrightarrow> x \<notin> \<rat>"
   apply (rule notI)
   apply (erule Rats_abs_nat_div_natE)
   apply (simp del: real_of_nat_mult
@@ -65,6 +65,6 @@
   done
 
 lemmas two_sqrt_irrational =
-  prime_sqrt_irrational [OF two_is_prime]
+  prime_sqrt_irrational [OF two_is_prime_nat]
 
 end
--- a/src/HOL/ex/Termination.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/HOL/ex/Termination.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -1,5 +1,4 @@
 (* Title:       HOL/ex/Termination.thy
-   ID:          $Id$
    Author:      Lukas Bulwahn, TU Muenchen
    Author:      Alexander Krauss, TU Muenchen
 *)
@@ -10,12 +9,33 @@
 imports Main Multiset
 begin
 
+subsection {* Manually giving termination relations using @{text relation} and
+@{term measure} *}
+
+function sum :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+where
+  "sum i N = (if i > N then 0 else i + sum (Suc i) N)"
+by pat_completeness auto
+
+termination by (relation "measure (\<lambda>(i,N). N + 1 - i)") auto
+
+function foo :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+where
+  "foo i N = (if i > N 
+              then (if N = 0 then 0 else foo 0 (N - 1))
+              else i + foo (Suc i) N)"
+by pat_completeness auto
+
+termination by (relation "measures [\<lambda>(i, N). N, \<lambda>(i,N). N + 1 - i]") auto
+
+
+subsection {* @{text lexicographic_order}: Trivial examples *}
+
 text {*
-  The @{text fun} command uses the method @{text lexicographic_order} by default.
+  The @{text fun} command uses the method @{text lexicographic_order} by default,
+  so it is not explicitly invoked.
 *}
 
-subsection {* Trivial examples *}
-
 fun identity :: "nat \<Rightarrow> nat"
 where
   "identity n = n"
--- a/src/HOL/ex/mirabelle.ML	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,318 +0,0 @@
-(* Title:  mirabelle.ML
-   Author: Jasmin Blanchette and Sascha Boehme
-*)
-
-signature MIRABELLE =
-sig
-  type action
-  type settings
-  val register : string -> action -> theory -> theory
-  val invoke : string -> settings -> theory -> theory
-
-  val timeout : int Config.T
-  val verbose : bool Config.T
-  val set_logfile : string -> theory -> theory
-
-  val setup : theory -> theory
-
-  val step_hook : Toplevel.transition -> Toplevel.state -> Toplevel.state ->
-    unit
-
-  val goal_thm_of : Proof.state -> thm
-  val can_apply : (Proof.context -> int -> tactic) -> Proof.state -> bool
-  val theorems_in_proof_term : Thm.thm -> Thm.thm list
-  val theorems_of_sucessful_proof : Toplevel.state -> Thm.thm list
-  val get_setting : settings -> string * string -> string
-  val get_int_setting : settings -> string * int -> int
-
-(* FIXME  val refute_action : action *)
-  val quickcheck_action : action
-  val arith_action : action
-  val sledgehammer_action : action
-  val metis_action : action
-end
-
-
-
-structure Mirabelle (*: MIRABELLE*) =
-struct
-
-(* Mirabelle core *)
-
-type settings = (string * string) list
-type invoked = {pre: Proof.state, post: Toplevel.state option} -> string option
-type action = settings -> invoked
-
-structure Registered = TheoryDataFun
-(
-  type T = action Symtab.table
-  val empty = Symtab.empty
-  val copy = I
-  val extend = I
-  fun merge _ = Symtab.merge (K true)
-)
-
-fun register name act = Registered.map (Symtab.update_new (name, act))
-
-
-structure Invoked = TheoryDataFun
-(
-  type T = (string * invoked) list
-  val empty = []
-  val copy = I
-  val extend = I
-  fun merge _ = Library.merge (K true)
-)
-
-fun invoke name sts thy = 
-  let 
-    val act = 
-      (case Symtab.lookup (Registered.get thy) name of
-        SOME act => act
-      | NONE => error ("The invoked action " ^ quote name ^ 
-          " is not registered."))
-  in Invoked.map (cons (name, act sts)) thy end
-
-val (logfile, setup1) = Attrib.config_string "mirabelle_logfile" ""
-val (timeout, setup2) = Attrib.config_int "mirabelle_timeout" 30
-val (verbose, setup3) = Attrib.config_bool "mirabelle_verbose" true
-val (start_line, setup4) = Attrib.config_int "mirabelle_start_line" 0
-val (end_line, setup5) = Attrib.config_int "mirabelle_end_line" ~1
-
-val setup_config = setup1 #> setup2 #> setup3 #> setup4 #> setup5
-
-fun set_logfile name =
-  let val _ = File.write (Path.explode name) ""   (* erase file content *)
-  in Config.put_thy logfile name end
-
-local
-
-fun log thy s =
-  let fun append_to n = if n = "" then K () else File.append (Path.explode n)
-  in append_to (Config.get_thy thy logfile) (s ^ "\n") end
-  (* FIXME: with multithreading and parallel proofs enabled, we might need to
-     encapsulate this inside a critical section *)
-
-fun verbose_msg verbose msg = if verbose then SOME msg else NONE
-
-fun with_time_limit (verb, secs) f x = TimeLimit.timeLimit secs f x
-  handle TimeLimit.TimeOut => verbose_msg verb "time out"
-       | ERROR msg => verbose_msg verb ("error: " ^ msg)
-
-fun capture_exns verb f x =
-  (case try f x of NONE => verbose_msg verb "exception" | SOME msg => msg)
-
-fun apply_action (c as (verb, _)) st (name, invoked) =
-  Option.map (pair name) (capture_exns verb (with_time_limit c invoked) st)
-
-fun in_range _ _ NONE = true
-  | in_range l r (SOME i) = (l <= i andalso (r < 0 orelse i <= r))
-
-fun only_within_range thy pos f x =
-  let val l = Config.get_thy thy start_line and r = Config.get_thy thy end_line
-  in if in_range l r (Position.line_of pos) then f x else [] end
-
-fun pretty_print verbose pos name msgs =
-  let
-    val file = the_default "unknown file" (Position.file_of pos)
-
-    val str0 = string_of_int o the_default 0
-    val loc = str0 (Position.line_of pos) ^ ":" ^ str0 (Position.column_of pos)
-
-    val full_loc = if verbose then file ^ ":" ^ loc else "at " ^ loc
-    val head = full_loc ^ " (" ^ name ^ "):"
-
-    fun pretty_msg (name, msg) = Pretty.block (map Pretty.str [name, ": ", msg])
-  in
-    Pretty.string_of (Pretty.big_list head (map pretty_msg msgs))
-  end
-
-in
-
-fun basic_hook tr pre post =
-  let
-    val thy = Proof.theory_of pre
-    val pos = Toplevel.pos_of tr
-    val name = Toplevel.name_of tr
-    val verb = Config.get_thy thy verbose
-    val secs = Time.fromSeconds (Config.get_thy thy timeout)
-    val st = {pre=pre, post=post}
-  in
-    Invoked.get thy
-    |> only_within_range thy pos (map_filter (apply_action (verb, secs) st))
-    |> (fn [] => () | msgs => log thy (pretty_print verb pos name msgs))
-  end
-
-end
-
-fun step_hook tr pre post =
- (* FIXME: might require wrapping into "interruptible" *)
-  if can (Proof.assert_backward o Toplevel.proof_of) pre andalso
-     not (member (op =) ["disable_pr", "enable_pr"] (Toplevel.name_of tr))
-  then basic_hook tr (Toplevel.proof_of pre) (SOME post)
-  else ()   (* FIXME: add theory_hook here *)
-
-
-
-(* Mirabelle utility functions *)
-
-val goal_thm_of = snd o snd o Proof.get_goal
-
-fun can_apply tac st =
-  let val (ctxt, (facts, goal)) = Proof.get_goal st
-  in
-    (case Seq.pull (HEADGOAL (Method.insert_tac facts THEN' tac ctxt) goal) of
-      SOME (thm, _) => true
-    | NONE => false)
-  end
-
-local
-
-fun fold_body_thms f =
-  let
-    fun app n (PBody {thms, ...}) = thms |> fold (fn (i, (name, prop, body)) =>
-      fn (x, seen) =>
-        if Inttab.defined seen i then (x, seen)
-        else
-          let
-            val body' = Future.join body
-            val (x', seen') = app (n + (if name = "" then 0 else 1)) body'
-              (x, Inttab.update (i, ()) seen)
-        in (x' |> n = 0 ? f (name, prop, body'), seen') end)
-  in fn bodies => fn x => #1 (fold (app 0) bodies (x, Inttab.empty)) end
-
-in
-
-fun theorems_in_proof_term thm =
-  let
-    val all_thms = PureThy.all_thms_of (Thm.theory_of_thm thm)
-    fun collect (s, _, _) = if s <> "" then insert (op =) s else I
-    fun member_of xs (x, y) = if member (op =) xs x then SOME y else NONE
-    fun resolve_thms names = map_filter (member_of names) all_thms
-  in
-    resolve_thms (fold_body_thms collect [Thm.proof_body_of thm] [])
-  end
-
-end
-
-fun theorems_of_sucessful_proof state =
-  (case state of
-    NONE => []
-  | SOME st =>
-      if not (Toplevel.is_proof st) then []
-      else theorems_in_proof_term (goal_thm_of (Toplevel.proof_of st)))
-
-fun get_setting settings (key, default) =
-  the_default default (AList.lookup (op =) settings key)
-
-fun get_int_setting settings (key, default) =
-  (case Option.map Int.fromString (AList.lookup (op =) settings key) of
-    SOME (SOME i) => i
-  | SOME NONE => error ("bad option: " ^ key)
-  | NONE => default)
-
-
-
-(* Mirabelle actions *)
-
-(* FIXME
-fun refute_action settings {pre=st, ...} = 
-  let
-    val params   = [("minsize", "2") (*"maxsize", "2"*)]
-    val subgoal = 0
-    val thy     = Proof.theory_of st
-    val thm = goal_thm_of st
-
-    val _ = Refute.refute_subgoal thy parms thm subgoal
-  in
-    val writ_log = Substring.full (the (Symtab.lookup tab "writeln"))
-    val warn_log = Substring.full (the (Symtab.lookup tab "warning"))
-
-    val r =
-      if Substring.isSubstring "model found" writ_log
-      then
-        if Substring.isSubstring "spurious" warn_log
-        then SOME "potential counterexample"
-        else SOME "real counterexample (bug?)"
-      else
-        if Substring.isSubstring "time limit" writ_log
-        then SOME "no counterexample (time out)"
-        else if Substring.isSubstring "Search terminated" writ_log
-        then SOME "no counterexample (normal termination)"
-        else SOME "no counterexample (unknown)"
-  in r end
-*)
-
-fun quickcheck_action settings {pre=st, ...} =
-  let
-    val has_valid_key = member (op =) ["iterations", "size", "generator"] o fst
-    val args = filter has_valid_key settings
-  in
-    (case Quickcheck.quickcheck args 1 st of
-      NONE => SOME "no counterexample"
-    | SOME _ => SOME "counterexample found")
-  end
-
-
-fun arith_action _ {pre=st, ...} = 
-  if can_apply Arith_Data.arith_tac st
-  then SOME "succeeded"
-  else NONE
-
-
-fun sledgehammer_action settings {pre=st, ...} =
-  let
-    val prover_name = hd (space_explode " " (AtpManager.get_atps ()))
-    val thy = Proof.theory_of st
- 
-    val prover = the (AtpManager.get_prover prover_name thy)
-    val timeout = AtpManager.get_timeout () 
-
-    val (success, message) =
-      let
-        val (success, message, _, _, _) =
-          prover timeout NONE NONE prover_name 1 (Proof.get_goal st)
-      in (success, message) end
-      handle ResHolClause.TOO_TRIVIAL => (true, "trivial")
-           | ERROR msg => (false, "error: " ^ msg)
-  in
-    if success
-    then SOME ("success (" ^ prover_name ^ ": " ^ message ^ ")")
-    else NONE
-  end
-
-
-fun metis_action settings {pre, post} =
-  let
-    val thms = theorems_of_sucessful_proof post
-    val names = map Thm.get_name thms
-
-    val facts = Facts.props (ProofContext.facts_of (Proof.context_of pre))
-
-    fun metis ctxt = MetisTools.metis_tac ctxt (thms @ facts)
-  in
-    (if can_apply metis pre then "succeeded" else "failed")
-    |> suffix (" (" ^ commas names ^ ")")
-    |> SOME
-  end
-
-
-
-(* Mirabelle setup *)
-
-val setup =
-  setup_config #>
-(* FIXME  register "refute" refute_action #> *)
-  register "quickcheck" quickcheck_action #>
-  register "arith" arith_action #>
-  register "sledgehammer" sledgehammer_action #>
-  register "metis" metis_action (* #> FIXME:
-  Context.theory_map (Specification.add_theorem_hook theorem_hook) *)
-
-end
-
-val _ = Toplevel.add_hook Mirabelle.step_hook
-
-(* no multithreading, no parallel proofs *)
-val _ = Multithreading.max_threads := 1
-val _ = Goal.parallel_proofs := 0
--- a/src/Pure/Concurrent/future.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/Concurrent/future.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -237,7 +237,7 @@
               val total = length (! workers);
               val active = count_active ();
             in
-              "SCHEDULE: " ^
+              "SCHEDULE " ^ string_of_int (Time.toMilliseconds now) ^ ": " ^
                 string_of_int ready ^ " ready, " ^
                 string_of_int pending ^ " pending, " ^
                 string_of_int running ^ " running; " ^
--- a/src/Pure/General/alist.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/General/alist.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -122,6 +122,6 @@
   in coal end;
 
 fun group eq xs =
-  fold_rev (fn (k, v) => default eq (k, []) #> map_entry eq k (cons v)) xs [];
+  fold_rev (fn (k, v) => map_default eq (k, []) (cons v)) xs [];
 
 end;
--- a/src/Pure/General/event_bus.scala	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/General/event_bus.scala	Fri Sep 11 15:57:16 2009 +1000
@@ -1,38 +1,35 @@
 /*  Title:      Pure/General/event_bus.scala
     Author:     Makarius
 
-Generic event bus with multiple handlers and optional exception
-logging.
+Generic event bus with multiple receiving actors.
 */
 
 package isabelle
 
+import scala.actors.Actor, Actor._
 import scala.collection.mutable.ListBuffer
 
 
-class EventBus[Event]
+class Event_Bus[Event]
 {
-  /* event handlers */
+  /* receivers */
+
+  private val receivers = new ListBuffer[Actor]
+
+  def += (r: Actor) { synchronized { receivers += r } }
+  def + (r: Actor): Event_Bus[Event] = { this += r; this }
 
-  type Handler = Event => Unit
-  private val handlers = new ListBuffer[Handler]
+  def += (f: Event => Unit) {
+    this += actor { loop { react { case x: Event => f(x) } } }
+  }
 
-  def += (h: Handler) = synchronized { handlers += h }
-  def + (h: Handler) = { this += h; this }
+  def + (f: Event => Unit): Event_Bus[Event] = { this += f; this }
 
-  def -= (h: Handler) = synchronized { handlers -= h }
-  def - (h: Handler) = { this -= h; this }
+  def -= (r: Actor) { synchronized { receivers -= r } }
+  def - (r: Actor) = { this -= r; this }
 
 
   /* event invocation */
 
-  var logger: Throwable => Unit = throw _
-
-  def event(x: Event) = {
-    val log = logger
-    for (h <- synchronized { handlers.toList }) {
-      try { h(x) }
-      catch { case e: Throwable => log(e) }
-    }
-  }
+  def event(x: Event) { synchronized { receivers.foreach(_ ! x) } }
 }
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/General/linear_set.scala	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,118 @@
+/*  Title:      Pure/General/linear_set.scala
+    Author:     Makarius
+    Author:     Fabian Immler TU Munich
+
+Sets with canonical linear order, or immutable linked-lists.
+*/
+package isabelle
+
+
+object Linear_Set
+{
+  private case class Rep[A](
+    val first_elem: Option[A], val last_elem: Option[A], val body: Map[A, A])
+
+  private def empty_rep[A] = Rep[A](None, None, Map())
+
+  private def make[A](first_elem: Option[A], last_elem: Option[A], body: Map[A, A]): Linear_Set[A] =
+    new Linear_Set[A] { override val rep = Rep(first_elem, last_elem, body) }
+
+
+  def empty[A]: Linear_Set[A] = new Linear_Set
+  def apply[A](elems: A*): Linear_Set[A] = empty[A] ++ elems
+
+  class Duplicate(s: String) extends Exception(s)
+  class Undefined(s: String) extends Exception(s)
+}
+
+
+class Linear_Set[A] extends scala.collection.immutable.Set[A]
+{
+  /* representation */
+
+  protected val rep = Linear_Set.empty_rep[A]
+
+
+  /* auxiliary operations */
+
+  def next(elem: A) = rep.body.get(elem)
+  def prev(elem: A) = rep.body.find(_._2 == elem).map(_._1)   // slow
+
+  private def _insert_after(hook: Option[A], elem: A): Linear_Set[A] =
+    if (contains(elem)) throw new Linear_Set.Duplicate(elem.toString)
+    else hook match {
+      case Some(hook) =>
+        if (!contains(hook)) throw new Linear_Set.Undefined(hook.toString)
+        else if (rep.body.isDefinedAt(hook))
+          Linear_Set.make(rep.first_elem, rep.last_elem,
+            rep.body - hook + (hook -> elem) + (elem -> rep.body(hook)))
+        else Linear_Set.make(rep.first_elem, Some(elem), rep.body + (hook -> elem))
+      case None =>
+        if (isEmpty) Linear_Set.make(Some(elem), Some(elem), Map())
+        else Linear_Set.make(Some(elem), rep.last_elem, rep.body + (elem -> rep.first_elem.get))
+    }
+
+  def insert_after(hook: Option[A], elems: Seq[A]): Linear_Set[A] =
+    elems.reverse.foldLeft (this) ((ls, elem) => ls._insert_after(hook, elem))
+
+  def delete_after(elem: Option[A]): Linear_Set[A] =
+    elem match {
+      case Some(elem) =>
+        if (!contains(elem)) throw new Linear_Set.Undefined(elem.toString)
+        else if (!rep.body.isDefinedAt(elem)) throw new Linear_Set.Undefined(null)
+        else if (rep.body(elem) == rep.last_elem.get)
+          Linear_Set.make(rep.first_elem, Some(elem), rep.body - elem)
+        else Linear_Set.make(rep.first_elem, rep.last_elem,
+          rep.body - elem - rep.body(elem) + (elem -> rep.body(rep.body(elem))))
+      case None =>
+        if (isEmpty) throw new Linear_Set.Undefined(null)
+        else if (size == 1) empty
+        else Linear_Set.make(Some(rep.body(rep.first_elem.get)), rep.last_elem, rep.body - rep.first_elem.get)
+    }
+
+  def delete_between(from: Option[A], to: Option[A]): Linear_Set[A] = {
+    if(!rep.first_elem.isDefined) this
+    else {
+      val next =
+        if (from == rep.last_elem) None
+        else if (from == None) rep.first_elem
+        else from.map(rep.body(_))
+      if (next == to) this
+      else delete_after(from).delete_between(from, to)
+    }
+  }
+
+
+  /* Set methods */
+
+  override def stringPrefix = "Linear_Set"
+
+  def empty[B]: Linear_Set[B] = Linear_Set.empty
+
+  override def isEmpty: Boolean = !rep.last_elem.isDefined
+  def size: Int = if (isEmpty) 0 else rep.body.size + 1
+
+  def elements = new Iterator[A] {
+    private var next_elem = rep.first_elem
+    def hasNext = next_elem.isDefined
+    def next = {
+      val elem = next_elem.get
+      next_elem = if (rep.body.isDefinedAt(elem)) Some(rep.body(elem)) else None
+      elem
+    }
+  }
+
+  def contains(elem: A): Boolean =
+    !isEmpty && (rep.last_elem.get == elem || rep.body.isDefinedAt(elem))
+
+  def + (elem: A): Linear_Set[A] = _insert_after(rep.last_elem, elem)
+
+  override def + (elem1: A, elem2: A, elems: A*): Linear_Set[A] =
+    this + elem1 + elem2 ++ elems
+  override def ++ (elems: Iterable[A]): Linear_Set[A] = (this /: elems) ((s, elem) => s + elem)
+  override def ++ (elems: Iterator[A]): Linear_Set[A] = (this /: elems) ((s, elem) => s + elem)
+
+  def - (elem: A): Linear_Set[A] =
+    if (!contains(elem)) throw new Linear_Set.Undefined(elem.toString)
+    else delete_after(prev(elem))
+}
\ No newline at end of file
--- a/src/Pure/General/markup.scala	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/General/markup.scala	Fri Sep 11 15:57:16 2009 +1000
@@ -6,8 +6,9 @@
 
 package isabelle
 
-object Markup {
 
+object Markup
+{
   /* name */
 
   val NAME = "name"
@@ -25,7 +26,8 @@
   val FILE = "file"
   val ID = "id"
 
-  val POSITION_PROPERTIES = Set(LINE, COLUMN, OFFSET, END_LINE, END_COLUMN, END_OFFSET, FILE, ID)
+  val POSITION_PROPERTIES =
+    Set(LINE, COLUMN, OFFSET, END_LINE, END_COLUMN, END_OFFSET, FILE, ID)
 
   val POSITION = "position"
   val LOCATION = "location"
--- a/src/Pure/General/position.scala	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/General/position.scala	Fri Sep 11 15:57:16 2009 +1000
@@ -6,11 +6,9 @@
 
 package isabelle
 
-import java.util.Properties
 
-
-object Position {
-
+object Position
+{
   type T = List[(String, String)]
 
   private def get_string(name: String, pos: T): Option[String] =
@@ -41,5 +39,4 @@
     val end = end_offset_of(pos)
     (begin, if (end.isDefined) end else begin.map(_ + 1))
   }
-
 }
--- a/src/Pure/General/scan.scala	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/General/scan.scala	Fri Sep 11 15:57:16 2009 +1000
@@ -11,7 +11,6 @@
 
 object Scan
 {
-
   /** Lexicon -- position tree **/
 
   object Lexicon
--- a/src/Pure/General/swing_thread.scala	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/General/swing_thread.scala	Fri Sep 11 15:57:16 2009 +1000
@@ -13,36 +13,43 @@
 
 object Swing_Thread
 {
+  /* checks */
+
+  def assert() = Predef.assert(SwingUtilities.isEventDispatchThread())
+  def require() = Predef.require(SwingUtilities.isEventDispatchThread())
+
+
   /* main dispatch queue */
 
   def now[A](body: => A): A = {
     var result: Option[A] = None
-    if (SwingUtilities.isEventDispatchThread) { result = Some(body) }
+    if (SwingUtilities.isEventDispatchThread()) { result = Some(body) }
     else SwingUtilities.invokeAndWait(new Runnable { def run = { result = Some(body) } })
     result.get
   }
 
   def later(body: => Unit) {
-    if (SwingUtilities.isEventDispatchThread) body
+    if (SwingUtilities.isEventDispatchThread()) body
     else SwingUtilities.invokeLater(new Runnable { def run = body })
   }
 
 
   /* delayed actions */
 
-  // turn multiple invocations into single action within time span
-  def delay(time_span: Int)(action: => Unit): () => Unit =
+  private def delayed_action(first: Boolean)(time_span: Int)(action: => Unit): () => Unit =
   {
     val listener =
       new ActionListener { override def actionPerformed(e: ActionEvent) { action } }
     val timer = new Timer(time_span, listener)
-    def invoke()
-    {
-      if (!timer.isRunning()) {
-        timer.setRepeats(false)
-        timer.start()
-      }
-    }
+    timer.setRepeats(false)
+
+    def invoke() { if (first) timer.start() else timer.restart() }
     invoke _
   }
+
+  // delayed action after first invocation
+  def delay_first = delayed_action(true) _
+
+  // delayed action after last invocation
+  def delay_last = delayed_action(false) _
 }
--- a/src/Pure/General/yxml.scala	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/General/yxml.scala	Fri Sep 11 15:57:16 2009 +1000
@@ -7,8 +7,8 @@
 package isabelle
 
 
-object YXML {
-
+object YXML
+{
   /* chunk markers */
 
   private val X = '\5'
@@ -22,7 +22,8 @@
   private def chunks(sep: Char, source: CharSequence) = new Iterator[CharSequence] {
     private val end = source.length
     private var state = if (end == 0) None else get_chunk(-1)
-    private def get_chunk(i: Int) = {
+    private def get_chunk(i: Int) =
+    {
       if (i < end) {
         var j = i; do j += 1 while (j < end && source.charAt(j) != sep)
         Some((source.subSequence(i + 1, j), j))
@@ -55,8 +56,8 @@
   }
 
 
-  def parse_body(source: CharSequence) = {
-
+  def parse_body(source: CharSequence): List[XML.Tree] =
+  {
     /* stack operations */
 
     var stack: List[((String, XML.Attributes), List[XML.Tree])] = null
@@ -94,7 +95,7 @@
     }
   }
 
-  def parse(source: CharSequence) =
+  def parse(source: CharSequence): XML.Tree =
     parse_body(source) match {
       case List(result) => result
       case Nil => XML.Text("")
@@ -108,14 +109,15 @@
     XML.Elem (Markup.BAD, Nil,
       List(XML.Text(source.toString.replace(X_string, "\\<^X>").replace(Y_string, "\\<^Y>"))))
 
-  def parse_body_failsafe(source: CharSequence) = {
+  def parse_body_failsafe(source: CharSequence): List[XML.Tree] =
+  {
     try { parse_body(source) }
     catch { case _: RuntimeException => List(markup_failsafe(source)) }
   }
 
-  def parse_failsafe(source: CharSequence) = {
+  def parse_failsafe(source: CharSequence): XML.Tree =
+  {
     try { parse(source) }
     catch { case _: RuntimeException => markup_failsafe(source) }
   }
-
 }
--- a/src/Pure/IsaMakefile	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/IsaMakefile	Fri Sep 11 15:57:16 2009 +1000
@@ -117,15 +117,14 @@
 
 ## Scala material
 
-SCALA_FILES = General/event_bus.scala General/markup.scala		\
-  General/position.scala General/scan.scala General/swing_thread.scala	\
-  General/symbol.scala General/xml.scala General/yxml.scala		\
-  Isar/isar.scala Isar/isar_document.scala Isar/outer_keyword.scala	\
+SCALA_FILES = General/event_bus.scala General/linear_set.scala		\
+  General/markup.scala General/position.scala General/scan.scala	\
+  General/swing_thread.scala General/symbol.scala General/xml.scala	\
+  General/yxml.scala Isar/isar_document.scala Isar/outer_keyword.scala	\
   System/cygwin.scala System/gui_setup.scala				\
-  System/isabelle_process.scala System/isabelle_system.scala		\
-  System/platform.scala Thy/completion.scala Thy/thy_header.scala	\
-  Tools/isabelle_syntax.scala
-
+  System/isabelle_process.scala System/isabelle_syntax.scala		\
+  System/isabelle_system.scala System/platform.scala			\
+  Thy/completion.scala Thy/thy_header.scala
 
 JAR_DIR = $(ISABELLE_HOME)/lib/classes
 PURE_JAR = $(JAR_DIR)/Pure.jar
@@ -135,7 +134,7 @@
 
 $(FULL_JAR): $(SCALA_FILES)
 	@rm -rf classes && mkdir classes
-	"$(SCALA_HOME)/bin/scalac" -deprecation -d classes -target jvm-1.5 $(SCALA_FILES)
+	"$(SCALA_HOME)/bin/scalac" -unchecked -deprecation -d classes -target jvm-1.5 $(SCALA_FILES)
 	"$(SCALA_HOME)/bin/scaladoc" -d classes $(SCALA_FILES)
 	@cp $(SCALA_FILES) classes/isabelle
 	@mkdir -p "$(JAR_DIR)"
--- a/src/Pure/Isar/code.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/Isar/code.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -19,11 +19,6 @@
   val constrset_of_consts: theory -> (string * typ) list
     -> string * ((string * sort) list * (string * typ list) list)
 
-  (*constant aliasses*)
-  val add_const_alias: thm -> theory -> theory
-  val triv_classes: theory -> class list
-  val resubst_alias: theory -> string -> string
-
   (*code equations*)
   val mk_eqn: theory -> thm * bool -> thm * bool
   val mk_eqn_warning: theory -> thm -> (thm * bool) option
@@ -169,7 +164,6 @@
 
 datatype spec = Spec of {
   history_concluded: bool,
-  aliasses: ((string * string) * thm) list * class list,
   eqns: ((bool * eqns) * (serial * eqns) list) Symtab.table
     (*with explicit history*),
   dtyps: ((serial * ((string * sort) list * (string * typ list) list)) list) Symtab.table
@@ -177,19 +171,16 @@
   cases: (int * (int * string list)) Symtab.table * unit Symtab.table
 };
 
-fun make_spec ((history_concluded, aliasses), (eqns, (dtyps, cases))) =
-  Spec { history_concluded = history_concluded, aliasses = aliasses,
-    eqns = eqns, dtyps = dtyps, cases = cases };
-fun map_spec f (Spec { history_concluded = history_concluded, aliasses = aliasses, eqns = eqns,
+fun make_spec (history_concluded, (eqns, (dtyps, cases))) =
+  Spec { history_concluded = history_concluded, eqns = eqns, dtyps = dtyps, cases = cases };
+fun map_spec f (Spec { history_concluded = history_concluded, eqns = eqns,
   dtyps = dtyps, cases = cases }) =
-  make_spec (f ((history_concluded, aliasses), (eqns, (dtyps, cases))));
-fun merge_spec (Spec { history_concluded = _, aliasses = aliasses1, eqns = eqns1,
+  make_spec (f (history_concluded, (eqns, (dtyps, cases))));
+fun merge_spec (Spec { history_concluded = _, eqns = eqns1,
     dtyps = dtyps1, cases = (cases1, undefs1) },
-  Spec { history_concluded = _, aliasses = aliasses2, eqns = eqns2,
+  Spec { history_concluded = _, eqns = eqns2,
     dtyps = dtyps2, cases = (cases2, undefs2) }) =
   let
-    val aliasses = (Library.merge (eq_snd Thm.eq_thm_prop) (pairself fst (aliasses1, aliasses2)),
-      Library.merge (op =) (pairself snd (aliasses1, aliasses2)));
     fun merge_eqns ((_, history1), (_, history2)) =
       let
         val raw_history = AList.merge (op = : serial * serial -> bool)
@@ -202,15 +193,13 @@
     val dtyps = Symtab.join (K (AList.merge (op =) (K true))) (dtyps1, dtyps2);
     val cases = (Symtab.merge (K true) (cases1, cases2),
       Symtab.merge (K true) (undefs1, undefs2));
-  in make_spec ((false, aliasses), (eqns, (dtyps, cases))) end;
+  in make_spec (false, (eqns, (dtyps, cases))) end;
 
 fun history_concluded (Spec { history_concluded, ... }) = history_concluded;
-fun the_aliasses (Spec { aliasses, ... }) = aliasses;
 fun the_eqns (Spec { eqns, ... }) = eqns;
 fun the_dtyps (Spec { dtyps, ... }) = dtyps;
 fun the_cases (Spec { cases, ... }) = cases;
-val map_history_concluded = map_spec o apfst o apfst;
-val map_aliasses = map_spec o apfst o apsnd;
+val map_history_concluded = map_spec o apfst;
 val map_eqns = map_spec o apsnd o apfst;
 val map_dtyps = map_spec o apsnd o apsnd o apfst;
 val map_cases = map_spec o apsnd o apsnd o apsnd;
@@ -264,7 +253,7 @@
 structure Code_Data = TheoryDataFun
 (
   type T = spec * data ref;
-  val empty = (make_spec ((false, ([], [])),
+  val empty = (make_spec (false,
     (Symtab.empty, (Symtab.empty, (Symtab.empty, Symtab.empty)))), ref empty_data);
   fun copy (spec, data) = (spec, ref (! data));
   val extend = copy;
@@ -358,24 +347,6 @@
 end; (*local*)
 
 
-(** retrieval interfaces **)
-
-(* constant aliasses *)
-
-fun resubst_alias thy =
-  let
-    val alias = (fst o the_aliasses o the_exec) thy;
-    val subst_inst_param = Option.map fst o AxClass.inst_of_param thy;
-    fun subst_alias c =
-      get_first (fn ((c', c''), _) => if c = c'' then SOME c' else NONE) alias;
-  in
-    perhaps subst_inst_param
-    #> perhaps subst_alias
-  end;
-
-val triv_classes = snd o the_aliasses o the_exec;
-
-
 (** foundation **)
 
 (* constants *)
@@ -669,38 +640,6 @@
 
 (** declaring executable ingredients **)
 
-(* constant aliasses *)
-
-fun add_const_alias thm thy =
-  let
-    val (ofclass, eqn) = case try Logic.dest_equals (Thm.prop_of thm)
-     of SOME ofclass_eq => ofclass_eq
-      | _ => error ("Bad certificate: " ^ Display.string_of_thm_global thy thm);
-    val (T, class) = case try Logic.dest_of_class ofclass
-     of SOME T_class => T_class
-      | _ => error ("Bad certificate: " ^ Display.string_of_thm_global thy thm);
-    val tvar = case try Term.dest_TVar T
-     of SOME (tvar as (_, sort)) => if null (filter (can (AxClass.get_info thy)) sort)
-          then tvar
-          else error ("Bad sort: " ^ Display.string_of_thm_global thy thm)
-      | _ => error ("Bad type: " ^ Display.string_of_thm_global thy thm);
-    val _ = if Term.add_tvars eqn [] = [tvar] then ()
-      else error ("Inconsistent type: " ^ Display.string_of_thm_global thy thm);
-    val lhs_rhs = case try Logic.dest_equals eqn
-     of SOME lhs_rhs => lhs_rhs
-      | _ => error ("Not an equation: " ^ Syntax.string_of_term_global thy eqn);
-    val c_c' = case try (pairself (check_const thy)) lhs_rhs
-     of SOME c_c' => c_c'
-      | _ => error ("Not an equation with two constants: "
-          ^ Syntax.string_of_term_global thy eqn);
-    val _ = if the_list (AxClass.class_of_param thy (snd c_c')) = [class] then ()
-      else error ("Inconsistent class: " ^ Display.string_of_thm_global thy thm);
-  in thy |>
-    (map_exec_purge NONE o map_aliasses) (fn (alias, classes) =>
-      ((c_c', thm) :: alias, insert (op =) class classes))
-  end;
-
-
 (* datatypes *)
 
 structure Type_Interpretation = InterpretationFun(type T = string * serial val eq = eq_snd (op =) : T * T -> bool);
--- a/src/Pure/Isar/expression.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/Isar/expression.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -839,7 +839,7 @@
 fun gen_sublocale prep_expr intern raw_target expression thy =
   let
     val target = intern thy raw_target;
-    val target_ctxt = Locale.init target thy;
+    val target_ctxt = TheoryTarget.init (SOME target) thy;
     val ((propss, deps, export), goal_ctxt) = prep_expr expression target_ctxt;
     fun after_qed witss = ProofContext.theory
       (fold (fn ((dep, morph), wits) => Locale.add_dependency
--- a/src/Pure/Isar/isar.scala	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,26 +0,0 @@
-/*  Title:      Pure/Isar/isar.scala
-    Author:     Makarius
-
-Isar document model.
-*/
-
-package isabelle
-
-
-class Isar(isabelle_system: Isabelle_System,
-    results: EventBus[Isabelle_Process.Result], args: String*)
-  extends Isabelle_Process(isabelle_system, results, args: _*)
-{
-  /* basic editor commands */
-
-  def create_command(id: String, text: String) =
-    output_sync("Isar.command " + IsabelleSyntax.encode_string(id) + " " +
-      IsabelleSyntax.encode_string(text))
-
-  def insert_command(prev: String, id: String) =
-    output_sync("Isar.insert " + IsabelleSyntax.encode_string(prev) + " " +
-      IsabelleSyntax.encode_string(id))
-
-  def remove_command(id: String) =
-    output_sync("Isar.remove " + IsabelleSyntax.encode_string(id))
-}
--- a/src/Pure/Isar/isar_document.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/Isar/isar_document.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -15,7 +15,7 @@
   val edit_document: document_id -> document_id -> (command_id * command_id option) list -> unit
 end;
 
-structure IsarDocument: ISAR_DOCUMENT =
+structure Isar_Document: ISAR_DOCUMENT =
 struct
 
 (* unique identifiers *)
--- a/src/Pure/Isar/isar_document.scala	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/Isar/isar_document.scala	Fri Sep 11 15:57:16 2009 +1000
@@ -6,7 +6,9 @@
 
 package isabelle
 
-object IsarDocument {
+
+object Isar_Document
+{
   /* unique identifiers */
 
   type State_ID = String
@@ -14,28 +16,29 @@
   type Document_ID = String
 }
 
-trait IsarDocument extends Isabelle_Process
+
+trait Isar_Document extends Isabelle_Process
 {
-  import IsarDocument._
+  import Isar_Document._
 
 
   /* commands */
 
   def define_command(id: Command_ID, text: String) {
-    output_sync("Isar.define_command " + IsabelleSyntax.encode_string(id) + " " +
-      IsabelleSyntax.encode_string(text))
+    output_sync("Isar.define_command " + Isabelle_Syntax.encode_string(id) + " " +
+      Isabelle_Syntax.encode_string(text))
   }
 
 
   /* documents */
 
   def begin_document(id: Document_ID, path: String) {
-    output_sync("Isar.begin_document " + IsabelleSyntax.encode_string(id) + " " +
-      IsabelleSyntax.encode_string(path))
+    output_sync("Isar.begin_document " + Isabelle_Syntax.encode_string(id) + " " +
+      Isabelle_Syntax.encode_string(path))
   }
 
   def end_document(id: Document_ID) {
-    output_sync("Isar.end_document " + IsabelleSyntax.encode_string(id))
+    output_sync("Isar.end_document " + Isabelle_Syntax.encode_string(id))
   }
 
   def edit_document(old_id: Document_ID, new_id: Document_ID,
@@ -44,21 +47,21 @@
     def append_edit(edit: (Command_ID, Option[Command_ID]), result: StringBuilder)
     {
       edit match {
-        case (id, None) => IsabelleSyntax.append_string(id, result)
+        case (id, None) => Isabelle_Syntax.append_string(id, result)
         case (id, Some(id2)) =>
-          IsabelleSyntax.append_string(id, result)
+          Isabelle_Syntax.append_string(id, result)
           result.append(" ")
-          IsabelleSyntax.append_string(id2, result)
+          Isabelle_Syntax.append_string(id2, result)
       }
     }
 
     val buf = new StringBuilder(40)
     buf.append("Isar.edit_document ")
-    IsabelleSyntax.append_string(old_id, buf)
+    Isabelle_Syntax.append_string(old_id, buf)
     buf.append(" ")
-    IsabelleSyntax.append_string(new_id, buf)
+    Isabelle_Syntax.append_string(new_id, buf)
     buf.append(" ")
-    IsabelleSyntax.append_list(append_edit, edits, buf)
+    Isabelle_Syntax.append_list(append_edit, edits, buf)
     output_sync(buf.toString)
   }
 }
--- a/src/Pure/Isar/isar_syn.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/Isar/isar_syn.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -30,7 +30,7 @@
 
 val _ =
   OuterSyntax.command "theory" "begin theory" (K.tag_theory K.thy_begin)
-    (ThyHeader.args >> (Toplevel.print oo IsarCmd.init_theory));
+    (Thy_Header.args >> (Toplevel.print oo IsarCmd.init_theory));
 
 val _ =
   OuterSyntax.command "end" "end (local) theory" (K.tag_theory K.thy_end)
--- a/src/Pure/Isar/outer_keyword.scala	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/Isar/outer_keyword.scala	Fri Sep 11 15:57:16 2009 +1000
@@ -6,8 +6,9 @@
 
 package isabelle
 
-object OuterKeyword {
 
+object OuterKeyword
+{
   val MINOR = "minor"
   val CONTROL = "control"
   val DIAG = "diag"
--- a/src/Pure/ML/ml_compiler_polyml-5.3.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/ML/ml_compiler_polyml-5.3.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -1,7 +1,7 @@
 (*  Title:      Pure/ML/ml_compiler_polyml-5.3.ML
     Author:     Makarius
 
-Advanced runtime compilation for Poly/ML 5.3 (SVN 773).
+Advanced runtime compilation for Poly/ML 5.3 (SVN 839).
 *)
 
 signature ML_COMPILER =
@@ -74,8 +74,8 @@
     (* input *)
 
     val location_props =
-      Markup.markup (Markup.position |> Markup.properties
-          (filter (member (op =) [Markup.idN, Markup.fileN] o #1) (Position.properties_of pos))) "";
+      op ^ (YXML.output_markup (Markup.position |> Markup.properties
+            (filter (member (op =) [Markup.idN, Markup.fileN] o #1) (Position.properties_of pos))));
 
     val input = toks |> maps (fn tok =>
       let
@@ -111,9 +111,8 @@
     fun put s = change output_buffer (Buffer.add s);
 
     fun put_message {message, hard, location, context = _} =
-     (put (if hard then "Error: " else "Warning: ");
-      put (Pretty.string_of (Pretty.from_ML (pretty_ml message)));
-      put (Position.str_of (position_of location) ^ "\n"));
+     (put ((if hard then "Error" else "Warning") ^ Position.str_of (position_of location) ^ ":\n");
+      put (Pretty.string_of (Pretty.from_ML (pretty_ml message)) ^ "\n"));
 
 
     (* results *)
--- a/src/Pure/ProofGeneral/pgip_parser.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/ProofGeneral/pgip_parser.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -20,7 +20,7 @@
 fun badcmd text = [D.Badcmd {text = text}];
 
 fun thy_begin text =
-  (case try (ThyHeader.read Position.none) (Source.of_string text) of
+  (case try (Thy_Header.read Position.none) (Source.of_string text) of
     NONE => D.Opentheory {thyname = NONE, parentnames = [], text = text}
   | SOME (name, imports, _) =>
        D.Opentheory {thyname = SOME name, parentnames = imports, text = text})
--- a/src/Pure/System/isabelle_process.scala	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/System/isabelle_process.scala	Fri Sep 11 15:57:16 2009 +1000
@@ -11,9 +11,12 @@
 import java.io.{BufferedReader, BufferedWriter, InputStreamReader, OutputStreamWriter,
   InputStream, OutputStream, IOException}
 
+import scala.actors.Actor
+import Actor._
 
-object Isabelle_Process {
 
+object Isabelle_Process
+{
   /* results */
 
   object Kind extends Enumeration {
@@ -104,8 +107,7 @@
 }
 
 
-class Isabelle_Process(isabelle_system: Isabelle_System,
-  results: EventBus[Isabelle_Process.Result], args: String*)
+class Isabelle_Process(isabelle_system: Isabelle_System, receiver: Actor, args: String*)
 {
   import Isabelle_Process._
 
@@ -113,7 +115,8 @@
   /* demo constructor */
 
   def this(args: String*) =
-    this(new Isabelle_System, new EventBus[Isabelle_Process.Result] + Console.println, args: _*)
+    this(new Isabelle_System,
+      new Actor { def act = loop { react { case res => Console.println(res) } } }.start, args: _*)
 
 
   /* process information */
@@ -127,11 +130,6 @@
 
   /* results */
 
-  def parse_message(result: Result): XML.Tree =
-    Isabelle_Process.parse_message(isabelle_system, result)
-
-  private val result_queue = new LinkedBlockingQueue[Result]
-
   private def put_result(kind: Kind.Value, props: List[(String, String)], result: String)
   {
     if (kind == Kind.INIT) {
@@ -139,24 +137,7 @@
       if (map.isDefinedAt(Markup.PID)) pid = map(Markup.PID)
       if (map.isDefinedAt(Markup.SESSION)) the_session = map(Markup.SESSION)
     }
-    result_queue.put(new Result(kind, props, result))
-  }
-
-  private class ResultThread extends Thread("isabelle: results") {
-    override def run() = {
-      var finished = false
-      while (!finished) {
-        val result =
-          try { result_queue.take }
-          catch { case _: NullPointerException => null }
-
-        if (result != null) {
-          results.event(result)
-          if (result.kind == Kind.EXIT) finished = true
-        }
-        else finished = true
-      }
-    }
+    receiver ! new Result(kind, props, result)
   }
 
 
@@ -204,14 +185,14 @@
 
 
   def command(text: String) =
-    output_sync("Isabelle.command " + IsabelleSyntax.encode_string(text))
+    output_sync("Isabelle.command " + Isabelle_Syntax.encode_string(text))
 
   def command(props: List[(String, String)], text: String) =
-    output_sync("Isabelle.command " + IsabelleSyntax.encode_properties(props) + " " +
-      IsabelleSyntax.encode_string(text))
+    output_sync("Isabelle.command " + Isabelle_Syntax.encode_properties(props) + " " +
+      Isabelle_Syntax.encode_string(text))
 
   def ML(text: String) =
-    output_sync("ML_val " + IsabelleSyntax.encode_string(text))
+    output_sync("ML_val " + Isabelle_Syntax.encode_string(text))
 
   def close() = synchronized {    // FIXME watchdog/timeout
     output_raw("\u0000")
@@ -396,8 +377,6 @@
     val message_thread = new MessageThread(message_fifo)
     message_thread.start
 
-    new ResultThread().start
-
 
     /* exec process */
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/System/isabelle_syntax.scala	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,74 @@
+/*  Title:      Pure/System/isabelle_syntax.scala
+    Author:     Makarius
+
+Isabelle outer syntax.
+*/
+
+package isabelle
+
+
+object Isabelle_Syntax
+{
+  /* string token */
+
+  def append_string(str: String, result: StringBuilder)
+  {
+    result.append("\"")
+    for (c <- str) {
+      if (c < 32 || c == '\\' || c == '\"') {
+        result.append("\\")
+        if (c < 10) result.append('0')
+        if (c < 100) result.append('0')
+        result.append(c.asInstanceOf[Int].toString)
+      }
+      else result.append(c)
+    }
+    result.append("\"")
+  }
+
+  def encode_string(str: String) =
+  {
+    val result = new StringBuilder(str.length + 10)
+    append_string(str, result)
+    result.toString
+  }
+
+
+  /* list */
+
+  def append_list[A](append_elem: (A, StringBuilder) => Unit, body: Iterable[A],
+    result: StringBuilder)
+  {
+    result.append("(")
+    val elems = body.elements
+    if (elems.hasNext) append_elem(elems.next, result)
+    while (elems.hasNext) {
+      result.append(", ")
+      append_elem(elems.next, result)
+    }
+    result.append(")")
+  }
+
+  def encode_list[A](append_elem: (A, StringBuilder) => Unit, elems: Iterable[A]) =
+  {
+    val result = new StringBuilder
+    append_list(append_elem, elems, result)
+    result.toString
+  }
+
+
+  /* properties */
+
+  def append_properties(props: List[(String, String)], result: StringBuilder)
+  {
+    append_list((p: (String, String), res) =>
+      { append_string(p._1, res); res.append(" = "); append_string(p._2, res) }, props, result)
+  }
+
+  def encode_properties(props: List[(String, String)]) =
+  {
+    val result = new StringBuilder
+    append_properties(props, result)
+    result.toString
+  }
+}
--- a/src/Pure/System/isabelle_system.scala	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/System/isabelle_system.scala	Fri Sep 11 15:57:16 2009 +1000
@@ -42,13 +42,11 @@
     val rc = proc.waitFor
     (output, rc)
   }
-
 }
 
 
 class Isabelle_System
 {
-
   /** unique ids **/
 
   private var id_count: BigInt = 0
@@ -244,6 +242,7 @@
   }
 
 
+
   /** system tools **/
 
   /* external processes */
@@ -296,6 +295,7 @@
   }
 
 
+
   /** Isabelle resources **/
 
   /* components */
--- a/src/Pure/System/isar.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/System/isar.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -22,12 +22,6 @@
   val toplevel_loop: {init: bool, welcome: bool, sync: bool, secure: bool} -> unit
   val loop: unit -> unit
   val main: unit -> unit
-
-  type id = string
-  val no_id: id
-  val create_command: Toplevel.transition -> id
-  val insert_command: id -> id -> unit
-  val remove_command: id -> unit
 end;
 
 structure Isar: ISAR =
@@ -145,7 +139,7 @@
 
 fun toplevel_loop {init = do_init, welcome, sync, secure} =
  (Context.set_thread_data NONE;
-  if do_init then init () else ();  (* FIXME init editor model *)
+  if do_init then init () else ();
   if welcome then writeln (Session.welcome ()) else ();
   uninterruptible (fn _ => fn () => raw_loop secure (OuterSyntax.isar sync)) ());
 
@@ -159,198 +153,6 @@
 
 
 
-(** individual toplevel commands **)
-
-(* unique identification *)
-
-type id = string;
-val no_id : id = "";
-
-
-(* command category *)
-
-datatype category = Empty | Theory | Proof | Diag | Control;
-
-fun category_of tr =
-  let val name = Toplevel.name_of tr in
-    if name = "" then Empty
-    else if OuterKeyword.is_theory name then Theory
-    else if OuterKeyword.is_proof name then Proof
-    else if OuterKeyword.is_diag name then Diag
-    else Control
-  end;
-
-val is_theory = fn Theory => true | _ => false;
-val is_proper = fn Theory => true | Proof => true | _ => false;
-val is_regular = fn Control => false | _ => true;
-
-
-(* command status *)
-
-datatype status =
-  Unprocessed |
-  Running |
-  Failed of exn * string |
-  Finished of Toplevel.state;
-
-fun status_markup Unprocessed = Markup.unprocessed
-  | status_markup Running = (Markup.runningN, [])
-  | status_markup (Failed _) = Markup.failed
-  | status_markup (Finished _) = Markup.finished;
-
-fun run int tr state =
-  (case Toplevel.transition int tr state of
-    NONE => NONE
-  | SOME (_, SOME err) => (Toplevel.error_msg tr err; SOME (Failed err))
-  | SOME (state', NONE) => SOME (Finished state'));
-
-
-(* datatype command *)
-
-datatype command = Command of
- {category: category,
-  transition: Toplevel.transition,
-  status: status};
-
-fun make_command (category, transition, status) =
-  Command {category = category, transition = transition, status = status};
-
-val empty_command =
-  make_command (Empty, Toplevel.empty, Finished Toplevel.toplevel);
-
-fun map_command f (Command {category, transition, status}) =
-  make_command (f (category, transition, status));
-
-fun map_status f = map_command (fn (category, transition, status) =>
-  (category, transition, f status));
-
-
-(* global collection of identified commands *)
-
-fun err_dup id = sys_error ("Duplicate command " ^ quote id);
-fun err_undef id = sys_error ("Unknown command " ^ quote id);
-
-local val global_commands = ref (Graph.empty: command Graph.T) in
-
-fun change_commands f = NAMED_CRITICAL "Isar" (fn () => change global_commands f)
-  handle Graph.DUP bad => err_dup bad | Graph.UNDEF bad => err_undef bad;
-
-fun get_commands () = NAMED_CRITICAL "Isar" (fn () => ! global_commands);
-
-end;
-
-fun add_edge (id1, id2) =
-  if id1 = no_id orelse id2 = no_id then I else Graph.add_edge (id1, id2);
-
-
-fun init_commands () = change_commands (K Graph.empty);
-
-fun the_command id =
-  let val Command cmd =
-    if id = no_id then empty_command
-    else (Graph.get_node (get_commands ()) id handle Graph.UNDEF bad => err_undef bad)
-  in cmd end;
-
-fun prev_command id =
-  if id = no_id then no_id
-  else
-    (case Graph.imm_preds (get_commands ()) id handle Graph.UNDEF bad => err_undef bad of
-      [] => no_id
-    | [prev] => prev
-    | _ => sys_error ("Non-linear command dependency " ^ quote id));
-
-fun next_commands id =
-  if id = no_id then []
-  else Graph.imm_succs (get_commands ()) id handle Graph.UNDEF bad => err_undef bad;
-
-fun descendant_commands ids =
-  Graph.all_succs (get_commands ()) (distinct (op =) (filter_out (fn id => id = no_id) ids))
-    handle Graph.UNDEF bad => err_undef bad;
-
-
-(* maintain status *)
-
-fun report_status markup id = Toplevel.status (#transition (the_command id)) markup;
-
-fun update_status status id = change_commands (Graph.map_node id (map_status (K status)));
-
-fun report_update_status status id =
-  change_commands (Graph.map_node id (map_status (fn old_status =>
-    let val markup = status_markup status
-    in if markup <> status_markup old_status then report_status markup id else (); status end)));
-
-
-(* create and dispose commands *)
-
-fun create_command raw_tr =
-  let
-    val (id, tr) =
-      (case Toplevel.get_id raw_tr of
-        SOME id => (id, raw_tr)
-      | NONE =>
-          let val id =
-            if ! Toplevel.debug then "isabelle:" ^ Toplevel.name_of raw_tr ^ serial_string ()
-            else "i" ^ serial_string ()
-          in (id, Toplevel.put_id id raw_tr) end);
-
-    val cmd = make_command (category_of tr, tr, Unprocessed);
-    val _ = change_commands (Graph.new_node (id, cmd));
-  in id end;
-
-fun dispose_commands ids =
-  let
-    val desc = descendant_commands ids;
-    val _ = List.app (report_status Markup.disposed) desc;
-    val _ = change_commands (Graph.del_nodes desc);
-  in () end;
-
-
-(* final state *)
-
-fun the_state id =
-  (case the_command id of
-    {status = Finished state, ...} => state
-  | {transition, ...} => error ("Unfinished command " ^ Toplevel.str_of transition));
-
-
-
-(** editor model **)
-
-(* run commands *)
-
-fun try_run id =
-  (case try the_state (prev_command id) of
-    NONE => ()
-  | SOME state =>
-      (case run true (#transition (the_command id)) state of
-        NONE => ()
-      | SOME status => report_update_status status id));
-
-fun rerun_commands ids =
-  (List.app (report_update_status Unprocessed) ids; List.app try_run ids);
-
-
-(* modify document *)
-
-fun insert_command prev id = NAMED_CRITICAL "Isar" (fn () =>
-  let
-    val nexts = next_commands prev;
-    val _ = change_commands
-     (fold (fn next => Graph.del_edge (prev, next)) nexts #> add_edge (prev, id) #>
-      fold (fn next => Graph.add_edge (id, next)) nexts);
-  in descendant_commands [id] end) |> rerun_commands;
-
-fun remove_command id = NAMED_CRITICAL "Isar" (fn () =>
-  let
-    val prev = prev_command id;
-    val nexts = next_commands id;
-    val _ = change_commands
-     (fold (fn next => Graph.del_edge (id, next)) nexts #>
-      fold (fn next => add_edge (prev, next)) nexts);
-  in descendant_commands nexts end) |> rerun_commands;
-
-
-
 (** command syntax **)
 
 local
@@ -392,24 +194,6 @@
   OuterSyntax.improper_command "kill" "kill partial proof or theory development" K.control
     (Scan.succeed (Toplevel.no_timing o Toplevel.imperative kill));
 
-
-(* editor model *)
-
-val _ =
-  OuterSyntax.internal_command "Isar.command"
-    (P.string -- P.string >> (fn (id, text) =>
-      Toplevel.imperative (fn () =>
-        ignore (create_command (OuterSyntax.prepare_command (Position.id id) text)))));
-
-val _ =
-  OuterSyntax.internal_command "Isar.insert"
-    (P.string -- P.string >> (fn (prev, id) =>
-      Toplevel.imperative (fn () => insert_command prev id)));
-
-val _ =
-  OuterSyntax.internal_command "Isar.remove"
-    (P.string >> (fn id => Toplevel.imperative (fn () => remove_command id)));
-
 end;
 
 end;
--- a/src/Pure/Thy/thy_header.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/Thy/thy_header.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -11,7 +11,7 @@
   val read: Position.T -> (string, 'a) Source.source -> string * string list * (string * bool) list
 end;
 
-structure ThyHeader: THY_HEADER =
+structure Thy_Header: THY_HEADER =
 struct
 
 structure T = OuterLex;
--- a/src/Pure/Thy/thy_header.scala	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/Thy/thy_header.scala	Fri Sep 11 15:57:16 2009 +1000
@@ -6,8 +6,9 @@
 
 package isabelle
 
-object ThyHeader {
 
+object Thy_Header
+{
   val HEADER = "header"
   val THEORY = "theory"
   val IMPORTS = "imports"
--- a/src/Pure/Thy/thy_load.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/Thy/thy_load.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -109,7 +109,7 @@
     val master as (path, _) = check_thy dir name;
     val text = explode (File.read path);
     val (name', imports, uses) =
-      ThyHeader.read (Path.position path) (Source.of_list_limited 8000 text);
+      Thy_Header.read (Path.position path) (Source.of_list_limited 8000 text);
     val _ = check_name name name';
     val uses = map (Path.explode o #1) uses;
   in {master = master, text = text, imports = imports, uses = uses} end;
--- a/src/Pure/Tools/isabelle_syntax.scala	Fri Sep 11 15:56:51 2009 +1000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,74 +0,0 @@
-/*  Title:      Pure/Tools/isabelle_syntax.scala
-    Author:     Makarius
-
-Isabelle outer syntax.
-*/
-
-package isabelle
-
-
-object IsabelleSyntax {
-
-  /* string token */
-
-  def append_string(str: String, result: StringBuilder)
-  {
-    result.append("\"")
-    for (c <- str) {
-      if (c < 32 || c == '\\' || c == '\"') {
-        result.append("\\")
-        if (c < 10) result.append('0')
-        if (c < 100) result.append('0')
-        result.append(c.asInstanceOf[Int].toString)
-      }
-      else result.append(c)
-    }
-    result.append("\"")
-  }
-
-  def encode_string(str: String) =
-  {
-    val result = new StringBuilder(str.length + 10)
-    append_string(str, result)
-    result.toString
-  }
-
-
-  /* list */
-
-  def append_list[A](append_elem: (A, StringBuilder) => Unit, body: Iterable[A],
-    result: StringBuilder)
-  {
-    result.append("(")
-    val elems = body.elements
-    if (elems.hasNext) append_elem(elems.next, result)
-    while (elems.hasNext) {
-      result.append(", ")
-      append_elem(elems.next, result)
-    }
-    result.append(")")
-  }
-
-  def encode_list[A](append_elem: (A, StringBuilder) => Unit, elems: Iterable[A]) =
-  {
-    val result = new StringBuilder
-    append_list(append_elem, elems, result)
-    result.toString
-  }
-
-
-  /* properties */
-
-  def append_properties(props: List[(String, String)], result: StringBuilder)
-  {
-    append_list((p: (String, String), res) =>
-      { append_string(p._1, res); res.append(" = "); append_string(p._2, res) }, props, result)
-  }
-
-  def encode_properties(props: List[(String, String)]) =
-  {
-    val result = new StringBuilder
-    append_properties(props, result)
-    result.toString
-  }
-}
--- a/src/Pure/display.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/display.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -27,12 +27,6 @@
   val string_of_thm_without_context: thm -> string
   val pretty_thms_aux: Proof.context -> bool -> thm list -> Pretty.T
   val pretty_thms: Proof.context -> thm list -> Pretty.T
-  val pretty_ctyp: ctyp -> Pretty.T
-  val string_of_ctyp: ctyp -> string
-  val print_ctyp: ctyp -> unit
-  val pretty_cterm: cterm -> Pretty.T
-  val string_of_cterm: cterm -> string
-  val print_cterm: cterm -> unit
   val print_syntax: theory -> unit
   val pretty_full_theory: bool -> theory -> Pretty.T list
 end;
@@ -121,17 +115,6 @@
 fun pretty_thms ctxt = pretty_thms_aux ctxt true;
 
 
-(* other printing commands *)
-
-fun pretty_ctyp cT = Syntax.pretty_typ_global (Thm.theory_of_ctyp cT) (Thm.typ_of cT);
-fun string_of_ctyp cT = Syntax.string_of_typ_global (Thm.theory_of_ctyp cT) (Thm.typ_of cT);
-val print_ctyp = writeln o string_of_ctyp;
-
-fun pretty_cterm ct = Syntax.pretty_term_global (Thm.theory_of_cterm ct) (Thm.term_of ct);
-fun string_of_cterm ct = Syntax.string_of_term_global (Thm.theory_of_cterm ct) (Thm.term_of ct);
-val print_cterm = writeln o string_of_cterm;
-
-
 
 (** print theory **)
 
--- a/src/Pure/old_goals.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Pure/old_goals.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -362,10 +362,7 @@
           (case Seq.pull (tac st0) of
                SOME(st,_) => st
              | _ => error ("prove_goal: tactic failed"))
-  in  mkresult (check, cond_timeit (!Output.timing) "" statef)  end
-  handle e => (print_sign_exn_unit (Thm.theory_of_cterm chorn) e;
-               writeln ("The exception above was raised for\n" ^
-                      Display.string_of_cterm chorn); raise e);
+  in  mkresult (check, cond_timeit (!Output.timing) "" statef)  end;
 
 (*Two variants: one checking the result, one not.
   Neither prints runtime messages: they are for internal packages.*)
--- a/src/Tools/Code/code_preproc.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Tools/Code/code_preproc.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -23,7 +23,6 @@
   val all: code_graph -> string list
   val pretty: theory -> code_graph -> Pretty.T
   val obtain: theory -> string list -> term list -> code_algebra * code_graph
-  val resubst_triv_consts: theory -> term -> term
   val eval_conv: theory
     -> (code_algebra -> code_graph -> (string * sort) list -> term -> cterm -> thm) -> cterm -> thm
   val eval: theory -> ((term -> term) -> 'a -> 'a)
@@ -73,10 +72,8 @@
   if AList.defined (op =) xs key then AList.delete (op =) key xs
   else error ("No such " ^ msg ^ ": " ^ quote key);
 
-fun map_data f thy =
-  thy
-  |> Code.purge_data
-  |> (Code_Preproc_Data.map o map_thmproc) f;
+fun map_data f = Code.purge_data
+  #> (Code_Preproc_Data.map o map_thmproc) f;
 
 val map_pre_post = map_data o apfst;
 val map_pre = map_pre_post o apfst;
@@ -163,10 +160,7 @@
     |> rhs_conv (Simplifier.rewrite post)
   end;
 
-fun resubst_triv_consts thy = map_aterms (fn t as Const (c, ty) => Const (Code.resubst_alias thy c, ty)
-  | t => t);
-
-fun postprocess_term thy = term_of_conv thy (postprocess_conv thy) #> resubst_triv_consts thy;
+fun postprocess_term thy = term_of_conv thy (postprocess_conv thy);
 
 fun print_codeproc thy =
   let
@@ -489,17 +483,6 @@
 fun obtain thy cs ts = apsnd snd
   (Wellsorted.change_yield thy (extend_arities_eqngr thy cs ts));
 
-fun prepare_sorts_typ prep_sort
-  = map_type_tfree (fn (v, sort) => TFree (v, prep_sort sort));
-
-fun prepare_sorts prep_sort (Const (c, ty)) =
-      Const (c, prepare_sorts_typ prep_sort ty)
-  | prepare_sorts prep_sort (t1 $ t2) =
-      prepare_sorts prep_sort t1 $ prepare_sorts prep_sort t2
-  | prepare_sorts prep_sort (Abs (v, ty, t)) =
-      Abs (v, prepare_sorts_typ prep_sort ty, prepare_sorts prep_sort t)
-  | prepare_sorts _ (t as Bound _) = t;
-
 fun gen_eval thy cterm_of conclude_evaluation evaluator proto_ct =
   let
     val pp = Syntax.pp_global thy;
@@ -512,12 +495,8 @@
     val vs = Term.add_tfrees t' [];
     val consts = fold_aterms
       (fn Const (c, _) => insert (op =) c | _ => I) t' [];
-
-    val add_triv_classes = curry (Sorts.inter_sort (Sign.classes_of thy))
-      (Code.triv_classes thy);
-    val t'' = prepare_sorts add_triv_classes t';
-    val (algebra', eqngr') = obtain thy consts [t''];
-  in conclude_evaluation (evaluator algebra' eqngr' vs t'' ct') thm end;
+    val (algebra', eqngr') = obtain thy consts [t'];
+  in conclude_evaluation (evaluator algebra' eqngr' vs t' ct') thm end;
 
 fun simple_evaluator evaluator algebra eqngr vs t ct =
   evaluator algebra eqngr vs t;
--- a/src/Tools/Code/code_thingol.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Tools/Code/code_thingol.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -429,7 +429,7 @@
 
 fun desymbolize_all_vars thy = desymbolize_tvars thy #> map (desymbolize_vars thy);
 
-fun clean_thms thy = same_arity thy #> desymbolize_all_vars thy;
+fun clean_thms thy = map (Thm.transfer thy) #> same_arity thy #> desymbolize_all_vars thy;
 
 
 (** statements, abstract programs **)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Code/etc/settings	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,2 @@
+
+ISABELLE_TOOLS="$ISABELLE_TOOLS:$COMPONENT/lib/Tools"
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Code/lib/Tools/codegen	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,65 @@
+#!/usr/bin/env bash
+#
+# Author: Florian Haftmann, TUM
+#
+# DESCRIPTION: issue code generation from shell
+
+
+PRG="$(basename "$0")"
+
+function usage()
+{
+  echo
+  echo "Usage: isabelle $PRG [OPTIONS] IMAGE THY CMD"
+  echo
+  echo "  Options are:"
+  echo "    -q    run in quick'n'dirty mode"
+  echo
+  echo "  Issues code generation using image IMAGE,"
+  echo "  theory THY,"
+  echo "  with Isar command 'export_code CMD'"
+  echo
+  exit 1
+}
+
+## process command line
+
+QUICK_AND_DIRTY=0
+
+while getopts "q" OPT
+do
+  case "$OPT" in
+    q)
+      QUICK_AND_DIRTY=1
+      ;;
+    \?)
+      usage
+      ;;
+  esac
+done
+
+shift $(($OPTIND - 1))
+
+[ "$#" -ne 3 ] && usage
+
+IMAGE="$1"; shift
+THY="$1"; shift
+CMD="$1"
+
+
+## main
+
+CODE_CMD=$(echo $CMD | perl -pe 's/\\/\\\\/g; s/"/\\\"/g')
+
+if [ "$QUICK_AND_DIRTY" -eq 1 ]
+then
+  QND_CMD="set"
+else
+  QND_CMD="reset"
+fi
+
+CTXT_CMD="ML_Context.eval_in (SOME (ProofContext.init (theory \"HOL\"))) false Position.none \"Code_Target.shell_command thyname cmd\";"
+
+FULL_CMD="$QND_CMD quick_and_dirty; val thyname = \"$THY\"; val cmd = \"$CODE_CMD\"; $CTXT_CMD"
+
+"$ISABELLE" -q -e "$FULL_CMD" "$IMAGE" || exit 1
--- a/src/Tools/induct.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Tools/induct.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -288,21 +288,21 @@
 
 (* prep_inst *)
 
-fun prep_inst thy align tune (tm, ts) =
+fun prep_inst ctxt align tune (tm, ts) =
   let
-    val cert = Thm.cterm_of thy;
+    val cert = Thm.cterm_of (ProofContext.theory_of ctxt);
     fun prep_var (x, SOME t) =
           let
             val cx = cert x;
             val xT = #T (Thm.rep_cterm cx);
             val ct = cert (tune t);
-            val tT = Thm.ctyp_of_term ct;
+            val tT = #T (Thm.rep_cterm ct);
           in
-            if Type.could_unify (Thm.typ_of tT, xT) then SOME (cx, ct)
+            if Type.could_unify (tT, xT) then SOME (cx, ct)
             else error (Pretty.string_of (Pretty.block
              [Pretty.str "Ill-typed instantiation:", Pretty.fbrk,
-              Display.pretty_cterm ct, Pretty.str " ::", Pretty.brk 1,
-              Display.pretty_ctyp (#T (Thm.crep_cterm ct))]))
+              Syntax.pretty_term ctxt (Thm.term_of ct), Pretty.str " ::", Pretty.brk 1,
+              Syntax.pretty_typ ctxt tT]))
           end
       | prep_var (_, NONE) = NONE;
     val xs = vars_of tm;
@@ -342,12 +342,11 @@
 fun cases_tac ctxt insts opt_rule facts =
   let
     val thy = ProofContext.theory_of ctxt;
-    val cert = Thm.cterm_of thy;
 
     fun inst_rule r =
       if null insts then `RuleCases.get r
       else (align_left "Rule has fewer premises than arguments given" (Thm.prems_of r) insts
-        |> maps (prep_inst thy align_left I)
+        |> maps (prep_inst ctxt align_left I)
         |> Drule.cterm_instantiate) r |> pair (RuleCases.get r);
 
     val ruleq =
@@ -411,8 +410,8 @@
 
 (* prepare rule *)
 
-fun rule_instance thy inst rule =
-  Drule.cterm_instantiate (prep_inst thy align_left I (Thm.prop_of rule, inst)) rule;
+fun rule_instance ctxt inst rule =
+  Drule.cterm_instantiate (prep_inst ctxt align_left I (Thm.prop_of rule, inst)) rule;
 
 fun internalize k th =
   th |> Thm.permute_prems 0 k
@@ -589,7 +588,7 @@
       (if null insts then `RuleCases.get r
        else (align_left "Rule has fewer conclusions than arguments given"
           (map Logic.strip_imp_concl (Logic.dest_conjunctions (Thm.concl_of r))) insts
-        |> maps (prep_inst thy align_right (atomize_term thy))
+        |> maps (prep_inst ctxt align_right (atomize_term thy))
         |> Drule.cterm_instantiate) r |> pair (RuleCases.get r))
       |> (fn ((cases, consumes), th) => (((cases, concls), consumes), th));
 
@@ -619,7 +618,7 @@
           THEN' inner_atomize_tac) j))
         THEN' atomize_tac) i st |> Seq.maps (fn st' =>
             guess_instance ctxt (internalize more_consumes rule) i st'
-            |> Seq.map (rule_instance thy (burrow_options (Variable.polymorphic ctxt) taking))
+            |> Seq.map (rule_instance ctxt (burrow_options (Variable.polymorphic ctxt) taking))
             |> Seq.maps (fn rule' =>
               CASES (rule_cases rule' cases)
                 (Tactic.rtac rule' i THEN
@@ -657,7 +656,7 @@
 
     fun inst_rule r =
       if null inst then `RuleCases.get r
-      else Drule.cterm_instantiate (prep_inst thy align_right I (main_prop_of r, inst)) r
+      else Drule.cterm_instantiate (prep_inst ctxt align_right I (main_prop_of r, inst)) r
         |> pair (RuleCases.get r);
 
     fun ruleq goal =
@@ -673,7 +672,7 @@
       |> Seq.maps (RuleCases.consume [] facts)
       |> Seq.maps (fn ((cases, (_, more_facts)), rule) =>
         guess_instance ctxt rule i st
-        |> Seq.map (rule_instance thy (burrow_options (Variable.polymorphic ctxt) taking))
+        |> Seq.map (rule_instance ctxt (burrow_options (Variable.polymorphic ctxt) taking))
         |> Seq.maps (fn rule' =>
           CASES (RuleCases.make_common false (thy, Thm.prop_of rule') cases)
             (Method.insert_tac more_facts i THEN Tactic.rtac rule' i) st)))
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/more_conv.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -0,0 +1,61 @@
+(* Title:  Tools/more_conv.ML
+   Author: Sascha Boehme
+
+Further conversions and conversionals.
+*)
+
+signature MORE_CONV =
+sig
+  val rewrs_conv: thm list -> conv
+
+  val sub_conv: (Proof.context -> conv) -> Proof.context -> conv
+  val bottom_conv: (Proof.context -> conv) -> Proof.context -> conv
+  val top_conv: (Proof.context -> conv) -> Proof.context -> conv
+  val top_sweep_conv: (Proof.context -> conv) -> Proof.context -> conv
+
+  val binder_conv: (Proof.context -> conv) -> Proof.context -> conv
+
+  val cache_conv: conv -> conv
+end
+
+
+
+structure More_Conv : MORE_CONV =
+struct
+
+
+fun rewrs_conv eqs = Conv.first_conv (map Conv.rewr_conv eqs)
+
+
+fun sub_conv conv ctxt =
+  Conv.comb_conv (conv ctxt) else_conv
+  Conv.abs_conv (fn (_, cx) => conv cx) ctxt else_conv
+  Conv.all_conv
+
+fun bottom_conv conv ctxt ct =
+  (sub_conv (bottom_conv conv) ctxt then_conv conv ctxt) ct
+
+fun top_conv conv ctxt ct =
+  (conv ctxt then_conv sub_conv (top_conv conv) ctxt) ct
+
+fun top_sweep_conv conv ctxt ct =
+  (conv ctxt else_conv sub_conv (top_sweep_conv conv) ctxt) ct
+
+
+fun binder_conv cv ctxt =
+  Conv.arg_conv (Conv.abs_conv (fn (_, cx) => cv cx) ctxt)
+
+
+fun cache_conv conv =
+  let 
+    val tab = ref Termtab.empty
+    fun add_result t thm =
+      let val _ = change tab (Termtab.insert Thm.eq_thm (t, thm))
+      in thm end
+    fun cconv ct =  
+      (case Termtab.lookup (!tab) (Thm.term_of ct) of
+        SOME thm => thm
+      | NONE => add_result (Thm.term_of ct) (conv ct))
+  in cconv end
+
+end
--- a/src/Tools/nbe.ML	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/Tools/nbe.ML	Fri Sep 11 15:57:16 2009 +1000
@@ -23,6 +23,7 @@
   val trace: bool ref
 
   val setup: theory -> theory
+  val add_const_alias: thm -> theory -> theory
 end;
 
 structure Nbe: NBE =
@@ -31,10 +32,107 @@
 (* generic non-sense *)
 
 val trace = ref false;
-fun tracing f x = if !trace then (Output.tracing (f x); x) else x;
+fun traced f x = if !trace then (tracing (f x); x) else x;
 
 
-(** the semantical universe **)
+(** certificates and oracle for "trivial type classes" **)
+
+structure Triv_Class_Data = TheoryDataFun
+(
+  type T = (class * thm) list;
+  val empty = [];
+  val copy = I;
+  val extend = I;
+  fun merge pp = AList.merge (op =) (K true);
+);
+
+fun add_const_alias thm thy =
+  let
+    val (ofclass, eqn) = case try Logic.dest_equals (Thm.prop_of thm)
+     of SOME ofclass_eq => ofclass_eq
+      | _ => error ("Bad certificate: " ^ Display.string_of_thm_global thy thm);
+    val (T, class) = case try Logic.dest_of_class ofclass
+     of SOME T_class => T_class
+      | _ => error ("Bad certificate: " ^ Display.string_of_thm_global thy thm);
+    val tvar = case try Term.dest_TVar T
+     of SOME (tvar as (_, sort)) => if null (filter (can (AxClass.get_info thy)) sort)
+          then tvar
+          else error ("Bad sort: " ^ Display.string_of_thm_global thy thm)
+      | _ => error ("Bad type: " ^ Display.string_of_thm_global thy thm);
+    val _ = if Term.add_tvars eqn [] = [tvar] then ()
+      else error ("Inconsistent type: " ^ Display.string_of_thm_global thy thm);
+    val lhs_rhs = case try Logic.dest_equals eqn
+     of SOME lhs_rhs => lhs_rhs
+      | _ => error ("Not an equation: " ^ Syntax.string_of_term_global thy eqn);
+    val c_c' = case try (pairself (Code.check_const thy)) lhs_rhs
+     of SOME c_c' => c_c'
+      | _ => error ("Not an equation with two constants: "
+          ^ Syntax.string_of_term_global thy eqn);
+    val _ = if the_list (AxClass.class_of_param thy (snd c_c')) = [class] then ()
+      else error ("Inconsistent class: " ^ Display.string_of_thm_global thy thm);
+  in Triv_Class_Data.map (AList.update (op =) (class, thm)) thy end;
+
+local
+
+val get_triv_classes = map fst o Triv_Class_Data.get;
+
+val (_, triv_of_class) = Context.>>> (Context.map_theory_result
+  (Thm.add_oracle (Binding.name "triv_of_class", fn (thy, (v, T), class) =>
+    Thm.cterm_of thy (Logic.mk_of_class (T, class)))));
+
+in
+
+fun lift_triv_classes_conv thy conv ct =
+  let
+    val algebra = Sign.classes_of thy;
+    val triv_classes = get_triv_classes thy;
+    val certT = Thm.ctyp_of thy;
+    fun critical_classes sort = filter_out (fn class => Sign.subsort thy (sort, [class])) triv_classes;
+    val vs = Term.add_tfrees (Thm.term_of ct) []
+      |> map_filter (fn (v, sort) => case critical_classes sort
+          of [] => NONE
+           | classes => SOME (v, ((sort, classes), Sorts.inter_sort algebra (triv_classes, sort))));
+    val of_classes = maps (fn (v, ((sort, classes), _)) => map (fn class =>
+      ((v, class), triv_of_class (thy, (v, TVar ((v, 0), sort)), class))) classes
+      @ map (fn class => ((v, class), Thm.of_class (certT (TVar ((v, 0), sort)), class)))
+        sort) vs;
+    fun strip_of_class thm =
+      let
+        val prem_props = (Logic.strip_imp_prems o Thm.prop_of) thm;
+        val prem_thms = map (the o AList.lookup (op =) of_classes
+          o apfst (fst o fst o dest_TVar) o Logic.dest_of_class) prem_props;
+      in Drule.implies_elim_list thm prem_thms end;
+  in ct
+    |> Drule.cterm_rule Thm.varifyT
+    |> Thm.instantiate_cterm (Thm.certify_inst thy (map (fn (v, ((sort, _), sort')) =>
+        (((v, 0), sort), TFree (v, sort'))) vs, []))
+    |> Drule.cterm_rule Thm.freezeT
+    |> conv
+    |> Thm.varifyT
+    |> fold (fn (v, (_, sort')) => Thm.unconstrainT (certT (TVar ((v, 0), sort')))) vs
+    |> Thm.certify_instantiate (map (fn (v, ((sort, _), _)) =>
+        (((v, 0), []), TVar ((v, 0), sort))) vs, [])
+    |> strip_of_class
+    |> Thm.freezeT
+  end;
+
+fun lift_triv_classes_rew thy rew t =
+  let
+    val algebra = Sign.classes_of thy;
+    val triv_classes = get_triv_classes thy;
+    val vs = Term.add_tfrees t [];
+  in t
+    |> (map_types o map_type_tfree)
+        (fn (v, sort) => TFree (v, Sorts.inter_sort algebra (sort, triv_classes)))
+    |> rew
+    |> (map_types o map_type_tfree)
+        (fn (v, _) => TFree (v, the (AList.lookup (op =) vs v)))
+  end;
+
+end;
+
+
+(** the semantic universe **)
 
 (*
    Functions are given by their semantical function value. To avoid
@@ -275,7 +373,7 @@
         val cs = map fst eqnss;
       in
         s
-        |> tracing (fn s => "\n--- code to be evaluated:\n" ^ s)
+        |> traced (fn s => "\n--- code to be evaluated:\n" ^ s)
         |> ML_Context.evaluate ctxt (!trace) univs_cookie
         |> (fn f => f deps_vals)
         |> (fn univs => cs ~~ univs)
@@ -450,14 +548,14 @@
     val string_of_term = setmp show_types true (Syntax.string_of_term_global thy);
   in
     compile_eval thy naming program (vs, t) deps
-    |> Code_Preproc.resubst_triv_consts thy
-    |> tracing (fn t => "Normalized:\n" ^ string_of_term t)
+    |> traced (fn t => "Normalized:\n" ^ string_of_term t)
     |> type_infer
-    |> tracing (fn t => "Types inferred:\n" ^ string_of_term t)
+    |> traced (fn t => "Types inferred:\n" ^ string_of_term t)
     |> check_tvars
-    |> tracing (fn t => "---\n")
+    |> traced (fn t => "---\n")
   end;
 
+
 (* evaluation oracle *)
 
 fun mk_equals thy lhs raw_rhs =
@@ -500,9 +598,9 @@
 val norm_conv = no_frees_conv (fn ct =>
   let
     val thy = Thm.theory_of_cterm ct;
-  in Code_Thingol.eval_conv thy (norm_oracle thy) ct end);
+  in lift_triv_classes_conv thy (Code_Thingol.eval_conv thy (norm_oracle thy)) ct end);
 
-fun norm thy = no_frees_rew (Code_Thingol.eval thy I (normalize thy));
+fun norm thy = lift_triv_classes_rew thy (no_frees_rew (Code_Thingol.eval thy I (normalize thy)));
 
 (* evaluation command *)
 
--- a/src/ZF/Datatype_ZF.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/ZF/Datatype_ZF.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -14,9 +14,9 @@
 (*Typechecking rules for most datatypes involving univ*)
 structure Data_Arg =
   struct
-  val intrs = 
+  val intrs =
       [@{thm SigmaI}, @{thm InlI}, @{thm InrI},
-       @{thm Pair_in_univ}, @{thm Inl_in_univ}, @{thm Inr_in_univ}, 
+       @{thm Pair_in_univ}, @{thm Inl_in_univ}, @{thm Inr_in_univ},
        @{thm zero_in_univ}, @{thm A_into_univ}, @{thm nat_into_univ}, @{thm UnCI}];
 
 
@@ -25,7 +25,7 @@
   end;
 
 
-structure Data_Package = 
+structure Data_Package =
   Add_datatype_def_Fun
    (structure Fp=Lfp and Pr=Standard_Prod and CP=Standard_CP
     and Su=Standard_Sum
@@ -37,16 +37,16 @@
 (*Typechecking rules for most codatatypes involving quniv*)
 structure CoData_Arg =
   struct
-  val intrs = 
+  val intrs =
       [@{thm QSigmaI}, @{thm QInlI}, @{thm QInrI},
-       @{thm QPair_in_quniv}, @{thm QInl_in_quniv}, @{thm QInr_in_quniv}, 
+       @{thm QPair_in_quniv}, @{thm QInl_in_quniv}, @{thm QInr_in_quniv},
        @{thm zero_in_quniv}, @{thm A_into_quniv}, @{thm nat_into_quniv}, @{thm UnCI}];
 
   val elims = [make_elim @{thm QInlD}, make_elim @{thm QInrD},   (*for mutual recursion*)
                @{thm QSigmaE}, @{thm qsumE}];                    (*allows * and + in spec*)
   end;
 
-structure CoData_Package = 
+structure CoData_Package =
   Add_datatype_def_Fun
    (structure Fp=Gfp and Pr=Quine_Prod and CP=Quine_CP
     and Su=Quine_Sum
@@ -69,9 +69,9 @@
  val datatype_ss = @{simpset};
 
  fun proc sg ss old =
-   let val _ = if !trace then writeln ("data_free: OLD = " ^ 
-                                       Display.string_of_cterm (cterm_of sg old))
-               else ()
+   let val _ =
+         if !trace then writeln ("data_free: OLD = " ^ Syntax.string_of_term_global sg old)
+         else ()
        val (lhs,rhs) = FOLogic.dest_eq old
        val (lhead, largs) = strip_comb lhs
        and (rhead, rargs) = strip_comb rhs
@@ -81,15 +81,15 @@
          handle Option => raise Match;
        val rcon_info = the (Symtab.lookup (ConstructorsData.get sg) rname)
          handle Option => raise Match;
-       val new = 
-           if #big_rec_name lcon_info = #big_rec_name rcon_info 
+       val new =
+           if #big_rec_name lcon_info = #big_rec_name rcon_info
                andalso not (null (#free_iffs lcon_info)) then
                if lname = rname then mk_new (largs, rargs)
                else Const("False",FOLogic.oT)
            else raise Match
-       val _ = if !trace then 
-                 writeln ("NEW = " ^ Display.string_of_cterm (Thm.cterm_of sg new))
-               else ();
+       val _ =
+         if !trace then writeln ("NEW = " ^ Syntax.string_of_term_global sg new)
+         else ();
        val goal = Logic.mk_equals (old, new)
        val thm = Goal.prove (Simplifier.the_context ss) [] [] goal
          (fn _ => rtac iff_reflection 1 THEN
--- a/src/ZF/ex/Limit.thy	Fri Sep 11 15:56:51 2009 +1000
+++ b/src/ZF/ex/Limit.thy	Fri Sep 11 15:57:16 2009 +1000
@@ -488,18 +488,24 @@
       and Mfun:  "M \<in> nat->nat->set(D)"
       and cpoD:  "cpo(D)"
   shows "matrix(D,M)"
-apply (simp add: matrix_def, safe)
-apply (rule Mfun)
-apply (cut_tac y1 = m and n = n in yprem [THEN chain_rel], simp+)
-apply (simp add: chain_rel xprem)
-apply (rule cpo_trans [OF cpoD])
-apply (cut_tac y1 = m and n = n in yprem [THEN chain_rel], simp+)
-apply (simp_all add: chain_fun [THEN apply_type] xprem)
-done
-
-lemma lemma_rel_rel:
-     "[|m \<in> nat; rel(D, (\<lambda>n \<in> nat. M`n`n)`m, y)|] ==> rel(D,M`m`m, y)"
-by simp
+proof -
+  {
+    fix n m assume "n : nat" "m : nat"
+    with chain_rel [OF yprem]
+    have "rel(D, M ` n ` m, M ` succ(n) ` m)" by simp
+  } note rel_succ = this
+  show "matrix(D,M)"
+  proof (simp add: matrix_def Mfun rel_succ, intro conjI ballI)
+    fix n m assume n: "n : nat" and m: "m : nat"
+    thus "rel(D, M ` n ` m, M ` n ` succ(m))"
+      by (simp add: chain_rel xprem)
+  next
+    fix n m assume n: "n : nat" and m: "m : nat"
+    thus "rel(D, M ` n ` m, M ` succ(n) ` succ(m))"
+      by (rule cpo_trans [OF cpoD rel_succ], 
+          simp_all add: chain_fun [THEN apply_type] xprem)
+  qed
+qed
 
 lemma lemma2:
      "[|x \<in> nat; m \<in> nat; rel(D,(\<lambda>n \<in> nat. M`n`m1)`x,(\<lambda>n \<in> nat. M`n`m1)`m)|]
@@ -509,65 +515,72 @@
 lemma isub_lemma:
     "[|isub(D, \<lambda>n \<in> nat. M`n`n, y); matrix(D,M); cpo(D)|] 
      ==> isub(D, \<lambda>n \<in> nat. lub(D,\<lambda>m \<in> nat. M`n`m), y)"
-apply (unfold isub_def, safe)
-apply (simp (no_asm_simp))
-apply (frule matrix_fun [THEN apply_type], assumption)
-apply (simp (no_asm_simp))
-apply (rule matrix_chain_left [THEN cpo_lub, THEN islub_least], assumption+)
-apply (unfold isub_def, safe)
-(*???VERY indirect proof: beta-redexes could be simplified now!*)
-apply (rename_tac k n)
-apply (case_tac "k le n")
-apply (rule cpo_trans, assumption)
-apply (rule lemma2)
-apply (rule_tac [4] lemma_rel_rel)
-prefer 5 apply blast
-apply (assumption | rule chain_rel_gen matrix_chain_right matrix_in isubD1)+
-txt{*opposite case*}
-apply (rule cpo_trans, assumption)
-apply (rule not_le_iff_lt [THEN iffD1, THEN leI, THEN chain_rel_gen])
-prefer 3 apply assumption
-apply (assumption | rule nat_into_Ord matrix_chain_left)+
-apply (rule lemma_rel_rel)
-apply (simp_all add: matrix_in)
-done
+proof (simp add: isub_def, safe)
+  fix n
+  assume DM: "matrix(D, M)" and D: "cpo(D)" and n: "n \<in> nat" and y: "y \<in> set(D)" 
+     and rel: "\<forall>n\<in>nat. rel(D, M ` n ` n, y)"
+  have "rel(D, lub(D, M ` n), y)"
+  proof (rule matrix_chain_left [THEN cpo_lub, THEN islub_least], simp_all add: n D DM)
+    show "isub(D, M ` n, y)" 
+      proof (unfold isub_def, intro conjI ballI y)
+	fix k assume k: "k \<in> nat"
+	show "rel(D, M ` n ` k, y)"
+	  proof (cases "n le k")
+	    case True 
+	    hence yy: "rel(D, M`n`k, M`k`k)" 
+	      by (blast intro: lemma2 n k y DM D chain_rel_gen matrix_chain_right) 
+	    show "?thesis"
+	      by (rule cpo_trans [OF D yy], 
+                  simp_all add: k rel n y DM matrix_in)
+	  next
+	    case False
+	    hence le: "k le n" 
+	      by (blast intro: not_le_iff_lt [THEN iffD1, THEN leI] nat_into_Ord n k) 
+	    show "?thesis"
+	      by (rule cpo_trans [OF D chain_rel_gen [OF le]], 
+		  simp_all add: n y k rel DM D matrix_chain_left)
+	  qed
+	qed
+  qed
+  moreover
+  have "M ` n \<in> nat \<rightarrow> set(D)" by (blast intro: DM n matrix_fun [THEN apply_type])
+  ultimately show "rel(D, lub(D, Lambda(nat, op `(M ` n))), y)"  by simp
+qed
 
 lemma matrix_chain_lub:
     "[|matrix(D,M); cpo(D)|] ==> chain(D,\<lambda>n \<in> nat. lub(D,\<lambda>m \<in> nat. M`n`m))"
-apply (simp add: chain_def, safe)
-apply (rule lam_type)
-apply (rule islub_in)
-apply (rule cpo_lub)
-prefer 2 apply assumption
-apply (rule chainI)
-apply (rule lam_type)
-apply (simp_all add: matrix_in)
-apply (rule matrix_rel_0_1, assumption+)
-apply (simp add: matrix_chain_left [THEN chain_fun, THEN eta])
-apply (rule dominate_islub)
-apply (rule_tac [3] cpo_lub)
-apply (rule_tac [2] cpo_lub)
-apply (simp add: dominate_def)
-apply (blast intro: matrix_rel_1_0)
-apply (simp_all add: matrix_chain_left nat_succI chain_fun)
-done
+proof (simp add: chain_def, intro conjI ballI)
+  assume "matrix(D, M)" "cpo(D)"
+  thus "(\<lambda>x\<in>nat. lub(D, Lambda(nat, op `(M ` x)))) \<in> nat \<rightarrow> set(D)"
+    by (force intro: islub_in cpo_lub chainI lam_type matrix_in matrix_rel_0_1) 
+next
+  fix n
+  assume DD: "matrix(D, M)" "cpo(D)" "n \<in> nat"
+  hence "dominate(D, M ` n, M ` succ(n))"
+    by (force simp add: dominate_def intro: matrix_rel_1_0)
+  with DD show "rel(D, lub(D, Lambda(nat, op `(M ` n))),
+               lub(D, Lambda(nat, op `(M ` succ(n)))))"
+    by (simp add: matrix_chain_left [THEN chain_fun, THEN eta] 
+      dominate_islub cpo_lub matrix_chain_left chain_fun)
+qed
 
 lemma isub_eq:
-    "[|matrix(D,M); cpo(D)|] 
-     ==> isub(D,(\<lambda>n \<in> nat. lub(D,\<lambda>m \<in> nat. M`n`m)),y) <->
-         isub(D,(\<lambda>n \<in> nat. M`n`n),y)"
-apply (rule iffI)
-apply (rule dominate_isub)
-prefer 2 apply assumption
-apply (simp add: dominate_def)
-apply (rule ballI)
-apply (rule bexI, auto)
-apply (simp add: matrix_chain_left [THEN chain_fun, THEN eta])
-apply (rule islub_ub)
-apply (rule cpo_lub)
-apply (simp_all add: matrix_chain_left matrix_chain_diag chain_fun 
-                     matrix_chain_lub isub_lemma)
-done
+  assumes DM: "matrix(D, M)" and D: "cpo(D)"
+  shows "isub(D,(\<lambda>n \<in> nat. lub(D,\<lambda>m \<in> nat. M`n`m)),y) <-> isub(D,(\<lambda>n \<in> nat. M`n`n),y)"
+proof
+  assume isub: "isub(D, \<lambda>n\<in>nat. lub(D, Lambda(nat, op `(M ` n))), y)"
+  hence dom: "dominate(D, \<lambda>n\<in>nat. M ` n ` n, \<lambda>n\<in>nat. lub(D, Lambda(nat, op `(M ` n))))" 
+    using DM D
+    by (simp add: dominate_def, intro ballI bexI,
+        simp_all add: matrix_chain_left [THEN chain_fun, THEN eta] islub_ub cpo_lub matrix_chain_left)
+  thus "isub(D, \<lambda>n\<in>nat. M ` n ` n, y)" using DM D
+    by - (rule dominate_isub [OF dom isub], 
+          simp_all add: matrix_chain_diag chain_fun matrix_chain_lub)
+next
+  assume isub: "isub(D, \<lambda>n\<in>nat. M ` n ` n, y)" 
+  thus "isub(D, \<lambda>n\<in>nat. lub(D, Lambda(nat, op `(M ` n))), y)"  using DM D
+    by (simp add: isub_lemma)
+qed
 
 lemma lub_matrix_diag_aux1:
     "lub(D,(\<lambda>n \<in> nat. lub(D,\<lambda>m \<in> nat. M`n`m))) =
@@ -695,34 +708,43 @@
     "[|chain(cf(D,E),X); chain(D,Xa); cpo(D); cpo(E) |] 
      ==> matrix(E,\<lambda>x \<in> nat. \<lambda>xa \<in> nat. X`x`(Xa`xa))"
 apply (rule matrix_chainI, auto)
-apply (rule chainI)
-apply (blast intro: lam_type apply_funtype cont_fun cf_cont chain_in, simp)
-apply (blast intro: cont_mono nat_succI chain_rel cf_cont chain_in)
-apply (rule chainI)
-apply (blast intro: lam_type apply_funtype cont_fun cf_cont chain_in, simp)
-apply (rule rel_cf)
-apply (simp_all add: chain_in chain_rel)
+apply (force intro: chainI lam_type apply_funtype cont_fun cf_cont cont_mono)
+apply (force intro: chainI lam_type apply_funtype cont_fun cf_cont rel_cf)
 apply (blast intro: lam_type apply_funtype cont_fun cf_cont chain_in)
 done
 
 lemma chain_cf_lub_cont:
-    "[|chain(cf(D,E),X); cpo(D); cpo(E) |] 
-     ==> (\<lambda>x \<in> set(D). lub(E, \<lambda>n \<in> nat. X ` n ` x)) \<in> cont(D, E)"
-apply (rule contI)
-apply (rule lam_type)
-apply (assumption | rule chain_cf [THEN cpo_lub, THEN islub_in])+
-apply simp
-apply (rule dominate_islub)
-apply (erule_tac [2] chain_cf [THEN cpo_lub], simp_all)+
-apply (rule dominateI, assumption, simp)
-apply (assumption | rule chain_in [THEN cf_cont, THEN cont_mono])+
-apply (assumption | rule chain_cf [THEN chain_fun])+
-apply (simp add: cpo_lub [THEN islub_in] 
-                 chain_in [THEN cf_cont, THEN cont_lub])
-apply (frule matrix_lemma [THEN lub_matrix_diag], assumption+)
-apply (simp add: chain_in [THEN beta])
-apply (drule matrix_lemma [THEN lub_matrix_diag_sym], auto)
-done
+  assumes ch: "chain(cf(D,E),X)" and D: "cpo(D)" and E: "cpo(E)"
+  shows   "(\<lambda>x \<in> set(D). lub(E, \<lambda>n \<in> nat. X ` n ` x)) \<in> cont(D, E)"
+proof (rule contI)
+  show "(\<lambda>x\<in>set(D). lub(E, \<lambda>n\<in>nat. X ` n ` x)) \<in> set(D) \<rightarrow> set(E)"
+    by (blast intro: lam_type chain_cf [THEN cpo_lub, THEN islub_in] ch E) 
+next
+  fix x y
+  assume xy: "rel(D, x, y)" "x \<in> set(D)" "y \<in> set(D)"
+  hence dom: "dominate(E, \<lambda>n\<in>nat. X ` n ` x, \<lambda>n\<in>nat. X ` n ` y)"
+    by (force intro: dominateI  chain_in [OF ch, THEN cf_cont, THEN cont_mono])
+  note chE = chain_cf [OF ch] 
+  from xy show "rel(E, (\<lambda>x\<in>set(D). lub(E, \<lambda>n\<in>nat. X ` n ` x)) ` x,
+                       (\<lambda>x\<in>set(D). lub(E, \<lambda>n\<in>nat. X ` n ` x)) ` y)"
+    by (simp add: dominate_islub [OF dom] cpo_lub [OF chE] E chain_fun [OF chE])
+next
+  fix Y
+  assume chDY: "chain(D,Y)"
+  have "lub(E, \<lambda>x\<in>nat. lub(E, \<lambda>y\<in>nat. X ` x ` (Y ` y))) =
+        lub(E, \<lambda>x\<in>nat. X ` x ` (Y ` x))"
+    using matrix_lemma [THEN lub_matrix_diag, OF ch chDY]
+    by (simp add: D E)
+   also have "... =  lub(E, \<lambda>x\<in>nat. lub(E, \<lambda>n\<in>nat. X ` n ` (Y ` x)))"
+    using  matrix_lemma [THEN lub_matrix_diag_sym, OF ch chDY]
+    by (simp add: D E) 
+  finally have "lub(E, \<lambda>x\<in>nat. lub(E, \<lambda>n\<in>nat. X ` x ` (Y ` n))) =
+                lub(E, \<lambda>x\<in>nat. lub(E, \<lambda>n\<in>nat. X ` n ` (Y ` x)))" .
+  thus "(\<lambda>x\<in>set(D). lub(E, \<lambda>n\<in>nat. X ` n ` x)) ` lub(D, Y) =
+        lub(E, \<lambda>n\<in>nat. (\<lambda>x\<in>set(D). lub(E, \<lambda>n\<in>nat. X ` n ` x)) ` (Y ` n))"
+    by (simp add: cpo_lub [THEN islub_in]  D chDY
+                  chain_in [THEN cf_cont, THEN cont_lub, OF ch])
+  qed
 
 lemma islub_cf:
     "[| chain(cf(D,E),X); cpo(D); cpo(E)|]