merged
authorpaulson
Wed, 26 Aug 2009 12:52:01 +0100
changeset 32407 9c1b3e2f1b88
parent 32402 5731300da417 (diff)
parent 32406 9ea59bd1397a (current diff)
child 32409 7e38dedf3f7d
merged
--- a/bin/isabelle	Wed Aug 26 12:51:38 2009 +0100
+++ b/bin/isabelle	Wed Aug 26 12:52:01 2009 +0100
@@ -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	Wed Aug 26 12:51:38 2009 +0100
+++ b/bin/isabelle-process	Wed Aug 26 12:52:01 2009 +0100
@@ -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/etc/components	Wed Aug 26 12:51:38 2009 +0100
+++ b/etc/components	Wed Aug 26 12:52:01 2009 +0100
@@ -12,5 +12,5 @@
 src/Sequents
 #misc components
 src/HOL/Tools/ATP_Manager
+src/HOL/Tools/Mirabelle
 src/HOL/Library/Sum_Of_Squares
-src/HOL/Tools/Mirabelle
--- a/etc/settings	Wed Aug 26 12:51:38 2009 +0100
+++ b/etc/settings	Wed Aug 26 12:52:01 2009 +0100
@@ -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/doc	Wed Aug 26 12:51:38 2009 +0100
+++ b/lib/Tools/doc	Wed Aug 26 12:52:01 2009 +0100
@@ -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	Wed Aug 26 12:51:38 2009 +0100
+++ b/lib/Tools/document	Wed Aug 26 12:52:01 2009 +0100
@@ -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	Wed Aug 26 12:51:38 2009 +0100
+++ b/lib/Tools/findlogics	Wed Aug 26 12:52:01 2009 +0100
@@ -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	Wed Aug 26 12:51:38 2009 +0100
+++ /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	Wed Aug 26 12:51:38 2009 +0100
+++ b/lib/Tools/makeall	Wed Aug 26 12:52:01 2009 +0100
@@ -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	Wed Aug 26 12:51:38 2009 +0100
+++ b/lib/scripts/getsettings	Wed Aug 26 12:52:01 2009 +0100
@@ -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/src/HOL/HOL.thy	Wed Aug 26 12:51:38 2009 +0100
+++ b/src/HOL/HOL.thy	Wed Aug 26 12:52:01 2009 +0100
@@ -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} *}
--- a/src/HOL/IsaMakefile	Wed Aug 26 12:51:38 2009 +0100
+++ b/src/HOL/IsaMakefile	Wed Aug 26 12:52:01 2009 +0100
@@ -108,6 +108,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 \
@@ -900,8 +901,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
 
 
--- a/src/HOL/Library/normarith.ML	Wed Aug 26 12:51:38 2009 +0100
+++ b/src/HOL/Library/normarith.ML	Wed Aug 26 12:52:01 2009 +0100
@@ -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	Wed Aug 26 12:51:38 2009 +0100
+++ b/src/HOL/Library/positivstellensatz.ML	Wed Aug 26 12:52:01 2009 +0100
@@ -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/SetInterval.thy	Wed Aug 26 12:51:38 2009 +0100
+++ b/src/HOL/SetInterval.thy	Wed Aug 26 12:52:01 2009 +0100
@@ -186,26 +186,60 @@
   seems to take forever (more than one hour). *}
 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
 
 subsection {* Intervals of natural numbers *}
--- a/src/HOL/Tools/Function/fundef_core.ML	Wed Aug 26 12:51:38 2009 +0100
+++ b/src/HOL/Tools/Function/fundef_core.ML	Wed Aug 26 12:52:01 2009 +0100
@@ -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/Mirabelle/Tools/mirabelle.ML	Wed Aug 26 12:51:38 2009 +0100
+++ b/src/HOL/Tools/Mirabelle/Tools/mirabelle.ML	Wed Aug 26 12:52:01 2009 +0100
@@ -7,11 +7,11 @@
   type action
   val register : string * action -> theory -> theory
 
+  val logfile : string Config.T
   val timeout : int Config.T
   val verbose : bool Config.T
   val start_line : int Config.T
   val end_line : int Config.T
-  val set_logfile : string -> theory -> theory
 
   val goal_thm_of : Proof.state -> thm
   val can_apply : (Proof.context -> int -> tactic) -> Proof.state -> bool
@@ -59,10 +59,6 @@
 
 val setup = 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 =
@@ -90,15 +86,11 @@
   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 =
+fun pretty_print 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 ^ "):"
+    val head = "at " ^ loc ^ " (" ^ name ^ "):"
 
     fun pretty_msg (name, msg) = Pretty.block (map Pretty.str [name, ": ", msg])
   in
@@ -119,7 +111,7 @@
     Actions.get thy
     |> Symtab.dest
     |> only_within_range thy pos (map_filter (apply_action (verb, secs) st))
-    |> (fn [] => () | msgs => log thy (pretty_print verb pos name msgs))
+    |> (fn [] => () | msgs => log thy (pretty_print pos name msgs))
   end
 
 end
--- a/src/HOL/Tools/Mirabelle/etc/settings	Wed Aug 26 12:51:38 2009 +0100
+++ b/src/HOL/Tools/Mirabelle/etc/settings	Wed Aug 26 12:52:01 2009 +0100
@@ -6,4 +6,4 @@
 MIRABELLE_TIMEOUT=30
 MIRABELLE_VERBOSE=false
 
