merged
authorhaftmann
Fri, 26 Nov 2010 23:12:01 +0100
changeset 40727 29885c9be6ae
parent 40725 02b3fab953c9 (diff)
parent 40726 16dcfedc4eb7 (current diff)
child 40728 aef83e8fa2a4
child 40732 4375cc6f1a7e
merged
--- a/NEWS	Fri Nov 26 22:33:21 2010 +0100
+++ b/NEWS	Fri Nov 26 23:12:01 2010 +0100
@@ -104,14 +104,13 @@
 avoid confusion with finite sets.  INCOMPATIBILITY.
 
 * Quickcheck's generator for random generation is renamed from "code" to
-"random".  INCOMPATIBILITY. 
+"random".  INCOMPATIBILITY.
 
 * Theory Multiset provides stable quicksort implementation of sort_key.
 
 * Quickcheck now has a configurable time limit which is set to 30 seconds
 by default. This can be changed by adding [timeout = n] to the quickcheck
-command. The time limit for auto quickcheck is still set independently,
-by default to 5 seconds.
+command. The time limit for Auto Quickcheck is still set independently.
 
 * New command 'partial_function' provides basic support for recursive
 function definitions over complete partial orders. Concrete instances
@@ -357,11 +356,26 @@
     (and "ms" and "min" are no longer supported)
     INCOMPATIBILITY.
 
+* Metis and Meson now have configuration options "meson_trace", "metis_trace",
+  and "metis_verbose" that can be enabled to diagnose these tools. E.g.
+
+    using [[metis_trace = true]]
+
 * Nitpick:
   - Renamed options:
     nitpick [timeout = 77 s] ~> nitpick [timeout = 77]
     nitpick [tac_timeout = 777 ms] ~> nitpick [tac_timeout = 0.777]
     INCOMPATIBILITY.
+  - Now requires Kodkodi 1.2.9. INCOMPATIBILITY.
+  - Added support for partial quotient types.
+  - Added local versions of the "Nitpick.register_xxx" functions.
+  - Added "whack" option.
+  - Allow registration of quotient types as codatatypes.
+  - Improved "merge_type_vars" option to merge more types.
+  - Removed unsound "fast_descrs" option.
+  - Added custom symmetry breaking for datatypes, making it possible to reach
+    higher cardinalities.
+  - Prevent the expansion of too large definitions.
 
 * Changed SMT configuration options:
   - Renamed:
@@ -650,7 +664,7 @@
 
 Tracing is then active for all invocations of the simplifier in
 subsequent goal refinement steps. Tracing may also still be enabled or
-disabled via the ProofGeneral settings menu.
+disabled via the Proof General settings menu.
 
 * Separate commands 'hide_class', 'hide_type', 'hide_const',
 'hide_fact' replace the former 'hide' KIND command.  Minor
--- a/src/HOL/Tools/Meson/meson.ML	Fri Nov 26 22:33:21 2010 +0100
+++ b/src/HOL/Tools/Meson/meson.ML	Fri Nov 26 23:12:01 2010 +0100
@@ -46,7 +46,7 @@
 structure Meson : MESON =
 struct
 
-val (trace, trace_setup) = Attrib.config_bool "trace_meson" (K false)
+val (trace, trace_setup) = Attrib.config_bool "meson_trace" (K false)
 
 fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else ()
 
--- a/src/HOL/Tools/Metis/metis_reconstruct.ML	Fri Nov 26 22:33:21 2010 +0100
+++ b/src/HOL/Tools/Metis/metis_reconstruct.ML	Fri Nov 26 23:12:01 2010 +0100
@@ -31,8 +31,8 @@
 
 open Metis_Translate
 
-val (trace, trace_setup) = Attrib.config_bool "trace_metis" (K false)
-val (verbose, verbose_setup) = Attrib.config_bool "verbose_metis" (K true)
+val (trace, trace_setup) = Attrib.config_bool "metis_trace" (K false)
+val (verbose, verbose_setup) = Attrib.config_bool "metis_verbose" (K true)
 
 fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else ()
 fun verbose_warning ctxt msg =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Fri Nov 26 22:33:21 2010 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Fri Nov 26 23:12:01 2010 +0100
@@ -440,7 +440,7 @@
 val smt_iter_timeout_divisor = 2
 val smt_monomorph_limit = 4
 
