merged;
authorwenzelm
Sat, 17 Mar 2012 11:57:03 +0100
changeset 46983 216a839841bc
parent 46982 144d94446378 (current diff)
parent 46979 ef4b0d6b2fb6 (diff)
child 46984 0f1e85fcf5f4
merged;
NEWS
--- a/NEWS	Sat Mar 17 11:35:18 2012 +0100
+++ b/NEWS	Sat Mar 17 11:57:03 2012 +0100
@@ -48,6 +48,10 @@
 
 *** Pure ***
 
+* Command 'definition' no longer exports the foundational "raw_def"
+into the user context.  Minor INCOMPATIBILITY, may use the regular
+"def" result with attribute "abs_def" to imitate the old version.
+
 * Attribute "abs_def" turns an equation of the form "f x y == t" into
 "f == %x y. t", which ensures that "simp" or "unfold" steps always
 expand it.  This also works for object-logic equality.  (Formerly
--- a/src/HOL/Orderings.thy	Sat Mar 17 11:35:18 2012 +0100
+++ b/src/HOL/Orderings.thy	Sat Mar 17 11:57:03 2012 +0100
@@ -1314,7 +1314,6 @@
 
 definition
   [no_atp]: "\<top> = (\<lambda>x. \<top>)"
-declare top_fun_def_raw [no_atp]
 
 lemma top_apply [simp] (* CANDIDATE [code] *):
   "\<top> x = \<top>"
--- a/src/HOL/Quickcheck.thy	Sat Mar 17 11:35:18 2012 +0100
+++ b/src/HOL/Quickcheck.thy	Sat Mar 17 11:57:03 2012 +0100
@@ -43,8 +43,9 @@
 instantiation itself :: (typerep) random
 begin
 
-definition random_itself :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a itself \<times> (unit \<Rightarrow> term)) \<times> Random.seed" where
-  "random_itself _ = Pair (Code_Evaluation.valtermify TYPE('a))"
+definition
+  random_itself :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a itself \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
+where "random_itself _ = Pair (Code_Evaluation.valtermify TYPE('a))"
 
 instance ..
 
@@ -73,7 +74,9 @@
 instantiation nat :: random
 begin
 
-definition random_nat :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> (nat \<times> (unit \<Rightarrow> Code_Evaluation.term)) \<times> Random.seed" where
+definition random_nat :: "code_numeral \<Rightarrow> Random.seed
+  \<Rightarrow> (nat \<times> (unit \<Rightarrow> Code_Evaluation.term)) \<times> Random.seed"
+where
   "random_nat i = Random.range (i + 1) \<circ>\<rightarrow> (\<lambda>k. Pair (
      let n = Code_Numeral.nat_of k
      in (n, \<lambda>_. Code_Evaluation.term_of n)))"
@@ -100,18 +103,22 @@
 text {* Towards @{typ "'a \<Rightarrow> 'b"} *}
 
 axiomatization random_fun_aux :: "typerep \<Rightarrow> typerep \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> term)
