# HG changeset patch # User blanchet # Date 1342593843 -7200 # Node ID 6cf5e58f11850d9768c75de31ffab02146f34ab5 # Parent e5c5037a3104d7370730f658daaf31f2bb3625a6 more implementation work on MaSh diff -r e5c5037a3104 -r 6cf5e58f1185 src/HOL/TPTP/ATP_Theory_Export.thy --- a/src/HOL/TPTP/ATP_Theory_Export.thy Wed Jul 18 08:44:03 2012 +0200 +++ b/src/HOL/TPTP/ATP_Theory_Export.thy Wed Jul 18 08:44:03 2012 +0200 @@ -5,7 +5,7 @@ header {* ATP Theory Exporter *} theory ATP_Theory_Export -imports (* ### Complex_Main *) "~~/src/HOL/Sledgehammer2d" +imports Complex_Main uses "atp_theory_export.ML" begin @@ -15,7 +15,7 @@ *} ML {* -val do_it = true; (* switch to "true" to generate the files *) +val do_it = false; (* switch to "true" to generate the files *) val thy = @{theory List}; val ctxt = @{context} *} diff -r e5c5037a3104 -r 6cf5e58f1185 src/HOL/TPTP/mash_export.ML --- a/src/HOL/TPTP/mash_export.ML Wed Jul 18 08:44:03 2012 +0200 +++ b/src/HOL/TPTP/mash_export.ML Wed Jul 18 08:44:03 2012 +0200 @@ -33,7 +33,7 @@ fun do_fact ((_, (_, status)), th) prevs = let val name = Thm.get_name_hint th - val feats = features_of thy (status, th) + val feats = features_of thy status [prop_of th] val deps = isabelle_dependencies_of all_names th val kind = Thm.legacy_get_kind th val name = fact_name_of name diff -r e5c5037a3104 -r 6cf5e58f1185 src/HOL/Tools/ATP/atp_problem_generate.ML --- a/src/HOL/Tools/ATP/atp_problem_generate.ML Wed Jul 18 08:44:03 2012 +0200 +++ b/src/HOL/Tools/ATP/atp_problem_generate.ML Wed Jul 18 08:44:03 2012 +0200 @@ -94,7 +94,7 @@ val is_type_enc_sound : type_enc -> bool val type_enc_from_string : strictness -> string -> type_enc val adjust_type_enc : atp_format -> type_enc -> type_enc - val has_no_lambdas : term -> bool + val is_lambda_free : term -> bool val mk_aconns : connective -> ('a, 'b, 'c, 'd) formula list -> ('a, 'b, 'c, 'd) formula val unmangled_const : string -> string * (string, 'b) ho_term list @@ -699,22 +699,22 @@ (if is_type_enc_sound type_enc then Tags else Guards) stuff | adjust_type_enc _ type_enc = type_enc -fun has_no_lambdas t = +fun is_lambda_free t = case t of - @{const Not} $ t1 => has_no_lambdas t1 - | Const (@{const_name All}, _) $ Abs (_, _, t') => has_no_lambdas t' - | Const (@{const_name All}, _) $ t1 => has_no_lambdas t1 - | Const (@{const_name Ex}, _) $ Abs (_, _, t') => has_no_lambdas t' - | Const (@{const_name Ex}, _) $ t1 => has_no_lambdas t1 - | @{const HOL.conj} $ t1 $ t2 => has_no_lambdas t1 andalso has_no_lambdas t2 - | @{const HOL.disj} $ t1 $ t2 => has_no_lambdas t1 andalso has_no_lambdas t2 - | @{const HOL.implies} $ t1 $ t2 => has_no_lambdas t1 andalso has_no_lambdas t2 + @{const Not} $ t1 => is_lambda_free t1 + | Const (@{const_name All}, _) $ Abs (_, _, t') => is_lambda_free t' + | Const (@{const_name All}, _) $ t1 => is_lambda_free t1 + | Const (@{const_name Ex}, _) $ Abs (_, _, t') => is_lambda_free t' + | Const (@{const_name Ex}, _) $ t1 => is_lambda_free t1 + | @{const HOL.conj} $ t1 $ t2 => is_lambda_free t1 andalso is_lambda_free t2 + | @{const HOL.disj} $ t1 $ t2 => is_lambda_free t1 andalso is_lambda_free t2 + | @{const HOL.implies} $ t1 $ t2 => is_lambda_free t1 andalso is_lambda_free t2 | Const (@{const_name HOL.eq}, Type (_, [@{typ bool}, _])) $ t1 $ t2 => - has_no_lambdas t1 andalso has_no_lambdas t2 + is_lambda_free t1 andalso is_lambda_free t2 | _ => not (exists_subterm (fn Abs _ => true | _ => false) t) fun simple_translate_lambdas do_lambdas ctxt t = - if has_no_lambdas t then + if is_lambda_free t then t else let diff -r e5c5037a3104 -r 6cf5e58f1185 src/HOL/Tools/Sledgehammer/sledgehammer_filter_mash.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer_filter_mash.ML Wed Jul 18 08:44:03 2012 +0200 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_filter_mash.ML Wed Jul 18 08:44:03 2012 +0200 @@ -22,7 +22,7 @@ val thms_by_thy : ('a * thm) list -> (string * thm list) list val has_thy : theory -> thm -> bool val parent_facts : (string * thm list) list -> theory -> string list - val features_of : theory -> status * thm -> string list + val features_of : theory -> status -> term list -> string list val isabelle_dependencies_of : string list -> thm -> string list val goal_of_thm : theory -> thm -> thm val run_prover : Proof.context -> params -> fact list -> thm -> prover_result @@ -32,10 +32,14 @@ Proof.context -> theory -> bool -> string -> unit val generate_atp_dependencies : Proof.context -> params -> theory -> bool -> string -> unit + val mash_RESET : unit -> unit + val mash_ADD : string -> string list -> string list -> string list -> unit + val mash_DEL : string list -> string list -> unit + val mash_SUGGEST : string list -> string list -> string list val mash_reset : unit -> unit val mash_can_suggest_facts : unit -> bool val mash_suggest_facts : - theory -> params -> string -> int -> term list -> term -> fact list + Proof.context -> params -> string -> int -> term list -> term -> fact list -> fact list * fact list val mash_can_learn_thy : theory -> bool val mash_learn_thy : theory -> real -> unit @@ -59,25 +63,11 @@ val model_file = "model" val state_file = "state" -fun mash_path file = getenv "ISABELLE_HOME_USER" ^ "/" ^ mash_dir ^ "/" ^ file - - -(*** Low-level communication with MaSh ***) - -fun mash_RESET () = - warning "MaSh RESET" +fun mk_path file = + getenv "ISABELLE_HOME_USER" ^ "/" ^ mash_dir ^ "/" ^ file + |> Path.explode -fun mash_ADD fact access feats deps = - warning ("MaSh ADD " ^ fact ^ ": " ^ space_implode " " access ^ "; " ^ - space_implode " " feats ^ "; " ^ space_implode " " deps) - -fun mash_DEL fact = - warning ("MaSh DEL " ^ fact) - -fun mash_SUGGEST access feats = - warning ("MaSh SUGGEST " ^ space_implode " " access ^ "; " ^ - space_implode " " feats) - +val fresh_fact_prefix = Long_Name.separator (*** Isabelle helpers ***) @@ -145,7 +135,7 @@ end -fun interesting_terms_types_and_classes term_max_depth type_max_depth t = +fun interesting_terms_types_and_classes term_max_depth type_max_depth ts = let val bad_types = [@{type_name prop}, @{type_name bool}, @{type_name fun}] val bad_consts = atp_widely_irrelevant_consts @@ -180,10 +170,10 @@ | _ => I) #> fold add_patterns args end - in [] |> add_patterns t |> sort string_ord end + in [] |> fold add_patterns ts |> sort string_ord end fun is_likely_tautology th = - null (interesting_terms_types_and_classes 0 ~1 (prop_of th)) andalso + null (interesting_terms_types_and_classes 0 ~1 [prop_of th]) andalso not (Thm.eq_thm_prop (@{thm ext}, th)) fun is_too_meta thy th = @@ -227,22 +217,20 @@ val type_max_depth = 1 (* TODO: Generate type classes for types? *) -fun features_of thy (status, th) = - let val t = Thm.prop_of th in - thy_name_of (thy_name_of_thm th) :: - interesting_terms_types_and_classes term_max_depth type_max_depth t - |> not (has_no_lambdas t) ? cons "lambdas" - |> exists_Const is_exists t ? cons "skolems" - |> not (is_fo_term thy t) ? cons "ho" - |> (case status of - General => I - | Induction => cons "induction" - | Intro => cons "intro" - | Inductive => cons "inductive" - | Elim => cons "elim" - | Simp => cons "simp" - | Def => cons "def") - end +fun features_of thy status ts = + thy_name_of (Context.theory_name thy) :: + interesting_terms_types_and_classes term_max_depth type_max_depth ts + |> exists (not o is_lambda_free) ts ? cons "lambdas" + |> exists (exists_Const is_exists) ts ? cons "skolems" + |> exists (not o is_fo_term thy) ts ? cons "ho" + |> (case status of + General => I + | Induction => cons "induction" + | Intro => cons "intro" + | Inductive => cons "inductive" + | Elim => cons "elim" + | Simp => cons "simp" + | Def => cons "def") fun isabelle_dependencies_of all_facts = thms_in_proof (SOME all_facts) @@ -303,7 +291,7 @@ fun do_fact ((_, (_, status)), th) = let val name = Thm.get_name_hint th - val feats = features_of thy (status, th) + val feats = features_of (theory_of_thm th) status [prop_of th] val s = fact_name_of name ^ ": " ^ space_implode " " feats ^ "\n" in File.append path s end in List.app do_fact facts end @@ -375,68 +363,106 @@ in List.app do_thm ths end +(*** Low-level communication with MaSh ***) + +fun mash_RESET () = (warning "MaSh RESET"; File.rm (mk_path model_file)) + +fun mash_ADD fact access feats deps = + warning ("MaSh ADD " ^ fact ^ ": " ^ space_implode " " access ^ "; " ^ + space_implode " " feats ^ "; " ^ space_implode " " deps) + +fun mash_DEL facts feats = + warning ("MaSh DEL " ^ space_implode " " facts ^ "; " ^ + space_implode " " feats) + +fun mash_SUGGEST access feats = + (warning ("MaSh SUGGEST " ^ space_implode " " access ^ "; " ^ + space_implode " " feats); + []) + + (*** High-level communication with MaSh ***) type mash_state = - {completed_thys : unit Symtab.table, + {fresh : int, + completed_thys : unit Symtab.table, thy_facts : string list Symtab.table} val mash_zero = - {completed_thys = Symtab.empty, + {fresh = 0, + completed_thys = Symtab.empty, thy_facts = Symtab.empty} local fun mash_load (state as (true, _)) = state | mash_load _ = - (true, - case mash_path state_file |> Path.explode |> File.read_lines of - [] => mash_zero - | comp_line :: facts_lines => - let - fun comp_thys_of_line comp_line = - Symtab.make (comp_line |> space_explode " " |> map (rpair ())) - fun add_facts_line line = - case space_explode " " line of - thy :: facts => Symtab.update_new (thy, facts) - | _ => I (* shouldn't happen *) - in - {completed_thys = comp_thys_of_line comp_line, - thy_facts = fold add_facts_line facts_lines Symtab.empty} - end) + let + val path = mk_path state_file + val _ = Isabelle_System.mkdir (path |> Path.dir) + in + (true, + case try File.read_lines path of + SOME (fresh_line :: comp_line :: facts_lines) => + let + fun comp_thys_of_line comp_line = + Symtab.make (comp_line |> space_explode " " |> map (rpair ())) + fun add_facts_line line = + case space_explode " " line of + thy :: facts => Symtab.update_new (thy, facts) + | _ => I (* shouldn't happen *) + in + {fresh = Int.fromString fresh_line |> the_default 0, + completed_thys = comp_thys_of_line comp_line, + thy_facts = fold add_facts_line facts_lines Symtab.empty} + end + | _ => mash_zero) + end -fun mash_save ({completed_thys, thy_facts} : mash_state) = +fun mash_save ({fresh, completed_thys, thy_facts} : mash_state) = let - val path = mash_path state_file |> Path.explode + val path = mk_path state_file val comp_line = (completed_thys |> Symtab.keys |> space_implode " ") ^ "\n" fun fact_line_for (thy, facts) = space_implode " " (thy :: facts) ^ "\n" in - File.write path comp_line; + File.write path (string_of_int fresh ^ "\n" ^ comp_line); Symtab.fold (fn thy_fact => fn () => File.append path (fact_line_for thy_fact)) thy_facts end -val state = - Synchronized.var "Sledgehammer_Filter_MaSh.state" (false, mash_zero) +val global_state = + Synchronized.var "Sledgehammer_Filter_MaSh.global_state" (false, mash_zero) in fun mash_change f = - Synchronized.change state (apsnd (tap mash_save o f) o mash_load) + Synchronized.change global_state (mash_load ##> (f #> tap mash_save)) + +fun mash_peek f = Synchronized.change global_state (mash_load ##> tap f) -fun mash_value () = Synchronized.change_result state (`snd o mash_load) +fun mash_value () = Synchronized.change_result global_state (mash_load #> `snd) + +fun mash_reset () = + Synchronized.change global_state (fn _ => + (mash_RESET (); File.rm (mk_path state_file); (true, mash_zero))) end -fun mash_reset () = mash_change (fn _ => (mash_RESET (); mash_zero)) - fun mash_can_suggest_facts () = not (Symtab.is_empty (#thy_facts (mash_value ()))) +fun accessibility_of thy thy_facts = + let + val _ = 0 + in + [] (*###*) + end + fun mash_suggest_facts ctxt params prover max_facts hyp_ts concl_t facts = let - val access = [] - val feats = [] + val thy = Proof_Context.theory_of ctxt + val access = accessibility_of thy (#thy_facts (mash_value ())) + val feats = features_of thy General (concl_t :: hyp_ts) val suggs = mash_SUGGEST access feats in (facts, []) end @@ -445,14 +471,20 @@ (Context.theory_name thy)) fun mash_learn_thy thy timeout = () +(* ### *) -fun mash_learn_proof thy t ths = () -(*### - let - in - mash_ADD - end -*) +fun mash_learn_proof thy t ths = + mash_change (fn {fresh, completed_thys, thy_facts} => + let + val fact = fresh_fact_prefix ^ string_of_int fresh + val access = accessibility_of thy thy_facts + val feats = features_of thy General [t] + val deps = ths |> map (fact_name_of o Thm.get_name_hint) + in + mash_ADD fact access feats deps; + {fresh = fresh + 1, completed_thys = completed_thys, + thy_facts = thy_facts} + end) fun relevant_facts ctxt params prover max_facts ({add, only, ...} : fact_override) hyp_ts concl_t facts = @@ -472,7 +504,7 @@ concl_t facts val (mash_facts, mash_rejected) = mash_suggest_facts ctxt params prover max_facts hyp_ts concl_t facts - val mesh_facts = iter_facts + val mesh_facts = iter_facts (* ### *) in mesh_facts |> not (null add_ths) ? prepend_facts add_ths diff -r e5c5037a3104 -r 6cf5e58f1185 src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML --- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML Wed Jul 18 08:44:03 2012 +0200 +++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML Wed Jul 18 08:44:03 2012 +0200 @@ -39,12 +39,6 @@ val mash_resetN = "mash_reset" val mash_learnN = "mash_learn" -(* experimental *) -val mash_RESET_N = "mash_RESET" -val mash_ADD_N = "mash_ADD" -val mash_DEL_N = "mash_DEL" -val mash_SUGGEST_N = "mash_SUGGEST" - val auto = Unsynchronized.ref false val _ = @@ -387,14 +381,6 @@ mash_reset () else if subcommand = mash_learnN then () (* TODO: implement *) - else if subcommand = mash_RESET_N then - () (* TODO: implement *) - else if subcommand = mash_ADD_N then - () (* TODO: implement *) - else if subcommand = mash_DEL_N then - () (* TODO: implement *) - else if subcommand = mash_SUGGEST_N then - () (* TODO: implement *) else error ("Unknown subcommand: " ^ quote subcommand ^ ".") end