-fun smt_filter_loop ({debug, verbose, timeout, ...} : params) remote state i =
+fun smt_filter_loop ({verbose, timeout, ...} : params) remote state i =
   let
     val ctxt = Proof.context_of state
     fun iter timeout iter_num outcome0 msecs_so_far facts =
@@ -534,6 +534,7 @@
     val smt_fact = Option.map (apsnd (Thm.transfer thy)) o try dest_Untranslated
     val {outcome, used_facts, run_time_in_msecs} =
       smt_filter_loop params remote state subgoal (map_filter smt_fact facts)
+    val (chained_lemmas, other_lemmas) = split_used_facts (map fst used_facts)
     val outcome = outcome |> Option.map failure_from_smt_failure
     val message =
       case outcome of
@@ -547,8 +548,9 @@
         in
           try_command_line (proof_banner auto)
                            (apply_on_subgoal subgoal subgoal_count ^
-                            command_call method (map (fst o fst) used_facts)) ^
-          minimize_line minimize_command (map (fst o fst) used_facts)
+                            command_call method (map fst other_lemmas)) ^
+          minimize_line minimize_command
+                        (map fst (other_lemmas @ chained_lemmas))
         end
       | SOME failure => string_for_failure "SMT solver" failure
   in
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_atp_reconstruct.ML	Fri Nov 26 22:33:21 2010 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_atp_reconstruct.ML	Fri Nov 26 23:12:01 2010 +0100
@@ -24,7 +24,9 @@
   val command_call : string -> string list -> string
   val try_command_line : string -> string -> string
   val minimize_line : ('a list -> string) -> 'a list -> string
-  val metis_proof_text : metis_params -> text_result
+  val split_used_facts :
+    (string * locality) list
+    -> (string * locality) list * (string * locality) list
   val isar_proof_text : isar_params -> metis_params -> text_result
   val proof_text : bool -> isar_params -> metis_params -> text_result
 end;
@@ -159,15 +161,15 @@
 fun used_facts_in_tstplike_proof fact_names =
   atp_proof_from_tstplike_string #> rpair [] #-> fold (add_fact fact_names)
 
-fun used_facts fact_names =
-  used_facts_in_tstplike_proof fact_names
-  #> List.partition (curry (op =) Chained o snd)
+val split_used_facts =
+  List.partition (curry (op =) Chained o snd)
   #> pairself (sort_distinct (string_ord o pairself fst))
 
 fun metis_proof_text (banner, full_types, minimize_command,
                       tstplike_proof, fact_names, goal, i) =
   let
-    val (chained_lemmas, other_lemmas) = used_facts fact_names tstplike_proof
+    val (chained_lemmas, other_lemmas) =
+      split_used_facts (used_facts_in_tstplike_proof fact_names tstplike_proof)
     val n = Logic.count_prems (prop_of goal)
   in
     (metis_line banner full_types i n (map fst other_lemmas) ^
@@ -912,14 +914,14 @@
   in do_proof end
 
 fun isar_proof_text (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
-                    (other_params as (_, full_types, _, tstplike_proof,
+                    (metis_params as (_, full_types, _, tstplike_proof,
                                       fact_names, goal, i)) =
   let
     val (params, hyp_ts, concl_t) = strip_subgoal goal i
     val frees = fold Term.add_frees (concl_t :: hyp_ts) []
     val tfrees = fold Term.add_tfrees (concl_t :: hyp_ts) []
     val n = Logic.count_prems (prop_of goal)
-    val (one_line_proof, lemma_names) = metis_proof_text other_params
+    val (one_line_proof, lemma_names) = metis_proof_text metis_params
     fun isar_proof_for () =
       case isar_proof_from_tstplike_proof pool ctxt full_types tfrees
                isar_shrink_factor tstplike_proof conjecture_shape fact_names
@@ -940,8 +942,8 @@
         |> the_default "\nWarning: The Isar proof construction failed."
   in (one_line_proof ^ isar_proof, lemma_names) end
 
-fun proof_text isar_proof isar_params other_params =
+fun proof_text isar_proof isar_params metis_params =
   (if isar_proof then isar_proof_text isar_params else metis_proof_text)
-      other_params
+      metis_params
 
 end;