-  \<Rightarrow> (Random.seed \<Rightarrow> ('b \<times> (unit \<Rightarrow> term)) \<times> Random.seed) \<Rightarrow> (Random.seed \<Rightarrow> Random.seed \<times> Random.seed)
+  \<Rightarrow> (Random.seed \<Rightarrow> ('b \<times> (unit \<Rightarrow> term)) \<times> Random.seed)
+  \<Rightarrow> (Random.seed \<Rightarrow> Random.seed \<times> Random.seed)
   \<Rightarrow> Random.seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
 
 definition random_fun_lift :: "(Random.seed \<Rightarrow> ('b \<times> (unit \<Rightarrow> term)) \<times> Random.seed)
-  \<Rightarrow> Random.seed \<Rightarrow> (('a\<Colon>term_of \<Rightarrow> 'b\<Colon>typerep) \<times> (unit \<Rightarrow> term)) \<times> Random.seed" where
-  "random_fun_lift f = random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Evaluation.term_of f Random.split_seed"
+  \<Rightarrow> Random.seed \<Rightarrow> (('a\<Colon>term_of \<Rightarrow> 'b\<Colon>typerep) \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
+where
+  "random_fun_lift f =
+    random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Evaluation.term_of f Random.split_seed"
 
 instantiation "fun" :: ("{equal, term_of}", random) random
 begin
 
-definition random_fun :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> Random.seed" where
-  "random i = random_fun_lift (random i)"
+definition
+  random_fun :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
+  where "random i = random_fun_lift (random i)"
 
 instance ..
 
@@ -119,19 +126,21 @@
 
 text {* Towards type copies and datatypes *}
 
-definition collapse :: "('a \<Rightarrow> ('a \<Rightarrow> 'b \<times> 'a) \<times> 'a) \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a" where
-  "collapse f = (f \<circ>\<rightarrow> id)"
+definition collapse :: "('a \<Rightarrow> ('a \<Rightarrow> 'b \<times> 'a) \<times> 'a) \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a"
+  where "collapse f = (f \<circ>\<rightarrow> id)"
 
-definition beyond :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" where
-  "beyond k l = (if l > k then l else 0)"
+definition beyond :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
+  where "beyond k l = (if l > k then l else 0)"
 
-lemma beyond_zero:
-  "beyond k 0 = 0"
+lemma beyond_zero: "beyond k 0 = 0"
   by (simp add: beyond_def)
 
 
-definition (in term_syntax) [code_unfold]: "valterm_emptyset = Code_Evaluation.valtermify ({} :: ('a :: typerep) set)"
-definition (in term_syntax) [code_unfold]: "valtermify_insert x s = Code_Evaluation.valtermify insert {\<cdot>} (x :: ('a :: typerep * _)) {\<cdot>} s"
+definition (in term_syntax) [code_unfold]:
+  "valterm_emptyset = Code_Evaluation.valtermify ({} :: ('a :: typerep) set)"
+
+definition (in term_syntax) [code_unfold]:
+  "valtermify_insert x s = Code_Evaluation.valtermify insert {\<cdot>} (x :: ('a :: typerep * _)) {\<cdot>} s"
 
 instantiation set :: (random) random
 begin
@@ -139,12 +148,17 @@
 primrec random_aux_set
 where
   "random_aux_set 0 j = collapse (Random.select_weight [(1, Pair valterm_emptyset)])"
-| "random_aux_set (Code_Numeral.Suc i) j = collapse (Random.select_weight [(1, Pair valterm_emptyset), (Code_Numeral.Suc i, random j \<circ>\<rightarrow> (%x. random_aux_set i j \<circ>\<rightarrow> (%s. Pair (valtermify_insert x s))))])"
+| "random_aux_set (Code_Numeral.Suc i) j =
+    collapse (Random.select_weight
+      [(1, Pair valterm_emptyset),
+       (Code_Numeral.Suc i,
+        random j \<circ>\<rightarrow> (%x. random_aux_set i j \<circ>\<rightarrow> (%s. Pair (valtermify_insert x s))))])"
 
 lemma [code]:
-  "random_aux_set i j = collapse (Random.select_weight [(1, Pair valterm_emptyset), (i, random j \<circ>\<rightarrow> (%x. random_aux_set (i - 1) j \<circ>\<rightarrow> (%s. Pair (valtermify_insert x s))))])"
+  "random_aux_set i j =
+    collapse (Random.select_weight [(1, Pair valterm_emptyset),
+      (i, random j \<circ>\<rightarrow> (%x. random_aux_set (i - 1) j \<circ>\<rightarrow> (%s. Pair (valtermify_insert x s))))])"
 proof (induct i rule: code_numeral.induct)
-print_cases
   case zero
   show ?case by (subst select_weight_drop_zero[symmetric])
     (simp add: filter.simps random_aux_set.simps[simplified])
@@ -153,9 +167,7 @@
   show ?case by (simp only: random_aux_set.simps(2)[of "i"] Suc_code_numeral_minus_one)
 qed
 
-definition random_set
-where
-  "random_set i = random_aux_set i i" 
+definition "random_set i = random_aux_set i i"
 
 instance ..
 
@@ -190,13 +202,15 @@
 subsection {* The Random-Predicate Monad *} 
 
 fun iter' ::
-  "'a itself => code_numeral => code_numeral => code_numeral * code_numeral => ('a::random) Predicate.pred"
+  "'a itself => code_numeral => code_numeral => code_numeral * code_numeral
+    => ('a::random) Predicate.pred"
 where
   "iter' T nrandom sz seed = (if nrandom = 0 then bot_class.bot else
      let ((x, _), seed') = random sz seed
    in Predicate.Seq (%u. Predicate.Insert x (iter' T (nrandom - 1) sz seed')))"
 
-definition iter :: "code_numeral => code_numeral => code_numeral * code_numeral => ('a::random) Predicate.pred"
+definition iter :: "code_numeral => code_numeral => code_numeral * code_numeral
+  => ('a::random) Predicate.pred"
 where
   "iter nrandom sz seed = iter' (TYPE('a)) nrandom sz seed"
 
@@ -248,22 +262,25 @@
   where "map f P = bind P (single o f)"
 
 hide_fact
-  random_bool_def random_bool_def_raw
-  random_itself_def random_itself_def_raw
-  random_char_def random_char_def_raw
-  random_literal_def random_literal_def_raw
-  random_nat_def random_nat_def_raw
-  random_int_def random_int_def_raw
-  random_fun_lift_def random_fun_lift_def_raw
-  random_fun_def random_fun_def_raw
-  collapse_def collapse_def_raw
-  beyond_def beyond_def_raw beyond_zero
+  random_bool_def
+  random_itself_def
+  random_char_def
+  random_literal_def
+  random_nat_def
+  random_int_def
+  random_fun_lift_def
+  random_fun_def
+  collapse_def
+  beyond_def
+  beyond_zero
   random_aux_rec
 
 hide_const (open) catch_match random collapse beyond random_fun_aux random_fun_lift
 
-hide_fact (open) iter'.simps iter_def empty_def single_def bind_def union_def if_randompred_def iterate_upto_def not_randompred_def Random_def map_def 
+hide_fact (open) iter'.simps iter_def empty_def single_def bind_def union_def
+  if_randompred_def iterate_upto_def not_randompred_def Random_def map_def 
 hide_type (open) randompred
-hide_const (open) iter' iter empty single bind union if_randompred iterate_upto not_randompred Random map
+hide_const (open) iter' iter empty single bind union if_randompred
+  iterate_upto not_randompred Random map
 
 end
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_filter.ML	Sat Mar 17 11:35:18 2012 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_filter.ML	Sat Mar 17 11:57:03 2012 +0100
@@ -867,8 +867,7 @@
             (not (Config.get ctxt ignore_no_atp) andalso
              is_package_def name0) orelse
             exists (fn s => String.isSuffix s name0)
-                   (multi_base_blacklist ctxt ho_atp) orelse
-            String.isSuffix "_def_raw" (* FIXME: crude hack *) name0) then
+                   (multi_base_blacklist ctxt ho_atp)) then
           I
         else
           let
--- a/src/Pure/Isar/locale.ML	Sat Mar 17 11:35:18 2012 +0100
+++ b/src/Pure/Isar/locale.ML	Sat Mar 17 11:57:03 2012 +0100
@@ -406,7 +406,7 @@
       | SOME export => collect_mixins context (name, morph $> export) $> export);
     val morph' = transfer input $> morph $> mixin;
     val notes' =
-      grouped 50 Par_List.map
+      grouped 100 Par_List.map
         (Element.transform_ctxt morph' o Notes o #1) (#notes (the_locale thy name));
   in
     fold_rev (fn elem => fn res => activ_elem (Element.transform_ctxt (transfer res) elem) res)
--- a/src/Pure/Tools/find_consts.ML	Sat Mar 17 11:35:18 2012 +0100
+++ b/src/Pure/Tools/find_consts.ML	Sat Mar 17 11:57:03 2012 +0100
@@ -36,11 +36,11 @@
 fun opt_not f (c as (_, (ty, _))) =
   if is_some (f c) then NONE else SOME (Term.size_of_typ ty);
 
-fun filter_const _ NONE = NONE
-  | filter_const f (SOME (c, r)) =
+fun filter_const _ _ NONE = NONE
+  | filter_const c f (SOME rank) =
       (case f c of
         NONE => NONE
-      | SOME i => SOME (c, Int.min (r, i)));
+      | SOME i => SOME (Int.min (rank, i)));
 
 
 (* pretty results *)
@@ -103,14 +103,13 @@
     val consts = Sign.consts_of thy;
     val (_, consts_tab) = #constants (Consts.dest consts);
     fun eval_entry c =
-      fold filter_const (user_visible consts :: criteria)
-        (SOME (c, low_ranking));
+      fold (filter_const c) (user_visible consts :: criteria) (SOME low_ranking);
 
     val matches =
-      Symtab.fold (cons o eval_entry) consts_tab []
-      |> map_filter I
-      |> sort (rev_order o int_ord o pairself snd)
-      |> map (apsnd fst o fst);
+      Symtab.fold (fn c => (case eval_entry c of NONE => I | SOME rank => cons (rank, c)))
+        consts_tab []
+      |> sort (prod_ord (rev_order o int_ord) (string_ord o pairself fst))
+      |> map (apsnd fst o snd);
   in
     Pretty.big_list "searched for:" (map pretty_criterion raw_criteria) ::
     Pretty.str "" ::
--- a/src/Pure/Tools/find_theorems.ML	Sat Mar 17 11:35:18 2012 +0100
+++ b/src/Pure/Tools/find_theorems.ML	Sat Mar 17 11:57:03 2012 +0100
@@ -420,7 +420,10 @@
       then by the substitution size*)
     fun result_ord (((p0, s0), _), ((p1, s1), _)) =
       prod_ord int_ord int_ord ((p1, s1), (p0, s0));
-  in map_filter eval_filters theorems |> sort result_ord |> map #2 end;
+  in
+    grouped 100 Par_List.map eval_filters theorems
+    |> map_filter I |> sort result_ord |> map #2
+  end;
 
 fun lazy_filter filters =
   let
@@ -581,7 +584,7 @@
     (if null theorems then [Pretty.str "nothing found"]
      else
       [Pretty.str (tally_msg ^ ":"), Pretty.str ""] @
-        map (pretty_theorem ctxt) theorems)
+        grouped 10 Par_List.map (pretty_theorem ctxt) theorems)
   end |> Pretty.chunks |> Pretty.writeln;
 
 fun print_theorems ctxt =