# HG changeset patch # User huffman # Date 1267755532 28800 # Node ID 374c638a796e9729aefa99c31d332c5b1aa19183 # Parent ee5df989b7c4e540f645e88002efa14ef76ad65a# Parent 4af56a1c4c7defd3951a4634d0ebd20c61ad92cf merged diff -r ee5df989b7c4 -r 374c638a796e Admin/isatest/settings/at-poly-test --- a/Admin/isatest/settings/at-poly-test Thu Mar 04 10:01:39 2010 -0800 +++ b/Admin/isatest/settings/at-poly-test Thu Mar 04 18:18:52 2010 -0800 @@ -1,6 +1,6 @@ # -*- shell-script -*- :mode=shellscript: - POLYML_HOME="/home/polyml/polyml-5.3.0" + POLYML_HOME="/home/polyml/polyml-svn" ML_SYSTEM="polyml-5.3.0" ML_PLATFORM="x86-linux" ML_HOME="$POLYML_HOME/$ML_PLATFORM" diff -r ee5df989b7c4 -r 374c638a796e README_REPOSITORY --- a/README_REPOSITORY Thu Mar 04 10:01:39 2010 -0800 +++ b/README_REPOSITORY Thu Mar 04 18:18:52 2010 -0800 @@ -83,6 +83,8 @@ See also the fine documentation for further details, especially the book http://hgbook.red-bean.com/ +There is also a nice tutorial at http://hginit.com/ + Shared pull/push access ----------------------- diff -r ee5df989b7c4 -r 374c638a796e src/HOL/Auth/Message.thy --- a/src/HOL/Auth/Message.thy Thu Mar 04 10:01:39 2010 -0800 +++ b/src/HOL/Auth/Message.thy Thu Mar 04 18:18:52 2010 -0800 @@ -236,7 +236,7 @@ by blast lemma parts_subset_iff [simp]: "(parts G \ parts H) = (G \ parts H)" -by (metis equalityE parts_idem parts_increasing parts_mono subset_trans) +by (metis parts_idem parts_increasing parts_mono subset_trans) lemma parts_trans: "[| X\ parts G; G \ parts H |] ==> X\ parts H" by (drule parts_mono, blast) @@ -611,7 +611,7 @@ by blast lemma synth_subset_iff [simp]: "(synth G \ synth H) = (G \ synth H)" -by (metis equalityE subset_trans synth_idem synth_increasing synth_mono) +by (metis subset_trans synth_idem synth_increasing synth_mono) lemma synth_trans: "[| X\ synth G; G \ synth H |] ==> X\ synth H" by (drule synth_mono, blast) @@ -674,8 +674,7 @@ lemma parts_insert_subset_Un: "X\ G ==> parts(insert X H) \ parts G \ parts H" by (metis UnCI Un_upper2 insert_subset parts_Un parts_mono) -text{*More specifically for Fake. Very occasionally we could do with a version - of the form @{term"parts{X} \ synth (analz H) \ parts H"} *} +text{*More specifically for Fake. See also @{text Fake_parts_sing} below *} lemma Fake_parts_insert: "X \ synth (analz H) ==> parts (insert X H) \ synth (analz H) \ parts H" @@ -884,17 +883,17 @@ lemma Fake_analz_eq [simp]: "X \ synth(analz H) ==> synth (analz (insert X H)) = synth (analz H)" -by (metis Fake_analz_insert Un_absorb Un_absorb1 Un_commute equalityI +by (metis Fake_analz_insert Un_absorb Un_absorb1 Un_commute subset_insertI synth_analz_mono synth_increasing synth_subset_iff) text{*Two generalizations of @{text analz_insert_eq}*} lemma gen_analz_insert_eq [rule_format]: - "X \ analz H ==> ALL G. H \ G --> analz (insert X G) = analz G"; + "X \ analz H ==> ALL G. H \ G --> analz (insert X G) = analz G" by (blast intro: analz_cut analz_insertI analz_mono [THEN [2] rev_subsetD]) lemma synth_analz_insert_eq [rule_format]: "X \ synth (analz H) - ==> ALL G. H \ G --> (Key K \ analz (insert X G)) = (Key K \ analz G)"; + ==> ALL G. H \ G --> (Key K \ analz (insert X G)) = (Key K \ analz G)" apply (erule synth.induct) apply (simp_all add: gen_analz_insert_eq subset_trans [OF _ subset_insertI]) done diff -r ee5df989b7c4 -r 374c638a796e src/HOL/Map.thy --- a/src/HOL/Map.thy Thu Mar 04 10:01:39 2010 -0800 +++ b/src/HOL/Map.thy Thu Mar 04 18:18:52 2010 -0800 @@ -11,11 +11,11 @@ imports List begin -types ('a,'b) "~=>" = "'a => 'b option" (infixr "~=>" 0) +types ('a,'b) "map" = "'a => 'b option" (infixr "~=>" 0) translations (type) "'a ~=> 'b" <= (type) "'a => 'b option" type_notation (xsymbols) - "~=>" (infixr "\" 0) + "map" (infixr "\" 0) abbreviation empty :: "'a ~=> 'b" where @@ -650,5 +650,52 @@ thus "\v. f = [x \ v]" by blast qed + +subsection {* Various *} + +lemma set_map_of_compr: + assumes distinct: "distinct (map fst xs)" + shows "set xs = {(k, v). map_of xs k = Some v}" +using assms proof (induct xs) + case Nil then show ?case by simp +next + case (Cons x xs) + obtain k v where "x = (k, v)" by (cases x) blast + with Cons.prems have "k \ dom (map_of xs)" + by (simp add: dom_map_of_conv_image_fst) + then have *: "insert (k, v) {(k, v). map_of xs k = Some v} = + {(k', v'). (map_of xs(k \ v)) k' = Some v'}" + by (auto split: if_splits) + from Cons have "set xs = {(k, v). map_of xs k = Some v}" by simp + with * `x = (k, v)` show ?case by simp +qed + +lemma map_of_inject_set: + assumes distinct: "distinct (map fst xs)" "distinct (map fst ys)" + shows "map_of xs = map_of ys \ set xs = set ys" (is "?lhs \ ?rhs") +proof + assume ?lhs + moreover from `distinct (map fst xs)` have "set xs = {(k, v). map_of xs k = Some v}" + by (rule set_map_of_compr) + moreover from `distinct (map fst ys)` have "set ys = {(k, v). map_of ys k = Some v}" + by (rule set_map_of_compr) + ultimately show ?rhs by simp +next + assume ?rhs show ?lhs proof + fix k + show "map_of xs k = map_of ys k" proof (cases "map_of xs k") + case None + moreover with `?rhs` have "map_of ys k = None" + by (simp add: map_of_eq_None_iff) + ultimately show ?thesis by simp + next + case (Some v) + moreover with distinct `?rhs` have "map_of ys k = Some v" + by simp + ultimately show ?thesis by simp + qed + qed +qed + end diff -r ee5df989b7c4 -r 374c638a796e src/HOL/Tools/ATP_Manager/atp_manager.ML --- a/src/HOL/Tools/ATP_Manager/atp_manager.ML Thu Mar 04 10:01:39 2010 -0800 +++ b/src/HOL/Tools/ATP_Manager/atp_manager.ML Thu Mar 04 18:18:52 2010 -0800 @@ -89,7 +89,7 @@ fun unregister message thread = Synchronized.change global_state (fn state as {manager, timeout_heap, active, cancelling, messages, store} => (case lookup_thread active thread of - SOME (birth_time, _, description) => + SOME (_, _, description) => let val active' = delete_thread thread active; val cancelling' = (thread, (Time.now (), description)) :: cancelling; @@ -267,7 +267,7 @@ "Empty clause: Try this command: " ^ Markup.markup Markup.sendback "apply metis" | ERROR msg => ("Error: " ^ msg); val _ = unregister message (Thread.self ()); - in () end) + in () end); in () end); diff -r ee5df989b7c4 -r 374c638a796e src/HOL/Tools/ATP_Manager/atp_wrapper.ML --- a/src/HOL/Tools/ATP_Manager/atp_wrapper.ML Thu Mar 04 10:01:39 2010 -0800 +++ b/src/HOL/Tools/ATP_Manager/atp_wrapper.ML Thu Mar 04 18:18:52 2010 -0800 @@ -111,14 +111,14 @@ |> Exn.release |> tap (after path); -fun external_prover relevance_filter prepare write cmd args find_failure produce_answer - axiom_clauses filtered_clauses name subgoalno goal = +fun external_prover relevance_filter prepare write cmd args produce_answer name + ({with_full_types, subgoal, goal, axiom_clauses, filtered_clauses}: problem) = let (* get clauses and prepare them for writing *) val (ctxt, (chain_ths, th)) = goal; val thy = ProofContext.theory_of ctxt; val chain_ths = map (Thm.put_name_hint Res_Reconstruct.chained_hint) chain_ths; - val goal_cls = #1 (Res_Axioms.neg_conjecture_clauses ctxt th subgoalno); + val goal_cls = #1 (Res_Axioms.neg_conjecture_clauses ctxt th subgoal); val the_filtered_clauses = (case filtered_clauses of NONE => relevance_filter goal goal_cls @@ -138,8 +138,8 @@ Path.basic (problem_prefix' ^ serial_string () ^ "_" ^ string_of_int nr) in if destdir' = "" then File.tmp_path probfile - else if File.exists (Path.explode (destdir')) - then Path.append (Path.explode (destdir')) probfile + else if File.exists (Path.explode destdir') + then Path.append (Path.explode destdir') probfile else error ("No such directory: " ^ destdir') end; @@ -167,7 +167,7 @@ if Config.get ctxt measure_runtime then split_time s else (s, 0) fun run_on probfile = if File.exists cmd then - write probfile clauses + write with_full_types probfile clauses |> pair (apfst split_time' (bash_output (cmd_line probfile))) else error ("Bad executable: " ^ Path.implode cmd); @@ -178,16 +178,16 @@ 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); + with_path cleanup export run_on (prob_pathname subgoal); (* check for success and print out some information on failure *) - val failure = find_failure proof; + val failure = Res_Reconstruct.find_failure proof; val success = rc = 0 andalso is_none failure; val (message, real_thm_names) = 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)); + (produce_answer name (proof, thm_names, conj_pos, ctxt, th, subgoal)); in {success = success, message = message, theorem_names = real_thm_names, runtime = time, proof = proof, @@ -201,22 +201,17 @@ let val {max_new_clauses, insert_theory_const, emit_structured_proof, command, arguments} = prover_config; - val {with_full_types, subgoal, goal, axiom_clauses, filtered_clauses} = problem; in external_prover (Res_ATP.get_relevant max_new_clauses insert_theory_const) (Res_ATP.prepare_clauses false) - (Res_HOL_Clause.tptp_write_file with_full_types) + Res_HOL_Clause.tptp_write_file command (arguments timeout) - Res_Reconstruct.find_failure (if emit_structured_proof then Res_Reconstruct.structured_proof else Res_Reconstruct.lemma_list false) - axiom_clauses - filtered_clauses name - subgoal - goal + problem end; fun tptp_prover (name, config) = (name, gen_tptp_prover (name, config)); @@ -276,22 +271,17 @@ fun gen_dfg_prover (name, prover_config: prover_config) timeout problem = let - val {max_new_clauses, insert_theory_const, command, arguments, ...} = prover_config - val {with_full_types, subgoal, goal, axiom_clauses, filtered_clauses} = problem + val {max_new_clauses, insert_theory_const, command, arguments, ...} = prover_config; in external_prover (Res_ATP.get_relevant max_new_clauses insert_theory_const) (Res_ATP.prepare_clauses true) - (Res_HOL_Clause.dfg_write_file with_full_types) + Res_HOL_Clause.dfg_write_file command (arguments timeout) - Res_Reconstruct.find_failure (Res_Reconstruct.lemma_list true) - axiom_clauses - filtered_clauses name - subgoal - goal + problem end; fun dfg_prover (name, config) = (name, gen_dfg_prover (name, config)); diff -r ee5df989b7c4 -r 374c638a796e src/HOL/Tools/metis_tools.ML --- a/src/HOL/Tools/metis_tools.ML Thu Mar 04 10:01:39 2010 -0800 +++ b/src/HOL/Tools/metis_tools.ML Thu Mar 04 18:18:52 2010 -0800 @@ -714,12 +714,12 @@ let val _ = trace_msg (fn () => "Metis called with theorems " ^ cat_lines (map (Display.string_of_thm ctxt) ths)) in - if exists_type Res_Axioms.type_has_empty_sort (prop_of st0) - then (error "metis: Proof state contains the empty sort"; Seq.empty) - else - (Meson.MESON Res_Axioms.neg_clausify - (fn cls => resolve_tac (FOL_SOLVE mode ctxt cls ths) 1) ctxt i - THEN Res_Axioms.expand_defs_tac st0) st0 + if exists_type Res_Axioms.type_has_topsort (prop_of st0) + then raise METIS "Metis: Proof state contains the universal sort {}" + else + (Meson.MESON Res_Axioms.neg_clausify + (fn cls => resolve_tac (FOL_SOLVE mode ctxt cls ths) 1) ctxt i + THEN Res_Axioms.expand_defs_tac st0) st0 end handle METIS s => (warning ("Metis: " ^ s); Seq.empty); @@ -734,7 +734,7 @@ type_lits_setup #> 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 @{binding metisFT} FT "METIS with fully-typed translation" #> Method.setup @{binding finish_clausify} (Scan.succeed (K (SIMPLE_METHOD (Res_Axioms.expand_defs_tac refl)))) "cleanup after conversion to clauses"; diff -r ee5df989b7c4 -r 374c638a796e src/HOL/Tools/res_axioms.ML --- a/src/HOL/Tools/res_axioms.ML Thu Mar 04 10:01:39 2010 -0800 +++ b/src/HOL/Tools/res_axioms.ML Thu Mar 04 18:18:52 2010 -0800 @@ -12,7 +12,7 @@ val pairname: thm -> string * thm val multi_base_blacklist: string list val bad_for_atp: thm -> bool - val type_has_empty_sort: typ -> bool + val type_has_topsort: typ -> bool val cnf_rules_pairs: theory -> (string * thm) list -> (thm * (string * int)) list val neg_clausify: thm list -> thm list val expand_defs_tac: thm -> tactic @@ -31,10 +31,10 @@ fun freeze_thm th = #1 (Drule.legacy_freeze_thaw th); -fun type_has_empty_sort (TFree (_, [])) = true - | type_has_empty_sort (TVar (_, [])) = true - | type_has_empty_sort (Type (_, Ts)) = exists type_has_empty_sort Ts - | type_has_empty_sort _ = false; +val type_has_topsort = Term.exists_subtype + (fn TFree (_, []) => true + | TVar (_, []) => true + | _ => false); (**** Transformation of Elimination Rules into First-Order Formulas****) @@ -321,7 +321,7 @@ fun bad_for_atp th = too_complex (prop_of th) - orelse exists_type type_has_empty_sort (prop_of th) + orelse exists_type type_has_topsort (prop_of th) orelse is_strange_thm th; val multi_base_blacklist =