--- a/doc-src/more_antiquote.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/doc-src/more_antiquote.ML Fri Sep 17 17:31:20 2010 +0200
@@ -6,47 +6,11 @@
signature MORE_ANTIQUOTE =
sig
- val typewriter: string -> string
end;
structure More_Antiquote : MORE_ANTIQUOTE =
struct
-(* printing typewriter lines *)
-
-fun typewriter s =
- let
- val parse = Scan.repeat
- (Scan.this_string "\n" |-- Scan.succeed "\\\\\n\\hspace*{0pt}"
- || (Scan.this_string " "
- || Scan.this_string "."
- || Scan.this_string ","
- || Scan.this_string ":"
- || Scan.this_string ";"
- || Scan.this_string "\"" |-- Scan.succeed "{\\char34}"
- || Scan.this_string "#" |-- Scan.succeed "{\\char35}"
- || Scan.this_string "$" |-- Scan.succeed "{\\char36}"
- || Scan.this_string "%" |-- Scan.succeed "{\\char37}"
- || Scan.this_string "&" |-- Scan.succeed "{\\char38}"
- || Scan.this_string "\\" |-- Scan.succeed "{\\char92}"
- || Scan.this_string "^" |-- Scan.succeed "{\\char94}"
- || Scan.this_string "_" |-- Scan.succeed "{\\char95}"
- || Scan.this_string "{" |-- Scan.succeed "{\\char123}"
- || Scan.this_string "}" |-- Scan.succeed "{\\char125}"
- || Scan.this_string "~" |-- Scan.succeed "{\\char126}")
- -- Scan.repeat (Scan.this_string " ")
- >> (fn (x, xs) => x ^ replicate_string (length xs) "~")
- || Scan.one Symbol.not_eof);
- fun is_newline s = (s = "\n");
- val cs = explode s |> take_prefix is_newline |> snd
- |> take_suffix is_newline |> fst;
- val (texts, []) = Scan.finite Symbol.stopper parse cs
- in if null texts
- then ""
- else implode ("\\isatypewriter%\n\\noindent%\n\\hspace*{0pt}" :: texts)
- end
-
-
(* code theorem antiquotation *)
local
@@ -81,36 +45,4 @@
end;
-
-(* code antiquotation *)
-
-local
-
- val parse_const_terms = Scan.repeat1 Args.term
- >> (fn ts => fn thy => map (Code.check_const thy) ts);
- val parse_consts = Scan.lift (Args.parens (Args.$$$ "consts")) |-- parse_const_terms
- >> (fn mk_cs => fn thy => fn naming => map_filter (Code_Thingol.lookup_const naming) (mk_cs thy));
- val parse_types = Scan.lift (Args.parens (Args.$$$ "types") |-- Scan.repeat1 Args.name)
- >> (fn tycos => fn thy => fn naming => map_filter (Code_Thingol.lookup_tyco naming o Sign.intern_type thy) tycos);
- val parse_classes = Scan.lift (Args.parens (Args.$$$ "classes") |-- Scan.repeat1 Args.name)
- >> (fn classes => fn thy => fn naming => map_filter (Code_Thingol.lookup_class naming o Sign.intern_class thy) classes);
- val parse_instances = Scan.lift (Args.parens (Args.$$$ "instances") |-- Scan.repeat1 (Args.name --| Args.$$$ "::" -- Args.name))
- >> (fn insts => fn thy => fn naming => map_filter (Code_Thingol.lookup_instance naming o apsnd (Sign.intern_type thy) o apfst (Sign.intern_class thy) o swap) insts);
- val parse_names = parse_consts || parse_types || parse_classes || parse_instances;
-
-in
-
-val _ = Thy_Output.antiquotation "code_stmts"
- (parse_const_terms -- Scan.repeat parse_names
- -- Scan.lift (Args.parens (Args.name -- Scan.option Parse.int)))
- (fn {context = ctxt, ...} => fn ((mk_cs, mk_stmtss), (target, some_width)) =>
- let val thy = ProofContext.theory_of ctxt in
- Code_Target.present_code thy (mk_cs thy)
- (fn naming => maps (fn f => f thy naming) mk_stmtss)
- target some_width "Example" []
- |> typewriter
- end);
-
end;
-
-end;
--- a/etc/components Fri Sep 17 17:11:43 2010 +0200
+++ b/etc/components Fri Sep 17 17:31:20 2010 +0200
@@ -17,3 +17,4 @@
src/HOL/Mirabelle
src/HOL/Library/Sum_Of_Squares
src/HOL/Tools/SMT
+src/HOL/Tools/Predicate_Compile
--- a/src/HOL/Code_Evaluation.thy Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Code_Evaluation.thy Fri Sep 17 17:31:20 2010 +0200
@@ -315,8 +315,8 @@
fun tracing s x = (Output.tracing s; x);
-fun eval_term thy t = Code_Runtime.value NONE (Evaluation.get, put_term, "Code_Evaluation.put_term")
- I thy (HOLogic.mk_term_of (fastype_of t) t) [];
+fun eval_term thy t = Code_Runtime.dynamic_value_strict (Evaluation.get, put_term, "Code_Evaluation.put_term")
+ thy NONE I (HOLogic.mk_term_of (fastype_of t) t) [];
end
*}
--- a/src/HOL/Divides.thy Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Divides.thy Fri Sep 17 17:31:20 2010 +0200
@@ -2183,6 +2183,18 @@
done
+lemma zdiv_eq_0_iff:
+ "(i::int) div k = 0 \<longleftrightarrow> k=0 \<or> 0\<le>i \<and> i<k \<or> i\<le>0 \<and> k<i" (is "?L = ?R")
+proof
+ assume ?L
+ have "?L \<longrightarrow> ?R" by (rule split_zdiv[THEN iffD2]) simp
+ with `?L` show ?R by blast
+next
+ assume ?R thus ?L
+ by(auto simp: div_pos_pos_trivial div_neg_neg_trivial)
+qed
+
+
subsubsection{*Quotients of Signs*}
lemma div_neg_pos_less0: "[| a < (0::int); 0 < b |] ==> a div b < 0"
@@ -2220,6 +2232,11 @@
lemma pos_imp_zdiv_neg_iff: "(0::int) < b ==> (a div b < 0) = (a < 0)"
by (simp add: linorder_not_le [symmetric] pos_imp_zdiv_nonneg_iff)
+lemma pos_imp_zdiv_pos_iff:
+ "0<k \<Longrightarrow> 0 < (i::int) div k \<longleftrightarrow> k \<le> i"
+using pos_imp_zdiv_nonneg_iff[of k i] zdiv_eq_0_iff[of i k]
+by arith
+
(*Again the law fails for \<le>: consider a = -1, b = -2 when a div b = 0*)
lemma neg_imp_zdiv_neg_iff: "b < (0::int) ==> (a div b < 0) = (0 < a)"
by (simp add: linorder_not_le [symmetric] neg_imp_zdiv_nonneg_iff)
@@ -2235,6 +2252,12 @@
done
+lemma zmod_le_nonneg_dividend: "(m::int) \<ge> 0 ==> m mod k \<le> m"
+apply (rule split_zmod[THEN iffD2])
+apply(fastsimp dest: q_pos_lemma intro: split_mult_pos_le)
+done
+
+
subsubsection {* The Divides Relation *}
lemmas zdvd_iff_zmod_eq_0_number_of [simp] =
--- a/src/HOL/HOL.thy Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/HOL.thy Fri Sep 17 17:31:20 2010 +0200
@@ -1958,42 +1958,17 @@
subsubsection {* Evaluation and normalization by evaluation *}
-text {* Avoid some named infixes in evaluation environment *}
-
-code_reserved Eval oo ooo oooo upto downto orf andf
-
setup {*
Value.add_evaluator ("SML", Codegen.eval_term o ProofContext.theory_of)
*}
ML {*
-structure Eval_Method = Proof_Data(
- type T = unit -> bool
- fun init thy () = error "Eval_Method"
-)
-*}
-
-oracle eval_oracle = {* fn ct =>
- let
- val thy = Thm.theory_of_cterm ct;
- val t = Thm.term_of ct;
- val dummy = @{cprop True};
- in case try HOLogic.dest_Trueprop t
- of SOME t' => if Code_Runtime.value NONE
- (Eval_Method.get, Eval_Method.put, "Eval_Method.put") (K I) thy t' []
- then Thm.capply (Thm.capply @{cterm "op \<equiv> \<Colon> prop \<Rightarrow> prop \<Rightarrow> prop"} ct) dummy
- else dummy
- | NONE => dummy
- end
-*}
-
-ML {*
fun gen_eval_method conv ctxt = SIMPLE_METHOD'
(CONVERSION (Conv.params_conv (~1) (K (Conv.concl_conv (~1) conv)) ctxt)
THEN' rtac TrueI)
*}
-method_setup eval = {* Scan.succeed (gen_eval_method eval_oracle) *}
+method_setup eval = {* Scan.succeed (gen_eval_method Code_Runtime.dynamic_holds_conv) *}
"solve goal by evaluation"
method_setup evaluation = {* Scan.succeed (gen_eval_method Codegen.evaluation_conv) *}
--- a/src/HOL/IsaMakefile Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/IsaMakefile Fri Sep 17 17:31:20 2010 +0200
@@ -268,8 +268,10 @@
$(SRC)/Tools/Metis/metis.ML \
Tools/async_manager.ML \
Tools/ATP/atp_problem.ML \
+ Tools/ATP/atp_proof.ML \
Tools/ATP/atp_systems.ML \
Tools/choice_specification.ML \
+ Tools/Datatype/datatype_selectors.ML \
Tools/int_arith.ML \
Tools/groebner.ML \
Tools/list_code.ML \
@@ -315,7 +317,8 @@
Tools/semiring_normalizer.ML \
Tools/Sledgehammer/clausifier.ML \
Tools/Sledgehammer/meson_tactic.ML \
- Tools/Sledgehammer/metis_clauses.ML \
+ Tools/Sledgehammer/metis_reconstruct.ML \
+ Tools/Sledgehammer/metis_translate.ML \
Tools/Sledgehammer/metis_tactics.ML \
Tools/Sledgehammer/sledgehammer.ML \
Tools/Sledgehammer/sledgehammer_filter.ML \
--- a/src/HOL/Library/Eval_Witness.thy Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Library/Eval_Witness.thy Fri Sep 17 17:31:20 2010 +0200
@@ -44,6 +44,13 @@
instance bool :: ml_equiv ..
instance list :: (ml_equiv) ml_equiv ..
+ML {*
+structure Eval_Method = Proof_Data (
+ type T = unit -> bool
+ fun init _ () = error "Eval_Method"
+)
+*}
+
oracle eval_witness_oracle = {* fn (cgoal, ws) =>
let
val thy = Thm.theory_of_cterm cgoal;
@@ -59,7 +66,7 @@
| dest_exs _ _ = sys_error "dest_exs";
val t = dest_exs (length ws) (HOLogic.dest_Trueprop goal);
in
- if Code_Runtime.value NONE (Eval_Method.get, Eval_Method.put, "Eval_Method.put") (K I) thy t ws
+ if Code_Runtime.dynamic_value_strict (Eval_Method.get, Eval_Method.put, "Eval_Method.put") thy NONE (K I) t ws
then Thm.cterm_of thy goal
else @{cprop True} (*dummy*)
end
--- a/src/HOL/Mirabelle/Tools/mirabelle.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Mirabelle/Tools/mirabelle.ML Fri Sep 17 17:31:20 2010 +0200
@@ -159,9 +159,9 @@
val {context = ctxt, facts, goal} = Proof.goal st
val full_tac = HEADGOAL (Method.insert_tac facts THEN' tac ctxt)
in
- (case TimeLimit.timeLimit time (Seq.pull o full_tac) goal of
- SOME _ => true
- | NONE => false)
+ (case try (TimeLimit.timeLimit time (Seq.pull o full_tac)) goal of
+ SOME (SOME _) => true
+ | _ => false)
end
local
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Fri Sep 17 17:31:20 2010 +0200
@@ -434,7 +434,7 @@
fun run_minimize args named_thms id ({pre=st, log, ...}: Mirabelle.run_args) =
let
- open Metis_Clauses
+ open Metis_Translate
val thy = Proof.theory_of st
val n0 = length (these (!named_thms))
val (prover_name, _) = get_atp thy args
--- a/src/HOL/Nominal/Examples/Standardization.thy Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Nominal/Examples/Standardization.thy Fri Sep 17 17:31:20 2010 +0200
@@ -169,14 +169,9 @@
lemma ex_head_tail:
"\<exists>ts h. t = h \<degree>\<degree> ts \<and> ((\<exists>n. h = Var n) \<or> (\<exists>x u. h = (Lam [x].u)))"
apply (induct t rule: lam.induct)
- apply (rule_tac x = "[]" in exI)
- apply (simp add: lam.inject)
- apply clarify
- apply (rename_tac ts1 ts2 h1 h2)
- apply (rule_tac x = "ts1 @ [h2 \<degree>\<degree> ts2]" in exI)
- apply (simp add: lam.inject)
- apply simp
- apply blast
+ apply (metis foldl_Nil)
+ apply (metis foldl_Cons foldl_Nil foldl_append)
+ apply (metis foldl_Nil)
done
lemma size_apps [simp]:
@@ -218,15 +213,11 @@
prefer 2
apply (erule allE, erule impE, rule refl, erule spec)
apply simp
- apply (rule lem0)
- apply force
- apply (rule elem_le_sum)
- apply force
+ apply (metis add_strict_increasing le_add1 length_pos_if_in_set listsum listsum_map_remove1 nat_add_commute)
apply clarify
apply (subgoal_tac "\<exists>y::name. y \<sharp> (x, u, z)")
- prefer 2
- apply (rule exists_fresh')
- apply (rule fin_supp)
+ prefer 2
+ apply (blast intro: exists_fresh' fin_supp)
apply (erule exE)
apply (subgoal_tac "(Lam [x].u) = (Lam [y].([(y, x)] \<bullet> u))")
prefer 2
@@ -241,13 +232,8 @@
apply clarify
apply (erule allE, erule impE)
prefer 2
- apply (erule allE, erule impE, rule refl, erule spec)
- apply simp
- apply (rule le_imp_less_Suc)
- apply (rule trans_le_add1)
- apply (rule trans_le_add2)
- apply (rule elem_le_sum)
- apply force
+ apply blast
+ apply (force intro: le_imp_less_Suc trans_le_add1 trans_le_add2 elem_le_sum)
done
theorem Apps_lam_induct:
@@ -441,13 +427,9 @@
assumes xy: "listrelp f (x::'a::pt_name list) y"
shows "listrelp ((pi::name prm) \<bullet> f) (pi \<bullet> x) (pi \<bullet> y)" using xy
apply induct
- apply simp
- apply (rule listrelp.intros)
+ apply (simp add: listrelp.intros)
apply simp
- apply (rule listrelp.intros)
- apply (drule_tac pi=pi in perm_boolI)
- apply perm_simp
- apply assumption
+ apply (metis listrelp.Cons in_eqvt mem_def perm_app pt_set_bij3)
done
inductive
@@ -745,12 +727,7 @@
case (2 x u ts)
show ?case
proof (cases ts)
- case Nil
- from 2 have "\<forall>u'. \<not> u \<rightarrow>\<^sub>\<beta> u'"
- by (auto intro: apps_preserves_beta)
- then have "NF u" by (rule 2)
- then have "NF (Lam [x].u)" by (rule NF.Abs)
- with Nil show ?thesis by simp
+ case Nil thus ?thesis by (metis 2 NF.Abs abs foldl_Nil)
next
case (Cons r rs)
have "(Lam [x].u) \<degree> r \<rightarrow>\<^sub>\<beta> u[x::=r]" ..
@@ -841,7 +818,7 @@
case Nil
show ?case by (rule listrelp.Nil)
next
- case (Cons x y xs ys)
+ case (Cons x y xs ys)
hence "x \<rightarrow>\<^sub>l y" and "xs [\<rightarrow>\<^sub>l] ys" by (auto del: in_listspD)
thus ?case by (rule listrelp.Cons)
qed
--- a/src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy Fri Sep 17 17:31:20 2010 +0200
@@ -15,12 +15,12 @@
limited_types = [],
limited_predicates = [],
replacing = [],
- manual_reorder = [],
- timeout = Time.fromSeconds 10,
- prolog_system = Code_Prolog.SWI_PROLOG}) *}
+ manual_reorder = []}) *}
values "{(x, y, z). append x y z}"
+values 4 "{(z, x, y). append x y ((1::nat) # (2 # (3 # z)))}"
+
values 3 "{(x, y, z). append x y z}"
setup {* Code_Prolog.map_code_options (K
@@ -28,9 +28,7 @@
limited_types = [],
limited_predicates = [],
replacing = [],
- manual_reorder = [],
- timeout = Time.fromSeconds 10,
- prolog_system = Code_Prolog.YAP}) *}
+ manual_reorder = []}) *}
values "{(x, y, z). append x y z}"
@@ -39,9 +37,7 @@
limited_types = [],
limited_predicates = [],
replacing = [],
- manual_reorder = [],
- timeout = Time.fromSeconds 10,
- prolog_system = Code_Prolog.SWI_PROLOG}) *}
+ manual_reorder = []}) *}
section {* Example queens *}
@@ -209,9 +205,7 @@
limited_types = [],
limited_predicates = [],
replacing = [],
- manual_reorder = [],
- timeout = Time.fromSeconds 10,
- prolog_system = Code_Prolog.SWI_PROLOG}) *}
+ manual_reorder = []}) *}
values 2 "{y. notB y}"
--- a/src/HOL/Predicate_Compile_Examples/Context_Free_Grammar_Example.thy Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Context_Free_Grammar_Example.thy Fri Sep 17 17:31:20 2010 +0200
@@ -19,7 +19,7 @@
declare size_list_def[symmetric, code_pred_inline]
-setup {* Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck) *}
+setup {* Context.theory_map (Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck)) *}
datatype alphabet = a | b
@@ -61,9 +61,7 @@
limited_types = [],
limited_predicates = [(["s1", "a1", "b1"], 2)],
replacing = [(("s1", "limited_s1"), "quickcheck")],
- manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])],
- timeout = Time.fromSeconds 10,
- prolog_system = Code_Prolog.SWI_PROLOG}) *}
+ manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])]}) *}
theorem S\<^isub>1_sound:
@@ -86,9 +84,7 @@
limited_types = [],
limited_predicates = [(["s2", "a2", "b2"], 3)],
replacing = [(("s2", "limited_s2"), "quickcheck")],
- manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])],
- timeout = Time.fromSeconds 10,
- prolog_system = Code_Prolog.SWI_PROLOG}) *}
+ manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])]}) *}
theorem S\<^isub>2_sound:
@@ -110,9 +106,7 @@
limited_types = [],
limited_predicates = [(["s3", "a3", "b3"], 6)],
replacing = [(("s3", "limited_s3"), "quickcheck")],
- manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])],
- timeout = Time.fromSeconds 10,
- prolog_system = Code_Prolog.SWI_PROLOG}) *}
+ manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])]}) *}
lemma S\<^isub>3_sound:
"w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
@@ -152,9 +146,7 @@
limited_types = [],
limited_predicates = [(["s4", "a4", "b4"], 6)],
replacing = [(("s4", "limited_s4"), "quickcheck")],
- manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])],
- timeout = Time.fromSeconds 10,
- prolog_system = Code_Prolog.SWI_PROLOG}) *}
+ manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])]}) *}
theorem S\<^isub>4_sound:
--- a/src/HOL/Predicate_Compile_Examples/Hotel_Example.thy Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Hotel_Example.thy Fri Sep 17 17:31:20 2010 +0200
@@ -89,14 +89,12 @@
limited_types = [],
limited_predicates = [],
replacing = [],
- manual_reorder = [],
- timeout = Time.fromSeconds 10,
- prolog_system = Code_Prolog.SWI_PROLOG}) *}
+ manual_reorder = []}) *}
values 40 "{s. hotel s}"
-setup {* Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck) *}
+setup {* Context.theory_map (Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck)) *}
lemma "\<lbrakk> hotel s; g \<in> isin s r \<rbrakk> \<Longrightarrow> owns s r = Some g"
quickcheck[generator = code, iterations = 100000, report]
@@ -119,9 +117,7 @@
limited_types = [],
limited_predicates = [(["hotel"], 5)],
replacing = [(("hotel", "limited_hotel"), "quickcheck")],
- manual_reorder = [],
- timeout = Time.fromSeconds 10,
- prolog_system = Code_Prolog.SWI_PROLOG}) *}
+ manual_reorder = []}) *}
lemma
"hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
--- a/src/HOL/Predicate_Compile_Examples/Lambda_Example.thy Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Lambda_Example.thy Fri Sep 17 17:31:20 2010 +0200
@@ -79,7 +79,7 @@
section {* Manual setup to find counterexample *}
-setup {* Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck) *}
+setup {* Context.theory_map (Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck)) *}
setup {* Code_Prolog.map_code_options (K
{ ensure_groundness = true,
@@ -87,9 +87,7 @@
limited_predicates = [(["typing"], 2), (["nthel1"], 2)],
replacing = [(("typing", "limited_typing"), "quickcheck"),
(("nthel1", "limited_nthel1"), "lim_typing")],
- manual_reorder = [],
- timeout = Time.fromSeconds 10,
- prolog_system = Code_Prolog.SWI_PROLOG}) *}
+ manual_reorder = []}) *}
lemma
"\<Gamma> \<turnstile> t : U \<Longrightarrow> t \<rightarrow>\<^sub>\<beta> t' \<Longrightarrow> \<Gamma> \<turnstile> t' : U"
--- a/src/HOL/Predicate_Compile_Examples/List_Examples.thy Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Predicate_Compile_Examples/List_Examples.thy Fri Sep 17 17:31:20 2010 +0200
@@ -2,7 +2,7 @@
imports Main "Predicate_Compile_Quickcheck" "Code_Prolog"
begin
-setup {* Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck) *}
+setup {* Context.theory_map (Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck)) *}
setup {* Code_Prolog.map_code_options (K
{ensure_groundness = true,
@@ -12,9 +12,7 @@
[(("appendP", "limited_appendP"), "quickcheck"),
(("revP", "limited_revP"), "quickcheck"),
(("appendP", "limited_appendP"), "lim_revP")],
- manual_reorder = [],
- timeout = Time.fromSeconds 10,
- prolog_system = Code_Prolog.SWI_PROLOG}) *}
+ manual_reorder = []}) *}
lemma "(xs :: nat list) = ys @ ys --> rev xs = xs"
quickcheck[generator = code, iterations = 200000, expect = counterexample]
--- a/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Examples.thy Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Examples.thy Fri Sep 17 17:31:20 2010 +0200
@@ -1310,6 +1310,10 @@
values "{xs::int list. test_append (replicate 1000 1) (replicate 1000 2) xs}"
values "{xs::int list. test_append (replicate 2000 1) (replicate 2000 2) xs}"
+text {* Redeclaring append.equation as code equation *}
+
+declare appendP.equation[code]
+
subsection {* Function with tuples *}
fun append'
@@ -1580,6 +1584,56 @@
thm big_step.equation
+section {* General Context Free Grammars *}
+text {* a contribution by Aditi Barthwal *}
+
+datatype ('nts,'ts) symbol = NTS 'nts
+ | TS 'ts
+
+
+datatype ('nts,'ts) rule = rule 'nts "('nts,'ts) symbol list"
+
+types ('nts,'ts) grammar = "('nts,'ts) rule set * 'nts"
+
+fun rules :: "('nts,'ts) grammar => ('nts,'ts) rule set"
+where
+ "rules (r, s) = r"
+
+definition derives
+where
+"derives g = { (lsl,rsl). \<exists>s1 s2 lhs rhs.
+ (s1 @ [NTS lhs] @ s2 = lsl) \<and>
+ (s1 @ rhs @ s2) = rsl \<and>
+ (rule lhs rhs) \<in> fst g }"
+
+abbreviation "example_grammar ==
+({ rule ''S'' [NTS ''A'', NTS ''B''],
+ rule ''S'' [TS ''a''],
+ rule ''A'' [TS ''b'']}, ''S'')"
+
+
+code_pred [inductify,skip_proof] derives .
+
+thm derives.equation
+
+definition "test = { rhs. derives example_grammar ([NTS ''S''], rhs) }"
+
+code_pred (modes: o \<Rightarrow> bool) [inductify, show_modes, show_intermediate_results, skip_proof] test .
+thm test.equation
+
+values "{rhs. test rhs}"
+
+declare rtrancl.intros(1)[code_pred_def] converse_rtrancl_into_rtrancl[code_pred_def]
+
+code_pred [inductify] rtrancl .
+
+definition "test2 = { rhs. ([NTS ''S''],rhs) \<in> (derives example_grammar)^* }"
+
+code_pred [inductify, skip_proof] test2 .
+
+values "{rhs. test2 rhs}"
+
+
section {* Examples for detecting switches *}
inductive detect_switches1 where
--- a/src/HOL/Predicate_Compile_Examples/Reg_Exp_Example.thy Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Reg_Exp_Example.thy Fri Sep 17 17:31:20 2010 +0200
@@ -96,7 +96,7 @@
oops
-setup {* Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck) *}
+setup {* Context.theory_map (Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck)) *}
setup {* Code_Prolog.map_code_options (K
{ensure_groundness = true,
@@ -108,9 +108,7 @@
(("subP", "limited_subP"), "repIntP"),
(("repIntPa", "limited_repIntPa"), "repIntP"),
(("accepts", "limited_accepts"), "quickcheck")],
- manual_reorder = [],
- timeout = Time.fromSeconds 10,
- prolog_system = Code_Prolog.SWI_PROLOG}) *}
+ manual_reorder = []}) *}
text {* Finding the counterexample still seems out of reach for the
prolog-style generation. *}
--- a/src/HOL/SMT.thy Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/SMT.thy Fri Sep 17 17:31:20 2010 +0200
@@ -7,6 +7,7 @@
theory SMT
imports List
uses
+ "Tools/Datatype/datatype_selectors.ML"
("Tools/SMT/smt_monomorph.ML")
("Tools/SMT/smt_normalize.ML")
("Tools/SMT/smt_translate.ML")
@@ -323,4 +324,13 @@
hide_const Pattern term_eq
hide_const (open) trigger pat nopat fun_app z3div z3mod
+
+
+subsection {* Selectors for datatypes *}
+
+setup {* Datatype_Selectors.setup *}
+
+declare [[ selector Pair 1 = fst, selector Pair 2 = snd ]]
+declare [[ selector Cons 1 = hd, selector Cons 2 = tl ]]
+
end
--- a/src/HOL/SMT_Examples/SMT_Tests.thy Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/SMT_Examples/SMT_Tests.thy Fri Sep 17 17:31:20 2010 +0200
@@ -607,7 +607,7 @@
-section {* Pairs *}
+section {* Pairs *} (* FIXME: tests for datatypes and records *)
lemma
"x = fst (x, y)"
--- a/src/HOL/Sledgehammer.thy Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Sledgehammer.thy Fri Sep 17 17:31:20 2010 +0200
@@ -11,11 +11,13 @@
imports Plain Hilbert_Choice
uses
("Tools/ATP/atp_problem.ML")
+ ("Tools/ATP/atp_proof.ML")
("Tools/ATP/atp_systems.ML")
("~~/src/Tools/Metis/metis.ML")
("Tools/Sledgehammer/clausifier.ML")
("Tools/Sledgehammer/meson_tactic.ML")
- ("Tools/Sledgehammer/metis_clauses.ML")
+ ("Tools/Sledgehammer/metis_translate.ML")
+ ("Tools/Sledgehammer/metis_reconstruct.ML")
("Tools/Sledgehammer/metis_tactics.ML")
("Tools/Sledgehammer/sledgehammer_util.ML")
("Tools/Sledgehammer/sledgehammer_filter.ML")
@@ -92,6 +94,7 @@
done
use "Tools/ATP/atp_problem.ML"
+use "Tools/ATP/atp_proof.ML"
use "Tools/ATP/atp_systems.ML"
setup ATP_Systems.setup
@@ -100,7 +103,8 @@
use "Tools/Sledgehammer/meson_tactic.ML"
setup Meson_Tactic.setup
-use "Tools/Sledgehammer/metis_clauses.ML"
+use "Tools/Sledgehammer/metis_translate.ML"
+use "Tools/Sledgehammer/metis_reconstruct.ML"
use "Tools/Sledgehammer/metis_tactics.ML"
setup Metis_Tactics.setup
--- a/src/HOL/Tools/ATP/atp_problem.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Tools/ATP/atp_problem.ML Fri Sep 17 17:31:20 2010 +0200
@@ -2,7 +2,7 @@
Author: Jia Meng, Cambridge University Computer Laboratory and NICTA
Author: Jasmin Blanchette, TU Muenchen
-TPTP syntax.
+Abstract representation of ATP problems and TPTP syntax.
*)
signature ATP_PROBLEM =
@@ -14,16 +14,17 @@
AQuant of quantifier * 'a list * ('a, 'b) formula |
AConn of connective * ('a, 'b) formula list |
AAtom of 'b
+ type 'a uniform_formula = ('a, 'a fo_term) formula
datatype kind = Axiom | Hypothesis | Conjecture
datatype 'a problem_line = Fof of string * kind * ('a, 'a fo_term) formula
type 'a problem = (string * 'a problem_line list) list
val timestamp : unit -> string
- val is_tptp_variable : string -> bool
- val strings_for_tptp_problem :
+ val is_atp_variable : string -> bool
+ val tptp_strings_for_atp_problem :
bool -> (string * string problem_line list) list -> string list
- val nice_tptp_problem :
+ val nice_atp_problem :
bool -> ('a * (string * string) problem_line list) list
-> ('a * string problem_line list) list
* (string Symtab.table * string Symtab.table) option
@@ -41,6 +42,7 @@
AQuant of quantifier * 'a list * ('a, 'b) formula |
AConn of connective * ('a, 'b) formula list |
AAtom of 'b
+type 'a uniform_formula = ('a, 'a fo_term) formula
datatype kind = Axiom | Hypothesis | Conjecture
datatype 'a problem_line = Fof of string * kind * ('a, 'a fo_term) formula
@@ -90,7 +92,7 @@
"fof(" ^ ident ^ ", " ^ string_for_kind kind ^ ",\n (" ^
string_for_formula phi ^ ")).\n"
end
-fun strings_for_tptp_problem use_conjecture_for_hypotheses problem =
+fun tptp_strings_for_atp_problem use_conjecture_for_hypotheses problem =
"% This file was generated by Isabelle (most likely Sledgehammer)\n\
\% " ^ timestamp () ^ "\n" ::
maps (fn (_, []) => []
@@ -99,7 +101,7 @@
map (string_for_problem_line use_conjecture_for_hypotheses) lines)
problem
-fun is_tptp_variable s = Char.isUpper (String.sub (s, 0))
+fun is_atp_variable s = Char.isUpper (String.sub (s, 0))
(** Nice names **)
@@ -163,7 +165,7 @@
fun nice_problem problem =
pool_map (fn (heading, lines) =>
pool_map nice_problem_line lines #>> pair heading) problem
-fun nice_tptp_problem readable_names problem =
+fun nice_atp_problem readable_names problem =
nice_problem problem (empty_name_pool readable_names)
end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/ATP/atp_proof.ML Fri Sep 17 17:31:20 2010 +0200
@@ -0,0 +1,357 @@
+(* Title: HOL/Tools/ATP/atp_proof.ML
+ Author: Lawrence C. Paulson, Cambridge University Computer Laboratory
+ Author: Claire Quigley, Cambridge University Computer Laboratory
+ Author: Jasmin Blanchette, TU Muenchen
+
+Abstract representation of ATP proofs and TSTP/Vampire/SPASS syntax.
+*)
+
+signature ATP_PROOF =
+sig
+ type 'a fo_term = 'a ATP_Problem.fo_term
+ type 'a uniform_formula = 'a ATP_Problem.uniform_formula
+
+ datatype failure =
+ Unprovable | IncompleteUnprovable | CantConnect | TimedOut |
+ OutOfResources | SpassTooOld | VampireTooOld | NoPerl | NoLibwwwPerl |
+ MalformedInput | MalformedOutput | Interrupted | InternalError |
+ UnknownError
+
+ type step_name = string * string option
+
+ datatype 'a step =
+ Definition of step_name * 'a * 'a |
+ Inference of step_name * 'a * step_name list
+
+ type 'a proof = 'a uniform_formula step list
+
+ val strip_spaces : (char -> bool) -> string -> string
+ val string_for_failure : failure -> string
+ val extract_important_message : string -> string
+ val extract_known_failure :
+ (failure * string) list -> string -> failure option
+ val extract_tstplike_proof_and_outcome :
+ bool -> int -> (string * string) list -> (failure * string) list -> string
+ -> string * failure option
+ val is_same_step : step_name * step_name -> bool
+ val atp_proof_from_tstplike_string : string -> string proof
+ val map_term_names_in_atp_proof :
+ (string -> string) -> string proof -> string proof
+ val nasty_atp_proof : string Symtab.table -> string proof -> string proof
+end;
+
+structure ATP_Proof : ATP_PROOF =
+struct
+
+open ATP_Problem
+
+datatype failure =
+ Unprovable | IncompleteUnprovable | CantConnect | TimedOut | OutOfResources |
+ SpassTooOld | VampireTooOld | NoPerl | NoLibwwwPerl | MalformedInput |
+ MalformedOutput | Interrupted | InternalError | UnknownError
+
+fun strip_spaces_in_list _ [] = []
+ | strip_spaces_in_list _ [c1] = if Char.isSpace c1 then [] else [str c1]
+ | strip_spaces_in_list is_evil [c1, c2] =
+ strip_spaces_in_list is_evil [c1] @ strip_spaces_in_list is_evil [c2]
+ | strip_spaces_in_list is_evil (c1 :: c2 :: c3 :: cs) =
+ if Char.isSpace c1 then
+ strip_spaces_in_list is_evil (c2 :: c3 :: cs)
+ else if Char.isSpace c2 then
+ if Char.isSpace c3 then
+ strip_spaces_in_list is_evil (c1 :: c3 :: cs)
+ else
+ str c1 :: (if forall is_evil [c1, c3] then [" "] else []) @
+ strip_spaces_in_list is_evil (c3 :: cs)
+ else
+ str c1 :: strip_spaces_in_list is_evil (c2 :: c3 :: cs)
+fun strip_spaces is_evil =
+ implode o strip_spaces_in_list is_evil o String.explode
+
+fun is_ident_char c = Char.isAlphaNum c orelse c = #"_"
+val strip_spaces_except_between_ident_chars = strip_spaces is_ident_char
+
+val missing_message_tail =
+ " appears to be missing. You will need to install it if you want to run \
+ \ATPs remotely."
+
+fun string_for_failure Unprovable = "The ATP problem is unprovable."
+ | string_for_failure IncompleteUnprovable =
+ "The ATP cannot prove the problem."
+ | string_for_failure CantConnect = "Can't connect to remote server."
+ | string_for_failure TimedOut = "Timed out."
+ | string_for_failure OutOfResources = "The ATP ran out of resources."
+ | string_for_failure SpassTooOld =
+ "Isabelle requires a more recent version of SPASS with support for the \
+ \TPTP syntax. To install it, download and extract the package \
+ \\"http://isabelle.in.tum.de/dist/contrib/spass-3.7.tar.gz\" and add the \
+ \\"spass-3.7\" directory's absolute path to " ^
+ quote (Path.implode (Path.expand (Path.appends
+ (Path.variable "ISABELLE_HOME_USER" ::
+ map Path.basic ["etc", "components"])))) ^
+ " on a line of its own."
+ | string_for_failure VampireTooOld =
+ "Isabelle requires a more recent version of Vampire. To install it, follow \
+ \the instructions from the Sledgehammer manual (\"isabelle doc\
+ \ sledgehammer\")."
+ | string_for_failure NoPerl = "Perl" ^ missing_message_tail
+ | string_for_failure NoLibwwwPerl =
+ "The Perl module \"libwww-perl\"" ^ missing_message_tail
+ | string_for_failure MalformedInput =
+ "The ATP problem is malformed. Please report this to the Isabelle \
+ \developers."
+ | string_for_failure MalformedOutput = "The ATP output is malformed."
+ | string_for_failure Interrupted = "The ATP was interrupted."
+ | string_for_failure InternalError = "An internal ATP error occurred."
+ | string_for_failure UnknownError = "An unknown ATP error occurred."
+
+fun extract_delimited (begin_delim, end_delim) output =
+ output |> first_field begin_delim |> the |> snd
+ |> first_field end_delim |> the |> fst
+ |> first_field "\n" |> the |> snd
+ handle Option.Option => ""
+
+val tstp_important_message_delims =
+ ("% SZS start RequiredInformation", "% SZS end RequiredInformation")
+
+fun extract_important_message output =
+ case extract_delimited tstp_important_message_delims output of
+ "" => ""
+ | s => s |> space_explode "\n" |> filter_out (curry (op =) "")
+ |> map (perhaps (try (unprefix "%")))
+ |> map (perhaps (try (unprefix " ")))
+ |> space_implode "\n " |> quote
+
+(* Splits by the first possible of a list of delimiters. *)
+fun extract_tstplike_proof delims output =
+ case pairself (find_first (fn s => String.isSubstring s output))
+ (ListPair.unzip delims) of
+ (SOME begin_delim, SOME end_delim) =>
+ extract_delimited (begin_delim, end_delim) output
+ | _ => ""
+
+fun extract_known_failure known_failures output =
+ known_failures
+ |> find_first (fn (_, pattern) => String.isSubstring pattern output)
+ |> Option.map fst
+
+fun extract_tstplike_proof_and_outcome complete res_code proof_delims
+ known_failures output =
+ case extract_known_failure known_failures output of
+ NONE => (case extract_tstplike_proof proof_delims output of
+ "" => ("", SOME (if res_code = 0 andalso output = "" then
+ Interrupted
+ else
+ UnknownError))
+ | tstplike_proof => if res_code = 0 then (tstplike_proof, NONE)
+ else ("", SOME UnknownError))
+ | SOME failure =>
+ ("", SOME (if failure = IncompleteUnprovable andalso complete then
+ Unprovable
+ else
+ failure))
+
+fun mk_anot (AConn (ANot, [phi])) = phi
+ | mk_anot phi = AConn (ANot, [phi])
+fun mk_aconn c (phi1, phi2) = AConn (c, [phi1, phi2])
+
+type step_name = string * string option
+
+fun is_same_step p = p |> pairself fst |> op =
+
+fun step_name_ord p =
+ let val q = pairself fst p in
+ (* The "unprefix" part is to cope with remote Vampire's output. The proper
+ solution would be to perform a topological sort, e.g. using the nice
+ "Graph" functor. *)
+ case pairself (Int.fromString o perhaps (try (unprefix "f"))) q of
+ (NONE, NONE) => string_ord q
+ | (NONE, SOME _) => LESS
+ | (SOME _, NONE) => GREATER
+ | (SOME i, SOME j) => int_ord (i, j)
+ end
+
+datatype 'a step =
+ Definition of step_name * 'a * 'a |
+ Inference of step_name * 'a * step_name list
+
+type 'a proof = 'a uniform_formula step list
+
+fun step_name (Definition (name, _, _)) = name
+ | step_name (Inference (name, _, _)) = name
+
+(**** PARSING OF TSTP FORMAT ****)
+
+(*Strings enclosed in single quotes, e.g. filenames*)
+val scan_general_id =
+ $$ "'" |-- Scan.repeat (~$$ "'") --| $$ "'" >> implode
+ || Scan.repeat ($$ "$") -- Scan.many1 Symbol.is_letdig
+ >> (fn (ss1, ss2) => implode ss1 ^ implode ss2)
+
+(* Generalized first-order terms, which include file names, numbers, etc. *)
+fun parse_annotation strict x =
+ ((scan_general_id ::: Scan.repeat ($$ " " |-- scan_general_id)
+ >> (strict ? filter (is_some o Int.fromString)))
+ -- Scan.optional (parse_annotation strict) [] >> op @
+ || $$ "(" |-- parse_annotations strict --| $$ ")"
+ || $$ "[" |-- parse_annotations strict --| $$ "]") x
+and parse_annotations strict x =
+ (Scan.optional (parse_annotation strict
+ ::: Scan.repeat ($$ "," |-- parse_annotation strict)) []
+ >> flat) x
+
+(* Vampire proof lines sometimes contain needless information such as "(0:3)",
+ which can be hard to disambiguate from function application in an LL(1)
+ parser. As a workaround, we extend the TPTP term syntax with such detritus
+ and ignore it. *)
+fun parse_vampire_detritus x =
+ (scan_general_id |-- $$ ":" --| scan_general_id >> K []) x
+
+fun parse_term x =
+ (scan_general_id
+ -- Scan.optional ($$ "(" |-- (parse_vampire_detritus || parse_terms)
+ --| $$ ")") []
+ --| Scan.optional ($$ "(" |-- parse_vampire_detritus --| $$ ")") []
+ >> ATerm) x
+and parse_terms x = (parse_term ::: Scan.repeat ($$ "," |-- parse_term)) x
+
+val parse_atom =
+ parse_term -- Scan.option (Scan.option ($$ "!") --| $$ "=" -- parse_term)
+ >> (fn (u1, NONE) => AAtom u1
+ | (u1, SOME (NONE, u2)) => AAtom (ATerm ("c_equal", [u1, u2]))
+ | (u1, SOME (SOME _, u2)) =>
+ mk_anot (AAtom (ATerm ("c_equal", [u1, u2]))))
+
+fun fo_term_head (ATerm (s, _)) = s
+
+(* TPTP formulas are fully parenthesized, so we don't need to worry about
+ operator precedence. *)
+fun parse_formula x =
+ (($$ "(" |-- parse_formula --| $$ ")"
+ || ($$ "!" >> K AForall || $$ "?" >> K AExists)
+ --| $$ "[" -- parse_terms --| $$ "]" --| $$ ":" -- parse_formula
+ >> (fn ((q, ts), phi) => AQuant (q, map fo_term_head ts, phi))
+ || $$ "~" |-- parse_formula >> mk_anot
+ || parse_atom)
+ -- Scan.option ((Scan.this_string "=>" >> K AImplies
+ || Scan.this_string "<=>" >> K AIff
+ || Scan.this_string "<~>" >> K ANotIff
+ || Scan.this_string "<=" >> K AIf
+ || $$ "|" >> K AOr || $$ "&" >> K AAnd)
+ -- parse_formula)
+ >> (fn (phi1, NONE) => phi1
+ | (phi1, SOME (c, phi2)) => mk_aconn c (phi1, phi2))) x
+
+val parse_tstp_extra_arguments =
+ Scan.optional ($$ "," |-- parse_annotation false
+ --| Scan.option ($$ "," |-- parse_annotations false)) []
+
+(* Syntax: (fof|cnf)\(<num>, <formula_role>, <formula> <extra_arguments>\).
+ The <num> could be an identifier, but we assume integers. *)
+val parse_tstp_line =
+ ((Scan.this_string "fof" || Scan.this_string "cnf") -- $$ "(")
+ |-- scan_general_id --| $$ "," -- Symbol.scan_id --| $$ ","
+ -- parse_formula -- parse_tstp_extra_arguments --| $$ ")" --| $$ "."
+ >> (fn (((num, role), phi), deps) =>
+ let
+ val (name, deps) =
+ case deps of
+ ["file", _, s] => ((num, SOME s), [])
+ | _ => ((num, NONE), deps)
+ in
+ case role of
+ "definition" =>
+ (case phi of
+ AConn (AIff, [phi1 as AAtom _, phi2]) =>
+ Definition (name, phi1, phi2)
+ | AAtom (ATerm ("c_equal", _)) =>
+ (* Vampire's equality proxy axiom *)
+ Inference (name, phi, map (rpair NONE) deps)
+ | _ => raise Fail "malformed definition")
+ | _ => Inference (name, phi, map (rpair NONE) deps)
+ end)
+
+(**** PARSING OF VAMPIRE OUTPUT ****)
+
+(* Syntax: <num>. <formula> <annotation> *)
+val parse_vampire_line =
+ scan_general_id --| $$ "." -- parse_formula -- parse_annotation true
+ >> (fn ((num, phi), deps) =>
+ Inference ((num, NONE), phi, map (rpair NONE) deps))
+
+(**** PARSING OF SPASS OUTPUT ****)
+
+(* SPASS returns clause references of the form "x.y". We ignore "y", whose role
+ is not clear anyway. *)
+val parse_dot_name = scan_general_id --| $$ "." --| scan_general_id
+
+val parse_spass_annotations =
+ Scan.optional ($$ ":" |-- Scan.repeat (parse_dot_name
+ --| Scan.option ($$ ","))) []
+
+(* It is not clear why some literals are followed by sequences of stars and/or
+ pluses. We ignore them. *)
+val parse_decorated_atom =
+ parse_atom --| Scan.repeat ($$ "*" || $$ "+" || $$ " ")
+
+fun mk_horn ([], []) = AAtom (ATerm ("c_False", []))
+ | mk_horn ([], pos_lits) = foldr1 (mk_aconn AOr) pos_lits
+ | mk_horn (neg_lits, []) = mk_anot (foldr1 (mk_aconn AAnd) neg_lits)
+ | mk_horn (neg_lits, pos_lits) =
+ mk_aconn AImplies (foldr1 (mk_aconn AAnd) neg_lits,
+ foldr1 (mk_aconn AOr) pos_lits)
+
+val parse_horn_clause =
+ Scan.repeat parse_decorated_atom --| $$ "|" --| $$ "|"
+ -- Scan.repeat parse_decorated_atom --| $$ "-" --| $$ ">"
+ -- Scan.repeat parse_decorated_atom
+ >> (mk_horn o apfst (op @))
+
+(* Syntax: <num>[0:<inference><annotations>]
+ <atoms> || <atoms> -> <atoms>. *)
+val parse_spass_line =
+ scan_general_id --| $$ "[" --| $$ "0" --| $$ ":" --| Symbol.scan_id
+ -- parse_spass_annotations --| $$ "]" -- parse_horn_clause --| $$ "."
+ >> (fn ((num, deps), u) => Inference ((num, NONE), u, map (rpair NONE) deps))
+
+val parse_line = parse_tstp_line || parse_vampire_line || parse_spass_line
+val parse_proof =
+ fst o Scan.finite Symbol.stopper
+ (Scan.error (!! (fn _ => raise Fail "unrecognized ATP output")
+ (Scan.repeat1 parse_line)))
+ o explode o strip_spaces_except_between_ident_chars
+
+fun clean_up_dependency seen dep = find_first (curry is_same_step dep) seen
+fun clean_up_dependencies _ [] = []
+ | clean_up_dependencies seen ((step as Definition (name, _, _)) :: steps) =
+ step :: clean_up_dependencies (name :: seen) steps
+ | clean_up_dependencies seen (Inference (name, u, deps) :: steps) =
+ Inference (name, u, map_filter (clean_up_dependency seen) deps) ::
+ clean_up_dependencies (name :: seen) steps
+
+val atp_proof_from_tstplike_string =
+ suffix "$" (* the $ sign acts as a sentinel (FIXME: needed?) *)
+ #> parse_proof
+ #> sort (step_name_ord o pairself step_name)
+ #> clean_up_dependencies []
+
+fun map_term_names_in_term f (ATerm (s, ts)) =
+ ATerm (f s, map (map_term_names_in_term f) ts)
+fun map_term_names_in_formula f (AQuant (q, xs, phi)) =
+ AQuant (q, xs, map_term_names_in_formula f phi)
+ | map_term_names_in_formula f (AConn (c, phis)) =
+ AConn (c, map (map_term_names_in_formula f) phis)
+ | map_term_names_in_formula f (AAtom t) = AAtom (map_term_names_in_term f t)
+fun map_term_names_in_step f (Definition (name, phi1, phi2)) =
+ Definition (name, map_term_names_in_formula f phi1,
+ map_term_names_in_formula f phi2)
+ | map_term_names_in_step f (Inference (name, phi, deps)) =
+ Inference (name, map_term_names_in_formula f phi, deps)
+fun map_term_names_in_atp_proof f = map (map_term_names_in_step f)
+
+fun nasty_name pool s = s |> Symtab.lookup pool |> the_default s
+fun nasty_atp_proof pool =
+ if Symtab.is_empty pool then I
+ else map_term_names_in_atp_proof (nasty_name pool)
+
+end;
--- a/src/HOL/Tools/ATP/atp_systems.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Tools/ATP/atp_systems.ML Fri Sep 17 17:31:20 2010 +0200
@@ -7,11 +7,7 @@
signature ATP_SYSTEMS =
sig
- datatype failure =
- Unprovable | IncompleteUnprovable | CantConnect | TimedOut |
- OutOfResources | SpassTooOld | VampireTooOld | NoPerl | NoLibwwwPerl |
- MalformedInput | MalformedOutput | Interrupted | InternalError |
- UnknownError
+ type failure = ATP_Proof.failure
type prover_config =
{exec: string * string,
@@ -24,9 +20,6 @@
explicit_forall: bool,
use_conjecture_for_hypotheses: bool}
- val string_for_failure : failure -> string
- val known_failure_in_output :
- string -> (failure * string) list -> failure option
val add_prover: string * prover_config -> theory -> theory
val get_prover: theory -> string -> prover_config
val available_atps: theory -> unit
@@ -38,12 +31,9 @@
structure ATP_Systems : ATP_SYSTEMS =
struct
-(* prover configuration *)
+open ATP_Proof
-datatype failure =
- Unprovable | IncompleteUnprovable | CantConnect | TimedOut | OutOfResources |
- SpassTooOld | VampireTooOld | NoPerl | NoLibwwwPerl | MalformedInput |
- MalformedOutput | Interrupted | InternalError | UnknownError
+(* prover configuration *)
type prover_config =
{exec: string * string,
@@ -56,44 +46,6 @@
explicit_forall: bool,
use_conjecture_for_hypotheses: bool}
-val missing_message_tail =
- " appears to be missing. You will need to install it if you want to run \
- \ATPs remotely."
-
-fun string_for_failure Unprovable = "The ATP problem is unprovable."
- | string_for_failure IncompleteUnprovable =
- "The ATP cannot prove the problem."
- | string_for_failure CantConnect = "Can't connect to remote server."
- | string_for_failure TimedOut = "Timed out."
- | string_for_failure OutOfResources = "The ATP ran out of resources."
- | string_for_failure SpassTooOld =
- "Isabelle requires a more recent version of SPASS with support for the \
- \TPTP syntax. To install it, download and extract the package \
- \\"http://isabelle.in.tum.de/dist/contrib/spass-3.7.tar.gz\" and add the \
- \\"spass-3.7\" directory's absolute path to " ^
- quote (Path.implode (Path.expand (Path.appends
- (Path.variable "ISABELLE_HOME_USER" ::
- map Path.basic ["etc", "components"])))) ^
- " on a line of its own."
- | string_for_failure VampireTooOld =
- "Isabelle requires a more recent version of Vampire. To install it, follow \
- \the instructions from the Sledgehammer manual (\"isabelle doc\
- \ sledgehammer\")."
- | string_for_failure NoPerl = "Perl" ^ missing_message_tail
- | string_for_failure NoLibwwwPerl =
- "The Perl module \"libwww-perl\"" ^ missing_message_tail
- | string_for_failure MalformedInput =
- "The ATP problem is malformed. Please report this to the Isabelle \
- \developers."
- | string_for_failure MalformedOutput = "The ATP output is malformed."
- | string_for_failure Interrupted = "The ATP was interrupted."
- | string_for_failure InternalError = "An internal ATP error occurred."
- | string_for_failure UnknownError = "An unknown ATP error occurred."
-
-fun known_failure_in_output output =
- find_first (fn (_, pattern) => String.isSubstring pattern output)
- #> Option.map fst
-
val known_perl_failures =
[(CantConnect, "HTTP error"),
(NoPerl, "env: perl"),
@@ -124,6 +76,7 @@
fun to_secs bonus time = (Time.toMilliseconds time + bonus + 999) div 1000
+
(* E prover *)
(* Give older versions of E an extra second, because the "eproof" script wrongly
@@ -163,6 +116,8 @@
val e = ("e", e_config)
+(* SPASS *)
+
(* The "-VarWeight=3" option helps the higher-order problems, probably by
counteracting the presence of "hAPP". *)
val spass_config : prover_config =
@@ -226,11 +181,11 @@
fun get_systems () =
case bash_output "\"$ISABELLE_ATP/scripts/remote_atp\" -w 2>&1" of
- (answer, 0) => split_lines answer
- | (answer, _) =>
- error (case known_failure_in_output answer known_perl_failures of
+ (output, 0) => split_lines output
+ | (output, _) =>
+ error (case extract_known_failure known_perl_failures output of
SOME failure => string_for_failure failure
- | NONE => perhaps (try (unsuffix "\n")) answer ^ ".")
+ | NONE => perhaps (try (unsuffix "\n")) output ^ ".")
fun refresh_systems_on_tptp () =
Synchronized.change systems (fn _ => get_systems ())
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Datatype/datatype_selectors.ML Fri Sep 17 17:31:20 2010 +0200
@@ -0,0 +1,83 @@
+(* Title: HOL/Tools/Datatype/datatype_selectors.ML
+ Author: Sascha Boehme, TU Muenchen
+
+Selector functions for datatype constructor arguments.
+*)
+
+signature DATATYPE_SELECTORS =
+sig
+ val add_selector: ((string * typ) * int) * (string * typ) ->
+ Context.generic -> Context.generic
+ val lookup_selector: Proof.context -> string * int -> string option
+ val setup: theory -> theory
+end
+
+structure Datatype_Selectors: DATATYPE_SELECTORS =
+struct
+
+structure Stringinttab = Table
+(
+ type key = string * int
+ val ord = prod_ord fast_string_ord int_ord
+)
+
+structure Data = Generic_Data
+(
+ type T = string Stringinttab.table
+ val empty = Stringinttab.empty
+ val extend = I
+ val merge = Stringinttab.merge (K true)
+)
+
+fun pretty_term context = Syntax.pretty_term (Context.proof_of context)
+
+fun sanity_check context (((con as (n, _), i), sel as (m, _))) =
+ let
+ val thy = Context.theory_of context
+ val varify_const =
+ Const #> Type.varify_global [] #> snd #> Term.dest_Const #>
+ snd #> Term.strip_type
+
+ val (Ts, T) = varify_const con
+ val (Us, U) = varify_const sel
+ val _ = (0 < i andalso i <= length Ts) orelse
+ error (Pretty.string_of (Pretty.block [
+ Pretty.str "The constructor ",
+ Pretty.quote (pretty_term context (Const con)),
+ Pretty.str " has no argument position ",
+ Pretty.str (string_of_int i),
+ Pretty.str "."]))
+ val _ = length Us = 1 orelse
+ error (Pretty.string_of (Pretty.block [
+ Pretty.str "The term ", Pretty.quote (pretty_term context (Const sel)),
+ Pretty.str " might not be a selector ",
+ Pretty.str "(it accepts more than one argument)."]))
+ val _ =
+ (Sign.typ_equiv thy (T, hd Us) andalso
+ Sign.typ_equiv thy (nth Ts (i-1), U)) orelse
+ error (Pretty.string_of (Pretty.block [
+ Pretty.str "The types of the constructor ",
+ Pretty.quote (pretty_term context (Const con)),
+ Pretty.str " and of the selector ",
+ Pretty.quote (pretty_term context (Const sel)),
+ Pretty.str " do not fit to each other."]))
+ in ((n, i), m) end
+
+fun add_selector (entry as ((con as (n, _), i), (_, T))) context =
+ (case Stringinttab.lookup (Data.get context) (n, i) of
+ NONE => Data.map (Stringinttab.update (sanity_check context entry)) context
+ | SOME c => error (Pretty.string_of (Pretty.block [
+ Pretty.str "There is already a selector assigned to constructor ",
+ Pretty.quote (pretty_term context (Const con)), Pretty.str ", namely ",
+ Pretty.quote (pretty_term context (Const (c, T))), Pretty.str "."])))
+
+fun lookup_selector ctxt = Stringinttab.lookup (Data.get (Context.Proof ctxt))
+
+val setup =
+ Attrib.setup @{binding selector}
+ ((Args.term >> Term.dest_Const) -- Scan.lift (Parse.nat) --|
+ Scan.lift (Parse.$$$ "=") -- (Args.term >> Term.dest_Const) >>
+ (Thm.declaration_attribute o K o add_selector))
+ "assign a selector function to a datatype constructor argument"
+
+end
--- a/src/HOL/Tools/Nitpick/kodkod.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Tools/Nitpick/kodkod.ML Fri Sep 17 17:31:20 2010 +0200
@@ -2,7 +2,7 @@
Author: Jasmin Blanchette, TU Muenchen
Copyright 2008, 2009, 2010
-ML interface on top of Kodkod.
+ML interface for Kodkod.
*)
signature KODKOD =
--- a/src/HOL/Tools/Nitpick/nitpick_kodkod.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_kodkod.ML Fri Sep 17 17:31:20 2010 +0200
@@ -1019,7 +1019,7 @@
fun kodkod_formula_from_nut ofs
(kk as {kk_all, kk_exist, kk_formula_let, kk_formula_if, kk_or, kk_not,
kk_iff, kk_implies, kk_and, kk_subset, kk_rel_eq, kk_no,
- kk_lone, kk_one, kk_some, kk_rel_let, kk_rel_if, kk_union,
+ kk_lone, kk_some, kk_rel_let, kk_rel_if, kk_union,
kk_difference, kk_intersect, kk_product, kk_join, kk_closure,
kk_comprehension, kk_project, kk_project_seq, kk_not3,
kk_nat_less, kk_int_less, ...}) u =
--- a/src/HOL/Tools/Predicate_Compile/code_prolog.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/code_prolog.ML Fri Sep 17 17:31:20 2010 +0200
@@ -12,12 +12,13 @@
limited_types : (typ * int) list,
limited_predicates : (string list * int) list,
replacing : ((string * string) * string) list,
- manual_reorder : ((string * int) * int list) list,
- timeout : Time.time,
- prolog_system : prolog_system}
+ manual_reorder : ((string * int) * int list) list}
+ val set_ensure_groundness : code_options -> code_options
+ val map_limit_predicates : ((string list * int) list -> (string list * int) list)
+ -> code_options -> code_options
val code_options_of : theory -> code_options
val map_code_options : (code_options -> code_options) -> theory -> theory
-
+
datatype arith_op = Plus | Minus
datatype prol_term = Var of string | Cons of string | AppF of string * prol_term list
| Number of int | ArithOp of arith_op * prol_term list;
@@ -30,12 +31,13 @@
type clause = ((string * prol_term list) * prem);
type logic_program = clause list;
type constant_table = (string * string) list
-
- val generate : bool -> Proof.context -> string -> (logic_program * constant_table)
+
+ val generate : Predicate_Compile_Aux.mode option * bool ->
+ Proof.context -> string -> (logic_program * constant_table)
val write_program : logic_program -> string
- val run : (Time.time * prolog_system) -> logic_program -> string -> string list -> int option -> prol_term list list
-
- val quickcheck : Proof.context -> bool -> term -> int -> term list option * (bool list * bool)
+ val run : (Time.time * prolog_system) -> logic_program -> (string * prol_term list) -> string list -> int option -> prol_term list list
+
+ val quickcheck : Proof.context -> term -> int -> term list option * (bool list * bool)
val trace : bool Unsynchronized.ref
@@ -53,44 +55,70 @@
(* code generation options *)
-datatype prolog_system = SWI_PROLOG | YAP
type code_options =
{ensure_groundness : bool,
limited_types : (typ * int) list,
limited_predicates : (string list * int) list,
replacing : ((string * string) * string) list,
- manual_reorder : ((string * int) * int list) list,
- timeout : Time.time,
- prolog_system : prolog_system}
+ manual_reorder : ((string * int) * int list) list}
+
+fun set_ensure_groundness {ensure_groundness, limited_types, limited_predicates,
+ replacing, manual_reorder} =
+ {ensure_groundness = true, limited_types = limited_types,
+ limited_predicates = limited_predicates, replacing = replacing,
+ manual_reorder = manual_reorder}
+
+fun map_limit_predicates f {ensure_groundness, limited_types, limited_predicates,
+ replacing, manual_reorder} =
+ {ensure_groundness = ensure_groundness, limited_types = limited_types,
+ limited_predicates = f limited_predicates, replacing = replacing,
+ manual_reorder = manual_reorder}
+
structure Options = Theory_Data
(
type T = code_options
val empty = {ensure_groundness = false,
- limited_types = [], limited_predicates = [], replacing = [], manual_reorder = [],
- timeout = Time.fromSeconds 10, prolog_system = SWI_PROLOG}
+ limited_types = [], limited_predicates = [], replacing = [], manual_reorder = []}
val extend = I;
fun merge
({ensure_groundness = ensure_groundness1, limited_types = limited_types1,
limited_predicates = limited_predicates1, replacing = replacing1,
- manual_reorder = manual_reorder1, timeout = timeout1, prolog_system = prolog_system1},
+ manual_reorder = manual_reorder1},
{ensure_groundness = ensure_groundness2, limited_types = limited_types2,
limited_predicates = limited_predicates2, replacing = replacing2,
- manual_reorder = manual_reorder2, timeout = timeout2, prolog_system = prolog_system2}) =
+ manual_reorder = manual_reorder2}) =
{ensure_groundness = ensure_groundness1 orelse ensure_groundness2,
limited_types = AList.merge (op =) (K true) (limited_types1, limited_types2),
limited_predicates = AList.merge (op =) (K true) (limited_predicates1, limited_predicates2),
manual_reorder = AList.merge (op =) (K true) (manual_reorder1, manual_reorder2),
- replacing = Library.merge (op =) (replacing1, replacing2),
- timeout = timeout1,
- prolog_system = prolog_system1};
+ replacing = Library.merge (op =) (replacing1, replacing2)};
);
val code_options_of = Options.get
val map_code_options = Options.map
+(* system configuration *)
+
+datatype prolog_system = SWI_PROLOG | YAP
+
+fun string_of_system SWI_PROLOG = "swiprolog"
+ | string_of_system YAP = "yap"
+
+type system_configuration = {timeout : Time.time, prolog_system : prolog_system}
+
+structure System_Config = Generic_Data
+(
+ type T = system_configuration
+ val empty = {timeout = Time.fromSeconds 10, prolog_system = SWI_PROLOG}
+ val extend = I;
+ fun merge ({timeout = timeout1, prolog_system = prolog_system1},
+ {timeout = timeout2, prolog_system = prolog_system2}) =
+ {timeout = timeout1, prolog_system = prolog_system1}
+)
+
(* general string functions *)
val first_upper = implode o nth_map 0 Symbol.to_ascii_upper o explode;
@@ -176,7 +204,6 @@
type constant_table = (string * string) list
-(* assuming no clashing *)
fun declare_consts consts constant_table =
let
fun update' c table =
@@ -281,7 +308,6 @@
val intros = map (preprocess_intro (ProofContext.theory_of ctxt)) (Graph.get_node gr const)
val (intros', ctxt') = Variable.import_terms true (map prop_of intros) ctxt
val constant_table' = declare_consts (fold Term.add_const_names intros' []) constant_table
- |> declare_consts [@{const_name "Groups.zero_class.zero"}, @{const_name "Suc"}]
fun translate_intro intro =
let
val head = HOLogic.dest_Trueprop (Logic.strip_imp_concl intro)
@@ -316,18 +342,153 @@
tracing (cat_lines (map (fn const =>
"Constant " ^ const ^ "has intros:\n" ^
cat_lines (map (Display.string_of_thm ctxt) (Graph.get_node gr const))) consts))
-
-fun generate ensure_groundness ctxt const =
+
+(* translation of moded predicates *)
+
+(** generating graph of moded predicates **)
+
+(* could be moved to Predicate_Compile_Core *)
+fun requires_modes polarity cls =
+ let
+ fun req_mode_of pol (t, derivation) =
+ (case fst (strip_comb t) of
+ Const (c, _) => SOME (c, (pol, Predicate_Compile_Core.head_mode_of derivation))
+ | _ => NONE)
+ fun req (Predicate_Compile_Aux.Prem t, derivation) = req_mode_of polarity (t, derivation)
+ | req (Predicate_Compile_Aux.Negprem t, derivation) = req_mode_of (not polarity) (t, derivation)
+ | req _ = NONE
+ in
+ maps (fn (_, prems) => map_filter req prems) cls
+ end
+
+structure Mode_Graph = Graph(type key = string * (bool * Predicate_Compile_Aux.mode)
+ val ord = prod_ord fast_string_ord (prod_ord bool_ord Predicate_Compile_Aux.mode_ord));
+
+fun mk_moded_clauses_graph ctxt scc gr =
+ let
+ val options = Predicate_Compile_Aux.default_options
+ val mode_analysis_options =
+ {use_random = true, reorder_premises = true, infer_pos_and_neg_modes = true}
+ fun infer prednames (gr, (pos_modes, neg_modes, random)) =
+ let
+ val (lookup_modes, lookup_neg_modes, needs_random) =
+ ((fn s => the (AList.lookup (op =) pos_modes s)),
+ (fn s => the (AList.lookup (op =) neg_modes s)),
+ (fn s => member (op =) (the (AList.lookup (op =) random s))))
+ val (preds, all_vs, param_vs, all_modes, clauses) =
+ Predicate_Compile_Core.prepare_intrs options ctxt prednames
+ (maps (Predicate_Compile_Core.intros_of ctxt) prednames)
+ val ((moded_clauses, random'), _) =
+ Predicate_Compile_Core.infer_modes mode_analysis_options options
+ (lookup_modes, lookup_neg_modes, needs_random) ctxt preds all_modes param_vs clauses
+ val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
+ val pos_modes' = map (apsnd (map_filter (fn (true, m) => SOME m | _ => NONE))) modes
+ val neg_modes' = map (apsnd (map_filter (fn (false, m) => SOME m | _ => NONE))) modes
+ val _ = tracing ("Inferred modes:\n" ^
+ cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
+ (fn (p, m) => Predicate_Compile_Aux.string_of_mode m ^ (if p then "pos" else "neg")) ms)) modes))
+ val gr' = gr
+ |> fold (fn (p, mps) => fold (fn (mode, cls) =>
+ Mode_Graph.new_node ((p, mode), cls)) mps)
+ moded_clauses
+ |> fold (fn (p, mps) => fold (fn (mode, cls) => fold (fn req =>
+ Mode_Graph.add_edge ((p, mode), req)) (requires_modes (fst mode) cls)) mps)
+ moded_clauses
+ in
+ (gr', (AList.merge (op =) (op =) (pos_modes, pos_modes'),
+ AList.merge (op =) (op =) (neg_modes, neg_modes'),
+ AList.merge (op =) (op =) (random, random')))
+ end
+ in
+ fst (fold infer (rev scc) (Mode_Graph.empty, ([], [], [])))
+ end
+
+fun declare_moded_predicate moded_preds table =
+ let
+ fun update' (p as (pred, (pol, mode))) table =
+ if AList.defined (op =) table p then table else
+ let
+ val name = Long_Name.base_name pred ^ (if pol then "p" else "n")
+ ^ Predicate_Compile_Aux.ascii_string_of_mode mode
+ val p' = mk_conform first_lower "pred" (map snd table) name
+ in
+ AList.update (op =) (p, p') table
+ end
+ in
+ fold update' moded_preds table
+ end
+
+fun mk_program ctxt moded_gr moded_preds (prog, (moded_pred_table, constant_table)) =
+ let
+ val moded_pred_table' = declare_moded_predicate moded_preds moded_pred_table
+ fun mk_literal pol derivation constant_table' t =
+ let
+ val (p, args) = strip_comb t
+ val mode = Predicate_Compile_Core.head_mode_of derivation
+ val name = fst (dest_Const p)
+
+ val p' = the (AList.lookup (op =) moded_pred_table' (name, (pol, mode)))
+ val args' = map (translate_term ctxt constant_table') args
+ in
+ Rel (p', args')
+ end
+ fun mk_prem pol (indprem, derivation) constant_table =
+ case indprem of
+ Predicate_Compile_Aux.Generator (s, T) => (Ground (s, T), constant_table)
+ | _ =>
+ declare_consts (Term.add_const_names (Predicate_Compile_Aux.dest_indprem indprem) []) constant_table
+ |> (fn constant_table' =>
+ (case indprem of Predicate_Compile_Aux.Negprem t =>
+ NegRel_of (mk_literal (not pol) derivation constant_table' t)
+ | _ =>
+ mk_literal pol derivation constant_table' (Predicate_Compile_Aux.dest_indprem indprem), constant_table'))
+ fun mk_clause pred_name pol (ts, prems) (prog, constant_table) =
+ let
+ val constant_table' = declare_consts (fold Term.add_const_names ts []) constant_table
+ val args = map (translate_term ctxt constant_table') ts
+ val (prems', constant_table'') = fold_map (mk_prem pol) prems constant_table'
+ in
+ (((pred_name, args), Conj prems') :: prog, constant_table'')
+ end
+ fun mk_clauses (pred, mode as (pol, _)) =
+ let
+ val clauses = Mode_Graph.get_node moded_gr (pred, mode)
+ val pred_name = the (AList.lookup (op =) moded_pred_table' (pred, mode))
+ in
+ fold (mk_clause pred_name pol) clauses
+ end
+ in
+ apsnd (pair moded_pred_table') (fold mk_clauses moded_preds (prog, constant_table))
+ end
+
+fun generate (use_modes, ensure_groundness) ctxt const =
let
fun strong_conn_of gr keys =
Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
val gr = Predicate_Compile_Core.intros_graph_of ctxt
val gr' = add_edges depending_preds_of const gr
val scc = strong_conn_of gr' [const]
- val _ = print_intros ctxt gr (flat scc)
- val constant_table = declare_consts (flat scc) []
+ val initial_constant_table =
+ declare_consts [@{const_name "Groups.zero_class.zero"}, @{const_name "Suc"}] []
in
- apfst flat (fold_map (translate_intros ensure_groundness ctxt gr) (flat scc) constant_table)
+ case use_modes of
+ SOME mode =>
+ let
+ val moded_gr = mk_moded_clauses_graph ctxt scc gr
+ val moded_gr' = Mode_Graph.subgraph
+ (member (op =) (Mode_Graph.all_succs moded_gr [(const, (true, mode))])) moded_gr
+ val scc = Mode_Graph.strong_conn moded_gr'
+ in
+ apfst rev (apsnd snd
+ (fold (mk_program ctxt moded_gr') (rev scc) ([], ([], initial_constant_table))))
+ end
+ | NONE =>
+ let
+ val _ = print_intros ctxt gr (flat scc)
+ val constant_table = declare_consts (flat scc) initial_constant_table
+ in
+ apfst flat (fold_map (translate_intros ensure_groundness ctxt gr) (flat scc) constant_table)
+ end
end
(* implementation for fully enumerating predicates and
@@ -480,6 +641,7 @@
in
fst (fold_map reorder' p [])
end
+
(* rename variables to prolog-friendly names *)
fun rename_vars_term renaming = map_vars (fn v => the (AList.lookup (op =) renaming v))
@@ -521,6 +683,7 @@
| write_prem (NotEq (l, r)) = write_term l ^ " \\= " ^ write_term r
| write_prem (ArithEq (l, r)) = write_term l ^ " is " ^ write_term r
| write_prem (NotArithEq (l, r)) = write_term l ^ " =\\= " ^ write_term r
+ | write_prem _ = raise Fail "Not a valid prolog premise"
fun write_clause (head, prem) =
write_rel head ^ (if prem = Conj [] then "." else " :- " ^ write_prem prem ^ ".")
@@ -532,14 +695,14 @@
(** query and prelude for swi-prolog **)
-fun swi_prolog_query_first rel vnames =
- "eval :- once(" ^ rel ^ "(" ^ space_implode ", " vnames ^ ")),\n" ^
+fun swi_prolog_query_first (rel, args) vnames =
+ "eval :- once(" ^ rel ^ "(" ^ space_implode ", " (map write_term args) ^ ")),\n" ^
"writef('" ^ space_implode ";" (map (fn v => v ^ " = %w") vnames) ^
"\\n', [" ^ space_implode ", " vnames ^ "]).\n"
-fun swi_prolog_query_firstn n rel vnames =
+fun swi_prolog_query_firstn n (rel, args) vnames =
"eval :- findnsols(" ^ string_of_int n ^ ", (" ^ space_implode ", " vnames ^ "), " ^
- rel ^ "(" ^ space_implode ", " vnames ^ "), Sols), writelist(Sols).\n" ^
+ rel ^ "(" ^ space_implode ", " (map write_term args) ^ "), Sols), writelist(Sols).\n" ^
"writelist([]).\n" ^
"writelist([(" ^ space_implode ", " vnames ^ ")|T]) :- " ^
"writef('" ^ space_implode ";" (map (fn v => v ^ " = %w") vnames) ^
@@ -556,8 +719,8 @@
(** query and prelude for yap **)
-fun yap_query_first rel vnames =
- "eval :- once(" ^ rel ^ "(" ^ space_implode ", " vnames ^ ")),\n" ^
+fun yap_query_first (rel, args) vnames =
+ "eval :- once(" ^ rel ^ "(" ^ space_implode ", " (map write_term args) ^ ")),\n" ^
"format('" ^ space_implode ";" (map (fn v => v ^ " = ~w") vnames) ^
"\\n', [" ^ space_implode ", " vnames ^ "]).\n"
@@ -580,9 +743,16 @@
fun invoke system file_name =
let
+ val env_var =
+ (case system of SWI_PROLOG => "EXEC_SWIPL"| YAP => "EXEC_YAP")
+ val prog = getenv env_var
val cmd =
- case system of SWI_PROLOG => "/usr/local/bin/swipl -f " | YAP => "/usr/local/bin/yap -L "
- in fst (bash_output (cmd ^ file_name)) end
+ case system of SWI_PROLOG => prog ^ " -f " | YAP => prog ^ " -L "
+ in
+ if prog = "" then
+ (warning (env_var ^ " not set; could not execute code for " ^ string_of_system system); "")
+ else fst (bash_output (cmd ^ file_name))
+ end
(* parsing prolog solution *)
@@ -637,13 +807,14 @@
(* calling external interpreter and getting results *)
-fun run (timeout, system) p query_rel vnames nsols =
+fun run (timeout, system) p (query_rel, args) vnames nsols =
let
val p' = rename_vars_program p
val _ = tracing "Renaming variable names..."
- val renaming = fold mk_renaming vnames []
+ val renaming = fold mk_renaming (fold add_vars args vnames) []
val vnames' = map (fn v => the (AList.lookup (op =) renaming v)) vnames
- val prog = prelude system ^ query system nsols query_rel vnames' ^ write_program p'
+ val args' = map (rename_vars_term renaming) args
+ val prog = prelude system ^ query system nsols (query_rel, args') vnames' ^ write_program p'
val _ = tracing ("Generated prolog program:\n" ^ prog)
val solution = TimeLimit.timeLimit timeout (fn prog => Cache_IO.with_tmp_file "prolog_file" (fn prolog_file =>
(File.write prolog_file prog; invoke system (Path.implode prolog_file)))) prog
@@ -672,6 +843,17 @@
map (restore_term ctxt constant_table) (args ~~ argsT'))
end
+
+(* restore numerals in natural numbers *)
+
+fun restore_nat_numerals t =
+ if fastype_of t = @{typ nat} andalso is_some (try HOLogic.dest_nat t) then
+ HOLogic.mk_number @{typ nat} (HOLogic.dest_nat t)
+ else
+ (case t of
+ t1 $ t2 => restore_nat_numerals t1 $ restore_nat_numerals t2
+ | t => t)
+
(* values command *)
val preprocess_options = Predicate_Compile_Aux.Options {
@@ -711,10 +893,6 @@
case strip_comb body of
(Const (name, T), all_args) => (Const (name, T), all_args)
| (head, _) => error ("Not a constant: " ^ Syntax.string_of_term ctxt head)
- val vnames =
- case try (map (fst o dest_Free)) all_args of
- SOME vs => vs
- | NONE => error ("Not only free variables in " ^ commas (map (Syntax.string_of_term ctxt) all_args))
val _ = tracing "Preprocessing specification..."
val T = Sign.the_const_type (ProofContext.theory_of ctxt) name
val t = Const (name, T)
@@ -723,16 +901,19 @@
|> Predicate_Compile.preprocess preprocess_options t
val ctxt' = ProofContext.init_global thy'
val _ = tracing "Generating prolog program..."
- val (p, constant_table) = generate (#ensure_groundness options) ctxt' name
+ val (p, constant_table) = generate (NONE, #ensure_groundness options) ctxt' name (* FIXME *)
|> (if #ensure_groundness options then
add_ground_predicates ctxt' (#limited_types options)
else I)
|> add_limited_predicates (#limited_predicates options)
|> apfst (fold replace (#replacing options))
|> apfst (reorder_manually (#manual_reorder options))
+ val constant_table' = declare_consts (fold Term.add_const_names all_args []) constant_table
+ val args' = map (translate_term ctxt constant_table') all_args
val _ = tracing "Running prolog program..."
- val tss = run (#timeout options, #prolog_system options)
- p (translate_const constant_table name) (map first_upper vnames) soln
+ val system_config = System_Config.get (Context.Proof ctxt)
+ val tss = run (#timeout system_config, #prolog_system system_config)
+ p (translate_const constant_table' name, args') output_names soln
val _ = tracing "Restoring terms..."
val empty = Const("Orderings.bot_class.bot", fastype_of t_compr)
fun mk_insert x S =
@@ -759,7 +940,8 @@
end
in
foldl1 (HOLogic.mk_binop @{const_name sup}) (mk_set_compr []
- (map (fn ts => HOLogic.mk_tuple (map (restore_term ctxt' constant_table) (ts ~~ Ts))) tss) [])
+ (map (fn ts => HOLogic.mk_tuple
+ (map (restore_nat_numerals o restore_term ctxt' constant_table) (ts ~~ Ts))) tss) [])
end
fun values_cmd print_modes soln raw_t state =
@@ -798,7 +980,7 @@
fun strip_horn A = (strip_imp_prems A, strip_imp_concl A);
-fun quickcheck ctxt report t size =
+fun quickcheck ctxt t size =
let
val options = code_options_of (ProofContext.theory_of ctxt)
val thy = Theory.copy (ProofContext.theory_of ctxt)
@@ -821,14 +1003,15 @@
val thy3 = Predicate_Compile.preprocess preprocess_options const thy2
val ctxt' = ProofContext.init_global thy3
val _ = tracing "Generating prolog program..."
- val (p, constant_table) = generate true ctxt' full_constname
+ val (p, constant_table) = generate (NONE, true) ctxt' full_constname
|> add_ground_predicates ctxt' (#limited_types options)
|> add_limited_predicates (#limited_predicates options)
|> apfst (fold replace (#replacing options))
|> apfst (reorder_manually (#manual_reorder options))
val _ = tracing "Running prolog program..."
- val tss = run (#timeout options, #prolog_system options)
- p (translate_const constant_table full_constname) (map fst vs') (SOME 1)
+ val system_config = System_Config.get (Context.Proof ctxt)
+ val tss = run (#timeout system_config, #prolog_system system_config)
+ p (translate_const constant_table full_constname, map (Var o fst) vs') (map fst vs') (SOME 1)
val _ = tracing "Restoring terms..."
val res =
case tss of
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/etc/settings Fri Sep 17 17:31:20 2010 +0200
@@ -0,0 +1,13 @@
+
+EXEC_SWIPL=$(choosefrom \
+ "$ISABELLE_HOME/contrib/swipl/bin/swipl" \
+ "$ISABELLE_HOME/contrib_devel/swipl/bin/swipl" \
+ "$ISABELLE_HOME/../swipl" \
+ $(type -p swipl) \
+ "")
+
+EXEC_YAP=$(choosefrom \
+ "$ISABELLE_HOME/contrib/yap" \
+ "$ISABELLE_HOME/../yap" \
+ $(type -p yap) \
+ "")
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Fri Sep 17 17:31:20 2010 +0200
@@ -71,6 +71,7 @@
type mode_analysis_options = {use_random : bool, reorder_premises : bool, infer_pos_and_neg_modes : bool}
datatype mode_derivation = Mode_App of mode_derivation * mode_derivation | Context of mode
| Mode_Pair of mode_derivation * mode_derivation | Term of mode
+ val head_mode_of : mode_derivation -> mode
type moded_clause = term list * (Predicate_Compile_Aux.indprem * mode_derivation) list
type 'a pred_mode_table = (string * ((bool * mode) * 'a) list) list
@@ -3277,16 +3278,16 @@
val [nrandom, size, depth] = arguments
in
rpair NONE (fst (DSequence.yieldn k
- (Code_Runtime.value NONE (Dseq_Random_Result.get, put_dseq_random_result, "Predicate_Compile_Core.put_dseq_random_result")
- (fn proc => fn g => fn nrandom => fn size => fn s => g nrandom size s |>> DSequence.map proc)
- thy t' [] nrandom size
+ (Code_Runtime.dynamic_value_strict (Dseq_Random_Result.get, put_dseq_random_result, "Predicate_Compile_Core.put_dseq_random_result")
+ thy NONE (fn proc => fn g => fn nrandom => fn size => fn s => g nrandom size s |>> DSequence.map proc)
+ t' [] nrandom size
|> Random_Engine.run)
depth true))
end
| DSeq =>
rpair NONE (fst (DSequence.yieldn k
- (Code_Runtime.value NONE (Dseq_Result.get, put_dseq_result, "Predicate_Compile_Core.put_dseq_result")
- DSequence.map thy t' []) (the_single arguments) true))
+ (Code_Runtime.dynamic_value_strict (Dseq_Result.get, put_dseq_result, "Predicate_Compile_Core.put_dseq_result")
+ thy NONE DSequence.map t' []) (the_single arguments) true))
| New_Pos_Random_DSeq =>
let
val [nrandom, size, depth] = arguments
@@ -3295,23 +3296,25 @@
if stats then
apsnd (SOME o accumulate) (split_list
(fst (Lazy_Sequence.yieldn k
- (Code_Runtime.value NONE
+ (Code_Runtime.dynamic_value_strict
(Lseq_Random_Stats_Result.get, put_lseq_random_stats_result, "Predicate_Compile_Core.put_lseq_random_stats_result")
+ thy NONE
(fn proc => fn g => fn nrandom => fn size => fn s => fn depth => g nrandom size s depth
|> Lazy_Sequence.mapa (apfst proc))
- thy t' [] nrandom size seed depth))))
+ t' [] nrandom size seed depth))))
else rpair NONE
(fst (Lazy_Sequence.yieldn k
- (Code_Runtime.value NONE
+ (Code_Runtime.dynamic_value_strict
(Lseq_Random_Result.get, put_lseq_random_result, "Predicate_Compile_Core.put_lseq_random_result")
+ thy NONE
(fn proc => fn g => fn nrandom => fn size => fn s => fn depth => g nrandom size s depth
|> Lazy_Sequence.mapa proc)
- thy t' [] nrandom size seed depth)))
+ t' [] nrandom size seed depth)))
end
| _ =>
rpair NONE (fst (Predicate.yieldn k
- (Code_Runtime.value NONE (Pred_Result.get, put_pred_result, "Predicate_Compile_Core.put_pred_result")
- Predicate.map thy t' [])))
+ (Code_Runtime.dynamic_value_strict (Pred_Result.get, put_pred_result, "Predicate_Compile_Core.put_pred_result")
+ thy NONE Predicate.map t' [])))
in ((T, ts), statistics) end;
fun values ctxt param_user_modes ((raw_expected, stats), comp_options) k t_compr =
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML Fri Sep 17 17:31:20 2010 +0200
@@ -333,11 +333,20 @@
((full_constname, [definition])::new_defs, thy'))
end
| replace_abs_arg arg (new_defs, thy) =
- if (is_predT (fastype_of arg)) then
+ if is_some (try HOLogic.dest_prodT (fastype_of arg)) then
+ (case try HOLogic.dest_prod arg of
+ SOME (t1, t2) =>
+ (new_defs, thy)
+ |> process constname t1
+ ||>> process constname t2
+ |>> HOLogic.mk_prod
+ | NONE => (warning ("Replacing higher order arguments " ^
+ "is not applied in an undestructable product type"); (arg, (new_defs, thy))))
+ else if (is_predT (fastype_of arg)) then
process constname arg (new_defs, thy)
else
(arg, (new_defs, thy))
-
+
val (args', (new_defs', thy')) = fold_map replace_abs_arg
(map Envir.beta_eta_contract args) (new_defs, thy)
in
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML Fri Sep 17 17:31:20 2010 +0200
@@ -272,9 +272,10 @@
Pos_Random_DSeq =>
let
val compiled_term =
- Code_Runtime.value (SOME target) (Dseq_Result.get, put_dseq_result, "Predicate_Compile_Quickcheck.put_dseq_result")
+ Code_Runtime.dynamic_value_strict (Dseq_Result.get, put_dseq_result, "Predicate_Compile_Quickcheck.put_dseq_result")
+ thy4 (SOME target)
(fn proc => fn g => fn n => fn size => fn s => g n size s |>> (DSequence.map o map) proc)
- thy4 qc_term []
+ qc_term []
in
(fn size => fn nrandom => fn depth =>
Option.map fst (DSequence.yield
@@ -283,11 +284,12 @@
| New_Pos_Random_DSeq =>
let
val compiled_term =
- Code_Runtime.value (SOME target)
+ Code_Runtime.dynamic_value_strict
(Lseq_Result.get, put_lseq_result, "Predicate_Compile_Quickcheck.put_lseq_result")
+ thy4 (SOME target)
(fn proc => fn g => fn nrandom => fn size => fn s => fn depth =>
g nrandom size s depth |> (Lazy_Sequence.mapa o map) proc)
- thy4 qc_term []
+ qc_term []
in
fn size => fn nrandom => fn depth => Option.map fst (Lazy_Sequence.yield
(
@@ -297,11 +299,11 @@
end
| Depth_Limited_Random =>
let
- val compiled_term = Code_Runtime.value (SOME target)
+ val compiled_term = Code_Runtime.dynamic_value_strict
(Pred_Result.get, put_pred_result, "Predicate_Compile_Quickcheck.put_pred_result")
- (fn proc => fn g => fn depth => fn nrandom => fn size => fn seed =>
+ thy4 (SOME target) (fn proc => fn g => fn depth => fn nrandom => fn size => fn seed =>
g depth nrandom size seed |> (Predicate.map o map) proc)
- thy4 qc_term []
+ qc_term []
in
fn size => fn nrandom => fn depth => Option.map fst (Predicate.yield
(compiled_term depth nrandom size (Random_Engine.run (fn s => (s, s)))))
--- a/src/HOL/Tools/SMT/smt_normalize.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Tools/SMT/smt_normalize.ML Fri Sep 17 17:31:20 2010 +0200
@@ -9,14 +9,16 @@
* embed natural numbers into integers,
* add extra rules specifying types and constants which occur frequently,
* fully translate into object logic, add universal closure,
+ * monomorphize (create instances of schematic rules),
* lift lambda terms,
* make applications explicit for functions with varying number of arguments.
+ * add (hypothetical definitions for) missing datatype selectors,
*)
signature SMT_NORMALIZE =
sig
type extra_norm = thm list -> Proof.context -> thm list * Proof.context
- val normalize: extra_norm -> thm list -> Proof.context ->
+ val normalize: extra_norm -> bool -> thm list -> Proof.context ->
thm list * Proof.context
val atomize_conv: Proof.context -> conv
val eta_expand_conv: (Proof.context -> conv) -> Proof.context -> conv
@@ -427,13 +429,60 @@
+(* add missing datatype selectors via hypothetical definitions *)
+
+local
+ val add = (fn Type (n, _) => Symtab.update (n, ()) | _ => I)
+
+ fun collect t =
+ (case Term.strip_comb t of
+ (Abs (_, T, t), _) => add T #> collect t
+ | (Const (_, T), ts) => collects T ts
+ | (Free (_, T), ts) => collects T ts
+ | _ => I)
+ and collects T ts =
+ let val ((Ts, Us), U) = Term.strip_type T |> apfst (chop (length ts))
+ in fold add Ts #> add (Us ---> U) #> fold collect ts end
+
+ fun add_constructors thy n =
+ (case Datatype.get_info thy n of
+ NONE => I
+ | SOME {descr, ...} => fold (fn (_, (_, _, cs)) => fold (fn (n, ds) =>
+ fold (insert (op =) o pair n) (1 upto length ds)) cs) descr)
+
+ fun add_selector (c as (n, i)) ctxt =
+ (case Datatype_Selectors.lookup_selector ctxt c of
+ SOME _ => ctxt
+ | NONE =>
+ let
+ val T = Sign.the_const_type (ProofContext.theory_of ctxt) n
+ val U = Term.body_type T --> nth (Term.binder_types T) (i-1)
+ in
+ ctxt
+ |> yield_singleton Variable.variant_fixes Name.uu
+ |>> pair ((n, T), i) o rpair U
+ |-> Context.proof_map o Datatype_Selectors.add_selector
+ end)
+in
+
+fun datatype_selectors thms ctxt =
+ let
+ val ns = Symtab.keys (fold (collect o Thm.prop_of) thms Symtab.empty)
+ val cs = fold (add_constructors (ProofContext.theory_of ctxt)) ns []
+ in (thms, fold add_selector cs ctxt) end
+ (* FIXME: also generate hypothetical definitions for the selectors *)
+
+end
+
+
+
(* combined normalization *)
type extra_norm = thm list -> Proof.context -> thm list * Proof.context
fun with_context f thms ctxt = (f ctxt thms, ctxt)
-fun normalize extra_norm thms ctxt =
+fun normalize extra_norm with_datatypes thms ctxt =
thms
|> trivial_distinct ctxt
|> rewrite_bool_cases ctxt
@@ -445,5 +494,6 @@
|-> SMT_Monomorph.monomorph
|-> lift_lambdas
|-> with_context explicit_application
+ |-> (if with_datatypes then datatype_selectors else pair)
end
--- a/src/HOL/Tools/SMT/smt_solver.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Tools/SMT/smt_solver.ML Fri Sep 17 17:31:20 2010 +0200
@@ -195,9 +195,11 @@
("timeout: " ^ string_of_int (Config.get ctxt timeout)) ::
"arguments:" :: arguments
val {extra_norm, translate} = interface
+ val {builtins, ...} = translate
+ val {has_datatypes, ...} = builtins
in
(prems, ctxt)
- |-> SMT_Normalize.normalize extra_norm
+ |-> SMT_Normalize.normalize extra_norm has_datatypes
|-> invoke translate comments command arguments
|-> reconstruct
|-> (fn thm => fn ctxt' => thm
--- a/src/HOL/Tools/SMT/smt_translate.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Tools/SMT/smt_translate.ML Fri Sep 17 17:31:20 2010 +0200
@@ -291,28 +291,31 @@
SOME (f, _) => (f, cx)
| NONE => new_fun func_prefix t ss cx)
-fun inst_const f Ts t =
- let
- val (n, T) = Term.dest_Const (snd (Type.varify_global [] t))
- val inst = map Term.dest_TVar (snd (Term.dest_Type (f T))) ~~ Ts
- in Const (n, Term_Subst.instantiateT inst T) end
+fun mk_type (_, Tfs) (d as Datatype.DtTFree _) = the (AList.lookup (op =) Tfs d)
+ | mk_type Ts (Datatype.DtType (n, ds)) = Type (n, map (mk_type Ts) ds)
+ | mk_type (Tds, _) (Datatype.DtRec i) = nth Tds i
-fun inst_constructor Ts = inst_const Term.body_type Ts
-fun inst_selector Ts = inst_const Term.domain_type Ts
+fun mk_selector ctxt Ts T n (i, d) =
+ (case Datatype_Selectors.lookup_selector ctxt (n, i+1) of
+ NONE => raise Fail ("missing selector for datatype constructor " ^ quote n)
+ | SOME m => mk_type Ts d |> (fn U => (Const (m, T --> U), U)))
+
+fun mk_constructor ctxt Ts T (n, args) =
+ let val (sels, Us) = split_list (map_index (mk_selector ctxt Ts T n) args)
+ in (Const (n, Us ---> T), sels) end
-fun lookup_datatype ctxt n Ts = (* FIXME: use Datatype/Record.get_info *)
- if n = @{type_name prod}
- then SOME [
- (Type (n, Ts), [
- (inst_constructor Ts @{term Pair},
- map (inst_selector Ts) [@{term fst}, @{term snd}])])]
- else if n = @{type_name list}
- then SOME [
- (Type (n, Ts), [
- (inst_constructor Ts @{term Nil}, []),
- (inst_constructor Ts @{term Cons},
- map (inst_selector Ts) [@{term hd}, @{term tl}])])]
- else NONE
+fun lookup_datatype ctxt n Ts = if n = @{type_name bool} then NONE else
+ (case Datatype.get_info (ProofContext.theory_of ctxt) n of
+ NONE => NONE (* FIXME: also use Record.get_info *)
+ | SOME {descr, ...} =>
+ let
+ val Tds = map (fn (_, (tn, _, _)) => Type (tn, Ts))
+ (sort (int_ord o pairself fst) descr)
+ val Tfs = (case hd descr of (_, (_, tfs, _)) => tfs ~~ Ts)
+ in
+ SOME (descr |> map (fn (i, (_, _, cs)) =>
+ (nth Tds i, map (mk_constructor ctxt (Tds, Tfs) (nth Tds i)) cs)))
+ end)
fun relaxed thms = (([], thms), map prop_of thms)
--- a/src/HOL/Tools/Sledgehammer/metis_clauses.ML Fri Sep 17 17:11:43 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,524 +0,0 @@
-(* Title: HOL/Tools/Sledgehammer/metis_clauses.ML
- Author: Jia Meng, Cambridge University Computer Laboratory and NICTA
- Author: Jasmin Blanchette, TU Muenchen
-
-Storing/printing FOL clauses and arity clauses. Typed equality is
-treated differently.
-*)
-
-signature METIS_CLAUSES =
-sig
- type name = string * string
- datatype type_literal =
- TyLitVar of name * name |
- TyLitFree of name * name
- datatype arLit =
- TConsLit of name * name * name list |
- TVarLit of name * name
- datatype arity_clause =
- ArityClause of {name: string, conclLit: arLit, premLits: arLit list}
- datatype class_rel_clause =
- ClassRelClause of {name: string, subclass: name, superclass: name}
- datatype combtyp =
- CombTVar of name |
- CombTFree of name |
- CombType of name * combtyp list
- datatype combterm =
- CombConst of name * combtyp * combtyp list (* Const and Free *) |
- CombVar of name * combtyp |
- CombApp of combterm * combterm
- datatype fol_literal = FOLLiteral of bool * combterm
-
- val type_wrapper_name : string
- val bound_var_prefix : string
- val schematic_var_prefix: string
- val fixed_var_prefix: string
- val tvar_prefix: string
- val tfree_prefix: string
- val const_prefix: string
- val type_const_prefix: string
- val class_prefix: string
- val invert_const: string -> string
- val ascii_of: string -> string
- val unascii_of: string -> string
- val strip_prefix_and_unascii: string -> string -> string option
- val make_bound_var : string -> string
- val make_schematic_var : string * int -> string
- val make_fixed_var : string -> string
- val make_schematic_type_var : string * int -> string
- val make_fixed_type_var : string -> string
- val make_fixed_const : string -> string
- val make_fixed_type_const : string -> string
- val make_type_class : string -> string
- val skolem_theory_name: string
- val skolem_prefix: string
- val skolem_infix: string
- val is_skolem_const_name: string -> bool
- val num_type_args: theory -> string -> int
- val type_literals_for_types : typ list -> type_literal list
- val make_class_rel_clauses: theory -> class list -> class list -> class_rel_clause list
- val make_arity_clauses: theory -> string list -> class list -> class list * arity_clause list
- val combtyp_of : combterm -> combtyp
- val strip_combterm_comb : combterm -> combterm * combterm list
- val combterm_from_term :
- theory -> (string * typ) list -> term -> combterm * typ list
- val literals_of_term : theory -> term -> fol_literal list * typ list
- val conceal_skolem_terms :
- int -> (string * term) list -> term -> (string * term) list * term
- val reveal_skolem_terms : (string * term) list -> term -> term
- val tfree_classes_of_terms : term list -> string list
- val tvar_classes_of_terms : term list -> string list
- val type_consts_of_terms : theory -> term list -> string list
-end
-
-structure Metis_Clauses : METIS_CLAUSES =
-struct
-
-val type_wrapper_name = "ti"
-
-val bound_var_prefix = "B_"
-val schematic_var_prefix = "V_"
-val fixed_var_prefix = "v_"
-
-val tvar_prefix = "T_";
-val tfree_prefix = "t_";
-
-val const_prefix = "c_";
-val type_const_prefix = "tc_";
-val class_prefix = "class_";
-
-fun union_all xss = fold (union (op =)) xss []
-
-(* Readable names for the more common symbolic functions. Do not mess with the
- last nine entries of the table unless you know what you are doing. *)
-val const_trans_table =
- Symtab.make [(@{type_name Product_Type.prod}, "prod"),
- (@{type_name Sum_Type.sum}, "sum"),
- (@{const_name HOL.eq}, "equal"),
- (@{const_name HOL.conj}, "and"),
- (@{const_name HOL.disj}, "or"),
- (@{const_name HOL.implies}, "implies"),
- (@{const_name Set.member}, "member"),
- (@{const_name fequal}, "fequal"),
- (@{const_name COMBI}, "COMBI"),
- (@{const_name COMBK}, "COMBK"),
- (@{const_name COMBB}, "COMBB"),
- (@{const_name COMBC}, "COMBC"),
- (@{const_name COMBS}, "COMBS"),
- (@{const_name True}, "True"),
- (@{const_name False}, "False"),
- (@{const_name If}, "If")]
-
-(* Invert the table of translations between Isabelle and ATPs. *)
-val const_trans_table_inv =
- Symtab.update ("fequal", @{const_name HOL.eq})
- (Symtab.make (map swap (Symtab.dest const_trans_table)))
-
-val invert_const = perhaps (Symtab.lookup const_trans_table_inv)
-
-(*Escaping of special characters.
- Alphanumeric characters are left unchanged.
- The character _ goes to __
- Characters in the range ASCII space to / go to _A to _P, respectively.
- Other characters go to _nnn where nnn is the decimal ASCII code.*)
-val A_minus_space = Char.ord #"A" - Char.ord #" ";
-
-fun stringN_of_int 0 _ = ""
- | stringN_of_int k n = stringN_of_int (k-1) (n div 10) ^ Int.toString (n mod 10);
-
-fun ascii_of_c c =
- if Char.isAlphaNum c then String.str c
- else if c = #"_" then "__"
- else if #" " <= c andalso c <= #"/"
- then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space))
- else ("_" ^ stringN_of_int 3 (Char.ord c)) (*fixed width, in case more digits follow*)
-
-val ascii_of = String.translate ascii_of_c;
-
-(** Remove ASCII armouring from names in proof files **)
-
-(*We don't raise error exceptions because this code can run inside the watcher.
- Also, the errors are "impossible" (hah!)*)
-fun unascii_aux rcs [] = String.implode(rev rcs)
- | unascii_aux rcs [#"_"] = unascii_aux (#"_"::rcs) [] (*ERROR*)
- (*Three types of _ escapes: __, _A to _P, _nnn*)
- | unascii_aux rcs (#"_" :: #"_" :: cs) = unascii_aux (#"_"::rcs) cs
- | unascii_aux rcs (#"_" :: c :: cs) =
- if #"A" <= c andalso c<= #"P" (*translation of #" " to #"/"*)
- then unascii_aux (Char.chr(Char.ord c - A_minus_space) :: rcs) cs
- else
- let val digits = List.take (c::cs, 3) handle Subscript => []
- in
- case Int.fromString (String.implode digits) of
- NONE => unascii_aux (c:: #"_"::rcs) cs (*ERROR*)
- | SOME n => unascii_aux (Char.chr n :: rcs) (List.drop (cs, 2))
- end
- | unascii_aux rcs (c::cs) = unascii_aux (c::rcs) cs
-val unascii_of = unascii_aux [] o String.explode
-
-(* If string s has the prefix s1, return the result of deleting it,
- un-ASCII'd. *)
-fun strip_prefix_and_unascii s1 s =
- if String.isPrefix s1 s then
- SOME (unascii_of (String.extract (s, size s1, NONE)))
- else
- NONE
-
-(*Remove the initial ' character from a type variable, if it is present*)
-fun trim_type_var s =
- if s <> "" andalso String.sub(s,0) = #"'" then String.extract(s,1,NONE)
- else error ("trim_type: Malformed type variable encountered: " ^ s);
-
-fun ascii_of_indexname (v,0) = ascii_of v
- | ascii_of_indexname (v,i) = ascii_of v ^ "_" ^ Int.toString i;
-
-fun make_bound_var x = bound_var_prefix ^ ascii_of x
-fun make_schematic_var v = schematic_var_prefix ^ ascii_of_indexname v
-fun make_fixed_var x = fixed_var_prefix ^ ascii_of x
-
-fun make_schematic_type_var (x,i) =
- tvar_prefix ^ (ascii_of_indexname (trim_type_var x,i));
-fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (trim_type_var x));
-
-fun lookup_const c =
- case Symtab.lookup const_trans_table c of
- SOME c' => c'
- | NONE => ascii_of c
-
-(* HOL.eq MUST BE "equal" because it's built into ATPs. *)
-fun make_fixed_const @{const_name HOL.eq} = "equal"
- | make_fixed_const c = const_prefix ^ lookup_const c
-
-fun make_fixed_type_const c = type_const_prefix ^ lookup_const c
-
-fun make_type_class clas = class_prefix ^ ascii_of clas;
-
-val skolem_theory_name = "Sledgehammer" ^ Long_Name.separator ^ "Sko"
-val skolem_prefix = "sko_"
-val skolem_infix = "$"
-
-(* Hack: Could return false positives (e.g., a user happens to declare a
- constant called "SomeTheory.sko_means_shoe_in_$wedish". *)
-val is_skolem_const_name =
- Long_Name.base_name
- #> String.isPrefix skolem_prefix andf String.isSubstring skolem_infix
-
-(* The number of type arguments of a constant, zero if it's monomorphic. For
- (instances of) Skolem pseudoconstants, this information is encoded in the
- constant name. *)
-fun num_type_args thy s =
- if String.isPrefix skolem_theory_name s then
- s |> unprefix skolem_theory_name
- |> space_explode skolem_infix |> hd
- |> space_explode "_" |> List.last |> Int.fromString |> the
- else
- (s, Sign.the_const_type thy s) |> Sign.const_typargs thy |> length
-
-(**** Definitions and functions for FOL clauses for TPTP format output ****)
-
-type name = string * string
-
-(**** Isabelle FOL clauses ****)
-
-(* The first component is the type class; the second is a TVar or TFree. *)
-datatype type_literal =
- TyLitVar of name * name |
- TyLitFree of name * name
-
-exception CLAUSE of string * term;
-
-(*Make literals for sorted type variables*)
-fun sorts_on_typs_aux (_, []) = []
- | sorts_on_typs_aux ((x,i), s::ss) =
- let val sorts = sorts_on_typs_aux ((x,i), ss)
- in
- if s = "HOL.type" then sorts
- else if i = ~1 then TyLitFree (`make_type_class s, `make_fixed_type_var x) :: sorts
- else TyLitVar (`make_type_class s, (make_schematic_type_var (x,i), x)) :: sorts
- end;
-
-fun sorts_on_typs (TFree (a,s)) = sorts_on_typs_aux ((a,~1),s)
- | sorts_on_typs (TVar (v,s)) = sorts_on_typs_aux (v,s);
-
-(*Given a list of sorted type variables, return a list of type literals.*)
-fun type_literals_for_types Ts =
- fold (union (op =)) (map sorts_on_typs Ts) []
-
-(** make axiom and conjecture clauses. **)
-
-(**** Isabelle arities ****)
-
-datatype arLit =
- TConsLit of name * name * name list |
- TVarLit of name * name
-
-datatype arity_clause =
- ArityClause of {name: string, conclLit: arLit, premLits: arLit list}
-
-
-fun gen_TVars 0 = []
- | gen_TVars n = ("T_" ^ Int.toString n) :: gen_TVars (n-1);
-
-fun pack_sort(_,[]) = []
- | pack_sort(tvar, "HOL.type"::srt) = pack_sort (tvar, srt) (*IGNORE sort "type"*)
- | pack_sort(tvar, cls::srt) =
- (`make_type_class cls, (tvar, tvar)) :: pack_sort (tvar, srt)
-
-(*Arity of type constructor tcon :: (arg1,...,argN)res*)
-fun make_axiom_arity_clause (tcons, name, (cls,args)) =
- let
- val tvars = gen_TVars (length args)
- val tvars_srts = ListPair.zip (tvars, args)
- in
- ArityClause {name = name,
- conclLit = TConsLit (`make_type_class cls,
- `make_fixed_type_const tcons,
- tvars ~~ tvars),
- premLits = map TVarLit (union_all (map pack_sort tvars_srts))}
- end
-
-
-(**** Isabelle class relations ****)
-
-datatype class_rel_clause =
- ClassRelClause of {name: string, subclass: name, superclass: name}
-
-(*Generate all pairs (sub,super) such that sub is a proper subclass of super in theory thy.*)
-fun class_pairs _ [] _ = []
- | class_pairs thy subs supers =
- let
- val class_less = Sorts.class_less (Sign.classes_of thy)
- fun add_super sub super = class_less (sub, super) ? cons (sub, super)
- fun add_supers sub = fold (add_super sub) supers
- in fold add_supers subs [] end
-
-fun make_class_rel_clause (sub,super) =
- ClassRelClause {name = sub ^ "_" ^ super,
- subclass = `make_type_class sub,
- superclass = `make_type_class super}
-
-fun make_class_rel_clauses thy subs supers =
- map make_class_rel_clause (class_pairs thy subs supers);
-
-
-(** Isabelle arities **)
-
-fun arity_clause _ _ (_, []) = []
- | arity_clause seen n (tcons, ("HOL.type",_)::ars) = (*ignore*)
- arity_clause seen n (tcons,ars)
- | arity_clause seen n (tcons, (ar as (class,_)) :: ars) =
- if member (op =) seen class then (*multiple arities for the same tycon, class pair*)
- make_axiom_arity_clause (tcons, lookup_const tcons ^ "_" ^ class ^ "_" ^ Int.toString n, ar) ::
- arity_clause seen (n+1) (tcons,ars)
- else
- make_axiom_arity_clause (tcons, lookup_const tcons ^ "_" ^ class, ar) ::
- arity_clause (class::seen) n (tcons,ars)
-
-fun multi_arity_clause [] = []
- | multi_arity_clause ((tcons, ars) :: tc_arlists) =
- arity_clause [] 1 (tcons, ars) @ multi_arity_clause tc_arlists
-
-(*Generate all pairs (tycon,class,sorts) such that tycon belongs to class in theory thy
- provided its arguments have the corresponding sorts.*)
-fun type_class_pairs thy tycons classes =
- let val alg = Sign.classes_of thy
- fun domain_sorts tycon = Sorts.mg_domain alg tycon o single
- fun add_class tycon class =
- cons (class, domain_sorts tycon class)
- handle Sorts.CLASS_ERROR _ => I
- fun try_classes tycon = (tycon, fold (add_class tycon) classes [])
- in map try_classes tycons end;
-
-(*Proving one (tycon, class) membership may require proving others, so iterate.*)
-fun iter_type_class_pairs _ _ [] = ([], [])
- | iter_type_class_pairs thy tycons classes =
- let val cpairs = type_class_pairs thy tycons classes
- val newclasses = union_all (union_all (union_all (map (map #2 o #2) cpairs)))
- |> subtract (op =) classes |> subtract (op =) HOLogic.typeS
- val (classes', cpairs') = iter_type_class_pairs thy tycons newclasses
- in (union (op =) classes' classes, union (op =) cpairs' cpairs) end;
-
-fun make_arity_clauses thy tycons classes =
- let val (classes', cpairs) = iter_type_class_pairs thy tycons classes
- in (classes', multi_arity_clause cpairs) end;
-
-datatype combtyp =
- CombTVar of name |
- CombTFree of name |
- CombType of name * combtyp list
-
-datatype combterm =
- CombConst of name * combtyp * combtyp list (* Const and Free *) |
- CombVar of name * combtyp |
- CombApp of combterm * combterm
-
-datatype fol_literal = FOLLiteral of bool * combterm
-
-(*********************************************************************)
-(* convert a clause with type Term.term to a clause with type clause *)
-(*********************************************************************)
-
-(*Result of a function type; no need to check that the argument type matches.*)
-fun result_type (CombType (_, [_, tp2])) = tp2
- | result_type _ = raise Fail "non-function type"
-
-fun combtyp_of (CombConst (_, tp, _)) = tp
- | combtyp_of (CombVar (_, tp)) = tp
- | combtyp_of (CombApp (t1, _)) = result_type (combtyp_of t1)
-
-(*gets the head of a combinator application, along with the list of arguments*)
-fun strip_combterm_comb u =
- let fun stripc (CombApp(t,u), ts) = stripc (t, u::ts)
- | stripc x = x
- in stripc(u,[]) end
-
-fun type_of (Type (a, Ts)) =
- let val (folTypes,ts) = types_of Ts in
- (CombType (`make_fixed_type_const a, folTypes), ts)
- end
- | type_of (tp as TFree (a, _)) = (CombTFree (`make_fixed_type_var a), [tp])
- | type_of (tp as TVar (x, _)) =
- (CombTVar (make_schematic_type_var x, string_of_indexname x), [tp])
-and types_of Ts =
- let val (folTyps, ts) = ListPair.unzip (map type_of Ts) in
- (folTyps, union_all ts)
- end
-
-(* same as above, but no gathering of sort information *)
-fun simp_type_of (Type (a, Ts)) =
- CombType (`make_fixed_type_const a, map simp_type_of Ts)
- | simp_type_of (TFree (a, _)) = CombTFree (`make_fixed_type_var a)
- | simp_type_of (TVar (x, _)) =
- CombTVar (make_schematic_type_var x, string_of_indexname x)
-
-(* Converts a term (with combinators) into a combterm. Also accummulates sort
- infomation. *)
-fun combterm_from_term thy bs (P $ Q) =
- let val (P', tsP) = combterm_from_term thy bs P
- val (Q', tsQ) = combterm_from_term thy bs Q
- in (CombApp (P', Q'), union (op =) tsP tsQ) end
- | combterm_from_term thy _ (Const (c, T)) =
- let
- val (tp, ts) = type_of T
- val tvar_list =
- (if String.isPrefix skolem_theory_name c then
- [] |> Term.add_tvarsT T |> map TVar
- else
- (c, T) |> Sign.const_typargs thy)
- |> map simp_type_of
- val c' = CombConst (`make_fixed_const c, tp, tvar_list)
- in (c',ts) end
- | combterm_from_term _ _ (Free (v, T)) =
- let val (tp,ts) = type_of T
- val v' = CombConst (`make_fixed_var v, tp, [])
- in (v',ts) end
- | combterm_from_term _ _ (Var (v, T)) =
- let val (tp,ts) = type_of T
- val v' = CombVar ((make_schematic_var v, string_of_indexname v), tp)
- in (v',ts) end
- | combterm_from_term _ bs (Bound j) =
- let
- val (s, T) = nth bs j
- val (tp, ts) = type_of T
- val v' = CombConst (`make_bound_var s, tp, [])
- in (v', ts) end
- | combterm_from_term _ _ (Abs _) = raise Fail "HOL clause: Abs"
-
-fun predicate_of thy ((@{const Not} $ P), pos) = predicate_of thy (P, not pos)
- | predicate_of thy (t, pos) =
- (combterm_from_term thy [] (Envir.eta_contract t), pos)
-
-fun literals_of_term1 args thy (@{const Trueprop} $ P) =
- literals_of_term1 args thy P
- | literals_of_term1 args thy (@{const HOL.disj} $ P $ Q) =
- literals_of_term1 (literals_of_term1 args thy P) thy Q
- | literals_of_term1 (lits, ts) thy P =
- let val ((pred, ts'), pol) = predicate_of thy (P, true) in
- (FOLLiteral (pol, pred) :: lits, union (op =) ts ts')
- end
-val literals_of_term = literals_of_term1 ([], [])
-
-fun skolem_name i j num_T_args =
- skolem_prefix ^ (space_implode "_" (map Int.toString [i, j, num_T_args])) ^
- skolem_infix ^ "g"
-
-fun conceal_skolem_terms i skolems t =
- if exists_Const (curry (op =) @{const_name skolem} o fst) t then
- let
- fun aux skolems
- (t as (Const (@{const_name skolem}, Type (_, [_, T])) $ _)) =
- let
- val (skolems, s) =
- if i = ~1 then
- (skolems, @{const_name undefined})
- else case AList.find (op aconv) skolems t of
- s :: _ => (skolems, s)
- | [] =>
- let
- val s = skolem_theory_name ^ "." ^
- skolem_name i (length skolems)
- (length (Term.add_tvarsT T []))
- in ((s, t) :: skolems, s) end
- in (skolems, Const (s, T)) end
- | aux skolems (t1 $ t2) =
- let
- val (skolems, t1) = aux skolems t1
- val (skolems, t2) = aux skolems t2
- in (skolems, t1 $ t2) end
- | aux skolems (Abs (s, T, t')) =
- let val (skolems, t') = aux skolems t' in
- (skolems, Abs (s, T, t'))
- end
- | aux skolems t = (skolems, t)
- in aux skolems t end
- else
- (skolems, t)
-
-fun reveal_skolem_terms skolems =
- map_aterms (fn t as Const (s, _) =>
- if String.isPrefix skolem_theory_name s then
- AList.lookup (op =) skolems s |> the
- |> map_types Type_Infer.paramify_vars
- else
- t
- | t => t)
-
-
-(***************************************************************)
-(* Type Classes Present in the Axiom or Conjecture Clauses *)
-(***************************************************************)
-
-fun set_insert (x, s) = Symtab.update (x, ()) s
-
-fun add_classes (sorts, cset) = List.foldl set_insert cset (flat sorts)
-
-(*Remove this trivial type class*)
-fun delete_type cset = Symtab.delete_safe (the_single @{sort HOL.type}) cset;
-
-fun tfree_classes_of_terms ts =
- let val sorts_list = map (map #2 o OldTerm.term_tfrees) ts
- in Symtab.keys (delete_type (List.foldl add_classes Symtab.empty sorts_list)) end;
-
-fun tvar_classes_of_terms ts =
- let val sorts_list = map (map #2 o OldTerm.term_tvars) ts
- in Symtab.keys (delete_type (List.foldl add_classes Symtab.empty sorts_list)) end;
-
-(*fold type constructors*)
-fun fold_type_consts f (Type (a, Ts)) x = fold (fold_type_consts f) Ts (f (a,x))
- | fold_type_consts _ _ x = x;
-
-(*Type constructors used to instantiate overloaded constants are the only ones needed.*)
-fun add_type_consts_in_term thy =
- let
- fun aux (Const x) =
- fold (fold_type_consts set_insert) (Sign.const_typargs thy x)
- | aux (Abs (_, _, u)) = aux u
- | aux (Const (@{const_name skolem}, _) $ _) = I
- | aux (t $ u) = aux t #> aux u
- | aux _ = I
- in aux end
-
-fun type_consts_of_terms thy ts =
- Symtab.keys (fold (add_type_consts_in_term thy) ts Symtab.empty);
-
-end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Sledgehammer/metis_reconstruct.ML Fri Sep 17 17:31:20 2010 +0200
@@ -0,0 +1,514 @@
+(* Title: HOL/Tools/Sledgehammer/metis_reconstruct.ML
+ Author: Kong W. Susanto, Cambridge University Computer Laboratory
+ Author: Lawrence C. Paulson, Cambridge University Computer Laboratory
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright Cambridge University 2007
+
+Proof reconstruction for Metis.
+*)
+
+signature METIS_RECONSTRUCT =
+sig
+ type mode = Metis_Translate.mode
+
+ val trace: bool Unsynchronized.ref
+ val lookth : (Metis_Thm.thm * 'a) list -> Metis_Thm.thm -> 'a
+ val replay_one_inference :
+ Proof.context -> mode -> (string * term) list
+ -> Metis_Thm.thm * Metis_Proof.inference -> (Metis_Thm.thm * thm) list
+ -> (Metis_Thm.thm * thm) list
+end;
+
+structure Metis_Reconstruct : METIS_RECONSTRUCT =
+struct
+
+open Metis_Translate
+
+val trace = Unsynchronized.ref false
+fun trace_msg msg = if !trace then tracing (msg ()) else ()
+
+datatype term_or_type = SomeTerm of term | SomeType of typ
+
+fun terms_of [] = []
+ | terms_of (SomeTerm t :: tts) = t :: terms_of tts
+ | terms_of (SomeType _ :: tts) = terms_of tts;
+
+fun types_of [] = []
+ | types_of (SomeTerm (Var ((a,idx), _)) :: tts) =
+ if String.isPrefix "_" a then
+ (*Variable generated by Metis, which might have been a type variable.*)
+ TVar (("'" ^ a, idx), HOLogic.typeS) :: types_of tts
+ else types_of tts
+ | types_of (SomeTerm _ :: tts) = types_of tts
+ | types_of (SomeType T :: tts) = T :: types_of tts;
+
+fun apply_list rator nargs rands =
+ let val trands = terms_of rands
+ in if length trands = nargs then SomeTerm (list_comb(rator, trands))
+ else raise Fail
+ ("apply_list: wrong number of arguments: " ^ Syntax.string_of_term_global Pure.thy rator ^
+ " expected " ^ Int.toString nargs ^
+ " received " ^ commas (map (Syntax.string_of_term_global Pure.thy) trands))
+ end;
+
+fun infer_types ctxt =
+ Syntax.check_terms (ProofContext.set_mode ProofContext.mode_pattern ctxt);
+
+(*We use 1 rather than 0 because variable references in clauses may otherwise conflict
+ with variable constraints in the goal...at least, type inference often fails otherwise.
+ SEE ALSO axiom_inf below.*)
+fun mk_var (w, T) = Var ((w, 1), T)
+
+(*include the default sort, if available*)
+fun mk_tfree ctxt w =
+ let val ww = "'" ^ w
+ in TFree(ww, the_default HOLogic.typeS (Variable.def_sort ctxt (ww, ~1))) end;
+
+(*Remove the "apply" operator from an HO term*)
+fun strip_happ args (Metis_Term.Fn(".",[t,u])) = strip_happ (u::args) t
+ | strip_happ args x = (x, args);
+
+fun make_tvar s = TVar (("'" ^ s, 0), HOLogic.typeS)
+
+fun smart_invert_const "fequal" = @{const_name HOL.eq}
+ | smart_invert_const s = invert_const s
+
+fun hol_type_from_metis_term _ (Metis_Term.Var v) =
+ (case strip_prefix_and_unascii tvar_prefix v of
+ SOME w => make_tvar w
+ | NONE => make_tvar v)
+ | hol_type_from_metis_term ctxt (Metis_Term.Fn(x, tys)) =
+ (case strip_prefix_and_unascii type_const_prefix x of
+ SOME tc => Type (smart_invert_const tc,
+ map (hol_type_from_metis_term ctxt) tys)
+ | NONE =>
+ case strip_prefix_and_unascii tfree_prefix x of
+ SOME tf => mk_tfree ctxt tf
+ | NONE => raise Fail ("hol_type_from_metis_term: " ^ x));
+
+(*Maps metis terms to isabelle terms*)
+fun hol_term_from_metis_PT ctxt fol_tm =
+ let val thy = ProofContext.theory_of ctxt
+ val _ = trace_msg (fn () => "hol_term_from_metis_PT: " ^
+ Metis_Term.toString fol_tm)
+ fun tm_to_tt (Metis_Term.Var v) =
+ (case strip_prefix_and_unascii tvar_prefix v of
+ SOME w => SomeType (make_tvar w)
+ | NONE =>
+ case strip_prefix_and_unascii schematic_var_prefix v of
+ SOME w => SomeTerm (mk_var (w, HOLogic.typeT))
+ | NONE => SomeTerm (mk_var (v, HOLogic.typeT)) )
+ (*Var from Metis with a name like _nnn; possibly a type variable*)
+ | tm_to_tt (Metis_Term.Fn ("{}", [arg])) = tm_to_tt arg (*hBOOL*)
+ | tm_to_tt (t as Metis_Term.Fn (".",_)) =
+ let val (rator,rands) = strip_happ [] t
+ in case rator of
+ Metis_Term.Fn(fname,ts) => applic_to_tt (fname, ts @ rands)
+ | _ => case tm_to_tt rator of
+ SomeTerm t => SomeTerm (list_comb(t, terms_of (map tm_to_tt rands)))
+ | _ => raise Fail "tm_to_tt: HO application"
+ end
+ | tm_to_tt (Metis_Term.Fn (fname, args)) = applic_to_tt (fname,args)
+ and applic_to_tt ("=",ts) =
+ SomeTerm (list_comb(Const (@{const_name HOL.eq}, HOLogic.typeT), terms_of (map tm_to_tt ts)))
+ | applic_to_tt (a,ts) =
+ case strip_prefix_and_unascii const_prefix a of
+ SOME b =>
+ let val c = smart_invert_const b
+ val ntypes = num_type_args thy c
+ val nterms = length ts - ntypes
+ val tts = map tm_to_tt ts
+ val tys = types_of (List.take(tts,ntypes))
+ in if length tys = ntypes then
+ apply_list (Const (c, dummyT)) nterms (List.drop(tts,ntypes))
+ else
+ raise Fail ("Constant " ^ c ^ " expects " ^ Int.toString ntypes ^
+ " but gets " ^ Int.toString (length tys) ^
+ " type arguments\n" ^
+ cat_lines (map (Syntax.string_of_typ ctxt) tys) ^
+ " the terms are \n" ^
+ cat_lines (map (Syntax.string_of_term ctxt) (terms_of tts)))
+ end
+ | NONE => (*Not a constant. Is it a type constructor?*)
+ case strip_prefix_and_unascii type_const_prefix a of
+ SOME b =>
+ SomeType (Type (smart_invert_const b, types_of (map tm_to_tt ts)))
+ | NONE => (*Maybe a TFree. Should then check that ts=[].*)
+ case strip_prefix_and_unascii tfree_prefix a of
+ SOME b => SomeType (mk_tfree ctxt b)
+ | NONE => (*a fixed variable? They are Skolem functions.*)
+ case strip_prefix_and_unascii fixed_var_prefix a of
+ SOME b =>
+ let val opr = Free (b, HOLogic.typeT)
+ in apply_list opr (length ts) (map tm_to_tt ts) end
+ | NONE => raise Fail ("unexpected metis function: " ^ a)
+ in
+ case tm_to_tt fol_tm of
+ SomeTerm t => t
+ | SomeType T => raise TYPE ("fol_tm_to_tt: Term expected", [T], [])
+ end
+
+(*Maps fully-typed metis terms to isabelle terms*)
+fun hol_term_from_metis_FT ctxt fol_tm =
+ let val _ = trace_msg (fn () => "hol_term_from_metis_FT: " ^
+ Metis_Term.toString fol_tm)
+ fun cvt (Metis_Term.Fn ("ti", [Metis_Term.Var v, _])) =
+ (case strip_prefix_and_unascii schematic_var_prefix v of
+ SOME w => mk_var(w, dummyT)
+ | NONE => mk_var(v, dummyT))
+ | cvt (Metis_Term.Fn ("ti", [Metis_Term.Fn ("=",[]), _])) =
+ Const (@{const_name HOL.eq}, HOLogic.typeT)
+ | cvt (Metis_Term.Fn ("ti", [Metis_Term.Fn (x,[]), ty])) =
+ (case strip_prefix_and_unascii const_prefix x of
+ SOME c => Const (smart_invert_const c, dummyT)
+ | NONE => (*Not a constant. Is it a fixed variable??*)
+ case strip_prefix_and_unascii fixed_var_prefix x of
+ SOME v => Free (v, hol_type_from_metis_term ctxt ty)
+ | NONE => raise Fail ("hol_term_from_metis_FT bad constant: " ^ x))
+ | cvt (Metis_Term.Fn ("ti", [Metis_Term.Fn (".",[tm1,tm2]), _])) =
+ cvt tm1 $ cvt tm2
+ | cvt (Metis_Term.Fn (".",[tm1,tm2])) = (*untyped application*)
+ cvt tm1 $ cvt tm2
+ | cvt (Metis_Term.Fn ("{}", [arg])) = cvt arg (*hBOOL*)
+ | cvt (Metis_Term.Fn ("=", [tm1,tm2])) =
+ list_comb(Const (@{const_name HOL.eq}, HOLogic.typeT), map cvt [tm1,tm2])
+ | cvt (t as Metis_Term.Fn (x, [])) =
+ (case strip_prefix_and_unascii const_prefix x of
+ SOME c => Const (smart_invert_const c, dummyT)
+ | NONE => (*Not a constant. Is it a fixed variable??*)
+ case strip_prefix_and_unascii fixed_var_prefix x of
+ SOME v => Free (v, dummyT)
+ | NONE => (trace_msg (fn () => "hol_term_from_metis_FT bad const: " ^ x);
+ hol_term_from_metis_PT ctxt t))
+ | cvt t = (trace_msg (fn () => "hol_term_from_metis_FT bad term: " ^ Metis_Term.toString t);
+ hol_term_from_metis_PT ctxt t)
+ in fol_tm |> cvt end
+
+fun hol_term_from_metis FT = hol_term_from_metis_FT
+ | hol_term_from_metis _ = hol_term_from_metis_PT
+
+fun hol_terms_from_fol ctxt mode skolems fol_tms =
+ let val ts = map (hol_term_from_metis mode ctxt) fol_tms
+ val _ = trace_msg (fn () => " calling type inference:")
+ val _ = app (fn t => trace_msg (fn () => Syntax.string_of_term ctxt t)) ts
+ val ts' = ts |> map (reveal_skolem_terms skolems) |> infer_types ctxt
+ val _ = app (fn t => trace_msg
+ (fn () => " final term: " ^ Syntax.string_of_term ctxt t ^
+ " of type " ^ Syntax.string_of_typ ctxt (type_of t)))
+ ts'
+ in ts' end;
+
+(* ------------------------------------------------------------------------- *)
+(* FOL step Inference Rules *)
+(* ------------------------------------------------------------------------- *)
+
+(*for debugging only*)
+(*
+fun print_thpair (fth,th) =
+ (trace_msg (fn () => "=============================================");
+ trace_msg (fn () => "Metis: " ^ Metis_Thm.toString fth);
+ trace_msg (fn () => "Isabelle: " ^ Display.string_of_thm_without_context th));
+*)
+
+fun lookth thpairs (fth : Metis_Thm.thm) =
+ the (AList.lookup (uncurry Metis_Thm.equal) thpairs fth)
+ handle Option.Option =>
+ raise Fail ("Failed to find Metis theorem " ^ Metis_Thm.toString fth)
+
+fun cterm_incr_types thy idx = cterm_of thy o (map_types (Logic.incr_tvar idx));
+
+(* INFERENCE RULE: AXIOM *)
+
+fun axiom_inf thpairs th = Thm.incr_indexes 1 (lookth thpairs th);
+ (*This causes variables to have an index of 1 by default. SEE ALSO mk_var above.*)
+
+(* INFERENCE RULE: ASSUME *)
+
+val EXCLUDED_MIDDLE = @{lemma "P ==> ~ P ==> False" by (rule notE)}
+
+fun inst_excluded_middle thy i_atm =
+ let val th = EXCLUDED_MIDDLE
+ val [vx] = Term.add_vars (prop_of th) []
+ val substs = [(cterm_of thy (Var vx), cterm_of thy i_atm)]
+ in cterm_instantiate substs th end;
+
+fun assume_inf ctxt mode skolems atm =
+ inst_excluded_middle
+ (ProofContext.theory_of ctxt)
+ (singleton (hol_terms_from_fol ctxt mode skolems) (Metis_Term.Fn atm))
+
+(* INFERENCE RULE: INSTANTIATE (SUBST). Type instantiations are ignored. Trying
+ to reconstruct them admits new possibilities of errors, e.g. concerning
+ sorts. Instead we try to arrange that new TVars are distinct and that types
+ can be inferred from terms. *)
+
+fun inst_inf ctxt mode skolems thpairs fsubst th =
+ let val thy = ProofContext.theory_of ctxt
+ val i_th = lookth thpairs th
+ val i_th_vars = Term.add_vars (prop_of i_th) []
+ fun find_var x = the (List.find (fn ((a,_),_) => a=x) i_th_vars)
+ fun subst_translation (x,y) =
+ let val v = find_var x
+ (* We call "reveal_skolem_terms" and "infer_types" below. *)
+ val t = hol_term_from_metis mode ctxt y
+ in SOME (cterm_of thy (Var v), t) end
+ handle Option.Option =>
+ (trace_msg (fn () => "\"find_var\" failed for " ^ x ^
+ " in " ^ Display.string_of_thm ctxt i_th);
+ NONE)
+ | TYPE _ =>
+ (trace_msg (fn () => "\"hol_term_from_metis\" failed for " ^ x ^
+ " in " ^ Display.string_of_thm ctxt i_th);
+ NONE)
+ fun remove_typeinst (a, t) =
+ case strip_prefix_and_unascii schematic_var_prefix a of
+ SOME b => SOME (b, t)
+ | NONE => case strip_prefix_and_unascii tvar_prefix a of
+ SOME _ => NONE (*type instantiations are forbidden!*)
+ | NONE => SOME (a,t) (*internal Metis var?*)
+ val _ = trace_msg (fn () => " isa th: " ^ Display.string_of_thm ctxt i_th)
+ val substs = map_filter remove_typeinst (Metis_Subst.toList fsubst)
+ val (vars,rawtms) = ListPair.unzip (map_filter subst_translation substs)
+ val tms = rawtms |> map (reveal_skolem_terms skolems) |> infer_types ctxt
+ val ctm_of = cterm_incr_types thy (1 + Thm.maxidx_of i_th)
+ val substs' = ListPair.zip (vars, map ctm_of tms)
+ val _ = trace_msg (fn () =>
+ cat_lines ("subst_translations:" ::
+ (substs' |> map (fn (x, y) =>
+ Syntax.string_of_term ctxt (term_of x) ^ " |-> " ^
+ Syntax.string_of_term ctxt (term_of y)))));
+ in cterm_instantiate substs' i_th end
+ handle THM (msg, _, _) =>
+ error ("Cannot replay Metis proof in Isabelle:\n" ^ msg)
+
+(* INFERENCE RULE: RESOLVE *)
+
+(* Like RSN, but we rename apart only the type variables. Vars here typically
+ have an index of 1, and the use of RSN would increase this typically to 3.
+ Instantiations of those Vars could then fail. See comment on "mk_var". *)
+fun resolve_inc_tyvars thy tha i thb =
+ let
+ val tha = Drule.incr_type_indexes (1 + Thm.maxidx_of thb) tha
+ fun aux tha thb =
+ case Thm.bicompose false (false, tha, nprems_of tha) i thb
+ |> Seq.list_of |> distinct Thm.eq_thm of
+ [th] => th
+ | _ => raise THM ("resolve_inc_tyvars: unique result expected", i,
+ [tha, thb])
+ in
+ aux tha thb
+ handle TERM z =>
+ (* The unifier, which is invoked from "Thm.bicompose", will sometimes
+ refuse to unify "?a::?'a" with "?a::?'b" or "?a::nat" and throw a
+ "TERM" exception (with "add_ffpair" as first argument). We then
+ perform unification of the types of variables by hand and try
+ again. We could do this the first time around but this error
+ occurs seldom and we don't want to break existing proofs in subtle
+ ways or slow them down needlessly. *)
+ case [] |> fold (Term.add_vars o prop_of) [tha, thb]
+ |> AList.group (op =)
+ |> maps (fn ((s, _), T :: Ts) =>
+ map (fn T' => (Free (s, T), Free (s, T'))) Ts)
+ |> rpair (Envir.empty ~1)
+ |-> fold (Pattern.unify thy)
+ |> Envir.type_env |> Vartab.dest
+ |> map (fn (x, (S, T)) =>
+ pairself (ctyp_of thy) (TVar (x, S), T)) of
+ [] => raise TERM z
+ | ps => aux (instantiate (ps, []) tha) (instantiate (ps, []) thb)
+ end
+
+fun mk_not (Const (@{const_name Not}, _) $ b) = b
+ | mk_not b = HOLogic.mk_not b
+
+(* Match untyped terms. *)
+fun untyped_aconv (Const (a, _)) (Const(b, _)) = (a = b)
+ | untyped_aconv (Free (a, _)) (Free (b, _)) = (a = b)
+ | untyped_aconv (Var ((a, _), _)) (Var ((b, _), _)) =
+ (a = b) (* The index is ignored, for some reason. *)
+ | untyped_aconv (Bound i) (Bound j) = (i = j)
+ | untyped_aconv (Abs (_, _, t)) (Abs (_, _, u)) = untyped_aconv t u
+ | untyped_aconv (t1 $ t2) (u1 $ u2) =
+ untyped_aconv t1 u1 andalso untyped_aconv t2 u2
+ | untyped_aconv _ _ = false
+
+(* Finding the relative location of an untyped term within a list of terms *)
+fun literal_index lit =
+ let
+ val lit = Envir.eta_contract lit
+ fun get _ [] = raise Empty
+ | get n (x :: xs) =
+ if untyped_aconv lit (Envir.eta_contract (HOLogic.dest_Trueprop x)) then
+ n
+ else
+ get (n+1) xs
+ in get 1 end
+
+fun resolve_inf ctxt mode skolems thpairs atm th1 th2 =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val i_th1 = lookth thpairs th1 and i_th2 = lookth thpairs th2
+ val _ = trace_msg (fn () => " isa th1 (pos): " ^ Display.string_of_thm ctxt i_th1)
+ val _ = trace_msg (fn () => " isa th2 (neg): " ^ Display.string_of_thm ctxt i_th2)
+ in
+ (* Trivial cases where one operand is type info *)
+ if Thm.eq_thm (TrueI, i_th1) then
+ i_th2
+ else if Thm.eq_thm (TrueI, i_th2) then
+ i_th1
+ else
+ let
+ val i_atm = singleton (hol_terms_from_fol ctxt mode skolems)
+ (Metis_Term.Fn atm)
+ val _ = trace_msg (fn () => " atom: " ^ Syntax.string_of_term ctxt i_atm)
+ val prems_th1 = prems_of i_th1
+ val prems_th2 = prems_of i_th2
+ val index_th1 = literal_index (mk_not i_atm) prems_th1
+ handle Empty => raise Fail "Failed to find literal in th1"
+ val _ = trace_msg (fn () => " index_th1: " ^ Int.toString index_th1)
+ val index_th2 = literal_index i_atm prems_th2
+ handle Empty => raise Fail "Failed to find literal in th2"
+ val _ = trace_msg (fn () => " index_th2: " ^ Int.toString index_th2)
+ in
+ resolve_inc_tyvars thy (Meson.select_literal index_th1 i_th1) index_th2
+ i_th2
+ end
+ end;
+
+(* INFERENCE RULE: REFL *)
+
+val REFL_THM = Thm.incr_indexes 2 @{lemma "t ~= t ==> False" by simp}
+
+val refl_x = cterm_of @{theory} (Var (hd (Term.add_vars (prop_of REFL_THM) [])));
+val refl_idx = 1 + Thm.maxidx_of REFL_THM;
+
+fun refl_inf ctxt mode skolems t =
+ let val thy = ProofContext.theory_of ctxt
+ val i_t = singleton (hol_terms_from_fol ctxt mode skolems) t
+ val _ = trace_msg (fn () => " term: " ^ Syntax.string_of_term ctxt i_t)
+ val c_t = cterm_incr_types thy refl_idx i_t
+ in cterm_instantiate [(refl_x, c_t)] REFL_THM end;
+
+(* INFERENCE RULE: EQUALITY *)
+
+val subst_em = @{lemma "s = t ==> P s ==> ~ P t ==> False" by simp}
+val ssubst_em = @{lemma "s = t ==> P t ==> ~ P s ==> False" by simp}
+
+val metis_eq = Metis_Term.Fn ("=", []);
+
+fun get_ty_arg_size _ (Const (@{const_name HOL.eq}, _)) = 0 (*equality has no type arguments*)
+ | get_ty_arg_size thy (Const (c, _)) = (num_type_args thy c handle TYPE _ => 0)
+ | get_ty_arg_size _ _ = 0;
+
+fun equality_inf ctxt mode skolems (pos, atm) fp fr =
+ let val thy = ProofContext.theory_of ctxt
+ val m_tm = Metis_Term.Fn atm
+ val [i_atm,i_tm] = hol_terms_from_fol ctxt mode skolems [m_tm, fr]
+ val _ = trace_msg (fn () => "sign of the literal: " ^ Bool.toString pos)
+ fun replace_item_list lx 0 (_::ls) = lx::ls
+ | replace_item_list lx i (l::ls) = l :: replace_item_list lx (i-1) ls
+ fun path_finder_FO tm [] = (tm, Bound 0)
+ | path_finder_FO tm (p::ps) =
+ let val (tm1,args) = strip_comb tm
+ val adjustment = get_ty_arg_size thy tm1
+ val p' = if adjustment > p then p else p-adjustment
+ val tm_p = List.nth(args,p')
+ handle Subscript =>
+ error ("Cannot replay Metis proof in Isabelle:\n" ^
+ "equality_inf: " ^ Int.toString p ^ " adj " ^
+ Int.toString adjustment ^ " term " ^
+ Syntax.string_of_term ctxt tm)
+ val _ = trace_msg (fn () => "path_finder: " ^ Int.toString p ^
+ " " ^ Syntax.string_of_term ctxt tm_p)
+ val (r,t) = path_finder_FO tm_p ps
+ in
+ (r, list_comb (tm1, replace_item_list t p' args))
+ end
+ fun path_finder_HO tm [] = (tm, Bound 0)
+ | path_finder_HO (t$u) (0::ps) = (fn(x,y) => (x, y$u)) (path_finder_HO t ps)
+ | path_finder_HO (t$u) (_::ps) = (fn(x,y) => (x, t$y)) (path_finder_HO u ps)
+ | path_finder_HO tm ps =
+ raise Fail ("Cannot replay Metis proof in Isabelle:\n" ^
+ "equality_inf, path_finder_HO: path = " ^
+ space_implode " " (map Int.toString ps) ^
+ " isa-term: " ^ Syntax.string_of_term ctxt tm)
+ fun path_finder_FT tm [] _ = (tm, Bound 0)
+ | path_finder_FT tm (0::ps) (Metis_Term.Fn ("ti", [t1, _])) =
+ path_finder_FT tm ps t1
+ | path_finder_FT (t$u) (0::ps) (Metis_Term.Fn (".", [t1, _])) =
+ (fn(x,y) => (x, y$u)) (path_finder_FT t ps t1)
+ | path_finder_FT (t$u) (1::ps) (Metis_Term.Fn (".", [_, t2])) =
+ (fn(x,y) => (x, t$y)) (path_finder_FT u ps t2)
+ | path_finder_FT tm ps t =
+ raise Fail ("Cannot replay Metis proof in Isabelle:\n" ^
+ "equality_inf, path_finder_FT: path = " ^
+ space_implode " " (map Int.toString ps) ^
+ " isa-term: " ^ Syntax.string_of_term ctxt tm ^
+ " fol-term: " ^ Metis_Term.toString t)
+ fun path_finder FO tm ps _ = path_finder_FO tm ps
+ | path_finder HO (tm as Const(@{const_name HOL.eq},_) $ _ $ _) (p::ps) _ =
+ (*equality: not curried, as other predicates are*)
+ if p=0 then path_finder_HO tm (0::1::ps) (*select first operand*)
+ else path_finder_HO tm (p::ps) (*1 selects second operand*)
+ | path_finder HO tm (_ :: ps) (Metis_Term.Fn ("{}", [_])) =
+ path_finder_HO tm ps (*if not equality, ignore head to skip hBOOL*)
+ | path_finder FT (tm as Const(@{const_name HOL.eq}, _) $ _ $ _) (p::ps)
+ (Metis_Term.Fn ("=", [t1,t2])) =
+ (*equality: not curried, as other predicates are*)
+ if p=0 then path_finder_FT tm (0::1::ps)
+ (Metis_Term.Fn (".", [Metis_Term.Fn (".", [metis_eq,t1]), t2]))
+ (*select first operand*)
+ else path_finder_FT tm (p::ps)
+ (Metis_Term.Fn (".", [metis_eq,t2]))
+ (*1 selects second operand*)
+ | path_finder FT tm (_ :: ps) (Metis_Term.Fn ("{}", [t1])) = path_finder_FT tm ps t1
+ (*if not equality, ignore head to skip the hBOOL predicate*)
+ | path_finder FT tm ps t = path_finder_FT tm ps t (*really an error case!*)
+ fun path_finder_lit ((nt as Const (@{const_name Not}, _)) $ tm_a) idx =
+ let val (tm, tm_rslt) = path_finder mode tm_a idx m_tm
+ in (tm, nt $ tm_rslt) end
+ | path_finder_lit tm_a idx = path_finder mode tm_a idx m_tm
+ val (tm_subst, body) = path_finder_lit i_atm fp
+ val tm_abs = Abs ("x", type_of tm_subst, body)
+ val _ = trace_msg (fn () => "abstraction: " ^ Syntax.string_of_term ctxt tm_abs)
+ val _ = trace_msg (fn () => "i_tm: " ^ Syntax.string_of_term ctxt i_tm)
+ val _ = trace_msg (fn () => "located term: " ^ Syntax.string_of_term ctxt tm_subst)
+ val imax = maxidx_of_term (i_tm $ tm_abs $ tm_subst) (*ill typed but gives right max*)
+ val subst' = Thm.incr_indexes (imax+1) (if pos then subst_em else ssubst_em)
+ val _ = trace_msg (fn () => "subst' " ^ Display.string_of_thm ctxt subst')
+ val eq_terms = map (pairself (cterm_of thy))
+ (ListPair.zip (OldTerm.term_vars (prop_of subst'), [tm_abs, tm_subst, i_tm]))
+ in cterm_instantiate eq_terms subst' end;
+
+val factor = Seq.hd o distinct_subgoals_tac;
+
+fun step ctxt mode skolems thpairs p =
+ case p of
+ (fol_th, Metis_Proof.Axiom _) => factor (axiom_inf thpairs fol_th)
+ | (_, Metis_Proof.Assume f_atm) => assume_inf ctxt mode skolems f_atm
+ | (_, Metis_Proof.Metis_Subst (f_subst, f_th1)) =>
+ factor (inst_inf ctxt mode skolems thpairs f_subst f_th1)
+ | (_, Metis_Proof.Resolve(f_atm, f_th1, f_th2)) =>
+ factor (resolve_inf ctxt mode skolems thpairs f_atm f_th1 f_th2)
+ | (_, Metis_Proof.Refl f_tm) => refl_inf ctxt mode skolems f_tm
+ | (_, Metis_Proof.Equality (f_lit, f_p, f_r)) =>
+ equality_inf ctxt mode skolems f_lit f_p f_r
+
+fun is_real_literal (_, (c, _)) = not (String.isPrefix class_prefix c)
+
+fun replay_one_inference ctxt mode skolems (fol_th, inf) thpairs =
+ let
+ val _ = trace_msg (fn () => "=============================================")
+ val _ = trace_msg (fn () => "METIS THM: " ^ Metis_Thm.toString fol_th)
+ val _ = trace_msg (fn () => "INFERENCE: " ^ Metis_Proof.inferenceToString inf)
+ val th = Meson.flexflex_first_order (step ctxt mode skolems
+ thpairs (fol_th, inf))
+ val _ = trace_msg (fn () => "ISABELLE THM: " ^ Display.string_of_thm ctxt th)
+ val _ = trace_msg (fn () => "=============================================")
+ val n_metis_lits =
+ length (filter is_real_literal (Metis_LiteralSet.toList (Metis_Thm.clause fol_th)))
+ val _ = if nprems_of th = n_metis_lits then ()
+ else error "Cannot replay Metis proof in Isabelle: Out of sync."
+ in (fol_th, th) :: thpairs end
+
+end;
--- a/src/HOL/Tools/Sledgehammer/metis_tactics.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/metis_tactics.ML Fri Sep 17 17:31:20 2010 +0200
@@ -9,744 +9,53 @@
signature METIS_TACTICS =
sig
- val trace: bool Unsynchronized.ref
- val type_lits: bool Config.T
- val metis_tac: Proof.context -> thm list -> int -> tactic
- val metisF_tac: Proof.context -> thm list -> int -> tactic
- val metisFT_tac: Proof.context -> thm list -> int -> tactic
- val setup: theory -> theory
+ val trace : bool Unsynchronized.ref
+ val type_lits : bool Config.T
+ val metis_tac : Proof.context -> thm list -> int -> tactic
+ val metisF_tac : Proof.context -> thm list -> int -> tactic
+ val metisFT_tac : Proof.context -> thm list -> int -> tactic
+ val setup : theory -> theory
end
structure Metis_Tactics : METIS_TACTICS =
struct
-open Metis_Clauses
+open Metis_Translate
+open Metis_Reconstruct
-val trace = Unsynchronized.ref false;
-fun trace_msg msg = if !trace then tracing (msg ()) else ();
+val trace = Unsynchronized.ref false
+fun trace_msg msg = if !trace then tracing (msg ()) else ()
val (type_lits, type_lits_setup) = Attrib.config_bool "metis_type_lits" (K true);
-datatype mode = FO | HO | FT (* first-order, higher-order, fully-typed *)
-
-(* ------------------------------------------------------------------------- *)
-(* Useful Theorems *)
-(* ------------------------------------------------------------------------- *)
-val EXCLUDED_MIDDLE = @{lemma "P ==> ~ P ==> False" by (rule notE)}
-val REFL_THM = Thm.incr_indexes 2 @{lemma "t ~= t ==> False" by simp}
-val subst_em = @{lemma "s = t ==> P s ==> ~ P t ==> False" by simp}
-val ssubst_em = @{lemma "s = t ==> P t ==> ~ P s ==> False" by simp}
-
-(* ------------------------------------------------------------------------- *)
-(* Useful Functions *)
-(* ------------------------------------------------------------------------- *)
-
-(* Match untyped terms. *)
-fun untyped_aconv (Const (a, _)) (Const(b, _)) = (a = b)
- | untyped_aconv (Free (a, _)) (Free (b, _)) = (a = b)
- | untyped_aconv (Var ((a, _), _)) (Var ((b, _), _)) =
- (a = b) (* The index is ignored, for some reason. *)
- | untyped_aconv (Bound i) (Bound j) = (i = j)
- | untyped_aconv (Abs (_, _, t)) (Abs (_, _, u)) = untyped_aconv t u
- | untyped_aconv (t1 $ t2) (u1 $ u2) =
- untyped_aconv t1 u1 andalso untyped_aconv t2 u2
- | untyped_aconv _ _ = false
-
-(* Finding the relative location of an untyped term within a list of terms *)
-fun get_index lit =
- let val lit = Envir.eta_contract lit
- fun get _ [] = raise Empty
- | get n (x::xs) = if untyped_aconv lit (Envir.eta_contract (HOLogic.dest_Trueprop x))
- then n else get (n+1) xs
- in get 1 end;
-
-(* ------------------------------------------------------------------------- *)
-(* HOL to FOL (Isabelle to Metis) *)
-(* ------------------------------------------------------------------------- *)
-
-fun fn_isa_to_met_sublevel "equal" = "=" (* FIXME: "c_fequal" *)
- | fn_isa_to_met_sublevel x = x
-fun fn_isa_to_met_toplevel "equal" = "="
- | fn_isa_to_met_toplevel x = x
-
-fun metis_lit b c args = (b, (c, args));
-
-fun metis_term_from_combtyp (CombTVar (s, _)) = Metis_Term.Var s
- | metis_term_from_combtyp (CombTFree (s, _)) = Metis_Term.Fn (s, [])
- | metis_term_from_combtyp (CombType ((s, _), tps)) =
- Metis_Term.Fn (s, map metis_term_from_combtyp tps);
-
-(*These two functions insert type literals before the real literals. That is the
- opposite order from TPTP linkup, but maybe OK.*)
-
-fun hol_term_to_fol_FO tm =
- case strip_combterm_comb tm of
- (CombConst ((c, _), _, tys), tms) =>
- let val tyargs = map metis_term_from_combtyp tys
- val args = map hol_term_to_fol_FO tms
- in Metis_Term.Fn (c, tyargs @ args) end
- | (CombVar ((v, _), _), []) => Metis_Term.Var v
- | _ => raise Fail "non-first-order combterm"
-
-fun hol_term_to_fol_HO (CombConst ((a, _), _, tylist)) =
- Metis_Term.Fn (fn_isa_to_met_sublevel a, map metis_term_from_combtyp tylist)
- | hol_term_to_fol_HO (CombVar ((s, _), _)) = Metis_Term.Var s
- | hol_term_to_fol_HO (CombApp (tm1, tm2)) =
- Metis_Term.Fn (".", map hol_term_to_fol_HO [tm1, tm2]);
-
-(*The fully-typed translation, to avoid type errors*)
-fun wrap_type (tm, ty) = Metis_Term.Fn("ti", [tm, metis_term_from_combtyp ty]);
-
-fun hol_term_to_fol_FT (CombVar ((s, _), ty)) = wrap_type (Metis_Term.Var s, ty)
- | hol_term_to_fol_FT (CombConst((a, _), ty, _)) =
- wrap_type (Metis_Term.Fn(fn_isa_to_met_sublevel a, []), ty)
- | hol_term_to_fol_FT (tm as CombApp(tm1,tm2)) =
- wrap_type (Metis_Term.Fn(".", map hol_term_to_fol_FT [tm1,tm2]),
- combtyp_of tm)
-
-fun hol_literal_to_fol FO (FOLLiteral (pos, tm)) =
- let val (CombConst((p, _), _, tys), tms) = strip_combterm_comb tm
- val tylits = if p = "equal" then [] else map metis_term_from_combtyp tys
- val lits = map hol_term_to_fol_FO tms
- in metis_lit pos (fn_isa_to_met_toplevel p) (tylits @ lits) end
- | hol_literal_to_fol HO (FOLLiteral (pos, tm)) =
- (case strip_combterm_comb tm of
- (CombConst(("equal", _), _, _), tms) =>
- metis_lit pos "=" (map hol_term_to_fol_HO tms)
- | _ => metis_lit pos "{}" [hol_term_to_fol_HO tm]) (*hBOOL*)
- | hol_literal_to_fol FT (FOLLiteral (pos, tm)) =
- (case strip_combterm_comb tm of
- (CombConst(("equal", _), _, _), tms) =>
- metis_lit pos "=" (map hol_term_to_fol_FT tms)
- | _ => metis_lit pos "{}" [hol_term_to_fol_FT tm]) (*hBOOL*);
-
-fun literals_of_hol_term thy mode t =
- let val (lits, types_sorts) = literals_of_term thy t
- in (map (hol_literal_to_fol mode) lits, types_sorts) end;
-
-(*Sign should be "true" for conjecture type constraints, "false" for type lits in clauses.*)
-fun metis_of_type_literals pos (TyLitVar ((s, _), (s', _))) =
- metis_lit pos s [Metis_Term.Var s']
- | metis_of_type_literals pos (TyLitFree ((s, _), (s', _))) =
- metis_lit pos s [Metis_Term.Fn (s',[])]
-
-fun default_sort _ (TVar _) = false
- | default_sort ctxt (TFree (x, s)) = (s = the_default [] (Variable.def_sort ctxt (x, ~1)));
-
-fun metis_of_tfree tf =
- Metis_Thm.axiom (Metis_LiteralSet.singleton (metis_of_type_literals true tf));
-
-fun hol_thm_to_fol is_conjecture ctxt mode j skolems th =
- let
- val thy = ProofContext.theory_of ctxt
- val (skolems, (mlits, types_sorts)) =
- th |> prop_of |> conceal_skolem_terms j skolems
- ||> (HOLogic.dest_Trueprop #> literals_of_hol_term thy mode)
- in
- if is_conjecture then
- (Metis_Thm.axiom (Metis_LiteralSet.fromList mlits),
- type_literals_for_types types_sorts, skolems)
- else
- let val tylits = filter_out (default_sort ctxt) types_sorts
- |> type_literals_for_types
- val mtylits = if Config.get ctxt type_lits
- then map (metis_of_type_literals false) tylits else []
- in
- (Metis_Thm.axiom (Metis_LiteralSet.fromList(mtylits @ mlits)), [],
- skolems)
- end
- end;
-
-(* ARITY CLAUSE *)
-
-fun m_arity_cls (TConsLit ((c, _), (t, _), args)) =
- metis_lit true c [Metis_Term.Fn(t, map (Metis_Term.Var o fst) args)]
- | m_arity_cls (TVarLit ((c, _), (s, _))) =
- metis_lit false c [Metis_Term.Var s]
-
-(*TrueI is returned as the Isabelle counterpart because there isn't any.*)
-fun arity_cls (ArityClause {conclLit, premLits, ...}) =
- (TrueI,
- Metis_Thm.axiom (Metis_LiteralSet.fromList (map m_arity_cls (conclLit :: premLits))));
-
-(* CLASSREL CLAUSE *)
-
-fun m_class_rel_cls (subclass, _) (superclass, _) =
- [metis_lit false subclass [Metis_Term.Var "T"], metis_lit true superclass [Metis_Term.Var "T"]];
-
-fun class_rel_cls (ClassRelClause {subclass, superclass, ...}) =
- (TrueI, Metis_Thm.axiom (Metis_LiteralSet.fromList (m_class_rel_cls subclass superclass)));
-
-(* ------------------------------------------------------------------------- *)
-(* FOL to HOL (Metis to Isabelle) *)
-(* ------------------------------------------------------------------------- *)
-
-datatype term_or_type = Term of Term.term | Type of Term.typ;
-
-fun terms_of [] = []
- | terms_of (Term t :: tts) = t :: terms_of tts
- | terms_of (Type _ :: tts) = terms_of tts;
-
-fun types_of [] = []
- | types_of (Term (Term.Var ((a,idx), _)) :: tts) =
- if String.isPrefix "_" a then
- (*Variable generated by Metis, which might have been a type variable.*)
- TVar (("'" ^ a, idx), HOLogic.typeS) :: types_of tts
- else types_of tts
- | types_of (Term _ :: tts) = types_of tts
- | types_of (Type T :: tts) = T :: types_of tts;
-
-fun apply_list rator nargs rands =
- let val trands = terms_of rands
- in if length trands = nargs then Term (list_comb(rator, trands))
- else raise Fail
- ("apply_list: wrong number of arguments: " ^ Syntax.string_of_term_global Pure.thy rator ^
- " expected " ^ Int.toString nargs ^
- " received " ^ commas (map (Syntax.string_of_term_global Pure.thy) trands))
- end;
-
-fun infer_types ctxt =
- Syntax.check_terms (ProofContext.set_mode ProofContext.mode_pattern ctxt);
-
-(*We use 1 rather than 0 because variable references in clauses may otherwise conflict
- with variable constraints in the goal...at least, type inference often fails otherwise.
- SEE ALSO axiom_inf below.*)
-fun mk_var (w,T) = Term.Var((w,1), T);
-
-(*include the default sort, if available*)
-fun mk_tfree ctxt w =
- let val ww = "'" ^ w
- in TFree(ww, the_default HOLogic.typeS (Variable.def_sort ctxt (ww, ~1))) end;
-
-(*Remove the "apply" operator from an HO term*)
-fun strip_happ args (Metis_Term.Fn(".",[t,u])) = strip_happ (u::args) t
- | strip_happ args x = (x, args);
-
-fun make_tvar s = TVar (("'" ^ s, 0), HOLogic.typeS)
-
-fun smart_invert_const "fequal" = @{const_name HOL.eq}
- | smart_invert_const s = invert_const s
-
-fun hol_type_from_metis_term _ (Metis_Term.Var v) =
- (case strip_prefix_and_unascii tvar_prefix v of
- SOME w => make_tvar w
- | NONE => make_tvar v)
- | hol_type_from_metis_term ctxt (Metis_Term.Fn(x, tys)) =
- (case strip_prefix_and_unascii type_const_prefix x of
- SOME tc => Term.Type (smart_invert_const tc,
- map (hol_type_from_metis_term ctxt) tys)
- | NONE =>
- case strip_prefix_and_unascii tfree_prefix x of
- SOME tf => mk_tfree ctxt tf
- | NONE => raise Fail ("hol_type_from_metis_term: " ^ x));
-
-(*Maps metis terms to isabelle terms*)
-fun hol_term_from_metis_PT ctxt fol_tm =
- let val thy = ProofContext.theory_of ctxt
- val _ = trace_msg (fn () => "hol_term_from_metis_PT: " ^
- Metis_Term.toString fol_tm)
- fun tm_to_tt (Metis_Term.Var v) =
- (case strip_prefix_and_unascii tvar_prefix v of
- SOME w => Type (make_tvar w)
- | NONE =>
- case strip_prefix_and_unascii schematic_var_prefix v of
- SOME w => Term (mk_var (w, HOLogic.typeT))
- | NONE => Term (mk_var (v, HOLogic.typeT)) )
- (*Var from Metis with a name like _nnn; possibly a type variable*)
- | tm_to_tt (Metis_Term.Fn ("{}", [arg])) = tm_to_tt arg (*hBOOL*)
- | tm_to_tt (t as Metis_Term.Fn (".",_)) =
- let val (rator,rands) = strip_happ [] t
- in case rator of
- Metis_Term.Fn(fname,ts) => applic_to_tt (fname, ts @ rands)
- | _ => case tm_to_tt rator of
- Term t => Term (list_comb(t, terms_of (map tm_to_tt rands)))
- | _ => raise Fail "tm_to_tt: HO application"
- end
- | tm_to_tt (Metis_Term.Fn (fname, args)) = applic_to_tt (fname,args)
- and applic_to_tt ("=",ts) =
- Term (list_comb(Const (@{const_name HOL.eq}, HOLogic.typeT), terms_of (map tm_to_tt ts)))
- | applic_to_tt (a,ts) =
- case strip_prefix_and_unascii const_prefix a of
- SOME b =>
- let val c = smart_invert_const b
- val ntypes = num_type_args thy c
- val nterms = length ts - ntypes
- val tts = map tm_to_tt ts
- val tys = types_of (List.take(tts,ntypes))
- in if length tys = ntypes then
- apply_list (Const (c, dummyT)) nterms (List.drop(tts,ntypes))
- else
- raise Fail ("Constant " ^ c ^ " expects " ^ Int.toString ntypes ^
- " but gets " ^ Int.toString (length tys) ^
- " type arguments\n" ^
- cat_lines (map (Syntax.string_of_typ ctxt) tys) ^
- " the terms are \n" ^
- cat_lines (map (Syntax.string_of_term ctxt) (terms_of tts)))
- end
- | NONE => (*Not a constant. Is it a type constructor?*)
- case strip_prefix_and_unascii type_const_prefix a of
- SOME b =>
- Type (Term.Type (smart_invert_const b, types_of (map tm_to_tt ts)))
- | NONE => (*Maybe a TFree. Should then check that ts=[].*)
- case strip_prefix_and_unascii tfree_prefix a of
- SOME b => Type (mk_tfree ctxt b)
- | NONE => (*a fixed variable? They are Skolem functions.*)
- case strip_prefix_and_unascii fixed_var_prefix a of
- SOME b =>
- let val opr = Term.Free(b, HOLogic.typeT)
- in apply_list opr (length ts) (map tm_to_tt ts) end
- | NONE => raise Fail ("unexpected metis function: " ^ a)
- in
- case tm_to_tt fol_tm of
- Term t => t
- | _ => raise Fail "fol_tm_to_tt: Term expected"
- end
+fun is_false t = t aconv (HOLogic.mk_Trueprop HOLogic.false_const);
-(*Maps fully-typed metis terms to isabelle terms*)
-fun hol_term_from_metis_FT ctxt fol_tm =
- let val _ = trace_msg (fn () => "hol_term_from_metis_FT: " ^
- Metis_Term.toString fol_tm)
- fun cvt (Metis_Term.Fn ("ti", [Metis_Term.Var v, _])) =
- (case strip_prefix_and_unascii schematic_var_prefix v of
- SOME w => mk_var(w, dummyT)
- | NONE => mk_var(v, dummyT))
- | cvt (Metis_Term.Fn ("ti", [Metis_Term.Fn ("=",[]), _])) =
- Const (@{const_name HOL.eq}, HOLogic.typeT)
- | cvt (Metis_Term.Fn ("ti", [Metis_Term.Fn (x,[]), ty])) =
- (case strip_prefix_and_unascii const_prefix x of
- SOME c => Const (smart_invert_const c, dummyT)
- | NONE => (*Not a constant. Is it a fixed variable??*)
- case strip_prefix_and_unascii fixed_var_prefix x of
- SOME v => Free (v, hol_type_from_metis_term ctxt ty)
- | NONE => raise Fail ("hol_term_from_metis_FT bad constant: " ^ x))
- | cvt (Metis_Term.Fn ("ti", [Metis_Term.Fn (".",[tm1,tm2]), _])) =
- cvt tm1 $ cvt tm2
- | cvt (Metis_Term.Fn (".",[tm1,tm2])) = (*untyped application*)
- cvt tm1 $ cvt tm2
- | cvt (Metis_Term.Fn ("{}", [arg])) = cvt arg (*hBOOL*)
- | cvt (Metis_Term.Fn ("=", [tm1,tm2])) =
- list_comb(Const (@{const_name HOL.eq}, HOLogic.typeT), map cvt [tm1,tm2])
- | cvt (t as Metis_Term.Fn (x, [])) =
- (case strip_prefix_and_unascii const_prefix x of
- SOME c => Const (smart_invert_const c, dummyT)
- | NONE => (*Not a constant. Is it a fixed variable??*)
- case strip_prefix_and_unascii fixed_var_prefix x of
- SOME v => Free (v, dummyT)
- | NONE => (trace_msg (fn () => "hol_term_from_metis_FT bad const: " ^ x);
- hol_term_from_metis_PT ctxt t))
- | cvt t = (trace_msg (fn () => "hol_term_from_metis_FT bad term: " ^ Metis_Term.toString t);
- hol_term_from_metis_PT ctxt t)
- in fol_tm |> cvt end
-
-fun hol_term_from_metis FT = hol_term_from_metis_FT
- | hol_term_from_metis _ = hol_term_from_metis_PT
-
-fun hol_terms_from_fol ctxt mode skolems fol_tms =
- let val ts = map (hol_term_from_metis mode ctxt) fol_tms
- val _ = trace_msg (fn () => " calling type inference:")
- val _ = app (fn t => trace_msg (fn () => Syntax.string_of_term ctxt t)) ts
- val ts' = ts |> map (reveal_skolem_terms skolems) |> infer_types ctxt
- val _ = app (fn t => trace_msg
- (fn () => " final term: " ^ Syntax.string_of_term ctxt t ^
- " of type " ^ Syntax.string_of_typ ctxt (type_of t)))
- ts'
- in ts' end;
-
-fun mk_not (Const (@{const_name Not}, _) $ b) = b
- | mk_not b = HOLogic.mk_not b;
-
-val metis_eq = Metis_Term.Fn ("=", []);
-
-(* ------------------------------------------------------------------------- *)
-(* FOL step Inference Rules *)
-(* ------------------------------------------------------------------------- *)
-
-(*for debugging only*)
-(*
-fun print_thpair (fth,th) =
- (trace_msg (fn () => "=============================================");
- trace_msg (fn () => "Metis: " ^ Metis_Thm.toString fth);
- trace_msg (fn () => "Isabelle: " ^ Display.string_of_thm_without_context th));
-*)
-
-fun lookth thpairs (fth : Metis_Thm.thm) =
- the (AList.lookup (uncurry Metis_Thm.equal) thpairs fth)
- handle Option =>
- raise Fail ("Failed to find a Metis theorem " ^ Metis_Thm.toString fth);
-
-fun is_TrueI th = Thm.eq_thm(TrueI,th);
-
-fun cterm_incr_types thy idx = cterm_of thy o (map_types (Logic.incr_tvar idx));
-
-fun inst_excluded_middle thy i_atm =
- let val th = EXCLUDED_MIDDLE
- val [vx] = Term.add_vars (prop_of th) []
- val substs = [(cterm_of thy (Var vx), cterm_of thy i_atm)]
- in cterm_instantiate substs th end;
-
-(* INFERENCE RULE: AXIOM *)
-fun axiom_inf thpairs th = Thm.incr_indexes 1 (lookth thpairs th);
- (*This causes variables to have an index of 1 by default. SEE ALSO mk_var above.*)
-
-(* INFERENCE RULE: ASSUME *)
-fun assume_inf ctxt mode skolems atm =
- inst_excluded_middle
- (ProofContext.theory_of ctxt)
- (singleton (hol_terms_from_fol ctxt mode skolems) (Metis_Term.Fn atm))
-
-(* INFERENCE RULE: INSTANTIATE (Subst). Type instantiations are ignored. Trying to reconstruct
- them admits new possibilities of errors, e.g. concerning sorts. Instead we try to arrange
- that new TVars are distinct and that types can be inferred from terms.*)
-fun inst_inf ctxt mode skolems thpairs fsubst th =
- let val thy = ProofContext.theory_of ctxt
- val i_th = lookth thpairs th
- val i_th_vars = Term.add_vars (prop_of i_th) []
- fun find_var x = the (List.find (fn ((a,_),_) => a=x) i_th_vars)
- fun subst_translation (x,y) =
- let val v = find_var x
- (* We call "reveal_skolem_terms" and "infer_types" below. *)
- val t = hol_term_from_metis mode ctxt y
- in SOME (cterm_of thy (Var v), t) end
- handle Option =>
- (trace_msg (fn() => "\"find_var\" failed for the variable " ^ x ^
- " in " ^ Display.string_of_thm ctxt i_th);
- NONE)
- fun remove_typeinst (a, t) =
- case strip_prefix_and_unascii schematic_var_prefix a of
- SOME b => SOME (b, t)
- | NONE => case strip_prefix_and_unascii tvar_prefix a of
- SOME _ => NONE (*type instantiations are forbidden!*)
- | NONE => SOME (a,t) (*internal Metis var?*)
- val _ = trace_msg (fn () => " isa th: " ^ Display.string_of_thm ctxt i_th)
- val substs = map_filter remove_typeinst (Metis_Subst.toList fsubst)
- val (vars,rawtms) = ListPair.unzip (map_filter subst_translation substs)
- val tms = rawtms |> map (reveal_skolem_terms skolems) |> infer_types ctxt
- val ctm_of = cterm_incr_types thy (1 + Thm.maxidx_of i_th)
- val substs' = ListPair.zip (vars, map ctm_of tms)
- val _ = trace_msg (fn () =>
- cat_lines ("subst_translations:" ::
- (substs' |> map (fn (x, y) =>
- Syntax.string_of_term ctxt (term_of x) ^ " |-> " ^
- Syntax.string_of_term ctxt (term_of y)))));
- in cterm_instantiate substs' i_th end
- handle THM (msg, _, _) =>
- error ("Cannot replay Metis proof in Isabelle:\n" ^ msg)
-
-(* INFERENCE RULE: RESOLVE *)
-
-(* Like RSN, but we rename apart only the type variables. Vars here typically
- have an index of 1, and the use of RSN would increase this typically to 3.
- Instantiations of those Vars could then fail. See comment on "mk_var". *)
-fun resolve_inc_tyvars thy tha i thb =
- let
- val tha = Drule.incr_type_indexes (1 + Thm.maxidx_of thb) tha
- fun aux tha thb =
- case Thm.bicompose false (false, tha, nprems_of tha) i thb
- |> Seq.list_of |> distinct Thm.eq_thm of
- [th] => th
- | _ => raise THM ("resolve_inc_tyvars: unique result expected", i,
- [tha, thb])
- in
- aux tha thb
- handle TERM z =>
- (* The unifier, which is invoked from "Thm.bicompose", will sometimes
- refuse to unify "?a::?'a" with "?a::?'b" or "?a::nat" and throw a
- "TERM" exception (with "add_ffpair" as first argument). We then
- perform unification of the types of variables by hand and try
- again. We could do this the first time around but this error
- occurs seldom and we don't want to break existing proofs in subtle
- ways or slow them down needlessly. *)
- case [] |> fold (Term.add_vars o prop_of) [tha, thb]
- |> AList.group (op =)
- |> maps (fn ((s, _), T :: Ts) =>
- map (fn T' => (Free (s, T), Free (s, T'))) Ts)
- |> rpair (Envir.empty ~1)
- |-> fold (Pattern.unify thy)
- |> Envir.type_env |> Vartab.dest
- |> map (fn (x, (S, T)) =>
- pairself (ctyp_of thy) (TVar (x, S), T)) of
- [] => raise TERM z
- | ps => aux (instantiate (ps, []) tha) (instantiate (ps, []) thb)
- end
-
-fun resolve_inf ctxt mode skolems thpairs atm th1 th2 =
- let
- val thy = ProofContext.theory_of ctxt
- val i_th1 = lookth thpairs th1 and i_th2 = lookth thpairs th2
- val _ = trace_msg (fn () => " isa th1 (pos): " ^ Display.string_of_thm ctxt i_th1)
- val _ = trace_msg (fn () => " isa th2 (neg): " ^ Display.string_of_thm ctxt i_th2)
- in
- if is_TrueI i_th1 then i_th2 (*Trivial cases where one operand is type info*)
- else if is_TrueI i_th2 then i_th1
- else
- let
- val i_atm = singleton (hol_terms_from_fol ctxt mode skolems)
- (Metis_Term.Fn atm)
- val _ = trace_msg (fn () => " atom: " ^ Syntax.string_of_term ctxt i_atm)
- val prems_th1 = prems_of i_th1
- val prems_th2 = prems_of i_th2
- val index_th1 = get_index (mk_not i_atm) prems_th1
- handle Empty => raise Fail "Failed to find literal in th1"
- val _ = trace_msg (fn () => " index_th1: " ^ Int.toString index_th1)
- val index_th2 = get_index i_atm prems_th2
- handle Empty => raise Fail "Failed to find literal in th2"
- val _ = trace_msg (fn () => " index_th2: " ^ Int.toString index_th2)
- in
- resolve_inc_tyvars thy (Meson.select_literal index_th1 i_th1) index_th2
- i_th2
- end
- end;
-
-(* INFERENCE RULE: REFL *)
-val refl_x = cterm_of @{theory} (Var (hd (Term.add_vars (prop_of REFL_THM) [])));
-val refl_idx = 1 + Thm.maxidx_of REFL_THM;
-
-fun refl_inf ctxt mode skolems t =
- let val thy = ProofContext.theory_of ctxt
- val i_t = singleton (hol_terms_from_fol ctxt mode skolems) t
- val _ = trace_msg (fn () => " term: " ^ Syntax.string_of_term ctxt i_t)
- val c_t = cterm_incr_types thy refl_idx i_t
- in cterm_instantiate [(refl_x, c_t)] REFL_THM end;
-
-fun get_ty_arg_size _ (Const (@{const_name HOL.eq}, _)) = 0 (*equality has no type arguments*)
- | get_ty_arg_size thy (Const (c, _)) = (num_type_args thy c handle TYPE _ => 0)
- | get_ty_arg_size _ _ = 0;
-
-(* INFERENCE RULE: EQUALITY *)
-fun equality_inf ctxt mode skolems (pos, atm) fp fr =
- let val thy = ProofContext.theory_of ctxt
- val m_tm = Metis_Term.Fn atm
- val [i_atm,i_tm] = hol_terms_from_fol ctxt mode skolems [m_tm, fr]
- val _ = trace_msg (fn () => "sign of the literal: " ^ Bool.toString pos)
- fun replace_item_list lx 0 (_::ls) = lx::ls
- | replace_item_list lx i (l::ls) = l :: replace_item_list lx (i-1) ls
- fun path_finder_FO tm [] = (tm, Term.Bound 0)
- | path_finder_FO tm (p::ps) =
- let val (tm1,args) = strip_comb tm
- val adjustment = get_ty_arg_size thy tm1
- val p' = if adjustment > p then p else p-adjustment
- val tm_p = List.nth(args,p')
- handle Subscript =>
- error ("Cannot replay Metis proof in Isabelle:\n" ^
- "equality_inf: " ^ Int.toString p ^ " adj " ^
- Int.toString adjustment ^ " term " ^
- Syntax.string_of_term ctxt tm)
- val _ = trace_msg (fn () => "path_finder: " ^ Int.toString p ^
- " " ^ Syntax.string_of_term ctxt tm_p)
- val (r,t) = path_finder_FO tm_p ps
- in
- (r, list_comb (tm1, replace_item_list t p' args))
- end
- fun path_finder_HO tm [] = (tm, Term.Bound 0)
- | path_finder_HO (t$u) (0::ps) = (fn(x,y) => (x, y$u)) (path_finder_HO t ps)
- | path_finder_HO (t$u) (_::ps) = (fn(x,y) => (x, t$y)) (path_finder_HO u ps)
- | path_finder_HO tm ps =
- raise Fail ("equality_inf, path_finder_HO: path = " ^
- space_implode " " (map Int.toString ps) ^
- " isa-term: " ^ Syntax.string_of_term ctxt tm)
- fun path_finder_FT tm [] _ = (tm, Term.Bound 0)
- | path_finder_FT tm (0::ps) (Metis_Term.Fn ("ti", [t1, _])) =
- path_finder_FT tm ps t1
- | path_finder_FT (t$u) (0::ps) (Metis_Term.Fn (".", [t1, _])) =
- (fn(x,y) => (x, y$u)) (path_finder_FT t ps t1)
- | path_finder_FT (t$u) (1::ps) (Metis_Term.Fn (".", [_, t2])) =
- (fn(x,y) => (x, t$y)) (path_finder_FT u ps t2)
- | path_finder_FT tm ps t =
- raise Fail ("equality_inf, path_finder_FT: path = " ^
- space_implode " " (map Int.toString ps) ^
- " isa-term: " ^ Syntax.string_of_term ctxt tm ^
- " fol-term: " ^ Metis_Term.toString t)
- fun path_finder FO tm ps _ = path_finder_FO tm ps
- | path_finder HO (tm as Const(@{const_name HOL.eq},_) $ _ $ _) (p::ps) _ =
- (*equality: not curried, as other predicates are*)
- if p=0 then path_finder_HO tm (0::1::ps) (*select first operand*)
- else path_finder_HO tm (p::ps) (*1 selects second operand*)
- | path_finder HO tm (_ :: ps) (Metis_Term.Fn ("{}", [_])) =
- path_finder_HO tm ps (*if not equality, ignore head to skip hBOOL*)
- | path_finder FT (tm as Const(@{const_name HOL.eq}, _) $ _ $ _) (p::ps)
- (Metis_Term.Fn ("=", [t1,t2])) =
- (*equality: not curried, as other predicates are*)
- if p=0 then path_finder_FT tm (0::1::ps)
- (Metis_Term.Fn (".", [Metis_Term.Fn (".", [metis_eq,t1]), t2]))
- (*select first operand*)
- else path_finder_FT tm (p::ps)
- (Metis_Term.Fn (".", [metis_eq,t2]))
- (*1 selects second operand*)
- | path_finder FT tm (_ :: ps) (Metis_Term.Fn ("{}", [t1])) = path_finder_FT tm ps t1
- (*if not equality, ignore head to skip the hBOOL predicate*)
- | path_finder FT tm ps t = path_finder_FT tm ps t (*really an error case!*)
- fun path_finder_lit ((nt as Const (@{const_name Not}, _)) $ tm_a) idx =
- let val (tm, tm_rslt) = path_finder mode tm_a idx m_tm
- in (tm, nt $ tm_rslt) end
- | path_finder_lit tm_a idx = path_finder mode tm_a idx m_tm
- val (tm_subst, body) = path_finder_lit i_atm fp
- val tm_abs = Term.Abs("x", Term.type_of tm_subst, body)
- val _ = trace_msg (fn () => "abstraction: " ^ Syntax.string_of_term ctxt tm_abs)
- val _ = trace_msg (fn () => "i_tm: " ^ Syntax.string_of_term ctxt i_tm)
- val _ = trace_msg (fn () => "located term: " ^ Syntax.string_of_term ctxt tm_subst)
- val imax = maxidx_of_term (i_tm $ tm_abs $ tm_subst) (*ill typed but gives right max*)
- val subst' = Thm.incr_indexes (imax+1) (if pos then subst_em else ssubst_em)
- val _ = trace_msg (fn () => "subst' " ^ Display.string_of_thm ctxt subst')
- val eq_terms = map (pairself (cterm_of thy))
- (ListPair.zip (OldTerm.term_vars (prop_of subst'), [tm_abs, tm_subst, i_tm]))
- in cterm_instantiate eq_terms subst' end;
-
-val factor = Seq.hd o distinct_subgoals_tac;
-
-fun step ctxt mode skolems thpairs p =
- case p of
- (fol_th, Metis_Proof.Axiom _) => factor (axiom_inf thpairs fol_th)
- | (_, Metis_Proof.Assume f_atm) => assume_inf ctxt mode skolems f_atm
- | (_, Metis_Proof.Metis_Subst (f_subst, f_th1)) =>
- factor (inst_inf ctxt mode skolems thpairs f_subst f_th1)
- | (_, Metis_Proof.Resolve(f_atm, f_th1, f_th2)) =>
- factor (resolve_inf ctxt mode skolems thpairs f_atm f_th1 f_th2)
- | (_, Metis_Proof.Refl f_tm) => refl_inf ctxt mode skolems f_tm
- | (_, Metis_Proof.Equality (f_lit, f_p, f_r)) =>
- equality_inf ctxt mode skolems f_lit f_p f_r
-
-fun real_literal (_, (c, _)) = not (String.isPrefix class_prefix c);
-
-fun translate_one ctxt mode skolems (fol_th, inf) thpairs =
- let
- val _ = trace_msg (fn () => "=============================================")
- val _ = trace_msg (fn () => "METIS THM: " ^ Metis_Thm.toString fol_th)
- val _ = trace_msg (fn () => "INFERENCE: " ^ Metis_Proof.inferenceToString inf)
- val th = Meson.flexflex_first_order (step ctxt mode skolems
- thpairs (fol_th, inf))
- val _ = trace_msg (fn () => "ISABELLE THM: " ^ Display.string_of_thm ctxt th)
- val _ = trace_msg (fn () => "=============================================")
- val n_metis_lits =
- length (filter real_literal (Metis_LiteralSet.toList (Metis_Thm.clause fol_th)))
- val _ = if nprems_of th = n_metis_lits then ()
- else error "Cannot replay Metis proof in Isabelle."
- in (fol_th, th) :: thpairs end
+fun have_common_thm ths1 ths2 =
+ exists (member Thm.eq_thm ths1) (map Meson.make_meta_clause ths2)
(*Determining which axiom clauses are actually used*)
fun used_axioms axioms (th, Metis_Proof.Axiom _) = SOME (lookth axioms th)
| used_axioms _ _ = NONE;
-(* ------------------------------------------------------------------------- *)
-(* Translation of HO Clauses *)
-(* ------------------------------------------------------------------------- *)
-
-fun type_ext thy tms =
- let val subs = tfree_classes_of_terms tms
- val supers = tvar_classes_of_terms tms
- and tycons = type_consts_of_terms thy tms
- val (supers', arity_clauses) = make_arity_clauses thy tycons supers
- val class_rel_clauses = make_class_rel_clauses thy subs supers'
- in map class_rel_cls class_rel_clauses @ map arity_cls arity_clauses
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Logic maps manage the interface between HOL and first-order logic. *)
-(* ------------------------------------------------------------------------- *)
-
-type logic_map =
- {axioms: (Metis_Thm.thm * thm) list,
- tfrees: type_literal list,
- skolems: (string * term) list}
-
-fun const_in_metis c (pred, tm_list) =
- let
- fun in_mterm (Metis_Term.Var _) = false
- | in_mterm (Metis_Term.Fn (".", tm_list)) = exists in_mterm tm_list
- | in_mterm (Metis_Term.Fn (nm, tm_list)) = c=nm orelse exists in_mterm tm_list
- in c = pred orelse exists in_mterm tm_list end;
-
-(*Extract TFree constraints from context to include as conjecture clauses*)
-fun init_tfrees ctxt =
- let fun add ((a,i),s) Ts = if i = ~1 then TFree(a,s) :: Ts else Ts in
- Vartab.fold add (#2 (Variable.constraints_of ctxt)) []
- |> type_literals_for_types
- end;
-
-(*transform isabelle type / arity clause to metis clause *)
-fun add_type_thm [] lmap = lmap
- | add_type_thm ((ith, mth) :: cls) {axioms, tfrees, skolems} =
- add_type_thm cls {axioms = (mth, ith) :: axioms, tfrees = tfrees,
- skolems = skolems}
-
-(*Insert non-logical axioms corresponding to all accumulated TFrees*)
-fun add_tfrees {axioms, tfrees, skolems} : logic_map =
- {axioms = map (rpair TrueI o metis_of_tfree) (distinct (op =) tfrees) @
- axioms,
- tfrees = tfrees, skolems = skolems}
-
-fun string_of_mode FO = "FO"
- | string_of_mode HO = "HO"
- | string_of_mode FT = "FT"
-
-val helpers =
- [("c_COMBI", (false, map (`I) @{thms COMBI_def})),
- ("c_COMBK", (false, map (`I) @{thms COMBK_def})),
- ("c_COMBB", (false, map (`I) @{thms COMBB_def})),
- ("c_COMBC", (false, map (`I) @{thms COMBC_def})),
- ("c_COMBS", (false, map (`I) @{thms COMBS_def})),
- ("c_fequal", (false, map (rpair @{thm equal_imp_equal})
- @{thms fequal_imp_equal equal_imp_fequal})),
- ("c_True", (true, map (`I) @{thms True_or_False})),
- ("c_False", (true, map (`I) @{thms True_or_False})),
- ("c_If", (true, map (`I) @{thms if_True if_False True_or_False}))]
-
-fun is_quasi_fol_clause thy =
- Meson.is_fol_term thy o snd o conceal_skolem_terms ~1 [] o prop_of
-
-(* Function to generate metis clauses, including comb and type clauses *)
-fun build_map mode0 ctxt cls ths =
- let val thy = ProofContext.theory_of ctxt
- (*The modes FO and FT are sticky. HO can be downgraded to FO.*)
- fun set_mode FO = FO
- | set_mode HO =
- if forall (is_quasi_fol_clause thy) (cls @ ths) then FO else HO
- | set_mode FT = FT
- val mode = set_mode mode0
- (*transform isabelle clause to metis clause *)
- fun add_thm is_conjecture (metis_ith, isa_ith) {axioms, tfrees, skolems}
- : logic_map =
- let
- val (mth, tfree_lits, skolems) =
- hol_thm_to_fol is_conjecture ctxt mode (length axioms) skolems
- metis_ith
- in
- {axioms = (mth, Meson.make_meta_clause isa_ith) :: axioms,
- tfrees = union (op =) tfree_lits tfrees, skolems = skolems}
- end;
- val lmap = {axioms = [], tfrees = init_tfrees ctxt, skolems = []}
- |> fold (add_thm true o `I) cls
- |> add_tfrees
- |> fold (add_thm false o `I) ths
- val clause_lists = map (Metis_Thm.clause o #1) (#axioms lmap)
- fun is_used c =
- exists (Metis_LiteralSet.exists (const_in_metis c o #2)) clause_lists
- val lmap =
- if mode = FO then
- lmap
- else
- let
- val helper_ths =
- helpers |> filter (is_used o fst)
- |> maps (fn (c, (needs_full_types, thms)) =>
- if not (is_used c) orelse
- needs_full_types andalso mode <> FT then
- []
- else
- thms)
- in lmap |> fold (add_thm false) helper_ths end
- in (mode, add_type_thm (type_ext thy (map prop_of (cls @ ths))) lmap) end
-
-fun refute cls =
- Metis_Resolution.loop (Metis_Resolution.new Metis_Resolution.default {axioms = cls, conjecture = []});
-
-fun is_false t = t aconv (HOLogic.mk_Trueprop HOLogic.false_const);
-
-fun common_thm ths1 ths2 = exists (member Thm.eq_thm ths1) (map Meson.make_meta_clause ths2);
-
+val clause_params =
+ {ordering = Metis_KnuthBendixOrder.default,
+ orderLiterals = Metis_Clause.UnsignedLiteralOrder,
+ orderTerms = true}
+val active_params =
+ {clause = clause_params,
+ prefactor = #prefactor Metis_Active.default,
+ postfactor = #postfactor Metis_Active.default}
+val waiting_params =
+ {symbolsWeight = 1.0,
+ variablesWeight = 0.0,
+ literalsWeight = 0.0,
+ models = []}
+val resolution_params = {active = active_params, waiting = waiting_params}
(* Main function to start Metis proof and reconstruction *)
fun FOL_SOLVE mode ctxt cls ths0 =
let val thy = ProofContext.theory_of ctxt
+ val type_lits = Config.get ctxt type_lits
val th_cls_pairs =
map (fn th => (Thm.get_name_hint th, Clausifier.cnf_axiom thy th)) ths0
val ths = maps #2 th_cls_pairs
@@ -754,7 +63,8 @@
val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) cls
val _ = trace_msg (fn () => "THEOREM CLAUSES")
val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) ths
- val (mode, {axioms, tfrees, skolems}) = build_map mode ctxt cls ths
+ val (mode, {axioms, tfrees, skolems}) =
+ build_logic_map mode ctxt type_lits cls ths
val _ = if null tfrees then ()
else (trace_msg (fn () => "TFREE CLAUSES");
app (fn TyLitFree ((s, _), (s', _)) =>
@@ -768,21 +78,22 @@
case filter (is_false o prop_of) cls of
false_th::_ => [false_th RS @{thm FalseE}]
| [] =>
- case refute thms of
+ case Metis_Resolution.new resolution_params {axioms = thms, conjecture = []}
+ |> Metis_Resolution.loop of
Metis_Resolution.Contradiction mth =>
let val _ = trace_msg (fn () => "METIS RECONSTRUCTION START: " ^
Metis_Thm.toString mth)
val ctxt' = fold Variable.declare_constraints (map prop_of cls) ctxt
(*add constraints arising from converting goal to clause form*)
val proof = Metis_Proof.proof mth
- val result = fold (translate_one ctxt' mode skolems) proof axioms
+ val result = fold (replay_one_inference ctxt' mode skolems) proof axioms
and used = map_filter (used_axioms axioms) proof
val _ = trace_msg (fn () => "METIS COMPLETED...clauses actually used:")
val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) used
val unused = th_cls_pairs |> map_filter (fn (name, cls) =>
- if common_thm used cls then NONE else SOME name)
+ if have_common_thm used cls then NONE else SOME name)
in
- if not (null cls) andalso not (common_thm used cls) then
+ if not (null cls) andalso not (have_common_thm used cls) then
warning "Metis: The assumptions are inconsistent."
else
();
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Sledgehammer/metis_translate.ML Fri Sep 17 17:31:20 2010 +0200
@@ -0,0 +1,738 @@
+(* Title: HOL/Tools/Sledgehammer/metis_translate.ML
+ Author: Jia Meng, Cambridge University Computer Laboratory and NICTA
+ Author: Kong W. Susanto, Cambridge University Computer Laboratory
+ Author: Lawrence C. Paulson, Cambridge University Computer Laboratory
+ Author: Jasmin Blanchette, TU Muenchen
+
+Translation of HOL to FOL for Metis.
+*)
+
+signature METIS_TRANSLATE =
+sig
+ type name = string * string
+ datatype type_literal =
+ TyLitVar of name * name |
+ TyLitFree of name * name
+ datatype arLit =
+ TConsLit of name * name * name list |
+ TVarLit of name * name
+ datatype arity_clause =
+ ArityClause of {name: string, conclLit: arLit, premLits: arLit list}
+ datatype class_rel_clause =
+ ClassRelClause of {name: string, subclass: name, superclass: name}
+ datatype combtyp =
+ CombTVar of name |
+ CombTFree of name |
+ CombType of name * combtyp list
+ datatype combterm =
+ CombConst of name * combtyp * combtyp list (* Const and Free *) |
+ CombVar of name * combtyp |
+ CombApp of combterm * combterm
+ datatype fol_literal = FOLLiteral of bool * combterm
+
+ datatype mode = FO | HO | FT
+ type logic_map =
+ {axioms: (Metis_Thm.thm * thm) list,
+ tfrees: type_literal list,
+ skolems: (string * term) list}
+
+ val type_wrapper_name : string
+ val bound_var_prefix : string
+ val schematic_var_prefix: string
+ val fixed_var_prefix: string
+ val tvar_prefix: string
+ val tfree_prefix: string
+ val const_prefix: string
+ val type_const_prefix: string
+ val class_prefix: string
+ val invert_const: string -> string
+ val ascii_of: string -> string
+ val unascii_of: string -> string
+ val strip_prefix_and_unascii: string -> string -> string option
+ val make_bound_var : string -> string
+ val make_schematic_var : string * int -> string
+ val make_fixed_var : string -> string
+ val make_schematic_type_var : string * int -> string
+ val make_fixed_type_var : string -> string
+ val make_fixed_const : string -> string
+ val make_fixed_type_const : string -> string
+ val make_type_class : string -> string
+ val num_type_args: theory -> string -> int
+ val type_literals_for_types : typ list -> type_literal list
+ val make_class_rel_clauses :
+ theory -> class list -> class list -> class_rel_clause list
+ val make_arity_clauses :
+ theory -> string list -> class list -> class list * arity_clause list
+ val combtyp_of : combterm -> combtyp
+ val strip_combterm_comb : combterm -> combterm * combterm list
+ val combterm_from_term :
+ theory -> (string * typ) list -> term -> combterm * typ list
+ val reveal_skolem_terms : (string * term) list -> term -> term
+ val tfree_classes_of_terms : term list -> string list
+ val tvar_classes_of_terms : term list -> string list
+ val type_consts_of_terms : theory -> term list -> string list
+ val string_of_mode : mode -> string
+ val build_logic_map :
+ mode -> Proof.context -> bool -> thm list -> thm list -> mode * logic_map
+end
+
+structure Metis_Translate : METIS_TRANSLATE =
+struct
+
+val type_wrapper_name = "ti"
+
+val bound_var_prefix = "B_"
+val schematic_var_prefix = "V_"
+val fixed_var_prefix = "v_"
+
+val tvar_prefix = "T_";
+val tfree_prefix = "t_";
+
+val const_prefix = "c_";
+val type_const_prefix = "tc_";
+val class_prefix = "class_";
+
+fun union_all xss = fold (union (op =)) xss []
+
+(* Readable names for the more common symbolic functions. Do not mess with the
+ last nine entries of the table unless you know what you are doing. *)
+val const_trans_table =
+ Symtab.make [(@{type_name Product_Type.prod}, "prod"),
+ (@{type_name Sum_Type.sum}, "sum"),
+ (@{const_name HOL.eq}, "equal"),
+ (@{const_name HOL.conj}, "and"),
+ (@{const_name HOL.disj}, "or"),
+ (@{const_name HOL.implies}, "implies"),
+ (@{const_name Set.member}, "member"),
+ (@{const_name fequal}, "fequal"),
+ (@{const_name COMBI}, "COMBI"),
+ (@{const_name COMBK}, "COMBK"),
+ (@{const_name COMBB}, "COMBB"),
+ (@{const_name COMBC}, "COMBC"),
+ (@{const_name COMBS}, "COMBS"),
+ (@{const_name True}, "True"),
+ (@{const_name False}, "False"),
+ (@{const_name If}, "If")]
+
+(* Invert the table of translations between Isabelle and ATPs. *)
+val const_trans_table_inv =
+ Symtab.update ("fequal", @{const_name HOL.eq})
+ (Symtab.make (map swap (Symtab.dest const_trans_table)))
+
+val invert_const = perhaps (Symtab.lookup const_trans_table_inv)
+
+(*Escaping of special characters.
+ Alphanumeric characters are left unchanged.
+ The character _ goes to __
+ Characters in the range ASCII space to / go to _A to _P, respectively.
+ Other characters go to _nnn where nnn is the decimal ASCII code.*)
+val A_minus_space = Char.ord #"A" - Char.ord #" ";
+
+fun stringN_of_int 0 _ = ""
+ | stringN_of_int k n = stringN_of_int (k-1) (n div 10) ^ Int.toString (n mod 10);
+
+fun ascii_of_c c =
+ if Char.isAlphaNum c then String.str c
+ else if c = #"_" then "__"
+ else if #" " <= c andalso c <= #"/"
+ then "_" ^ String.str (Char.chr (Char.ord c + A_minus_space))
+ else ("_" ^ stringN_of_int 3 (Char.ord c)) (*fixed width, in case more digits follow*)
+
+val ascii_of = String.translate ascii_of_c;
+
+(** Remove ASCII armouring from names in proof files **)
+
+(*We don't raise error exceptions because this code can run inside the watcher.
+ Also, the errors are "impossible" (hah!)*)
+fun unascii_aux rcs [] = String.implode(rev rcs)
+ | unascii_aux rcs [#"_"] = unascii_aux (#"_"::rcs) [] (*ERROR*)
+ (*Three types of _ escapes: __, _A to _P, _nnn*)
+ | unascii_aux rcs (#"_" :: #"_" :: cs) = unascii_aux (#"_"::rcs) cs
+ | unascii_aux rcs (#"_" :: c :: cs) =
+ if #"A" <= c andalso c<= #"P" (*translation of #" " to #"/"*)
+ then unascii_aux (Char.chr(Char.ord c - A_minus_space) :: rcs) cs
+ else
+ let val digits = List.take (c::cs, 3) handle Subscript => []
+ in
+ case Int.fromString (String.implode digits) of
+ NONE => unascii_aux (c:: #"_"::rcs) cs (*ERROR*)
+ | SOME n => unascii_aux (Char.chr n :: rcs) (List.drop (cs, 2))
+ end
+ | unascii_aux rcs (c::cs) = unascii_aux (c::rcs) cs
+val unascii_of = unascii_aux [] o String.explode
+
+(* If string s has the prefix s1, return the result of deleting it,
+ un-ASCII'd. *)
+fun strip_prefix_and_unascii s1 s =
+ if String.isPrefix s1 s then
+ SOME (unascii_of (String.extract (s, size s1, NONE)))
+ else
+ NONE
+
+(*Remove the initial ' character from a type variable, if it is present*)
+fun trim_type_var s =
+ if s <> "" andalso String.sub(s,0) = #"'" then String.extract(s,1,NONE)
+ else error ("trim_type: Malformed type variable encountered: " ^ s);
+
+fun ascii_of_indexname (v,0) = ascii_of v
+ | ascii_of_indexname (v,i) = ascii_of v ^ "_" ^ Int.toString i;
+
+fun make_bound_var x = bound_var_prefix ^ ascii_of x
+fun make_schematic_var v = schematic_var_prefix ^ ascii_of_indexname v
+fun make_fixed_var x = fixed_var_prefix ^ ascii_of x
+
+fun make_schematic_type_var (x,i) =
+ tvar_prefix ^ (ascii_of_indexname (trim_type_var x,i));
+fun make_fixed_type_var x = tfree_prefix ^ (ascii_of (trim_type_var x));
+
+fun lookup_const c =
+ case Symtab.lookup const_trans_table c of
+ SOME c' => c'
+ | NONE => ascii_of c
+
+(* HOL.eq MUST BE "equal" because it's built into ATPs. *)
+fun make_fixed_const @{const_name HOL.eq} = "equal"
+ | make_fixed_const c = const_prefix ^ lookup_const c
+
+fun make_fixed_type_const c = type_const_prefix ^ lookup_const c
+
+fun make_type_class clas = class_prefix ^ ascii_of clas;
+
+val skolem_prefix = "Sledgehammer" ^ Long_Name.separator ^ "Sko"
+
+(* The number of type arguments of a constant, zero if it's monomorphic. For
+ (instances of) Skolem pseudoconstants, this information is encoded in the
+ constant name. *)
+fun num_type_args thy s =
+ if String.isPrefix skolem_prefix s then
+ s |> space_explode Long_Name.separator |> List.last |> Int.fromString |> the
+ else
+ (s, Sign.the_const_type thy s) |> Sign.const_typargs thy |> length
+
+(**** Definitions and functions for FOL clauses for TPTP format output ****)
+
+type name = string * string
+
+(**** Isabelle FOL clauses ****)
+
+(* The first component is the type class; the second is a TVar or TFree. *)
+datatype type_literal =
+ TyLitVar of name * name |
+ TyLitFree of name * name
+
+(*Make literals for sorted type variables*)
+fun sorts_on_typs_aux (_, []) = []
+ | sorts_on_typs_aux ((x,i), s::ss) =
+ let val sorts = sorts_on_typs_aux ((x,i), ss)
+ in
+ if s = "HOL.type" then sorts
+ else if i = ~1 then TyLitFree (`make_type_class s, `make_fixed_type_var x) :: sorts
+ else TyLitVar (`make_type_class s, (make_schematic_type_var (x,i), x)) :: sorts
+ end;
+
+fun sorts_on_typs (TFree (a,s)) = sorts_on_typs_aux ((a,~1),s)
+ | sorts_on_typs (TVar (v,s)) = sorts_on_typs_aux (v,s);
+
+(*Given a list of sorted type variables, return a list of type literals.*)
+fun type_literals_for_types Ts =
+ fold (union (op =)) (map sorts_on_typs Ts) []
+
+(** make axiom and conjecture clauses. **)
+
+(**** Isabelle arities ****)
+
+datatype arLit =
+ TConsLit of name * name * name list |
+ TVarLit of name * name
+
+datatype arity_clause =
+ ArityClause of {name: string, conclLit: arLit, premLits: arLit list}
+
+
+fun gen_TVars 0 = []
+ | gen_TVars n = ("T_" ^ Int.toString n) :: gen_TVars (n-1);
+
+fun pack_sort(_,[]) = []
+ | pack_sort(tvar, "HOL.type"::srt) = pack_sort (tvar, srt) (*IGNORE sort "type"*)
+ | pack_sort(tvar, cls::srt) =
+ (`make_type_class cls, (tvar, tvar)) :: pack_sort (tvar, srt)
+
+(*Arity of type constructor tcon :: (arg1,...,argN)res*)
+fun make_axiom_arity_clause (tcons, name, (cls,args)) =
+ let
+ val tvars = gen_TVars (length args)
+ val tvars_srts = ListPair.zip (tvars, args)
+ in
+ ArityClause {name = name,
+ conclLit = TConsLit (`make_type_class cls,
+ `make_fixed_type_const tcons,
+ tvars ~~ tvars),
+ premLits = map TVarLit (union_all (map pack_sort tvars_srts))}
+ end
+
+
+(**** Isabelle class relations ****)
+
+datatype class_rel_clause =
+ ClassRelClause of {name: string, subclass: name, superclass: name}
+
+(*Generate all pairs (sub,super) such that sub is a proper subclass of super in theory thy.*)
+fun class_pairs _ [] _ = []
+ | class_pairs thy subs supers =
+ let
+ val class_less = Sorts.class_less (Sign.classes_of thy)
+ fun add_super sub super = class_less (sub, super) ? cons (sub, super)
+ fun add_supers sub = fold (add_super sub) supers
+ in fold add_supers subs [] end
+
+fun make_class_rel_clause (sub,super) =
+ ClassRelClause {name = sub ^ "_" ^ super,
+ subclass = `make_type_class sub,
+ superclass = `make_type_class super}
+
+fun make_class_rel_clauses thy subs supers =
+ map make_class_rel_clause (class_pairs thy subs supers);
+
+
+(** Isabelle arities **)
+
+fun arity_clause _ _ (_, []) = []
+ | arity_clause seen n (tcons, ("HOL.type",_)::ars) = (*ignore*)
+ arity_clause seen n (tcons,ars)
+ | arity_clause seen n (tcons, (ar as (class,_)) :: ars) =
+ if member (op =) seen class then (*multiple arities for the same tycon, class pair*)
+ make_axiom_arity_clause (tcons, lookup_const tcons ^ "_" ^ class ^ "_" ^ Int.toString n, ar) ::
+ arity_clause seen (n+1) (tcons,ars)
+ else
+ make_axiom_arity_clause (tcons, lookup_const tcons ^ "_" ^ class, ar) ::
+ arity_clause (class::seen) n (tcons,ars)
+
+fun multi_arity_clause [] = []
+ | multi_arity_clause ((tcons, ars) :: tc_arlists) =
+ arity_clause [] 1 (tcons, ars) @ multi_arity_clause tc_arlists
+
+(*Generate all pairs (tycon,class,sorts) such that tycon belongs to class in theory thy
+ provided its arguments have the corresponding sorts.*)
+fun type_class_pairs thy tycons classes =
+ let val alg = Sign.classes_of thy
+ fun domain_sorts tycon = Sorts.mg_domain alg tycon o single
+ fun add_class tycon class =
+ cons (class, domain_sorts tycon class)
+ handle Sorts.CLASS_ERROR _ => I
+ fun try_classes tycon = (tycon, fold (add_class tycon) classes [])
+ in map try_classes tycons end;
+
+(*Proving one (tycon, class) membership may require proving others, so iterate.*)
+fun iter_type_class_pairs _ _ [] = ([], [])
+ | iter_type_class_pairs thy tycons classes =
+ let val cpairs = type_class_pairs thy tycons classes
+ val newclasses = union_all (union_all (union_all (map (map #2 o #2) cpairs)))
+ |> subtract (op =) classes |> subtract (op =) HOLogic.typeS
+ val (classes', cpairs') = iter_type_class_pairs thy tycons newclasses
+ in (union (op =) classes' classes, union (op =) cpairs' cpairs) end;
+
+fun make_arity_clauses thy tycons classes =
+ let val (classes', cpairs) = iter_type_class_pairs thy tycons classes
+ in (classes', multi_arity_clause cpairs) end;
+
+datatype combtyp =
+ CombTVar of name |
+ CombTFree of name |
+ CombType of name * combtyp list
+
+datatype combterm =
+ CombConst of name * combtyp * combtyp list (* Const and Free *) |
+ CombVar of name * combtyp |
+ CombApp of combterm * combterm
+
+datatype fol_literal = FOLLiteral of bool * combterm
+
+(*********************************************************************)
+(* convert a clause with type Term.term to a clause with type clause *)
+(*********************************************************************)
+
+(*Result of a function type; no need to check that the argument type matches.*)
+fun result_type (CombType (_, [_, tp2])) = tp2
+ | result_type _ = raise Fail "non-function type"
+
+fun combtyp_of (CombConst (_, tp, _)) = tp
+ | combtyp_of (CombVar (_, tp)) = tp
+ | combtyp_of (CombApp (t1, _)) = result_type (combtyp_of t1)
+
+(*gets the head of a combinator application, along with the list of arguments*)
+fun strip_combterm_comb u =
+ let fun stripc (CombApp(t,u), ts) = stripc (t, u::ts)
+ | stripc x = x
+ in stripc(u,[]) end
+
+fun combtype_of (Type (a, Ts)) =
+ let val (folTypes, ts) = combtypes_of Ts in
+ (CombType (`make_fixed_type_const a, folTypes), ts)
+ end
+ | combtype_of (tp as TFree (a, _)) = (CombTFree (`make_fixed_type_var a), [tp])
+ | combtype_of (tp as TVar (x, _)) =
+ (CombTVar (make_schematic_type_var x, string_of_indexname x), [tp])
+and combtypes_of Ts =
+ let val (folTyps, ts) = ListPair.unzip (map combtype_of Ts) in
+ (folTyps, union_all ts)
+ end
+
+(* same as above, but no gathering of sort information *)
+fun simple_combtype_of (Type (a, Ts)) =
+ CombType (`make_fixed_type_const a, map simple_combtype_of Ts)
+ | simple_combtype_of (TFree (a, _)) = CombTFree (`make_fixed_type_var a)
+ | simple_combtype_of (TVar (x, _)) =
+ CombTVar (make_schematic_type_var x, string_of_indexname x)
+
+(* Converts a term (with combinators) into a combterm. Also accummulates sort
+ infomation. *)
+fun combterm_from_term thy bs (P $ Q) =
+ let val (P', tsP) = combterm_from_term thy bs P
+ val (Q', tsQ) = combterm_from_term thy bs Q
+ in (CombApp (P', Q'), union (op =) tsP tsQ) end
+ | combterm_from_term thy _ (Const (c, T)) =
+ let
+ val (tp, ts) = combtype_of T
+ val tvar_list =
+ (if String.isPrefix skolem_prefix c then
+ [] |> Term.add_tvarsT T |> map TVar
+ else
+ (c, T) |> Sign.const_typargs thy)
+ |> map simple_combtype_of
+ val c' = CombConst (`make_fixed_const c, tp, tvar_list)
+ in (c',ts) end
+ | combterm_from_term _ _ (Free (v, T)) =
+ let val (tp, ts) = combtype_of T
+ val v' = CombConst (`make_fixed_var v, tp, [])
+ in (v',ts) end
+ | combterm_from_term _ _ (Var (v, T)) =
+ let val (tp,ts) = combtype_of T
+ val v' = CombVar ((make_schematic_var v, string_of_indexname v), tp)
+ in (v',ts) end
+ | combterm_from_term _ bs (Bound j) =
+ let
+ val (s, T) = nth bs j
+ val (tp, ts) = combtype_of T
+ val v' = CombConst (`make_bound_var s, tp, [])
+ in (v', ts) end
+ | combterm_from_term _ _ (Abs _) = raise Fail "HOL clause: Abs"
+
+fun predicate_of thy ((@{const Not} $ P), pos) = predicate_of thy (P, not pos)
+ | predicate_of thy (t, pos) =
+ (combterm_from_term thy [] (Envir.eta_contract t), pos)
+
+fun literals_of_term1 args thy (@{const Trueprop} $ P) =
+ literals_of_term1 args thy P
+ | literals_of_term1 args thy (@{const HOL.disj} $ P $ Q) =
+ literals_of_term1 (literals_of_term1 args thy P) thy Q
+ | literals_of_term1 (lits, ts) thy P =
+ let val ((pred, ts'), pol) = predicate_of thy (P, true) in
+ (FOLLiteral (pol, pred) :: lits, union (op =) ts ts')
+ end
+val literals_of_term = literals_of_term1 ([], [])
+
+fun skolem_name i j num_T_args =
+ skolem_prefix ^ Long_Name.separator ^
+ (space_implode Long_Name.separator (map Int.toString [i, j, num_T_args]))
+
+fun conceal_skolem_terms i skolems t =
+ if exists_Const (curry (op =) @{const_name skolem} o fst) t then
+ let
+ fun aux skolems
+ (t as (Const (@{const_name skolem}, Type (_, [_, T])) $ _)) =
+ let
+ val (skolems, s) =
+ if i = ~1 then
+ (skolems, @{const_name undefined})
+ else case AList.find (op aconv) skolems t of
+ s :: _ => (skolems, s)
+ | [] =>
+ let
+ val s = skolem_name i (length skolems)
+ (length (Term.add_tvarsT T []))
+ in ((s, t) :: skolems, s) end
+ in (skolems, Const (s, T)) end
+ | aux skolems (t1 $ t2) =
+ let
+ val (skolems, t1) = aux skolems t1
+ val (skolems, t2) = aux skolems t2
+ in (skolems, t1 $ t2) end
+ | aux skolems (Abs (s, T, t')) =
+ let val (skolems, t') = aux skolems t' in
+ (skolems, Abs (s, T, t'))
+ end
+ | aux skolems t = (skolems, t)
+ in aux skolems t end
+ else
+ (skolems, t)
+
+fun reveal_skolem_terms skolems =
+ map_aterms (fn t as Const (s, _) =>
+ if String.isPrefix skolem_prefix s then
+ AList.lookup (op =) skolems s |> the
+ |> map_types Type_Infer.paramify_vars
+ else
+ t
+ | t => t)
+
+
+(***************************************************************)
+(* Type Classes Present in the Axiom or Conjecture Clauses *)
+(***************************************************************)
+
+fun set_insert (x, s) = Symtab.update (x, ()) s
+
+fun add_classes (sorts, cset) = List.foldl set_insert cset (flat sorts)
+
+(*Remove this trivial type class*)
+fun delete_type cset = Symtab.delete_safe (the_single @{sort HOL.type}) cset;
+
+fun tfree_classes_of_terms ts =
+ let val sorts_list = map (map #2 o OldTerm.term_tfrees) ts
+ in Symtab.keys (delete_type (List.foldl add_classes Symtab.empty sorts_list)) end;
+
+fun tvar_classes_of_terms ts =
+ let val sorts_list = map (map #2 o OldTerm.term_tvars) ts
+ in Symtab.keys (delete_type (List.foldl add_classes Symtab.empty sorts_list)) end;
+
+(*fold type constructors*)
+fun fold_type_consts f (Type (a, Ts)) x = fold (fold_type_consts f) Ts (f (a,x))
+ | fold_type_consts _ _ x = x;
+
+(*Type constructors used to instantiate overloaded constants are the only ones needed.*)
+fun add_type_consts_in_term thy =
+ let
+ fun aux (Const x) =
+ fold (fold_type_consts set_insert) (Sign.const_typargs thy x)
+ | aux (Abs (_, _, u)) = aux u
+ | aux (Const (@{const_name skolem}, _) $ _) = I
+ | aux (t $ u) = aux t #> aux u
+ | aux _ = I
+ in aux end
+
+fun type_consts_of_terms thy ts =
+ Symtab.keys (fold (add_type_consts_in_term thy) ts Symtab.empty);
+
+(* ------------------------------------------------------------------------- *)
+(* HOL to FOL (Isabelle to Metis) *)
+(* ------------------------------------------------------------------------- *)
+
+datatype mode = FO | HO | FT (* first-order, higher-order, fully-typed *)
+
+fun string_of_mode FO = "FO"
+ | string_of_mode HO = "HO"
+ | string_of_mode FT = "FT"
+
+fun fn_isa_to_met_sublevel "equal" = "=" (* FIXME: "c_fequal" *)
+ | fn_isa_to_met_sublevel x = x
+fun fn_isa_to_met_toplevel "equal" = "="
+ | fn_isa_to_met_toplevel x = x
+
+fun metis_lit b c args = (b, (c, args));
+
+fun metis_term_from_combtyp (CombTVar (s, _)) = Metis_Term.Var s
+ | metis_term_from_combtyp (CombTFree (s, _)) = Metis_Term.Fn (s, [])
+ | metis_term_from_combtyp (CombType ((s, _), tps)) =
+ Metis_Term.Fn (s, map metis_term_from_combtyp tps);
+
+(*These two functions insert type literals before the real literals. That is the
+ opposite order from TPTP linkup, but maybe OK.*)
+
+fun hol_term_to_fol_FO tm =
+ case strip_combterm_comb tm of
+ (CombConst ((c, _), _, tys), tms) =>
+ let val tyargs = map metis_term_from_combtyp tys
+ val args = map hol_term_to_fol_FO tms
+ in Metis_Term.Fn (c, tyargs @ args) end
+ | (CombVar ((v, _), _), []) => Metis_Term.Var v
+ | _ => raise Fail "non-first-order combterm"
+
+fun hol_term_to_fol_HO (CombConst ((a, _), _, tylist)) =
+ Metis_Term.Fn (fn_isa_to_met_sublevel a, map metis_term_from_combtyp tylist)
+ | hol_term_to_fol_HO (CombVar ((s, _), _)) = Metis_Term.Var s
+ | hol_term_to_fol_HO (CombApp (tm1, tm2)) =
+ Metis_Term.Fn (".", map hol_term_to_fol_HO [tm1, tm2]);
+
+(*The fully-typed translation, to avoid type errors*)
+fun wrap_type (tm, ty) = Metis_Term.Fn("ti", [tm, metis_term_from_combtyp ty]);
+
+fun hol_term_to_fol_FT (CombVar ((s, _), ty)) = wrap_type (Metis_Term.Var s, ty)
+ | hol_term_to_fol_FT (CombConst((a, _), ty, _)) =
+ wrap_type (Metis_Term.Fn(fn_isa_to_met_sublevel a, []), ty)
+ | hol_term_to_fol_FT (tm as CombApp(tm1,tm2)) =
+ wrap_type (Metis_Term.Fn(".", map hol_term_to_fol_FT [tm1,tm2]),
+ combtyp_of tm)
+
+fun hol_literal_to_fol FO (FOLLiteral (pos, tm)) =
+ let val (CombConst((p, _), _, tys), tms) = strip_combterm_comb tm
+ val tylits = if p = "equal" then [] else map metis_term_from_combtyp tys
+ val lits = map hol_term_to_fol_FO tms
+ in metis_lit pos (fn_isa_to_met_toplevel p) (tylits @ lits) end
+ | hol_literal_to_fol HO (FOLLiteral (pos, tm)) =
+ (case strip_combterm_comb tm of
+ (CombConst(("equal", _), _, _), tms) =>
+ metis_lit pos "=" (map hol_term_to_fol_HO tms)
+ | _ => metis_lit pos "{}" [hol_term_to_fol_HO tm]) (*hBOOL*)
+ | hol_literal_to_fol FT (FOLLiteral (pos, tm)) =
+ (case strip_combterm_comb tm of
+ (CombConst(("equal", _), _, _), tms) =>
+ metis_lit pos "=" (map hol_term_to_fol_FT tms)
+ | _ => metis_lit pos "{}" [hol_term_to_fol_FT tm]) (*hBOOL*);
+
+fun literals_of_hol_term thy mode t =
+ let val (lits, types_sorts) = literals_of_term thy t
+ in (map (hol_literal_to_fol mode) lits, types_sorts) end;
+
+(*Sign should be "true" for conjecture type constraints, "false" for type lits in clauses.*)
+fun metis_of_type_literals pos (TyLitVar ((s, _), (s', _))) =
+ metis_lit pos s [Metis_Term.Var s']
+ | metis_of_type_literals pos (TyLitFree ((s, _), (s', _))) =
+ metis_lit pos s [Metis_Term.Fn (s',[])]
+
+fun default_sort _ (TVar _) = false
+ | default_sort ctxt (TFree (x, s)) = (s = the_default [] (Variable.def_sort ctxt (x, ~1)));
+
+fun metis_of_tfree tf =
+ Metis_Thm.axiom (Metis_LiteralSet.singleton (metis_of_type_literals true tf));
+
+fun hol_thm_to_fol is_conjecture ctxt type_lits mode j skolems th =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val (skolems, (mlits, types_sorts)) =
+ th |> prop_of |> conceal_skolem_terms j skolems
+ ||> (HOLogic.dest_Trueprop #> literals_of_hol_term thy mode)
+ in
+ if is_conjecture then
+ (Metis_Thm.axiom (Metis_LiteralSet.fromList mlits),
+ type_literals_for_types types_sorts, skolems)
+ else
+ let
+ val tylits = filter_out (default_sort ctxt) types_sorts
+ |> type_literals_for_types
+ val mtylits =
+ if type_lits then map (metis_of_type_literals false) tylits else []
+ in
+ (Metis_Thm.axiom (Metis_LiteralSet.fromList(mtylits @ mlits)), [],
+ skolems)
+ end
+ end;
+
+val helpers =
+ [("c_COMBI", (false, map (`I) @{thms COMBI_def})),
+ ("c_COMBK", (false, map (`I) @{thms COMBK_def})),
+ ("c_COMBB", (false, map (`I) @{thms COMBB_def})),
+ ("c_COMBC", (false, map (`I) @{thms COMBC_def})),
+ ("c_COMBS", (false, map (`I) @{thms COMBS_def})),
+ ("c_fequal", (false, map (rpair @{thm equal_imp_equal})
+ @{thms fequal_imp_equal equal_imp_fequal})),
+ ("c_True", (true, map (`I) @{thms True_or_False})),
+ ("c_False", (true, map (`I) @{thms True_or_False})),
+ ("c_If", (true, map (`I) @{thms if_True if_False True_or_False}))]
+
+(* ------------------------------------------------------------------------- *)
+(* Logic maps manage the interface between HOL and first-order logic. *)
+(* ------------------------------------------------------------------------- *)
+
+type logic_map =
+ {axioms: (Metis_Thm.thm * thm) list,
+ tfrees: type_literal list,
+ skolems: (string * term) list}
+
+fun is_quasi_fol_clause thy =
+ Meson.is_fol_term thy o snd o conceal_skolem_terms ~1 [] o prop_of
+
+(*Extract TFree constraints from context to include as conjecture clauses*)
+fun init_tfrees ctxt =
+ let fun add ((a,i),s) Ts = if i = ~1 then TFree(a,s) :: Ts else Ts in
+ Vartab.fold add (#2 (Variable.constraints_of ctxt)) []
+ |> type_literals_for_types
+ end;
+
+(*Insert non-logical axioms corresponding to all accumulated TFrees*)
+fun add_tfrees {axioms, tfrees, skolems} : logic_map =
+ {axioms = map (rpair TrueI o metis_of_tfree) (distinct (op =) tfrees) @
+ axioms,
+ tfrees = tfrees, skolems = skolems}
+
+(*transform isabelle type / arity clause to metis clause *)
+fun add_type_thm [] lmap = lmap
+ | add_type_thm ((ith, mth) :: cls) {axioms, tfrees, skolems} =
+ add_type_thm cls {axioms = (mth, ith) :: axioms, tfrees = tfrees,
+ skolems = skolems}
+
+fun const_in_metis c (pred, tm_list) =
+ let
+ fun in_mterm (Metis_Term.Var _) = false
+ | in_mterm (Metis_Term.Fn (".", tm_list)) = exists in_mterm tm_list
+ | in_mterm (Metis_Term.Fn (nm, tm_list)) = c=nm orelse exists in_mterm tm_list
+ in c = pred orelse exists in_mterm tm_list end;
+
+(* ARITY CLAUSE *)
+fun m_arity_cls (TConsLit ((c, _), (t, _), args)) =
+ metis_lit true c [Metis_Term.Fn(t, map (Metis_Term.Var o fst) args)]
+ | m_arity_cls (TVarLit ((c, _), (s, _))) =
+ metis_lit false c [Metis_Term.Var s]
+(*TrueI is returned as the Isabelle counterpart because there isn't any.*)
+fun arity_cls (ArityClause {conclLit, premLits, ...}) =
+ (TrueI,
+ Metis_Thm.axiom (Metis_LiteralSet.fromList (map m_arity_cls (conclLit :: premLits))));
+
+(* CLASSREL CLAUSE *)
+fun m_class_rel_cls (subclass, _) (superclass, _) =
+ [metis_lit false subclass [Metis_Term.Var "T"], metis_lit true superclass [Metis_Term.Var "T"]];
+fun class_rel_cls (ClassRelClause {subclass, superclass, ...}) =
+ (TrueI, Metis_Thm.axiom (Metis_LiteralSet.fromList (m_class_rel_cls subclass superclass)));
+
+fun type_ext thy tms =
+ let val subs = tfree_classes_of_terms tms
+ val supers = tvar_classes_of_terms tms
+ and tycons = type_consts_of_terms thy tms
+ val (supers', arity_clauses) = make_arity_clauses thy tycons supers
+ val class_rel_clauses = make_class_rel_clauses thy subs supers'
+ in map class_rel_cls class_rel_clauses @ map arity_cls arity_clauses
+ end;
+
+(* Function to generate metis clauses, including comb and type clauses *)
+fun build_logic_map mode0 ctxt type_lits cls ths =
+ let val thy = ProofContext.theory_of ctxt
+ (*The modes FO and FT are sticky. HO can be downgraded to FO.*)
+ fun set_mode FO = FO
+ | set_mode HO =
+ if forall (is_quasi_fol_clause thy) (cls @ ths) then FO else HO
+ | set_mode FT = FT
+ val mode = set_mode mode0
+ (*transform isabelle clause to metis clause *)
+ fun add_thm is_conjecture (metis_ith, isa_ith) {axioms, tfrees, skolems}
+ : logic_map =
+ let
+ val (mth, tfree_lits, skolems) =
+ hol_thm_to_fol is_conjecture ctxt type_lits mode (length axioms)
+ skolems metis_ith
+ in
+ {axioms = (mth, Meson.make_meta_clause isa_ith) :: axioms,
+ tfrees = union (op =) tfree_lits tfrees, skolems = skolems}
+ end;
+ val lmap = {axioms = [], tfrees = init_tfrees ctxt, skolems = []}
+ |> fold (add_thm true o `I) cls
+ |> add_tfrees
+ |> fold (add_thm false o `I) ths
+ val clause_lists = map (Metis_Thm.clause o #1) (#axioms lmap)
+ fun is_used c =
+ exists (Metis_LiteralSet.exists (const_in_metis c o #2)) clause_lists
+ val lmap =
+ if mode = FO then
+ lmap
+ else
+ let
+ val helper_ths =
+ helpers |> filter (is_used o fst)
+ |> maps (fn (c, (needs_full_types, thms)) =>
+ if not (is_used c) orelse
+ needs_full_types andalso mode <> FT then
+ []
+ else
+ thms)
+ in lmap |> fold (add_thm false) helper_ths end
+ in (mode, add_type_thm (type_ext thy (map prop_of (cls @ ths))) lmap) end
+
+end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer.ML Fri Sep 17 17:31:20 2010 +0200
@@ -13,6 +13,7 @@
type relevance_override = Sledgehammer_Filter.relevance_override
type fol_formula = Sledgehammer_Translate.fol_formula
type minimize_command = Sledgehammer_Reconstruct.minimize_command
+
type params =
{blocking: bool,
debug: bool,
@@ -27,12 +28,14 @@
isar_shrink_factor: int,
timeout: Time.time,
expect: string}
+
type problem =
{state: Proof.state,
goal: thm,
subgoal: int,
axioms: (term * ((string * locality) * fol_formula) option) list,
only: bool}
+
type prover_result =
{outcome: failure option,
message: string,
@@ -40,9 +43,10 @@
used_thm_names: (string * locality) list,
atp_run_time_in_msecs: int,
output: string,
- proof: string,
+ tstplike_proof: string,
axiom_names: (string * locality) list vector,
conjecture_shape: int list list}
+
type prover = params -> minimize_command -> problem -> prover_result
val dest_dir : string Config.T
@@ -62,8 +66,9 @@
struct
open ATP_Problem
+open ATP_Proof
open ATP_Systems
-open Metis_Clauses
+open Metis_Translate
open Sledgehammer_Util
open Sledgehammer_Filter
open Sledgehammer_Translate
@@ -111,7 +116,7 @@
used_thm_names: (string * locality) list,
atp_run_time_in_msecs: int,
output: string,
- proof: string,
+ tstplike_proof: string,
axiom_names: (string * locality) list vector,
conjecture_shape: int list list}
@@ -135,102 +140,11 @@
|> Exn.release
|> tap (after path)
-fun extract_delimited (begin_delim, end_delim) output =
- output |> first_field begin_delim |> the |> snd
- |> first_field end_delim |> the |> fst
- |> first_field "\n" |> the |> snd
- handle Option.Option => ""
-
-val tstp_important_message_delims =
- ("% SZS start RequiredInformation", "% SZS end RequiredInformation")
-
-fun extract_important_message output =
- case extract_delimited tstp_important_message_delims output of
- "" => ""
- | s => s |> space_explode "\n" |> filter_out (curry (op =) "")
- |> map (perhaps (try (unprefix "%")))
- |> map (perhaps (try (unprefix " ")))
- |> space_implode "\n " |> quote
-
-(* Splits by the first possible of a list of delimiters. *)
-fun extract_proof delims output =
- case pairself (find_first (fn s => String.isSubstring s output))
- (ListPair.unzip delims) of
- (SOME begin_delim, SOME end_delim) =>
- extract_delimited (begin_delim, end_delim) output
- | _ => ""
-
-fun extract_proof_and_outcome complete res_code proof_delims known_failures
- output =
- case known_failure_in_output output known_failures of
- NONE => (case extract_proof proof_delims output of
- "" => ("", SOME (if res_code = 0 andalso output = "" then
- Interrupted
- else
- UnknownError))
- | proof => if res_code = 0 then (proof, NONE)
- else ("", SOME UnknownError))
- | SOME failure =>
- ("", SOME (if failure = IncompleteUnprovable andalso complete then
- Unprovable
- else
- failure))
+(* generic TPTP-based provers *)
-fun extract_clause_sequence output =
- let
- val tokens_of = String.tokens (not o Char.isAlphaNum)
- fun extract_num ("clause" :: (ss as _ :: _)) =
- Int.fromString (List.last ss)
- | extract_num _ = NONE
- in output |> split_lines |> map_filter (extract_num o tokens_of) end
-
-val set_ClauseFormulaRelationN = "set_ClauseFormulaRelation"
-
-val parse_clause_formula_pair =
- $$ "(" |-- scan_integer --| $$ ","
- -- (Symbol.scan_id ::: Scan.repeat ($$ "," |-- Symbol.scan_id)) --| $$ ")"
- --| Scan.option ($$ ",")
-val parse_clause_formula_relation =
- Scan.this_string set_ClauseFormulaRelationN |-- $$ "("
- |-- Scan.repeat parse_clause_formula_pair
-val extract_clause_formula_relation =
- Substring.full #> Substring.position set_ClauseFormulaRelationN
- #> snd #> Substring.position "." #> fst #> Substring.string
- #> explode #> filter_out Symbol.is_blank #> parse_clause_formula_relation
- #> fst
-
-(* TODO: move to "Sledgehammer_Reconstruct" *)
-fun repair_conjecture_shape_and_theorem_names output conjecture_shape
- axiom_names =
- if String.isSubstring set_ClauseFormulaRelationN output then
- (* This is a hack required for keeping track of axioms after they have been
- clausified by SPASS's Flotter tool. The "ATP/scripts/spass" script is
- also part of this hack. *)
- let
- val j0 = hd (hd conjecture_shape)
- val seq = extract_clause_sequence output
- val name_map = extract_clause_formula_relation output
- fun renumber_conjecture j =
- conjecture_prefix ^ string_of_int (j - j0)
- |> AList.find (fn (s, ss) => member (op =) ss s) name_map
- |> map (fn s => find_index (curry (op =) s) seq + 1)
- fun names_for_number j =
- j |> AList.lookup (op =) name_map |> these
- |> map_filter (try (unprefix axiom_prefix)) |> map unascii_of
- |> map (fn name =>
- (name, name |> find_first_in_list_vector axiom_names
- |> the)
- handle Option.Option =>
- error ("No such fact: " ^ quote name ^ "."))
- in
- (conjecture_shape |> map (maps renumber_conjecture),
- seq |> map names_for_number |> Vector.fromList)
- end
- else
- (conjecture_shape, axiom_names)
-
-
-(* generic TPTP-based provers *)
+(* Important messages are important but not so important that users want to see
+ them each time. *)
+val important_message_keep_factor = 0.1
fun prover_fun auto atp_name
{exec, required_execs, arguments, has_incomplete_mode, proof_delims,
@@ -297,16 +211,16 @@
else
I)
|>> (if measure_run_time then split_time else rpair 0)
- val (proof, outcome) =
- extract_proof_and_outcome complete res_code proof_delims
- known_failures output
- in (output, msecs, proof, outcome) end
+ val (tstplike_proof, outcome) =
+ extract_tstplike_proof_and_outcome complete res_code
+ proof_delims known_failures output
+ in (output, msecs, tstplike_proof, outcome) end
val readable_names = debug andalso overlord
val (problem, pool, conjecture_offset, axiom_names) =
prepare_problem ctxt readable_names explicit_forall full_types
explicit_apply hyp_ts concl_t axioms
- val ss = strings_for_tptp_problem use_conjecture_for_hypotheses
- problem
+ val ss = tptp_strings_for_atp_problem use_conjecture_for_hypotheses
+ problem
val _ = File.write_list probfile ss
val conjecture_shape =
conjecture_offset + 1 upto conjecture_offset + length hyp_ts + 1
@@ -322,8 +236,8 @@
|> run_twice
? (fn (_, msecs0, _, SOME _) =>
run true (Time.- (timeout, Timer.checkRealTimer timer))
- |> (fn (output, msecs, proof, outcome) =>
- (output, msecs0 + msecs, proof, outcome))
+ |> (fn (output, msecs, tstplike_proof, outcome) =>
+ (output, msecs0 + msecs, tstplike_proof, outcome))
| result => result)
in ((pool, conjecture_shape, axiom_names), result) end
else
@@ -339,12 +253,16 @@
else
File.write (Path.explode (Path.implode probfile ^ "_proof")) output
val ((pool, conjecture_shape, axiom_names),
- (output, msecs, proof, outcome)) =
+ (output, msecs, tstplike_proof, outcome)) =
with_path cleanup export run_on problem_path_name
val (conjecture_shape, axiom_names) =
- repair_conjecture_shape_and_theorem_names output conjecture_shape
- axiom_names
- val important_message = extract_important_message output
+ repair_conjecture_shape_and_axiom_names output conjecture_shape
+ axiom_names
+ val important_message =
+ if random () <= important_message_keep_factor then
+ extract_important_message output
+ else
+ ""
val banner = if auto then "Sledgehammer found a proof"
else "Try this command"
val (message, used_thm_names) =
@@ -352,8 +270,8 @@
NONE =>
proof_text isar_proof
(pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
- (banner, full_types, minimize_command, proof, axiom_names, goal,
- subgoal)
+ (banner, full_types, minimize_command, tstplike_proof, axiom_names,
+ goal, subgoal)
|>> (fn message =>
message ^ (if verbose then
"\nATP real CPU time: " ^ string_of_int msecs ^
@@ -369,14 +287,14 @@
in
{outcome = outcome, message = message, pool = pool,
used_thm_names = used_thm_names, atp_run_time_in_msecs = msecs,
- output = output, proof = proof, axiom_names = axiom_names,
- conjecture_shape = conjecture_shape}
+ output = output, tstplike_proof = tstplike_proof,
+ axiom_names = axiom_names, conjecture_shape = conjecture_shape}
end
fun get_prover_fun thy name = prover_fun false name (get_prover thy name)
-fun run_prover (params as {blocking, verbose, max_relevant, timeout, expect,
- ...})
+fun run_prover (params as {blocking, debug, verbose, max_relevant, timeout,
+ expect, ...})
auto i n minimize_command (problem as {state, goal, axioms, ...})
(prover as {default_max_relevant, ...}, atp_name) =
let
@@ -398,14 +316,19 @@
"\n" ^ Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i)))
fun go () =
let
- val (outcome_code, message) =
+ fun really_go () =
prover_fun auto atp_name prover params (minimize_command atp_name)
problem
|> (fn {outcome, message, ...} =>
(if is_some outcome then "none" else "some", message))
- handle ERROR message => ("unknown", "Error: " ^ message ^ "\n")
- | exn => ("unknown", "Internal error:\n" ^
- ML_Compiler.exn_message exn ^ "\n")
+ val (outcome_code, message) =
+ if debug then
+ really_go ()
+ else
+ (really_go ()
+ handle ERROR message => ("unknown", "Error: " ^ message ^ "\n")
+ | exn => ("unknown", "Internal error:\n" ^
+ ML_Compiler.exn_message exn ^ "\n"))
val _ =
if expect = "" orelse outcome_code = expect then
()
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML Fri Sep 17 17:31:20 2010 +0200
@@ -20,7 +20,7 @@
structure Sledgehammer_Minimize : SLEDGEHAMMER_MINIMIZE =
struct
-open ATP_Systems
+open ATP_Proof
open Sledgehammer_Util
open Sledgehammer_Filter
open Sledgehammer_Translate
@@ -117,7 +117,7 @@
val new_timeout =
Int.min (msecs, Time.toMilliseconds time + fudge_msecs)
|> Time.fromMilliseconds
- val (min_thms, {proof, axiom_names, ...}) =
+ val (min_thms, {tstplike_proof, axiom_names, ...}) =
sublinear_minimize (do_test new_timeout)
(filter_used_facts used_thm_names axioms) ([], result)
val n = length min_thms
@@ -130,8 +130,8 @@
(SOME min_thms,
proof_text isar_proof
(pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
- ("Minimized command", full_types, K "", proof, axiom_names, goal,
- i) |> fst)
+ ("Minimized command", full_types, K "", tstplike_proof,
+ axiom_names, goal, i) |> fst)
end
| {outcome = SOME TimedOut, ...} =>
(NONE, "Timeout: You can increase the time limit using the \"timeout\" \
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML Fri Sep 17 17:31:20 2010 +0200
@@ -3,7 +3,7 @@
Author: Claire Quigley, Cambridge University Computer Laboratory
Author: Jasmin Blanchette, TU Muenchen
-Transfer of proofs from external provers.
+Proof reconstruction for Sledgehammer.
*)
signature SLEDGEHAMMER_RECONSTRUCT =
@@ -17,6 +17,9 @@
string Symtab.table * bool * int * Proof.context * int list list
type text_result = string * (string * locality) list
+ val repair_conjecture_shape_and_axiom_names :
+ string -> int list list -> (string * locality) list vector
+ -> int list list * (string * locality) list vector
val metis_proof_text : metis_params -> text_result
val isar_proof_text : isar_params -> metis_params -> text_result
val proof_text : bool -> isar_params -> metis_params -> text_result
@@ -26,7 +29,8 @@
struct
open ATP_Problem
-open Metis_Clauses
+open ATP_Proof
+open Metis_Translate
open Sledgehammer_Util
open Sledgehammer_Filter
open Sledgehammer_Translate
@@ -39,6 +43,139 @@
string Symtab.table * bool * int * Proof.context * int list list
type text_result = string * (string * locality) list
+fun is_head_digit s = Char.isDigit (String.sub (s, 0))
+val scan_integer = Scan.many1 is_head_digit >> (the o Int.fromString o implode)
+
+fun find_first_in_list_vector vec key =
+ Vector.foldl (fn (ps, NONE) => AList.lookup (op =) ps key
+ | (_, value) => value) NONE vec
+
+
+(** SPASS's Flotter hack **)
+
+(* This is a hack required for keeping track of axioms after they have been
+ clausified by SPASS's Flotter tool. The "ATP/scripts/spass" script is also
+ part of this hack. *)
+
+val set_ClauseFormulaRelationN = "set_ClauseFormulaRelation"
+
+fun extract_clause_sequence output =
+ let
+ val tokens_of = String.tokens (not o Char.isAlphaNum)
+ fun extract_num ("clause" :: (ss as _ :: _)) =
+ Int.fromString (List.last ss)
+ | extract_num _ = NONE
+ in output |> split_lines |> map_filter (extract_num o tokens_of) end
+
+val parse_clause_formula_pair =
+ $$ "(" |-- scan_integer --| $$ ","
+ -- (Symbol.scan_id ::: Scan.repeat ($$ "," |-- Symbol.scan_id)) --| $$ ")"
+ --| Scan.option ($$ ",")
+val parse_clause_formula_relation =
+ Scan.this_string set_ClauseFormulaRelationN |-- $$ "("
+ |-- Scan.repeat parse_clause_formula_pair
+val extract_clause_formula_relation =
+ Substring.full #> Substring.position set_ClauseFormulaRelationN
+ #> snd #> Substring.position "." #> fst #> Substring.string
+ #> explode #> filter_out Symbol.is_blank #> parse_clause_formula_relation
+ #> fst
+
+fun repair_conjecture_shape_and_axiom_names output conjecture_shape
+ axiom_names =
+ if String.isSubstring set_ClauseFormulaRelationN output then
+ let
+ val j0 = hd (hd conjecture_shape)
+ val seq = extract_clause_sequence output
+ val name_map = extract_clause_formula_relation output
+ fun renumber_conjecture j =
+ conjecture_prefix ^ string_of_int (j - j0)
+ |> AList.find (fn (s, ss) => member (op =) ss s) name_map
+ |> map (fn s => find_index (curry (op =) s) seq + 1)
+ fun names_for_number j =
+ j |> AList.lookup (op =) name_map |> these
+ |> map_filter (try (unprefix axiom_prefix)) |> map unascii_of
+ |> map (fn name =>
+ (name, name |> find_first_in_list_vector axiom_names
+ |> the)
+ handle Option.Option =>
+ error ("No such fact: " ^ quote name ^ "."))
+ in
+ (conjecture_shape |> map (maps renumber_conjecture),
+ seq |> map names_for_number |> Vector.fromList)
+ end
+ else
+ (conjecture_shape, axiom_names)
+
+
+(** Soft-core proof reconstruction: Metis one-liner **)
+
+fun string_for_label (s, num) = s ^ string_of_int num
+
+fun metis_using [] = ""
+ | metis_using ls =
+ "using " ^ space_implode " " (map string_for_label ls) ^ " "
+fun metis_apply _ 1 = "by "
+ | metis_apply 1 _ = "apply "
+ | metis_apply i _ = "prefer " ^ string_of_int i ^ " apply "
+fun metis_name full_types = if full_types then "metisFT" else "metis"
+fun metis_call full_types [] = metis_name full_types
+ | metis_call full_types ss =
+ "(" ^ metis_name full_types ^ " " ^ space_implode " " ss ^ ")"
+fun metis_command full_types i n (ls, ss) =
+ metis_using ls ^ metis_apply i n ^ metis_call full_types ss
+fun metis_line banner full_types i n ss =
+ banner ^ ": " ^
+ Markup.markup Markup.sendback (metis_command full_types i n ([], ss)) ^ "."
+fun minimize_line _ [] = ""
+ | minimize_line minimize_command ss =
+ case minimize_command ss of
+ "" => ""
+ | command =>
+ "\nTo minimize the number of lemmas, try this: " ^
+ Markup.markup Markup.sendback command ^ "."
+
+fun resolve_axiom axiom_names ((_, SOME s)) =
+ (case strip_prefix_and_unascii axiom_prefix s of
+ SOME s' => (case find_first_in_list_vector axiom_names s' of
+ SOME x => [(s', x)]
+ | NONE => [])
+ | NONE => [])
+ | resolve_axiom axiom_names (num, NONE) =
+ case Int.fromString num of
+ SOME j =>
+ if j > 0 andalso j <= Vector.length axiom_names then
+ Vector.sub (axiom_names, j - 1)
+ else
+ []
+ | NONE => []
+
+fun add_fact axiom_names (Inference (name, _, [])) =
+ append (resolve_axiom axiom_names name)
+ | add_fact _ _ = I
+
+fun used_facts_in_tstplike_proof axiom_names =
+ atp_proof_from_tstplike_string #> rpair [] #-> fold (add_fact axiom_names)
+
+fun used_facts axiom_names =
+ used_facts_in_tstplike_proof axiom_names
+ #> 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, axiom_names, goal, i) =
+ let
+ val (chained_lemmas, other_lemmas) =
+ used_facts axiom_names tstplike_proof
+ val n = Logic.count_prems (prop_of goal)
+ in
+ (metis_line banner full_types i n (map fst other_lemmas) ^
+ minimize_line minimize_command (map fst (other_lemmas @ chained_lemmas)),
+ other_lemmas @ chained_lemmas)
+ end
+
+
+(** Hard-core proof reconstruction: structured Isar proofs **)
+
(* Simple simplifications to ensure that sort annotations don't leave a trail of
spurious "True"s. *)
fun s_not @{const False} = @{const True}
@@ -58,61 +195,9 @@
| s_iff (t1, @{const True}) = t1
| s_iff (t1, t2) = HOLogic.eq_const HOLogic.boolT $ t1 $ t2
-fun mk_anot (AConn (ANot, [phi])) = phi
- | mk_anot phi = AConn (ANot, [phi])
-fun mk_aconn c (phi1, phi2) = AConn (c, [phi1, phi2])
-
fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda v t
fun exists_of v t = HOLogic.exists_const (fastype_of v) $ lambda v t
-datatype raw_step_name = Str of string * string | Num of string
-
-fun raw_step_num (Str (num, _)) = num
- | raw_step_num (Num num) = num
-
-fun same_raw_step s t = raw_step_num s = raw_step_num t
-
-fun raw_step_name_ord p =
- let val q = pairself raw_step_num p in
- (* The "unprefix" part is to cope with remote Vampire's output. The proper
- solution would be to perform a topological sort, e.g. using the nice
- "Graph" functor. *)
- case pairself (Int.fromString o perhaps (try (unprefix "f"))) q of
- (NONE, NONE) => string_ord q
- | (NONE, SOME _) => LESS
- | (SOME _, NONE) => GREATER
- | (SOME i, SOME j) => int_ord (i, j)
- end
-
-fun index_in_shape x = find_index (exists (curry (op =) x))
-fun resolve_axiom axiom_names (Str (_, s)) =
- (case strip_prefix_and_unascii axiom_prefix s of
- SOME s' => (case find_first_in_list_vector axiom_names s' of
- SOME x => [(s', x)]
- | NONE => [])
- | NONE => [])
- | resolve_axiom axiom_names (Num num) =
- case Int.fromString num of
- SOME j =>
- if j > 0 andalso j <= Vector.length axiom_names then
- Vector.sub (axiom_names, j - 1)
- else
- []
- | NONE => []
-val is_axiom = not o null oo resolve_axiom
-
-fun resolve_conjecture conjecture_shape (Str (num, s)) =
- let
- val k = case try (unprefix conjecture_prefix) s of
- SOME s => Int.fromString s |> the_default ~1
- | NONE => case Int.fromString num of
- SOME j => index_in_shape j conjecture_shape
- | NONE => ~1
- in if k >= 0 then [k] else [] end
- | resolve_conjecture conjecture_shape (Num num) =
- resolve_conjecture conjecture_shape (Str (num, "")) (* HACK *)
-val is_conjecture = not o null oo resolve_conjecture
-
fun negate_term (Const (@{const_name All}, T) $ Abs (s, T', t')) =
Const (@{const_name Ex}, T) $ Abs (s, T', negate_term t')
| negate_term (Const (@{const_name Ex}, T) $ Abs (s, T', t')) =
@@ -126,174 +211,32 @@
| negate_term (@{const Not} $ t) = t
| negate_term t = @{const Not} $ t
-datatype ('a, 'b, 'c) raw_step =
- Definition of raw_step_name * 'a * 'b |
- Inference of raw_step_name * 'c * raw_step_name list
-
-(**** PARSING OF TSTP FORMAT ****)
-
-(*Strings enclosed in single quotes, e.g. filenames*)
-val scan_general_id =
- $$ "'" |-- Scan.repeat (~$$ "'") --| $$ "'" >> implode
- || Scan.repeat ($$ "$") -- Scan.many1 Symbol.is_letdig
- >> (fn (ss1, ss2) => implode ss1 ^ implode ss2)
-
-fun repair_name _ "$true" = "c_True"
- | repair_name _ "$false" = "c_False"
- | repair_name _ "$$e" = "c_equal" (* seen in Vampire proofs *)
- | repair_name _ "equal" = "c_equal" (* needed by SPASS? *)
- | repair_name pool s =
- case Symtab.lookup pool s of
- SOME s' => s'
- | NONE =>
- if String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s then
- "c_equal" (* seen in Vampire proofs *)
- else
- s
-(* Generalized first-order terms, which include file names, numbers, etc. *)
-fun parse_annotation strict x =
- ((scan_general_id ::: Scan.repeat ($$ " " |-- scan_general_id)
- >> (strict ? filter (is_some o Int.fromString)))
- -- Scan.optional (parse_annotation strict) [] >> op @
- || $$ "(" |-- parse_annotations strict --| $$ ")"
- || $$ "[" |-- parse_annotations strict --| $$ "]") x
-and parse_annotations strict x =
- (Scan.optional (parse_annotation strict
- ::: Scan.repeat ($$ "," |-- parse_annotation strict)) []
- >> flat) x
+val indent_size = 2
+val no_label = ("", ~1)
-(* Vampire proof lines sometimes contain needless information such as "(0:3)",
- which can be hard to disambiguate from function application in an LL(1)
- parser. As a workaround, we extend the TPTP term syntax with such detritus
- and ignore it. *)
-fun parse_vampire_detritus x =
- (scan_general_id |-- $$ ":" --| scan_general_id >> K []) x
-
-fun parse_term pool x =
- ((scan_general_id >> repair_name pool)
- -- Scan.optional ($$ "(" |-- (parse_vampire_detritus || parse_terms pool)
- --| $$ ")") []
- --| Scan.optional ($$ "(" |-- parse_vampire_detritus --| $$ ")") []
- >> ATerm) x
-and parse_terms pool x =
- (parse_term pool ::: Scan.repeat ($$ "," |-- parse_term pool)) x
-
-fun parse_atom pool =
- parse_term pool -- Scan.option (Scan.option ($$ "!") --| $$ "="
- -- parse_term pool)
- >> (fn (u1, NONE) => AAtom u1
- | (u1, SOME (NONE, u2)) => AAtom (ATerm ("c_equal", [u1, u2]))
- | (u1, SOME (SOME _, u2)) =>
- mk_anot (AAtom (ATerm ("c_equal", [u1, u2]))))
-
-fun fo_term_head (ATerm (s, _)) = s
-
-(* TPTP formulas are fully parenthesized, so we don't need to worry about
- operator precedence. *)
-fun parse_formula pool x =
- (($$ "(" |-- parse_formula pool --| $$ ")"
- || ($$ "!" >> K AForall || $$ "?" >> K AExists)
- --| $$ "[" -- parse_terms pool --| $$ "]" --| $$ ":"
- -- parse_formula pool
- >> (fn ((q, ts), phi) => AQuant (q, map fo_term_head ts, phi))
- || $$ "~" |-- parse_formula pool >> mk_anot
- || parse_atom pool)
- -- Scan.option ((Scan.this_string "=>" >> K AImplies
- || Scan.this_string "<=>" >> K AIff
- || Scan.this_string "<~>" >> K ANotIff
- || Scan.this_string "<=" >> K AIf
- || $$ "|" >> K AOr || $$ "&" >> K AAnd)
- -- parse_formula pool)
- >> (fn (phi1, NONE) => phi1
- | (phi1, SOME (c, phi2)) => mk_aconn c (phi1, phi2))) x
+val raw_prefix = "X"
+val assum_prefix = "A"
+val fact_prefix = "F"
-val parse_tstp_extra_arguments =
- Scan.optional ($$ "," |-- parse_annotation false
- --| Scan.option ($$ "," |-- parse_annotations false)) []
-
-(* Syntax: (fof|cnf)\(<num>, <formula_role>, <formula> <extra_arguments>\).
- The <num> could be an identifier, but we assume integers. *)
- fun parse_tstp_line pool =
- ((Scan.this_string "fof" || Scan.this_string "cnf") -- $$ "(")
- |-- scan_general_id --| $$ "," -- Symbol.scan_id --| $$ ","
- -- parse_formula pool -- parse_tstp_extra_arguments --| $$ ")" --| $$ "."
- >> (fn (((num, role), phi), deps) =>
- let
- val (name, deps) =
- case deps of
- ["file", _, s] => (Str (num, s), [])
- | _ => (Num num, deps)
- in
- case role of
- "definition" =>
- (case phi of
- AConn (AIff, [phi1 as AAtom _, phi2]) =>
- Definition (name, phi1, phi2)
- | AAtom (ATerm ("c_equal", _)) =>
- (* Vampire's equality proxy axiom *)
- Inference (name, phi, map Num deps)
- | _ => raise Fail "malformed definition")
- | _ => Inference (name, phi, map Num deps)
- end)
-
-(**** PARSING OF VAMPIRE OUTPUT ****)
-
-(* Syntax: <num>. <formula> <annotation> *)
-fun parse_vampire_line pool =
- scan_general_id --| $$ "." -- parse_formula pool -- parse_annotation true
- >> (fn ((num, phi), deps) => Inference (Num num, phi, map Num deps))
-
-(**** PARSING OF SPASS OUTPUT ****)
-
-(* SPASS returns clause references of the form "x.y". We ignore "y", whose role
- is not clear anyway. *)
-val parse_dot_name = scan_general_id --| $$ "." --| scan_general_id
+fun resolve_conjecture conjecture_shape (num, s_opt) =
+ let
+ val k = case try (unprefix conjecture_prefix) (the_default "" s_opt) of
+ SOME s => Int.fromString s |> the_default ~1
+ | NONE => case Int.fromString num of
+ SOME j => find_index (exists (curry (op =) j))
+ conjecture_shape
+ | NONE => ~1
+ in if k >= 0 then [k] else [] end
-val parse_spass_annotations =
- Scan.optional ($$ ":" |-- Scan.repeat (parse_dot_name
- --| Scan.option ($$ ","))) []
-
-(* It is not clear why some literals are followed by sequences of stars and/or
- pluses. We ignore them. *)
-fun parse_decorated_atom pool =
- parse_atom pool --| Scan.repeat ($$ "*" || $$ "+" || $$ " ")
-
-fun mk_horn ([], []) = AAtom (ATerm ("c_False", []))
- | mk_horn ([], pos_lits) = foldr1 (mk_aconn AOr) pos_lits
- | mk_horn (neg_lits, []) = mk_anot (foldr1 (mk_aconn AAnd) neg_lits)
- | mk_horn (neg_lits, pos_lits) =
- mk_aconn AImplies (foldr1 (mk_aconn AAnd) neg_lits,
- foldr1 (mk_aconn AOr) pos_lits)
-
-fun parse_horn_clause pool =
- Scan.repeat (parse_decorated_atom pool) --| $$ "|" --| $$ "|"
- -- Scan.repeat (parse_decorated_atom pool) --| $$ "-" --| $$ ">"
- -- Scan.repeat (parse_decorated_atom pool)
- >> (mk_horn o apfst (op @))
+val is_axiom = not o null oo resolve_axiom
+val is_conjecture = not o null oo resolve_conjecture
-(* Syntax: <num>[0:<inference><annotations>]
- <atoms> || <atoms> -> <atoms>. *)
-fun parse_spass_line pool =
- scan_general_id --| $$ "[" --| $$ "0" --| $$ ":" --| Symbol.scan_id
- -- parse_spass_annotations --| $$ "]" -- parse_horn_clause pool --| $$ "."
- >> (fn ((num, deps), u) => Inference (Num num, u, map Num deps))
-
-fun parse_line pool =
- parse_tstp_line pool || parse_vampire_line pool || parse_spass_line pool
-fun parse_lines pool = Scan.repeat1 (parse_line pool)
-fun parse_proof pool =
- fst o Scan.finite Symbol.stopper
- (Scan.error (!! (fn _ => raise Fail "unrecognized ATP output")
- (parse_lines pool)))
- o explode o strip_spaces_except_between_ident_chars
-
-fun clean_up_dependencies _ [] = []
- | clean_up_dependencies seen ((step as Definition (name, _, _)) :: steps) =
- step :: clean_up_dependencies (name :: seen) steps
- | clean_up_dependencies seen (Inference (name, u, deps) :: steps) =
- Inference (name, u,
- map_filter (fn dep => find_first (same_raw_step dep) seen) deps) ::
- clean_up_dependencies (name :: seen) steps
+fun raw_label_for_name conjecture_shape name =
+ case resolve_conjecture conjecture_shape name of
+ [j] => (conjecture_prefix, j)
+ | _ => case Int.fromString (fst name) of
+ SOME j => (raw_prefix, j)
+ | NONE => (raw_prefix ^ fst name, 0)
(**** INTERPRETATION OF TSTP SYNTAX TREES ****)
@@ -411,7 +354,7 @@
case strip_prefix_and_unascii schematic_var_prefix a of
SOME b => Var ((b, 0), T)
| NONE =>
- if is_tptp_variable a then
+ if is_atp_variable a then
Var ((repair_atp_variable_name Char.toLower a, 0), T)
else
(* Skolem constants? *)
@@ -537,7 +480,7 @@
val is_only_type_information = curry (op aconv) HOLogic.true_const
fun replace_one_dependency (old, new) dep =
- if raw_step_num dep = raw_step_num old then new else [dep]
+ if is_same_step (dep, old) then new else [dep]
fun replace_dependencies_in_line _ (line as Definition _) = line
| replace_dependencies_in_line p (Inference (name, t, deps)) =
Inference (name, t, fold (union (op =) o replace_one_dependency p) deps [])
@@ -607,104 +550,6 @@
else
map (replace_dependencies_in_line (name, deps)) lines) (* drop line *)
-(** EXTRACTING LEMMAS **)
-
-(* Like "split_line" but splits on ".\n" (for TSTP and SPASS) or "]\n" (for
- Vampire). *)
-val split_proof_lines =
- let
- fun aux [] [] = []
- | aux line [] = [implode (rev line)]
- | aux line ("." :: "\n" :: rest) = aux line [] @ aux [] rest
- | aux line ("]" :: "\n" :: rest) = aux line [] @ aux [] rest
- | aux line (s :: rest) = aux (s :: line) rest
- in aux [] o explode end
-
-(* A list consisting of the first number in each line is returned. For TSTP,
- interesting lines have the form "fof(108, axiom, ...)", where the number
- (108) is extracted. For SPASS, lines have the form "108[0:Inp] ...", where
- the first number (108) is extracted. For Vampire, we look for
- "108. ... [input]". *)
-fun used_facts_in_atp_proof axiom_names atp_proof =
- let
- val tokens_of =
- String.tokens (fn c => not (Char.isAlphaNum c) andalso c <> #"_")
- fun do_line (tag :: num :: "axiom" :: (rest as _ :: _)) =
- if tag = "cnf" orelse tag = "fof" then
- (case strip_prefix_and_unascii axiom_prefix (List.last rest) of
- SOME name =>
- if member (op =) rest "file" then
- ([(name, name |> find_first_in_list_vector axiom_names |> the)]
- handle Option.Option =>
- error ("No such fact: " ^ quote name ^ "."))
- else
- resolve_axiom axiom_names (Num num)
- | NONE => resolve_axiom axiom_names (Num num))
- else
- []
- | do_line (num :: "0" :: "Inp" :: _) = resolve_axiom axiom_names (Num num)
- | do_line (num :: rest) =
- (case List.last rest of
- "input" => resolve_axiom axiom_names (Num num)
- | _ => [])
- | do_line _ = []
- in atp_proof |> split_proof_lines |> maps (do_line o tokens_of) end
-
-val indent_size = 2
-val no_label = ("", ~1)
-
-val raw_prefix = "X"
-val assum_prefix = "A"
-val fact_prefix = "F"
-
-fun string_for_label (s, num) = s ^ string_of_int num
-
-fun raw_label_for_name conjecture_shape name =
- case resolve_conjecture conjecture_shape name of
- [j] => (conjecture_prefix, j)
- | _ => case Int.fromString (raw_step_num name) of
- SOME j => (raw_prefix, j)
- | NONE => (raw_prefix ^ raw_step_num name, 0)
-
-fun metis_using [] = ""
- | metis_using ls =
- "using " ^ space_implode " " (map string_for_label ls) ^ " "
-fun metis_apply _ 1 = "by "
- | metis_apply 1 _ = "apply "
- | metis_apply i _ = "prefer " ^ string_of_int i ^ " apply "
-fun metis_name full_types = if full_types then "metisFT" else "metis"
-fun metis_call full_types [] = metis_name full_types
- | metis_call full_types ss =
- "(" ^ metis_name full_types ^ " " ^ space_implode " " ss ^ ")"
-fun metis_command full_types i n (ls, ss) =
- metis_using ls ^ metis_apply i n ^ metis_call full_types ss
-fun metis_line banner full_types i n ss =
- banner ^ ": " ^
- Markup.markup Markup.sendback (metis_command full_types i n ([], ss)) ^ "."
-fun minimize_line _ [] = ""
- | minimize_line minimize_command ss =
- case minimize_command ss of
- "" => ""
- | command =>
- "\nTo minimize the number of lemmas, try this: " ^
- Markup.markup Markup.sendback command ^ "."
-
-fun used_facts axiom_names =
- used_facts_in_atp_proof axiom_names
- #> List.partition (curry (op =) Chained o snd)
- #> pairself (sort_distinct (string_ord o pairself fst))
-
-fun metis_proof_text (banner, full_types, minimize_command, atp_proof,
- axiom_names, goal, i) =
- let
- val (chained_lemmas, other_lemmas) = used_facts axiom_names atp_proof
- val n = Logic.count_prems (prop_of goal)
- in
- (metis_line banner full_types i n (map fst other_lemmas) ^
- minimize_line minimize_command (map fst (other_lemmas @ chained_lemmas)),
- other_lemmas @ chained_lemmas)
- end
-
(** Isar proof construction and manipulation **)
fun merge_fact_sets (ls1, ss1) (ls2, ss2) =
@@ -713,16 +558,16 @@
type label = string * int
type facts = label list * string list
-datatype qualifier = Show | Then | Moreover | Ultimately
+datatype isar_qualifier = Show | Then | Moreover | Ultimately
-datatype step =
+datatype isar_step =
Fix of (string * typ) list |
Let of term * term |
Assume of label * term |
- Have of qualifier list * label * term * byline
+ Have of isar_qualifier list * label * term * byline
and byline =
ByMetis of facts |
- CaseSplit of step list list * facts
+ CaseSplit of isar_step list list * facts
fun smart_case_split [] facts = ByMetis facts
| smart_case_split proofs facts = CaseSplit (proofs, facts)
@@ -743,17 +588,24 @@
ByMetis (fold (add_fact_from_dependency conjecture_shape axiom_names)
deps ([], [])))
-fun raw_step_name (Definition (name, _, _)) = name
- | raw_step_name (Inference (name, _, _)) = name
+fun repair_name "$true" = "c_True"
+ | repair_name "$false" = "c_False"
+ | repair_name "$$e" = "c_equal" (* seen in Vampire proofs *)
+ | repair_name "equal" = "c_equal" (* needed by SPASS? *)
+ | repair_name s =
+ if String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s then
+ "c_equal" (* seen in Vampire proofs *)
+ else
+ s
-fun proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor
- atp_proof conjecture_shape axiom_names params frees =
+fun isar_proof_from_tstplike_proof pool ctxt full_types tfrees isar_shrink_factor
+ tstplike_proof conjecture_shape axiom_names params frees =
let
val lines =
- atp_proof ^ "$" (* the $ sign acts as a sentinel (FIXME: needed?) *)
- |> parse_proof pool
- |> sort (raw_step_name_ord o pairself raw_step_name)
- |> clean_up_dependencies []
+ tstplike_proof
+ |> atp_proof_from_tstplike_string
+ |> nasty_atp_proof pool
+ |> map_term_names_in_atp_proof repair_name
|> decode_lines ctxt full_types tfrees
|> rpair [] |-> fold_rev (add_line conjecture_shape axiom_names)
|> rpair [] |-> fold_rev add_nontrivial_line
@@ -1049,14 +901,13 @@
and do_proof [Have (_, _, _, ByMetis _)] = ""
| do_proof proof =
(if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^
- do_indent 0 ^ "proof -\n" ^
- do_steps "" "" 1 proof ^
- do_indent 0 ^ (if n <> 1 then "next" else "qed")
+ do_indent 0 ^ "proof -\n" ^ do_steps "" "" 1 proof ^ do_indent 0 ^
+ (if n <> 1 then "next" else "qed")
in do_proof end
fun isar_proof_text (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
- (other_params as (_, full_types, _, atp_proof, axiom_names,
- goal, i)) =
+ (other_params as (_, full_types, _, tstplike_proof,
+ axiom_names, goal, i)) =
let
val (params, hyp_ts, concl_t) = strip_subgoal goal i
val frees = fold Term.add_frees (concl_t :: hyp_ts) []
@@ -1064,9 +915,9 @@
val n = Logic.count_prems (prop_of goal)
val (one_line_proof, lemma_names) = metis_proof_text other_params
fun isar_proof_for () =
- case proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor
- atp_proof conjecture_shape axiom_names params
- frees
+ case isar_proof_from_tstplike_proof pool ctxt full_types tfrees
+ isar_shrink_factor tstplike_proof conjecture_shape axiom_names
+ params frees
|> redirect_proof hyp_ts concl_t
|> kill_duplicate_assumptions_in_proof
|> then_chain_proof
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_translate.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_translate.ML Fri Sep 17 17:31:20 2010 +0200
@@ -3,7 +3,7 @@
Author: Makarius
Author: Jasmin Blanchette, TU Muenchen
-Translation of HOL to FOL.
+Translation of HOL to FOL for Sledgehammer.
*)
signature SLEDGEHAMMER_TRANSLATE =
@@ -30,7 +30,7 @@
struct
open ATP_Problem
-open Metis_Clauses
+open Metis_Translate
open Sledgehammer_Util
val axiom_prefix = "ax_"
@@ -376,7 +376,7 @@
type const_info = {min_arity: int, max_arity: int, sub_level: bool}
fun consider_term top_level (ATerm ((s, _), ts)) =
- (if is_tptp_variable s then
+ (if is_atp_variable s then
I
else
let val n = length ts in
@@ -468,7 +468,7 @@
fun close_universally phi =
let
fun term_vars bounds (ATerm (name as (s, _), tms)) =
- (is_tptp_variable s andalso not (member (op =) bounds name))
+ (is_atp_variable s andalso not (member (op =) bounds name))
? insert (op =) name
#> fold (term_vars bounds) tms
fun formula_vars bounds (AQuant (_, xs, phi)) =
@@ -524,7 +524,7 @@
("Conjectures", conjecture_lines),
("Type variables", tfree_lines)]
|> repair_problem thy explicit_forall full_types explicit_apply
- val (problem, pool) = nice_tptp_problem readable_names problem
+ val (problem, pool) = nice_atp_problem readable_names problem
val conjecture_offset =
length axiom_lines + length class_rel_lines + length arity_lines
+ length helper_lines
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML Fri Sep 17 17:31:20 2010 +0200
@@ -6,14 +6,11 @@
signature SLEDGEHAMMER_UTIL =
sig
- val find_first_in_list_vector : (''a * 'b) list vector -> ''a -> 'b option
val plural_s : int -> string
val serial_commas : string -> string list -> string list
val simplify_spaces : string -> string
- val strip_spaces_except_between_ident_chars : string -> string
val parse_bool_option : bool -> string -> string -> bool option
val parse_time_option : string -> string -> Time.time option
- val scan_integer : string list -> int * string list
val nat_subscript : int -> string
val unyxml : string -> string
val maybe_quote : string -> string
@@ -29,10 +26,6 @@
structure Sledgehammer_Util : SLEDGEHAMMER_UTIL =
struct
-fun find_first_in_list_vector vec key =
- Vector.foldl (fn (ps, NONE) => AList.lookup (op =) ps key
- | (_, value) => value) NONE vec
-
fun plural_s n = if n = 1 then "" else "s"
fun serial_commas _ [] = ["??"]
@@ -41,28 +34,7 @@
| serial_commas conj [s1, s2, s3] = [s1 ^ ",", s2 ^ ",", conj, s3]
| serial_commas conj (s :: ss) = s ^ "," :: serial_commas conj ss
-fun strip_spaces_in_list _ [] = []
- | strip_spaces_in_list _ [c1] = if Char.isSpace c1 then [] else [str c1]
- | strip_spaces_in_list is_evil [c1, c2] =
- strip_spaces_in_list is_evil [c1] @ strip_spaces_in_list is_evil [c2]
- | strip_spaces_in_list is_evil (c1 :: c2 :: c3 :: cs) =
- if Char.isSpace c1 then
- strip_spaces_in_list is_evil (c2 :: c3 :: cs)
- else if Char.isSpace c2 then
- if Char.isSpace c3 then
- strip_spaces_in_list is_evil (c1 :: c3 :: cs)
- else
- str c1 :: (if forall is_evil [c1, c3] then [" "] else []) @
- strip_spaces_in_list is_evil (c3 :: cs)
- else
- str c1 :: strip_spaces_in_list is_evil (c2 :: c3 :: cs)
-fun strip_spaces is_evil =
- implode o strip_spaces_in_list is_evil o String.explode
-
-val simplify_spaces = strip_spaces (K true)
-
-fun is_ident_char c = Char.isAlphaNum c orelse c = #"_"
-val strip_spaces_except_between_ident_chars = strip_spaces is_ident_char
+val simplify_spaces = ATP_Proof.strip_spaces (K true)
fun parse_bool_option option name s =
(case s of
@@ -94,9 +66,6 @@
SOME (Time.fromMilliseconds msecs)
end
-fun is_head_digit s = Char.isDigit (String.sub (s, 0))
-val scan_integer = Scan.many1 is_head_digit >> (the o Int.fromString o implode)
-
val subscript = implode o map (prefix "\<^isub>") o explode
fun nat_subscript n =
n |> string_of_int |> print_mode_active Symbol.xsymbolsN ? subscript
--- a/src/HOL/Tools/quickcheck_generators.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/HOL/Tools/quickcheck_generators.ML Fri Sep 17 17:31:20 2010 +0200
@@ -392,16 +392,16 @@
in if Quickcheck.report ctxt then
let
val t' = mk_reporting_generator_expr thy t Ts;
- val compile = Code_Runtime.value (SOME target)
+ val compile = Code_Runtime.dynamic_value_strict
(Counterexample_Report.get, put_counterexample_report, "Quickcheck_Generators.put_counterexample_report")
- (fn proc => fn g => fn s => g s #>> (apfst o Option.map o map) proc) thy t' [];
+ thy (SOME target) (fn proc => fn g => fn s => g s #>> (apfst o Option.map o map) proc) t' [];
in compile #> Random_Engine.run end
else
let
val t' = mk_generator_expr thy t Ts;
- val compile = Code_Runtime.value (SOME target)
+ val compile = Code_Runtime.dynamic_value_strict
(Counterexample.get, put_counterexample, "Quickcheck_Generators.put_counterexample")
- (fn proc => fn g => fn s => g s #>> (Option.map o map) proc) thy t' [];
+ thy (SOME target) (fn proc => fn g => fn s => g s #>> (Option.map o map) proc) t' [];
val dummy_report = ([], false)
in compile #> Random_Engine.run #> rpair dummy_report end
end;
--- a/src/Pure/General/exn.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Pure/General/exn.ML Fri Sep 17 17:31:20 2010 +0200
@@ -11,6 +11,7 @@
val get_exn: 'a result -> exn option
val capture: ('a -> 'b) -> 'a -> 'b result
val release: 'a result -> 'a
+ val map_result: ('a -> 'a) -> 'a result -> 'a result
exception Interrupt
val interrupt: unit -> 'a
val is_interrupt: exn -> bool
@@ -43,6 +44,9 @@
fun release (Result y) = y
| release (Exn e) = reraise e;
+fun map_result f (Result x) = Result (f x)
+ | map_result f e = e;
+
(* interrupts *)
--- a/src/Pure/Thy/latex.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Pure/Thy/latex.ML Fri Sep 17 17:31:20 2010 +0200
@@ -13,6 +13,7 @@
val output_markup: string -> string -> string
val output_markup_env: string -> string -> string
val output_verbatim: string -> string
+ val output_typewriter: string -> string
val markup_true: string
val markup_false: string
val begin_delim: string -> string
@@ -97,6 +98,38 @@
end;
+fun output_typewriter s =
+ let
+ val parse = Scan.repeat
+ (Scan.this_string "\n" |-- Scan.succeed "\\\\\n\\hspace*{0pt}"
+ || (Scan.this_string " "
+ || Scan.this_string "."
+ || Scan.this_string ","
+ || Scan.this_string ":"
+ || Scan.this_string ";"
+ || Scan.this_string "\"" |-- Scan.succeed "{\\char34}"
+ || Scan.this_string "#" |-- Scan.succeed "{\\char35}"
+ || Scan.this_string "$" |-- Scan.succeed "{\\char36}"
+ || Scan.this_string "%" |-- Scan.succeed "{\\char37}"
+ || Scan.this_string "&" |-- Scan.succeed "{\\char38}"
+ || Scan.this_string "\\" |-- Scan.succeed "{\\char92}"
+ || Scan.this_string "^" |-- Scan.succeed "{\\char94}"
+ || Scan.this_string "_" |-- Scan.succeed "{\\char95}"
+ || Scan.this_string "{" |-- Scan.succeed "{\\char123}"
+ || Scan.this_string "}" |-- Scan.succeed "{\\char125}"
+ || Scan.this_string "~" |-- Scan.succeed "{\\char126}")
+ -- Scan.repeat (Scan.this_string " ")
+ >> (fn (x, xs) => x ^ replicate_string (length xs) "~")
+ || Scan.one Symbol.not_eof);
+ fun is_newline s = (s = "\n");
+ val cs = explode s |> take_prefix is_newline |> snd
+ |> take_suffix is_newline |> fst;
+ val (texts, []) = Scan.finite Symbol.stopper parse cs
+ in if null texts
+ then ""
+ else implode ("\\isatypewriter%\n\\noindent%\n\\hspace*{0pt}" :: texts)
+ end
+
(* token output *)
--- a/src/Tools/Code/code_preproc.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Code/code_preproc.ML Fri Sep 17 17:31:20 2010 +0200
@@ -29,7 +29,9 @@
val dynamic_eval_value: theory -> ((term -> term) -> 'a -> 'a)
-> (code_algebra -> code_graph -> (string * sort) list -> term -> 'a) -> term -> 'a
val static_eval_conv: theory -> string list
- -> (code_algebra -> code_graph -> (string * sort) list -> term -> conv) -> conv
+ -> (code_algebra -> code_graph -> theory -> (string * sort) list -> term -> conv) -> conv
+ val static_eval_value: theory -> ((term -> term) -> 'a -> 'a) -> string list
+ -> (code_algebra -> code_graph -> theory -> (string * sort) list -> term -> 'a) -> term -> 'a
val setup: theory -> theory
end
@@ -138,6 +140,8 @@
fun preprocess_conv thy ct =
let
+ val ctxt = ProofContext.init_global thy;
+ val _ = (Sign.no_frees ctxt o map_types (K dummyT) o Sign.no_vars ctxt) (Thm.term_of ct);
val pre = (Simplifier.global_context thy o #pre o the_thmproc) thy;
in
ct
@@ -145,6 +149,8 @@
|> trans_conv_rule (AxClass.unoverload_conv thy)
end;
+fun preprocess_term thy = term_of_conv thy (preprocess_conv thy);
+
fun postprocess_conv thy ct =
let
val post = (Simplifier.global_context thy o #post o the_thmproc) thy;
@@ -427,43 +433,56 @@
fun dest_cterm ct = let val t = Thm.term_of ct in (Term.add_tfrees t [], t) end;
-fun dynamic_eval thy cterm_of conclude_evaluation evaluator proto_ct =
+fun dynamic_eval_conv thy conv ct =
let
- val ctxt = Syntax.init_pretty_global thy;
- val ct = cterm_of proto_ct;
- val _ = (Sign.no_frees ctxt o map_types (K dummyT) o Sign.no_vars ctxt) (Thm.term_of ct);
- val thm = preprocess_conv thy ct;
- val ct' = Thm.rhs_of thm;
+ val thm1 = preprocess_conv thy ct;
+ val ct' = Thm.rhs_of thm1;
val (vs', t') = dest_cterm ct';
val consts = fold_aterms
(fn Const (c, _) => insert (op =) c | _ => I) t' [];
val (algebra', eqngr') = obtain false thy consts [t'];
- in conclude_evaluation (evaluator algebra' eqngr' vs' t' ct') thm end;
+ val thm2 = conv algebra' eqngr' vs' t' ct';
+ val thm3 = postprocess_conv thy (Thm.rhs_of thm2);
+ in
+ Thm.transitive thm1 (Thm.transitive thm2 thm3) handle THM _ =>
+ error ("could not construct evaluation proof:\n"
+ ^ (cat_lines o map (Display.string_of_thm_global thy)) [thm1, thm2, thm3])
+ end;
-fun dynamic_eval_conv thy =
+fun dynamic_eval_value thy postproc evaluator t =
let
- fun conclude_evaluation thm2 thm1 =
- let
- val thm3 = postprocess_conv thy (Thm.rhs_of thm2);
- in
- Thm.transitive thm1 (Thm.transitive thm2 thm3) handle THM _ =>
- error ("could not construct evaluation proof:\n"
- ^ (cat_lines o map (Display.string_of_thm_global thy)) [thm1, thm2, thm3])
- end;
- in dynamic_eval thy I conclude_evaluation end;
-
-fun dynamic_eval_value thy postproc evaluator = dynamic_eval thy (Thm.cterm_of thy)
- (K o postproc (postprocess_term thy)) (K oooo evaluator);
+ val t' = preprocess_term thy t;
+ val vs' = Term.add_tfrees t' [];
+ val consts = fold_aterms
+ (fn Const (c, _) => insert (op =) c | _ => I) t' [];
+ val (algebra', eqngr') = obtain false thy consts [t'];
+ val result = evaluator algebra' eqngr' vs' t';
+ in
+ evaluator algebra' eqngr' vs' t'
+ |> postproc (postprocess_term thy)
+ end;
fun static_eval_conv thy consts conv =
let
val (algebra, eqngr) = obtain true thy consts [];
+ val conv' = conv algebra eqngr;
in
Conv.tap_thy (fn thy => (preprocess_conv thy)
- then_conv (fn ct => uncurry (conv algebra eqngr) (dest_cterm ct) ct)
+ then_conv (fn ct => uncurry (conv' thy) (dest_cterm ct) ct)
then_conv (postprocess_conv thy))
end;
+fun static_eval_value thy postproc consts evaluator =
+ let
+ val (algebra, eqngr) = obtain true thy consts [];
+ val evaluator' = evaluator algebra eqngr;
+ in fn t =>
+ t
+ |> preprocess_term thy
+ |> (fn t => evaluator' thy (Term.add_tfrees t []) t)
+ |> postproc (postprocess_term thy)
+ end;
+
(** setup **)
--- a/src/Tools/Code/code_runtime.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Code/code_runtime.ML Fri Sep 17 17:31:20 2010 +0200
@@ -7,41 +7,151 @@
signature CODE_RUNTIME =
sig
val target: string
- val value: string option
- -> (Proof.context -> unit -> 'a) * ((unit -> 'a) -> Proof.context -> Proof.context) * string
- -> ((term -> term) -> 'a -> 'a) -> theory -> term -> string list -> 'a
+ type 'a cookie = (Proof.context -> unit -> 'a) * ((unit -> 'a) -> Proof.context -> Proof.context) * string
+ val dynamic_value: 'a cookie -> theory -> string option
+ -> ((term -> term) -> 'a -> 'a) -> term -> string list -> 'a option
+ val dynamic_value_strict: 'a cookie -> theory -> string option
+ -> ((term -> term) -> 'a -> 'a) -> term -> string list -> 'a
+ val dynamic_value_exn: 'a cookie -> theory -> string option
+ -> ((term -> term) -> 'a -> 'a) -> term -> string list -> 'a Exn.result
+ val static_value: 'a cookie -> theory -> string option
+ -> ((term -> term) -> 'a -> 'a) -> string list -> term -> 'a option
+ val static_value_strict: 'a cookie -> theory -> string option
+ -> ((term -> term) -> 'a -> 'a) -> string list -> term -> 'a
+ val static_value_exn: 'a cookie -> theory -> string option
+ -> ((term -> term) -> 'a -> 'a) -> string list -> term -> 'a Exn.result
+ val dynamic_holds_conv: conv
+ val static_holds_conv: theory -> string list -> conv
val code_reflect: (string * string list) list -> string list -> string -> string option
-> theory -> theory
- val setup: theory -> theory
+ datatype truth = Holds
+ val put_truth: (unit -> truth) -> Proof.context -> Proof.context
end;
structure Code_Runtime : CODE_RUNTIME =
struct
+open Basic_Code_Thingol;
+
(** evaluation **)
+(* technical prerequisites *)
+
+val this = "Code_Runtime";
+val s_truth = Long_Name.append this "truth";
+val s_Holds = Long_Name.append this "Holds";
+
val target = "Eval";
-val truth_struct = "Code_Truth";
+datatype truth = Holds;
-fun value some_target cookie postproc thy t args =
+val _ = Context.>> (Context.map_theory
+ (Code_Target.extend_target (target, (Code_ML.target_SML, K I))
+ #> Code_Target.add_tyco_syntax target @{type_name prop}
+ (SOME (0, (K o K o K) (Code_Printer.str s_truth)))
+ #> Code_Target.add_const_syntax target @{const_name Code_Generator.holds}
+ (SOME (Code_Printer.plain_const_syntax s_Holds))
+ #> Code_Target.add_reserved target this
+ #> fold (Code_Target.add_reserved target) ["oo", "ooo", "oooo", "upto", "downto", "orf", "andf"]));
+ (*avoid further pervasive infix names*)
+
+
+(* evaluation into target language values *)
+
+type 'a cookie = (Proof.context -> unit -> 'a) * ((unit -> 'a) -> Proof.context -> Proof.context) * string;
+
+fun obtain_serializer thy some_target = Code_Target.produce_code_for thy
+ (the_default target some_target) NONE "Code" [];
+
+fun base_evaluator cookie serializer naming thy program ((vs, ty), t) deps args =
let
val ctxt = ProofContext.init_global thy;
- fun evaluator naming program ((_, (_, ty)), t) deps =
- let
- val _ = if Code_Thingol.contains_dictvar t then
- error "Term to be evaluated contains free dictionaries" else ();
- val value_name = "Value.VALUE.value"
- val program' = program
- |> Graph.new_node (value_name,
- Code_Thingol.Fun (Term.dummy_patternN, ((([], ty), [(([], t), (NONE, true))]), NONE)))
- |> fold (curry Graph.add_edge value_name) deps;
- val (program_code, [SOME value_name']) = Code_Target.produce_code_for thy
- (the_default target some_target) NONE "Code" [] naming program' [value_name];
- val value_code = space_implode " "
- (value_name' :: map (enclose "(" ")") args);
- in ML_Context.value ctxt cookie (program_code, value_code) end;
- in Code_Thingol.dynamic_eval_value thy postproc evaluator t end;
+ val _ = if Code_Thingol.contains_dictvar t then
+ error "Term to be evaluated contains free dictionaries" else ();
+ val v' = Name.variant (map fst vs) "a";
+ val vs' = (v', []) :: vs
+ val ty' = Code_Thingol.fun_tyco `%% [ITyVar v', ty];
+ val value_name = "Value.value.value"
+ val program' = program
+ |> Graph.new_node (value_name,
+ Code_Thingol.Fun (Term.dummy_patternN, (((vs', ty'), [(([IVar NONE], t), (NONE, true))]), NONE)))
+ |> fold (curry Graph.add_edge value_name) deps;
+ val (program_code, [SOME value_name']) = serializer naming program' [value_name];
+ val value_code = space_implode " "
+ (value_name' :: "()" :: map (enclose "(" ")") args);
+ in Exn.capture (ML_Context.value ctxt cookie) (program_code, value_code) end;
+
+fun partiality_as_none e = SOME (Exn.release e)
+ handle General.Match => NONE
+ | General.Bind => NONE
+ | General.Fail _ => NONE;
+
+fun dynamic_value_exn cookie thy some_target postproc t args =
+ let
+ fun evaluator naming program ((_, vs_ty), t) deps =
+ base_evaluator cookie (obtain_serializer thy some_target) naming thy program (vs_ty, t) deps args;
+ in Code_Thingol.dynamic_eval_value thy (Exn.map_result o postproc) evaluator t end;
+
+fun dynamic_value_strict cookie thy some_target postproc t args =
+ Exn.release (dynamic_value_exn cookie thy some_target postproc t args);
+
+fun dynamic_value cookie thy some_target postproc t args =
+ partiality_as_none (dynamic_value_exn cookie thy some_target postproc t args);
+
+fun static_value_exn cookie thy some_target postproc consts =
+ let
+ val serializer = obtain_serializer thy some_target;
+ fun evaluator naming program thy ((_, vs_ty), t) deps =
+ base_evaluator cookie serializer naming thy program (vs_ty, t) deps [];
+ in Code_Thingol.static_eval_value thy (Exn.map_result o postproc) consts evaluator end;
+
+fun static_value_strict cookie thy some_target postproc consts t =
+ Exn.release (static_value_exn cookie thy some_target postproc consts t);
+
+fun static_value cookie thy some_target postproc consts t =
+ partiality_as_none (static_value_exn cookie thy some_target postproc consts t);
+
+
+(* evaluation for truth or nothing *)
+
+structure Truth_Result = Proof_Data (
+ type T = unit -> truth
+ fun init _ () = error "Truth_Result"
+);
+val put_truth = Truth_Result.put;
+val truth_cookie = (Truth_Result.get, put_truth, Long_Name.append this "put_truth");
+
+fun check_holds serializer naming thy program vs_t deps ct =
+ let
+ val t = Thm.term_of ct;
+ val _ = if fastype_of t <> propT
+ then error ("Not a proposition: " ^ Syntax.string_of_term_global thy t)
+ else ();
+ val iff = Thm.cterm_of thy (Term.Const ("==", propT --> propT --> propT));
+ val result = case partiality_as_none (base_evaluator truth_cookie serializer naming thy program vs_t deps [])
+ of SOME Holds => true
+ | _ => false;
+ in
+ Thm.mk_binop iff ct (if result then @{cprop "PROP Code_Generator.holds"} else ct)
+ end;
+
+val (_, raw_check_holds_oracle) = Context.>>> (Context.map_theory_result
+ (Thm.add_oracle (Binding.name "holds_by_evaluation",
+ fn (serializer, naming, thy, program, vs_t, deps, ct) => check_holds serializer naming thy program vs_t deps ct)));
+
+fun check_holds_oracle serializer naming thy program ((_, vs_ty), t) deps ct =
+ raw_check_holds_oracle (serializer, naming, thy, program, (vs_ty, t), deps, ct);
+
+val dynamic_holds_conv = Conv.tap_thy (fn thy => Code_Thingol.dynamic_eval_conv thy
+ (fn naming => check_holds_oracle (obtain_serializer thy NONE) naming thy));
+
+fun static_holds_conv thy consts =
+ let
+ val serializer = obtain_serializer thy NONE;
+ in
+ Code_Thingol.static_eval_conv thy consts
+ (fn naming => fn program => fn thy => check_holds_oracle serializer naming thy program)
+ end;
(** instrumentalization **)
@@ -155,7 +265,7 @@
fun add_eval_const (const, const') = Code_Target.add_const_syntax target
const (SOME (Code_Printer.simple_const_syntax (0, (K o K o K) const')));
-fun process (code_body, (tyco_map, (constr_map, const_map))) module_name NONE thy =
+fun process_reflection (code_body, (tyco_map, (constr_map, const_map))) module_name NONE thy =
thy
|> Code_Target.add_reserved target module_name
|> Context.theory_map
@@ -163,7 +273,7 @@
|> fold (add_eval_tyco o apsnd Code_Printer.str) tyco_map
|> fold (add_eval_constr o apsnd Code_Printer.str) constr_map
|> fold (add_eval_const o apsnd Code_Printer.str) const_map
- | process (code_body, _) _ (SOME file_name) thy =
+ | process_reflection (code_body, _) _ (SOME file_name) thy =
let
val preamble =
"(* Generated from " ^ Path.implode (Thy_Header.thy_path (Context.theory_name thy))
@@ -185,7 +295,7 @@
|> (apsnd o apsnd) (chop (length constrs));
in
thy
- |> process result module_name some_file
+ |> process_reflection result module_name some_file
end;
val code_reflect = gen_code_reflect Code_Target.cert_tyco (K I);
@@ -221,6 +331,4 @@
end; (*local*)
-val setup = Code_Target.extend_target (target, (Code_ML.target_SML, K I));
-
end; (*struct*)
--- a/src/Tools/Code/code_simp.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Code/code_simp.ML Fri Sep 17 17:31:20 2010 +0200
@@ -66,7 +66,7 @@
fun conclude_tac thy some_ss = Simplifier.full_simp_tac (simpset_default thy some_ss);
-(* evaluation with current code context *)
+(* evaluation with dynamic code context *)
fun dynamic_eval_conv thy = no_frees_conv (Code_Thingol.dynamic_eval_conv thy
(fn naming => fn program => fn t => fn deps => rewrite_modulo thy NONE program));
@@ -81,10 +81,11 @@
#> Value.add_evaluator ("simp", dynamic_eval_value o ProofContext.theory_of);
-(* evaluation with freezed code context *)
+(* evaluation with static code context *)
fun static_eval_conv thy some_ss consts = no_frees_conv
- (Code_Thingol.static_eval_conv_simple thy consts (rewrite_modulo thy some_ss));
+ (Code_Thingol.static_eval_conv_simple thy consts
+ (fn program => fn thy => fn _ => fn _ => rewrite_modulo thy some_ss program));
fun static_eval_tac thy some_ss consts = CONVERSION (static_eval_conv thy some_ss consts)
THEN' conclude_tac thy some_ss;
--- a/src/Tools/Code/code_target.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Code/code_target.ML Fri Sep 17 17:31:20 2010 +0200
@@ -241,7 +241,7 @@
local
-fun activate_target_for thy target naming program =
+fun activate_target thy target =
let
val ((targets, abortable), default_width) = Targets.get thy;
fun collapse_hierarchy target =
@@ -250,13 +250,13 @@
of SOME data => data
| NONE => error ("Unknown code target language: " ^ quote target);
in case the_description data
- of Fundamental _ => (I, data)
+ of Fundamental _ => (K I, data)
| Extension (super, modify) => let
val (modify', data') = collapse_hierarchy super
- in (modify' #> modify naming, merge_target false target (data', data)) end
+ in (fn naming => modify' naming #> modify naming, merge_target false target (data', data)) end
end;
val (modify, data) = collapse_hierarchy target;
- in (default_width, abortable, data, modify program) end;
+ in (default_width, abortable, data, modify) end;
fun activate_syntax lookup_name src_tab = Symtab.empty
|> fold_map (fn thing_identifier => fn tab => case lookup_name thing_identifier
@@ -303,7 +303,7 @@
val program4 = Graph.subgraph (member (op =) names4) program3;
in (names4, program4) end;
-fun invoke_serializer thy abortable serializer literals reserved abs_includes
+fun invoke_serializer thy abortable serializer literals reserved all_includes
module_alias proto_class_syntax proto_instance_syntax proto_tyco_syntax proto_const_syntax
module_name args naming proto_program names =
let
@@ -311,7 +311,12 @@
activate_symbol_syntax thy literals naming
proto_class_syntax proto_instance_syntax proto_tyco_syntax proto_const_syntax;
val (names_all, program) = project_program thy abortable names_hidden names proto_program;
- val includes = abs_includes names_all;
+ fun select_include (name, (content, cs)) =
+ if null cs orelse exists (fn c => case Code_Thingol.lookup_const naming c
+ of SOME name => member (op =) names_all name
+ | NONE => false) cs
+ then SOME (name, content) else NONE;
+ val includes = map_filter select_include (Symtab.dest all_includes);
in
serializer args {
labelled_name = Code_Thingol.labelled_name thy proto_program,
@@ -324,29 +329,20 @@
program = program }
end;
-fun mount_serializer thy target some_width module_name args naming proto_program names =
+fun mount_serializer thy target some_width module_name args =
let
- val (default_width, abortable, data, program) =
- activate_target_for thy target naming proto_program;
+ val (default_width, abortable, data, modify) = activate_target thy target;
val serializer = case the_description data
of Fundamental seri => #serializer seri;
val reserved = the_reserved data;
- fun select_include names_all (name, (content, cs)) =
- if null cs then SOME (name, content)
- else if exists (fn c => case Code_Thingol.lookup_const naming c
- of SOME name => member (op =) names_all name
- | NONE => false) cs
- then SOME (name, content) else NONE;
- fun includes names_all = map_filter (select_include names_all)
- ((Symtab.dest o the_includes) data);
val module_alias = the_module_alias data
val { class, instance, tyco, const } = the_symbol_syntax data;
val literals = the_literals thy target;
val width = the_default default_width some_width;
- in
+ in fn naming => fn program => fn names =>
invoke_serializer thy abortable serializer literals reserved
- includes module_alias class instance tyco const module_name args
- naming program names width
+ (the_includes data) module_alias class instance tyco const module_name args
+ naming (modify naming program) names width
end;
fun assert_module_name "" = error ("Empty module name not allowed.")
@@ -354,16 +350,22 @@
in
-fun export_code_for thy some_path target some_width some_module_name args naming program names =
- export some_path (mount_serializer thy target some_width some_module_name args naming program names);
+fun export_code_for thy some_path target some_width module_name args =
+ export some_path ooo mount_serializer thy target some_width module_name args;
-fun produce_code_for thy target some_width module_name args naming program names =
+fun produce_code_for thy target some_width module_name args =
let
- val (s, deresolve) = produce (mount_serializer thy target some_width (assert_module_name module_name) args naming program names)
- in (s, map deresolve names) end;
+ val serializer = mount_serializer thy target some_width (assert_module_name module_name) args;
+ in fn naming => fn program => fn names =>
+ produce (serializer naming program names) |> apsnd (fn deresolve => map deresolve names)
+ end;
-fun present_code_for thy target some_width module_name args naming program (names, selects) =
- present selects (mount_serializer thy target some_width (assert_module_name module_name) args naming program names);
+fun present_code_for thy target some_width module_name args =
+ let
+ val serializer = mount_serializer thy target some_width (assert_module_name module_name) args;
+ in fn naming => fn program => fn (names, selects) =>
+ present selects (serializer naming program names)
+ end;
fun check_code_for thy target strict args naming program names_cs =
let
@@ -441,6 +443,43 @@
fun check_code_cmd raw_cs seris thy = check_code thy (read_const_exprs thy raw_cs) seris;
+local
+
+val parse_const_terms = Scan.repeat1 Args.term
+ >> (fn ts => fn thy => map (Code.check_const thy) ts);
+
+fun parse_names category parse internalize lookup =
+ Scan.lift (Args.parens (Args.$$$ category)) |-- Scan.repeat1 parse
+ >> (fn xs => fn thy => fn naming => map_filter (lookup naming o internalize thy) xs);
+
+val parse_consts = parse_names "consts" Args.term
+ Code.check_const Code_Thingol.lookup_const ;
+
+val parse_types = parse_names "types" (Scan.lift Args.name)
+ Sign.intern_type Code_Thingol.lookup_tyco;
+
+val parse_classes = parse_names "classes" (Scan.lift Args.name)
+ Sign.intern_class Code_Thingol.lookup_class;
+
+val parse_instances = parse_names "instances" (Scan.lift (Args.name --| Args.$$$ "::" -- Args.name))
+ (fn thy => fn (raw_tyco, raw_class) => (Sign.intern_class thy raw_class, Sign.intern_type thy raw_tyco))
+ Code_Thingol.lookup_instance;
+
+in
+
+val _ = Thy_Output.antiquotation "code_stmts"
+ (parse_const_terms -- Scan.repeat (parse_consts || parse_types || parse_classes || parse_instances)
+ -- Scan.lift (Args.parens (Args.name -- Scan.option Parse.int)))
+ (fn {context = ctxt, ...} => fn ((mk_cs, mk_stmtss), (target, some_width)) =>
+ let val thy = ProofContext.theory_of ctxt in
+ present_code thy (mk_cs thy)
+ (fn naming => maps (fn f => f thy naming) mk_stmtss)
+ target some_width "Example" []
+ |> Latex.output_typewriter
+ end);
+
+end;
+
(** serializer configuration **)
--- a/src/Tools/Code/code_thingol.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Code/code_thingol.ML Fri Sep 17 17:31:20 2010 +0200
@@ -92,17 +92,21 @@
val read_const_exprs: theory -> string list -> string list * string list
val consts_program: theory -> bool -> string list -> string list * (naming * program)
- val dynamic_eval_conv: theory
- -> (naming -> program -> ((string * sort) list * typscheme) * iterm -> string list -> conv)
+ val dynamic_eval_conv: theory -> (naming -> program
+ -> ((string * sort) list * typscheme) * iterm -> string list -> conv)
-> conv
- val dynamic_eval_value: theory -> ((term -> term) -> 'a -> 'a)
- -> (naming -> program -> ((string * sort) list * typscheme) * iterm -> string list -> 'a)
+ val dynamic_eval_value: theory -> ((term -> term) -> 'a -> 'a) -> (naming -> program
+ -> ((string * sort) list * typscheme) * iterm -> string list -> 'a)
-> term -> 'a
- val static_eval_conv: theory -> string list
- -> (naming -> program -> ((string * sort) list * typscheme) * iterm -> string list -> conv)
+ val static_eval_conv: theory -> string list -> (naming -> program
+ -> theory -> ((string * sort) list * typscheme) * iterm -> string list -> conv)
-> conv
val static_eval_conv_simple: theory -> string list
- -> (program -> conv) -> conv
+ -> (program -> theory -> (string * sort) list -> term -> conv) -> conv
+ val static_eval_value: theory -> ((term -> term) -> 'a -> 'a) -> string list ->
+ (naming -> program
+ -> theory -> ((string * sort) list * typscheme) * iterm -> string list -> 'a)
+ -> term -> 'a
end;
structure Code_Thingol: CODE_THINGOL =
@@ -841,17 +845,17 @@
(* program generation *)
-fun consts_program thy permissive cs =
+fun consts_program thy permissive consts =
let
- fun project_consts cs (naming, program) = (*FIXME only necessary for cache_generation*)
- let
- val cs_all = Graph.all_succs program cs;
- in (cs, (naming, Graph.subgraph (member (op =) cs_all) program)) end;
+ fun project_consts consts (naming, program) =
+ if permissive then (consts, (naming, program))
+ else (consts, (naming, Graph.subgraph
+ (member (op =) (Graph.all_succs program consts)) program));
fun generate_consts thy algebra eqngr =
fold_map (ensure_const thy algebra eqngr permissive);
in
- invoke_generation (not permissive) thy (Code_Preproc.obtain false thy cs [])
- generate_consts cs
+ invoke_generation permissive thy (Code_Preproc.obtain false thy consts [])
+ generate_consts consts
|-> project_consts
end;
@@ -884,24 +888,46 @@
#> term_value
end;
-fun base_evaluator thy evaluator algebra eqngr vs t =
+fun original_sorts vs =
+ map (fn (v, _) => (v, (the o AList.lookup (op =) vs o prefix "'") v));
+
+fun dynamic_evaluator thy evaluator algebra eqngr vs t =
let
val (((naming, program), (((vs', ty'), t'), deps)), _) =
invoke_generation false thy (algebra, eqngr) ensure_value t;
- val vs'' = map (fn (v, _) => (v, (the o AList.lookup (op =) vs o prefix "'") v)) vs';
- in evaluator naming program ((vs'', (vs', ty')), t') deps end;
+ in evaluator naming program ((original_sorts vs vs', (vs', ty')), t') deps end;
+
+fun dynamic_eval_conv thy evaluator =
+ Code_Preproc.dynamic_eval_conv thy (dynamic_evaluator thy evaluator);
+
+fun dynamic_eval_value thy postproc evaluator =
+ Code_Preproc.dynamic_eval_value thy postproc (dynamic_evaluator thy evaluator);
-fun dynamic_eval_conv thy = Code_Preproc.dynamic_eval_conv thy o base_evaluator thy;
-fun dynamic_eval_value thy postproc = Code_Preproc.dynamic_eval_value thy postproc o base_evaluator thy;
+fun provide_program thy consts f algebra eqngr =
+ let
+ fun generate_consts thy algebra eqngr =
+ fold_map (ensure_const thy algebra eqngr false);
+ val (_, (naming, program)) = invoke_generation true thy (algebra, eqngr)
+ generate_consts consts;
+ in f algebra eqngr naming program end;
+
+fun static_evaluator evaluator algebra eqngr naming program thy vs t =
+ let
+ val (((_, program'), (((vs', ty'), t'), deps)), _) =
+ ensure_value thy algebra eqngr t (NONE, (naming, program));
+ in evaluator naming program' thy ((original_sorts vs vs', (vs', ty')), t') deps end;
fun static_eval_conv thy consts conv =
- Code_Preproc.static_eval_conv thy consts (base_evaluator thy conv); (*FIXME avoid re-generation*)
+ Code_Preproc.static_eval_conv thy consts
+ (provide_program thy consts (static_evaluator conv));
fun static_eval_conv_simple thy consts conv =
- Code_Preproc.static_eval_conv thy consts (fn algebra => fn eqngr => fn _ => fn _ => fn ct =>
- conv ((NONE, (empty_naming, Graph.empty)) (*FIXME provide abstraction for this kind of invocation*)
- |> fold_map (ensure_const thy algebra eqngr false) consts
- |> (snd o snd o snd)) ct);
+ Code_Preproc.static_eval_conv thy consts
+ (provide_program thy consts ((K o K o K) conv));
+
+fun static_eval_value thy postproc consts evaluator =
+ Code_Preproc.static_eval_value thy postproc consts
+ (provide_program thy consts (static_evaluator evaluator));
(** diagnostic commands **)
--- a/src/Tools/Code_Generator.thy Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Code_Generator.thy Fri Sep 17 17:31:20 2010 +0200
@@ -21,8 +21,8 @@
"~~/src/Tools/Code/code_ml.ML"
"~~/src/Tools/Code/code_haskell.ML"
"~~/src/Tools/Code/code_scala.ML"
- "~~/src/Tools/Code/code_runtime.ML"
"~~/src/Tools/nbe.ML"
+ ("~~/src/Tools/Code/code_runtime.ML")
begin
setup {*
@@ -32,7 +32,6 @@
#> Code_ML.setup
#> Code_Haskell.setup
#> Code_Scala.setup
- #> Code_Runtime.setup
#> Nbe.setup
#> Quickcheck.setup
*}
@@ -64,6 +63,8 @@
by rule (rule holds)+
qed
+use "~~/src/Tools/Code/code_runtime.ML"
+
hide_const (open) holds
end
--- a/src/Tools/Metis/Makefile Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/Makefile Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
###############################################################################
# METIS MAKEFILE
-# Copyright (c) 2001 Joe Hurd, distributed under the GNU GPL version 2
+# Copyright (c) 2001 Joe Hurd, distributed under the BSD License
###############################################################################
.SUFFIXES:
--- a/src/Tools/Metis/Makefile.FILES Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/Makefile.FILES Fri Sep 17 17:31:20 2010 +0200
@@ -4,7 +4,7 @@
echo > $@
refresh_FILES:
echo $(POLYML_SRC) | \
- sed "s/src\///g" | \
- sed "s/ Tptp\.s[a-z][a-z]//g" | \
- sed "s/ Options\.s[a-z][a-z]//g" \
+ sed "s/src\/PortablePolyml/PortableIsabelle/g" | \
+ sed "s/ src\/Tptp\.s[a-z][a-z]//g" | \
+ sed "s/ src\/Options\.s[a-z][a-z]//g" \
> FILES
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Metis/PortableIsabelle.sml Fri Sep 17 17:31:20 2010 +0200
@@ -0,0 +1,26 @@
+(* Title: Tools/Metis/PortableIsabelle.sml
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2010
+
+Isabelle-specific setup for Metis. Based on "src/PortablePolyml.sml".
+*)
+
+structure Portable :> Portable =
+struct
+
+val ml = "isabelle"
+
+fun pointerEqual (x : 'a, y : 'a) = pointer_eq (x, y)
+
+fun critical e () = NAMED_CRITICAL "metis" e
+
+val randomWord = Random.nextWord
+val randomBool = Random.nextBool
+val randomInt = Random.nextInt
+val randomReal = Random.nextReal
+
+fun time f x = f x (* dummy *)
+
+end
+
+datatype 'a frag = QUOTE of string | ANTIQUOTE of 'a
--- a/src/Tools/Metis/README Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/README Fri Sep 17 17:31:20 2010 +0200
@@ -5,57 +5,43 @@
unfortunately somewhat involved and frustrating, and hopefully
temporary.
- 1. The file "Makefile" and the directory "src/" and "script/" were
- initially copied from Joe Hurd's Metis package. (His "script/"
- directory has many files in it, but we only need "mlpp".) The
- package that was used when these notes where written was an
- unnumbered version of Metis, more recent than 2.2 and dated 19
- July 2010.
+ 1. The files "Makefile" and "script/mlpp" and the directory "src/"
+ were initially copied from Joe Hurd's official Metis package. The
+ package that was used when these notes where written was Metis 2.3
+ (16 Sept. 2010).
2. The license in each source file will probably not be something we
- can use in Isabelle. Lawrence C. Paulson's command
-
- perl -p -i~ -w -e 's/MIT license/BSD License/g' *sig *sml
-
- run in the "src/" directory should do the trick. In a 13 Sept.
- 2010 email to Gerwin Klein, Joe Hurd, the sole copyright holder of
- Metis, wrote:
+ can use in Isabelle. The "fix_metis_license" script can be run to
+ replace all occurrences of "MIT license" with "BSD License". In a
+ 13 Sept. 2010 email to Gerwin Klein, Joe Hurd, the sole copyright
+ holder of Metis, wrote:
I hereby give permission to the Isabelle team to release Metis
as part of Isabelle, with the Metis code covered under the
Isabelle BSD license.
- 3. Some modifications have to be done manually to the source files.
- Some of these are necessary so that the sources compile at all
- with Isabelle, some are necessary to avoid race conditions in a
- multithreaded environment, and some affect the defaults of Metis's
- heuristics and are needed for backward compatibility with older
- Isabelle proofs and for performance reasons. These changes should
- be identified by a comment
+ 3. Some modifications might have to be done manually to the source
+ files. The ultimate way to track them down is to use Mercurial.
+ The command
- (* MODIFIED by ?Joe ?Blow *)
+ hg diff -rcffceed8e7fa: src
- but the ultimate way to track them down is to use Mercurial. The
- command
-
- hg diff -r2d0a4361c3ef: src
-
- should do the trick. (You might need to specify a different
+ should do the trick. You might need to specify a different
revision number if somebody updated the Metis sources without
- updating these instructions.)
+ updating these instructions; consult the history in case of doubt.
- 4. Isabelle itself only cares about "metis.ML", which is generated
+ 4. Isabelle itself cares only about "metis.ML", which is generated
from the files in "src/" by the script "make_metis". The script
- relies on "Makefile", "src/", and "scripts/", as well as a special
- file "Makefile.FILES" used to determine all the files needed to be
- included in "metis.ML".
+ relies on "Makefile", "scripts/mlpp", and "src/", as well as
+ the pseudo-makefile "Makefile.FILES" used to determine all the
+ files needed to be included in "metis.ML".
5. The output of "make_metis" should then work as is with Isabelle,
- but there are of course no guarantees. The script "make_metis" has
- a few evil regex hacks that could go wrong. It also produces a
- few temporary files ("FILES", "Makefile.dev", and
- "bin/mosml/Makefile.src") as side-effects; you can safely ignore
- them or delete them.
+ but there are of course no guarantees. The script "make_metis" and
+ the pseudo-makefile "Makefile.FILES" have a few regexes that could
+ go wrong. They also produce a few temporary files ("FILES",
+ "Makefile.dev", and "bin/mosml/Makefile.src") as side-effects; you
+ can safely ignore them or delete them.
6. Once you have successfully regenerated "metis.ML", build all of
Isabelle and keep an eye on the time taken by Metis.
@@ -69,4 +55,4 @@
Good luck!
Jasmin Blanchette
- 15 Sept. 2010
+ 17 Sept. 2010
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Metis/fix_metis_license Fri Sep 17 17:31:20 2010 +0200
@@ -0,0 +1,4 @@
+#!/usr/bin/env bash
+THIS=$(cd "$(dirname "$0")"; echo $PWD)
+(cd $THIS;
+ perl -p -i~ -w -e 's/MIT license/BSD License/g' Makefile src/*.s* scripts/mlpp)
--- a/src/Tools/Metis/make_metis Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/make_metis Fri Sep 17 17:31:20 2010 +0200
@@ -7,8 +7,8 @@
# compile within Isabelle on Poly/ML and SML/NJ.
THIS=$(cd "$(dirname "$0")"; echo $PWD)
-
make -f Makefile.FILES refresh_FILES
+FILES=$(cat "$THIS/FILES")
(
cat <<EOF
@@ -32,19 +32,17 @@
EOF
- for FILE in $(cat "$THIS/FILES")
+ for FILE in $FILES
do
echo
echo "(**** Original file: $FILE ****)"
echo
echo -e "$FILE" >&2
- "$THIS/scripts/mlpp" -c polyml "src/$FILE" | \
+ "$THIS/scripts/mlpp" -c polyml "$FILE" | \
perl -p -e \
's/type name$/type name = string/;'\
's/\bref\b/Unsynchronized.ref/g;'\
-'s/\bPolyML.pointerEq\b/pointer_eq/g;'\
-'s/\bRL\b/Metis_RL/g;'\
-"`grep "^\(signature\|structure\|functor\)" src/*.{sig,sml} | \
+"`grep "^\(signature\|structure\|functor\)" $FILES | \
sed "s/[^:]*:[a-z]* \([A-Za-z0-9_]*\).*/\1/" | \
tr " " "\n" | \
sort | \
--- a/src/Tools/Metis/metis.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/metis.ML Fri Sep 17 17:31:20 2010 +0200
@@ -17,7 +17,7 @@
print_depth 0;
-(**** Original file: Random.sig ****)
+(**** Original file: src/Random.sig ****)
(* Title: Tools/random_word.ML
Author: Makarius
@@ -29,6 +29,7 @@
signature Metis_Random =
sig
+
val nextWord : unit -> word
val nextBool : unit -> bool
@@ -39,7 +40,7 @@
end;
-(**** Original file: Random.sml ****)
+(**** Original file: src/Random.sml ****)
(* Title: Tools/random_word.ML
Author: Makarius
@@ -90,11 +91,11 @@
end;
-(**** Original file: Portable.sig ****)
-
-(* ========================================================================= *)
-(* ML SPECIFIC FUNCTIONS *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *)
+(**** Original file: src/Portable.sig ****)
+
+(* ========================================================================= *)
+(* ML COMPILER SPECIFIC FUNCTIONS *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Portable =
@@ -113,17 +114,10 @@
val pointerEqual : 'a * 'a -> bool
(* ------------------------------------------------------------------------- *)
-(* Timing function applications. *)
-(* ------------------------------------------------------------------------- *)
-
-val time : ('a -> 'b) -> 'a -> 'b
-
-(* ------------------------------------------------------------------------- *)
-(* Critical section markup (multiprocessing) *)
-(* ------------------------------------------------------------------------- *)
-
-(* MODIFIED by Jasmin Blanchette *)
-val CRITICAL: (unit -> 'a) -> 'a
+(* Marking critical sections of code. *)
+(* ------------------------------------------------------------------------- *)
+
+val critical : (unit -> 'a) -> unit -> 'a
(* ------------------------------------------------------------------------- *)
(* Generating random values. *)
@@ -137,97 +131,44 @@
val randomWord : unit -> Word.word
-end
-
-(**** Original file: PortablePolyml.sml ****)
-
-(* ========================================================================= *)
-(* POLY/ML SPECIFIC FUNCTIONS *)
-(* Copyright (c) 2008 Joe Hurd, distributed under the BSD License *)
-(* ========================================================================= *)
+(* ------------------------------------------------------------------------- *)
+(* Timing function applications. *)
+(* ------------------------------------------------------------------------- *)
+
+val time : ('a -> 'b) -> 'a -> 'b
+
+end
+
+(**** Original file: PortableIsabelle.sml ****)
+
+(* Title: Tools/Metis/PortableIsabelle.sml
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2010
+
+Isabelle-specific setup for Metis. Based on "src/PortablePolyml.sml".
+*)
structure Metis_Portable :> Metis_Portable =
struct
-(* ------------------------------------------------------------------------- *)
-(* The ML implementation. *)
-(* ------------------------------------------------------------------------- *)
-
-val ml = "polyml";
-
-(* ------------------------------------------------------------------------- *)
-(* Pointer equality using the run-time system. *)
-(* ------------------------------------------------------------------------- *)
-
-fun pointerEqual (x : 'a, y : 'a) = pointer_eq(x,y);
-
-(* ------------------------------------------------------------------------- *)
-(* Timing function applications. *)
-(* ------------------------------------------------------------------------- *)
-
-fun time f x =
- let
- fun p t =
- let
- val s = Time.fmt 3 t
- in
- case size (List.last (String.fields (fn x => x = #".") s)) of
- 3 => s
- | 2 => s ^ "0"
- | 1 => s ^ "00"
- | _ => raise Fail "Metis_Portable.time"
- end
-
- val c = Timer.startCPUTimer ()
-
- val r = Timer.startRealTimer ()
-
- fun pt () =
- let
- val {usr,sys} = Timer.checkCPUTimer c
- val real = Timer.checkRealTimer r
- in
- TextIO.print (* MODIFIED by Jasmin Blanchette *)
- ("User: " ^ p usr ^ " System: " ^ p sys ^
- " Real: " ^ p real ^ "\n")
- end
-
- val y = f x handle e => (pt (); raise e)
-
- val () = pt ()
- in
- y
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Critical section markup (multiprocessing) *)
-(* ------------------------------------------------------------------------- *)
-
-(* MODIFIED by Jasmin Blanchette *)
-fun CRITICAL e = NAMED_CRITICAL "metis" e;
-
-
-(* ------------------------------------------------------------------------- *)
-(* Generating random values. *)
-(* ------------------------------------------------------------------------- *)
-
-val randomWord = Metis_Random.nextWord;
-
-val randomBool = Metis_Random.nextBool;
-
-val randomInt = Metis_Random.nextInt;
-
-val randomReal = Metis_Random.nextReal;
-
-end
-
-(* ------------------------------------------------------------------------- *)
-(* Quotations a la Moscow ML. *)
-(* ------------------------------------------------------------------------- *)
-
-datatype 'a frag = QUOTE of string | ANTIQUOTE of 'a;
-
-(**** Original file: Useful.sig ****)
+val ml = "isabelle"
+
+fun pointerEqual (x : 'a, y : 'a) = pointer_eq (x, y)
+
+fun critical e () = NAMED_CRITICAL "metis" e
+
+val randomWord = Metis_Random.nextWord
+val randomBool = Metis_Random.nextBool
+val randomInt = Metis_Random.nextInt
+val randomReal = Metis_Random.nextReal
+
+fun time f x = f x (* dummy *)
+
+end
+
+datatype 'a frag = QUOTE of string | ANTIQUOTE of 'a
+
+(**** Original file: src/Useful.sig ****)
(* ========================================================================= *)
(* ML UTILITY FUNCTIONS *)
@@ -253,8 +194,6 @@
(* Tracing. *)
(* ------------------------------------------------------------------------- *)
-val print : string -> unit (* MODIFIED by Jasmin Blanchette *)
-
val tracePrint : (string -> unit) Unsynchronized.ref
val trace : string -> unit
@@ -340,10 +279,6 @@
(* Lists: note we count elements from 0. *)
(* ------------------------------------------------------------------------- *)
-val foldl : ('a * 'b -> 'b) -> 'b -> 'a list -> 'b (* MODIFIED by Jasmin Blanchette *)
-
-val foldr : ('a * 'b -> 'b) -> 'b -> 'a list -> 'b (* MODIFIED by Jasmin Blanchette *)
-
val cons : 'a -> 'a list -> 'a list
val hdTl : 'a list -> 'a * 'a list
@@ -448,10 +383,6 @@
(* Strings. *)
(* ------------------------------------------------------------------------- *)
-val implode : char list -> string (* MODIFIED by Jasmin Blanchette *)
-
-val explode : string -> char list (* MODIFIED by Jasmin Blanchette *)
-
val rot : int -> char -> char
val charToInt : char -> int option
@@ -554,7 +485,9 @@
val try : ('a -> 'b) -> 'a -> 'b
-val chat : string -> unit
+val chat : string -> unit (* stdout *)
+
+val chide : string -> unit (* stderr *)
val warn : string -> unit
@@ -568,7 +501,7 @@
end
-(**** Original file: Useful.sml ****)
+(**** Original file: src/Useful.sml ****)
(* ========================================================================= *)
(* ML UTILITY FUNCTIONS *)
@@ -622,9 +555,7 @@
(* Tracing. *)
(* ------------------------------------------------------------------------- *)
-val print = TextIO.print; (* MODIFIED by Jasmin Blanchette *)
-
-val tracePrint = Unsynchronized.ref print;
+val tracePrint = Unsynchronized.ref TextIO.print;
fun trace mesg = !tracePrint mesg;
@@ -744,10 +675,6 @@
(* Lists. *)
(* ------------------------------------------------------------------------- *)
-val foldl = List.foldl; (* MODIFIED by Jasmin Blanchette *)
-
-val foldr = List.foldr; (* MODIFIED by Jasmin Blanchette *)
-
fun cons x y = x :: y;
fun hdTl l = (hd l, tl l);
@@ -783,19 +710,22 @@
fun zip xs ys = zipWith pair xs ys;
-fun unzip ab =
- foldl (fn ((x, y), (xs, ys)) => (x :: xs, y :: ys)) ([], []) (rev ab);
+local
+ fun inc ((x,y),(xs,ys)) = (x :: xs, y :: ys);
+in
+ fun unzip ab = List.foldl inc ([],[]) (rev ab);
+end;
fun cartwith f =
- let
- fun aux _ res _ [] = res
- | aux xsCopy res [] (y :: yt) = aux xsCopy res xsCopy yt
- | aux xsCopy res (x :: xt) (ys as y :: _) =
- aux xsCopy (f x y :: res) xt ys
- in
- fn xs => fn ys =>
- let val xs' = rev xs in aux xs' [] xs' (rev ys) end
- end;
+ let
+ fun aux _ res _ [] = res
+ | aux xsCopy res [] (y :: yt) = aux xsCopy res xsCopy yt
+ | aux xsCopy res (x :: xt) (ys as y :: _) =
+ aux xsCopy (f x y :: res) xt ys
+ in
+ fn xs => fn ys =>
+ let val xs' = rev xs in aux xs' [] xs' (rev ys) end
+ end;
fun cart xs ys = cartwith pair xs ys;
@@ -914,15 +844,32 @@
fun delete x s = List.filter (not o equal x) s;
-fun setify s = rev (foldl (fn (v,x) => if mem v x then x else v :: x) [] s);
-
-fun union s t = foldl (fn (v,x) => if mem v t then x else v :: x) t (rev s);
+local
+ fun inc (v,x) = if mem v x then x else v :: x;
+in
+ fun setify s = rev (List.foldl inc [] s);
+end;
+
+fun union s t =
+ let
+ fun inc (v,x) = if mem v t then x else v :: x
+ in
+ List.foldl inc t (rev s)
+ end;
fun intersect s t =
- foldl (fn (v,x) => if mem v t then v :: x else x) [] (rev s);
+ let
+ fun inc (v,x) = if mem v t then v :: x else x
+ in
+ List.foldl inc [] (rev s)
+ end;
fun difference s t =
- foldl (fn (v,x) => if mem v t then x else v :: x) [] (rev s);
+ let
+ fun inc (v,x) = if mem v t then x else v :: x
+ in
+ List.foldl inc [] (rev s)
+ end;
fun subset s t = List.all (fn x => mem x t) s;
@@ -1020,66 +967,43 @@
end;
local
- fun calcPrimes ps n i =
- if List.exists (fn p => divides p i) ps then calcPrimes ps n (i + 1)
+ fun calcPrimes mode ps i n =
+ if n = 0 then ps
+ else if List.exists (fn p => divides p i) ps then
+ let
+ val i = i + 1
+ and n = if mode then n else n - 1
+ in
+ calcPrimes mode ps i n
+ end
else
let
val ps = ps @ [i]
+ and i = i + 1
and n = n - 1
in
- if n = 0 then ps else calcPrimes ps n (i + 1)
+ calcPrimes mode ps i n
end;
-
- val primesList = Unsynchronized.ref [2];
-in
- (* MODIFIED by Jasmin Blanchette *)
- fun primes n = CRITICAL (fn () =>
- let
- val Unsynchronized.ref ps = primesList
-
- val k = n - length ps
- in
- if k <= 0 then List.take (ps,n)
- else
- let
- val ps = calcPrimes ps k (List.last ps + 1)
-
- val () = primesList := ps
- in
- ps
- end
- end);
-end;
-
-(* MODIFIED by Jasmin Blanchette *)
-fun primesUpTo n = CRITICAL (fn () =>
- let
- fun f k =
- let
- val l = primes k
-
- val p = List.last l
- in
- if p < n then f (2 * k) else takeWhile (fn j => j <= n) l
- end
- in
- f 8
- end);
+in
+ fun primes n =
+ if n <= 0 then []
+ else calcPrimes true [2] 3 (n - 1);
+
+ fun primesUpTo n =
+ if n < 2 then []
+ else calcPrimes false [2] 3 (n - 2);
+end;
(* ------------------------------------------------------------------------- *)
(* Strings. *)
(* ------------------------------------------------------------------------- *)
-val implode = String.implode (* MODIFIED by Jasmin Blanchette *)
-
-val explode = String.explode (* MODIFIED by Jasmin Blanchette *)
-
local
fun len l = (length l, l)
- val upper = len (explode "ABCDEFGHIJKLMNOPQRSTUVWXYZ");
-
- val lower = len (explode "abcdefghijklmnopqrstuvwxyz");
+ val upper = len (String.explode "ABCDEFGHIJKLMNOPQRSTUVWXYZ");
+
+ val lower = len (String.explode "abcdefghijklmnopqrstuvwxyz");
fun rotate (n,l) c k =
List.nth (l, (k + Option.valOf (index (equal c) l)) mod n);
@@ -1118,7 +1042,7 @@
let
fun dup 0 l = l | dup n l = dup (n - 1) (x :: l)
in
- fn n => implode (dup n [])
+ fn n => String.implode (dup n [])
end;
fun chomp s =
@@ -1130,14 +1054,15 @@
end;
local
- fun chop [] = []
- | chop (l as (h :: t)) = if Char.isSpace h then chop t else l;
-in
- val trim = implode o chop o rev o chop o rev o explode;
-end;
-
-fun join _ [] = ""
- | join s (h :: t) = foldl (fn (x,y) => y ^ s ^ x) h t;
+ fun chop l =
+ case l of
+ [] => []
+ | h :: t => if Char.isSpace h then chop t else l;
+in
+ val trim = String.implode o chop o rev o chop o rev o String.explode;
+end;
+
+val join = String.concatWith;
local
fun match [] l = SOME l
@@ -1145,18 +1070,19 @@
| match (x :: xs) (y :: ys) = if x = y then match xs ys else NONE;
fun stringify acc [] = acc
- | stringify acc (h :: t) = stringify (implode h :: acc) t;
+ | stringify acc (h :: t) = stringify (String.implode h :: acc) t;
in
fun split sep =
let
val pat = String.explode sep
+
fun div1 prev recent [] = stringify [] (rev recent :: prev)
| div1 prev recent (l as h :: t) =
case match pat l of
NONE => div1 prev (h :: recent) t
| SOME rest => div1 (rev recent :: prev) [] rest
in
- fn s => div1 [] [] (explode s)
+ fn s => div1 [] [] (String.explode s)
end;
end;
@@ -1285,25 +1211,30 @@
local
val generator = Unsynchronized.ref 0
-in
- (* MODIFIED by Jasmin Blanchette *)
- fun newInt () = CRITICAL (fn () =>
+
+ fun newIntThunk () =
let
val n = !generator
+
val () = generator := n + 1
in
n
- end);
-
- fun newInts 0 = []
- (* MODIFIED by Jasmin Blanchette *)
- | newInts k = CRITICAL (fn () =>
+ end;
+
+ fun newIntsThunk k () =
let
val n = !generator
+
val () = generator := n + k
in
interval n k
- end);
+ end;
+in
+ fun newInt () = Metis_Portable.critical newIntThunk ();
+
+ fun newInts k =
+ if k <= 0 then []
+ else Metis_Portable.critical (newIntsThunk k) ();
end;
fun withRef (r,new) f x =
@@ -1383,10 +1314,12 @@
(* Profiling and error reporting. *)
(* ------------------------------------------------------------------------- *)
-fun chat s = TextIO.output (TextIO.stdErr, s ^ "\n");
-
-local
- fun err x s = chat (x ^ ": " ^ s);
+fun chat s = TextIO.output (TextIO.stdOut, s ^ "\n");
+
+fun chide s = TextIO.output (TextIO.stdErr, s ^ "\n");
+
+local
+ fun err x s = chide (x ^ ": " ^ s);
in
fun try f x = f x
handle e as Error _ => (err "try" (errorToString e); raise e)
@@ -1431,7 +1364,7 @@
end
-(**** Original file: Lazy.sig ****)
+(**** Original file: src/Lazy.sig ****)
(* ========================================================================= *)
(* SUPPORT FOR LAZY EVALUATION *)
@@ -1453,7 +1386,7 @@
end
-(**** Original file: Lazy.sml ****)
+(**** Original file: src/Lazy.sml ****)
(* ========================================================================= *)
(* SUPPORT FOR LAZY EVALUATION *)
@@ -1494,11 +1427,11 @@
end
-(**** Original file: Ordered.sig ****)
+(**** Original file: src/Ordered.sig ****)
(* ========================================================================= *)
(* ORDERED TYPES *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Ordered =
@@ -1530,11 +1463,11 @@
end
-(**** Original file: Ordered.sml ****)
+(**** Original file: src/Ordered.sml ****)
(* ========================================================================= *)
(* ORDERED TYPES *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_IntOrdered =
@@ -1556,7 +1489,7 @@
structure Metis_StringOrdered =
struct type t = string val compare = String.compare end;
-(**** Original file: Map.sig ****)
+(**** Original file: src/Map.sig ****)
(* ========================================================================= *)
(* FINITE MAPS IMPLEMENTED WITH RANDOMLY BALANCED TREES *)
@@ -1749,7 +1682,7 @@
end
-(**** Original file: Map.sml ****)
+(**** Original file: src/Map.sml ****)
(* ========================================================================= *)
(* FINITE MAPS IMPLEMENTED WITH RANDOMLY BALANCED TREES *)
@@ -2632,38 +2565,40 @@
(* ------------------------------------------------------------------------- *)
datatype ('key,'value) iterator =
- LR of ('key * 'value) * ('key,'value) tree * ('key,'value) node list
- | Metis_RL of ('key * 'value) * ('key,'value) tree * ('key,'value) node list;
-
-fun fromSpineLR nodes =
+ LeftToRightIterator of
+ ('key * 'value) * ('key,'value) tree * ('key,'value) node list
+ | RightToLeftIterator of
+ ('key * 'value) * ('key,'value) tree * ('key,'value) node list;
+
+fun fromSpineLeftToRightIterator nodes =
case nodes of
[] => NONE
| Node {key,value,right,...} :: nodes =>
- SOME (LR ((key,value),right,nodes));
-
-fun fromSpineRL nodes =
+ SOME (LeftToRightIterator ((key,value),right,nodes));
+
+fun fromSpineRightToLeftIterator nodes =
case nodes of
[] => NONE
| Node {key,value,left,...} :: nodes =>
- SOME (Metis_RL ((key,value),left,nodes));
-
-fun addLR nodes tree = fromSpineLR (treeLeftSpine nodes tree);
-
-fun addRL nodes tree = fromSpineRL (treeRightSpine nodes tree);
-
-fun treeMkIterator tree = addLR [] tree;
-
-fun treeMkRevIterator tree = addRL [] tree;
+ SOME (RightToLeftIterator ((key,value),left,nodes));
+
+fun addLeftToRightIterator nodes tree = fromSpineLeftToRightIterator (treeLeftSpine nodes tree);
+
+fun addRightToLeftIterator nodes tree = fromSpineRightToLeftIterator (treeRightSpine nodes tree);
+
+fun treeMkIterator tree = addLeftToRightIterator [] tree;
+
+fun treeMkRevIterator tree = addRightToLeftIterator [] tree;
fun readIterator iter =
case iter of
- LR (key_value,_,_) => key_value
- | Metis_RL (key_value,_,_) => key_value;
+ LeftToRightIterator (key_value,_,_) => key_value
+ | RightToLeftIterator (key_value,_,_) => key_value;
fun advanceIterator iter =
case iter of
- LR (_,tree,nodes) => addLR nodes tree
- | Metis_RL (_,tree,nodes) => addRL nodes tree;
+ LeftToRightIterator (_,tree,nodes) => addLeftToRightIterator nodes tree
+ | RightToLeftIterator (_,tree,nodes) => addRightToLeftIterator nodes tree;
fun foldIterator f acc io =
case io of
@@ -3177,7 +3112,7 @@
end
-(**** Original file: KeyMap.sig ****)
+(**** Original file: src/KeyMap.sig ****)
(* ========================================================================= *)
(* FINITE MAPS WITH A FIXED KEY TYPE *)
@@ -3376,7 +3311,7 @@
end
-(**** Original file: KeyMap.sml ****)
+(**** Original file: src/KeyMap.sml ****)
(* ========================================================================= *)
(* FINITE MAPS WITH A FIXED KEY TYPE *)
@@ -4267,38 +4202,40 @@
(* ------------------------------------------------------------------------- *)
datatype 'value iterator =
- LR of (key * 'value) * 'value tree * 'value node list
- | Metis_RL of (key * 'value) * 'value tree * 'value node list;
-
-fun fromSpineLR nodes =
+ LeftToRightIterator of
+ (key * 'value) * 'value tree * 'value node list
+ | RightToLeftIterator of
+ (key * 'value) * 'value tree * 'value node list;
+
+fun fromSpineLeftToRightIterator nodes =
case nodes of
[] => NONE
| Node {key,value,right,...} :: nodes =>
- SOME (LR ((key,value),right,nodes));
-
-fun fromSpineRL nodes =
+ SOME (LeftToRightIterator ((key,value),right,nodes));
+
+fun fromSpineRightToLeftIterator nodes =
case nodes of
[] => NONE
| Node {key,value,left,...} :: nodes =>
- SOME (Metis_RL ((key,value),left,nodes));
-
-fun addLR nodes tree = fromSpineLR (treeLeftSpine nodes tree);
-
-fun addRL nodes tree = fromSpineRL (treeRightSpine nodes tree);
-
-fun treeMkIterator tree = addLR [] tree;
-
-fun treeMkRevIterator tree = addRL [] tree;
+ SOME (RightToLeftIterator ((key,value),left,nodes));
+
+fun addLeftToRightIterator nodes tree = fromSpineLeftToRightIterator (treeLeftSpine nodes tree);
+
+fun addRightToLeftIterator nodes tree = fromSpineRightToLeftIterator (treeRightSpine nodes tree);
+
+fun treeMkIterator tree = addLeftToRightIterator [] tree;
+
+fun treeMkRevIterator tree = addRightToLeftIterator [] tree;
fun readIterator iter =
case iter of
- LR (key_value,_,_) => key_value
- | Metis_RL (key_value,_,_) => key_value;
+ LeftToRightIterator (key_value,_,_) => key_value
+ | RightToLeftIterator (key_value,_,_) => key_value;
fun advanceIterator iter =
case iter of
- LR (_,tree,nodes) => addLR nodes tree
- | Metis_RL (_,tree,nodes) => addRL nodes tree;
+ LeftToRightIterator (_,tree,nodes) => addLeftToRightIterator nodes tree
+ | RightToLeftIterator (_,tree,nodes) => addRightToLeftIterator nodes tree;
fun foldIterator f acc io =
case io of
@@ -4821,7 +4758,7 @@
structure Metis_StringMap =
Metis_KeyMap (Metis_StringOrdered);
-(**** Original file: Set.sig ****)
+(**** Original file: src/Set.sig ****)
(* ========================================================================= *)
(* FINITE SETS IMPLEMENTED WITH RANDOMLY BALANCED TREES *)
@@ -4993,7 +4930,7 @@
end
-(**** Original file: Set.sml ****)
+(**** Original file: src/Set.sml ****)
(* ========================================================================= *)
(* FINITE SETS IMPLEMENTED WITH RANDOMLY BALANCED TREES *)
@@ -5327,7 +5264,7 @@
end
-(**** Original file: ElementSet.sig ****)
+(**** Original file: src/ElementSet.sig ****)
(* ========================================================================= *)
(* FINITE SETS WITH A FIXED ELEMENT TYPE *)
@@ -5505,7 +5442,7 @@
end
-(**** Original file: ElementSet.sml ****)
+(**** Original file: src/ElementSet.sml ****)
(* ========================================================================= *)
(* FINITE SETS WITH A FIXED ELEMENT TYPE *)
@@ -5855,11 +5792,11 @@
structure Metis_StringSet =
Metis_ElementSet (Metis_StringMap);
-(**** Original file: Sharing.sig ****)
+(**** Original file: src/Sharing.sig ****)
(* ========================================================================= *)
(* PRESERVING SHARING OF ML VALUES *)
-(* Copyright (c) 2005-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2005 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Sharing =
@@ -5903,11 +5840,11 @@
end
-(**** Original file: Sharing.sml ****)
+(**** Original file: src/Sharing.sml ****)
(* ========================================================================= *)
(* PRESERVING SHARING OF ML VALUES *)
-(* Copyright (c) 2005-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2005 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Sharing :> Metis_Sharing =
@@ -6064,11 +6001,11 @@
end
-(**** Original file: Stream.sig ****)
+(**** Original file: src/Stream.sig ****)
(* ========================================================================= *)
(* A POSSIBLY-INFINITE STREAM DATATYPE FOR ML *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Stream =
@@ -6176,18 +6113,16 @@
end
-(**** Original file: Stream.sml ****)
+(**** Original file: src/Stream.sml ****)
(* ========================================================================= *)
(* A POSSIBLY-INFINITE STREAM DATATYPE FOR ML *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Stream :> Metis_Stream =
struct
-open Metis_Useful; (* MODIFIED by Jasmin Blanchette *)
-
val K = Metis_Useful.K;
val pair = Metis_Useful.pair;
@@ -6381,9 +6316,9 @@
fun listConcat s = concat (map fromList s);
-fun toString s = implode (toList s);
-
-fun fromString s = fromList (explode s);
+fun toString s = String.implode (toList s);
+
+fun fromString s = fromList (String.explode s);
fun toTextFile {filename = f} s =
let
@@ -6415,11 +6350,11 @@
end
-(**** Original file: Heap.sig ****)
+(**** Original file: src/Heap.sig ****)
(* ========================================================================= *)
(* A HEAP DATATYPE FOR ML *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Heap =
@@ -6449,11 +6384,11 @@
end
-(**** Original file: Heap.sml ****)
+(**** Original file: src/Heap.sml ****)
(* ========================================================================= *)
(* A HEAP DATATYPE FOR ML *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Heap :> Metis_Heap =
@@ -6529,32 +6464,44 @@
end
-(**** Original file: Print.sig ****)
+(**** Original file: src/Print.sig ****)
(* ========================================================================= *)
(* PRETTY-PRINTING *)
-(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2008 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Print =
sig
(* ------------------------------------------------------------------------- *)
+(* Escaping strings. *)
+(* ------------------------------------------------------------------------- *)
+
+val escapeString : {escape : char list} -> string -> string
+
+(* ------------------------------------------------------------------------- *)
+(* A type of strings annotated with their size. *)
+(* ------------------------------------------------------------------------- *)
+
+type stringSize = string * int
+
+val mkStringSize : string -> stringSize
+
+(* ------------------------------------------------------------------------- *)
(* A type of pretty-printers. *)
(* ------------------------------------------------------------------------- *)
datatype breakStyle = Consistent | Inconsistent
-datatype ppStep =
+datatype step =
BeginBlock of breakStyle * int
| EndBlock
- | AddString of string
+ | AddString of stringSize
| AddBreak of int
| AddNewline
-type ppstream = ppStep Metis_Stream.stream
-
-type 'a pp = 'a -> ppstream
+type ppstream = step Metis_Stream.stream
(* ------------------------------------------------------------------------- *)
(* Pretty-printer primitives. *)
@@ -6564,7 +6511,7 @@
val endBlock : ppstream
-val addString : string -> ppstream
+val addString : stringSize -> ppstream
val addBreak : int -> ppstream
@@ -6584,18 +6531,24 @@
val blockProgram : breakStyle -> int -> ppstream list -> ppstream
-val bracket : string -> string -> ppstream -> ppstream
-
-val field : string -> ppstream -> ppstream
-
-val record : (string * ppstream) list -> ppstream
+(* ------------------------------------------------------------------------- *)
+(* Executing pretty-printers to generate lines. *)
+(* ------------------------------------------------------------------------- *)
+
+val execute :
+ {lineLength : int} -> ppstream ->
+ {indent : int, line : string} Metis_Stream.stream
(* ------------------------------------------------------------------------- *)
(* Pretty-printer combinators. *)
(* ------------------------------------------------------------------------- *)
+type 'a pp = 'a -> ppstream
+
val ppMap : ('a -> 'b) -> 'b pp -> 'a pp
+val ppString : string pp
+
val ppBracket : string -> string -> 'a pp -> 'a pp
val ppOp : string -> ppstream
@@ -6614,8 +6567,6 @@
val ppChar : char pp
-val ppString : string pp
-
val ppEscapeString : {escape : char list} -> string pp
val ppUnit : unit pp
@@ -6644,7 +6595,7 @@
val ppBreakStyle : breakStyle pp
-val ppPpStep : ppStep pp
+val ppStep : step pp
val ppPpstream : ppstream pp
@@ -6652,28 +6603,27 @@
(* Pretty-printing infix operators. *)
(* ------------------------------------------------------------------------- *)
+type token = string
+
+datatype assoc =
+ LeftAssoc
+ | NonAssoc
+ | RightAssoc
+
datatype infixes =
Infixes of
- {token : string,
+ {token : token,
precedence : int,
- leftAssoc : bool} list
+ assoc : assoc} list
val tokensInfixes : infixes -> Metis_StringSet.set
-val layerInfixes :
- infixes ->
- {tokens : {leftSpaces : int, token : string, rightSpaces : int} list,
- leftAssoc : bool} list
+val layerInfixes : infixes -> {tokens : Metis_StringSet.set, assoc : assoc} list
val ppInfixes :
- infixes -> ('a -> (string * 'a * 'a) option) -> ('a * bool) pp ->
- ('a * bool) pp
-
-(* ------------------------------------------------------------------------- *)
-(* Executing pretty-printers to generate lines. *)
-(* ------------------------------------------------------------------------- *)
-
-val execute : {lineLength : int} -> ppstream -> string Metis_Stream.stream
+ infixes ->
+ ('a -> (token * 'a * 'a) option) -> ('a * token) pp ->
+ ('a * bool) pp -> ('a * bool) pp
(* ------------------------------------------------------------------------- *)
(* Executing pretty-printers with a global line length. *)
@@ -6689,11 +6639,11 @@
end
-(**** Original file: Print.sml ****)
+(**** Original file: src/Print.sml ****)
(* ========================================================================= *)
(* PRETTY-PRINTING *)
-(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2008 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Print :> Metis_Print =
@@ -6722,47 +6672,69 @@
| Metis_Stream.Cons (h,t) => revAppend h (revConcat o t);
local
- fun join current prev = (prev ^ "\n", current);
-in
- fun joinNewline strm =
- case strm of
- Metis_Stream.Nil => Metis_Stream.Nil
- | Metis_Stream.Cons (h,t) => Metis_Stream.maps join Metis_Stream.singleton h (t ());
-end;
-
-local
fun calcSpaces n = nChars #" " n;
- val cachedSpaces = Vector.tabulate (initialLineLength,calcSpaces);
+ val cacheSize = 2 * initialLineLength;
+
+ val cachedSpaces = Vector.tabulate (cacheSize,calcSpaces);
in
fun nSpaces n =
- if n < initialLineLength then Vector.sub (cachedSpaces,n)
+ if n < cacheSize then Vector.sub (cachedSpaces,n)
else calcSpaces n;
end;
(* ------------------------------------------------------------------------- *)
+(* Escaping strings. *)
+(* ------------------------------------------------------------------------- *)
+
+fun escapeString {escape} =
+ let
+ val escapeMap = map (fn c => (c, "\\" ^ str c)) escape
+
+ fun escapeChar c =
+ case c of
+ #"\\" => "\\\\"
+ | #"\n" => "\\n"
+ | #"\t" => "\\t"
+ | _ =>
+ case List.find (equal c o fst) escapeMap of
+ SOME (_,s) => s
+ | NONE => str c
+ in
+ String.translate escapeChar
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* A type of strings annotated with their size. *)
+(* ------------------------------------------------------------------------- *)
+
+type stringSize = string * int;
+
+fun mkStringSize s = (s, size s);
+
+val emptyStringSize = mkStringSize "";
+
+(* ------------------------------------------------------------------------- *)
(* A type of pretty-printers. *)
(* ------------------------------------------------------------------------- *)
datatype breakStyle = Consistent | Inconsistent;
-datatype ppStep =
+datatype step =
BeginBlock of breakStyle * int
| EndBlock
- | AddString of string
+ | AddString of stringSize
| AddBreak of int
| AddNewline;
-type ppstream = ppStep Metis_Stream.stream;
-
-type 'a pp = 'a -> ppstream;
+type ppstream = step Metis_Stream.stream;
fun breakStyleToString style =
case style of
Consistent => "Consistent"
| Inconsistent => "Inconsistent";
-fun ppStepToString step =
+fun stepToString step =
case step of
BeginBlock _ => "BeginBlock"
| EndBlock => "EndBlock"
@@ -6802,330 +6774,6 @@
fun blockProgram style indent pps = block style indent (program pps);
-fun bracket l r pp =
- blockProgram Inconsistent (size l)
- [addString l,
- pp,
- addString r];
-
-fun field f pp =
- blockProgram Inconsistent 2
- [addString (f ^ " ="),
- addBreak 1,
- pp];
-
-val record =
- let
- val sep = sequence (addString ",") (addBreak 1)
-
- fun recordField (f,pp) = field f pp
-
- fun sepField f = sequence sep (recordField f)
-
- fun fields [] = []
- | fields (f :: fs) = recordField f :: map sepField fs
- in
- bracket "{" "}" o blockProgram Consistent 0 o fields
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Pretty-printer combinators. *)
-(* ------------------------------------------------------------------------- *)
-
-fun ppMap f ppB a : ppstream = ppB (f a);
-
-fun ppBracket l r ppA a = bracket l r (ppA a);
-
-fun ppOp s = sequence (if s = "" then skip else addString s) (addBreak 1);
-
-fun ppOp2 ab ppA ppB (a,b) =
- blockProgram Inconsistent 0
- [ppA a,
- ppOp ab,
- ppB b];
-
-fun ppOp3 ab bc ppA ppB ppC (a,b,c) =
- blockProgram Inconsistent 0
- [ppA a,
- ppOp ab,
- ppB b,
- ppOp bc,
- ppC c];
-
-fun ppOpList s ppA =
- let
- fun ppOpA a = sequence (ppOp s) (ppA a)
- in
- fn [] => skip
- | h :: t => blockProgram Inconsistent 0 (ppA h :: map ppOpA t)
- end;
-
-fun ppOpStream s ppA =
- let
- fun ppOpA a = sequence (ppOp s) (ppA a)
- in
- fn Metis_Stream.Nil => skip
- | Metis_Stream.Cons (h,t) =>
- blockProgram Inconsistent 0
- [ppA h,
- Metis_Stream.concat (Metis_Stream.map ppOpA (t ()))]
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Pretty-printers for common types. *)
-(* ------------------------------------------------------------------------- *)
-
-fun ppChar c = addString (str c);
-
-val ppString = addString;
-
-fun ppEscapeString {escape} =
- let
- val escapeMap = map (fn c => (c, "\\" ^ str c)) escape
-
- fun escapeChar c =
- case c of
- #"\\" => "\\\\"
- | #"\n" => "\\n"
- | #"\t" => "\\t"
- | _ =>
- case List.find (equal c o fst) escapeMap of
- SOME (_,s) => s
- | NONE => str c
- in
- fn s => addString (String.translate escapeChar s)
- end;
-
-val ppUnit : unit pp = K (addString "()");
-
-fun ppBool b = addString (if b then "true" else "false");
-
-fun ppInt i = addString (Int.toString i);
-
-local
- val ppNeg = addString "~"
- and ppSep = addString ","
- and ppZero = addString "0"
- and ppZeroZero = addString "00";
-
- fun ppIntBlock i =
- if i < 10 then sequence ppZeroZero (ppInt i)
- else if i < 100 then sequence ppZero (ppInt i)
- else ppInt i;
-
- fun ppIntBlocks i =
- if i < 1000 then ppInt i
- else sequence (ppIntBlocks (i div 1000))
- (sequence ppSep (ppIntBlock (i mod 1000)));
-in
- fun ppPrettyInt i =
- if i < 0 then sequence ppNeg (ppIntBlocks (~i))
- else ppIntBlocks i;
-end;
-
-fun ppReal r = addString (Real.toString r);
-
-fun ppPercent p = addString (percentToString p);
-
-fun ppOrder x =
- addString
- (case x of
- LESS => "Less"
- | EQUAL => "Equal"
- | GREATER => "Greater");
-
-fun ppList ppA = ppBracket "[" "]" (ppOpList "," ppA);
-
-fun ppStream ppA = ppBracket "[" "]" (ppOpStream "," ppA);
-
-fun ppOption ppA ao =
- case ao of
- SOME a => ppA a
- | NONE => addString "-";
-
-fun ppPair ppA ppB = ppBracket "(" ")" (ppOp2 "," ppA ppB);
-
-fun ppTriple ppA ppB ppC = ppBracket "(" ")" (ppOp3 "," "," ppA ppB ppC);
-
-fun ppBreakStyle style = addString (breakStyleToString style);
-
-fun ppPpStep step =
- let
- val cmd = ppStepToString step
- in
- blockProgram Inconsistent 2
- (addString cmd ::
- (case step of
- BeginBlock style_indent =>
- [addBreak 1,
- ppPair ppBreakStyle ppInt style_indent]
- | EndBlock => []
- | AddString s =>
- [addBreak 1,
- addString ("\"" ^ s ^ "\"")]
- | AddBreak n =>
- [addBreak 1,
- ppInt n]
- | AddNewline => []))
- end;
-
-val ppPpstream = ppStream ppPpStep;
-
-(* ------------------------------------------------------------------------- *)
-(* Pretty-printing infix operators. *)
-(* ------------------------------------------------------------------------- *)
-
-datatype infixes =
- Infixes of
- {token : string,
- precedence : int,
- leftAssoc : bool} list;
-
-local
- fun chop l =
- case l of
- #" " :: l => let val (n,l) = chop l in (n + 1, l) end
- | _ => (0,l);
-in
- fun opSpaces tok =
- let
- val tok = explode tok
- val (r,tok) = chop (rev tok)
- val (l,tok) = chop (rev tok)
- val tok = implode tok
- in
- {leftSpaces = l, token = tok, rightSpaces = r}
- end;
-end;
-
-fun ppOpSpaces {leftSpaces,token,rightSpaces} =
- let
- val leftSpacesToken =
- if leftSpaces = 0 then token else nSpaces leftSpaces ^ token
- in
- sequence
- (addString leftSpacesToken)
- (addBreak rightSpaces)
- end;
-
-local
- fun new t l acc = {tokens = [opSpaces t], leftAssoc = l} :: acc;
-
- fun add t l acc =
- case acc of
- [] => raise Bug "Metis_Print.layerInfixOps.layer"
- | {tokens = ts, leftAssoc = l'} :: acc =>
- if l = l' then {tokens = opSpaces t :: ts, leftAssoc = l} :: acc
- else raise Bug "Metis_Print.layerInfixOps: mixed assocs";
-
- fun layer ({token = t, precedence = p, leftAssoc = l}, (acc,p')) =
- let
- val acc = if p = p' then add t l acc else new t l acc
- in
- (acc,p)
- end;
-in
- fun layerInfixes (Infixes i) =
- case sortMap #precedence Int.compare i of
- [] => []
- | {token = t, precedence = p, leftAssoc = l} :: i =>
- let
- val acc = new t l []
-
- val (acc,_) = List.foldl layer (acc,p) i
- in
- acc
- end;
-end;
-
-val tokensLayeredInfixes =
- let
- fun addToken ({leftSpaces = _, token = t, rightSpaces = _}, s) =
- Metis_StringSet.add s t
-
- fun addTokens ({tokens = t, leftAssoc = _}, s) =
- List.foldl addToken s t
- in
- List.foldl addTokens Metis_StringSet.empty
- end;
-
-val tokensInfixes = tokensLayeredInfixes o layerInfixes;
-
-local
- val mkTokenMap =
- let
- fun add (token,m) =
- let
- val {leftSpaces = _, token = t, rightSpaces = _} = token
- in
- Metis_StringMap.insert m (t, ppOpSpaces token)
- end
- in
- List.foldl add (Metis_StringMap.new ())
- end;
-in
- fun ppGenInfix {tokens,leftAssoc} =
- let
- val tokenMap = mkTokenMap tokens
- in
- fn dest => fn ppSub =>
- let
- fun dest' tm =
- case dest tm of
- NONE => NONE
- | SOME (t,a,b) =>
- case Metis_StringMap.peek tokenMap t of
- NONE => NONE
- | SOME p => SOME (p,a,b)
-
- fun ppGo (tmr as (tm,r)) =
- case dest' tm of
- NONE => ppSub tmr
- | SOME (p,a,b) =>
- program
- [(if leftAssoc then ppGo else ppSub) (a,true),
- p,
- (if leftAssoc then ppSub else ppGo) (b,r)]
- in
- fn tmr as (tm,_) =>
- if Option.isSome (dest' tm) then
- block Inconsistent 0 (ppGo tmr)
- else
- ppSub tmr
- end
- end;
-end
-
-fun ppInfixes ops =
- let
- val layeredOps = layerInfixes ops
-
- val toks = tokensLayeredInfixes layeredOps
-
- val iprinters = List.map ppGenInfix layeredOps
- in
- fn dest => fn ppSub =>
- let
- fun printer sub = foldl (fn (ip,p) => ip dest p) sub iprinters
-
- fun isOp t =
- case dest t of
- SOME (x,_,_) => Metis_StringSet.member x toks
- | NONE => false
-
- fun subpr (tmr as (tm,_)) =
- if isOp tm then
- blockProgram Inconsistent 1
- [addString "(",
- printer subpr (tm,false),
- addString ")"]
- else
- ppSub tmr
- in
- fn tmr => block Inconsistent 0 (printer subpr tmr)
- end
- end;
-
(* ------------------------------------------------------------------------- *)
(* Executing pretty-printers to generate lines. *)
(* ------------------------------------------------------------------------- *)
@@ -7189,18 +6837,22 @@
val sizeChunks = List.foldl (fn (c,z) => sizeChunk c + z) 0;
local
- fun render acc [] = acc
- | render acc (chunk :: chunks) =
- case chunk of
- StringChunk {string = s, ...} => render (acc ^ s) chunks
- | BreakChunk n => render (acc ^ nSpaces n) chunks
- | BlockChunk (Block {chunks = l, ...}) =>
- render acc (List.revAppend (l,chunks));
-in
- fun renderChunks indent chunks = render (nSpaces indent) (rev chunks);
-
- fun renderChunk indent chunk = renderChunks indent [chunk];
-end;
+ fun flatten acc chunks =
+ case chunks of
+ [] => rev acc
+ | chunk :: chunks =>
+ case chunk of
+ StringChunk {string = s, ...} => flatten (s :: acc) chunks
+ | BreakChunk n => flatten (nSpaces n :: acc) chunks
+ | BlockChunk (Block {chunks = l, ...}) =>
+ flatten acc (List.revAppend (l,chunks));
+in
+ fun renderChunks indent chunks =
+ {indent = indent,
+ line = String.concat (flatten [] (rev chunks))};
+end;
+
+fun renderChunk indent chunk = renderChunks indent [chunk];
fun isEmptyBlock block =
let
@@ -7216,6 +6868,7 @@
empty
end;
+(*BasicDebug
fun checkBlock ind block =
let
val Block {indent, style = _, size, chunks} = block
@@ -7251,6 +6904,7 @@
in
check initialIndent o rev
end;
+*)
val initialBlock =
let
@@ -7662,12 +7316,10 @@
(lines,state)
end;
- fun executeAddString lineLength s lines state =
+ fun executeAddString lineLength (s,n) lines state =
let
val State {blocks,lineIndent,lineSize} = state
- val n = size s
-
val blocks =
case blocks of
[] => raise Bug "Metis_Print.executeAddString: no block"
@@ -7754,10 +7406,13 @@
fun executeAddNewline lineLength lines state =
let
- val (lines,state) = executeAddString lineLength "" lines state
- val (lines,state) = executeBigBreak lineLength lines state
- in
- executeAddString lineLength "" lines state
+ val (lines,state) =
+ executeAddString lineLength emptyStringSize lines state
+
+ val (lines,state) =
+ executeBigBreak lineLength lines state
+ in
+ executeAddString lineLength emptyStringSize lines state
end;
fun final lineLength lines state =
@@ -7801,7 +7456,7 @@
(lines,state)
end
handle Bug bug =>
- raise Bug ("Metis_Print.advance: after " ^ ppStepToString step ^
+ raise Bug ("Metis_Print.advance: after " ^ stepToString step ^
" command:\n" ^ bug)
*)
in
@@ -7810,26 +7465,377 @@
end;
(* ------------------------------------------------------------------------- *)
+(* Pretty-printer combinators. *)
+(* ------------------------------------------------------------------------- *)
+
+type 'a pp = 'a -> ppstream;
+
+fun ppMap f ppB a : ppstream = ppB (f a);
+
+fun ppBracket' l r ppA a =
+ let
+ val (_,n) = l
+ in
+ blockProgram Inconsistent n
+ [addString l,
+ ppA a,
+ addString r]
+ end;
+
+fun ppOp' s = sequence (addString s) (addBreak 1);
+
+fun ppOp2' ab ppA ppB (a,b) =
+ blockProgram Inconsistent 0
+ [ppA a,
+ ppOp' ab,
+ ppB b];
+
+fun ppOp3' ab bc ppA ppB ppC (a,b,c) =
+ blockProgram Inconsistent 0
+ [ppA a,
+ ppOp' ab,
+ ppB b,
+ ppOp' bc,
+ ppC c];
+
+fun ppOpList' s ppA =
+ let
+ fun ppOpA a = sequence (ppOp' s) (ppA a)
+ in
+ fn [] => skip
+ | h :: t => blockProgram Inconsistent 0 (ppA h :: map ppOpA t)
+ end;
+
+fun ppOpStream' s ppA =
+ let
+ fun ppOpA a = sequence (ppOp' s) (ppA a)
+ in
+ fn Metis_Stream.Nil => skip
+ | Metis_Stream.Cons (h,t) =>
+ blockProgram Inconsistent 0
+ [ppA h,
+ Metis_Stream.concat (Metis_Stream.map ppOpA (t ()))]
+ end;
+
+fun ppBracket l r = ppBracket' (mkStringSize l) (mkStringSize r);
+
+fun ppOp s = ppOp' (mkStringSize s);
+
+fun ppOp2 ab = ppOp2' (mkStringSize ab);
+
+fun ppOp3 ab bc = ppOp3' (mkStringSize ab) (mkStringSize bc);
+
+fun ppOpList s = ppOpList' (mkStringSize s);
+
+fun ppOpStream s = ppOpStream' (mkStringSize s);
+
+(* ------------------------------------------------------------------------- *)
+(* Pretty-printers for common types. *)
+(* ------------------------------------------------------------------------- *)
+
+fun ppChar c = addString (str c, 1);
+
+fun ppString s = addString (mkStringSize s);
+
+fun ppEscapeString escape = ppMap (escapeString escape) ppString;
+
+local
+ val pp = ppString "()";
+in
+ fun ppUnit () = pp;
+end;
+
+local
+ val ppTrue = ppString "true"
+ and ppFalse = ppString "false";
+in
+ fun ppBool b = if b then ppTrue else ppFalse;
+end;
+
+val ppInt = ppMap Int.toString ppString;
+
+local
+ val ppNeg = ppString "~"
+ and ppSep = ppString ","
+ and ppZero = ppString "0"
+ and ppZeroZero = ppString "00";
+
+ fun ppIntBlock i =
+ if i < 10 then sequence ppZeroZero (ppInt i)
+ else if i < 100 then sequence ppZero (ppInt i)
+ else ppInt i;
+
+ fun ppIntBlocks i =
+ if i < 1000 then ppInt i
+ else sequence (ppIntBlocks (i div 1000))
+ (sequence ppSep (ppIntBlock (i mod 1000)));
+in
+ fun ppPrettyInt i =
+ if i < 0 then sequence ppNeg (ppIntBlocks (~i))
+ else ppIntBlocks i;
+end;
+
+val ppReal = ppMap Real.toString ppString;
+
+val ppPercent = ppMap percentToString ppString;
+
+local
+ val ppLess = ppString "Less"
+ and ppEqual = ppString "Equal"
+ and ppGreater = ppString "Greater";
+in
+ fun ppOrder ord =
+ case ord of
+ LESS => ppLess
+ | EQUAL => ppEqual
+ | GREATER => ppGreater;
+end;
+
+local
+ val left = mkStringSize "["
+ and right = mkStringSize "]"
+ and sep = mkStringSize ",";
+in
+ fun ppList ppX xs = ppBracket' left right (ppOpList' sep ppX) xs;
+
+ fun ppStream ppX xs = ppBracket' left right (ppOpStream' sep ppX) xs;
+end;
+
+local
+ val ppNone = ppString "-";
+in
+ fun ppOption ppX xo =
+ case xo of
+ SOME x => ppX x
+ | NONE => ppNone;
+end;
+
+local
+ val left = mkStringSize "("
+ and right = mkStringSize ")"
+ and sep = mkStringSize ",";
+in
+ fun ppPair ppA ppB =
+ ppBracket' left right (ppOp2' sep ppA ppB);
+
+ fun ppTriple ppA ppB ppC =
+ ppBracket' left right (ppOp3' sep sep ppA ppB ppC);
+end;
+
+val ppBreakStyle = ppMap breakStyleToString ppString;
+
+val ppStep = ppMap stepToString ppString;
+
+val ppStringSize =
+ let
+ val left = mkStringSize "\""
+ and right = mkStringSize "\""
+ and escape = {escape = [#"\""]}
+
+ val pp = ppBracket' left right (ppEscapeString escape)
+ in
+ fn (s,n) => if size s = n then pp s else ppPair pp ppInt (s,n)
+ end;
+
+fun ppStep step =
+ blockProgram Inconsistent 2
+ (ppStep step ::
+ (case step of
+ BeginBlock style_indent =>
+ [addBreak 1,
+ ppPair ppBreakStyle ppInt style_indent]
+ | EndBlock => []
+ | AddString s =>
+ [addBreak 1,
+ ppStringSize s]
+ | AddBreak n =>
+ [addBreak 1,
+ ppInt n]
+ | AddNewline => []));
+
+val ppPpstream = ppStream ppStep;
+
+(* ------------------------------------------------------------------------- *)
+(* Pretty-printing infix operators. *)
+(* ------------------------------------------------------------------------- *)
+
+type token = string;
+
+datatype assoc =
+ LeftAssoc
+ | NonAssoc
+ | RightAssoc;
+
+datatype infixes =
+ Infixes of
+ {token : token,
+ precedence : int,
+ assoc : assoc} list;
+
+local
+ fun comparePrecedence (io1,io2) =
+ let
+ val {token = _, precedence = p1, assoc = _} = io1
+ and {token = _, precedence = p2, assoc = _} = io2
+ in
+ Int.compare (p2,p1)
+ end;
+
+ fun equalAssoc a a' =
+ case a of
+ LeftAssoc => (case a' of LeftAssoc => true | _ => false)
+ | NonAssoc => (case a' of NonAssoc => true | _ => false)
+ | RightAssoc => (case a' of RightAssoc => true | _ => false);
+
+ fun new t a acc = {tokens = Metis_StringSet.singleton t, assoc = a} :: acc;
+
+ fun add t a' acc =
+ case acc of
+ [] => raise Bug "Metis_Print.layerInfixes: null"
+ | {tokens = ts, assoc = a} :: acc =>
+ if equalAssoc a a' then {tokens = Metis_StringSet.add ts t, assoc = a} :: acc
+ else raise Bug "Metis_Print.layerInfixes: mixed assocs";
+
+ fun layer ({token = t, precedence = p, assoc = a}, (acc,p')) =
+ let
+ val acc = if p = p' then add t a acc else new t a acc
+ in
+ (acc,p)
+ end;
+in
+ fun layerInfixes (Infixes ios) =
+ case sort comparePrecedence ios of
+ [] => []
+ | {token = t, precedence = p, assoc = a} :: ios =>
+ let
+ val acc = new t a []
+
+ val (acc,_) = List.foldl layer (acc,p) ios
+ in
+ acc
+ end;
+end;
+
+local
+ fun add ({tokens = ts, assoc = _}, tokens) = Metis_StringSet.union ts tokens;
+in
+ fun tokensLayeredInfixes l = List.foldl add Metis_StringSet.empty l;
+end;
+
+fun tokensInfixes ios = tokensLayeredInfixes (layerInfixes ios);
+
+fun destInfixOp dest tokens tm =
+ case dest tm of
+ NONE => NONE
+ | s as SOME (t,a,b) => if Metis_StringSet.member t tokens then s else NONE;
+
+fun ppLayeredInfixes dest ppTok {tokens,assoc} ppLower ppSub =
+ let
+ fun isLayer t = Metis_StringSet.member t tokens
+
+ fun ppTerm aligned (tm,r) =
+ case dest tm of
+ NONE => ppSub (tm,r)
+ | SOME (t,a,b) =>
+ if aligned andalso isLayer t then ppLayer (tm,t,a,b,r)
+ else ppLower (tm,t,a,b,r)
+
+ and ppLeft tm_r =
+ let
+ val aligned = case assoc of LeftAssoc => true | _ => false
+ in
+ ppTerm aligned tm_r
+ end
+
+ and ppLayer (tm,t,a,b,r) =
+ program
+ [ppLeft (a,true),
+ ppTok (tm,t),
+ ppRight (b,r)]
+
+ and ppRight tm_r =
+ let
+ val aligned = case assoc of RightAssoc => true | _ => false
+ in
+ ppTerm aligned tm_r
+ end
+ in
+ fn tm_t_a_b_r as (_,t,_,_,_) =>
+ if isLayer t then block Inconsistent 0 (ppLayer tm_t_a_b_r)
+ else ppLower tm_t_a_b_r
+ end;
+
+local
+ val leftBrack = mkStringSize "("
+ and rightBrack = mkStringSize ")";
+in
+ fun ppInfixes ops =
+ let
+ val layers = layerInfixes ops
+
+ val toks = tokensLayeredInfixes layers
+ in
+ fn dest => fn ppTok => fn ppSub =>
+ let
+ fun destOp tm = destInfixOp dest toks tm
+
+ fun ppInfix tm_t_a_b_r = ppLayers layers tm_t_a_b_r
+
+ and ppLayers ls (tm,t,a,b,r) =
+ case ls of
+ [] =>
+ ppBracket' leftBrack rightBrack ppInfix (tm,t,a,b,false)
+ | l :: ls =>
+ let
+ val ppLower = ppLayers ls
+ in
+ ppLayeredInfixes destOp ppTok l ppLower ppSub (tm,t,a,b,r)
+ end
+ in
+ fn (tm,r) =>
+ case destOp tm of
+ SOME (t,a,b) => ppInfix (tm,t,a,b,r)
+ | NONE => ppSub (tm,r)
+ end
+ end;
+end;
+
+(* ------------------------------------------------------------------------- *)
(* Executing pretty-printers with a global line length. *)
(* ------------------------------------------------------------------------- *)
val lineLength = Unsynchronized.ref initialLineLength;
fun toStream ppA a =
- Metis_Stream.map (fn s => s ^ "\n")
+ Metis_Stream.map (fn {indent,line} => nSpaces indent ^ line ^ "\n")
(execute {lineLength = !lineLength} (ppA a));
-fun toString ppA a =
- case execute {lineLength = !lineLength} (ppA a) of
- Metis_Stream.Nil => ""
- | Metis_Stream.Cons (h,t) => Metis_Stream.foldl (fn (s,z) => z ^ "\n" ^ s) h (t ());
-
-fun trace ppX nameX x =
- Metis_Useful.trace (toString (ppOp2 " =" ppString ppX) (nameX,x) ^ "\n");
-
-end
-
-(**** Original file: Parse.sig ****)
+local
+ fun inc {indent,line} acc = line :: nSpaces indent :: acc;
+
+ fun incn (indent_line,acc) = inc indent_line ("\n" :: acc);
+in
+ fun toString ppA a =
+ case execute {lineLength = !lineLength} (ppA a) of
+ Metis_Stream.Nil => ""
+ | Metis_Stream.Cons (h,t) =>
+ let
+ val lines = Metis_Stream.foldl incn (inc h []) (t ())
+ in
+ String.concat (rev lines)
+ end;
+end;
+
+local
+ val sep = mkStringSize " =";
+in
+ fun trace ppX nameX x =
+ Metis_Useful.trace (toString (ppOp2' sep ppString ppX) (nameX,x) ^ "\n");
+end;
+
+end
+
+(**** Original file: src/Parse.sig ****)
(* ========================================================================= *)
(* PARSING *)
@@ -7932,8 +7938,9 @@
(* ------------------------------------------------------------------------- *)
val parseInfixes :
- Metis_Print.infixes -> (string * 'a * 'a -> 'a) -> (string,'a) parser ->
- (string,'a) parser
+ Metis_Print.infixes ->
+ (Metis_Print.token * 'a * 'a -> 'a) -> ('b,Metis_Print.token) parser ->
+ ('b,'a) parser -> ('b,'a) parser
(* ------------------------------------------------------------------------- *)
(* Quotations. *)
@@ -7945,7 +7952,7 @@
end
-(**** Original file: Parse.sml ****)
+(**** Original file: src/Parse.sml ****)
(* ========================================================================= *)
(* PARSING *)
@@ -8083,7 +8090,7 @@
val Unsynchronized.ref (n,_,l2,l3) = lastLine
val () = lastLine := (n + 1, l2, l3, line)
in
- explode line
+ String.explode line
end
in
Metis_Stream.memoize (Metis_Stream.map saveLast lines)
@@ -8109,7 +8116,7 @@
[] => nothing
| c :: cs => (exactChar c ++ exactCharList cs) >> snd;
-fun exactString s = exactCharList (explode s);
+fun exactString s = exactCharList (String.explode s);
fun escapeString {escape} =
let
@@ -8132,7 +8139,7 @@
((exactChar #"\\" ++ escapeParser) >> snd) ||
some isNormal
in
- many charParser >> implode
+ many charParser >> String.implode
end;
local
@@ -8145,46 +8152,51 @@
val atLeastOneSpace = atLeastOne space >> K ();
end;
-fun fromString parser s = fromList parser (explode s);
+fun fromString parser s = fromList parser (String.explode s);
(* ------------------------------------------------------------------------- *)
(* Infix operators. *)
(* ------------------------------------------------------------------------- *)
-fun parseGenInfix update sof toks parse inp =
- let
- val (e,rest) = parse inp
-
- val continue =
- case rest of
- Metis_Stream.Nil => NONE
- | Metis_Stream.Cons (h_t as (h,_)) =>
- if Metis_StringSet.member h toks then SOME h_t else NONE
- in
- case continue of
- NONE => (sof e, rest)
- | SOME (h,t) => parseGenInfix update (update sof h e) toks parse (t ())
- end;
-
-local
- fun add ({leftSpaces = _, token = t, rightSpaces = _}, s) = Metis_StringSet.add s t;
-
- fun parse leftAssoc toks con =
- let
- val update =
- if leftAssoc then (fn f => fn t => fn a => fn b => con (t, f a, b))
- else (fn f => fn t => fn a => fn b => f (con (t, a, b)))
- in
- parseGenInfix update I toks
- end;
-in
- fun parseLayeredInfixes {tokens,leftAssoc} =
- let
- val toks = List.foldl add Metis_StringSet.empty tokens
- in
- parse leftAssoc toks
- end;
-end;
+fun parseLayeredInfixes {tokens,assoc} mk tokParser subParser =
+ let
+ fun layerTokParser inp =
+ let
+ val tok_rest as (tok,_) = tokParser inp
+ in
+ if Metis_StringSet.member tok tokens then tok_rest
+ else raise NoParse
+ end
+
+ fun layerMk (x,txs) =
+ case assoc of
+ Metis_Print.LeftAssoc =>
+ let
+ fun inc ((t,y),z) = mk (t,z,y)
+ in
+ List.foldl inc x txs
+ end
+ | Metis_Print.NonAssoc =>
+ (case txs of
+ [] => x
+ | [(t,y)] => mk (t,x,y)
+ | _ => raise NoParse)
+ | Metis_Print.RightAssoc =>
+ (case rev txs of
+ [] => x
+ | tx :: txs =>
+ let
+ fun inc ((t,y),(u,z)) = (t, mk (u,y,z))
+
+ val (t,y) = List.foldl inc tx txs
+ in
+ mk (t,x,y)
+ end)
+
+ val layerParser = subParser ++ many (layerTokParser ++ subParser)
+ in
+ layerParser >> layerMk
+ end;
fun parseInfixes ops =
let
@@ -8192,7 +8204,8 @@
val iparsers = List.map parseLayeredInfixes layeredOps
in
- fn con => fn subparser => foldl (fn (p,sp) => p con sp) subparser iparsers
+ fn mk => fn tokParser => fn subParser =>
+ List.foldr (fn (p,sp) => p mk tokParser sp) subParser iparsers
end;
(* ------------------------------------------------------------------------- *)
@@ -8206,14 +8219,14 @@
fun expand (QUOTE q, s) = s ^ q
| expand (ANTIQUOTE a, s) = s ^ printer a
- val string = foldl expand "" quote
+ val string = List.foldl expand "" quote
in
parser string
end;
end
-(**** Original file: Name.sig ****)
+(**** Original file: src/Name.sig ****)
(* ========================================================================= *)
(* NAMES *)
@@ -8261,7 +8274,7 @@
end
-(**** Original file: Name.sml ****)
+(**** Original file: src/Name.sml ****)
(* ========================================================================= *)
(* NAMES *)
@@ -8349,24 +8362,24 @@
structure Metis_NameSet =
struct
- local
- structure S = Metis_ElementSet (Metis_NameMap);
- in
- open S;
- end;
-
- val pp =
- Metis_Print.ppMap
- toList
- (Metis_Print.ppBracket "{" "}" (Metis_Print.ppOpList "," Metis_Name.pp));
-
-end
-
-(**** Original file: NameArity.sig ****)
+local
+ structure S = Metis_ElementSet (Metis_NameMap);
+in
+ open S;
+end;
+
+val pp =
+ Metis_Print.ppMap
+ toList
+ (Metis_Print.ppBracket "{" "}" (Metis_Print.ppOpList "," Metis_Name.pp));
+
+end
+
+(**** Original file: src/NameArity.sig ****)
(* ========================================================================= *)
(* NAME/ARITY PAIRS *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_NameArity =
@@ -8412,11 +8425,11 @@
end
-(**** Original file: NameArity.sml ****)
+(**** Original file: src/NameArity.sml ****)
(* ========================================================================= *)
(* NAME/ARITY PAIRS *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_NameArity :> Metis_NameArity =
@@ -8462,7 +8475,7 @@
fun pp (n,i) =
Metis_Print.blockProgram Metis_Print.Inconsistent 0
[Metis_Name.pp n,
- Metis_Print.addString "/",
+ Metis_Print.ppString "/",
Metis_Print.ppInt i];
end
@@ -8473,40 +8486,41 @@
structure Metis_NameArityMap =
struct
- local
- structure S = Metis_KeyMap (Metis_NameArityOrdered);
- in
- open S;
- end;
-
- fun compose m1 m2 =
- let
- fun pk ((_,a),n) = peek m2 (n,a)
- in
- mapPartial pk m1
- end;
+local
+ structure S = Metis_KeyMap (Metis_NameArityOrdered);
+in
+ open S;
+end;
+
+fun compose m1 m2 =
+ let
+ fun pk ((_,a),n) = peek m2 (n,a)
+ in
+ mapPartial pk m1
+ end;
end
structure Metis_NameAritySet =
struct
- local
- structure S = Metis_ElementSet (Metis_NameArityMap);
- in
- open S;
- end;
-
- val allNullary = all Metis_NameArity.nullary;
-
- val pp =
- Metis_Print.ppMap
- toList
- (Metis_Print.ppBracket "{" "}" (Metis_Print.ppOpList "," Metis_NameArity.pp));
-
-end
-
-(**** Original file: Term.sig ****)
+local
+ structure S = Metis_ElementSet (Metis_NameArityMap);
+in
+ open S;
+end;
+
+val allNullary = all Metis_NameArity.nullary;
+
+val pp =
+ Metis_Print.ppMap
+ toList
+ (Metis_Print.ppBracket "{" "}" (Metis_Print.ppOpList "," Metis_NameArity.pp));
+
+
+end
+
+(**** Original file: src/Term.sig ****)
(* ========================================================================= *)
(* FIRST ORDER LOGIC TERMS *)
@@ -8696,7 +8710,7 @@
end
-(**** Original file: Term.sml ****)
+(**** Original file: src/Term.sml ****)
(* ========================================================================= *)
(* FIRST ORDER LOGIC TERMS *)
@@ -9048,7 +9062,7 @@
val isApp = can destApp;
-fun listMkApp (f,l) = foldl mkApp f l;
+fun listMkApp (f,l) = List.foldl mkApp f l;
local
fun strip tms tm =
@@ -9068,38 +9082,38 @@
val infixes =
(Unsynchronized.ref o Metis_Print.Infixes)
[(* ML symbols *)
- {token = " / ", precedence = 7, leftAssoc = true},
- {token = " div ", precedence = 7, leftAssoc = true},
- {token = " mod ", precedence = 7, leftAssoc = true},
- {token = " * ", precedence = 7, leftAssoc = true},
- {token = " + ", precedence = 6, leftAssoc = true},
- {token = " - ", precedence = 6, leftAssoc = true},
- {token = " ^ ", precedence = 6, leftAssoc = true},
- {token = " @ ", precedence = 5, leftAssoc = false},
- {token = " :: ", precedence = 5, leftAssoc = false},
- {token = " = ", precedence = 4, leftAssoc = true},
- {token = " <> ", precedence = 4, leftAssoc = true},
- {token = " <= ", precedence = 4, leftAssoc = true},
- {token = " < ", precedence = 4, leftAssoc = true},
- {token = " >= ", precedence = 4, leftAssoc = true},
- {token = " > ", precedence = 4, leftAssoc = true},
- {token = " o ", precedence = 3, leftAssoc = true},
- {token = " -> ", precedence = 2, leftAssoc = false}, (* inferred prec *)
- {token = " : ", precedence = 1, leftAssoc = false}, (* inferred prec *)
- {token = ", ", precedence = 0, leftAssoc = false}, (* inferred prec *)
+ {token = "/", precedence = 7, assoc = Metis_Print.LeftAssoc},
+ {token = "div", precedence = 7, assoc = Metis_Print.LeftAssoc},
+ {token = "mod", precedence = 7, assoc = Metis_Print.LeftAssoc},
+ {token = "*", precedence = 7, assoc = Metis_Print.LeftAssoc},
+ {token = "+", precedence = 6, assoc = Metis_Print.LeftAssoc},
+ {token = "-", precedence = 6, assoc = Metis_Print.LeftAssoc},
+ {token = "^", precedence = 6, assoc = Metis_Print.LeftAssoc},
+ {token = "@", precedence = 5, assoc = Metis_Print.RightAssoc},
+ {token = "::", precedence = 5, assoc = Metis_Print.RightAssoc},
+ {token = "=", precedence = 4, assoc = Metis_Print.NonAssoc},
+ {token = "<>", precedence = 4, assoc = Metis_Print.NonAssoc},
+ {token = "<=", precedence = 4, assoc = Metis_Print.NonAssoc},
+ {token = "<", precedence = 4, assoc = Metis_Print.NonAssoc},
+ {token = ">=", precedence = 4, assoc = Metis_Print.NonAssoc},
+ {token = ">", precedence = 4, assoc = Metis_Print.NonAssoc},
+ {token = "o", precedence = 3, assoc = Metis_Print.LeftAssoc},
+ {token = "->", precedence = 2, assoc = Metis_Print.RightAssoc},
+ {token = ":", precedence = 1, assoc = Metis_Print.NonAssoc},
+ {token = ",", precedence = 0, assoc = Metis_Print.RightAssoc},
(* Logical connectives *)
- {token = " /\\ ", precedence = ~1, leftAssoc = false},
- {token = " \\/ ", precedence = ~2, leftAssoc = false},
- {token = " ==> ", precedence = ~3, leftAssoc = false},
- {token = " <=> ", precedence = ~4, leftAssoc = false},
+ {token = "/\\", precedence = ~1, assoc = Metis_Print.RightAssoc},
+ {token = "\\/", precedence = ~2, assoc = Metis_Print.RightAssoc},
+ {token = "==>", precedence = ~3, assoc = Metis_Print.RightAssoc},
+ {token = "<=>", precedence = ~4, assoc = Metis_Print.RightAssoc},
(* Other symbols *)
- {token = " . ", precedence = 9, leftAssoc = true}, (* function app *)
- {token = " ** ", precedence = 8, leftAssoc = true},
- {token = " ++ ", precedence = 6, leftAssoc = true},
- {token = " -- ", precedence = 6, leftAssoc = true},
- {token = " == ", precedence = 4, leftAssoc = true}];
+ {token = ".", precedence = 9, assoc = Metis_Print.LeftAssoc},
+ {token = "**", precedence = 8, assoc = Metis_Print.LeftAssoc},
+ {token = "++", precedence = 6, assoc = Metis_Print.LeftAssoc},
+ {token = "--", precedence = 6, assoc = Metis_Print.LeftAssoc},
+ {token = "==", precedence = 4, assoc = Metis_Print.NonAssoc}];
(* The negation symbol *)
@@ -9122,9 +9136,14 @@
and neg = !negation
and bracks = !brackets
- val bracks = map (fn (b1,b2) => (b1 ^ b2, b1, b2)) bracks
-
- val bTokens = map #2 bracks @ map #3 bracks
+ val bMap =
+ let
+ fun f (b1,b2) = (b1 ^ b2, b1, b2)
+ in
+ List.map f bracks
+ end
+
+ val bTokens = op@ (unzip bracks)
val iTokens = Metis_Print.tokensInfixes iOps
@@ -9138,7 +9157,15 @@
end
| _ => NONE
- val iPrinter = Metis_Print.ppInfixes iOps destI
+ fun isI tm = Option.isSome (destI tm)
+
+ fun iToken (_,tok) =
+ Metis_Print.program
+ [(if tok = "," then Metis_Print.skip else Metis_Print.ppString " "),
+ Metis_Print.ppString tok,
+ Metis_Print.addBreak 1];
+
+ val iPrinter = Metis_Print.ppInfixes iOps destI iToken
val specialTokens =
Metis_StringSet.addList iTokens (neg :: quants @ ["$","(",")"] @ bTokens)
@@ -9166,8 +9193,6 @@
fun functionName bv = Metis_Print.ppMap (checkFunctionName bv) Metis_Print.ppString
- fun isI tm = Option.isSome (destI tm)
-
fun stripNeg tm =
case tm of
Fn (f,[a]) =>
@@ -9194,7 +9219,7 @@
let
val s = Metis_Name.toString b
in
- case List.find (fn (n,_,_) => n = s) bracks of
+ case List.find (fn (n,_,_) => n = s) bMap of
NONE => NONE
| SOME (_,b1,b2) => SOME (b1,tm,b2)
end
@@ -9227,11 +9252,11 @@
val bv = Metis_StringSet.addList bv (map Metis_Name.toString (v :: vs))
in
Metis_Print.program
- [Metis_Print.addString q,
+ [Metis_Print.ppString q,
varName bv v,
Metis_Print.program
(map (Metis_Print.sequence (Metis_Print.addBreak 1) o varName bv) vs),
- Metis_Print.addString ".",
+ Metis_Print.ppString ".",
Metis_Print.addBreak 1,
innerQuant bv tm]
end
@@ -9245,7 +9270,7 @@
val (n,tm) = stripNeg tm
in
Metis_Print.blockProgram Metis_Print.Inconsistent n
- [Metis_Print.duplicate n (Metis_Print.addString neg),
+ [Metis_Print.duplicate n (Metis_Print.ppString neg),
if isI tm orelse (r andalso isQuant tm) then bracket bv tm
else quantifier bv tm]
end
@@ -9271,31 +9296,32 @@
val isAlphaNum =
let
- val alphaNumChars = explode "_'"
+ val alphaNumChars = String.explode "_'"
in
fn c => mem c alphaNumChars orelse Char.isAlphaNum c
end;
local
- val alphaNumToken = atLeastOne (some isAlphaNum) >> implode;
+ val alphaNumToken = atLeastOne (some isAlphaNum) >> String.implode;
val symbolToken =
let
fun isNeg c = str c = !negation
- val symbolChars = explode "<>=-*+/\\?@|!$%&#^:;~"
+ val symbolChars = String.explode "<>=-*+/\\?@|!$%&#^:;~"
fun isSymbol c = mem c symbolChars
fun isNonNegSymbol c = not (isNeg c) andalso isSymbol c
in
some isNeg >> str ||
- (some isNonNegSymbol ++ many (some isSymbol)) >> (implode o op::)
+ (some isNonNegSymbol ++ many (some isSymbol)) >>
+ (String.implode o op::)
end;
val punctToken =
let
- val punctChars = explode "()[]{}.,"
+ val punctChars = String.explode "()[]{}.,"
fun isPunct c = mem c punctChars
in
@@ -9327,8 +9353,9 @@
val iTokens = Metis_Print.tokensInfixes iOps
- val iParser =
- parseInfixes iOps (fn (f,a,b) => Fn (Metis_Name.fromString f, [a,b]))
+ fun iMk (f,a,b) = Fn (Metis_Name.fromString f, [a,b])
+
+ val iParser = parseInfixes iOps iMk any
val specialTokens =
Metis_StringSet.addList iTokens (neg :: quants @ ["$"] @ bTokens)
@@ -9367,7 +9394,7 @@
some (Metis_Useful.equal ".")) >>++
(fn (_,(vs,_)) =>
term (Metis_StringSet.addList bv vs) >>
- (fn body => foldr bind body vs))
+ (fn body => List.foldr bind body vs))
end
in
var ||
@@ -9396,7 +9423,7 @@
in
fun fromString input =
let
- val chars = Metis_Stream.fromList (explode input)
+ val chars = Metis_Stream.fromList (String.explode input)
val tokens = everything (lexer >> singleton) chars
@@ -9409,7 +9436,8 @@
end;
local
- val antiquotedTermToString = Metis_Print.toString (Metis_Print.ppBracket "(" ")" pp);
+ val antiquotedTermToString =
+ Metis_Print.toString (Metis_Print.ppBracket "(" ")" pp);
in
val parse = Metis_Parse.parseQuotation antiquotedTermToString fromString;
end;
@@ -9423,11 +9451,11 @@
structure Metis_TermSet = Metis_ElementSet (Metis_TermMap);
-(**** Original file: Subst.sig ****)
+(**** Original file: src/Subst.sig ****)
(* ========================================================================= *)
(* FIRST ORDER LOGIC SUBSTITUTIONS *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Subst =
@@ -9545,11 +9573,11 @@
end
-(**** Original file: Subst.sml ****)
+(**** Original file: src/Subst.sml ****)
(* ========================================================================= *)
(* FIRST ORDER LOGIC SUBSTITUTIONS *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Subst :> Metis_Subst =
@@ -9797,11 +9825,11 @@
end
-(**** Original file: Atom.sig ****)
+(**** Original file: src/Atom.sig ****)
(* ========================================================================= *)
(* FIRST ORDER LOGIC ATOMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Atom =
@@ -9941,11 +9969,11 @@
end
-(**** Original file: Atom.sml ****)
+(**** Original file: src/Atom.sml ****)
(* ========================================================================= *)
(* FIRST ORDER LOGIC ATOMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Atom :> Metis_Atom =
@@ -9979,14 +10007,14 @@
let
fun f (tm,acc) = Metis_NameAritySet.union (Metis_Term.functions tm) acc
in
- fn atm => foldl f Metis_NameAritySet.empty (arguments atm)
+ fn atm => List.foldl f Metis_NameAritySet.empty (arguments atm)
end;
val functionNames =
let
fun f (tm,acc) = Metis_NameSet.union (Metis_Term.functionNames tm) acc
in
- fn atm => foldl f Metis_NameSet.empty (arguments atm)
+ fn atm => List.foldl f Metis_NameSet.empty (arguments atm)
end;
(* Binary relations *)
@@ -10003,7 +10031,8 @@
(* The size of an atom in symbols. *)
(* ------------------------------------------------------------------------- *)
-fun symbols atm = foldl (fn (tm,z) => Metis_Term.symbols tm + z) 1 (arguments atm);
+fun symbols atm =
+ List.foldl (fn (tm,z) => Metis_Term.symbols tm + z) 1 (arguments atm);
(* ------------------------------------------------------------------------- *)
(* A total comparison function for atoms. *)
@@ -10030,7 +10059,7 @@
let
fun f ((n,tm),l) = map (fn (p,s) => (n :: p, s)) (Metis_Term.subterms tm) @ l
in
- foldl f [] (enumerate tms)
+ List.foldl f [] (enumerate tms)
end;
fun replace _ ([],_) = raise Bug "Metis_Atom.replace: empty path"
@@ -10065,7 +10094,7 @@
let
fun f (tm,acc) = Metis_NameSet.union (Metis_Term.freeVars tm) acc
in
- fn atm => foldl f Metis_NameSet.empty (arguments atm)
+ fn atm => List.foldl f Metis_NameSet.empty (arguments atm)
end;
(* ------------------------------------------------------------------------- *)
@@ -10091,7 +10120,7 @@
val _ = (Metis_Name.equal p1 p2 andalso length tms1 = length tms2) orelse
raise Error "Metis_Atom.match"
in
- foldl matchArg sub (zip tms1 tms2)
+ List.foldl matchArg sub (zip tms1 tms2)
end;
end;
@@ -10107,7 +10136,7 @@
val _ = (Metis_Name.equal p1 p2 andalso length tms1 = length tms2) orelse
raise Error "Metis_Atom.unify"
in
- foldl unifyArg sub (zip tms1 tms2)
+ List.foldl unifyArg sub (zip tms1 tms2)
end;
end;
@@ -10156,7 +10185,7 @@
(* ------------------------------------------------------------------------- *)
fun typedSymbols ((_,tms) : atom) =
- foldl (fn (tm,z) => Metis_Term.typedSymbols tm + z) 1 tms;
+ List.foldl (fn (tm,z) => Metis_Term.typedSymbols tm + z) 1 tms;
fun nonVarTypedSubterms (_,tms) =
let
@@ -10164,10 +10193,10 @@
let
fun addTm ((path,tm),acc) = (n :: path, tm) :: acc
in
- foldl addTm acc (Metis_Term.nonVarTypedSubterms arg)
- end
- in
- foldl addArg [] (enumerate tms)
+ List.foldl addTm acc (Metis_Term.nonVarTypedSubterms arg)
+ end
+ in
+ List.foldl addArg [] (enumerate tms)
end;
(* ------------------------------------------------------------------------- *)
@@ -10191,11 +10220,11 @@
structure Metis_AtomSet = Metis_ElementSet (Metis_AtomMap);
-(**** Original file: Formula.sig ****)
+(**** Original file: src/Formula.sig ****)
(* ========================================================================= *)
(* FIRST ORDER LOGIC FORMULAS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Formula =
@@ -10391,11 +10420,11 @@
end
-(**** Original file: Formula.sml ****)
+(**** Original file: src/Formula.sml ****)
(* ========================================================================= *)
(* FIRST ORDER LOGIC FORMULAS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Formula :> Metis_Formula =
@@ -10540,7 +10569,7 @@
(* Conjunctions *)
fun listMkConj fms =
- case rev fms of [] => True | fm :: fms => foldl And fm fms;
+ case rev fms of [] => True | fm :: fms => List.foldl And fm fms;
local
fun strip cs (And (p,q)) = strip (p :: cs) q
@@ -10563,7 +10592,7 @@
(* Disjunctions *)
fun listMkDisj fms =
- case rev fms of [] => False | fm :: fms => foldl Or fm fms;
+ case rev fms of [] => False | fm :: fms => List.foldl Or fm fms;
local
fun strip cs (Or (p,q)) = strip (p :: cs) q
@@ -10586,7 +10615,7 @@
(* Equivalences *)
fun listMkEquiv fms =
- case rev fms of [] => True | fm :: fms => foldl Iff fm fms;
+ case rev fms of [] => True | fm :: fms => List.foldl Iff fm fms;
local
fun strip cs (Iff (p,q)) = strip (p :: cs) q
@@ -10995,11 +11024,11 @@
structure Metis_FormulaSet = Metis_ElementSet (Metis_FormulaMap);
-(**** Original file: Literal.sig ****)
+(**** Original file: src/Literal.sig ****)
(* ========================================================================= *)
(* FIRST ORDER LOGIC LITERALS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Literal =
@@ -11163,11 +11192,11 @@
end
-(**** Original file: Literal.sml ****)
+(**** Original file: src/Literal.sml ****)
(* ========================================================================= *)
(* FIRST ORDER LOGIC LITERALS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Literal :> Metis_Literal =
@@ -11459,7 +11488,7 @@
structure Metis_LiteralSetSet = Metis_ElementSet (Metis_LiteralSetMap);
-(**** Original file: Thm.sig ****)
+(**** Original file: src/Thm.sig ****)
(* ========================================================================= *)
(* A LOGICAL KERNEL FOR FIRST ORDER CLAUSAL THEOREMS *)
@@ -11610,7 +11639,7 @@
end
-(**** Original file: Thm.sml ****)
+(**** Original file: src/Thm.sml ****)
(* ========================================================================= *)
(* A LOGICAL KERNEL FOR FIRST ORDER CLAUSAL THEOREMS *)
@@ -11725,7 +11754,7 @@
in
fun pp th =
Metis_Print.blockProgram Metis_Print.Inconsistent 3
- [Metis_Print.addString "|- ",
+ [Metis_Print.ppString "|- ",
Metis_Formula.pp (toFormula th)];
end;
@@ -11827,11 +11856,11 @@
end
-(**** Original file: Proof.sig ****)
+(**** Original file: src/Proof.sig ****)
(* ========================================================================= *)
(* PROOFS IN FIRST ORDER LOGIC *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Proof =
@@ -11891,11 +11920,11 @@
end
-(**** Original file: Proof.sml ****)
+(**** Original file: src/Proof.sml ****)
(* ========================================================================= *)
(* PROOFS IN FIRST ORDER LOGIC *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Proof :> Metis_Proof =
@@ -11934,40 +11963,40 @@
fun ppSubst ppThm (sub,thm) =
Metis_Print.sequence (Metis_Print.addBreak 1)
(Metis_Print.blockProgram Metis_Print.Inconsistent 1
- [Metis_Print.addString "{",
+ [Metis_Print.ppString "{",
Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Subst.pp ("sub",sub),
- Metis_Print.addString ",",
+ Metis_Print.ppString ",",
Metis_Print.addBreak 1,
Metis_Print.ppOp2 " =" Metis_Print.ppString ppThm ("thm",thm),
- Metis_Print.addString "}"]);
+ Metis_Print.ppString "}"]);
fun ppResolve ppThm (res,pos,neg) =
Metis_Print.sequence (Metis_Print.addBreak 1)
(Metis_Print.blockProgram Metis_Print.Inconsistent 1
- [Metis_Print.addString "{",
+ [Metis_Print.ppString "{",
Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Atom.pp ("res",res),
- Metis_Print.addString ",",
+ Metis_Print.ppString ",",
Metis_Print.addBreak 1,
Metis_Print.ppOp2 " =" Metis_Print.ppString ppThm ("pos",pos),
- Metis_Print.addString ",",
+ Metis_Print.ppString ",",
Metis_Print.addBreak 1,
Metis_Print.ppOp2 " =" Metis_Print.ppString ppThm ("neg",neg),
- Metis_Print.addString "}"]);
+ Metis_Print.ppString "}"]);
fun ppRefl tm = Metis_Print.sequence (Metis_Print.addBreak 1) (Metis_Term.pp tm);
fun ppEquality (lit,path,res) =
Metis_Print.sequence (Metis_Print.addBreak 1)
(Metis_Print.blockProgram Metis_Print.Inconsistent 1
- [Metis_Print.addString "{",
+ [Metis_Print.ppString "{",
Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Literal.pp ("lit",lit),
- Metis_Print.addString ",",
+ Metis_Print.ppString ",",
Metis_Print.addBreak 1,
Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Term.ppPath ("path",path),
- Metis_Print.addString ",",
+ Metis_Print.ppString ",",
Metis_Print.addBreak 1,
Metis_Print.ppOp2 " =" Metis_Print.ppString Metis_Term.pp ("res",res),
- Metis_Print.addString "}"]);
+ Metis_Print.ppString "}"]);
fun ppInf ppAxiom ppThm inf =
let
@@ -11975,7 +12004,7 @@
in
Metis_Print.block Metis_Print.Inconsistent 2
(Metis_Print.sequence
- (Metis_Print.addString infString)
+ (Metis_Print.ppString infString)
(case inf of
Axiom cl => ppAxiom cl
| Assume x => ppAssume x
@@ -12001,7 +12030,7 @@
val prf = enumerate prf
fun ppThm th =
- Metis_Print.addString
+ Metis_Print.ppString
let
val cl = Metis_Thm.clause th
@@ -12018,7 +12047,7 @@
in
Metis_Print.sequence
(Metis_Print.blockProgram Metis_Print.Consistent (1 + size s)
- [Metis_Print.addString (s ^ " "),
+ [Metis_Print.ppString (s ^ " "),
Metis_Thm.pp th,
Metis_Print.addBreak 2,
Metis_Print.ppBracket "[" "]" (ppInf (K Metis_Print.skip) ppThm) inf])
@@ -12026,10 +12055,10 @@
end
in
Metis_Print.blockProgram Metis_Print.Consistent 0
- [Metis_Print.addString "START OF PROOF",
+ [Metis_Print.ppString "START OF PROOF",
Metis_Print.addNewline,
Metis_Print.program (map ppStep prf),
- Metis_Print.addString "END OF PROOF"]
+ Metis_Print.ppString "END OF PROOF"]
end
(*MetisDebug
handle Error err => raise Bug ("Metis_Proof.pp: shouldn't fail:\n" ^ err);
@@ -12346,11 +12375,11 @@
end
-(**** Original file: Rule.sig ****)
+(**** Original file: src/Rule.sig ****)
(* ========================================================================= *)
(* DERIVED RULES FOR CREATING FIRST ORDER LOGIC THEOREMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Rule =
@@ -12623,11 +12652,11 @@
end
-(**** Original file: Rule.sml ****)
+(**** Original file: src/Rule.sml ****)
(* ========================================================================= *)
(* DERIVED RULES FOR CREATING FIRST ORDER LOGIC THEOREMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Rule :> Metis_Rule =
@@ -12783,17 +12812,19 @@
val noConv : conv = fn _ => raise Error "noConv";
+(*MetisDebug
fun traceConv s conv tm =
let
val res as (tm',th) = conv tm
- val () = print (s ^ ": " ^ Metis_Term.toString tm ^ " --> " ^
+ val () = trace (s ^ ": " ^ Metis_Term.toString tm ^ " --> " ^
Metis_Term.toString tm' ^ " " ^ Metis_Thm.toString th ^ "\n")
in
res
end
handle Error err =>
- (print (s ^ ": " ^ Metis_Term.toString tm ^ " --> Error: " ^ err ^ "\n");
+ (trace (s ^ ": " ^ Metis_Term.toString tm ^ " --> Error: " ^ err ^ "\n");
raise Error (s ^ ": " ^ err));
+*)
fun thenConvTrans tm (tm',th1) (tm'',th2) =
let
@@ -13082,7 +13113,7 @@
val reflTh = Metis_Thm.refl (Metis_Term.Fn (f,xs))
val reflLit = Metis_Thm.destUnit reflTh
in
- fst (foldl cong (reflTh,reflLit) (enumerate ys))
+ fst (List.foldl cong (reflTh,reflLit) (enumerate ys))
end;
(* ------------------------------------------------------------------------- *)
@@ -13109,7 +13140,7 @@
val assumeLit = (false,(R,xs))
val assumeTh = Metis_Thm.assume assumeLit
in
- fst (foldl cong (assumeTh,assumeLit) (enumerate ys))
+ fst (List.foldl cong (assumeTh,assumeLit) (enumerate ys))
end;
(* ------------------------------------------------------------------------- *)
@@ -13409,11 +13440,11 @@
end
-(**** Original file: Normalize.sig ****)
+(**** Original file: src/Normalize.sig ****)
(* ========================================================================= *)
(* NORMALIZING FORMULAS *)
-(* Copyright (c) 2001-2009 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Normalize =
@@ -13467,11 +13498,11 @@
end
-(**** Original file: Normalize.sml ****)
+(**** Original file: src/Normalize.sml ****)
(* ========================================================================= *)
(* NORMALIZING FORMULAS *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Normalize :> Metis_Normalize =
@@ -14156,23 +14187,23 @@
(* Basic conjunctive normal form. *)
(* ------------------------------------------------------------------------- *)
-val newSkolemFunction =
- let
- val counter : int Metis_StringMap.map Unsynchronized.ref = Unsynchronized.ref (Metis_StringMap.new ())
- in
- (* MODIFIED by Jasmin Blanchette *)
- fn n => CRITICAL (fn () =>
- let
- val Unsynchronized.ref m = counter
- val s = Metis_Name.toString n
- val i = Option.getOpt (Metis_StringMap.peek m s, 0)
- val () = counter := Metis_StringMap.insert m (s, i + 1)
- val i = if i = 0 then "" else "_" ^ Int.toString i
- val s = skolemPrefix ^ "_" ^ s ^ i
- in
- Metis_Name.fromString s
- end)
- end;
+local
+ val counter : int Metis_StringMap.map Unsynchronized.ref = Unsynchronized.ref (Metis_StringMap.new ());
+
+ fun new n () =
+ let
+ val Unsynchronized.ref m = counter
+ val s = Metis_Name.toString n
+ val i = Option.getOpt (Metis_StringMap.peek m s, 0)
+ val () = counter := Metis_StringMap.insert m (s, i + 1)
+ val i = if i = 0 then "" else "_" ^ Int.toString i
+ val s = skolemPrefix ^ "_" ^ s ^ i
+ in
+ Metis_Name.fromString s
+ end;
+in
+ fun newSkolemFunction n = Metis_Portable.critical (new n) ();
+end;
fun skolemize fv bv fm =
let
@@ -14286,8 +14317,8 @@
(c2', s2', (c1,s1,fm,c2,s2) :: l)
end
- val (c1,_,fms) = foldl fwd (count0,empty,[]) fms
- val (c2,_,fms) = foldl bwd (count0,empty,[]) fms
+ val (c1,_,fms) = List.foldl fwd (count0,empty,[]) fms
+ val (c2,_,fms) = List.foldl bwd (count0,empty,[]) fms
(*MetisDebug
val _ = countEquivish c1 c2 orelse
@@ -14332,10 +14363,10 @@
let
val fms = sortMap (measure o count) countCompare fms
in
- foldl breakSet1 best (cumulatives fms)
- end
-
- val best = foldl breakSing best (cumulatives fms)
+ List.foldl breakSet1 best (cumulatives fms)
+ end
+
+ val best = List.foldl breakSing best (cumulatives fms)
val best = breakSet I best
val best = breakSet countNegate best
val best = breakSet countClauses best
@@ -14711,19 +14742,19 @@
(* Definitions. *)
(* ------------------------------------------------------------------------- *)
-val newDefinitionRelation =
- let
- val counter : int Unsynchronized.ref = Unsynchronized.ref 0
- in
- (* MODIFIED by Jasmin Blanchette *)
- fn () => CRITICAL (fn () =>
- let
- val Unsynchronized.ref i = counter
- val () = counter := i + 1
- in
- definitionPrefix ^ "_" ^ Int.toString i
- end)
- end;
+local
+ val counter : int Unsynchronized.ref = Unsynchronized.ref 0;
+
+ fun new () =
+ let
+ val Unsynchronized.ref i = counter
+ val () = counter := i + 1
+ in
+ definitionPrefix ^ "_" ^ Int.toString i
+ end;
+in
+ fun newDefinitionRelation () = Metis_Portable.critical new ();
+end;
fun newDefinition def =
let
@@ -14858,7 +14889,7 @@
end
-(**** Original file: Model.sig ****)
+(**** Original file: src/Model.sig ****)
(* ========================================================================= *)
(* RANDOM FINITE MODELS *)
@@ -15138,7 +15169,7 @@
end
-(**** Original file: Model.sml ****)
+(**** Original file: src/Model.sml ****)
(* ========================================================================= *)
(* RANDOM FINITE MODELS *)
@@ -15414,10 +15445,10 @@
fun ppEntry (tag,source_arity,target) =
Metis_Print.blockProgram Metis_Print.Inconsistent 2
- [Metis_Print.addString tag,
+ [Metis_Print.ppString tag,
Metis_Print.addBreak 1,
Metis_NameArity.pp source_arity,
- Metis_Print.addString " ->",
+ Metis_Print.ppString " ->",
Metis_Print.addBreak 1,
Metis_Name.pp target];
in
@@ -16332,7 +16363,7 @@
let
fun add (y,acc) = FunctionPerturbation (func_xs,y) :: acc
in
- foldl add acc target
+ List.foldl add acc target
end
in
pertTerms M onTarget tms xs acc
@@ -16413,17 +16444,17 @@
fun pp M =
Metis_Print.program
- [Metis_Print.addString "Metis_Model{",
+ [Metis_Print.ppString "Metis_Model{",
Metis_Print.ppInt (size M),
- Metis_Print.addString "}"];
-
-end
-
-(**** Original file: Problem.sig ****)
+ Metis_Print.ppString "}"];
+
+end
+
+(**** Original file: src/Problem.sig ****)
(* ========================================================================= *)
(* CNF PROBLEMS *)
-(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Problem =
@@ -16485,11 +16516,11 @@
end
-(**** Original file: Problem.sml ****)
+(**** Original file: src/Problem.sml ****)
(* ========================================================================= *)
(* CNF PROBLEMS *)
-(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Problem :> Metis_Problem =
@@ -16518,9 +16549,9 @@
val cls = toClauses prob
in
{clauses = length cls,
- literals = foldl lits 0 cls,
- symbols = foldl syms 0 cls,
- typedSymbols = foldl typedSyms 0 cls}
+ literals = List.foldl lits 0 cls,
+ symbols = List.foldl syms 0 cls,
+ typedSymbols = List.foldl typedSyms 0 cls}
end;
fun freeVars {axioms,conjecture} =
@@ -16648,11 +16679,11 @@
end
-(**** Original file: TermNet.sig ****)
+(**** Original file: src/TermNet.sig ****)
(* ========================================================================= *)
(* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC TERMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_TermNet =
@@ -16701,11 +16732,11 @@
end
-(**** Original file: TermNet.sml ****)
+(**** Original file: src/TermNet.sml ****)
(* ========================================================================= *)
(* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC TERMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_TermNet :> Metis_TermNet =
@@ -16863,7 +16894,7 @@
fun null net = size net = 0;
-fun singles qtms a = foldr Single a qtms;
+fun singles qtms a = List.foldr Single a qtms;
local
fun pre NONE = (0,NONE)
@@ -16893,7 +16924,7 @@
handle Error _ => raise Bug "Metis_TermNet.insert: should never fail";
end;
-fun fromList parm l = foldl (fn (tm_a,n) => insert n tm_a) (new parm) l;
+fun fromList parm l = List.foldl (fn (tm_a,n) => insert n tm_a) (new parm) l;
fun filter pred =
let
@@ -17146,7 +17177,7 @@
local
fun inc (qtm, Result l, acc) =
- foldl (fn ((n,a),acc) => (n,(qtm,a)) :: acc) acc l
+ List.foldl (fn ((n,a),acc) => (n,(qtm,a)) :: acc) acc l
| inc _ = raise Bug "Metis_TermNet.pp.inc";
fun toList (Net (_,_,NONE)) = []
@@ -17159,11 +17190,11 @@
end
-(**** Original file: AtomNet.sig ****)
+(**** Original file: src/AtomNet.sig ****)
(* ========================================================================= *)
(* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC ATOMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_AtomNet =
@@ -17210,11 +17241,11 @@
end
-(**** Original file: AtomNet.sml ****)
+(**** Original file: src/AtomNet.sml ****)
(* ========================================================================= *)
(* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC ATOMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_AtomNet :> Metis_AtomNet =
@@ -17249,7 +17280,7 @@
fun insert net (atm,a) = Metis_TermNet.insert net (atomToTerm atm, a);
-fun fromList parm l = foldl (fn (atm_a,n) => insert n atm_a) (new parm) l;
+fun fromList parm l = List.foldl (fn (atm_a,n) => insert n atm_a) (new parm) l;
val filter = Metis_TermNet.filter;
@@ -17272,11 +17303,11 @@
end
-(**** Original file: LiteralNet.sig ****)
+(**** Original file: src/LiteralNet.sig ****)
(* ========================================================================= *)
(* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC LITERALS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_LiteralNet =
@@ -17325,11 +17356,11 @@
end
-(**** Original file: LiteralNet.sml ****)
+(**** Original file: src/LiteralNet.sml ****)
(* ========================================================================= *)
(* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC LITERALS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_LiteralNet :> Metis_LiteralNet =
@@ -17368,7 +17399,7 @@
| insert {positive,negative} ((false,atm),a) =
{positive = positive, negative = Metis_AtomNet.insert negative (atm,a)};
-fun fromList parm l = foldl (fn (lit_a,n) => insert n lit_a) (new parm) l;
+fun fromList parm l = List.foldl (fn (lit_a,n) => insert n lit_a) (new parm) l;
fun filter pred {positive,negative} =
{positive = Metis_AtomNet.filter pred positive,
@@ -17402,11 +17433,11 @@
end
-(**** Original file: Subsume.sig ****)
+(**** Original file: src/Subsume.sig ****)
(* ========================================================================= *)
(* SUBSUMPTION CHECKING FOR FIRST ORDER LOGIC CLAUSES *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Subsume =
@@ -17456,11 +17487,11 @@
end
-(**** Original file: Subsume.sml ****)
+(**** Original file: src/Subsume.sml ****)
(* ========================================================================= *)
(* SUBSUMPTION CHECKING FOR FIRST ORDER LOGIC CLAUSES *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Subsume :> Metis_Subsume =
@@ -17793,11 +17824,11 @@
end
-(**** Original file: KnuthBendixOrder.sig ****)
+(**** Original file: src/KnuthBendixOrder.sig ****)
(* ========================================================================= *)
(* THE KNUTH-BENDIX TERM ORDERING *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_KnuthBendixOrder =
@@ -17818,11 +17849,11 @@
end
-(**** Original file: KnuthBendixOrder.sml ****)
+(**** Original file: src/KnuthBendixOrder.sml ****)
(* ========================================================================= *)
(* KNUTH-BENDIX TERM ORDERING CONSTRAINTS *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_KnuthBendixOrder :> Metis_KnuthBendixOrder =
@@ -17921,7 +17952,7 @@
val ppWeightList =
let
fun ppCoeff n =
- if n < 0 then Metis_Print.sequence (Metis_Print.addString "~") (ppCoeff (~n))
+ if n < 0 then Metis_Print.sequence (Metis_Print.ppString "~") (ppCoeff (~n))
else if n = 1 then Metis_Print.skip
else Metis_Print.ppInt n
@@ -18020,11 +18051,11 @@
end
-(**** Original file: Rewrite.sig ****)
+(**** Original file: src/Rewrite.sig ****)
(* ========================================================================= *)
(* ORDERED REWRITING FOR FIRST ORDER TERMS *)
-(* Copyright (c) 2003-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2003 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Rewrite =
@@ -18122,11 +18153,11 @@
end
-(**** Original file: Rewrite.sml ****)
+(**** Original file: src/Rewrite.sml ****)
(* ========================================================================= *)
(* ORDERED REWRITING FOR FIRST ORDER TERMS *)
-(* Copyright (c) 2003-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2003 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Rewrite :> Metis_Rewrite =
@@ -18333,7 +18364,7 @@
rw
end;
-val addList = foldl (fn (eqn,rw) => add rw eqn);
+val addList = List.foldl (fn (eqn,rw) => add rw eqn);
(* ------------------------------------------------------------------------- *)
(* Rewriting (the order must be a refinement of the rewrite order). *)
@@ -18394,7 +18425,7 @@
val th = Metis_Rule.symmetryRule l r
in
fn tm =>
- if Metis_Term.equal tm r then (l,th) else raise Error "mkNeqConv: Metis_RL"
+ if Metis_Term.equal tm r then (l,th) else raise Error "mkNeqConv: RL"
end
| SOME EQUAL => raise Error "irreflexive"
| SOME GREATER =>
@@ -18560,8 +18591,9 @@
let
fun addSubterm b ((path,tm),net) = Metis_TermNet.insert net (tm,(id,b,path))
- val subterms = foldl (addSubterm true) subterms (Metis_Term.subterms l)
- val subterms = foldl (addSubterm false) subterms (Metis_Term.subterms r)
+ val subterms = List.foldl (addSubterm true) subterms (Metis_Term.subterms l)
+
+ val subterms = List.foldl (addSubterm false) subterms (Metis_Term.subterms r)
in
subterms
end;
@@ -18786,7 +18818,7 @@
in
fun orderedRewrite order ths =
let
- val rw = foldl addEqn (new order) (enumerate ths)
+ val rw = List.foldl addEqn (new order) (enumerate ths)
in
rewriteRule rw order
end;
@@ -18796,11 +18828,11 @@
end
-(**** Original file: Units.sig ****)
+(**** Original file: src/Units.sig ****)
(* ========================================================================= *)
(* A STORE FOR UNIT THEOREMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Units =
@@ -18848,11 +18880,11 @@
end
-(**** Original file: Units.sml ****)
+(**** Original file: src/Units.sml ****)
(* ========================================================================= *)
(* A STORE FOR UNIT THEOREMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Units :> Metis_Units =
@@ -18899,7 +18931,7 @@
end
end;
-val addList = foldl (fn (th,u) => add u th);
+val addList = List.foldl (fn (th,u) => add u th);
(* ------------------------------------------------------------------------- *)
(* Matching. *)
@@ -18957,11 +18989,11 @@
end
-(**** Original file: Clause.sig ****)
+(**** Original file: src/Clause.sig ****)
(* ========================================================================= *)
(* CLAUSE = ID + THEOREM *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Clause =
@@ -19067,11 +19099,11 @@
end
-(**** Original file: Clause.sml ****)
+(**** Original file: src/Clause.sml ****)
(* ========================================================================= *)
(* CLAUSE = ID + THEOREM *)
-(* Copyright (c) 2002-2004 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Clause :> Metis_Clause =
@@ -19086,10 +19118,17 @@
val newId =
let
val r = Unsynchronized.ref 0
- in
- (* MODIFIED by Jasmin Blanchette *)
- fn () => CRITICAL (fn () =>
- case r of Unsynchronized.ref n => let val () = r := n + 1 in n end)
+
+ fun new () =
+ let
+ val Unsynchronized.ref n = r
+
+ val () = r := n + 1
+ in
+ n
+ end
+ in
+ fn () => Metis_Portable.critical new ()
end;
(* ------------------------------------------------------------------------- *)
@@ -19133,7 +19172,7 @@
val default : parameters =
{ordering = Metis_KnuthBendixOrder.default,
- orderLiterals = UnsignedLiteralOrder (* PositiveLiteralOrder *) (* MODIFIED by Jasmin Blanchette *),
+ orderLiterals = UnsignedLiteralOrder,
orderTerms = true};
fun mk info = Metis_Clause info
@@ -19252,7 +19291,7 @@
let
fun addTm ((path,tm),acc) = (lit,path,tm) :: acc
in
- foldl addTm acc (Metis_Literal.nonVarTypedSubterms lit)
+ List.foldl addTm acc (Metis_Literal.nonVarTypedSubterms lit)
end;
in
fun largestSubterms cl = Metis_LiteralSet.foldl addLit [] (largestLiterals cl);
@@ -19432,11 +19471,11 @@
end
-(**** Original file: Active.sig ****)
+(**** Original file: src/Active.sig ****)
(* ========================================================================= *)
(* THE ACTIVE SET OF CLAUSES *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Active =
@@ -19490,11 +19529,11 @@
end
-(**** Original file: Active.sml ****)
+(**** Original file: src/Active.sml ****)
(* ========================================================================= *)
(* THE ACTIVE SET OF CLAUSES *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Active :> Metis_Active =
@@ -19519,7 +19558,7 @@
| NONE => rw
end
in
- foldl add (Metis_Rewrite.new (Metis_KnuthBendixOrder.compare ordering))
+ List.foldl add (Metis_Rewrite.new (Metis_KnuthBendixOrder.compare ordering))
end;
fun allFactors red =
@@ -19787,8 +19826,8 @@
end
val cls = clauses active
- val (cls,_) = foldl remove ([], Metis_Subsume.new ()) cls
- val (cls,subs) = foldl remove ([], Metis_Subsume.new ()) cls
+ val (cls,_) = List.foldl remove ([], Metis_Subsume.new ()) cls
+ val (cls,subs) = List.foldl remove ([], Metis_Subsume.new ()) cls
(*MetisDebug
val Metis_Active {parameters,...} = active
@@ -19815,7 +19854,7 @@
local
fun ppField f ppA a =
Metis_Print.blockProgram Metis_Print.Inconsistent 2
- [Metis_Print.addString (f ^ " ="),
+ [Metis_Print.ppString (f ^ " ="),
Metis_Print.addBreak 1,
ppA a];
@@ -19836,21 +19875,21 @@
in
fun pp (Metis_Active {clauses,rewrite,subterms,...}) =
Metis_Print.blockProgram Metis_Print.Inconsistent 2
- [Metis_Print.addString "Metis_Active",
+ [Metis_Print.ppString "Metis_Active",
Metis_Print.addBreak 1,
Metis_Print.blockProgram Metis_Print.Inconsistent 1
- [Metis_Print.addString "{",
+ [Metis_Print.ppString "{",
ppClauses clauses,
- Metis_Print.addString ",",
+ Metis_Print.ppString ",",
Metis_Print.addBreak 1,
ppRewrite rewrite,
(*MetisTrace5
- Metis_Print.addString ",",
+ Metis_Print.ppString ",",
Metis_Print.addBreak 1,
ppSubterms subterms,
*)
Metis_Print.skip],
- Metis_Print.addString "}"];
+ Metis_Print.ppString "}"];
end;
*)
@@ -19968,7 +20007,7 @@
fun add ((lit,ort,tm),equations) =
Metis_TermNet.insert equations (tm,(cl,lit,ort,tm))
in
- foldl add equations (Metis_Clause.largestEquations cl)
+ List.foldl add equations (Metis_Clause.largestEquations cl)
end;
fun addSubterms subterms cl =
@@ -19976,7 +20015,7 @@
fun add ((lit,path,tm),subterms) =
Metis_TermNet.insert subterms (tm,(cl,lit,path,tm))
in
- foldl add subterms (Metis_Clause.largestSubterms cl)
+ List.foldl add subterms (Metis_Clause.largestSubterms cl)
end;
fun addAllSubterms allSubterms cl =
@@ -19984,7 +20023,7 @@
fun add ((_,_,tm),allSubterms) =
Metis_TermNet.insert allSubterms (tm,(cl,tm))
in
- foldl add allSubterms (Metis_Clause.allSubterms cl)
+ List.foldl add allSubterms (Metis_Clause.allSubterms cl)
end;
fun addClause active cl =
@@ -20035,7 +20074,8 @@
*)
in
if Metis_Atom.isEq atm then acc
- else foldl resolve acc (Metis_LiteralNet.unify literals (Metis_Literal.negate lit))
+ else
+ List.foldl resolve acc (Metis_LiteralNet.unify literals (Metis_Literal.negate lit))
end;
fun deduceParamodulationWith subterms cl ((lit,ort,tm),acc) =
@@ -20045,7 +20085,7 @@
SOME cl' => cl' :: acc
| NONE => acc
in
- foldl para acc (Metis_TermNet.unify subterms tm)
+ List.foldl para acc (Metis_TermNet.unify subterms tm)
end;
fun deduceParamodulationInto equations cl ((lit,path,tm),acc) =
@@ -20055,7 +20095,7 @@
SOME cl' => cl' :: acc
| NONE => acc
in
- foldl para acc (Metis_TermNet.unify equations tm)
+ List.foldl para acc (Metis_TermNet.unify equations tm)
end;
fun deduce active cl =
@@ -20081,8 +20121,8 @@
val acc = []
val acc = Metis_LiteralSet.foldl (deduceResolution literals cl) acc lits
- val acc = foldl (deduceParamodulationWith subterms cl) acc eqns
- val acc = foldl (deduceParamodulationInto equations cl) acc subtms
+ val acc = List.foldl (deduceParamodulationWith subterms cl) acc eqns
+ val acc = List.foldl (deduceParamodulationInto equations cl) acc subtms
val acc = rev acc
(*MetisTrace5
@@ -20328,7 +20368,7 @@
let
val cls = sort_utilitywise (cl :: Metis_Clause.factor cl)
in
- foldl factor_add active_subsume_acc cls
+ List.foldl factor_add active_subsume_acc cls
end;
fun factor' active acc [] = (active, rev acc)
@@ -20336,7 +20376,7 @@
let
val cls = sort_utilitywise cls
val subsume = getSubsume active
- val (active,_,acc) = foldl factor1 (active,subsume,acc) cls
+ val (active,_,acc) = List.foldl factor1 (active,subsume,acc) cls
val (active,cls) = extract_rewritables active
in
factor' active acc cls
@@ -20404,11 +20444,11 @@
end
-(**** Original file: Waiting.sig ****)
+(**** Original file: src/Waiting.sig ****)
(* ========================================================================= *)
(* THE WAITING SET OF CLAUSES *)
-(* Copyright (c) 2002-2007 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Waiting =
@@ -20484,11 +20524,11 @@
end
-(**** Original file: Waiting.sml ****)
+(**** Original file: src/Waiting.sml ****)
(* ========================================================================= *)
(* THE WAITING SET OF CLAUSES *)
-(* Copyright (c) 2002-2007 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Waiting :> Metis_Waiting =
@@ -20528,17 +20568,16 @@
(* ------------------------------------------------------------------------- *)
val defaultModels : modelParameters list =
- [(* MODIFIED by Jasmin Blanchette
- {model = Metis_Model.default,
+ [{model = Metis_Model.default,
initialPerturbations = 100,
maxChecks = SOME 20,
perturbations = 0,
- weight = 1.0} *)];
+ weight = 1.0}];
val default : parameters =
{symbolsWeight = 1.0,
- literalsWeight = (* 1.0 *) 0.0, (* MODIFIED by Jasmin Blanchette *)
- variablesWeight = (* 1.0 *) 0.0, (* MODIFIED by Jasmin Blanchette *)
+ literalsWeight = 1.0,
+ variablesWeight = 1.0,
models = defaultModels};
fun size (Metis_Waiting {clauses,...}) = Metis_Heap.size clauses;
@@ -20651,7 +20690,7 @@
val symbolsW = Math.pow (clauseSymbols lits, symbolsWeight)
val variablesW = Math.pow (clauseVariables lits, variablesWeight)
val literalsW = Math.pow (clauseLiterals lits, literalsWeight)
- val modelsW = (* checkModels models mods mcl *) 1.0 (* MODIFIED by Jasmin Blanchette *)
+ val modelsW = checkModels models mods mcl
(*MetisTrace4
val () = trace ("Metis_Waiting.clauseWeight: dist = " ^
Real.toString dist ^ "\n")
@@ -20758,11 +20797,11 @@
end
-(**** Original file: Resolution.sig ****)
+(**** Original file: src/Resolution.sig ****)
(* ========================================================================= *)
(* THE RESOLUTION PROOF PROCEDURE *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Metis_Resolution =
@@ -20812,11 +20851,11 @@
end
-(**** Original file: Resolution.sml ****)
+(**** Original file: src/Resolution.sml ****)
(* ========================================================================= *)
(* THE RESOLUTION PROOF PROCEDURE *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Metis_Resolution :> Metis_Resolution =
--- a/src/Tools/Metis/scripts/mlpp Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/scripts/mlpp Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
#!/usr/bin/perl
-# Copyright (c) 2006 Joe Hurd, All Rights Reserved
+# Copyright (c) 2006 Joe Hurd, distributed under the BSD License
use strict;
use warnings;
@@ -78,6 +78,8 @@
$text =~ s/Array\.modifyi/Array_modifyi/g;
$text =~ s/OS\.Process\.isSuccess/OS_Process_isSuccess/g;
$text =~ s/PP\.ppstream/ppstream/g;
+ $text =~ s/String\.concatWith/String_concatWith/g;
+ $text =~ s/String\.isSubstring/String_isSubstring/g;
$text =~ s/String\.isSuffix/String_isSuffix/g;
$text =~ s/Substring\.full/Substring_full/g;
$text =~ s/TextIO\.inputLine/TextIO_inputLine/g;
--- a/src/Tools/Metis/src/Active.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Active.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* THE ACTIVE SET OF CLAUSES *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Active =
--- a/src/Tools/Metis/src/Active.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Active.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* THE ACTIVE SET OF CLAUSES *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Active :> Active =
@@ -25,7 +25,7 @@
| NONE => rw
end
in
- foldl add (Rewrite.new (KnuthBendixOrder.compare ordering))
+ List.foldl add (Rewrite.new (KnuthBendixOrder.compare ordering))
end;
fun allFactors red =
@@ -293,8 +293,8 @@
end
val cls = clauses active
- val (cls,_) = foldl remove ([], Subsume.new ()) cls
- val (cls,subs) = foldl remove ([], Subsume.new ()) cls
+ val (cls,_) = List.foldl remove ([], Subsume.new ()) cls
+ val (cls,subs) = List.foldl remove ([], Subsume.new ()) cls
(*MetisDebug
val Active {parameters,...} = active
@@ -321,7 +321,7 @@
local
fun ppField f ppA a =
Print.blockProgram Print.Inconsistent 2
- [Print.addString (f ^ " ="),
+ [Print.ppString (f ^ " ="),
Print.addBreak 1,
ppA a];
@@ -342,21 +342,21 @@
in
fun pp (Active {clauses,rewrite,subterms,...}) =
Print.blockProgram Print.Inconsistent 2
- [Print.addString "Active",
+ [Print.ppString "Active",
Print.addBreak 1,
Print.blockProgram Print.Inconsistent 1
- [Print.addString "{",
+ [Print.ppString "{",
ppClauses clauses,
- Print.addString ",",
+ Print.ppString ",",
Print.addBreak 1,
ppRewrite rewrite,
(*MetisTrace5
- Print.addString ",",
+ Print.ppString ",",
Print.addBreak 1,
ppSubterms subterms,
*)
Print.skip],
- Print.addString "}"];
+ Print.ppString "}"];
end;
*)
@@ -474,7 +474,7 @@
fun add ((lit,ort,tm),equations) =
TermNet.insert equations (tm,(cl,lit,ort,tm))
in
- foldl add equations (Clause.largestEquations cl)
+ List.foldl add equations (Clause.largestEquations cl)
end;
fun addSubterms subterms cl =
@@ -482,7 +482,7 @@
fun add ((lit,path,tm),subterms) =
TermNet.insert subterms (tm,(cl,lit,path,tm))
in
- foldl add subterms (Clause.largestSubterms cl)
+ List.foldl add subterms (Clause.largestSubterms cl)
end;
fun addAllSubterms allSubterms cl =
@@ -490,7 +490,7 @@
fun add ((_,_,tm),allSubterms) =
TermNet.insert allSubterms (tm,(cl,tm))
in
- foldl add allSubterms (Clause.allSubterms cl)
+ List.foldl add allSubterms (Clause.allSubterms cl)
end;
fun addClause active cl =
@@ -541,7 +541,8 @@
*)
in
if Atom.isEq atm then acc
- else foldl resolve acc (LiteralNet.unify literals (Literal.negate lit))
+ else
+ List.foldl resolve acc (LiteralNet.unify literals (Literal.negate lit))
end;
fun deduceParamodulationWith subterms cl ((lit,ort,tm),acc) =
@@ -551,7 +552,7 @@
SOME cl' => cl' :: acc
| NONE => acc
in
- foldl para acc (TermNet.unify subterms tm)
+ List.foldl para acc (TermNet.unify subterms tm)
end;
fun deduceParamodulationInto equations cl ((lit,path,tm),acc) =
@@ -561,7 +562,7 @@
SOME cl' => cl' :: acc
| NONE => acc
in
- foldl para acc (TermNet.unify equations tm)
+ List.foldl para acc (TermNet.unify equations tm)
end;
fun deduce active cl =
@@ -587,8 +588,8 @@
val acc = []
val acc = LiteralSet.foldl (deduceResolution literals cl) acc lits
- val acc = foldl (deduceParamodulationWith subterms cl) acc eqns
- val acc = foldl (deduceParamodulationInto equations cl) acc subtms
+ val acc = List.foldl (deduceParamodulationWith subterms cl) acc eqns
+ val acc = List.foldl (deduceParamodulationInto equations cl) acc subtms
val acc = rev acc
(*MetisTrace5
@@ -834,7 +835,7 @@
let
val cls = sort_utilitywise (cl :: Clause.factor cl)
in
- foldl factor_add active_subsume_acc cls
+ List.foldl factor_add active_subsume_acc cls
end;
fun factor' active acc [] = (active, rev acc)
@@ -842,7 +843,7 @@
let
val cls = sort_utilitywise cls
val subsume = getSubsume active
- val (active,_,acc) = foldl factor1 (active,subsume,acc) cls
+ val (active,_,acc) = List.foldl factor1 (active,subsume,acc) cls
val (active,cls) = extract_rewritables active
in
factor' active acc cls
--- a/src/Tools/Metis/src/Atom.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Atom.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* FIRST ORDER LOGIC ATOMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Atom =
--- a/src/Tools/Metis/src/Atom.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Atom.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* FIRST ORDER LOGIC ATOMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Atom :> Atom =
@@ -34,14 +34,14 @@
let
fun f (tm,acc) = NameAritySet.union (Term.functions tm) acc
in
- fn atm => foldl f NameAritySet.empty (arguments atm)
+ fn atm => List.foldl f NameAritySet.empty (arguments atm)
end;
val functionNames =
let
fun f (tm,acc) = NameSet.union (Term.functionNames tm) acc
in
- fn atm => foldl f NameSet.empty (arguments atm)
+ fn atm => List.foldl f NameSet.empty (arguments atm)
end;
(* Binary relations *)
@@ -58,7 +58,8 @@
(* The size of an atom in symbols. *)
(* ------------------------------------------------------------------------- *)
-fun symbols atm = foldl (fn (tm,z) => Term.symbols tm + z) 1 (arguments atm);
+fun symbols atm =
+ List.foldl (fn (tm,z) => Term.symbols tm + z) 1 (arguments atm);
(* ------------------------------------------------------------------------- *)
(* A total comparison function for atoms. *)
@@ -85,7 +86,7 @@
let
fun f ((n,tm),l) = map (fn (p,s) => (n :: p, s)) (Term.subterms tm) @ l
in
- foldl f [] (enumerate tms)
+ List.foldl f [] (enumerate tms)
end;
fun replace _ ([],_) = raise Bug "Atom.replace: empty path"
@@ -120,7 +121,7 @@
let
fun f (tm,acc) = NameSet.union (Term.freeVars tm) acc
in
- fn atm => foldl f NameSet.empty (arguments atm)
+ fn atm => List.foldl f NameSet.empty (arguments atm)
end;
(* ------------------------------------------------------------------------- *)
@@ -146,7 +147,7 @@
val _ = (Name.equal p1 p2 andalso length tms1 = length tms2) orelse
raise Error "Atom.match"
in
- foldl matchArg sub (zip tms1 tms2)
+ List.foldl matchArg sub (zip tms1 tms2)
end;
end;
@@ -162,7 +163,7 @@
val _ = (Name.equal p1 p2 andalso length tms1 = length tms2) orelse
raise Error "Atom.unify"
in
- foldl unifyArg sub (zip tms1 tms2)
+ List.foldl unifyArg sub (zip tms1 tms2)
end;
end;
@@ -211,7 +212,7 @@
(* ------------------------------------------------------------------------- *)
fun typedSymbols ((_,tms) : atom) =
- foldl (fn (tm,z) => Term.typedSymbols tm + z) 1 tms;
+ List.foldl (fn (tm,z) => Term.typedSymbols tm + z) 1 tms;
fun nonVarTypedSubterms (_,tms) =
let
@@ -219,10 +220,10 @@
let
fun addTm ((path,tm),acc) = (n :: path, tm) :: acc
in
- foldl addTm acc (Term.nonVarTypedSubterms arg)
+ List.foldl addTm acc (Term.nonVarTypedSubterms arg)
end
in
- foldl addArg [] (enumerate tms)
+ List.foldl addArg [] (enumerate tms)
end;
(* ------------------------------------------------------------------------- *)
--- a/src/Tools/Metis/src/AtomNet.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/AtomNet.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC ATOMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature AtomNet =
--- a/src/Tools/Metis/src/AtomNet.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/AtomNet.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC ATOMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure AtomNet :> AtomNet =
@@ -35,7 +35,7 @@
fun insert net (atm,a) = TermNet.insert net (atomToTerm atm, a);
-fun fromList parm l = foldl (fn (atm_a,n) => insert n atm_a) (new parm) l;
+fun fromList parm l = List.foldl (fn (atm_a,n) => insert n atm_a) (new parm) l;
val filter = TermNet.filter;
--- a/src/Tools/Metis/src/Clause.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Clause.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* CLAUSE = ID + THEOREM *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Clause =
--- a/src/Tools/Metis/src/Clause.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Clause.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* CLAUSE = ID + THEOREM *)
-(* Copyright (c) 2002-2004 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Clause :> Clause =
@@ -15,10 +15,17 @@
val newId =
let
val r = ref 0
+
+ fun new () =
+ let
+ val ref n = r
+
+ val () = r := n + 1
+ in
+ n
+ end
in
- (* MODIFIED by Jasmin Blanchette *)
- fn () => CRITICAL (fn () =>
- case r of ref n => let val () = r := n + 1 in n end)
+ fn () => Portable.critical new ()
end;
(* ------------------------------------------------------------------------- *)
@@ -62,7 +69,7 @@
val default : parameters =
{ordering = KnuthBendixOrder.default,
- orderLiterals = UnsignedLiteralOrder (* PositiveLiteralOrder *) (* MODIFIED by Jasmin Blanchette *),
+ orderLiterals = UnsignedLiteralOrder,
orderTerms = true};
fun mk info = Clause info
@@ -181,7 +188,7 @@
let
fun addTm ((path,tm),acc) = (lit,path,tm) :: acc
in
- foldl addTm acc (Literal.nonVarTypedSubterms lit)
+ List.foldl addTm acc (Literal.nonVarTypedSubterms lit)
end;
in
fun largestSubterms cl = LiteralSet.foldl addLit [] (largestLiterals cl);
--- a/src/Tools/Metis/src/Formula.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Formula.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* FIRST ORDER LOGIC FORMULAS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Formula =
--- a/src/Tools/Metis/src/Formula.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Formula.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* FIRST ORDER LOGIC FORMULAS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Formula :> Formula =
@@ -145,7 +145,7 @@
(* Conjunctions *)
fun listMkConj fms =
- case rev fms of [] => True | fm :: fms => foldl And fm fms;
+ case rev fms of [] => True | fm :: fms => List.foldl And fm fms;
local
fun strip cs (And (p,q)) = strip (p :: cs) q
@@ -168,7 +168,7 @@
(* Disjunctions *)
fun listMkDisj fms =
- case rev fms of [] => False | fm :: fms => foldl Or fm fms;
+ case rev fms of [] => False | fm :: fms => List.foldl Or fm fms;
local
fun strip cs (Or (p,q)) = strip (p :: cs) q
@@ -191,7 +191,7 @@
(* Equivalences *)
fun listMkEquiv fms =
- case rev fms of [] => True | fm :: fms => foldl Iff fm fms;
+ case rev fms of [] => True | fm :: fms => List.foldl Iff fm fms;
local
fun strip cs (Iff (p,q)) = strip (p :: cs) q
--- a/src/Tools/Metis/src/Heap.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Heap.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* A HEAP DATATYPE FOR ML *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Heap =
--- a/src/Tools/Metis/src/Heap.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Heap.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* A HEAP DATATYPE FOR ML *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Heap :> Heap =
--- a/src/Tools/Metis/src/KeyMap.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/KeyMap.sml Fri Sep 17 17:31:20 2010 +0200
@@ -887,38 +887,40 @@
(* ------------------------------------------------------------------------- *)
datatype 'value iterator =
- LR of (key * 'value) * 'value tree * 'value node list
- | RL of (key * 'value) * 'value tree * 'value node list;
+ LeftToRightIterator of
+ (key * 'value) * 'value tree * 'value node list
+ | RightToLeftIterator of
+ (key * 'value) * 'value tree * 'value node list;
-fun fromSpineLR nodes =
+fun fromSpineLeftToRightIterator nodes =
case nodes of
[] => NONE
| Node {key,value,right,...} :: nodes =>
- SOME (LR ((key,value),right,nodes));
+ SOME (LeftToRightIterator ((key,value),right,nodes));
-fun fromSpineRL nodes =
+fun fromSpineRightToLeftIterator nodes =
case nodes of
[] => NONE
| Node {key,value,left,...} :: nodes =>
- SOME (RL ((key,value),left,nodes));
+ SOME (RightToLeftIterator ((key,value),left,nodes));
-fun addLR nodes tree = fromSpineLR (treeLeftSpine nodes tree);
+fun addLeftToRightIterator nodes tree = fromSpineLeftToRightIterator (treeLeftSpine nodes tree);
-fun addRL nodes tree = fromSpineRL (treeRightSpine nodes tree);
+fun addRightToLeftIterator nodes tree = fromSpineRightToLeftIterator (treeRightSpine nodes tree);
-fun treeMkIterator tree = addLR [] tree;
+fun treeMkIterator tree = addLeftToRightIterator [] tree;
-fun treeMkRevIterator tree = addRL [] tree;
+fun treeMkRevIterator tree = addRightToLeftIterator [] tree;
fun readIterator iter =
case iter of
- LR (key_value,_,_) => key_value
- | RL (key_value,_,_) => key_value;
+ LeftToRightIterator (key_value,_,_) => key_value
+ | RightToLeftIterator (key_value,_,_) => key_value;
fun advanceIterator iter =
case iter of
- LR (_,tree,nodes) => addLR nodes tree
- | RL (_,tree,nodes) => addRL nodes tree;
+ LeftToRightIterator (_,tree,nodes) => addLeftToRightIterator nodes tree
+ | RightToLeftIterator (_,tree,nodes) => addRightToLeftIterator nodes tree;
fun foldIterator f acc io =
case io of
--- a/src/Tools/Metis/src/KnuthBendixOrder.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/KnuthBendixOrder.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* THE KNUTH-BENDIX TERM ORDERING *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature KnuthBendixOrder =
--- a/src/Tools/Metis/src/KnuthBendixOrder.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/KnuthBendixOrder.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* KNUTH-BENDIX TERM ORDERING CONSTRAINTS *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure KnuthBendixOrder :> KnuthBendixOrder =
@@ -99,7 +99,7 @@
val ppWeightList =
let
fun ppCoeff n =
- if n < 0 then Print.sequence (Print.addString "~") (ppCoeff (~n))
+ if n < 0 then Print.sequence (Print.ppString "~") (ppCoeff (~n))
else if n = 1 then Print.skip
else Print.ppInt n
--- a/src/Tools/Metis/src/Literal.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Literal.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* FIRST ORDER LOGIC LITERALS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Literal =
--- a/src/Tools/Metis/src/Literal.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Literal.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* FIRST ORDER LOGIC LITERALS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Literal :> Literal =
--- a/src/Tools/Metis/src/LiteralNet.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/LiteralNet.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC LITERALS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature LiteralNet =
--- a/src/Tools/Metis/src/LiteralNet.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/LiteralNet.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC LITERALS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure LiteralNet :> LiteralNet =
@@ -39,7 +39,7 @@
| insert {positive,negative} ((false,atm),a) =
{positive = positive, negative = AtomNet.insert negative (atm,a)};
-fun fromList parm l = foldl (fn (lit_a,n) => insert n lit_a) (new parm) l;
+fun fromList parm l = List.foldl (fn (lit_a,n) => insert n lit_a) (new parm) l;
fun filter pred {positive,negative} =
{positive = AtomNet.filter pred positive,
--- a/src/Tools/Metis/src/Map.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Map.sml Fri Sep 17 17:31:20 2010 +0200
@@ -879,38 +879,40 @@
(* ------------------------------------------------------------------------- *)
datatype ('key,'value) iterator =
- LR of ('key * 'value) * ('key,'value) tree * ('key,'value) node list
- | RL of ('key * 'value) * ('key,'value) tree * ('key,'value) node list;
+ LeftToRightIterator of
+ ('key * 'value) * ('key,'value) tree * ('key,'value) node list
+ | RightToLeftIterator of
+ ('key * 'value) * ('key,'value) tree * ('key,'value) node list;
-fun fromSpineLR nodes =
+fun fromSpineLeftToRightIterator nodes =
case nodes of
[] => NONE
| Node {key,value,right,...} :: nodes =>
- SOME (LR ((key,value),right,nodes));
+ SOME (LeftToRightIterator ((key,value),right,nodes));
-fun fromSpineRL nodes =
+fun fromSpineRightToLeftIterator nodes =
case nodes of
[] => NONE
| Node {key,value,left,...} :: nodes =>
- SOME (RL ((key,value),left,nodes));
+ SOME (RightToLeftIterator ((key,value),left,nodes));
-fun addLR nodes tree = fromSpineLR (treeLeftSpine nodes tree);
+fun addLeftToRightIterator nodes tree = fromSpineLeftToRightIterator (treeLeftSpine nodes tree);
-fun addRL nodes tree = fromSpineRL (treeRightSpine nodes tree);
+fun addRightToLeftIterator nodes tree = fromSpineRightToLeftIterator (treeRightSpine nodes tree);
-fun treeMkIterator tree = addLR [] tree;
+fun treeMkIterator tree = addLeftToRightIterator [] tree;
-fun treeMkRevIterator tree = addRL [] tree;
+fun treeMkRevIterator tree = addRightToLeftIterator [] tree;
fun readIterator iter =
case iter of
- LR (key_value,_,_) => key_value
- | RL (key_value,_,_) => key_value;
+ LeftToRightIterator (key_value,_,_) => key_value
+ | RightToLeftIterator (key_value,_,_) => key_value;
fun advanceIterator iter =
case iter of
- LR (_,tree,nodes) => addLR nodes tree
- | RL (_,tree,nodes) => addRL nodes tree;
+ LeftToRightIterator (_,tree,nodes) => addLeftToRightIterator nodes tree
+ | RightToLeftIterator (_,tree,nodes) => addRightToLeftIterator nodes tree;
fun foldIterator f acc io =
case io of
--- a/src/Tools/Metis/src/Model.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Model.sml Fri Sep 17 17:31:20 2010 +0200
@@ -272,10 +272,10 @@
fun ppEntry (tag,source_arity,target) =
Print.blockProgram Print.Inconsistent 2
- [Print.addString tag,
+ [Print.ppString tag,
Print.addBreak 1,
NameArity.pp source_arity,
- Print.addString " ->",
+ Print.ppString " ->",
Print.addBreak 1,
Name.pp target];
in
@@ -1190,7 +1190,7 @@
let
fun add (y,acc) = FunctionPerturbation (func_xs,y) :: acc
in
- foldl add acc target
+ List.foldl add acc target
end
in
pertTerms M onTarget tms xs acc
@@ -1271,8 +1271,8 @@
fun pp M =
Print.program
- [Print.addString "Model{",
+ [Print.ppString "Model{",
Print.ppInt (size M),
- Print.addString "}"];
+ Print.ppString "}"];
end
--- a/src/Tools/Metis/src/Name.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Name.sml Fri Sep 17 17:31:20 2010 +0200
@@ -84,15 +84,15 @@
structure NameSet =
struct
- local
- structure S = ElementSet (NameMap);
- in
- open S;
- end;
+local
+ structure S = ElementSet (NameMap);
+in
+ open S;
+end;
- val pp =
- Print.ppMap
- toList
- (Print.ppBracket "{" "}" (Print.ppOpList "," Name.pp));
+val pp =
+ Print.ppMap
+ toList
+ (Print.ppBracket "{" "}" (Print.ppOpList "," Name.pp));
end
--- a/src/Tools/Metis/src/NameArity.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/NameArity.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* NAME/ARITY PAIRS *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature NameArity =
--- a/src/Tools/Metis/src/NameArity.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/NameArity.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* NAME/ARITY PAIRS *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure NameArity :> NameArity =
@@ -46,7 +46,7 @@
fun pp (n,i) =
Print.blockProgram Print.Inconsistent 0
[Name.pp n,
- Print.addString "/",
+ Print.ppString "/",
Print.ppInt i];
end
@@ -57,35 +57,36 @@
structure NameArityMap =
struct
- local
- structure S = KeyMap (NameArityOrdered);
- in
- open S;
- end;
+local
+ structure S = KeyMap (NameArityOrdered);
+in
+ open S;
+end;
- fun compose m1 m2 =
- let
- fun pk ((_,a),n) = peek m2 (n,a)
- in
- mapPartial pk m1
- end;
+fun compose m1 m2 =
+ let
+ fun pk ((_,a),n) = peek m2 (n,a)
+ in
+ mapPartial pk m1
+ end;
end
structure NameAritySet =
struct
- local
- structure S = ElementSet (NameArityMap);
- in
- open S;
- end;
+local
+ structure S = ElementSet (NameArityMap);
+in
+ open S;
+end;
- val allNullary = all NameArity.nullary;
+val allNullary = all NameArity.nullary;
- val pp =
- Print.ppMap
- toList
- (Print.ppBracket "{" "}" (Print.ppOpList "," NameArity.pp));
+val pp =
+ Print.ppMap
+ toList
+ (Print.ppBracket "{" "}" (Print.ppOpList "," NameArity.pp));
+
end
--- a/src/Tools/Metis/src/Normalize.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Normalize.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* NORMALIZING FORMULAS *)
-(* Copyright (c) 2001-2009 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Normalize =
--- a/src/Tools/Metis/src/Normalize.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Normalize.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* NORMALIZING FORMULAS *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Normalize :> Normalize =
@@ -685,23 +685,23 @@
(* Basic conjunctive normal form. *)
(* ------------------------------------------------------------------------- *)
-val newSkolemFunction =
- let
- val counter : int StringMap.map ref = ref (StringMap.new ())
- in
- (* MODIFIED by Jasmin Blanchette *)
- fn n => CRITICAL (fn () =>
- let
- val ref m = counter
- val s = Name.toString n
- val i = Option.getOpt (StringMap.peek m s, 0)
- val () = counter := StringMap.insert m (s, i + 1)
- val i = if i = 0 then "" else "_" ^ Int.toString i
- val s = skolemPrefix ^ "_" ^ s ^ i
- in
- Name.fromString s
- end)
- end;
+local
+ val counter : int StringMap.map ref = ref (StringMap.new ());
+
+ fun new n () =
+ let
+ val ref m = counter
+ val s = Name.toString n
+ val i = Option.getOpt (StringMap.peek m s, 0)
+ val () = counter := StringMap.insert m (s, i + 1)
+ val i = if i = 0 then "" else "_" ^ Int.toString i
+ val s = skolemPrefix ^ "_" ^ s ^ i
+ in
+ Name.fromString s
+ end;
+in
+ fun newSkolemFunction n = Portable.critical (new n) ();
+end;
fun skolemize fv bv fm =
let
@@ -815,8 +815,8 @@
(c2', s2', (c1,s1,fm,c2,s2) :: l)
end
- val (c1,_,fms) = foldl fwd (count0,empty,[]) fms
- val (c2,_,fms) = foldl bwd (count0,empty,[]) fms
+ val (c1,_,fms) = List.foldl fwd (count0,empty,[]) fms
+ val (c2,_,fms) = List.foldl bwd (count0,empty,[]) fms
(*MetisDebug
val _ = countEquivish c1 c2 orelse
@@ -861,10 +861,10 @@
let
val fms = sortMap (measure o count) countCompare fms
in
- foldl breakSet1 best (cumulatives fms)
+ List.foldl breakSet1 best (cumulatives fms)
end
- val best = foldl breakSing best (cumulatives fms)
+ val best = List.foldl breakSing best (cumulatives fms)
val best = breakSet I best
val best = breakSet countNegate best
val best = breakSet countClauses best
@@ -1240,19 +1240,19 @@
(* Definitions. *)
(* ------------------------------------------------------------------------- *)
-val newDefinitionRelation =
- let
- val counter : int ref = ref 0
- in
- (* MODIFIED by Jasmin Blanchette *)
- fn () => CRITICAL (fn () =>
- let
- val ref i = counter
- val () = counter := i + 1
- in
- definitionPrefix ^ "_" ^ Int.toString i
- end)
- end;
+local
+ val counter : int ref = ref 0;
+
+ fun new () =
+ let
+ val ref i = counter
+ val () = counter := i + 1
+ in
+ definitionPrefix ^ "_" ^ Int.toString i
+ end;
+in
+ fun newDefinitionRelation () = Portable.critical new ();
+end;
fun newDefinition def =
let
--- a/src/Tools/Metis/src/Options.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Options.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* PROCESSING COMMAND LINE OPTIONS *)
-(* Copyright (c) 2003-2004 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2003 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Options =
--- a/src/Tools/Metis/src/Options.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Options.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* PROCESSING COMMAND LINE OPTIONS *)
-(* Copyright (c) 2003-2004 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2003 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Options :> Options =
@@ -146,7 +146,7 @@
fun indent (s, "" :: l) = indent (s ^ " ", l) | indent x = x
val (res,n) = indent (" ",n)
val res = res ^ join ", " n
- val res = foldl (fn (x,y) => y ^ " " ^ x) res r
+ val res = List.foldl (fn (x,y) => y ^ " " ^ x) res r
in
[res ^ " ...", " " ^ s]
end
@@ -185,7 +185,7 @@
exit allopts {message = SOME mesg, usage = true, success = false};
fun version allopts =
- (print (versionInformation allopts);
+ (TextIO.print (versionInformation allopts);
exit allopts {message = NONE, usage = false, success = true});
(* ------------------------------------------------------------------------- *)
@@ -220,7 +220,8 @@
| process ("-v" :: _) = version allopts
| process ("--version" :: _) = version allopts
| process (x :: xs) =
- if x = "" orelse x = "-" orelse hd (explode x) <> #"-" then ([], x :: xs)
+ if x = "" orelse x = "-" orelse hd (String.explode x) <> #"-" then
+ ([], x :: xs)
else
let
val (r,f) = findOption x
@@ -233,7 +234,8 @@
fn l =>
let
val (a,b) = process l
- val a = foldl (fn ((x,xs),ys) => x :: xs @ ys) [] (rev a)
+
+ val a = List.foldl (fn ((x,xs),ys) => x :: xs @ ys) [] (rev a)
in
(a,b)
end
--- a/src/Tools/Metis/src/Ordered.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Ordered.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* ORDERED TYPES *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Ordered =
--- a/src/Tools/Metis/src/Ordered.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Ordered.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* ORDERED TYPES *)
-(* Copyright (c) 2004-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2004 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure IntOrdered =
--- a/src/Tools/Metis/src/Parse.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Parse.sig Fri Sep 17 17:31:20 2010 +0200
@@ -99,8 +99,9 @@
(* ------------------------------------------------------------------------- *)
val parseInfixes :
- Print.infixes -> (string * 'a * 'a -> 'a) -> (string,'a) parser ->
- (string,'a) parser
+ Print.infixes ->
+ (Print.token * 'a * 'a -> 'a) -> ('b,Print.token) parser ->
+ ('b,'a) parser -> ('b,'a) parser
(* ------------------------------------------------------------------------- *)
(* Quotations. *)
--- a/src/Tools/Metis/src/Parse.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Parse.sml Fri Sep 17 17:31:20 2010 +0200
@@ -134,7 +134,7 @@
val ref (n,_,l2,l3) = lastLine
val () = lastLine := (n + 1, l2, l3, line)
in
- explode line
+ String.explode line
end
in
Stream.memoize (Stream.map saveLast lines)
@@ -160,7 +160,7 @@
[] => nothing
| c :: cs => (exactChar c ++ exactCharList cs) >> snd;
-fun exactString s = exactCharList (explode s);
+fun exactString s = exactCharList (String.explode s);
fun escapeString {escape} =
let
@@ -183,7 +183,7 @@
((exactChar #"\\" ++ escapeParser) >> snd) ||
some isNormal
in
- many charParser >> implode
+ many charParser >> String.implode
end;
local
@@ -196,46 +196,51 @@
val atLeastOneSpace = atLeastOne space >> K ();
end;
-fun fromString parser s = fromList parser (explode s);
+fun fromString parser s = fromList parser (String.explode s);
(* ------------------------------------------------------------------------- *)
(* Infix operators. *)
(* ------------------------------------------------------------------------- *)
-fun parseGenInfix update sof toks parse inp =
+fun parseLayeredInfixes {tokens,assoc} mk tokParser subParser =
let
- val (e,rest) = parse inp
-
- val continue =
- case rest of
- Stream.Nil => NONE
- | Stream.Cons (h_t as (h,_)) =>
- if StringSet.member h toks then SOME h_t else NONE
- in
- case continue of
- NONE => (sof e, rest)
- | SOME (h,t) => parseGenInfix update (update sof h e) toks parse (t ())
- end;
+ fun layerTokParser inp =
+ let
+ val tok_rest as (tok,_) = tokParser inp
+ in
+ if StringSet.member tok tokens then tok_rest
+ else raise NoParse
+ end
-local
- fun add ({leftSpaces = _, token = t, rightSpaces = _}, s) = StringSet.add s t;
+ fun layerMk (x,txs) =
+ case assoc of
+ Print.LeftAssoc =>
+ let
+ fun inc ((t,y),z) = mk (t,z,y)
+ in
+ List.foldl inc x txs
+ end
+ | Print.NonAssoc =>
+ (case txs of
+ [] => x
+ | [(t,y)] => mk (t,x,y)
+ | _ => raise NoParse)
+ | Print.RightAssoc =>
+ (case rev txs of
+ [] => x
+ | tx :: txs =>
+ let
+ fun inc ((t,y),(u,z)) = (t, mk (u,y,z))
- fun parse leftAssoc toks con =
- let
- val update =
- if leftAssoc then (fn f => fn t => fn a => fn b => con (t, f a, b))
- else (fn f => fn t => fn a => fn b => f (con (t, a, b)))
- in
- parseGenInfix update I toks
- end;
-in
- fun parseLayeredInfixes {tokens,leftAssoc} =
- let
- val toks = List.foldl add StringSet.empty tokens
- in
- parse leftAssoc toks
- end;
-end;
+ val (t,y) = List.foldl inc tx txs
+ in
+ mk (t,x,y)
+ end)
+
+ val layerParser = subParser ++ many (layerTokParser ++ subParser)
+ in
+ layerParser >> layerMk
+ end;
fun parseInfixes ops =
let
@@ -243,7 +248,8 @@
val iparsers = List.map parseLayeredInfixes layeredOps
in
- fn con => fn subparser => foldl (fn (p,sp) => p con sp) subparser iparsers
+ fn mk => fn tokParser => fn subParser =>
+ List.foldr (fn (p,sp) => p mk tokParser sp) subParser iparsers
end;
(* ------------------------------------------------------------------------- *)
@@ -257,7 +263,7 @@
fun expand (QUOTE q, s) = s ^ q
| expand (ANTIQUOTE a, s) = s ^ printer a
- val string = foldl expand "" quote
+ val string = List.foldl expand "" quote
in
parser string
end;
--- a/src/Tools/Metis/src/Portable.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Portable.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
-(* ML SPECIFIC FUNCTIONS *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *)
+(* ML COMPILER SPECIFIC FUNCTIONS *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Portable =
@@ -19,17 +19,10 @@
val pointerEqual : 'a * 'a -> bool
(* ------------------------------------------------------------------------- *)
-(* Timing function applications. *)
+(* Marking critical sections of code. *)
(* ------------------------------------------------------------------------- *)
-val time : ('a -> 'b) -> 'a -> 'b
-
-(* ------------------------------------------------------------------------- *)
-(* Critical section markup (multiprocessing) *)
-(* ------------------------------------------------------------------------- *)
-
-(* MODIFIED by Jasmin Blanchette *)
-val CRITICAL: (unit -> 'a) -> 'a
+val critical : (unit -> 'a) -> unit -> 'a
(* ------------------------------------------------------------------------- *)
(* Generating random values. *)
@@ -43,4 +36,10 @@
val randomWord : unit -> Word.word
+(* ------------------------------------------------------------------------- *)
+(* Timing function applications. *)
+(* ------------------------------------------------------------------------- *)
+
+val time : ('a -> 'b) -> 'a -> 'b
+
end
--- a/src/Tools/Metis/src/PortableMlton.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/PortableMlton.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* MLTON SPECIFIC FUNCTIONS *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Portable :> Portable =
@@ -19,6 +19,33 @@
val pointerEqual = MLton.eq;
(* ------------------------------------------------------------------------- *)
+(* Marking critical sections of code. *)
+(* ------------------------------------------------------------------------- *)
+
+fun critical f () = f ();
+
+(* ------------------------------------------------------------------------- *)
+(* Generating random values. *)
+(* ------------------------------------------------------------------------- *)
+
+fun randomWord () = MLton.Random.rand ();
+
+fun randomBool () = Word.andb (randomWord (),0w1) = 0w0;
+
+fun randomInt 1 = 0
+ | randomInt 2 = Word.toInt (Word.andb (randomWord (), 0w1))
+ | randomInt 4 = Word.toInt (Word.andb (randomWord (), 0w3))
+ | randomInt n = Word.toInt (Word.mod (randomWord (), Word.fromInt n));
+
+local
+ fun wordToReal w = Real.fromInt (Word.toInt (Word.>> (w,0w1)))
+
+ val normalizer = 1.0 / wordToReal (Word.notb 0w0);
+in
+ fun randomReal () = normalizer * wordToReal (randomWord ());
+end;
+
+(* ------------------------------------------------------------------------- *)
(* Timing function applications. *)
(* ------------------------------------------------------------------------- *)
@@ -56,34 +83,6 @@
y
end;
-(* ------------------------------------------------------------------------- *)
-(* Critical section markup (multiprocessing) *)
-(* ------------------------------------------------------------------------- *)
-
-(* MODIFIED by Jasmin Blanchette *)
-fun CRITICAL e = e (); (*dummy*)
-
-(* ------------------------------------------------------------------------- *)
-(* Generating random values. *)
-(* ------------------------------------------------------------------------- *)
-
-fun randomWord () = MLton.Random.rand ();
-
-fun randomBool () = Word.andb (randomWord (),0w1) = 0w0;
-
-fun randomInt 1 = 0
- | randomInt 2 = Word.toInt (Word.andb (randomWord (), 0w1))
- | randomInt 4 = Word.toInt (Word.andb (randomWord (), 0w3))
- | randomInt n = Word.toInt (Word.mod (randomWord (), Word.fromInt n));
-
-local
- fun wordToReal w = Real.fromInt (Word.toInt (Word.>> (w,0w1)))
-
- val normalizer = 1.0 / wordToReal (Word.notb 0w0);
-in
- fun randomReal () = normalizer * wordToReal (randomWord ());
-end;
-
end
(* ------------------------------------------------------------------------- *)
--- a/src/Tools/Metis/src/PortableMosml.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/PortableMosml.sml Fri Sep 17 17:31:20 2010 +0200
@@ -23,17 +23,10 @@
end;
(* ------------------------------------------------------------------------- *)
-(* Timing function applications. *)
+(* Marking critical sections of code. *)
(* ------------------------------------------------------------------------- *)
-val time = Mosml.time;
-
-(* ------------------------------------------------------------------------- *)
-(* Critical section markup (multiprocessing) *)
-(* ------------------------------------------------------------------------- *)
-
-(* MODIFIED by Jasmin Blanchette *)
-fun CRITICAL e = e (); (*dummy*)
+fun critical f () = f ();
(* ------------------------------------------------------------------------- *)
(* Generating random values. *)
@@ -57,6 +50,12 @@
Word.orb (Word.<< (h,0w16), l)
end;
+(* ------------------------------------------------------------------------- *)
+(* Timing function applications. *)
+(* ------------------------------------------------------------------------- *)
+
+val time = Mosml.time;
+
end
(* ------------------------------------------------------------------------- *)
@@ -68,6 +67,18 @@
val _ = catch_interrupt true;
(* ------------------------------------------------------------------------- *)
+(* Forcing fully qualified names of some key functions. *)
+(* ------------------------------------------------------------------------- *)
+
+(*BasicDebug
+val explode = ()
+and foldl = ()
+and foldr = ()
+and implode = ()
+and print = ();
+*)
+
+(* ------------------------------------------------------------------------- *)
(* Ad-hoc upgrading of the Moscow ML basis library. *)
(* ------------------------------------------------------------------------- *)
@@ -101,6 +112,28 @@
fun OS_Process_isSuccess s = s = OS.Process.success;
+fun String_concatWith s l =
+ case l of
+ [] => ""
+ | h :: t => List.foldl (fn (x,y) => y ^ s ^ x) h t;
+
+fun String_isSubstring p s =
+ let
+ val sizeP = size p
+ and sizeS = size s
+ in
+ if sizeP > sizeS then false
+ else if sizeP = sizeS then p = s
+ else
+ let
+ fun check i = String.substring (s,i,sizeP) = p
+
+ fun checkn i = check i orelse (i > 0 andalso checkn (i - 1))
+ in
+ checkn (sizeS - sizeP)
+ end
+ end;
+
fun String_isSuffix p s =
let
val sizeP = size p
--- a/src/Tools/Metis/src/PortablePolyml.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/PortablePolyml.sml Fri Sep 17 17:31:20 2010 +0200
@@ -19,6 +19,24 @@
fun pointerEqual (x : 'a, y : 'a) = PolyML.pointerEq(x,y);
(* ------------------------------------------------------------------------- *)
+(* Marking critical sections of code. *)
+(* ------------------------------------------------------------------------- *)
+
+fun critical f () = f ();
+
+(* ------------------------------------------------------------------------- *)
+(* Generating random values. *)
+(* ------------------------------------------------------------------------- *)
+
+val randomWord = Random.nextWord;
+
+val randomBool = Random.nextBool;
+
+val randomInt = Random.nextInt;
+
+val randomReal = Random.nextReal;
+
+(* ------------------------------------------------------------------------- *)
(* Timing function applications. *)
(* ------------------------------------------------------------------------- *)
@@ -44,7 +62,7 @@
val {usr,sys} = Timer.checkCPUTimer c
val real = Timer.checkRealTimer r
in
- TextIO.print (* MODIFIED by Jasmin Blanchette *)
+ print
("User: " ^ p usr ^ " System: " ^ p sys ^
" Real: " ^ p real ^ "\n")
end
@@ -56,26 +74,6 @@
y
end;
-(* ------------------------------------------------------------------------- *)
-(* Critical section markup (multiprocessing) *)
-(* ------------------------------------------------------------------------- *)
-
-(* MODIFIED by Jasmin Blanchette *)
-fun CRITICAL e = NAMED_CRITICAL "metis" e;
-
-
-(* ------------------------------------------------------------------------- *)
-(* Generating random values. *)
-(* ------------------------------------------------------------------------- *)
-
-val randomWord = Random.nextWord;
-
-val randomBool = Random.nextBool;
-
-val randomInt = Random.nextInt;
-
-val randomReal = Random.nextReal;
-
end
(* ------------------------------------------------------------------------- *)
--- a/src/Tools/Metis/src/Print.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Print.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,27 +1,39 @@
(* ========================================================================= *)
(* PRETTY-PRINTING *)
-(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2008 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Print =
sig
(* ------------------------------------------------------------------------- *)
+(* Escaping strings. *)
+(* ------------------------------------------------------------------------- *)
+
+val escapeString : {escape : char list} -> string -> string
+
+(* ------------------------------------------------------------------------- *)
+(* A type of strings annotated with their size. *)
+(* ------------------------------------------------------------------------- *)
+
+type stringSize = string * int
+
+val mkStringSize : string -> stringSize
+
+(* ------------------------------------------------------------------------- *)
(* A type of pretty-printers. *)
(* ------------------------------------------------------------------------- *)
datatype breakStyle = Consistent | Inconsistent
-datatype ppStep =
+datatype step =
BeginBlock of breakStyle * int
| EndBlock
- | AddString of string
+ | AddString of stringSize
| AddBreak of int
| AddNewline
-type ppstream = ppStep Stream.stream
-
-type 'a pp = 'a -> ppstream
+type ppstream = step Stream.stream
(* ------------------------------------------------------------------------- *)
(* Pretty-printer primitives. *)
@@ -31,7 +43,7 @@
val endBlock : ppstream
-val addString : string -> ppstream
+val addString : stringSize -> ppstream
val addBreak : int -> ppstream
@@ -51,18 +63,24 @@
val blockProgram : breakStyle -> int -> ppstream list -> ppstream
-val bracket : string -> string -> ppstream -> ppstream
+(* ------------------------------------------------------------------------- *)
+(* Executing pretty-printers to generate lines. *)
+(* ------------------------------------------------------------------------- *)
-val field : string -> ppstream -> ppstream
-
-val record : (string * ppstream) list -> ppstream
+val execute :
+ {lineLength : int} -> ppstream ->
+ {indent : int, line : string} Stream.stream
(* ------------------------------------------------------------------------- *)
(* Pretty-printer combinators. *)
(* ------------------------------------------------------------------------- *)
+type 'a pp = 'a -> ppstream
+
val ppMap : ('a -> 'b) -> 'b pp -> 'a pp
+val ppString : string pp
+
val ppBracket : string -> string -> 'a pp -> 'a pp
val ppOp : string -> ppstream
@@ -81,8 +99,6 @@
val ppChar : char pp
-val ppString : string pp
-
val ppEscapeString : {escape : char list} -> string pp
val ppUnit : unit pp
@@ -111,7 +127,7 @@
val ppBreakStyle : breakStyle pp
-val ppPpStep : ppStep pp
+val ppStep : step pp
val ppPpstream : ppstream pp
@@ -119,28 +135,27 @@
(* Pretty-printing infix operators. *)
(* ------------------------------------------------------------------------- *)
+type token = string
+
+datatype assoc =
+ LeftAssoc
+ | NonAssoc
+ | RightAssoc
+
datatype infixes =
Infixes of
- {token : string,
+ {token : token,
precedence : int,
- leftAssoc : bool} list
+ assoc : assoc} list
val tokensInfixes : infixes -> StringSet.set
-val layerInfixes :
- infixes ->
- {tokens : {leftSpaces : int, token : string, rightSpaces : int} list,
- leftAssoc : bool} list
+val layerInfixes : infixes -> {tokens : StringSet.set, assoc : assoc} list
val ppInfixes :
- infixes -> ('a -> (string * 'a * 'a) option) -> ('a * bool) pp ->
- ('a * bool) pp
-
-(* ------------------------------------------------------------------------- *)
-(* Executing pretty-printers to generate lines. *)
-(* ------------------------------------------------------------------------- *)
-
-val execute : {lineLength : int} -> ppstream -> string Stream.stream
+ infixes ->
+ ('a -> (token * 'a * 'a) option) -> ('a * token) pp ->
+ ('a * bool) pp -> ('a * bool) pp
(* ------------------------------------------------------------------------- *)
(* Executing pretty-printers with a global line length. *)
--- a/src/Tools/Metis/src/Print.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Print.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* PRETTY-PRINTING *)
-(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2008 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Print :> Print =
@@ -29,23 +29,47 @@
| Stream.Cons (h,t) => revAppend h (revConcat o t);
local
- fun join current prev = (prev ^ "\n", current);
+ fun calcSpaces n = nChars #" " n;
+
+ val cacheSize = 2 * initialLineLength;
+
+ val cachedSpaces = Vector.tabulate (cacheSize,calcSpaces);
in
- fun joinNewline strm =
- case strm of
- Stream.Nil => Stream.Nil
- | Stream.Cons (h,t) => Stream.maps join Stream.singleton h (t ());
+ fun nSpaces n =
+ if n < cacheSize then Vector.sub (cachedSpaces,n)
+ else calcSpaces n;
end;
-local
- fun calcSpaces n = nChars #" " n;
+(* ------------------------------------------------------------------------- *)
+(* Escaping strings. *)
+(* ------------------------------------------------------------------------- *)
+
+fun escapeString {escape} =
+ let
+ val escapeMap = map (fn c => (c, "\\" ^ str c)) escape
- val cachedSpaces = Vector.tabulate (initialLineLength,calcSpaces);
-in
- fun nSpaces n =
- if n < initialLineLength then Vector.sub (cachedSpaces,n)
- else calcSpaces n;
-end;
+ fun escapeChar c =
+ case c of
+ #"\\" => "\\\\"
+ | #"\n" => "\\n"
+ | #"\t" => "\\t"
+ | _ =>
+ case List.find (equal c o fst) escapeMap of
+ SOME (_,s) => s
+ | NONE => str c
+ in
+ String.translate escapeChar
+ end;
+
+(* ------------------------------------------------------------------------- *)
+(* A type of strings annotated with their size. *)
+(* ------------------------------------------------------------------------- *)
+
+type stringSize = string * int;
+
+fun mkStringSize s = (s, size s);
+
+val emptyStringSize = mkStringSize "";
(* ------------------------------------------------------------------------- *)
(* A type of pretty-printers. *)
@@ -53,23 +77,21 @@
datatype breakStyle = Consistent | Inconsistent;
-datatype ppStep =
+datatype step =
BeginBlock of breakStyle * int
| EndBlock
- | AddString of string
+ | AddString of stringSize
| AddBreak of int
| AddNewline;
-type ppstream = ppStep Stream.stream;
-
-type 'a pp = 'a -> ppstream;
+type ppstream = step Stream.stream;
fun breakStyleToString style =
case style of
Consistent => "Consistent"
| Inconsistent => "Inconsistent";
-fun ppStepToString step =
+fun stepToString step =
case step of
BeginBlock _ => "BeginBlock"
| EndBlock => "EndBlock"
@@ -109,330 +131,6 @@
fun blockProgram style indent pps = block style indent (program pps);
-fun bracket l r pp =
- blockProgram Inconsistent (size l)
- [addString l,
- pp,
- addString r];
-
-fun field f pp =
- blockProgram Inconsistent 2
- [addString (f ^ " ="),
- addBreak 1,
- pp];
-
-val record =
- let
- val sep = sequence (addString ",") (addBreak 1)
-
- fun recordField (f,pp) = field f pp
-
- fun sepField f = sequence sep (recordField f)
-
- fun fields [] = []
- | fields (f :: fs) = recordField f :: map sepField fs
- in
- bracket "{" "}" o blockProgram Consistent 0 o fields
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Pretty-printer combinators. *)
-(* ------------------------------------------------------------------------- *)
-
-fun ppMap f ppB a : ppstream = ppB (f a);
-
-fun ppBracket l r ppA a = bracket l r (ppA a);
-
-fun ppOp s = sequence (if s = "" then skip else addString s) (addBreak 1);
-
-fun ppOp2 ab ppA ppB (a,b) =
- blockProgram Inconsistent 0
- [ppA a,
- ppOp ab,
- ppB b];
-
-fun ppOp3 ab bc ppA ppB ppC (a,b,c) =
- blockProgram Inconsistent 0
- [ppA a,
- ppOp ab,
- ppB b,
- ppOp bc,
- ppC c];
-
-fun ppOpList s ppA =
- let
- fun ppOpA a = sequence (ppOp s) (ppA a)
- in
- fn [] => skip
- | h :: t => blockProgram Inconsistent 0 (ppA h :: map ppOpA t)
- end;
-
-fun ppOpStream s ppA =
- let
- fun ppOpA a = sequence (ppOp s) (ppA a)
- in
- fn Stream.Nil => skip
- | Stream.Cons (h,t) =>
- blockProgram Inconsistent 0
- [ppA h,
- Stream.concat (Stream.map ppOpA (t ()))]
- end;
-
-(* ------------------------------------------------------------------------- *)
-(* Pretty-printers for common types. *)
-(* ------------------------------------------------------------------------- *)
-
-fun ppChar c = addString (str c);
-
-val ppString = addString;
-
-fun ppEscapeString {escape} =
- let
- val escapeMap = map (fn c => (c, "\\" ^ str c)) escape
-
- fun escapeChar c =
- case c of
- #"\\" => "\\\\"
- | #"\n" => "\\n"
- | #"\t" => "\\t"
- | _ =>
- case List.find (equal c o fst) escapeMap of
- SOME (_,s) => s
- | NONE => str c
- in
- fn s => addString (String.translate escapeChar s)
- end;
-
-val ppUnit : unit pp = K (addString "()");
-
-fun ppBool b = addString (if b then "true" else "false");
-
-fun ppInt i = addString (Int.toString i);
-
-local
- val ppNeg = addString "~"
- and ppSep = addString ","
- and ppZero = addString "0"
- and ppZeroZero = addString "00";
-
- fun ppIntBlock i =
- if i < 10 then sequence ppZeroZero (ppInt i)
- else if i < 100 then sequence ppZero (ppInt i)
- else ppInt i;
-
- fun ppIntBlocks i =
- if i < 1000 then ppInt i
- else sequence (ppIntBlocks (i div 1000))
- (sequence ppSep (ppIntBlock (i mod 1000)));
-in
- fun ppPrettyInt i =
- if i < 0 then sequence ppNeg (ppIntBlocks (~i))
- else ppIntBlocks i;
-end;
-
-fun ppReal r = addString (Real.toString r);
-
-fun ppPercent p = addString (percentToString p);
-
-fun ppOrder x =
- addString
- (case x of
- LESS => "Less"
- | EQUAL => "Equal"
- | GREATER => "Greater");
-
-fun ppList ppA = ppBracket "[" "]" (ppOpList "," ppA);
-
-fun ppStream ppA = ppBracket "[" "]" (ppOpStream "," ppA);
-
-fun ppOption ppA ao =
- case ao of
- SOME a => ppA a
- | NONE => addString "-";
-
-fun ppPair ppA ppB = ppBracket "(" ")" (ppOp2 "," ppA ppB);
-
-fun ppTriple ppA ppB ppC = ppBracket "(" ")" (ppOp3 "," "," ppA ppB ppC);
-
-fun ppBreakStyle style = addString (breakStyleToString style);
-
-fun ppPpStep step =
- let
- val cmd = ppStepToString step
- in
- blockProgram Inconsistent 2
- (addString cmd ::
- (case step of
- BeginBlock style_indent =>
- [addBreak 1,
- ppPair ppBreakStyle ppInt style_indent]
- | EndBlock => []
- | AddString s =>
- [addBreak 1,
- addString ("\"" ^ s ^ "\"")]
- | AddBreak n =>
- [addBreak 1,
- ppInt n]
- | AddNewline => []))
- end;
-
-val ppPpstream = ppStream ppPpStep;
-
-(* ------------------------------------------------------------------------- *)
-(* Pretty-printing infix operators. *)
-(* ------------------------------------------------------------------------- *)
-
-datatype infixes =
- Infixes of
- {token : string,
- precedence : int,
- leftAssoc : bool} list;
-
-local
- fun chop l =
- case l of
- #" " :: l => let val (n,l) = chop l in (n + 1, l) end
- | _ => (0,l);
-in
- fun opSpaces tok =
- let
- val tok = explode tok
- val (r,tok) = chop (rev tok)
- val (l,tok) = chop (rev tok)
- val tok = implode tok
- in
- {leftSpaces = l, token = tok, rightSpaces = r}
- end;
-end;
-
-fun ppOpSpaces {leftSpaces,token,rightSpaces} =
- let
- val leftSpacesToken =
- if leftSpaces = 0 then token else nSpaces leftSpaces ^ token
- in
- sequence
- (addString leftSpacesToken)
- (addBreak rightSpaces)
- end;
-
-local
- fun new t l acc = {tokens = [opSpaces t], leftAssoc = l} :: acc;
-
- fun add t l acc =
- case acc of
- [] => raise Bug "Print.layerInfixOps.layer"
- | {tokens = ts, leftAssoc = l'} :: acc =>
- if l = l' then {tokens = opSpaces t :: ts, leftAssoc = l} :: acc
- else raise Bug "Print.layerInfixOps: mixed assocs";
-
- fun layer ({token = t, precedence = p, leftAssoc = l}, (acc,p')) =
- let
- val acc = if p = p' then add t l acc else new t l acc
- in
- (acc,p)
- end;
-in
- fun layerInfixes (Infixes i) =
- case sortMap #precedence Int.compare i of
- [] => []
- | {token = t, precedence = p, leftAssoc = l} :: i =>
- let
- val acc = new t l []
-
- val (acc,_) = List.foldl layer (acc,p) i
- in
- acc
- end;
-end;
-
-val tokensLayeredInfixes =
- let
- fun addToken ({leftSpaces = _, token = t, rightSpaces = _}, s) =
- StringSet.add s t
-
- fun addTokens ({tokens = t, leftAssoc = _}, s) =
- List.foldl addToken s t
- in
- List.foldl addTokens StringSet.empty
- end;
-
-val tokensInfixes = tokensLayeredInfixes o layerInfixes;
-
-local
- val mkTokenMap =
- let
- fun add (token,m) =
- let
- val {leftSpaces = _, token = t, rightSpaces = _} = token
- in
- StringMap.insert m (t, ppOpSpaces token)
- end
- in
- List.foldl add (StringMap.new ())
- end;
-in
- fun ppGenInfix {tokens,leftAssoc} =
- let
- val tokenMap = mkTokenMap tokens
- in
- fn dest => fn ppSub =>
- let
- fun dest' tm =
- case dest tm of
- NONE => NONE
- | SOME (t,a,b) =>
- case StringMap.peek tokenMap t of
- NONE => NONE
- | SOME p => SOME (p,a,b)
-
- fun ppGo (tmr as (tm,r)) =
- case dest' tm of
- NONE => ppSub tmr
- | SOME (p,a,b) =>
- program
- [(if leftAssoc then ppGo else ppSub) (a,true),
- p,
- (if leftAssoc then ppSub else ppGo) (b,r)]
- in
- fn tmr as (tm,_) =>
- if Option.isSome (dest' tm) then
- block Inconsistent 0 (ppGo tmr)
- else
- ppSub tmr
- end
- end;
-end
-
-fun ppInfixes ops =
- let
- val layeredOps = layerInfixes ops
-
- val toks = tokensLayeredInfixes layeredOps
-
- val iprinters = List.map ppGenInfix layeredOps
- in
- fn dest => fn ppSub =>
- let
- fun printer sub = foldl (fn (ip,p) => ip dest p) sub iprinters
-
- fun isOp t =
- case dest t of
- SOME (x,_,_) => StringSet.member x toks
- | NONE => false
-
- fun subpr (tmr as (tm,_)) =
- if isOp tm then
- blockProgram Inconsistent 1
- [addString "(",
- printer subpr (tm,false),
- addString ")"]
- else
- ppSub tmr
- in
- fn tmr => block Inconsistent 0 (printer subpr tmr)
- end
- end;
-
(* ------------------------------------------------------------------------- *)
(* Executing pretty-printers to generate lines. *)
(* ------------------------------------------------------------------------- *)
@@ -496,18 +194,22 @@
val sizeChunks = List.foldl (fn (c,z) => sizeChunk c + z) 0;
local
- fun render acc [] = acc
- | render acc (chunk :: chunks) =
- case chunk of
- StringChunk {string = s, ...} => render (acc ^ s) chunks
- | BreakChunk n => render (acc ^ nSpaces n) chunks
- | BlockChunk (Block {chunks = l, ...}) =>
- render acc (List.revAppend (l,chunks));
+ fun flatten acc chunks =
+ case chunks of
+ [] => rev acc
+ | chunk :: chunks =>
+ case chunk of
+ StringChunk {string = s, ...} => flatten (s :: acc) chunks
+ | BreakChunk n => flatten (nSpaces n :: acc) chunks
+ | BlockChunk (Block {chunks = l, ...}) =>
+ flatten acc (List.revAppend (l,chunks));
in
- fun renderChunks indent chunks = render (nSpaces indent) (rev chunks);
+ fun renderChunks indent chunks =
+ {indent = indent,
+ line = String.concat (flatten [] (rev chunks))};
+end;
- fun renderChunk indent chunk = renderChunks indent [chunk];
-end;
+fun renderChunk indent chunk = renderChunks indent [chunk];
fun isEmptyBlock block =
let
@@ -523,6 +225,7 @@
empty
end;
+(*BasicDebug
fun checkBlock ind block =
let
val Block {indent, style = _, size, chunks} = block
@@ -558,6 +261,7 @@
in
check initialIndent o rev
end;
+*)
val initialBlock =
let
@@ -969,12 +673,10 @@
(lines,state)
end;
- fun executeAddString lineLength s lines state =
+ fun executeAddString lineLength (s,n) lines state =
let
val State {blocks,lineIndent,lineSize} = state
- val n = size s
-
val blocks =
case blocks of
[] => raise Bug "Print.executeAddString: no block"
@@ -1061,10 +763,13 @@
fun executeAddNewline lineLength lines state =
let
- val (lines,state) = executeAddString lineLength "" lines state
- val (lines,state) = executeBigBreak lineLength lines state
+ val (lines,state) =
+ executeAddString lineLength emptyStringSize lines state
+
+ val (lines,state) =
+ executeBigBreak lineLength lines state
in
- executeAddString lineLength "" lines state
+ executeAddString lineLength emptyStringSize lines state
end;
fun final lineLength lines state =
@@ -1108,7 +813,7 @@
(lines,state)
end
handle Bug bug =>
- raise Bug ("Print.advance: after " ^ ppStepToString step ^
+ raise Bug ("Print.advance: after " ^ stepToString step ^
" command:\n" ^ bug)
*)
in
@@ -1117,21 +822,372 @@
end;
(* ------------------------------------------------------------------------- *)
+(* Pretty-printer combinators. *)
+(* ------------------------------------------------------------------------- *)
+
+type 'a pp = 'a -> ppstream;
+
+fun ppMap f ppB a : ppstream = ppB (f a);
+
+fun ppBracket' l r ppA a =
+ let
+ val (_,n) = l
+ in
+ blockProgram Inconsistent n
+ [addString l,
+ ppA a,
+ addString r]
+ end;
+
+fun ppOp' s = sequence (addString s) (addBreak 1);
+
+fun ppOp2' ab ppA ppB (a,b) =
+ blockProgram Inconsistent 0
+ [ppA a,
+ ppOp' ab,
+ ppB b];
+
+fun ppOp3' ab bc ppA ppB ppC (a,b,c) =
+ blockProgram Inconsistent 0
+ [ppA a,
+ ppOp' ab,
+ ppB b,
+ ppOp' bc,
+ ppC c];
+
+fun ppOpList' s ppA =
+ let
+ fun ppOpA a = sequence (ppOp' s) (ppA a)
+ in
+ fn [] => skip
+ | h :: t => blockProgram Inconsistent 0 (ppA h :: map ppOpA t)
+ end;
+
+fun ppOpStream' s ppA =
+ let
+ fun ppOpA a = sequence (ppOp' s) (ppA a)
+ in
+ fn Stream.Nil => skip
+ | Stream.Cons (h,t) =>
+ blockProgram Inconsistent 0
+ [ppA h,
+ Stream.concat (Stream.map ppOpA (t ()))]
+ end;
+
+fun ppBracket l r = ppBracket' (mkStringSize l) (mkStringSize r);
+
+fun ppOp s = ppOp' (mkStringSize s);
+
+fun ppOp2 ab = ppOp2' (mkStringSize ab);
+
+fun ppOp3 ab bc = ppOp3' (mkStringSize ab) (mkStringSize bc);
+
+fun ppOpList s = ppOpList' (mkStringSize s);
+
+fun ppOpStream s = ppOpStream' (mkStringSize s);
+
+(* ------------------------------------------------------------------------- *)
+(* Pretty-printers for common types. *)
+(* ------------------------------------------------------------------------- *)
+
+fun ppChar c = addString (str c, 1);
+
+fun ppString s = addString (mkStringSize s);
+
+fun ppEscapeString escape = ppMap (escapeString escape) ppString;
+
+local
+ val pp = ppString "()";
+in
+ fun ppUnit () = pp;
+end;
+
+local
+ val ppTrue = ppString "true"
+ and ppFalse = ppString "false";
+in
+ fun ppBool b = if b then ppTrue else ppFalse;
+end;
+
+val ppInt = ppMap Int.toString ppString;
+
+local
+ val ppNeg = ppString "~"
+ and ppSep = ppString ","
+ and ppZero = ppString "0"
+ and ppZeroZero = ppString "00";
+
+ fun ppIntBlock i =
+ if i < 10 then sequence ppZeroZero (ppInt i)
+ else if i < 100 then sequence ppZero (ppInt i)
+ else ppInt i;
+
+ fun ppIntBlocks i =
+ if i < 1000 then ppInt i
+ else sequence (ppIntBlocks (i div 1000))
+ (sequence ppSep (ppIntBlock (i mod 1000)));
+in
+ fun ppPrettyInt i =
+ if i < 0 then sequence ppNeg (ppIntBlocks (~i))
+ else ppIntBlocks i;
+end;
+
+val ppReal = ppMap Real.toString ppString;
+
+val ppPercent = ppMap percentToString ppString;
+
+local
+ val ppLess = ppString "Less"
+ and ppEqual = ppString "Equal"
+ and ppGreater = ppString "Greater";
+in
+ fun ppOrder ord =
+ case ord of
+ LESS => ppLess
+ | EQUAL => ppEqual
+ | GREATER => ppGreater;
+end;
+
+local
+ val left = mkStringSize "["
+ and right = mkStringSize "]"
+ and sep = mkStringSize ",";
+in
+ fun ppList ppX xs = ppBracket' left right (ppOpList' sep ppX) xs;
+
+ fun ppStream ppX xs = ppBracket' left right (ppOpStream' sep ppX) xs;
+end;
+
+local
+ val ppNone = ppString "-";
+in
+ fun ppOption ppX xo =
+ case xo of
+ SOME x => ppX x
+ | NONE => ppNone;
+end;
+
+local
+ val left = mkStringSize "("
+ and right = mkStringSize ")"
+ and sep = mkStringSize ",";
+in
+ fun ppPair ppA ppB =
+ ppBracket' left right (ppOp2' sep ppA ppB);
+
+ fun ppTriple ppA ppB ppC =
+ ppBracket' left right (ppOp3' sep sep ppA ppB ppC);
+end;
+
+val ppBreakStyle = ppMap breakStyleToString ppString;
+
+val ppStep = ppMap stepToString ppString;
+
+val ppStringSize =
+ let
+ val left = mkStringSize "\""
+ and right = mkStringSize "\""
+ and escape = {escape = [#"\""]}
+
+ val pp = ppBracket' left right (ppEscapeString escape)
+ in
+ fn (s,n) => if size s = n then pp s else ppPair pp ppInt (s,n)
+ end;
+
+fun ppStep step =
+ blockProgram Inconsistent 2
+ (ppStep step ::
+ (case step of
+ BeginBlock style_indent =>
+ [addBreak 1,
+ ppPair ppBreakStyle ppInt style_indent]
+ | EndBlock => []
+ | AddString s =>
+ [addBreak 1,
+ ppStringSize s]
+ | AddBreak n =>
+ [addBreak 1,
+ ppInt n]
+ | AddNewline => []));
+
+val ppPpstream = ppStream ppStep;
+
+(* ------------------------------------------------------------------------- *)
+(* Pretty-printing infix operators. *)
+(* ------------------------------------------------------------------------- *)
+
+type token = string;
+
+datatype assoc =
+ LeftAssoc
+ | NonAssoc
+ | RightAssoc;
+
+datatype infixes =
+ Infixes of
+ {token : token,
+ precedence : int,
+ assoc : assoc} list;
+
+local
+ fun comparePrecedence (io1,io2) =
+ let
+ val {token = _, precedence = p1, assoc = _} = io1
+ and {token = _, precedence = p2, assoc = _} = io2
+ in
+ Int.compare (p2,p1)
+ end;
+
+ fun equalAssoc a a' =
+ case a of
+ LeftAssoc => (case a' of LeftAssoc => true | _ => false)
+ | NonAssoc => (case a' of NonAssoc => true | _ => false)
+ | RightAssoc => (case a' of RightAssoc => true | _ => false);
+
+ fun new t a acc = {tokens = StringSet.singleton t, assoc = a} :: acc;
+
+ fun add t a' acc =
+ case acc of
+ [] => raise Bug "Print.layerInfixes: null"
+ | {tokens = ts, assoc = a} :: acc =>
+ if equalAssoc a a' then {tokens = StringSet.add ts t, assoc = a} :: acc
+ else raise Bug "Print.layerInfixes: mixed assocs";
+
+ fun layer ({token = t, precedence = p, assoc = a}, (acc,p')) =
+ let
+ val acc = if p = p' then add t a acc else new t a acc
+ in
+ (acc,p)
+ end;
+in
+ fun layerInfixes (Infixes ios) =
+ case sort comparePrecedence ios of
+ [] => []
+ | {token = t, precedence = p, assoc = a} :: ios =>
+ let
+ val acc = new t a []
+
+ val (acc,_) = List.foldl layer (acc,p) ios
+ in
+ acc
+ end;
+end;
+
+local
+ fun add ({tokens = ts, assoc = _}, tokens) = StringSet.union ts tokens;
+in
+ fun tokensLayeredInfixes l = List.foldl add StringSet.empty l;
+end;
+
+fun tokensInfixes ios = tokensLayeredInfixes (layerInfixes ios);
+
+fun destInfixOp dest tokens tm =
+ case dest tm of
+ NONE => NONE
+ | s as SOME (t,a,b) => if StringSet.member t tokens then s else NONE;
+
+fun ppLayeredInfixes dest ppTok {tokens,assoc} ppLower ppSub =
+ let
+ fun isLayer t = StringSet.member t tokens
+
+ fun ppTerm aligned (tm,r) =
+ case dest tm of
+ NONE => ppSub (tm,r)
+ | SOME (t,a,b) =>
+ if aligned andalso isLayer t then ppLayer (tm,t,a,b,r)
+ else ppLower (tm,t,a,b,r)
+
+ and ppLeft tm_r =
+ let
+ val aligned = case assoc of LeftAssoc => true | _ => false
+ in
+ ppTerm aligned tm_r
+ end
+
+ and ppLayer (tm,t,a,b,r) =
+ program
+ [ppLeft (a,true),
+ ppTok (tm,t),
+ ppRight (b,r)]
+
+ and ppRight tm_r =
+ let
+ val aligned = case assoc of RightAssoc => true | _ => false
+ in
+ ppTerm aligned tm_r
+ end
+ in
+ fn tm_t_a_b_r as (_,t,_,_,_) =>
+ if isLayer t then block Inconsistent 0 (ppLayer tm_t_a_b_r)
+ else ppLower tm_t_a_b_r
+ end;
+
+local
+ val leftBrack = mkStringSize "("
+ and rightBrack = mkStringSize ")";
+in
+ fun ppInfixes ops =
+ let
+ val layers = layerInfixes ops
+
+ val toks = tokensLayeredInfixes layers
+ in
+ fn dest => fn ppTok => fn ppSub =>
+ let
+ fun destOp tm = destInfixOp dest toks tm
+
+ fun ppInfix tm_t_a_b_r = ppLayers layers tm_t_a_b_r
+
+ and ppLayers ls (tm,t,a,b,r) =
+ case ls of
+ [] =>
+ ppBracket' leftBrack rightBrack ppInfix (tm,t,a,b,false)
+ | l :: ls =>
+ let
+ val ppLower = ppLayers ls
+ in
+ ppLayeredInfixes destOp ppTok l ppLower ppSub (tm,t,a,b,r)
+ end
+ in
+ fn (tm,r) =>
+ case destOp tm of
+ SOME (t,a,b) => ppInfix (tm,t,a,b,r)
+ | NONE => ppSub (tm,r)
+ end
+ end;
+end;
+
+(* ------------------------------------------------------------------------- *)
(* Executing pretty-printers with a global line length. *)
(* ------------------------------------------------------------------------- *)
val lineLength = ref initialLineLength;
fun toStream ppA a =
- Stream.map (fn s => s ^ "\n")
+ Stream.map (fn {indent,line} => nSpaces indent ^ line ^ "\n")
(execute {lineLength = !lineLength} (ppA a));
-fun toString ppA a =
- case execute {lineLength = !lineLength} (ppA a) of
- Stream.Nil => ""
- | Stream.Cons (h,t) => Stream.foldl (fn (s,z) => z ^ "\n" ^ s) h (t ());
+local
+ fun inc {indent,line} acc = line :: nSpaces indent :: acc;
-fun trace ppX nameX x =
- Useful.trace (toString (ppOp2 " =" ppString ppX) (nameX,x) ^ "\n");
+ fun incn (indent_line,acc) = inc indent_line ("\n" :: acc);
+in
+ fun toString ppA a =
+ case execute {lineLength = !lineLength} (ppA a) of
+ Stream.Nil => ""
+ | Stream.Cons (h,t) =>
+ let
+ val lines = Stream.foldl incn (inc h []) (t ())
+ in
+ String.concat (rev lines)
+ end;
+end;
+
+local
+ val sep = mkStringSize " =";
+in
+ fun trace ppX nameX x =
+ Useful.trace (toString (ppOp2' sep ppString ppX) (nameX,x) ^ "\n");
+end;
end
--- a/src/Tools/Metis/src/Problem.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Problem.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* CNF PROBLEMS *)
-(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Problem =
--- a/src/Tools/Metis/src/Problem.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Problem.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* CNF PROBLEMS *)
-(* Copyright (c) 2001-2008 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Problem :> Problem =
@@ -29,9 +29,9 @@
val cls = toClauses prob
in
{clauses = length cls,
- literals = foldl lits 0 cls,
- symbols = foldl syms 0 cls,
- typedSymbols = foldl typedSyms 0 cls}
+ literals = List.foldl lits 0 cls,
+ symbols = List.foldl syms 0 cls,
+ typedSymbols = List.foldl typedSyms 0 cls}
end;
fun freeVars {axioms,conjecture} =
--- a/src/Tools/Metis/src/Proof.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Proof.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* PROOFS IN FIRST ORDER LOGIC *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Proof =
--- a/src/Tools/Metis/src/Proof.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Proof.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* PROOFS IN FIRST ORDER LOGIC *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Proof :> Proof =
@@ -39,40 +39,40 @@
fun ppSubst ppThm (sub,thm) =
Print.sequence (Print.addBreak 1)
(Print.blockProgram Print.Inconsistent 1
- [Print.addString "{",
+ [Print.ppString "{",
Print.ppOp2 " =" Print.ppString Subst.pp ("sub",sub),
- Print.addString ",",
+ Print.ppString ",",
Print.addBreak 1,
Print.ppOp2 " =" Print.ppString ppThm ("thm",thm),
- Print.addString "}"]);
+ Print.ppString "}"]);
fun ppResolve ppThm (res,pos,neg) =
Print.sequence (Print.addBreak 1)
(Print.blockProgram Print.Inconsistent 1
- [Print.addString "{",
+ [Print.ppString "{",
Print.ppOp2 " =" Print.ppString Atom.pp ("res",res),
- Print.addString ",",
+ Print.ppString ",",
Print.addBreak 1,
Print.ppOp2 " =" Print.ppString ppThm ("pos",pos),
- Print.addString ",",
+ Print.ppString ",",
Print.addBreak 1,
Print.ppOp2 " =" Print.ppString ppThm ("neg",neg),
- Print.addString "}"]);
+ Print.ppString "}"]);
fun ppRefl tm = Print.sequence (Print.addBreak 1) (Term.pp tm);
fun ppEquality (lit,path,res) =
Print.sequence (Print.addBreak 1)
(Print.blockProgram Print.Inconsistent 1
- [Print.addString "{",
+ [Print.ppString "{",
Print.ppOp2 " =" Print.ppString Literal.pp ("lit",lit),
- Print.addString ",",
+ Print.ppString ",",
Print.addBreak 1,
Print.ppOp2 " =" Print.ppString Term.ppPath ("path",path),
- Print.addString ",",
+ Print.ppString ",",
Print.addBreak 1,
Print.ppOp2 " =" Print.ppString Term.pp ("res",res),
- Print.addString "}"]);
+ Print.ppString "}"]);
fun ppInf ppAxiom ppThm inf =
let
@@ -80,7 +80,7 @@
in
Print.block Print.Inconsistent 2
(Print.sequence
- (Print.addString infString)
+ (Print.ppString infString)
(case inf of
Axiom cl => ppAxiom cl
| Assume x => ppAssume x
@@ -106,7 +106,7 @@
val prf = enumerate prf
fun ppThm th =
- Print.addString
+ Print.ppString
let
val cl = Thm.clause th
@@ -123,7 +123,7 @@
in
Print.sequence
(Print.blockProgram Print.Consistent (1 + size s)
- [Print.addString (s ^ " "),
+ [Print.ppString (s ^ " "),
Thm.pp th,
Print.addBreak 2,
Print.ppBracket "[" "]" (ppInf (K Print.skip) ppThm) inf])
@@ -131,10 +131,10 @@
end
in
Print.blockProgram Print.Consistent 0
- [Print.addString "START OF PROOF",
+ [Print.ppString "START OF PROOF",
Print.addNewline,
Print.program (map ppStep prf),
- Print.addString "END OF PROOF"]
+ Print.ppString "END OF PROOF"]
end
(*MetisDebug
handle Error err => raise Bug ("Proof.pp: shouldn't fail:\n" ^ err);
--- a/src/Tools/Metis/src/Random.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Random.sig Fri Sep 17 17:31:20 2010 +0200
@@ -8,6 +8,7 @@
signature Random =
sig
+
val nextWord : unit -> word
val nextBool : unit -> bool
--- a/src/Tools/Metis/src/Resolution.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Resolution.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* THE RESOLUTION PROOF PROCEDURE *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Resolution =
--- a/src/Tools/Metis/src/Resolution.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Resolution.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* THE RESOLUTION PROOF PROCEDURE *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Resolution :> Resolution =
--- a/src/Tools/Metis/src/Rewrite.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Rewrite.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* ORDERED REWRITING FOR FIRST ORDER TERMS *)
-(* Copyright (c) 2003-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2003 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Rewrite =
--- a/src/Tools/Metis/src/Rewrite.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Rewrite.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* ORDERED REWRITING FOR FIRST ORDER TERMS *)
-(* Copyright (c) 2003-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2003 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Rewrite :> Rewrite =
@@ -207,7 +207,7 @@
rw
end;
-val addList = foldl (fn (eqn,rw) => add rw eqn);
+val addList = List.foldl (fn (eqn,rw) => add rw eqn);
(* ------------------------------------------------------------------------- *)
(* Rewriting (the order must be a refinement of the rewrite order). *)
@@ -434,8 +434,9 @@
let
fun addSubterm b ((path,tm),net) = TermNet.insert net (tm,(id,b,path))
- val subterms = foldl (addSubterm true) subterms (Term.subterms l)
- val subterms = foldl (addSubterm false) subterms (Term.subterms r)
+ val subterms = List.foldl (addSubterm true) subterms (Term.subterms l)
+
+ val subterms = List.foldl (addSubterm false) subterms (Term.subterms r)
in
subterms
end;
@@ -660,7 +661,7 @@
in
fun orderedRewrite order ths =
let
- val rw = foldl addEqn (new order) (enumerate ths)
+ val rw = List.foldl addEqn (new order) (enumerate ths)
in
rewriteRule rw order
end;
--- a/src/Tools/Metis/src/Rule.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Rule.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* DERIVED RULES FOR CREATING FIRST ORDER LOGIC THEOREMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Rule =
--- a/src/Tools/Metis/src/Rule.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Rule.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* DERIVED RULES FOR CREATING FIRST ORDER LOGIC THEOREMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Rule :> Rule =
@@ -156,17 +156,19 @@
val noConv : conv = fn _ => raise Error "noConv";
+(*MetisDebug
fun traceConv s conv tm =
let
val res as (tm',th) = conv tm
- val () = print (s ^ ": " ^ Term.toString tm ^ " --> " ^
+ val () = trace (s ^ ": " ^ Term.toString tm ^ " --> " ^
Term.toString tm' ^ " " ^ Thm.toString th ^ "\n")
in
res
end
handle Error err =>
- (print (s ^ ": " ^ Term.toString tm ^ " --> Error: " ^ err ^ "\n");
+ (trace (s ^ ": " ^ Term.toString tm ^ " --> Error: " ^ err ^ "\n");
raise Error (s ^ ": " ^ err));
+*)
fun thenConvTrans tm (tm',th1) (tm'',th2) =
let
@@ -455,7 +457,7 @@
val reflTh = Thm.refl (Term.Fn (f,xs))
val reflLit = Thm.destUnit reflTh
in
- fst (foldl cong (reflTh,reflLit) (enumerate ys))
+ fst (List.foldl cong (reflTh,reflLit) (enumerate ys))
end;
(* ------------------------------------------------------------------------- *)
@@ -482,7 +484,7 @@
val assumeLit = (false,(R,xs))
val assumeTh = Thm.assume assumeLit
in
- fst (foldl cong (assumeTh,assumeLit) (enumerate ys))
+ fst (List.foldl cong (assumeTh,assumeLit) (enumerate ys))
end;
(* ------------------------------------------------------------------------- *)
--- a/src/Tools/Metis/src/Sharing.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Sharing.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* PRESERVING SHARING OF ML VALUES *)
-(* Copyright (c) 2005-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2005 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Sharing =
--- a/src/Tools/Metis/src/Sharing.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Sharing.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* PRESERVING SHARING OF ML VALUES *)
-(* Copyright (c) 2005-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2005 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Sharing :> Sharing =
--- a/src/Tools/Metis/src/Stream.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Stream.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* A POSSIBLY-INFINITE STREAM DATATYPE FOR ML *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Stream =
--- a/src/Tools/Metis/src/Stream.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Stream.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,13 +1,11 @@
(* ========================================================================= *)
(* A POSSIBLY-INFINITE STREAM DATATYPE FOR ML *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Stream :> Stream =
struct
-open Useful; (* MODIFIED by Jasmin Blanchette *)
-
val K = Useful.K;
val pair = Useful.pair;
@@ -201,9 +199,9 @@
fun listConcat s = concat (map fromList s);
-fun toString s = implode (toList s);
+fun toString s = String.implode (toList s);
-fun fromString s = fromList (explode s);
+fun fromString s = fromList (String.explode s);
fun toTextFile {filename = f} s =
let
--- a/src/Tools/Metis/src/Subst.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Subst.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* FIRST ORDER LOGIC SUBSTITUTIONS *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Subst =
--- a/src/Tools/Metis/src/Subst.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Subst.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* FIRST ORDER LOGIC SUBSTITUTIONS *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Subst :> Subst =
--- a/src/Tools/Metis/src/Subsume.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Subsume.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* SUBSUMPTION CHECKING FOR FIRST ORDER LOGIC CLAUSES *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Subsume =
--- a/src/Tools/Metis/src/Subsume.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Subsume.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* SUBSUMPTION CHECKING FOR FIRST ORDER LOGIC CLAUSES *)
-(* Copyright (c) 2002-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Subsume :> Subsume =
--- a/src/Tools/Metis/src/Term.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Term.sml Fri Sep 17 17:31:20 2010 +0200
@@ -348,7 +348,7 @@
val isApp = can destApp;
-fun listMkApp (f,l) = foldl mkApp f l;
+fun listMkApp (f,l) = List.foldl mkApp f l;
local
fun strip tms tm =
@@ -368,38 +368,38 @@
val infixes =
(ref o Print.Infixes)
[(* ML symbols *)
- {token = " / ", precedence = 7, leftAssoc = true},
- {token = " div ", precedence = 7, leftAssoc = true},
- {token = " mod ", precedence = 7, leftAssoc = true},
- {token = " * ", precedence = 7, leftAssoc = true},
- {token = " + ", precedence = 6, leftAssoc = true},
- {token = " - ", precedence = 6, leftAssoc = true},
- {token = " ^ ", precedence = 6, leftAssoc = true},
- {token = " @ ", precedence = 5, leftAssoc = false},
- {token = " :: ", precedence = 5, leftAssoc = false},
- {token = " = ", precedence = 4, leftAssoc = true},
- {token = " <> ", precedence = 4, leftAssoc = true},
- {token = " <= ", precedence = 4, leftAssoc = true},
- {token = " < ", precedence = 4, leftAssoc = true},
- {token = " >= ", precedence = 4, leftAssoc = true},
- {token = " > ", precedence = 4, leftAssoc = true},
- {token = " o ", precedence = 3, leftAssoc = true},
- {token = " -> ", precedence = 2, leftAssoc = false}, (* inferred prec *)
- {token = " : ", precedence = 1, leftAssoc = false}, (* inferred prec *)
- {token = ", ", precedence = 0, leftAssoc = false}, (* inferred prec *)
+ {token = "/", precedence = 7, assoc = Print.LeftAssoc},
+ {token = "div", precedence = 7, assoc = Print.LeftAssoc},
+ {token = "mod", precedence = 7, assoc = Print.LeftAssoc},
+ {token = "*", precedence = 7, assoc = Print.LeftAssoc},
+ {token = "+", precedence = 6, assoc = Print.LeftAssoc},
+ {token = "-", precedence = 6, assoc = Print.LeftAssoc},
+ {token = "^", precedence = 6, assoc = Print.LeftAssoc},
+ {token = "@", precedence = 5, assoc = Print.RightAssoc},
+ {token = "::", precedence = 5, assoc = Print.RightAssoc},
+ {token = "=", precedence = 4, assoc = Print.NonAssoc},
+ {token = "<>", precedence = 4, assoc = Print.NonAssoc},
+ {token = "<=", precedence = 4, assoc = Print.NonAssoc},
+ {token = "<", precedence = 4, assoc = Print.NonAssoc},
+ {token = ">=", precedence = 4, assoc = Print.NonAssoc},
+ {token = ">", precedence = 4, assoc = Print.NonAssoc},
+ {token = "o", precedence = 3, assoc = Print.LeftAssoc},
+ {token = "->", precedence = 2, assoc = Print.RightAssoc},
+ {token = ":", precedence = 1, assoc = Print.NonAssoc},
+ {token = ",", precedence = 0, assoc = Print.RightAssoc},
(* Logical connectives *)
- {token = " /\\ ", precedence = ~1, leftAssoc = false},
- {token = " \\/ ", precedence = ~2, leftAssoc = false},
- {token = " ==> ", precedence = ~3, leftAssoc = false},
- {token = " <=> ", precedence = ~4, leftAssoc = false},
+ {token = "/\\", precedence = ~1, assoc = Print.RightAssoc},
+ {token = "\\/", precedence = ~2, assoc = Print.RightAssoc},
+ {token = "==>", precedence = ~3, assoc = Print.RightAssoc},
+ {token = "<=>", precedence = ~4, assoc = Print.RightAssoc},
(* Other symbols *)
- {token = " . ", precedence = 9, leftAssoc = true}, (* function app *)
- {token = " ** ", precedence = 8, leftAssoc = true},
- {token = " ++ ", precedence = 6, leftAssoc = true},
- {token = " -- ", precedence = 6, leftAssoc = true},
- {token = " == ", precedence = 4, leftAssoc = true}];
+ {token = ".", precedence = 9, assoc = Print.LeftAssoc},
+ {token = "**", precedence = 8, assoc = Print.LeftAssoc},
+ {token = "++", precedence = 6, assoc = Print.LeftAssoc},
+ {token = "--", precedence = 6, assoc = Print.LeftAssoc},
+ {token = "==", precedence = 4, assoc = Print.NonAssoc}];
(* The negation symbol *)
@@ -422,9 +422,14 @@
and neg = !negation
and bracks = !brackets
- val bracks = map (fn (b1,b2) => (b1 ^ b2, b1, b2)) bracks
+ val bMap =
+ let
+ fun f (b1,b2) = (b1 ^ b2, b1, b2)
+ in
+ List.map f bracks
+ end
- val bTokens = map #2 bracks @ map #3 bracks
+ val bTokens = op@ (unzip bracks)
val iTokens = Print.tokensInfixes iOps
@@ -438,7 +443,15 @@
end
| _ => NONE
- val iPrinter = Print.ppInfixes iOps destI
+ fun isI tm = Option.isSome (destI tm)
+
+ fun iToken (_,tok) =
+ Print.program
+ [(if tok = "," then Print.skip else Print.ppString " "),
+ Print.ppString tok,
+ Print.addBreak 1];
+
+ val iPrinter = Print.ppInfixes iOps destI iToken
val specialTokens =
StringSet.addList iTokens (neg :: quants @ ["$","(",")"] @ bTokens)
@@ -466,8 +479,6 @@
fun functionName bv = Print.ppMap (checkFunctionName bv) Print.ppString
- fun isI tm = Option.isSome (destI tm)
-
fun stripNeg tm =
case tm of
Fn (f,[a]) =>
@@ -494,7 +505,7 @@
let
val s = Name.toString b
in
- case List.find (fn (n,_,_) => n = s) bracks of
+ case List.find (fn (n,_,_) => n = s) bMap of
NONE => NONE
| SOME (_,b1,b2) => SOME (b1,tm,b2)
end
@@ -527,11 +538,11 @@
val bv = StringSet.addList bv (map Name.toString (v :: vs))
in
Print.program
- [Print.addString q,
+ [Print.ppString q,
varName bv v,
Print.program
(map (Print.sequence (Print.addBreak 1) o varName bv) vs),
- Print.addString ".",
+ Print.ppString ".",
Print.addBreak 1,
innerQuant bv tm]
end
@@ -545,7 +556,7 @@
val (n,tm) = stripNeg tm
in
Print.blockProgram Print.Inconsistent n
- [Print.duplicate n (Print.addString neg),
+ [Print.duplicate n (Print.ppString neg),
if isI tm orelse (r andalso isQuant tm) then bracket bv tm
else quantifier bv tm]
end
@@ -571,31 +582,32 @@
val isAlphaNum =
let
- val alphaNumChars = explode "_'"
+ val alphaNumChars = String.explode "_'"
in
fn c => mem c alphaNumChars orelse Char.isAlphaNum c
end;
local
- val alphaNumToken = atLeastOne (some isAlphaNum) >> implode;
+ val alphaNumToken = atLeastOne (some isAlphaNum) >> String.implode;
val symbolToken =
let
fun isNeg c = str c = !negation
- val symbolChars = explode "<>=-*+/\\?@|!$%&#^:;~"
+ val symbolChars = String.explode "<>=-*+/\\?@|!$%&#^:;~"
fun isSymbol c = mem c symbolChars
fun isNonNegSymbol c = not (isNeg c) andalso isSymbol c
in
some isNeg >> str ||
- (some isNonNegSymbol ++ many (some isSymbol)) >> (implode o op::)
+ (some isNonNegSymbol ++ many (some isSymbol)) >>
+ (String.implode o op::)
end;
val punctToken =
let
- val punctChars = explode "()[]{}.,"
+ val punctChars = String.explode "()[]{}.,"
fun isPunct c = mem c punctChars
in
@@ -627,8 +639,9 @@
val iTokens = Print.tokensInfixes iOps
- val iParser =
- parseInfixes iOps (fn (f,a,b) => Fn (Name.fromString f, [a,b]))
+ fun iMk (f,a,b) = Fn (Name.fromString f, [a,b])
+
+ val iParser = parseInfixes iOps iMk any
val specialTokens =
StringSet.addList iTokens (neg :: quants @ ["$"] @ bTokens)
@@ -667,7 +680,7 @@
some (Useful.equal ".")) >>++
(fn (_,(vs,_)) =>
term (StringSet.addList bv vs) >>
- (fn body => foldr bind body vs))
+ (fn body => List.foldr bind body vs))
end
in
var ||
@@ -696,7 +709,7 @@
in
fun fromString input =
let
- val chars = Stream.fromList (explode input)
+ val chars = Stream.fromList (String.explode input)
val tokens = everything (lexer >> singleton) chars
@@ -709,7 +722,8 @@
end;
local
- val antiquotedTermToString = Print.toString (Print.ppBracket "(" ")" pp);
+ val antiquotedTermToString =
+ Print.toString (Print.ppBracket "(" ")" pp);
in
val parse = Parse.parseQuotation antiquotedTermToString fromString;
end;
--- a/src/Tools/Metis/src/TermNet.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/TermNet.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC TERMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature TermNet =
--- a/src/Tools/Metis/src/TermNet.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/TermNet.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* MATCHING AND UNIFICATION FOR SETS OF FIRST ORDER LOGIC TERMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure TermNet :> TermNet =
@@ -158,7 +158,7 @@
fun null net = size net = 0;
-fun singles qtms a = foldr Single a qtms;
+fun singles qtms a = List.foldr Single a qtms;
local
fun pre NONE = (0,NONE)
@@ -188,7 +188,7 @@
handle Error _ => raise Bug "TermNet.insert: should never fail";
end;
-fun fromList parm l = foldl (fn (tm_a,n) => insert n tm_a) (new parm) l;
+fun fromList parm l = List.foldl (fn (tm_a,n) => insert n tm_a) (new parm) l;
fun filter pred =
let
@@ -441,7 +441,7 @@
local
fun inc (qtm, Result l, acc) =
- foldl (fn ((n,a),acc) => (n,(qtm,a)) :: acc) acc l
+ List.foldl (fn ((n,a),acc) => (n,(qtm,a)) :: acc) acc l
| inc _ = raise Bug "TermNet.pp.inc";
fun toList (Net (_,_,NONE)) = []
--- a/src/Tools/Metis/src/Thm.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Thm.sml Fri Sep 17 17:31:20 2010 +0200
@@ -111,7 +111,7 @@
in
fun pp th =
Print.blockProgram Print.Inconsistent 3
- [Print.addString "|- ",
+ [Print.ppString "|- ",
Formula.pp (toFormula th)];
end;
--- a/src/Tools/Metis/src/Tptp.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Tptp.sml Fri Sep 17 17:31:20 2010 +0200
@@ -158,7 +158,7 @@
fun isTptpChar #"_" = true
| isTptpChar c = Char.isAlphaNum c;
- fun isTptpName l s = nonEmptyPred (existsPred l) (explode s);
+ fun isTptpName l s = nonEmptyPred (existsPred l) (String.explode s);
fun isRegular (c,cs) =
Char.isLower c andalso List.all isTptpChar cs;
@@ -175,14 +175,14 @@
fun mkTptpVarName s =
let
val s =
- case List.filter isTptpChar (explode s) of
+ case List.filter isTptpChar (String.explode s) of
[] => [#"X"]
| l as c :: cs =>
if Char.isUpper c then l
else if Char.isLower c then Char.toUpper c :: cs
else #"X" :: l
in
- implode s
+ String.implode s
end;
val isTptpConstName = isTptpName [isRegular,isNumber,isDefined,isSystem]
@@ -615,10 +615,10 @@
let
val s = varToTptp mapping v
in
- Print.addString s
+ Print.ppString s
end;
-fun ppFnName mapping fa = Print.addString (fnToTptp mapping fa);
+fun ppFnName mapping fa = Print.ppString (fnToTptp mapping fa);
fun ppConst mapping c = ppFnName mapping (c,0);
@@ -633,14 +633,14 @@
| a =>
Print.blockProgram Print.Inconsistent 2
[ppFnName mapping (f,a),
- Print.addString "(",
+ Print.ppString "(",
Print.ppOpList "," term tms,
- Print.addString ")"]
+ Print.ppString ")"]
in
Print.block Print.Inconsistent 0 o term
end;
-fun ppRelName mapping ra = Print.addString (relToTptp mapping ra);
+fun ppRelName mapping ra = Print.ppString (relToTptp mapping ra);
fun ppProp mapping p = ppRelName mapping (p,0);
@@ -650,12 +650,12 @@
| a =>
Print.blockProgram Print.Inconsistent 2
[ppRelName mapping (r,a),
- Print.addString "(",
+ Print.ppString "(",
Print.ppOpList "," (ppTerm mapping) tms,
- Print.addString ")"];
+ Print.ppString ")"];
local
- val neg = Print.sequence (Print.addString "~") (Print.addBreak 1);
+ val neg = Print.sequence (Print.ppString "~") (Print.addBreak 1);
fun fof mapping fm =
case fm of
@@ -672,8 +672,8 @@
and unitary mapping fm =
case fm of
- Formula.True => Print.addString "$true"
- | Formula.False => Print.addString "$false"
+ Formula.True => Print.ppString "$true"
+ | Formula.False => Print.ppString "$false"
| Formula.Forall _ => quantified mapping ("!", Formula.stripForall fm)
| Formula.Exists _ => quantified mapping ("?", Formula.stripExists fm)
| Formula.Not _ =>
@@ -693,21 +693,21 @@
| NONE => ppAtom mapping atm)
| _ =>
Print.blockProgram Print.Inconsistent 1
- [Print.addString "(",
+ [Print.ppString "(",
fof mapping fm,
- Print.addString ")"]
+ Print.ppString ")"]
and quantified mapping (q,(vs,fm)) =
let
val mapping = addVarListMapping mapping vs
in
Print.blockProgram Print.Inconsistent 2
- [Print.addString q,
- Print.addString " ",
+ [Print.ppString q,
+ Print.ppString " ",
Print.blockProgram Print.Inconsistent (String.size q)
- [Print.addString "[",
+ [Print.ppString "[",
Print.ppOpList "," (ppVar mapping) vs,
- Print.addString "] :"],
+ Print.ppString "] :"],
Print.addBreak 1,
unitary mapping fm]
end;
@@ -735,7 +735,8 @@
infixr 7 >>
infixr 6 ||
- val alphaNumToken = atLeastOne (some isAlphaNum) >> (AlphaNum o implode);
+ val alphaNumToken =
+ atLeastOne (some isAlphaNum) >> (AlphaNum o String.implode);
val punctToken =
let
@@ -759,7 +760,7 @@
some (not o stopOn) >> singleton
in
exactChar #"'" ++ many quotedParser ++ exactChar #"'" >>
- (fn (_,(l,_)) => Quote (implode (List.concat l)))
+ (fn (_,(l,_)) => Quote (String.implode (List.concat l)))
end;
val lexToken = alphaNumToken || punctToken || quoteToken;
@@ -779,7 +780,7 @@
let
fun funcs (lit,acc) = NameAritySet.union (functionsLiteral lit) acc
in
- foldl funcs NameAritySet.empty
+ List.foldl funcs NameAritySet.empty
end;
val clauseRelations =
@@ -789,14 +790,14 @@
NONE => acc
| SOME r => NameAritySet.add acc r
in
- foldl rels NameAritySet.empty
+ List.foldl rels NameAritySet.empty
end;
val clauseFreeVars =
let
fun fvs (lit,acc) = NameSet.union (freeVarsLiteral lit) acc
in
- foldl fvs NameSet.empty
+ List.foldl fvs NameSet.empty
end;
fun clauseSubst sub lits = map (literalSubst sub) lits;
@@ -1013,7 +1014,7 @@
fun ppLiteralInf mapping (pol,atm) =
Print.sequence
- (if pol then Print.skip else Print.addString "~ ")
+ (if pol then Print.skip else Print.ppString "~ ")
(ppAtomInf mapping atm);
in
fun ppProofTerm mapping =
@@ -1032,7 +1033,7 @@
fun ppProof mapping inf =
Print.blockProgram Print.Inconsistent 1
- [Print.addString "[",
+ [Print.ppString "[",
(case inf of
Proof.Axiom _ => Print.skip
| Proof.Assume atm => ppProofAtom mapping atm
@@ -1042,13 +1043,13 @@
| Proof.Equality (lit,path,tm) =>
Print.program
[ppProofLiteral mapping lit,
- Print.addString ",",
+ Print.ppString ",",
Print.addBreak 1,
ppProofPath path,
- Print.addString ",",
+ Print.ppString ",",
Print.addBreak 1,
ppProofTerm mapping tm]),
- Print.addString "]"];
+ Print.ppString "]"];
val ppParent = ppFormulaName;
@@ -1073,16 +1074,16 @@
val name = nameStrip inference
in
Print.blockProgram Print.Inconsistent (size gen + 1)
- [Print.addString gen,
- Print.addString "(",
- Print.addString name,
- Print.addString ",",
+ [Print.ppString gen,
+ Print.ppString "(",
+ Print.ppString name,
+ Print.ppString ",",
Print.addBreak 1,
Print.ppBracket "[" "]" (ppStrip mapping) inference,
- Print.addString ",",
+ Print.ppString ",",
Print.addBreak 1,
Print.ppList ppParent parents,
- Print.addString ")"]
+ Print.ppString ")"]
end
| NormalizeFormulaSource {inference,parents} =>
let
@@ -1091,16 +1092,16 @@
val name = nameNormalize inference
in
Print.blockProgram Print.Inconsistent (size gen + 1)
- [Print.addString gen,
- Print.addString "(",
- Print.addString name,
- Print.addString ",",
+ [Print.ppString gen,
+ Print.ppString "(",
+ Print.ppString name,
+ Print.ppString ",",
Print.addBreak 1,
Print.ppBracket "[" "]" (ppNormalize mapping) inference,
- Print.addString ",",
+ Print.ppString ",",
Print.addBreak 1,
Print.ppList ppParent parents,
- Print.addString ")"]
+ Print.ppString ")"]
end
| ProofFormulaSource {inference,parents} =>
let
@@ -1121,28 +1122,28 @@
end
in
Print.blockProgram Print.Inconsistent (size gen + 1)
- ([Print.addString gen,
- Print.addString "("] @
+ ([Print.ppString gen,
+ Print.ppString "("] @
(if isTaut then
- [Print.addString "tautology",
- Print.addString ",",
+ [Print.ppString "tautology",
+ Print.ppString ",",
Print.addBreak 1,
Print.blockProgram Print.Inconsistent 1
- [Print.addString "[",
- Print.addString name,
- Print.addString ",",
+ [Print.ppString "[",
+ Print.ppString name,
+ Print.ppString ",",
Print.addBreak 1,
ppProof mapping inference,
- Print.addString "]"]]
+ Print.ppString "]"]]
else
- [Print.addString name,
- Print.addString ",",
+ [Print.ppString name,
+ Print.ppString ",",
Print.addBreak 1,
ppProof mapping inference,
- Print.addString ",",
+ Print.ppString ",",
Print.addBreak 1,
Print.ppList (ppProofParent mapping) parents]) @
- [Print.addString ")"])
+ [Print.ppString ")"])
end
end;
@@ -1208,14 +1209,14 @@
let
fun funcs (fm,acc) = NameAritySet.union (functionsFormula fm) acc
in
- foldl funcs NameAritySet.empty
+ List.foldl funcs NameAritySet.empty
end;
val formulasRelations =
let
fun rels (fm,acc) = NameAritySet.union (relationsFormula fm) acc
in
- foldl rels NameAritySet.empty
+ List.foldl rels NameAritySet.empty
end;
fun isCnfConjectureFormula fm =
@@ -1244,24 +1245,24 @@
| FofFormulaBody _ => "fof"
in
Print.blockProgram Print.Inconsistent (size gen + 1)
- ([Print.addString gen,
- Print.addString "(",
+ ([Print.ppString gen,
+ Print.ppString "(",
ppFormulaName name,
- Print.addString ",",
+ Print.ppString ",",
Print.addBreak 1,
ppRole role,
- Print.addString ",",
+ Print.ppString ",",
Print.addBreak 1,
Print.blockProgram Print.Consistent 1
- [Print.addString "(",
+ [Print.ppString "(",
ppFormulaBody mapping body,
- Print.addString ")"]] @
+ Print.ppString ")"]] @
(if isNoFormulaSource source then []
else
- [Print.addString ",",
+ [Print.ppString ",",
Print.addBreak 1,
ppFormulaSource mapping source]) @
- [Print.addString ")."])
+ [Print.ppString ")."])
end;
fun formulaToString mapping = Print.toString (ppFormula mapping);
@@ -1285,7 +1286,7 @@
val stringParser = lowerParser || upperParser;
- val numberParser = someAlphaNum (List.all Char.isDigit o explode);
+ val numberParser = someAlphaNum (List.all Char.isDigit o String.explode);
fun somePunct p =
maybe (fn Punct c => if p c then SOME c else NONE | _ => NONE);
@@ -1299,7 +1300,7 @@
| f [x] = x
| f (h :: t) = (h ++ f t) >> K ();
in
- fun symbolParser s = f (map punctParser (explode s));
+ fun symbolParser s = f (map punctParser (String.explode s));
end;
val definedParser =
@@ -1641,9 +1642,9 @@
fun ppInclude i =
Print.blockProgram Print.Inconsistent 2
- [Print.addString "include('",
- Print.addString i,
- Print.addString "')."];
+ [Print.ppString "include('",
+ Print.ppString i,
+ Print.ppString "')."];
val includeToString = Print.toString ppInclude;
@@ -1745,8 +1746,8 @@
fun destLineComment cs =
case cs of
[] => ""
- | #"%" :: #" " :: rest => implode rest
- | #"%" :: rest => implode rest
+ | #"%" :: #" " :: rest => String.implode rest
+ | #"%" :: rest => String.implode rest
| _ => raise Error "Tptp.destLineComment";
val isLineComment = can destLineComment;
--- a/src/Tools/Metis/src/Units.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Units.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* A STORE FOR UNIT THEOREMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Units =
--- a/src/Tools/Metis/src/Units.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Units.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* A STORE FOR UNIT THEOREMS *)
-(* Copyright (c) 2001-2006 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Units :> Units =
@@ -47,7 +47,7 @@
end
end;
-val addList = foldl (fn (th,u) => add u th);
+val addList = List.foldl (fn (th,u) => add u th);
(* ------------------------------------------------------------------------- *)
(* Matching. *)
--- a/src/Tools/Metis/src/Useful.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Useful.sig Fri Sep 17 17:31:20 2010 +0200
@@ -22,8 +22,6 @@
(* Tracing. *)
(* ------------------------------------------------------------------------- *)
-val print : string -> unit (* MODIFIED by Jasmin Blanchette *)
-
val tracePrint : (string -> unit) ref
val trace : string -> unit
@@ -109,10 +107,6 @@
(* Lists: note we count elements from 0. *)
(* ------------------------------------------------------------------------- *)
-val foldl : ('a * 'b -> 'b) -> 'b -> 'a list -> 'b (* MODIFIED by Jasmin Blanchette *)
-
-val foldr : ('a * 'b -> 'b) -> 'b -> 'a list -> 'b (* MODIFIED by Jasmin Blanchette *)
-
val cons : 'a -> 'a list -> 'a list
val hdTl : 'a list -> 'a * 'a list
@@ -217,10 +211,6 @@
(* Strings. *)
(* ------------------------------------------------------------------------- *)
-val implode : char list -> string (* MODIFIED by Jasmin Blanchette *)
-
-val explode : string -> char list (* MODIFIED by Jasmin Blanchette *)
-
val rot : int -> char -> char
val charToInt : char -> int option
@@ -323,7 +313,9 @@
val try : ('a -> 'b) -> 'a -> 'b
-val chat : string -> unit
+val chat : string -> unit (* stdout *)
+
+val chide : string -> unit (* stderr *)
val warn : string -> unit
--- a/src/Tools/Metis/src/Useful.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Useful.sml Fri Sep 17 17:31:20 2010 +0200
@@ -50,9 +50,7 @@
(* Tracing. *)
(* ------------------------------------------------------------------------- *)
-val print = TextIO.print; (* MODIFIED by Jasmin Blanchette *)
-
-val tracePrint = ref print;
+val tracePrint = ref TextIO.print;
fun trace mesg = !tracePrint mesg;
@@ -172,10 +170,6 @@
(* Lists. *)
(* ------------------------------------------------------------------------- *)
-val foldl = List.foldl; (* MODIFIED by Jasmin Blanchette *)
-
-val foldr = List.foldr; (* MODIFIED by Jasmin Blanchette *)
-
fun cons x y = x :: y;
fun hdTl l = (hd l, tl l);
@@ -211,19 +205,22 @@
fun zip xs ys = zipWith pair xs ys;
-fun unzip ab =
- foldl (fn ((x, y), (xs, ys)) => (x :: xs, y :: ys)) ([], []) (rev ab);
+local
+ fun inc ((x,y),(xs,ys)) = (x :: xs, y :: ys);
+in
+ fun unzip ab = List.foldl inc ([],[]) (rev ab);
+end;
fun cartwith f =
- let
- fun aux _ res _ [] = res
- | aux xsCopy res [] (y :: yt) = aux xsCopy res xsCopy yt
- | aux xsCopy res (x :: xt) (ys as y :: _) =
- aux xsCopy (f x y :: res) xt ys
- in
- fn xs => fn ys =>
- let val xs' = rev xs in aux xs' [] xs' (rev ys) end
- end;
+ let
+ fun aux _ res _ [] = res
+ | aux xsCopy res [] (y :: yt) = aux xsCopy res xsCopy yt
+ | aux xsCopy res (x :: xt) (ys as y :: _) =
+ aux xsCopy (f x y :: res) xt ys
+ in
+ fn xs => fn ys =>
+ let val xs' = rev xs in aux xs' [] xs' (rev ys) end
+ end;
fun cart xs ys = cartwith pair xs ys;
@@ -342,15 +339,32 @@
fun delete x s = List.filter (not o equal x) s;
-fun setify s = rev (foldl (fn (v,x) => if mem v x then x else v :: x) [] s);
+local
+ fun inc (v,x) = if mem v x then x else v :: x;
+in
+ fun setify s = rev (List.foldl inc [] s);
+end;
-fun union s t = foldl (fn (v,x) => if mem v t then x else v :: x) t (rev s);
+fun union s t =
+ let
+ fun inc (v,x) = if mem v t then x else v :: x
+ in
+ List.foldl inc t (rev s)
+ end;
fun intersect s t =
- foldl (fn (v,x) => if mem v t then v :: x else x) [] (rev s);
+ let
+ fun inc (v,x) = if mem v t then v :: x else x
+ in
+ List.foldl inc [] (rev s)
+ end;
fun difference s t =
- foldl (fn (v,x) => if mem v t then x else v :: x) [] (rev s);
+ let
+ fun inc (v,x) = if mem v t then x else v :: x
+ in
+ List.foldl inc [] (rev s)
+ end;
fun subset s t = List.all (fn x => mem x t) s;
@@ -448,66 +462,43 @@
end;
local
- fun calcPrimes ps n i =
- if List.exists (fn p => divides p i) ps then calcPrimes ps n (i + 1)
+ fun calcPrimes mode ps i n =
+ if n = 0 then ps
+ else if List.exists (fn p => divides p i) ps then
+ let
+ val i = i + 1
+ and n = if mode then n else n - 1
+ in
+ calcPrimes mode ps i n
+ end
else
let
val ps = ps @ [i]
+ and i = i + 1
and n = n - 1
in
- if n = 0 then ps else calcPrimes ps n (i + 1)
+ calcPrimes mode ps i n
end;
-
- val primesList = ref [2];
in
- (* MODIFIED by Jasmin Blanchette *)
- fun primes n = CRITICAL (fn () =>
- let
- val ref ps = primesList
-
- val k = n - length ps
- in
- if k <= 0 then List.take (ps,n)
- else
- let
- val ps = calcPrimes ps k (List.last ps + 1)
+ fun primes n =
+ if n <= 0 then []
+ else calcPrimes true [2] 3 (n - 1);
- val () = primesList := ps
- in
- ps
- end
- end);
+ fun primesUpTo n =
+ if n < 2 then []
+ else calcPrimes false [2] 3 (n - 2);
end;
-(* MODIFIED by Jasmin Blanchette *)
-fun primesUpTo n = CRITICAL (fn () =>
- let
- fun f k =
- let
- val l = primes k
-
- val p = List.last l
- in
- if p < n then f (2 * k) else takeWhile (fn j => j <= n) l
- end
- in
- f 8
- end);
-
(* ------------------------------------------------------------------------- *)
(* Strings. *)
(* ------------------------------------------------------------------------- *)
-val implode = String.implode (* MODIFIED by Jasmin Blanchette *)
-
-val explode = String.explode (* MODIFIED by Jasmin Blanchette *)
-
local
fun len l = (length l, l)
- val upper = len (explode "ABCDEFGHIJKLMNOPQRSTUVWXYZ");
+ val upper = len (String.explode "ABCDEFGHIJKLMNOPQRSTUVWXYZ");
- val lower = len (explode "abcdefghijklmnopqrstuvwxyz");
+ val lower = len (String.explode "abcdefghijklmnopqrstuvwxyz");
fun rotate (n,l) c k =
List.nth (l, (k + Option.valOf (index (equal c) l)) mod n);
@@ -546,7 +537,7 @@
let
fun dup 0 l = l | dup n l = dup (n - 1) (x :: l)
in
- fn n => implode (dup n [])
+ fn n => String.implode (dup n [])
end;
fun chomp s =
@@ -558,14 +549,15 @@
end;
local
- fun chop [] = []
- | chop (l as (h :: t)) = if Char.isSpace h then chop t else l;
+ fun chop l =
+ case l of
+ [] => []
+ | h :: t => if Char.isSpace h then chop t else l;
in
- val trim = implode o chop o rev o chop o rev o explode;
+ val trim = String.implode o chop o rev o chop o rev o String.explode;
end;
-fun join _ [] = ""
- | join s (h :: t) = foldl (fn (x,y) => y ^ s ^ x) h t;
+val join = String.concatWith;
local
fun match [] l = SOME l
@@ -573,18 +565,19 @@
| match (x :: xs) (y :: ys) = if x = y then match xs ys else NONE;
fun stringify acc [] = acc
- | stringify acc (h :: t) = stringify (implode h :: acc) t;
+ | stringify acc (h :: t) = stringify (String.implode h :: acc) t;
in
fun split sep =
let
val pat = String.explode sep
+
fun div1 prev recent [] = stringify [] (rev recent :: prev)
| div1 prev recent (l as h :: t) =
case match pat l of
NONE => div1 prev (h :: recent) t
| SOME rest => div1 (rev recent :: prev) [] rest
in
- fn s => div1 [] [] (explode s)
+ fn s => div1 [] [] (String.explode s)
end;
end;
@@ -713,25 +706,30 @@
local
val generator = ref 0
-in
- (* MODIFIED by Jasmin Blanchette *)
- fun newInt () = CRITICAL (fn () =>
+
+ fun newIntThunk () =
let
val n = !generator
+
val () = generator := n + 1
in
n
- end);
+ end;
- fun newInts 0 = []
- (* MODIFIED by Jasmin Blanchette *)
- | newInts k = CRITICAL (fn () =>
+ fun newIntsThunk k () =
let
val n = !generator
+
val () = generator := n + k
in
interval n k
- end);
+ end;
+in
+ fun newInt () = Portable.critical newIntThunk ();
+
+ fun newInts k =
+ if k <= 0 then []
+ else Portable.critical (newIntsThunk k) ();
end;
fun withRef (r,new) f x =
@@ -811,10 +809,12 @@
(* Profiling and error reporting. *)
(* ------------------------------------------------------------------------- *)
-fun chat s = TextIO.output (TextIO.stdErr, s ^ "\n");
+fun chat s = TextIO.output (TextIO.stdOut, s ^ "\n");
+
+fun chide s = TextIO.output (TextIO.stdErr, s ^ "\n");
local
- fun err x s = chat (x ^ ": " ^ s);
+ fun err x s = chide (x ^ ": " ^ s);
in
fun try f x = f x
handle e as Error _ => (err "try" (errorToString e); raise e)
--- a/src/Tools/Metis/src/Waiting.sig Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Waiting.sig Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* THE WAITING SET OF CLAUSES *)
-(* Copyright (c) 2002-2007 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
signature Waiting =
--- a/src/Tools/Metis/src/Waiting.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/Waiting.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* THE WAITING SET OF CLAUSES *)
-(* Copyright (c) 2002-2007 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2002 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
structure Waiting :> Waiting =
@@ -40,17 +40,16 @@
(* ------------------------------------------------------------------------- *)
val defaultModels : modelParameters list =
- [(* MODIFIED by Jasmin Blanchette
- {model = Model.default,
+ [{model = Model.default,
initialPerturbations = 100,
maxChecks = SOME 20,
perturbations = 0,
- weight = 1.0} *)];
+ weight = 1.0}];
val default : parameters =
{symbolsWeight = 1.0,
- literalsWeight = (* 1.0 *) 0.0, (* MODIFIED by Jasmin Blanchette *)
- variablesWeight = (* 1.0 *) 0.0, (* MODIFIED by Jasmin Blanchette *)
+ literalsWeight = 1.0,
+ variablesWeight = 1.0,
models = defaultModels};
fun size (Waiting {clauses,...}) = Heap.size clauses;
@@ -163,7 +162,7 @@
val symbolsW = Math.pow (clauseSymbols lits, symbolsWeight)
val variablesW = Math.pow (clauseVariables lits, variablesWeight)
val literalsW = Math.pow (clauseLiterals lits, literalsWeight)
- val modelsW = (* checkModels models mods mcl *) 1.0 (* MODIFIED by Jasmin Blanchette *)
+ val modelsW = checkModels models mods mcl
(*MetisTrace4
val () = trace ("Waiting.clauseWeight: dist = " ^
Real.toString dist ^ "\n")
--- a/src/Tools/Metis/src/metis.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/metis.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,21 +1,6 @@
(* ========================================================================= *)
(* METIS FIRST ORDER PROVER *)
-(* *)
-(* Copyright (c) 2001 Joe Hurd *)
-(* *)
-(* Metis is free software; you can redistribute it and/or modify *)
-(* it under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation; either version 2 of the License, or *)
-(* (at your option) any later version. *)
-(* *)
-(* Metis is distributed in the hope that it will be useful, *)
-(* but WITHOUT ANY WARRANTY; without even the implied warranty of *)
-(* MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the *)
-(* GNU General Public License for more details. *)
-(* *)
-(* You should have received a copy of the GNU General Public License *)
-(* along with Metis; if not, write to the Free Software *)
-(* Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 USA *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
open Useful;
@@ -26,9 +11,9 @@
val PROGRAM = "metis";
-val VERSION = "2.2";
+val VERSION = "2.3";
-val versionString = PROGRAM^" "^VERSION^" (release 20100825)"^"\n";
+val versionString = PROGRAM^" "^VERSION^" (release 20100916)"^"\n";
(* ------------------------------------------------------------------------- *)
(* Program options. *)
@@ -146,11 +131,11 @@
local
fun display_sep () =
if notshowing_any () then ()
- else print (nChars #"-" (!Print.lineLength) ^ "\n");
+ else TextIO.print (nChars #"-" (!Print.lineLength) ^ "\n");
fun display_name filename =
if notshowing "name" then ()
- else print ("Problem: " ^ filename ^ "\n\n");
+ else TextIO.print ("Problem: " ^ filename ^ "\n\n");
fun display_goal tptp =
if notshowing "goal" then ()
@@ -158,12 +143,12 @@
let
val goal = Tptp.goal tptp
in
- print ("Goal:\n" ^ Formula.toString goal ^ "\n\n")
+ TextIO.print ("Goal:\n" ^ Formula.toString goal ^ "\n\n")
end;
fun display_clauses cls =
if notshowing "clauses" then ()
- else print ("Clauses:\n" ^ Problem.toString cls ^ "\n\n");
+ else TextIO.print ("Clauses:\n" ^ Problem.toString cls ^ "\n\n");
fun display_size cls =
if notshowing "size" then ()
@@ -174,7 +159,7 @@
val {clauses,literals,symbols,typedSymbols} = Problem.size cls
in
- print
+ TextIO.print
("Size: " ^
plural clauses "clause" ^ ", " ^
plural literals "literal" ^ ", " ^
@@ -188,12 +173,12 @@
let
val cat = Problem.categorize cls
in
- print ("Category: " ^ Problem.categoryToString cat ^ ".\n\n")
+ TextIO.print ("Category: " ^ Problem.categoryToString cat ^ ".\n\n")
end;
local
fun display_proof_start filename =
- print ("\nSZS output start CNFRefutation for " ^ filename ^ "\n");
+ TextIO.print ("\nSZS output start CNFRefutation for " ^ filename ^ "\n");
fun display_proof_body problem proofs =
let
@@ -224,7 +209,7 @@
end;
fun display_proof_end filename =
- print ("SZS output end CNFRefutation for " ^ filename ^ "\n\n");
+ TextIO.print ("SZS output end CNFRefutation for " ^ filename ^ "\n\n");
in
fun display_proof filename problem proofs =
if notshowing "proof" then ()
@@ -264,17 +249,24 @@
filename = "saturation.tptp"}
end
*)
- val () = print ("\nSZS output start Saturation for " ^ filename ^ "\n")
- val () = app (fn th => print (Thm.toString th ^ "\n")) ths
- val () = print ("SZS output end Saturation for " ^ filename ^ "\n\n")
+ val () =
+ TextIO.print
+ ("\nSZS output start Saturation for " ^ filename ^ "\n")
+
+ val () = app (fn th => TextIO.print (Thm.toString th ^ "\n")) ths
+
+ val () =
+ TextIO.print
+ ("SZS output end Saturation for " ^ filename ^ "\n\n")
in
()
end;
fun display_status filename status =
if notshowing "status" then ()
- else print ("SZS status " ^ Tptp.toStringStatus status ^
- " for " ^ filename ^ "\n");
+ else
+ TextIO.print ("SZS status " ^ Tptp.toStringStatus status ^
+ " for " ^ filename ^ "\n");
fun display_problem filename cls =
let
--- a/src/Tools/Metis/src/problems.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/problems.sml Fri Sep 17 17:31:20 2010 +0200
@@ -1,6 +1,6 @@
(* ========================================================================= *)
(* SOME SAMPLE PROBLEMS TO TEST PROOF PROCEDURES *)
-(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License *)
+(* Copyright (c) 2001 Joe Hurd, distributed under the BSD License *)
(* ========================================================================= *)
(* ========================================================================= *)
--- a/src/Tools/Metis/src/problems2tptp.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/problems2tptp.sml Fri Sep 17 17:31:20 2010 +0200
@@ -39,7 +39,7 @@
dups names
end;
-fun listProblem {name, comments = _, goal = _} = print (name ^ "\n");
+fun listProblem {name, comments = _, goal = _} = TextIO.print (name ^ "\n");
fun outputProblem outputDir {name,comments,goal} =
let
--- a/src/Tools/Metis/src/selftest.sml Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/Metis/src/selftest.sml Fri Sep 17 17:31:20 2010 +0200
@@ -52,11 +52,11 @@
| partialOrderToString NONE = "NONE";
fun SAY s =
- print
+ TextIO.print
("-------------------------------------" ^
"-------------------------------------\n" ^ s ^ "\n\n");
-fun printval p x = (print (Print.toString p x ^ "\n\n"); x);
+fun printval p x = (TextIO.print (Print.toString p x ^ "\n\n"); x);
fun mkCl p th = Clause.mk {parameters = p, id = Clause.newId (), thm = th};
@@ -92,12 +92,13 @@
fun test_fun eq p r a =
if eq r a then p a ^ "\n" else
- (print ("\n\n" ^
- "test: should have\n-->" ^ p r ^ "<--\n\n" ^
- "test: actually have\n-->" ^ p a ^ "<--\n\n");
+ (TextIO.print
+ ("\n\n" ^
+ "test: should have\n-->" ^ p r ^ "<--\n\n" ^
+ "test: actually have\n-->" ^ p a ^ "<--\n\n");
raise Fail "test: failed a test");
-fun test eq p r a = print (test_fun eq p r a ^ "\n");
+fun test eq p r a = TextIO.print (test_fun eq p r a ^ "\n");
val test_tm = test Term.equal Term.toString o Term.parse;
@@ -123,7 +124,7 @@
(fn s => test_fun equal I s ((mini_print n o Formula.fromString) s))
(prep q);
-fun test_pp q = print (testlen_pp 40 q ^ "\n");
+fun test_pp q = TextIO.print (testlen_pp 40 q ^ "\n");
val () = test_pp `3 = f x`;
@@ -568,7 +569,7 @@
val rows = alignTable format table
- val () = print (join "\n" rows ^ "\n\n")
+ val () = TextIO.print (join "\n" rows ^ "\n\n")
in
()
end;
@@ -626,8 +627,8 @@
val fm = LiteralSet.disjoin cl
in
Print.blockProgram Print.Consistent ind
- [Print.addString p,
- Print.addString (nChars #" " (ind - size p)),
+ [Print.ppString p,
+ Print.ppString (nChars #" " (ind - size p)),
Formula.pp fm]
end;
@@ -1108,13 +1109,19 @@
val _ =
test_fun equal I g (mini_print (!Print.lineLength) p)
- handle e => (print ("Error in problem " ^ name ^ "\n\n"); raise e)
+ handle e =>
+ (TextIO.print ("Error in problem " ^ name ^ "\n\n");
+ raise e)
in
(name,p) :: acc
end;
in
fun check_syntax (p : problem list) =
- (foldl check [] p; print "ok\n\n");
+ let
+ val _ = List.foldl check [] p
+ in
+ TextIO.print "ok\n\n"
+ end;
end;
val () = check_syntax problems;
@@ -1125,11 +1132,11 @@
fun tptp f =
let
- val () = print ("parsing " ^ f ^ "... ")
+ val () = TextIO.print ("parsing " ^ f ^ "... ")
val filename = "tptp/" ^ f ^ ".tptp"
val mapping = Tptp.defaultMapping
val goal = Tptp.goal (Tptp.read {filename = filename, mapping = mapping})
- val () = print "ok\n"
+ val () = TextIO.print "ok\n"
in
pvFm goal
end;
--- a/src/Tools/nbe.ML Fri Sep 17 17:11:43 2010 +0200
+++ b/src/Tools/nbe.ML Fri Sep 17 17:31:20 2010 +0200
@@ -609,7 +609,8 @@
fun static_eval_conv thy consts = Code_Simp.no_frees_conv
(lift_triv_classes_conv thy (Code_Thingol.static_eval_conv thy consts
- (K (fn program => oracle thy program (compile true thy program)))));
+ (K (fn program => let val nbe_program = compile true thy program
+ in fn thy => oracle thy program nbe_program end))));
(** setup **)