-ISABELLE_TOOLS="$ISABELLE_TOOLS:$MIRABELLE_HOME/lib/Tools"
+ISABELLE_TOOLS="$ISABELLE_TOOLS:$COMPONENT/lib/Tools"
--- a/src/HOL/Tools/Mirabelle/lib/Tools/mirabelle	Wed Aug 26 12:51:38 2009 +0100
+++ b/src/HOL/Tools/Mirabelle/lib/Tools/mirabelle	Wed Aug 26 12:52:01 2009 +0100
@@ -7,15 +7,17 @@
 
 PRG="$(basename "$0")"
 
-function action_names() {
+function print_action_names() {
   TOOLS="$MIRABELLE_HOME/Tools/mirabelle_*.ML"
-  ACTION_NAMES=`find $TOOLS | sed 's/.*mirabelle_\(.*\)\.ML/\1/'`
+  for NAME in $TOOLS
+  do
+    echo $NAME | sed 's/.*mirabelle_\(.*\)\.ML/    \1/'
+  done
 }
 
 function usage() {
   out="$MIRABELLE_OUTPUT_PATH"
   timeout="$MIRABELLE_TIMEOUT"
-  action_names
   echo
   echo "Usage: isabelle $PRG [OPTIONS] ACTIONS FILES"
   echo
@@ -31,14 +33,11 @@
   echo
   echo "  ACTIONS is a colon-separated list of actions, where each action is"
   echo "  either NAME or NAME[KEY=VALUE,...,KEY=VALUE]. Available actions are:"
-  for NAME in $ACTION_NAMES
-  do
-    echo "    $NAME"
-  done
+  print_action_names
   echo
-  echo "  FILES is a space-separated list of theory files, where each file is"
-  echo "  either NAME.thy or NAME.thy[START:END] and START and END are numbers"
-  echo "  indicating the range the given actions are to be applied."
+  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
 }
@@ -48,7 +47,7 @@
 
 # options
 
-while getopts "L:T:O:vt:" OPT
+while getopts "L:T:O:vt:?" OPT
 do
   case "$OPT" in
     L)
@@ -61,7 +60,7 @@
       MIRABELLE_OUTPUT_PATH="$OPTARG"
       ;;
     v)
-      MIRABELLE_VERBOSE=true
+      MIRABELLE_VERBOSE="true"
       ;;
     t)
       MIRABELLE_TIMEOUT="$OPTARG"
@@ -81,13 +80,13 @@
 
 # setup
 
-mkdir -p $MIRABELLE_OUTPUT_PATH
+mkdir -p "$MIRABELLE_OUTPUT_PATH"
 
 
 ## main
 
 for FILE in "$@"
 do
-  perl -w $MIRABELLE_HOME/lib/scripts/mirabelle.pl $ACTIONS "$FILE"
+  perl -w "$MIRABELLE_HOME/lib/scripts/mirabelle.pl" "$ACTIONS" "$FILE"
 done
 
--- a/src/HOL/Tools/Mirabelle/lib/scripts/mirabelle.pl	Wed Aug 26 12:51:38 2009 +0100
+++ b/src/HOL/Tools/Mirabelle/lib/scripts/mirabelle.pl	Wed Aug 26 12:52:01 2009 +0100
@@ -43,9 +43,11 @@
 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 = "";
@@ -62,7 +64,7 @@
 begin
 
 setup {* 
-  Mirabelle.set_logfile "$log_file" #>
+  Config.put_thy Mirabelle.logfile "$log_file" #>
   Config.put_thy Mirabelle.timeout $timeout #>
   Config.put_thy Mirabelle.verbose $verbose #>
   Config.put_thy Mirabelle.start_line $start_line #>
@@ -99,7 +101,8 @@
 
 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 "$setup_thy_name" /gm;
+$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'";
@@ -114,14 +117,21 @@
 
 # run isabelle
 
-my $r = system "$isabelle_home/bin/isabelle-process " .
+open(LOG_FILE, ">$log_file");
+print LOG_FILE "Run of $new_thy_file with:\n";
+foreach $name (@action_names) {
+  print LOG_FILE "  $name\n";
+}
+print LOG_FILE "\n\n";
+close(LOG_FILE);
+
+my $r = system "\"$ISABELLE_PROCESS\" " .
   "-e 'use \"$root_file\";' -q $mirabelle_logic" . "\n";
 
 
 # cleanup
 
 unlink $root_file;
-unlink $new_thy_file;
 unlink $setup_file;
 
 exit $r;
--- a/src/HOL/Tools/Qelim/cooper.ML	Wed Aug 26 12:51:38 2009 +0100
+++ b/src/HOL/Tools/Qelim/cooper.ML	Wed Aug 26 12:52:01 2009 +0100
@@ -246,8 +246,10 @@
              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
--- a/src/HOL/ex/Termination.thy	Wed Aug 26 12:51:38 2009 +0100
+++ b/src/HOL/ex/Termination.thy	Wed Aug 26 12:52:01 2009 +0100
@@ -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"
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/more_conv.ML	Wed Aug 26 12:52:01 2009 +0100
@@ -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