--- a/src/CCL/Wfd.thy Thu Oct 29 16:22:14 2009 +0000
+++ b/src/CCL/Wfd.thy Thu Oct 29 18:53:58 2009 +0100
@@ -440,7 +440,7 @@
fun IHinst tac rls = SUBGOAL (fn (Bi,i) =>
let val bvs = bvars Bi []
- val ihs = List.filter could_IH (Logic.strip_assums_hyp Bi)
+ val ihs = filter could_IH (Logic.strip_assums_hyp Bi)
val rnames = map (fn x=>
let val (a,b) = get_bno [] 0 x
in (List.nth(bvs,a),b) end) ihs
--- a/src/FOLP/simp.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/FOLP/simp.ML Thu Oct 29 18:53:58 2009 +0100
@@ -215,7 +215,7 @@
fun add_norm_tags congs =
let val ccs = map cong_const congs
- val new_asms = List.filter (exists not o #2)
+ val new_asms = filter (exists not o #2)
(ccs ~~ (map (map atomic o prems_of) congs));
in add_norms(congs,ccs,new_asms) end;
--- a/src/HOL/ATP_Linkup.thy Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/ATP_Linkup.thy Thu Oct 29 18:53:58 2009 +0100
@@ -91,9 +91,9 @@
subsection {* Setup of external ATPs *}
-use "Tools/res_axioms.ML" setup ResAxioms.setup
+use "Tools/res_axioms.ML" setup Res_Axioms.setup
use "Tools/res_hol_clause.ML"
-use "Tools/res_reconstruct.ML" setup ResReconstruct.setup
+use "Tools/res_reconstruct.ML" setup Res_Reconstruct.setup
use "Tools/res_atp.ML"
use "Tools/ATP_Manager/atp_wrapper.ML" setup ATP_Wrapper.setup
--- a/src/HOL/Auth/KerberosIV.thy Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Auth/KerberosIV.thy Thu Oct 29 18:53:58 2009 +0100
@@ -780,7 +780,7 @@
lemma u_NotexpiredSK_NotexpiredAK:
"\<lbrakk> \<not> expiredSK Ts evs; servKlife + Ts <= authKlife + Ta \<rbrakk>
\<Longrightarrow> \<not> expiredAK Ta evs"
- by (metis nat_add_commute le_less_trans)
+ by (metis le_less_trans)
subsection{* Reliability: friendly agents send something if something else happened*}
--- a/src/HOL/HOL.thy Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/HOL.thy Thu Oct 29 18:53:58 2009 +0100
@@ -24,6 +24,7 @@
"~~/src/Tools/coherent.ML"
"~~/src/Tools/eqsubst.ML"
"~~/src/Provers/quantifier1.ML"
+ "Tools/res_blacklist.ML"
("Tools/simpdata.ML")
"~~/src/Tools/random_word.ML"
"~~/src/Tools/atomize_elim.ML"
@@ -35,6 +36,8 @@
setup {* Intuitionistic.method_setup @{binding iprover} *}
+setup Res_Blacklist.setup
+
subsection {* Primitive logic *}
@@ -833,19 +836,14 @@
val hyp_subst_tacs = [Hypsubst.hyp_subst_tac]
end);
-structure BasicClassical: BASIC_CLASSICAL = Classical;
-open BasicClassical;
+structure Basic_Classical: BASIC_CLASSICAL = Classical;
+open Basic_Classical;
ML_Antiquote.value "claset"
(Scan.succeed "Classical.claset_of (ML_Context.the_local_context ())");
-
-structure ResBlacklist = Named_Thms
- (val name = "noatp" val description = "theorems blacklisted for ATP");
*}
-text {*ResBlacklist holds theorems blacklisted to sledgehammer.
- These theorems typically produce clauses that are prolific (match too many equality or
- membership literals) and relate to seldom-used facts. Some duplicate other rules.*}
+setup Classical.setup
setup {*
let
@@ -856,8 +854,6 @@
in
Hypsubst.hypsubst_setup
#> ContextRules.addSWrapper (fn tac => hyp_subst_tac' ORELSE' tac)
- #> Classical.setup
- #> ResBlacklist.setup
end
*}
--- a/src/HOL/Import/import_syntax.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Import/import_syntax.ML Thu Oct 29 18:53:58 2009 +0100
@@ -61,7 +61,7 @@
val thyname = get_generating_thy thy
val segname = get_import_segment thy
val (int_thms,facts) = Replay.setup_int_thms thyname thy
- val thmnames = List.filter (not o should_ignore thyname thy) facts
+ val thmnames = filter_out (should_ignore thyname thy) facts
fun replay thy =
let
val _ = ImportRecorder.load_history thyname
--- a/src/HOL/Import/shuffler.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Import/shuffler.ML Thu Oct 29 18:53:58 2009 +0100
@@ -628,7 +628,7 @@
val all_thms =
map (`Thm.get_name_hint) (maps #2 (Facts.dest_static [] (PureThy.facts_of thy)))
in
- List.filter (match_consts ignored t) all_thms
+ filter (match_consts ignored t) all_thms
end
fun gen_shuffle_tac ctxt search thms i st =
--- a/src/HOL/IsaMakefile Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/IsaMakefile Thu Oct 29 18:53:58 2009 +0100
@@ -302,6 +302,7 @@
Tools/recdef.ML \
Tools/res_atp.ML \
Tools/res_axioms.ML \
+ Tools/res_blacklist.ML \
Tools/res_clause.ML \
Tools/res_hol_clause.ML \
Tools/res_reconstruct.ML \
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Thu Oct 29 18:53:58 2009 +0100
@@ -319,7 +319,7 @@
if success then (message, SH_OK (time_isa, time_atp, theorem_names))
else (message, SH_FAIL(time_isa, time_atp))
end
- handle ResHolClause.TOO_TRIVIAL => ("trivial", SH_OK (0, 0, []))
+ handle Res_HOL_Clause.TOO_TRIVIAL => ("trivial", SH_OK (0, 0, []))
| ERROR msg => ("error: " ^ msg, SH_ERROR)
| TimeLimit.TimeOut => ("timeout", SH_ERROR)
--- a/src/HOL/Nominal/nominal_datatype.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Nominal/nominal_datatype.ML Thu Oct 29 18:53:58 2009 +0100
@@ -66,7 +66,7 @@
fun mk_case_names_exhausts descr new =
map (RuleCases.case_names o exhaust_cases descr o #1)
- (List.filter (fn ((_, (name, _, _))) => name mem_string new) descr);
+ (filter (fn ((_, (name, _, _))) => name mem_string new) descr);
end;
@@ -1166,11 +1166,11 @@
fun make_ind_prem fsT f k T ((cname, cargs), idxs) =
let
- val recs = List.filter is_rec_type cargs;
+ val recs = filter is_rec_type cargs;
val Ts = map (typ_of_dtyp descr'' sorts) cargs;
val recTs' = map (typ_of_dtyp descr'' sorts) recs;
val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts);
- val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
+ val rec_tnames = map fst (filter (is_rec_type o snd) (tnames ~~ cargs));
val frees = tnames ~~ Ts;
val frees' = partition_cargs idxs frees;
val z = (Name.variant tnames "z", fsT);
--- a/src/HOL/Nominal/nominal_primrec.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Nominal/nominal_primrec.ML Thu Oct 29 18:53:58 2009 +0100
@@ -193,7 +193,7 @@
NONE =>
let
val dummy_fns = map (fn (_, cargs) => Const (@{const_name undefined},
- replicate ((length cargs) + (length (List.filter is_rec_type cargs)))
+ replicate (length cargs + length (filter is_rec_type cargs))
dummyT ---> HOLogic.unitT)) constrs;
val _ = warning ("No function definition for datatype " ^ quote tname)
in
--- a/src/HOL/Tools/ATP_Manager/atp_manager.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/ATP_Manager/atp_manager.ML Thu Oct 29 18:53:58 2009 +0100
@@ -86,7 +86,7 @@
(* unregister ATP thread *)
-fun unregister (success, message) thread = Synchronized.change global_state
+fun unregister message thread = Synchronized.change global_state
(fn state as {manager, timeout_heap, active, cancelling, messages, store} =>
(case lookup_thread active thread of
SOME (birth_time, _, description) =>
@@ -149,7 +149,7 @@
do
(Synchronized.timed_access global_state (SOME o time_limit o #timeout_heap) action
|> these
- |> List.app (unregister (false, "Interrupted (reached timeout)"));
+ |> List.app (unregister "Interrupted (reached timeout)");
print_new_messages ();
(*give threads some time to respond to interrupt*)
OS.Process.sleep min_wait_time)
@@ -263,14 +263,11 @@
let
val _ = register birth_time death_time (Thread.self (), desc);
val problem = ATP_Wrapper.problem_of_goal (! full_types) i (ctxt, (facts, goal));
- val result =
- let val {success, message, ...} = prover (! timeout) problem;
- in (success, message) end
- handle ResHolClause.TOO_TRIVIAL => (* FIXME !? *)
- (true, "Empty clause: Try this command: " ^
- Markup.markup Markup.sendback "apply metis")
- | ERROR msg => (false, "Error: " ^ msg);
- val _ = unregister result (Thread.self ());
+ val message = #message (prover (! timeout) problem)
+ handle Res_HOL_Clause.TOO_TRIVIAL => (* FIXME !? *)
+ "Empty clause: Try this command: " ^ Markup.markup Markup.sendback "apply metis"
+ | ERROR msg => ("Error: " ^ msg);
+ val _ = unregister message (Thread.self ());
in () end)
in () end);
--- a/src/HOL/Tools/ATP_Manager/atp_minimal.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/ATP_Manager/atp_minimal.ML Thu Oct 29 18:53:58 2009 +0100
@@ -103,7 +103,7 @@
let
val _ = priority ("Testing " ^ string_of_int (length name_thms_pairs) ^ " theorems... ")
val name_thm_pairs = maps (fn (n, ths) => map (pair n) ths) name_thms_pairs
- val axclauses = ResAxioms.cnf_rules_pairs (Proof.theory_of state) name_thm_pairs
+ val axclauses = Res_Axioms.cnf_rules_pairs (Proof.theory_of state) name_thm_pairs
val {context = ctxt, facts, goal} = Proof.raw_goal state (* FIXME ?? *)
val problem =
{with_full_types = ! ATP_Manager.full_types,
@@ -138,9 +138,9 @@
val to_use =
if length ordered_used < length name_thms_pairs then
filter (fn (name1, _) => List.exists (equal name1) ordered_used) name_thms_pairs
- else
- name_thms_pairs
- val (min_thms, n) = if null to_use then ([], 0)
+ else name_thms_pairs
+ val (min_thms, n) =
+ if null to_use then ([], 0)
else minimal (test_thms (SOME filtered)) to_use
val min_names = sort_distinct string_ord (map fst min_thms)
val _ = priority (cat_lines
@@ -157,7 +157,7 @@
(NONE, "Error in prover: " ^ msg)
| (Failure, _) =>
(NONE, "Failure: No proof with the theorems supplied"))
- handle ResHolClause.TOO_TRIVIAL =>
+ handle Res_HOL_Clause.TOO_TRIVIAL =>
(SOME ([], 0), "Trivial: Try this command: " ^ Markup.markup Markup.sendback "apply metis")
| ERROR msg => (NONE, "Error: " ^ msg)
end
--- a/src/HOL/Tools/ATP_Manager/atp_wrapper.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/ATP_Manager/atp_wrapper.ML Thu Oct 29 18:53:58 2009 +0100
@@ -117,8 +117,8 @@
(* get clauses and prepare them for writing *)
val (ctxt, (chain_ths, th)) = goal;
val thy = ProofContext.theory_of ctxt;
- val chain_ths = map (Thm.put_name_hint ResReconstruct.chained_hint) chain_ths;
- val goal_cls = #1 (ResAxioms.neg_conjecture_clauses ctxt th subgoalno);
+ val chain_ths = map (Thm.put_name_hint Res_Reconstruct.chained_hint) chain_ths;
+ val goal_cls = #1 (Res_Axioms.neg_conjecture_clauses ctxt th subgoalno);
val the_filtered_clauses =
(case filtered_clauses of
NONE => relevance_filter goal goal_cls
@@ -204,14 +204,14 @@
val {with_full_types, subgoal, goal, axiom_clauses, filtered_clauses} = problem;
in
external_prover
- (ResAtp.get_relevant max_new_clauses insert_theory_const)
- (ResAtp.prepare_clauses false)
- (ResHolClause.tptp_write_file with_full_types)
+ (Res_ATP.get_relevant max_new_clauses insert_theory_const)
+ (Res_ATP.prepare_clauses false)
+ (Res_HOL_Clause.tptp_write_file with_full_types)
command
(arguments timeout)
- ResReconstruct.find_failure
- (if emit_structured_proof then ResReconstruct.structured_proof
- else ResReconstruct.lemma_list false)
+ Res_Reconstruct.find_failure
+ (if emit_structured_proof then Res_Reconstruct.structured_proof
+ else Res_Reconstruct.lemma_list false)
axiom_clauses
filtered_clauses
name
@@ -280,13 +280,13 @@
val {with_full_types, subgoal, goal, axiom_clauses, filtered_clauses} = problem
in
external_prover
- (ResAtp.get_relevant max_new_clauses insert_theory_const)
- (ResAtp.prepare_clauses true)
- (ResHolClause.dfg_write_file with_full_types)
+ (Res_ATP.get_relevant max_new_clauses insert_theory_const)
+ (Res_ATP.prepare_clauses true)
+ (Res_HOL_Clause.dfg_write_file with_full_types)
command
(arguments timeout)
- ResReconstruct.find_failure
- (ResReconstruct.lemma_list true)
+ Res_Reconstruct.find_failure
+ (Res_Reconstruct.lemma_list true)
axiom_clauses
filtered_clauses
name
--- a/src/HOL/Tools/Datatype/datatype.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/Datatype/datatype.ML Thu Oct 29 18:53:58 2009 +0100
@@ -63,9 +63,11 @@
val get_all = #types o DatatypesData.get;
val get_info = Symtab.lookup o get_all;
-fun the_info thy name = (case get_info thy name of
- SOME info => info
- | NONE => error ("Unknown datatype " ^ quote name));
+
+fun the_info thy name =
+ (case get_info thy name of
+ SOME info => info
+ | NONE => error ("Unknown datatype " ^ quote name));
fun info_of_constr thy (c, T) =
let
@@ -94,6 +96,7 @@
cases = cases |> fold Symtab.update
(map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)});
+
(* complex queries *)
fun the_spec thy dtco =
@@ -155,6 +158,7 @@
| NONE => NONE;
+
(** various auxiliary **)
(* prepare datatype specifications *)
@@ -207,6 +211,7 @@
end;
+
(* translation rules for case *)
fun make_case ctxt = DatatypeCase.make_case
@@ -228,6 +233,7 @@
[], []);
+
(** document antiquotation **)
val _ = ThyOutput.antiquotation "datatype" Args.tyname
@@ -254,15 +260,17 @@
in ThyOutput.output (ThyOutput.maybe_pretty_source (K pretty_datatype) src [()]) end);
+
(** abstract theory extensions relative to a datatype characterisation **)
-structure DatatypeInterpretation = InterpretationFun
+structure Datatype_Interpretation = Interpretation
(type T = config * string list val eq: T * T -> bool = eq_snd op =);
-fun interpretation f = DatatypeInterpretation.interpretation (uncurry f);
+fun interpretation f = Datatype_Interpretation.interpretation (uncurry f);
fun make_dt_info alt_names descr sorts induct inducts rec_names rec_rewrites
(index, (((((((((((_, (tname, _, _))), inject), distinct),
- exhaust), nchotomy), case_name), case_rewrites), case_cong), weak_case_cong), (split, split_asm))) =
+ exhaust), nchotomy), case_name), case_rewrites), case_cong), weak_case_cong),
+ (split, split_asm))) =
(tname,
{index = index,
alt_names = alt_names,
@@ -309,7 +317,8 @@
config new_type_names descr sorts inject distinct exhaust case_rewrites thy8;
val inducts = Project_Rule.projections (ProofContext.init thy2) induct;
- val dt_infos = map_index (make_dt_info alt_names flat_descr sorts induct inducts rec_names rec_rewrites)
+ val dt_infos = map_index
+ (make_dt_info alt_names flat_descr sorts induct inducts rec_names rec_rewrites)
(hd descr ~~ inject ~~ distinct ~~ exhaust ~~ nchotomys ~~
case_names ~~ case_rewrites ~~ case_congs ~~ weak_case_congs ~~ splits);
val dt_names = map fst dt_infos;
@@ -319,7 +328,7 @@
[((Binding.empty, [nth inducts index]), [Induct.induct_type tname]),
((Binding.empty, [nth exhaust index]), [Induct.cases_type tname])]) dt_names);
val unnamed_rules = map (fn induct =>
- ((Binding.empty, [induct]), [Thm.kind_internal, Induct.induct_type ""]))
+ ((Binding.empty, [induct]), [RuleCases.inner_rule, Induct.induct_type ""]))
(Library.drop (length dt_names, inducts));
in
thy9
@@ -337,14 +346,16 @@
|> snd
|> add_case_tr' case_names
|> register dt_infos
- |> DatatypeInterpretation.data (config, dt_names)
+ |> Datatype_Interpretation.data (config, dt_names)
|> pair dt_names
end;
+
(** declare existing type as datatype **)
-fun prove_rep_datatype config dt_names alt_names descr sorts raw_inject half_distinct raw_induct thy1 =
+fun prove_rep_datatype config dt_names alt_names descr sorts
+ raw_inject half_distinct raw_induct thy1 =
let
val raw_distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct;
val new_type_names = map Long_Name.base_name (the_default dt_names alt_names);
@@ -417,7 +428,8 @@
(*FIXME somehow dubious*)
in
ProofContext.theory_result
- (prove_rep_datatype config dt_names alt_names descr vs raw_inject half_distinct raw_induct)
+ (prove_rep_datatype config dt_names alt_names descr vs
+ raw_inject half_distinct raw_induct)
#-> after_qed
end;
in
@@ -430,6 +442,7 @@
val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global default_config (K I);
+
(** definitional introduction of datatypes **)
fun gen_add_datatype prep_typ config new_type_names dts thy =
@@ -445,16 +458,20 @@
val (tyvars, _, _, _)::_ = dts;
val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) =>
let val full_tname = Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname)
- in (case duplicates (op =) tvs of
- [] => if eq_set (op =) (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
- else error ("Mutually recursive datatypes must have same type parameters")
- | dups => error ("Duplicate parameter(s) for datatype " ^ quote (Binding.str_of tname) ^
- " : " ^ commas dups))
+ in
+ (case duplicates (op =) tvs of
+ [] =>
+ if eq_set (op =) (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
+ else error ("Mutually recursive datatypes must have same type parameters")
+ | dups => error ("Duplicate parameter(s) for datatype " ^ quote (Binding.str_of tname) ^
+ " : " ^ commas dups))
end) dts);
val dt_names = map fst new_dts;
- val _ = (case duplicates (op =) (map fst new_dts) @ duplicates (op =) new_type_names of
- [] => () | dups => error ("Duplicate datatypes: " ^ commas dups));
+ val _ =
+ (case duplicates (op =) (map fst new_dts) @ duplicates (op =) new_type_names of
+ [] => ()
+ | dups => error ("Duplicate datatypes: " ^ commas dups));
fun prep_dt_spec ((tvs, tname, mx, constrs), tname') (dts', constr_syntax, sorts, i) =
let
@@ -508,13 +525,15 @@
val datatype_cmd = snd ooo gen_add_datatype read_typ default_config;
+
(** package setup **)
(* setup theory *)
val setup =
trfun_setup #>
- DatatypeInterpretation.init;
+ Datatype_Interpretation.init;
+
(* outer syntax *)
--- a/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Thu Oct 29 18:53:58 2009 +0100
@@ -292,7 +292,7 @@
val case_dummy_fns = map (fn (_, (_, _, constrs)) => map (fn (_, cargs) =>
let
val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val Ts' = map mk_dummyT (List.filter is_rec_type cargs)
+ val Ts' = map mk_dummyT (filter is_rec_type cargs)
in Const (@{const_name undefined}, Ts @ Ts' ---> T')
end) constrs) descr';
@@ -305,7 +305,7 @@
val (fns1, fns2) = split_list (map (fn ((_, cargs), j) =>
let
val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val Ts' = Ts @ map mk_dummyT (List.filter is_rec_type cargs);
+ val Ts' = Ts @ map mk_dummyT (filter is_rec_type cargs);
val frees' = map (uncurry (mk_Free "x")) (Ts' ~~ (1 upto length Ts'));
val frees = Library.take (length cargs, frees');
val free = mk_Free "f" (Ts ---> T') j
--- a/src/HOL/Tools/Datatype/datatype_aux.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/Datatype/datatype_aux.ML Thu Oct 29 18:53:58 2009 +0100
@@ -231,7 +231,7 @@
fun name_of_typ (Type (s, Ts)) =
let val s' = Long_Name.base_name s
- in space_implode "_" (List.filter (not o equal "") (map name_of_typ Ts) @
+ in space_implode "_" (filter_out (equal "") (map name_of_typ Ts) @
[if Syntax.is_identifier s' then s' else "x"])
end
| name_of_typ _ = "";
@@ -272,7 +272,7 @@
fun get_arities descr = fold (fn (_, (_, _, constrs)) =>
fold (fn (_, cargs) => fold (insert op =) (map (length o fst o strip_dtyp)
- (List.filter is_rec_type cargs))) constrs) descr [];
+ (filter is_rec_type cargs))) constrs) descr [];
(* interpret construction of datatype *)
--- a/src/HOL/Tools/Datatype/datatype_codegen.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/Datatype/datatype_codegen.ML Thu Oct 29 18:53:58 2009 +0100
@@ -42,8 +42,8 @@
fun add_dt_defs thy defs dep module (descr: Datatype.descr) sorts gr =
let
- val descr' = List.filter (can (map DatatypeAux.dest_DtTFree o #2 o snd)) descr;
- val rtnames = map (#1 o snd) (List.filter (fn (_, (_, _, cs)) =>
+ val descr' = filter (can (map DatatypeAux.dest_DtTFree o #2 o snd)) descr;
+ val rtnames = map (#1 o snd) (filter (fn (_, (_, _, cs)) =>
exists (exists DatatypeAux.is_rec_type o snd) cs) descr');
val (_, (tname, _, _)) :: _ = descr';
--- a/src/HOL/Tools/Datatype/datatype_prop.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/Datatype/datatype_prop.ML Thu Oct 29 18:53:58 2009 +0100
@@ -130,11 +130,11 @@
(make_pred (body_index dt) U $ app_bnds (Free (s, T)) (length Us)))
end;
- val recs = List.filter is_rec_type cargs;
+ val recs = filter is_rec_type cargs;
val Ts = map (typ_of_dtyp descr' sorts) cargs;
val recTs' = map (typ_of_dtyp descr' sorts) recs;
val tnames = Name.variant_list pnames (make_tnames Ts);
- val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
+ val rec_tnames = map fst (filter (is_rec_type o snd) (tnames ~~ cargs));
val frees = tnames ~~ Ts;
val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs');
@@ -190,7 +190,7 @@
map (fn (_, cargs) =>
let
val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val recs = List.filter (is_rec_type o fst) (cargs ~~ Ts);
+ val recs = filter (is_rec_type o fst) (cargs ~~ Ts);
fun mk_argT (dt, T) =
binder_types T ---> List.nth (rec_result_Ts, body_index dt);
@@ -223,11 +223,11 @@
fun make_primrec T comb_t (cname, cargs) (ts, f::fs) =
let
- val recs = List.filter is_rec_type cargs;
+ val recs = filter is_rec_type cargs;
val Ts = map (typ_of_dtyp descr' sorts) cargs;
val recTs' = map (typ_of_dtyp descr' sorts) recs;
val tnames = make_tnames Ts;
- val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
+ val rec_tnames = map fst (filter (is_rec_type o snd) (tnames ~~ cargs));
val frees = map Free (tnames ~~ Ts);
val frees' = map Free (rec_tnames ~~ recTs');
--- a/src/HOL/Tools/Predicate_Compile/pred_compile_fun.ML Thu Oct 29 16:22:14 2009 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,437 +0,0 @@
-(* Author: Lukas Bulwahn, TU Muenchen
-
-Preprocessing functions to predicates
-*)
-
-signature PREDICATE_COMPILE_FUN =
-sig
-val define_predicates : (string * thm list) list -> theory -> (string * thm list) list * theory
- val rewrite_intro : theory -> thm -> thm list
- val setup_oracle : theory -> theory
- val pred_of_function : theory -> string -> string option
-end;
-
-structure Predicate_Compile_Fun : PREDICATE_COMPILE_FUN =
-struct
-
-
-(* Oracle for preprocessing *)
-
-val (oracle : (string * (cterm -> thm)) option Unsynchronized.ref) = Unsynchronized.ref NONE;
-
-fun the_oracle () =
- case !oracle of
- NONE => error "Oracle is not setup"
- | SOME (_, oracle) => oracle
-
-val setup_oracle = Thm.add_oracle (Binding.name "pred_compile_preprocessing", I) #->
- (fn ora => fn thy => let val _ = (oracle := SOME ora) in thy end)
-
-
-fun is_funtype (Type ("fun", [_, _])) = true
- | is_funtype _ = false;
-
-fun is_Type (Type _) = true
- | is_Type _ = false
-
-(* returns true if t is an application of an datatype constructor *)
-(* which then consequently would be splitted *)
-(* else false *)
-(*
-fun is_constructor thy t =
- if (is_Type (fastype_of t)) then
- (case DatatypePackage.get_datatype thy ((fst o dest_Type o fastype_of) t) of
- NONE => false
- | SOME info => (let
- val constr_consts = maps (fn (_, (_, _, constrs)) => map fst constrs) (#descr info)
- val (c, _) = strip_comb t
- in (case c of
- Const (name, _) => name mem_string constr_consts
- | _ => false) end))
- else false
-*)
-
-(* must be exported in code.ML *)
-fun is_constr thy = is_some o Code.get_datatype_of_constr thy;
-
-(* Table from constant name (string) to term of inductive predicate *)
-structure Pred_Compile_Preproc = TheoryDataFun
-(
- type T = string Symtab.table;
- val empty = Symtab.empty;
- val copy = I;
- val extend = I;
- fun merge _ = Symtab.merge (op =);
-)
-
-fun pred_of_function thy name = Symtab.lookup (Pred_Compile_Preproc.get thy) name
-
-fun defined thy = Symtab.defined (Pred_Compile_Preproc.get thy)
-
-
-fun transform_ho_typ (T as Type ("fun", _)) =
- let
- val (Ts, T') = strip_type T
- in if T' = @{typ "bool"} then T else (Ts @ [T']) ---> HOLogic.boolT end
-| transform_ho_typ t = t
-
-fun transform_ho_arg arg =
- case (fastype_of arg) of
- (T as Type ("fun", _)) =>
- (case arg of
- Free (name, _) => Free (name, transform_ho_typ T)
- | _ => error "I am surprised")
-| _ => arg
-
-fun pred_type T =
- let
- val (Ts, T') = strip_type T
- val Ts' = map transform_ho_typ Ts
- in
- (Ts' @ [T']) ---> HOLogic.boolT
- end;
-
-(* FIXME: create new predicate name -- does not avoid nameclashing *)
-fun pred_of f =
- let
- val (name, T) = dest_Const f
- in
- if (body_type T = @{typ bool}) then
- (Free (Long_Name.base_name name ^ "P", T))
- else
- (Free (Long_Name.base_name name ^ "P", pred_type T))
- end
-
-fun mk_param thy lookup_pred (t as Free (v, _)) = lookup_pred t
- | mk_param thy lookup_pred t =
- let
- val _ = tracing ("called param with " ^ (Syntax.string_of_term_global thy t))
- in if Predicate_Compile_Aux.is_predT (fastype_of t) then
- t
- else
- let
- val (vs, body) = strip_abs t
- val names = Term.add_free_names body []
- val vs_names = Name.variant_list names (map fst vs)
- val vs' = map2 (curry Free) vs_names (map snd vs)
- val body' = subst_bounds (rev vs', body)
- val (f, args) = strip_comb body'
- val resname = Name.variant (vs_names @ names) "res"
- val resvar = Free (resname, body_type (fastype_of body'))
- (*val P = case try lookup_pred f of SOME P => P | NONE => error "mk_param"
- val pred_body = list_comb (P, args @ [resvar])
- *)
- val pred_body = HOLogic.mk_eq (body', resvar)
- val param = fold_rev lambda (vs' @ [resvar]) pred_body
- in param end
- end
-(* creates the list of premises for every intro rule *)
-(* theory -> term -> (string list, term list list) *)
-
-fun dest_code_eqn eqn = let
- val (lhs, rhs) = Logic.dest_equals (Logic.unvarify (Thm.prop_of eqn))
- val (func, args) = strip_comb lhs
-in ((func, args), rhs) end;
-
-fun string_of_typ T = Syntax.string_of_typ_global @{theory} T
-
-fun string_of_term t =
- case t of
- Const (c, T) => "Const (" ^ c ^ ", " ^ string_of_typ T ^ ")"
- | Free (c, T) => "Free (" ^ c ^ ", " ^ string_of_typ T ^ ")"
- | Var ((c, i), T) => "Var ((" ^ c ^ ", " ^ string_of_int i ^ "), " ^ string_of_typ T ^ ")"
- | Bound i => "Bound " ^ string_of_int i
- | Abs (x, T, t) => "Abs (" ^ x ^ ", " ^ string_of_typ T ^ ", " ^ string_of_term t ^ ")"
- | t1 $ t2 => "(" ^ string_of_term t1 ^ ") $ (" ^ string_of_term t2 ^ ")"
-
-fun ind_package_get_nparams thy name =
- case try (Inductive.the_inductive (ProofContext.init thy)) name of
- SOME (_, result) => length (Inductive.params_of (#raw_induct result))
- | NONE => error ("No such predicate: " ^ quote name)
-
-(* TODO: does not work with higher order functions yet *)
-fun mk_rewr_eq (func, pred) =
- let
- val (argTs, resT) = (strip_type (fastype_of func))
- val nctxt =
- Name.make_context (Term.fold_aterms (fn Free (x, _) => insert (op =) x | _ => I) (func $ pred) [])
- val (argnames, nctxt') = Name.variants (replicate (length argTs) "a") nctxt
- val ([resname], nctxt'') = Name.variants ["r"] nctxt'
- val args = map Free (argnames ~~ argTs)
- val res = Free (resname, resT)
- in Logic.mk_equals
- (HOLogic.mk_eq (res, list_comb (func, args)), list_comb (pred, args @ [res]))
- end;
-
-fun has_split_rule_cname @{const_name "nat_case"} = true
- | has_split_rule_cname @{const_name "list_case"} = true
- | has_split_rule_cname _ = false
-
-fun has_split_rule_term thy (Const (@{const_name "nat_case"}, _)) = true
- | has_split_rule_term thy (Const (@{const_name "list_case"}, _)) = true
- | has_split_rule_term thy _ = false
-
-fun has_split_rule_term' thy (Const (@{const_name "If"}, _)) = true
- | has_split_rule_term' thy (Const (@{const_name "Let"}, _)) = true
- | has_split_rule_term' thy c = has_split_rule_term thy c
-
-fun prepare_split_thm ctxt split_thm =
- (split_thm RS @{thm iffD2})
- |> LocalDefs.unfold ctxt [@{thm atomize_conjL[symmetric]},
- @{thm atomize_all[symmetric]}, @{thm atomize_imp[symmetric]}]
-
-fun find_split_thm thy (Const (name, typ)) =
- let
- fun split_name str =
- case first_field "." str
- of (SOME (field, rest)) => field :: split_name rest
- | NONE => [str]
- val splitted_name = split_name name
- in
- if length splitted_name > 0 andalso
- String.isSuffix "_case" (List.last splitted_name)
- then
- (List.take (splitted_name, length splitted_name - 1)) @ ["split"]
- |> space_implode "."
- |> PureThy.get_thm thy
- |> SOME
- handle ERROR msg => NONE
- else NONE
- end
- | find_split_thm _ _ = NONE
-
-fun find_split_thm' thy (Const (@{const_name "If"}, _)) = SOME @{thm split_if}
- | find_split_thm' thy (Const (@{const_name "Let"}, _)) = SOME @{thm refl} (* TODO *)
- | find_split_thm' thy c = find_split_thm thy c
-
-fun strip_all t = (Term.strip_all_vars t, Term.strip_all_body t)
-
-fun folds_map f xs y =
- let
- fun folds_map' acc [] y = [(rev acc, y)]
- | folds_map' acc (x :: xs) y =
- maps (fn (x, y) => folds_map' (x :: acc) xs y) (f x y)
- in
- folds_map' [] xs y
- end;
-
-fun mk_prems thy (lookup_pred, get_nparams) t (names, prems) =
- let
- fun mk_prems' (t as Const (name, T)) (names, prems) =
- if is_constr thy name orelse (is_none (try lookup_pred t)) then
- [(t, (names, prems))]
- else [(lookup_pred t, (names, prems))]
- | mk_prems' (t as Free (f, T)) (names, prems) =
- [(lookup_pred t, (names, prems))]
- | mk_prems' (t as Abs _) (names, prems) =
- if Predicate_Compile_Aux.is_predT (fastype_of t) then
- [(t, (names, prems))] else error "mk_prems': Abs "
- (* mk_param *)
- | mk_prems' t (names, prems) =
- if Predicate_Compile_Aux.is_constrt thy t then
- [(t, (names, prems))]
- else
- if has_split_rule_term' thy (fst (strip_comb t)) then
- let
- val (f, args) = strip_comb t
- val split_thm = prepare_split_thm (ProofContext.init thy) (the (find_split_thm' thy f))
- (* TODO: contextify things - this line is to unvarify the split_thm *)
- (*val ((_, [isplit_thm]), _) = Variable.import true [split_thm] (ProofContext.init thy)*)
- val (assms, concl) = Logic.strip_horn (Thm.prop_of split_thm)
- val (P, [split_t]) = strip_comb (HOLogic.dest_Trueprop concl)
- val subst = Pattern.match thy (split_t, t) (Vartab.empty, Vartab.empty)
- val (_, split_args) = strip_comb split_t
- val match = split_args ~~ args
- fun mk_prems_of_assm assm =
- let
- val (vTs, assm') = strip_all (Envir.beta_norm (Envir.subst_term subst assm))
- val var_names = Name.variant_list names (map fst vTs)
- val vars = map Free (var_names ~~ (map snd vTs))
- val (prems', pre_res) = Logic.strip_horn (subst_bounds (rev vars, assm'))
- val (_, [inner_t]) = strip_comb (HOLogic.dest_Trueprop pre_res)
- in
- mk_prems' inner_t (var_names @ names, prems' @ prems)
- end
- in
- maps mk_prems_of_assm assms
- end
- else
- let
- val (f, args) = strip_comb t
- (* TODO: special procedure for higher-order functions: split arguments in
- simple types and function types *)
- val resname = Name.variant names "res"
- val resvar = Free (resname, body_type (fastype_of t))
- val names' = resname :: names
- fun mk_prems'' (t as Const (c, _)) =
- if is_constr thy c orelse (is_none (try lookup_pred t)) then
- folds_map mk_prems' args (names', prems) |>
- map
- (fn (argvs, (names'', prems')) =>
- let
- val prem = HOLogic.mk_Trueprop (HOLogic.mk_eq (resvar, list_comb (f, argvs)))
- in (names'', prem :: prems') end)
- else
- let
- val pred = lookup_pred t
- val nparams = get_nparams pred
- val (params, args) = chop nparams args
- val params' = map (mk_param thy lookup_pred) params
- in
- folds_map mk_prems' args (names', prems)
- |> map (fn (argvs, (names'', prems')) =>
- let
- val prem = HOLogic.mk_Trueprop (list_comb (pred, params' @ argvs @ [resvar]))
- in (names'', prem :: prems') end)
- end
- | mk_prems'' (t as Free (_, _)) =
- let
- (* higher order argument call *)
- val pred = lookup_pred t
- in
- folds_map mk_prems' args (resname :: names, prems)
- |> map (fn (argvs, (names', prems')) =>
- let
- val prem = HOLogic.mk_Trueprop (list_comb (pred, argvs @ [resvar]))
- in (names', prem :: prems') end)
- end
- | mk_prems'' t =
- error ("Invalid term: " ^ Syntax.string_of_term_global thy t)
- in
- map (pair resvar) (mk_prems'' f)
- end
- in
- mk_prems' t (names, prems)
- end;
-
-(* assumption: mutual recursive predicates all have the same parameters. *)
-fun define_predicates specs thy =
- if forall (fn (const, _) => member (op =) (Symtab.keys (Pred_Compile_Preproc.get thy)) const) specs then
- ([], thy)
- else
- let
- val consts = map fst specs
- val eqns = maps snd specs
- (*val eqns = maps (Predicate_Compile_Preproc_Data.get_specification thy) consts*)
- (* create prednames *)
- val ((funs, argss), rhss) = map_split dest_code_eqn eqns |>> split_list
- val argss' = map (map transform_ho_arg) argss
- val pnames = map dest_Free (distinct (op =) (maps (filter (is_funtype o fastype_of)) argss'))
- val preds = map pred_of funs
- val prednames = map (fst o dest_Free) preds
- val funnames = map (fst o dest_Const) funs
- val fun_pred_names = (funnames ~~ prednames)
- (* mapping from term (Free or Const) to term *)
- fun lookup_pred (Const (name, T)) =
- (case (Symtab.lookup (Pred_Compile_Preproc.get thy) name) of
- SOME c => Const (c, pred_type T)
- | NONE =>
- (case AList.lookup op = fun_pred_names name of
- SOME f => Free (f, pred_type T)
- | NONE => Const (name, T)))
- | lookup_pred (Free (name, T)) =
- if member op = (map fst pnames) name then
- Free (name, transform_ho_typ T)
- else
- Free (name, T)
- | lookup_pred t =
- error ("lookup function is not defined for " ^ Syntax.string_of_term_global thy t)
-
- (* mapping from term (predicate term, not function term!) to int *)
- fun get_nparams (Const (name, _)) =
- the_default 0 (try (ind_package_get_nparams thy) name)
- | get_nparams (Free (name, _)) =
- (if member op = prednames name then
- length pnames
- else 0)
- | get_nparams t = error ("No parameters for " ^ (Syntax.string_of_term_global thy t))
-
- (* create intro rules *)
-
- fun mk_intros ((func, pred), (args, rhs)) =
- if (body_type (fastype_of func) = @{typ bool}) then
- (*TODO: preprocess predicate definition of rhs *)
- [Logic.list_implies ([HOLogic.mk_Trueprop rhs], HOLogic.mk_Trueprop (list_comb (pred, args)))]
- else
- let
- val names = Term.add_free_names rhs []
- in mk_prems thy (lookup_pred, get_nparams) rhs (names, [])
- |> map (fn (resultt, (names', prems)) =>
- Logic.list_implies (prems, HOLogic.mk_Trueprop (list_comb (pred, args @ [resultt]))))
- end
- fun mk_rewr_thm (func, pred) = @{thm refl}
- in
- case try (maps mk_intros) ((funs ~~ preds) ~~ (argss' ~~ rhss)) of
- NONE => ([], thy)
- | SOME intr_ts =>
- if is_some (try (map (cterm_of thy)) intr_ts) then
- let
- val (ind_result, thy') =
- Inductive.add_inductive_global (serial ())
- {quiet_mode = false, verbose = false, kind = Thm.internalK,
- alt_name = Binding.empty, coind = false, no_elim = false,
- no_ind = false, skip_mono = false, fork_mono = false}
- (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (distinct (op =) (map dest_Free preds)))
- pnames
- (map (fn x => (Attrib.empty_binding, x)) intr_ts)
- [] thy
- val prednames = map (fst o dest_Const) (#preds ind_result)
- (* val rewr_thms = map mk_rewr_eq ((distinct (op =) funs) ~~ (#preds ind_result)) *)
- (* add constants to my table *)
- val specs = map (fn predname => (predname, filter (Predicate_Compile_Aux.is_intro predname) (#intrs ind_result))) prednames
- val thy'' = Pred_Compile_Preproc.map (fold Symtab.update_new (consts ~~ prednames)) thy'
- in
- (specs, thy'')
- end
- else
- let
- val _ = tracing "Introduction rules of function_predicate are not welltyped"
- in ([], thy) end
- end
-
-(* preprocessing intro rules - uses oracle *)
-
-(* theory -> thm -> thm *)
-fun rewrite_intro thy intro =
- let
- fun lookup_pred (Const (name, T)) =
- (case (Symtab.lookup (Pred_Compile_Preproc.get thy) name) of
- SOME c => Const (c, pred_type T)
- | NONE => error ("Function " ^ name ^ " is not inductified"))
- | lookup_pred (Free (name, T)) = Free (name, T)
- | lookup_pred _ = error "lookup function is not defined!"
-
- fun get_nparams (Const (name, _)) =
- the_default 0 (try (ind_package_get_nparams thy) name)
- | get_nparams (Free _) = 0
- | get_nparams t = error ("No parameters for " ^ (Syntax.string_of_term_global thy t))
-
- val intro_t = (Logic.unvarify o prop_of) intro
- val (prems, concl) = Logic.strip_horn intro_t
- val frees = map fst (Term.add_frees intro_t [])
- fun rewrite prem names =
- let
- val t = (HOLogic.dest_Trueprop prem)
- val (lit, mk_lit) = case try HOLogic.dest_not t of
- SOME t => (t, HOLogic.mk_not)
- | NONE => (t, I)
- val (P, args) = (strip_comb lit)
- in
- folds_map (
- fn t => if (is_funtype (fastype_of t)) then (fn x => [(t, x)])
- else mk_prems thy (lookup_pred, get_nparams) t) args (names, [])
- |> map (fn (resargs, (names', prems')) =>
- let
- val prem' = HOLogic.mk_Trueprop (mk_lit (list_comb (P, resargs)))
- in (prem'::prems', names') end)
- end
- val intro_ts' = folds_map rewrite prems frees
- |> maps (fn (prems', frees') =>
- rewrite concl frees'
- |> map (fn (concl'::conclprems, _) =>
- Logic.list_implies ((flat prems') @ conclprems, concl')))
- in
- map (Drule.standard o the_oracle () o cterm_of thy) intro_ts'
- end;
-
-end;
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile.ML Thu Oct 29 18:53:58 2009 +0100
@@ -156,10 +156,35 @@
local structure P = OuterParse
in
+(*val parse_argmode' = P.nat >> rpair NONE || P.$$$ "(" |-- P.enum1 "," --| P.$$$ ")"*)
+datatype raw_argmode = Argmode of string | Argmode_Tuple of string list
+
+val parse_argmode' =
+ ((Args.$$$ "i" || Args.$$$ "o") >> Argmode) ||
+ (Args.$$$ "(" |-- P.enum1 "," (Args.$$$ "i" || Args.$$$ "o") --| Args.$$$ ")" >> Argmode_Tuple)
+
+fun mk_numeral_mode ss = flat (map_index (fn (i, s) => if s = "i" then [i + 1] else []) ss)
+
+val parse_smode' = P.$$$ "[" |-- P.enum1 "," parse_argmode' --| P.$$$ "]"
+ >> (fn m => flat (map_index
+ (fn (i, Argmode s) => if s = "i" then [(i + 1, NONE)] else []
+ | (i, Argmode_Tuple ss) => [(i + 1, SOME (mk_numeral_mode ss))]) m))
+
+val parse_smode = (P.$$$ "[" |-- P.enum "," P.nat --| P.$$$ "]")
+ >> map (rpair (NONE : int list option))
+
+fun gen_parse_mode smode_parser =
+ (Scan.optional
+ ((P.enum "=>" ((Args.$$$ "X" >> K NONE) || (smode_parser >> SOME))) --| Args.$$$ "==>") [])
+ -- smode_parser
+
+val parse_mode = gen_parse_mode parse_smode
+
+val parse_mode' = gen_parse_mode parse_smode'
+
val opt_modes =
Scan.optional (P.$$$ "(" |-- Args.$$$ "mode" |-- P.$$$ ":" |--
- P.enum1 "," (P.$$$ "[" |-- P.enum "," P.nat --| P.$$$ "]")
- --| P.$$$ ")" >> SOME) NONE
+ P.enum1 "," (parse_mode || parse_mode') --| P.$$$ ")" >> SOME) NONE
val scan_params =
let
@@ -170,8 +195,7 @@
val _ = OuterSyntax.local_theory_to_proof "code_pred"
"prove equations for predicate specified by intro/elim rules"
- OuterKeyword.thy_goal (opt_modes -- scan_params -- P.term_group >>
- code_pred_cmd)
+ OuterKeyword.thy_goal (opt_modes -- scan_params -- P.term_group >> code_pred_cmd)
end
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Thu Oct 29 18:53:58 2009 +0100
@@ -9,6 +9,29 @@
structure Predicate_Compile_Aux =
struct
+
+(* mode *)
+
+type smode = (int * int list option) list
+type mode = smode option list * smode
+datatype tmode = Mode of mode * smode * tmode option list;
+
+fun string_of_smode js =
+ commas (map
+ (fn (i, is) =>
+ string_of_int i ^ (case is of NONE => ""
+ | SOME is => "p" ^ enclose "[" "]" (commas (map string_of_int is)))) js)
+
+fun string_of_mode (iss, is) = space_implode " -> " (map
+ (fn NONE => "X"
+ | SOME js => enclose "[" "]" (string_of_smode js))
+ (iss @ [SOME is]));
+
+fun string_of_tmode (Mode (predmode, termmode, param_modes)) =
+ "predmode: " ^ (string_of_mode predmode) ^
+ (if null param_modes then "" else
+ "; " ^ "params: " ^ commas (map (the_default "NONE" o Option.map string_of_tmode) param_modes))
+
(* general syntactic functions *)
(*Like dest_conj, but flattens conjunctions however nested*)
@@ -136,7 +159,7 @@
(* Different options for compiler *)
datatype options = Options of {
- expected_modes : (string * int list list) option,
+ expected_modes : (string * mode list) option,
show_steps : bool,
show_proof_trace : bool,
show_intermediate_results : bool,
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Thu Oct 29 18:53:58 2009 +0100
@@ -9,24 +9,20 @@
val setup: theory -> theory
val code_pred: Predicate_Compile_Aux.options -> string -> Proof.context -> Proof.state
val code_pred_cmd: Predicate_Compile_Aux.options -> string -> Proof.context -> Proof.state
- type smode = (int * int list option) list
- type mode = smode option list * smode
- datatype tmode = Mode of mode * smode * tmode option list;
val register_predicate : (string * thm list * thm * int) -> theory -> theory
val register_intros : string * thm list -> theory -> theory
val is_registered : theory -> string -> bool
- val predfun_intro_of: theory -> string -> mode -> thm
- val predfun_elim_of: theory -> string -> mode -> thm
- val predfun_name_of: theory -> string -> mode -> string
+ val predfun_intro_of: theory -> string -> Predicate_Compile_Aux.mode -> thm
+ val predfun_elim_of: theory -> string -> Predicate_Compile_Aux.mode -> thm
+ val predfun_name_of: theory -> string -> Predicate_Compile_Aux.mode -> string
val all_preds_of : theory -> string list
- val modes_of: theory -> string -> mode list
- val depth_limited_modes_of: theory -> string -> mode list
- val depth_limited_function_name_of : theory -> string -> mode -> string
- val generator_modes_of: theory -> string -> mode list
- val generator_name_of : theory -> string -> mode -> string
- val all_modes_of : theory -> (string * mode list) list
- val all_generator_modes_of : theory -> (string * mode list) list
- val string_of_mode : mode -> string
+ val modes_of: theory -> string -> Predicate_Compile_Aux.mode list
+ val depth_limited_modes_of: theory -> string -> Predicate_Compile_Aux.mode list
+ val depth_limited_function_name_of : theory -> string -> Predicate_Compile_Aux.mode -> string
+ val generator_modes_of: theory -> string -> Predicate_Compile_Aux.mode list
+ val generator_name_of : theory -> string -> Predicate_Compile_Aux.mode -> string
+ val all_modes_of : theory -> (string * Predicate_Compile_Aux.mode list) list
+ val all_generator_modes_of : theory -> (string * Predicate_Compile_Aux.mode list) list
val intros_of: theory -> string -> thm list
val nparams_of: theory -> string -> int
val add_intro: thm -> theory -> theory
@@ -67,8 +63,6 @@
(* debug stuff *)
-fun tracing s = (if ! Toplevel.debug then Output.tracing s else ());
-
fun print_tac s = Seq.single;
fun print_tac' options s =
@@ -140,9 +134,6 @@
type mode = arg_mode list
type tmode = Mode of mode *
*)
-type smode = (int * int list option) list
-type mode = smode option list * smode;
-datatype tmode = Mode of mode * smode * tmode option list;
fun gen_split_smode (mk_tuple, strip_tuple) smode ts =
let
@@ -165,32 +156,16 @@
(split_smode' smode (i+1) ts)
in split_smode' smode 1 ts end
-val split_smode = gen_split_smode (HOLogic.mk_tuple, HOLogic.strip_tuple)
-val split_smodeT = gen_split_smode (HOLogic.mk_tupleT, HOLogic.strip_tupleT)
+fun split_smode smode ts = gen_split_smode (HOLogic.mk_tuple, HOLogic.strip_tuple) smode ts
+fun split_smodeT smode ts = gen_split_smode (HOLogic.mk_tupleT, HOLogic.strip_tupleT) smode ts
fun gen_split_mode split_smode (iss, is) ts =
let
val (t1, t2) = chop (length iss) ts
in (t1, split_smode is t2) end
-val split_mode = gen_split_mode split_smode
-val split_modeT = gen_split_mode split_smodeT
-
-fun string_of_smode js =
- commas (map
- (fn (i, is) =>
- string_of_int i ^ (case is of NONE => ""
- | SOME is => "p" ^ enclose "[" "]" (commas (map string_of_int is)))) js)
-
-fun string_of_mode (iss, is) = space_implode " -> " (map
- (fn NONE => "X"
- | SOME js => enclose "[" "]" (string_of_smode js))
- (iss @ [SOME is]));
-
-fun string_of_tmode (Mode (predmode, termmode, param_modes)) =
- "predmode: " ^ (string_of_mode predmode) ^
- (if null param_modes then "" else
- "; " ^ "params: " ^ commas (map (the_default "NONE" o Option.map string_of_tmode) param_modes))
+fun split_mode (iss, is) ts = gen_split_mode split_smode (iss, is) ts
+fun split_modeT (iss, is) ts = gen_split_mode split_smodeT (iss, is) ts
datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term
| Generator of (string * typ);
@@ -333,7 +308,7 @@
fun print_modes options modes =
if show_modes options then
- Output.tracing ("Inferred modes:\n" ^
+ tracing ("Inferred modes:\n" ^
cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
string_of_mode ms)) modes))
else ()
@@ -344,7 +319,7 @@
^ (string_of_entry pred mode entry)
fun print_pred (pred, modes) =
"predicate " ^ pred ^ ": " ^ cat_lines (map (print_mode pred) modes)
- val _ = Output.tracing (cat_lines (map print_pred pred_mode_table))
+ val _ = tracing (cat_lines (map print_pred pred_mode_table))
in () end;
fun string_of_prem thy (Prem (ts, p)) =
@@ -423,10 +398,10 @@
case expected_modes options of
SOME (s, ms) => (case AList.lookup (op =) modes s of
SOME modes =>
- if not (eq_set (op =) (map (map (rpair NONE)) ms, map snd modes)) then
+ if not (eq_set (op =) (ms, modes)) then
error ("expected modes were not inferred:\n"
- ^ "inferred modes for " ^ s ^ ": "
- ^ commas (map ((enclose "[" "]") o string_of_smode o snd) modes))
+ ^ "inferred modes for " ^ s ^ ": " ^ commas (map string_of_mode modes)
+ ^ "\n expected modes for " ^ s ^ ": " ^ commas (map string_of_mode ms))
else ()
| NONE => ())
| NONE => ()
@@ -444,7 +419,7 @@
val rec_consts = fold add_term_consts_2 cs' [];
val intr_consts = fold add_term_consts_2 intr_ts' [];
fun unify (cname, cT) =
- let val consts = map snd (List.filter (fn c => fst c = cname) intr_consts)
+ let val consts = map snd (filter (fn c => fst c = cname) intr_consts)
in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end;
val (env, _) = fold unify rec_consts (Vartab.empty, i');
val subst = map_types (Envir.norm_type env)
@@ -661,7 +636,7 @@
fun cons_intro gr =
case try (Graph.get_node gr) name of
SOME pred_data => Graph.map_node name (map_pred_data
- (apfst (fn (intros, elim, nparams) => (thm::intros, elim, nparams)))) gr
+ (apfst (fn (intros, elim, nparams) => (intros @ [thm], elim, nparams)))) gr
| NONE =>
let
val nparams = the_default (guess_nparams T) (try (#nparams o rep_pred_data o (fetch_pred_data thy)) name)
@@ -1052,16 +1027,16 @@
fun print_failed_mode options thy modes p m rs i =
if show_mode_inference options then
let
- val _ = Output.tracing ("Clause " ^ string_of_int (i + 1) ^ " of " ^
+ val _ = tracing ("Clause " ^ string_of_int (i + 1) ^ " of " ^
p ^ " violates mode " ^ string_of_mode m)
- val _ = Output.tracing (string_of_clause thy p (nth rs i))
+ val _ = tracing (string_of_clause thy p (nth rs i))
in () end
else ()
fun check_modes_pred options with_generator thy param_vs clauses modes gen_modes (p, ms) =
let
val rs = case AList.lookup (op =) clauses p of SOME rs => rs | NONE => []
- in (p, List.filter (fn m => case find_index
+ in (p, filter (fn m => case find_index
(is_none o check_mode_clause with_generator thy param_vs modes gen_modes m) rs of
~1 => true
| i => (print_failed_mode options thy modes p m rs i; false)) ms)
@@ -1191,6 +1166,28 @@
(t, names)
end;
+structure Comp_Mod =
+struct
+
+datatype comp_modifiers = Comp_Modifiers of
+{
+ const_name_of : theory -> string -> Predicate_Compile_Aux.mode -> string,
+ funT_of : compilation_funs -> mode -> typ -> typ,
+ additional_arguments : string list -> term list,
+ wrap_compilation : compilation_funs -> string -> typ -> mode -> term list -> term -> term,
+ transform_additional_arguments : indprem -> term list -> term list
+}
+
+fun dest_comp_modifiers (Comp_Modifiers c) = c
+
+val const_name_of = #const_name_of o dest_comp_modifiers
+val funT_of = #funT_of o dest_comp_modifiers
+val additional_arguments = #additional_arguments o dest_comp_modifiers
+val wrap_compilation = #wrap_compilation o dest_comp_modifiers
+val transform_additional_arguments = #transform_additional_arguments o dest_comp_modifiers
+
+end;
+
fun compile_arg compilation_modifiers compfuns additional_arguments thy param_vs iss arg =
let
fun map_params (t as Free (f, T)) =
@@ -1198,7 +1195,7 @@
case (the (AList.lookup (op =) (param_vs ~~ iss) f)) of
SOME is =>
let
- val T' = #funT_of compilation_modifiers compfuns ([], is) T
+ val T' = Comp_Mod.funT_of compilation_modifiers compfuns ([], is) T
in fst (mk_Eval_of additional_arguments ((Free (f, T'), T), SOME is) []) end
| NONE => t
else t
@@ -1248,9 +1245,9 @@
val params' = map (compile_param compilation_modifiers compfuns thy) (ms ~~ params)
val f' =
case f of
- Const (name, T) => Const (#const_name_of compilation_modifiers thy name mode,
- #funT_of compilation_modifiers compfuns mode T)
- | Free (name, T) => Free (name, #funT_of compilation_modifiers compfuns mode T)
+ Const (name, T) => Const (Comp_Mod.const_name_of compilation_modifiers thy name mode,
+ Comp_Mod.funT_of compilation_modifiers compfuns mode T)
+ | Free (name, T) => Free (name, Comp_Mod.funT_of compilation_modifiers compfuns mode T)
| _ => error ("PredicateCompiler: illegal parameter term")
in
list_comb (f', params' @ args')
@@ -1262,13 +1259,13 @@
let
val params' = map (compile_param compilation_modifiers compfuns thy) (ms ~~ params)
(*val mk_fun_of = if depth_limited then mk_depth_limited_fun_of else mk_fun_of*)
- val name' = #const_name_of compilation_modifiers thy name mode
- val T' = #funT_of compilation_modifiers compfuns mode T
+ val name' = Comp_Mod.const_name_of compilation_modifiers thy name mode
+ val T' = Comp_Mod.funT_of compilation_modifiers compfuns mode T
in
(list_comb (Const (name', T'), params' @ inargs @ additional_arguments))
end
| (Free (name, T), params) =>
- list_comb (Free (name, #funT_of compilation_modifiers compfuns mode T), params @ inargs @ additional_arguments)
+ list_comb (Free (name, Comp_Mod.funT_of compilation_modifiers compfuns mode T), params @ inargs @ additional_arguments)
fun compile_clause compilation_modifiers compfuns thy all_vs param_vs additional_arguments (iss, is) inp (ts, moded_ps) =
let
@@ -1302,7 +1299,7 @@
val (out_ts'', (names'', constr_vs')) = fold_map distinct_v
out_ts' ((names', map (rpair []) vs))
val additional_arguments' =
- #transform_additional_arguments compilation_modifiers p additional_arguments
+ Comp_Mod.transform_additional_arguments compilation_modifiers p additional_arguments
val (compiled_clause, rest) = case p of
Prem (us, t) =>
let
@@ -1356,7 +1353,7 @@
val (Ts1, Ts2) = chop (length (fst mode)) (binder_types T)
val (Us1, Us2) = split_smodeT (snd mode) Ts2
val Ts1' =
- map2 (fn NONE => I | SOME is => #funT_of compilation_modifiers compfuns ([], is)) (fst mode) Ts1
+ map2 (fn NONE => I | SOME is => Comp_Mod.funT_of compilation_modifiers compfuns ([], is)) (fst mode) Ts1
fun mk_input_term (i, NONE) =
[Free (Name.variant (all_vs @ param_vs) ("x" ^ string_of_int i), nth Ts2 (i - 1))]
| mk_input_term (i, SOME pis) = case HOLogic.strip_tupleT (nth Ts2 (i - 1)) of
@@ -1370,17 +1367,17 @@
else [HOLogic.mk_tuple (map Free (vnames ~~ map (fn j => nth Ts (j - 1)) pis))] end
val in_ts = maps mk_input_term (snd mode)
val params = map2 (fn s => fn T => Free (s, T)) param_vs Ts1'
- val additional_arguments = #additional_arguments compilation_modifiers (all_vs @ param_vs)
+ val additional_arguments = Comp_Mod.additional_arguments compilation_modifiers (all_vs @ param_vs)
val cl_ts =
map (compile_clause compilation_modifiers compfuns
thy all_vs param_vs additional_arguments mode (HOLogic.mk_tuple in_ts)) moded_cls;
- val compilation = #wrap_compilation compilation_modifiers compfuns s T mode additional_arguments
+ val compilation = Comp_Mod.wrap_compilation compilation_modifiers compfuns s T mode additional_arguments
(if null cl_ts then
mk_bot compfuns (HOLogic.mk_tupleT Us2)
else foldr1 (mk_sup compfuns) cl_ts)
val fun_const =
- Const (#const_name_of compilation_modifiers thy s mode,
- #funT_of compilation_modifiers compfuns mode T)
+ Const (Comp_Mod.const_name_of compilation_modifiers thy s mode,
+ Comp_Mod.funT_of compilation_modifiers compfuns mode T)
in
HOLogic.mk_Trueprop
(HOLogic.mk_eq (list_comb (fun_const, params @ in_ts @ additional_arguments), compilation))
@@ -2139,31 +2136,47 @@
(** main function of predicate compiler **)
+datatype steps = Steps of
+ {
+ compile_preds : theory -> string list -> string list -> (string * typ) list
+ -> (moded_clause list) pred_mode_table -> term pred_mode_table,
+ create_definitions: (string * typ) list -> string * mode list -> theory -> theory,
+ infer_modes : options -> theory -> (string * mode list) list -> (string * mode list) list
+ -> string list -> (string * (term list * indprem list) list) list
+ -> moded_clause list pred_mode_table,
+ prove : options -> theory -> (string * (term list * indprem list) list) list
+ -> (string * typ) list -> (string * mode list) list
+ -> moded_clause list pred_mode_table -> term pred_mode_table -> thm pred_mode_table,
+ are_not_defined : theory -> string list -> bool,
+ qname : bstring
+ }
+
+
fun add_equations_of steps options prednames thy =
let
+ fun dest_steps (Steps s) = s
val _ = print_step options ("Starting predicate compiler for predicates " ^ commas prednames ^ "...")
- val _ = tracing (commas (map (Display.string_of_thm_global thy) (maps (intros_of thy) prednames)))
(*val _ = check_intros_elim_match thy prednames*)
(*val _ = map (check_format_of_intro_rule thy) (maps (intros_of thy) prednames)*)
val (preds, nparams, all_vs, param_vs, extra_modes, clauses, all_modes) =
prepare_intrs thy prednames (maps (intros_of thy) prednames)
val _ = print_step options "Infering modes..."
- val moded_clauses = #infer_modes steps options thy extra_modes all_modes param_vs clauses
+ val moded_clauses = #infer_modes (dest_steps steps) options thy extra_modes all_modes param_vs clauses
val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
val _ = check_expected_modes options modes
val _ = print_modes options modes
(*val _ = print_moded_clauses thy moded_clauses*)
val _ = print_step options "Defining executable functions..."
- val thy' = fold (#create_definitions steps preds) modes thy
+ val thy' = fold (#create_definitions (dest_steps steps) preds) modes thy
|> Theory.checkpoint
val _ = print_step options "Compiling equations..."
val compiled_terms =
- (#compile_preds steps) thy' all_vs param_vs preds moded_clauses
+ #compile_preds (dest_steps steps) thy' all_vs param_vs preds moded_clauses
val _ = print_compiled_terms options thy' compiled_terms
val _ = print_step options "Proving equations..."
- val result_thms = #prove steps options thy' clauses preds (extra_modes @ modes)
+ val result_thms = #prove (dest_steps steps) options thy' clauses preds (extra_modes @ modes)
moded_clauses compiled_terms
- val qname = #qname steps
+ val qname = #qname (dest_steps steps)
val attrib = fn thy => Attrib.attribute_i thy (Attrib.internal (K (Thm.declaration_attribute
(fn thm => Context.mapping (Code.add_eqn thm) I))))
val thy'' = fold (fn (name, result_thms) => fn thy => snd (PureThy.add_thmss
@@ -2181,7 +2194,7 @@
SOME v => (G, v)
| NONE => (Graph.new_node (key, value_of key) G, value_of key)
val (G'', visited') = fold (extend' value_of edges_of) (subtract (op =) visited (edges_of (key, v)))
- (G', key :: visited)
+ (G', key :: visited)
in
(fold (Graph.add_edge o (pair key)) (edges_of (key, v)) G'', visited')
end;
@@ -2190,6 +2203,7 @@
fun gen_add_equations steps options names thy =
let
+ fun dest_steps (Steps s) = s
val thy' = PredData.map (fold (extend (fetch_pred_data thy) (depending_preds_of thy)) names) thy
|> Theory.checkpoint;
fun strong_conn_of gr keys =
@@ -2197,24 +2211,25 @@
val scc = strong_conn_of (PredData.get thy') names
val thy'' = fold_rev
(fn preds => fn thy =>
- if #are_not_defined steps thy preds then
+ if #are_not_defined (dest_steps steps) thy preds then
add_equations_of steps options preds thy else thy)
scc thy' |> Theory.checkpoint
in thy'' end
(* different instantiantions of the predicate compiler *)
-val predicate_comp_modifiers =
- {const_name_of = predfun_name_of,
- funT_of = funT_of,
+val predicate_comp_modifiers = Comp_Mod.Comp_Modifiers
+ {const_name_of = predfun_name_of : (theory -> string -> mode -> string),
+ funT_of = funT_of : (compilation_funs -> mode -> typ -> typ),
additional_arguments = K [],
- wrap_compilation = K (K (K (K (K I)))),
- transform_additional_arguments = K I
+ wrap_compilation = K (K (K (K (K I))))
+ : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
+ transform_additional_arguments = K I : (indprem -> term list -> term list)
}
-val depth_limited_comp_modifiers =
+val depth_limited_comp_modifiers = Comp_Mod.Comp_Modifiers
{const_name_of = depth_limited_function_name_of,
- funT_of = depth_limited_funT_of,
+ funT_of = depth_limited_funT_of : (compilation_funs -> mode -> typ -> typ),
additional_arguments = fn names =>
let
val [depth_name, polarity_name] = Name.variant_list names ["depth", "polarity"]
@@ -2245,38 +2260,38 @@
in [polarity', depth'] end
}
-val rpred_comp_modifiers =
+val rpred_comp_modifiers = Comp_Mod.Comp_Modifiers
{const_name_of = generator_name_of,
- funT_of = K generator_funT_of,
+ funT_of = K generator_funT_of : (compilation_funs -> mode -> typ -> typ),
additional_arguments = fn names => [Free (Name.variant names "size", @{typ code_numeral})],
- wrap_compilation = K (K (K (K (K I)))),
- transform_additional_arguments = K I
+ wrap_compilation = K (K (K (K (K I))))
+ : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
+ transform_additional_arguments = K I : (indprem -> term list -> term list)
}
-
val add_equations = gen_add_equations
- {infer_modes = infer_modes,
+ (Steps {infer_modes = infer_modes,
create_definitions = create_definitions,
compile_preds = compile_preds predicate_comp_modifiers PredicateCompFuns.compfuns,
prove = prove,
are_not_defined = fn thy => forall (null o modes_of thy),
- qname = "equation"}
+ qname = "equation"})
val add_depth_limited_equations = gen_add_equations
- {infer_modes = infer_modes,
+ (Steps {infer_modes = infer_modes,
create_definitions = create_definitions_of_depth_limited_functions,
compile_preds = compile_preds depth_limited_comp_modifiers PredicateCompFuns.compfuns,
prove = prove_by_skip,
are_not_defined = fn thy => forall (null o depth_limited_modes_of thy),
- qname = "depth_limited_equation"}
+ qname = "depth_limited_equation"})
val add_quickcheck_equations = gen_add_equations
- {infer_modes = infer_modes_with_generator,
+ (Steps {infer_modes = infer_modes_with_generator,
create_definitions = rpred_create_definitions,
compile_preds = compile_preds rpred_comp_modifiers RandomPredCompFuns.compfuns,
prove = prove_by_skip,
are_not_defined = fn thy => forall (null o rpred_modes_of thy),
- qname = "rpred_equation"}
+ qname = "rpred_equation"})
(** user interface **)
@@ -2307,7 +2322,7 @@
(extend (fetch_pred_data thy) (depending_preds_of thy) const)) lthy
|> LocalTheory.checkpoint
val thy' = ProofContext.theory_of lthy'
- val preds = Graph.all_preds (PredData.get thy') [const] |> filter_out (has_elim thy')
+ val preds = Graph.all_succs (PredData.get thy') [const] |> filter_out (has_elim thy')
fun mk_cases const =
let
val T = Sign.the_const_type thy const
--- a/src/HOL/Tools/TFL/post.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/TFL/post.ML Thu Oct 29 18:53:58 2009 +0100
@@ -71,7 +71,7 @@
| _ => r RS P_imp_P_eq_True
(*Is this the best way to invoke the simplifier??*)
-fun rewrite L = rewrite_rule (map mk_meta_eq (List.filter(not o id_thm) L))
+fun rewrite L = rewrite_rule (map mk_meta_eq (filter_out id_thm L))
fun join_assums th =
let val thy = Thm.theory_of_thm th
--- a/src/HOL/Tools/TFL/rules.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/TFL/rules.ML Thu Oct 29 18:53:58 2009 +0100
@@ -807,7 +807,7 @@
(prover names) (ss0 addsimps [cut_lemma'] addeqcongs congs) ctm
val th2 = equal_elim th1 th
in
- (th2, List.filter (not o restricted) (!tc_list))
+ (th2, filter_out restricted (!tc_list))
end;
--- a/src/HOL/Tools/TFL/tfl.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/TFL/tfl.ML Thu Oct 29 18:53:58 2009 +0100
@@ -423,13 +423,13 @@
end;
-fun givens pats = map pat_of (List.filter given pats);
+fun givens pats = map pat_of (filter given pats);
fun post_definition meta_tflCongs (theory, (def, pats)) =
let val tych = Thry.typecheck theory
val f = #lhs(S.dest_eq(concl def))
val corollary = R.MATCH_MP Thms.WFREC_COROLLARY def
- val pats' = List.filter given pats
+ val pats' = filter given pats
val given_pats = map pat_of pats'
val rows = map row_of_pat pats'
val WFR = #ant(S.dest_imp(concl corollary))
@@ -821,7 +821,7 @@
let val ex_tm = S.mk_exists{Bvar=v,Body=tm}
in (ex_tm, R.CHOOSE(tych v, R.ASSUME (tych ex_tm)) thm)
end
- val [veq] = List.filter (can S.dest_eq) (#1 (R.dest_thm thm))
+ val [veq] = filter (can S.dest_eq) (#1 (R.dest_thm thm))
val {lhs,rhs} = S.dest_eq veq
val L = S.free_vars_lr rhs
in #2 (fold_rev CHOOSER L (veq,thm)) end;
--- a/src/HOL/Tools/choice_specification.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/choice_specification.ML Thu Oct 29 18:53:58 2009 +0100
@@ -143,7 +143,7 @@
val (_, c') = Type.varify [] c
val (cname,ctyp) = dest_Const c'
in
- case List.filter (fn t => let val (name,typ) = dest_Const t
+ case filter (fn t => let val (name,typ) = dest_Const t
in name = cname andalso Sign.typ_equiv thy (typ, ctyp)
end) thawed_prop_consts of
[] => error ("Specification: No suitable instances of constant \"" ^ Syntax.string_of_term_global thy c ^ "\" found")
--- a/src/HOL/Tools/inductive.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/inductive.ML Thu Oct 29 18:53:58 2009 +0100
@@ -392,7 +392,7 @@
list_all (params',
Logic.list_implies (map (HOLogic.mk_Trueprop o HOLogic.mk_eq)
(frees ~~ us) @ ts, P));
- val c_intrs = (List.filter (equal c o #1 o #1 o #1) intrs);
+ val c_intrs = filter (equal c o #1 o #1 o #1) intrs;
val prems = HOLogic.mk_Trueprop (list_comb (c, params @ frees)) ::
map mk_elim_prem (map #1 c_intrs)
in
--- a/src/HOL/Tools/inductive_codegen.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/inductive_codegen.ML Thu Oct 29 18:53:58 2009 +0100
@@ -174,7 +174,7 @@
let val is' = map (fn i => i - k) (List.drop (is, k))
in map (fn x => Mode (m, is', x)) (cprods (map
(fn (NONE, _) => [NONE]
- | (SOME js, arg) => map SOME (List.filter
+ | (SOME js, arg) => map SOME (filter
(fn Mode (_, js', _) => js=js') (modes_of modes arg)))
(iss ~~ args1)))
end
@@ -237,7 +237,7 @@
fun check_modes_pred thy arg_vs preds modes (p, ms) =
let val SOME rs = AList.lookup (op =) preds p
- in (p, List.filter (fn m => case find_index
+ in (p, filter (fn m => case find_index
(not o check_mode_clause thy arg_vs modes m) rs of
~1 => true
| i => (message ("Clause " ^ string_of_int (i+1) ^ " of " ^
--- a/src/HOL/Tools/lin_arith.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/lin_arith.ML Thu Oct 29 18:53:58 2009 +0100
@@ -681,7 +681,7 @@
val beta_eta_norm = map (apsnd (map (Envir.eta_contract o Envir.beta_norm)))
split_goals
(* TRY (etac notE) THEN eq_assume_tac: *)
- val result = List.filter (not o negated_term_occurs_positively o snd)
+ val result = filter_out (negated_term_occurs_positively o snd)
beta_eta_norm
in
result
--- a/src/HOL/Tools/meson.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/meson.ML Thu Oct 29 18:53:58 2009 +0100
@@ -450,7 +450,7 @@
(*Is the given disjunction an all-negative support clause?*)
fun is_negative th = forall (not o #1) (literals (prop_of th));
-val neg_clauses = List.filter is_negative;
+val neg_clauses = filter is_negative;
(***** MESON PROOF PROCEDURE *****)
--- a/src/HOL/Tools/metis_tools.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/metis_tools.ML Thu Oct 29 18:53:58 2009 +0100
@@ -63,62 +63,62 @@
fun metis_lit b c args = (b, (c, args));
-fun hol_type_to_fol (ResClause.AtomV x) = Metis.Term.Var x
- | hol_type_to_fol (ResClause.AtomF x) = Metis.Term.Fn(x,[])
- | hol_type_to_fol (ResClause.Comp(tc,tps)) = Metis.Term.Fn(tc, map hol_type_to_fol tps);
+fun hol_type_to_fol (Res_Clause.AtomV x) = Metis.Term.Var x
+ | hol_type_to_fol (Res_Clause.AtomF x) = Metis.Term.Fn(x,[])
+ | hol_type_to_fol (Res_Clause.Comp(tc,tps)) = Metis.Term.Fn(tc, map hol_type_to_fol 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 ResHolClause.strip_comb tm of
- (ResHolClause.CombConst(c,_,tys), tms) =>
+ case Res_HOL_Clause.strip_comb tm of
+ (Res_HOL_Clause.CombConst(c,_,tys), tms) =>
let val tyargs = map hol_type_to_fol tys
val args = map hol_term_to_fol_FO tms
in Metis.Term.Fn (c, tyargs @ args) end
- | (ResHolClause.CombVar(v,_), []) => Metis.Term.Var v
+ | (Res_HOL_Clause.CombVar(v,_), []) => Metis.Term.Var v
| _ => error "hol_term_to_fol_FO";
-fun hol_term_to_fol_HO (ResHolClause.CombVar (a, _)) = Metis.Term.Var a
- | hol_term_to_fol_HO (ResHolClause.CombConst (a, _, tylist)) =
+fun hol_term_to_fol_HO (Res_HOL_Clause.CombVar (a, _)) = Metis.Term.Var a
+ | hol_term_to_fol_HO (Res_HOL_Clause.CombConst (a, _, tylist)) =
Metis.Term.Fn (fn_isa_to_met a, map hol_type_to_fol tylist)
- | hol_term_to_fol_HO (ResHolClause.CombApp (tm1, tm2)) =
+ | hol_term_to_fol_HO (Res_HOL_Clause.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, hol_type_to_fol ty]);
-fun hol_term_to_fol_FT (ResHolClause.CombVar(a, ty)) =
+fun hol_term_to_fol_FT (Res_HOL_Clause.CombVar(a, ty)) =
wrap_type (Metis.Term.Var a, ty)
- | hol_term_to_fol_FT (ResHolClause.CombConst(a, ty, _)) =
+ | hol_term_to_fol_FT (Res_HOL_Clause.CombConst(a, ty, _)) =
wrap_type (Metis.Term.Fn(fn_isa_to_met a, []), ty)
- | hol_term_to_fol_FT (tm as ResHolClause.CombApp(tm1,tm2)) =
+ | hol_term_to_fol_FT (tm as Res_HOL_Clause.CombApp(tm1,tm2)) =
wrap_type (Metis.Term.Fn(".", map hol_term_to_fol_FT [tm1,tm2]),
- ResHolClause.type_of_combterm tm);
+ Res_HOL_Clause.type_of_combterm tm);
-fun hol_literal_to_fol FO (ResHolClause.Literal (pol, tm)) =
- let val (ResHolClause.CombConst(p,_,tys), tms) = ResHolClause.strip_comb tm
+fun hol_literal_to_fol FO (Res_HOL_Clause.Literal (pol, tm)) =
+ let val (Res_HOL_Clause.CombConst(p,_,tys), tms) = Res_HOL_Clause.strip_comb tm
val tylits = if p = "equal" then [] else map hol_type_to_fol tys
val lits = map hol_term_to_fol_FO tms
in metis_lit pol (fn_isa_to_met p) (tylits @ lits) end
- | hol_literal_to_fol HO (ResHolClause.Literal (pol, tm)) =
- (case ResHolClause.strip_comb tm of
- (ResHolClause.CombConst("equal",_,_), tms) =>
+ | hol_literal_to_fol HO (Res_HOL_Clause.Literal (pol, tm)) =
+ (case Res_HOL_Clause.strip_comb tm of
+ (Res_HOL_Clause.CombConst("equal",_,_), tms) =>
metis_lit pol "=" (map hol_term_to_fol_HO tms)
| _ => metis_lit pol "{}" [hol_term_to_fol_HO tm]) (*hBOOL*)
- | hol_literal_to_fol FT (ResHolClause.Literal (pol, tm)) =
- (case ResHolClause.strip_comb tm of
- (ResHolClause.CombConst("equal",_,_), tms) =>
+ | hol_literal_to_fol FT (Res_HOL_Clause.Literal (pol, tm)) =
+ (case Res_HOL_Clause.strip_comb tm of
+ (Res_HOL_Clause.CombConst("equal",_,_), tms) =>
metis_lit pol "=" (map hol_term_to_fol_FT tms)
| _ => metis_lit pol "{}" [hol_term_to_fol_FT tm]) (*hBOOL*);
fun literals_of_hol_thm thy mode t =
- let val (lits, types_sorts) = ResHolClause.literals_of_term thy t
+ let val (lits, types_sorts) = Res_HOL_Clause.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_typeLit pos (ResClause.LTVar (s,x)) = metis_lit pos s [Metis.Term.Var x]
- | metis_of_typeLit pos (ResClause.LTFree (s,x)) = metis_lit pos s [Metis.Term.Fn(x,[])];
+fun metis_of_typeLit pos (Res_Clause.LTVar (s,x)) = metis_lit pos s [Metis.Term.Var x]
+ | metis_of_typeLit pos (Res_Clause.LTFree (s,x)) = metis_lit pos s [Metis.Term.Fn(x,[])];
fun default_sort _ (TVar _) = false
| default_sort ctxt (TFree (x, s)) = (s = the_default [] (Variable.def_sort ctxt (x, ~1)));
@@ -132,9 +132,9 @@
(literals_of_hol_thm thy mode o HOLogic.dest_Trueprop o prop_of) th
in
if is_conjecture then
- (Metis.Thm.axiom (Metis.LiteralSet.fromList mlits), ResClause.add_typs types_sorts)
+ (Metis.Thm.axiom (Metis.LiteralSet.fromList mlits), Res_Clause.add_typs types_sorts)
else
- let val tylits = ResClause.add_typs
+ let val tylits = Res_Clause.add_typs
(filter (not o default_sort ctxt) types_sorts)
val mtylits = if Config.get ctxt type_lits
then map (metis_of_typeLit false) tylits else []
@@ -145,13 +145,13 @@
(* ARITY CLAUSE *)
-fun m_arity_cls (ResClause.TConsLit (c,t,args)) =
- metis_lit true (ResClause.make_type_class c) [Metis.Term.Fn(t, map Metis.Term.Var args)]
- | m_arity_cls (ResClause.TVarLit (c,str)) =
- metis_lit false (ResClause.make_type_class c) [Metis.Term.Var str];
+fun m_arity_cls (Res_Clause.TConsLit (c,t,args)) =
+ metis_lit true (Res_Clause.make_type_class c) [Metis.Term.Fn(t, map Metis.Term.Var args)]
+ | m_arity_cls (Res_Clause.TVarLit (c,str)) =
+ metis_lit false (Res_Clause.make_type_class c) [Metis.Term.Var str];
(*TrueI is returned as the Isabelle counterpart because there isn't any.*)
-fun arity_cls (ResClause.ArityClause{conclLit,premLits,...}) =
+fun arity_cls (Res_Clause.ArityClause{conclLit,premLits,...}) =
(TrueI,
Metis.Thm.axiom (Metis.LiteralSet.fromList (map m_arity_cls (conclLit :: premLits))));
@@ -160,7 +160,7 @@
fun m_classrel_cls subclass superclass =
[metis_lit false subclass [Metis.Term.Var "T"], metis_lit true superclass [Metis.Term.Var "T"]];
-fun classrel_cls (ResClause.ClassrelClause {subclass, superclass, ...}) =
+fun classrel_cls (Res_Clause.ClassrelClause {subclass, superclass, ...}) =
(TrueI, Metis.Thm.axiom (Metis.LiteralSet.fromList (m_classrel_cls subclass superclass)));
(* ------------------------------------------------------------------------- *)
@@ -209,14 +209,14 @@
| strip_happ args x = (x, args);
fun fol_type_to_isa _ (Metis.Term.Var v) =
- (case ResReconstruct.strip_prefix ResClause.tvar_prefix v of
- SOME w => ResReconstruct.make_tvar w
- | NONE => ResReconstruct.make_tvar v)
+ (case Res_Reconstruct.strip_prefix Res_Clause.tvar_prefix v of
+ SOME w => Res_Reconstruct.make_tvar w
+ | NONE => Res_Reconstruct.make_tvar v)
| fol_type_to_isa ctxt (Metis.Term.Fn(x, tys)) =
- (case ResReconstruct.strip_prefix ResClause.tconst_prefix x of
- SOME tc => Term.Type (ResReconstruct.invert_type_const tc, map (fol_type_to_isa ctxt) tys)
+ (case Res_Reconstruct.strip_prefix Res_Clause.tconst_prefix x of
+ SOME tc => Term.Type (Res_Reconstruct.invert_type_const tc, map (fol_type_to_isa ctxt) tys)
| NONE =>
- case ResReconstruct.strip_prefix ResClause.tfree_prefix x of
+ case Res_Reconstruct.strip_prefix Res_Clause.tfree_prefix x of
SOME tf => mk_tfree ctxt tf
| NONE => error ("fol_type_to_isa: " ^ x));
@@ -225,10 +225,10 @@
let val thy = ProofContext.theory_of ctxt
val _ = trace_msg (fn () => "fol_term_to_hol: " ^ Metis.Term.toString fol_tm)
fun tm_to_tt (Metis.Term.Var v) =
- (case ResReconstruct.strip_prefix ResClause.tvar_prefix v of
- SOME w => Type (ResReconstruct.make_tvar w)
+ (case Res_Reconstruct.strip_prefix Res_Clause.tvar_prefix v of
+ SOME w => Type (Res_Reconstruct.make_tvar w)
| NONE =>
- case ResReconstruct.strip_prefix ResClause.schematic_var_prefix v of
+ case Res_Reconstruct.strip_prefix Res_Clause.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*)
@@ -245,14 +245,14 @@
and applic_to_tt ("=",ts) =
Term (list_comb(Const ("op =", HOLogic.typeT), terms_of (map tm_to_tt ts)))
| applic_to_tt (a,ts) =
- case ResReconstruct.strip_prefix ResClause.const_prefix a of
+ case Res_Reconstruct.strip_prefix Res_Clause.const_prefix a of
SOME b =>
- let val c = ResReconstruct.invert_const b
- val ntypes = ResReconstruct.num_typargs thy c
+ let val c = Res_Reconstruct.invert_const b
+ val ntypes = Res_Reconstruct.num_typargs thy c
val nterms = length ts - ntypes
val tts = map tm_to_tt ts
val tys = types_of (List.take(tts,ntypes))
- val ntyargs = ResReconstruct.num_typargs thy c
+ val ntyargs = Res_Reconstruct.num_typargs thy c
in if length tys = ntyargs then
apply_list (Const (c, dummyT)) nterms (List.drop(tts,ntypes))
else error ("Constant " ^ c ^ " expects " ^ Int.toString ntyargs ^
@@ -263,14 +263,14 @@
cat_lines (map (Syntax.string_of_term ctxt) (terms_of tts)))
end
| NONE => (*Not a constant. Is it a type constructor?*)
- case ResReconstruct.strip_prefix ResClause.tconst_prefix a of
+ case Res_Reconstruct.strip_prefix Res_Clause.tconst_prefix a of
SOME b =>
- Type (Term.Type (ResReconstruct.invert_type_const b, types_of (map tm_to_tt ts)))
+ Type (Term.Type (Res_Reconstruct.invert_type_const b, types_of (map tm_to_tt ts)))
| NONE => (*Maybe a TFree. Should then check that ts=[].*)
- case ResReconstruct.strip_prefix ResClause.tfree_prefix a of
+ case Res_Reconstruct.strip_prefix Res_Clause.tfree_prefix a of
SOME b => Type (mk_tfree ctxt b)
| NONE => (*a fixed variable? They are Skolem functions.*)
- case ResReconstruct.strip_prefix ResClause.fixed_var_prefix a of
+ case Res_Reconstruct.strip_prefix Res_Clause.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
@@ -281,16 +281,16 @@
fun fol_term_to_hol_FT ctxt fol_tm =
let val _ = trace_msg (fn () => "fol_term_to_hol_FT: " ^ Metis.Term.toString fol_tm)
fun cvt (Metis.Term.Fn ("ti", [Metis.Term.Var v, _])) =
- (case ResReconstruct.strip_prefix ResClause.schematic_var_prefix v of
+ (case Res_Reconstruct.strip_prefix Res_Clause.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 ("op =", HOLogic.typeT)
| cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn (x,[]), ty])) =
- (case ResReconstruct.strip_prefix ResClause.const_prefix x of
- SOME c => Const (ResReconstruct.invert_const c, dummyT)
+ (case Res_Reconstruct.strip_prefix Res_Clause.const_prefix x of
+ SOME c => Const (Res_Reconstruct.invert_const c, dummyT)
| NONE => (*Not a constant. Is it a fixed variable??*)
- case ResReconstruct.strip_prefix ResClause.fixed_var_prefix x of
+ case Res_Reconstruct.strip_prefix Res_Clause.fixed_var_prefix x of
SOME v => Free (v, fol_type_to_isa ctxt ty)
| NONE => error ("fol_term_to_hol_FT bad constant: " ^ x))
| cvt (Metis.Term.Fn ("ti", [Metis.Term.Fn (".",[tm1,tm2]), _])) =
@@ -301,10 +301,10 @@
| cvt (Metis.Term.Fn ("=", [tm1,tm2])) =
list_comb(Const ("op =", HOLogic.typeT), map cvt [tm1,tm2])
| cvt (t as Metis.Term.Fn (x, [])) =
- (case ResReconstruct.strip_prefix ResClause.const_prefix x of
- SOME c => Const (ResReconstruct.invert_const c, dummyT)
+ (case Res_Reconstruct.strip_prefix Res_Clause.const_prefix x of
+ SOME c => Const (Res_Reconstruct.invert_const c, dummyT)
| NONE => (*Not a constant. Is it a fixed variable??*)
- case ResReconstruct.strip_prefix ResClause.fixed_var_prefix x of
+ case Res_Reconstruct.strip_prefix Res_Clause.fixed_var_prefix x of
SOME v => Free (v, dummyT)
| NONE => (trace_msg (fn () => "fol_term_to_hol_FT bad const: " ^ x);
fol_term_to_hol_RAW ctxt t))
@@ -383,9 +383,9 @@
" in " ^ Display.string_of_thm ctxt i_th);
NONE)
fun remove_typeinst (a, t) =
- case ResReconstruct.strip_prefix ResClause.schematic_var_prefix a of
+ case Res_Reconstruct.strip_prefix Res_Clause.schematic_var_prefix a of
SOME b => SOME (b, t)
- | NONE => case ResReconstruct.strip_prefix ResClause.tvar_prefix a of
+ | NONE => case Res_Reconstruct.strip_prefix Res_Clause.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)
@@ -452,7 +452,7 @@
in cterm_instantiate [(refl_x, c_t)] REFL_THM end;
fun get_ty_arg_size _ (Const ("op =", _)) = 0 (*equality has no type arguments*)
- | get_ty_arg_size thy (Const (c, _)) = (ResReconstruct.num_typargs thy c handle TYPE _ => 0)
+ | get_ty_arg_size thy (Const (c, _)) = (Res_Reconstruct.num_typargs thy c handle TYPE _ => 0)
| get_ty_arg_size _ _ = 0;
(* INFERENCE RULE: EQUALITY *)
@@ -538,7 +538,7 @@
| step ctxt mode _ (_, Metis.Proof.Equality (f_lit, f_p, f_r)) =
equality_inf ctxt mode f_lit f_p f_r;
-fun real_literal (_, (c, _)) = not (String.isPrefix ResClause.class_prefix c);
+fun real_literal (_, (c, _)) = not (String.isPrefix Res_Clause.class_prefix c);
fun translate _ _ thpairs [] = thpairs
| translate mode ctxt thpairs ((fol_th, inf) :: infpairs) =
@@ -564,23 +564,23 @@
(* Translation of HO Clauses *)
(* ------------------------------------------------------------------------- *)
-fun cnf_th thy th = hd (ResAxioms.cnf_axiom thy th);
+fun cnf_th thy th = hd (Res_Axioms.cnf_axiom thy th);
val equal_imp_fequal' = cnf_th @{theory} @{thm equal_imp_fequal};
val fequal_imp_equal' = cnf_th @{theory} @{thm fequal_imp_equal};
-val comb_I = cnf_th @{theory} ResHolClause.comb_I;
-val comb_K = cnf_th @{theory} ResHolClause.comb_K;
-val comb_B = cnf_th @{theory} ResHolClause.comb_B;
-val comb_C = cnf_th @{theory} ResHolClause.comb_C;
-val comb_S = cnf_th @{theory} ResHolClause.comb_S;
+val comb_I = cnf_th @{theory} Res_HOL_Clause.comb_I;
+val comb_K = cnf_th @{theory} Res_HOL_Clause.comb_K;
+val comb_B = cnf_th @{theory} Res_HOL_Clause.comb_B;
+val comb_C = cnf_th @{theory} Res_HOL_Clause.comb_C;
+val comb_S = cnf_th @{theory} Res_HOL_Clause.comb_S;
fun type_ext thy tms =
- let val subs = ResAtp.tfree_classes_of_terms tms
- val supers = ResAtp.tvar_classes_of_terms tms
- and tycons = ResAtp.type_consts_of_terms thy tms
- val (supers', arity_clauses) = ResClause.make_arity_clauses thy tycons supers
- val classrel_clauses = ResClause.make_classrel_clauses thy subs supers'
+ let val subs = Res_ATP.tfree_classes_of_terms tms
+ val supers = Res_ATP.tvar_classes_of_terms tms
+ and tycons = Res_ATP.type_consts_of_terms thy tms
+ val (supers', arity_clauses) = Res_Clause.make_arity_clauses thy tycons supers
+ val classrel_clauses = Res_Clause.make_classrel_clauses thy subs supers'
in map classrel_cls classrel_clauses @ map arity_cls arity_clauses
end;
@@ -590,7 +590,7 @@
type logic_map =
{axioms : (Metis.Thm.thm * thm) list,
- tfrees : ResClause.type_literal list};
+ tfrees : Res_Clause.type_literal list};
fun const_in_metis c (pred, tm_list) =
let
@@ -602,7 +602,7 @@
(*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 ResClause.add_typs (Vartab.fold add (#2 (Variable.constraints_of ctxt)) []) end;
+ in Res_Clause.add_typs (Vartab.fold add (#2 (Variable.constraints_of ctxt)) []) end;
(*transform isabelle type / arity clause to metis clause *)
fun add_type_thm [] lmap = lmap
@@ -664,7 +664,7 @@
(* Main function to start metis prove and reconstruction *)
fun FOL_SOLVE mode ctxt cls ths0 =
let val thy = ProofContext.theory_of ctxt
- val th_cls_pairs = map (fn th => (Thm.get_name_hint th, ResAxioms.cnf_axiom thy th)) ths0
+ val th_cls_pairs = map (fn th => (Thm.get_name_hint th, Res_Axioms.cnf_axiom thy th)) ths0
val ths = maps #2 th_cls_pairs
val _ = trace_msg (fn () => "FOL_SOLVE: CONJECTURE CLAUSES")
val _ = app (fn th => trace_msg (fn () => Display.string_of_thm ctxt th)) cls
@@ -673,14 +673,14 @@
val (mode, {axioms,tfrees}) = build_map mode ctxt cls ths
val _ = if null tfrees then ()
else (trace_msg (fn () => "TFREE CLAUSES");
- app (fn tf => trace_msg (fn _ => ResClause.tptp_of_typeLit true tf)) tfrees)
+ app (fn tf => trace_msg (fn _ => Res_Clause.tptp_of_typeLit true tf)) tfrees)
val _ = trace_msg (fn () => "CLAUSES GIVEN TO METIS")
val thms = map #1 axioms
val _ = app (fn th => trace_msg (fn () => Metis.Thm.toString th)) thms
val _ = trace_msg (fn () => "mode = " ^ string_of_mode mode)
val _ = trace_msg (fn () => "START METIS PROVE PROCESS")
in
- case List.filter (is_false o prop_of) cls of
+ case filter (is_false o prop_of) cls of
false_th::_ => [false_th RS @{thm FalseE}]
| [] =>
case refute thms of
@@ -694,10 +694,11 @@
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 = filter (fn (_, cls) => not (common_thm used cls)) th_cls_pairs
+ val unused = th_cls_pairs |> map_filter (fn (name, cls) =>
+ if common_thm used cls then NONE else SOME name)
in
if null unused then ()
- else warning ("Metis: unused theorems " ^ commas_quote (map #1 unused));
+ else warning ("Metis: unused theorems " ^ commas_quote unused);
case result of
(_,ith)::_ =>
(trace_msg (fn () => "success: " ^ Display.string_of_thm ctxt ith);
@@ -714,12 +715,12 @@
let val _ = trace_msg (fn () =>
"Metis called with theorems " ^ cat_lines (map (Display.string_of_thm ctxt) ths))
in
- if exists_type ResAxioms.type_has_empty_sort (prop_of st0)
+ if exists_type Res_Axioms.type_has_empty_sort (prop_of st0)
then (warning "Proof state contains the empty sort"; Seq.empty)
else
- (Meson.MESON ResAxioms.neg_clausify
+ (Meson.MESON Res_Axioms.neg_clausify
(fn cls => resolve_tac (FOL_SOLVE mode ctxt cls ths) 1) ctxt i
- THEN ResAxioms.expand_defs_tac st0) st0
+ THEN Res_Axioms.expand_defs_tac st0) st0
end
handle METIS s => (warning ("Metis: " ^ s); Seq.empty);
@@ -736,7 +737,7 @@
method @{binding metisF} FO "METIS for FOL problems" #>
method @{binding metisFT} FT "METIS With-fully typed translation" #>
Method.setup @{binding finish_clausify}
- (Scan.succeed (K (SIMPLE_METHOD (ResAxioms.expand_defs_tac refl))))
+ (Scan.succeed (K (SIMPLE_METHOD (Res_Axioms.expand_defs_tac refl))))
"cleanup after conversion to clauses";
end;
--- a/src/HOL/Tools/old_primrec.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/old_primrec.ML Thu Oct 29 18:53:58 2009 +0100
@@ -195,7 +195,7 @@
NONE =>
let
val dummy_fns = map (fn (_, cargs) => Const ("HOL.undefined",
- replicate ((length cargs) + (length (List.filter is_rec_type cargs)))
+ replicate (length cargs + length (filter is_rec_type cargs))
dummyT ---> HOLogic.unitT)) constrs;
val _ = warning ("No function definition for datatype " ^ quote tname)
in
--- a/src/HOL/Tools/primrec.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/primrec.ML Thu Oct 29 18:53:58 2009 +0100
@@ -178,7 +178,7 @@
NONE =>
let
val dummy_fns = map (fn (_, cargs) => Const ("HOL.undefined",
- replicate ((length cargs) + (length (List.filter is_rec_type cargs)))
+ replicate (length cargs + length (filter is_rec_type cargs))
dummyT ---> HOLogic.unitT)) constrs;
val _ = warning ("No function definition for datatype " ^ quote tname)
in
--- a/src/HOL/Tools/refute.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/refute.ML Thu Oct 29 18:53:58 2009 +0100
@@ -394,7 +394,7 @@
(* (string * int) list *)
val sizes = map_filter
(fn (name, value) => Option.map (pair name) (Int.fromString value))
- (List.filter (fn (name, _) => name<>"minsize" andalso name<>"maxsize"
+ (filter (fn (name, _) => name<>"minsize" andalso name<>"maxsize"
andalso name<>"maxvars" andalso name<>"maxtime"
andalso name<>"satsolver") allparams)
in
@@ -1070,8 +1070,7 @@
(* continue search *)
next max (len+1) (sum+x1) (x2::xs)
(* only consider those types for which the size is not fixed *)
- val mutables = List.filter
- (not o (AList.defined (op =) sizes) o string_of_typ o fst) xs
+ val mutables = filter_out (AList.defined (op =) sizes o string_of_typ o fst) xs
(* subtract 'minsize' from every size (will be added again at the end) *)
val diffs = map (fn (_, n) => n-minsize) mutables
in
@@ -2552,7 +2551,7 @@
(* arguments *)
val (_, _, constrs) = the (AList.lookup (op =) descr idx)
val (_, dtyps) = List.nth (constrs, c)
- val rec_dtyps_args = List.filter
+ val rec_dtyps_args = filter
(DatatypeAux.is_rec_type o fst) (dtyps ~~ args)
(* map those indices to interpretations *)
val rec_dtyps_intrs = map (fn (dtyp, arg) =>
--- a/src/HOL/Tools/res_atp.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/res_atp.ML Thu Oct 29 18:53:58 2009 +0100
@@ -1,4 +1,6 @@
-(* Author: Jia Meng, Cambridge University Computer Laboratory, NICTA *)
+(* Title: HOL/Tools/res_atp.ML
+ Author: Jia Meng, Cambridge University Computer Laboratory, NICTA
+*)
signature RES_ATP =
sig
@@ -9,14 +11,14 @@
val get_relevant : int -> bool -> Proof.context * (thm list * 'a) -> thm list ->
(thm * (string * int)) list
val prepare_clauses : bool -> thm list -> thm list ->
- (thm * (ResHolClause.axiom_name * ResHolClause.clause_id)) list ->
- (thm * (ResHolClause.axiom_name * ResHolClause.clause_id)) list -> theory ->
- ResHolClause.axiom_name vector *
- (ResHolClause.clause list * ResHolClause.clause list * ResHolClause.clause list *
- ResHolClause.clause list * ResClause.classrelClause list * ResClause.arityClause list)
+ (thm * (Res_HOL_Clause.axiom_name * Res_HOL_Clause.clause_id)) list ->
+ (thm * (Res_HOL_Clause.axiom_name * Res_HOL_Clause.clause_id)) list -> theory ->
+ Res_HOL_Clause.axiom_name vector *
+ (Res_HOL_Clause.clause list * Res_HOL_Clause.clause list * Res_HOL_Clause.clause list *
+ Res_HOL_Clause.clause list * Res_Clause.classrelClause list * Res_Clause.arityClause list)
end;
-structure ResAtp: RES_ATP =
+structure Res_ATP: RES_ATP =
struct
@@ -236,10 +238,10 @@
let val cls = sort compare_pairs newpairs
val accepted = List.take (cls, max_new)
in
- ResAxioms.trace_msg (fn () => ("Number of candidates, " ^ Int.toString nnew ^
+ Res_Axioms.trace_msg (fn () => ("Number of candidates, " ^ Int.toString nnew ^
", exceeds the limit of " ^ Int.toString (max_new)));
- ResAxioms.trace_msg (fn () => ("Effective pass mark: " ^ Real.toString (#2 (List.last accepted))));
- ResAxioms.trace_msg (fn () => "Actually passed: " ^
+ Res_Axioms.trace_msg (fn () => ("Effective pass mark: " ^ Real.toString (#2 (List.last accepted))));
+ Res_Axioms.trace_msg (fn () => "Actually passed: " ^
space_implode ", " (map (fn (((_,(name,_)),_),_) => name) accepted));
(map #1 accepted, map #1 (List.drop (cls, max_new)))
@@ -254,7 +256,7 @@
val rel_consts' = List.foldl add_const_typ_table rel_consts new_consts
val newp = p + (1.0-p) / convergence
in
- ResAxioms.trace_msg (fn () => ("relevant this iteration: " ^ Int.toString (length newrels)));
+ Res_Axioms.trace_msg (fn () => "relevant this iteration: " ^ Int.toString (length newrels));
(map #1 newrels) @
(relevant_clauses max_new thy ctab newp rel_consts' (more_rejects@rejects))
end
@@ -262,12 +264,12 @@
let val weight = clause_weight ctab rel_consts consts_typs
in
if p <= weight orelse (follow_defs andalso defines thy (#1 clsthm) rel_consts)
- then (ResAxioms.trace_msg (fn () => (name ^ " clause " ^ Int.toString n ^
+ then (Res_Axioms.trace_msg (fn () => (name ^ " clause " ^ Int.toString n ^
" passes: " ^ Real.toString weight));
relevant ((ax,weight)::newrels, rejects) axs)
else relevant (newrels, ax::rejects) axs
end
- in ResAxioms.trace_msg (fn () => ("relevant_clauses, current pass mark = " ^ Real.toString p));
+ in Res_Axioms.trace_msg (fn () => ("relevant_clauses, current pass mark = " ^ Real.toString p));
relevant ([],[])
end;
@@ -276,12 +278,12 @@
then
let val const_tab = List.foldl (count_axiom_consts theory_const thy) Symtab.empty axioms
val goal_const_tab = get_goal_consts_typs thy goals
- val _ = ResAxioms.trace_msg (fn () => ("Initial constants: " ^
+ val _ = Res_Axioms.trace_msg (fn () => ("Initial constants: " ^
space_implode ", " (Symtab.keys goal_const_tab)));
val rels = relevant_clauses max_new thy const_tab (pass_mark)
goal_const_tab (map (pair_consts_typs_axiom theory_const thy) axioms)
in
- ResAxioms.trace_msg (fn () => ("Total relevant: " ^ Int.toString (length rels)));
+ Res_Axioms.trace_msg (fn () => ("Total relevant: " ^ Int.toString (length rels)));
rels
end
else axioms;
@@ -335,7 +337,7 @@
fun make_unique xs =
let val ht = mk_clause_table 7000
in
- ResAxioms.trace_msg (fn () => ("make_unique gets " ^ Int.toString (length xs) ^ " clauses"));
+ Res_Axioms.trace_msg (fn () => ("make_unique gets " ^ Int.toString (length xs) ^ " clauses"));
app (ignore o Polyhash.peekInsert ht) xs;
Polyhash.listItems ht
end;
@@ -352,11 +354,12 @@
fun valid_facts facts =
Facts.fold_static (fn (name, ths) =>
- if run_blacklist_filter andalso is_package_def name then I
+ if Facts.is_concealed facts name orelse
+ (run_blacklist_filter andalso is_package_def name) then I
else
let val xname = Facts.extern facts name in
if Name_Space.is_hidden xname then I
- else cons (xname, filter_out ResAxioms.bad_for_atp ths)
+ else cons (xname, filter_out Res_Axioms.bad_for_atp ths)
end) facts [];
fun all_valid_thms ctxt =
@@ -365,29 +368,29 @@
val local_facts = ProofContext.facts_of ctxt;
in valid_facts global_facts @ valid_facts local_facts end;
-fun multi_name a (th, (n,pairs)) =
- (n+1, (a ^ "(" ^ Int.toString n ^ ")", th) :: pairs)
+fun multi_name a th (n, pairs) =
+ (n + 1, (a ^ "(" ^ Int.toString n ^ ")", th) :: pairs);
-fun add_single_names ((a, []), pairs) = pairs
- | add_single_names ((a, [th]), pairs) = (a,th)::pairs
- | add_single_names ((a, ths), pairs) = #2 (List.foldl (multi_name a) (1,pairs) ths);
+fun add_single_names (a, []) pairs = pairs
+ | add_single_names (a, [th]) pairs = (a, th) :: pairs
+ | add_single_names (a, ths) pairs = #2 (fold (multi_name a) ths (1, pairs));
(*Ignore blacklisted basenames*)
-fun add_multi_names ((a, ths), pairs) =
- if (Long_Name.base_name a) mem_string ResAxioms.multi_base_blacklist then pairs
- else add_single_names ((a, ths), pairs);
+fun add_multi_names (a, ths) pairs =
+ if (Long_Name.base_name a) mem_string Res_Axioms.multi_base_blacklist then pairs
+ else add_single_names (a, ths) pairs;
fun is_multi (a, ths) = length ths > 1 orelse String.isSuffix ".axioms" a;
(*The single theorems go BEFORE the multiple ones. Blacklist is applied to all.*)
fun name_thm_pairs ctxt =
- let val (mults,singles) = List.partition is_multi (all_valid_thms ctxt)
- val ht = mk_clause_table 900 (*ht for blacklisted theorems*)
- fun blacklisted x = run_blacklist_filter andalso is_some (Polyhash.peek ht x)
+ let
+ val (mults, singles) = List.partition is_multi (all_valid_thms ctxt)
+ fun blacklisted (_, th) =
+ run_blacklist_filter andalso Res_Blacklist.blacklisted ctxt th
in
- app (fn th => ignore (Polyhash.peekInsert ht (th,()))) (ResBlacklist.get ctxt);
- filter (not o blacklisted o #2)
- (List.foldl add_single_names (List.foldl add_multi_names [] mults) singles)
+ filter_out blacklisted
+ (fold add_single_names singles (fold add_multi_names mults []))
end;
fun check_named ("", th) =
@@ -396,7 +399,7 @@
fun get_all_lemmas ctxt =
let val included_thms =
- tap (fn ths => ResAxioms.trace_msg
+ tap (fn ths => Res_Axioms.trace_msg
(fn () => ("Including all " ^ Int.toString (length ths) ^ " theorems")))
(name_thm_pairs ctxt)
in
@@ -499,17 +502,14 @@
| Fol => true
| Hol => false
-fun ths_to_cls thy ths =
- ResAxioms.cnf_rules_pairs thy (filter check_named (map ResAxioms.pairname ths))
-
fun get_relevant max_new theory_const (ctxt, (chain_ths, th)) goal_cls =
let
val thy = ProofContext.theory_of ctxt
val isFO = isFO thy goal_cls
- val included_thms = get_all_lemmas ctxt
- val included_cls = included_thms |> ResAxioms.cnf_rules_pairs thy |> make_unique
- |> restrict_to_logic thy isFO
- |> remove_unwanted_clauses
+ val included_cls = get_all_lemmas ctxt
+ |> Res_Axioms.cnf_rules_pairs thy |> make_unique
+ |> restrict_to_logic thy isFO
+ |> remove_unwanted_clauses
in
relevance_filter max_new theory_const thy included_cls (map prop_of goal_cls)
end;
@@ -519,24 +519,25 @@
fun prepare_clauses dfg goal_cls chain_ths axcls extra_cls thy =
let
(* add chain thms *)
- val chain_cls = ths_to_cls thy chain_ths
+ val chain_cls =
+ Res_Axioms.cnf_rules_pairs thy (filter check_named (map Res_Axioms.pairname chain_ths))
val axcls = chain_cls @ axcls
val extra_cls = chain_cls @ extra_cls
val isFO = isFO thy goal_cls
val ccls = subtract_cls goal_cls extra_cls
- val _ = app (fn th => ResAxioms.trace_msg (fn _ => Display.string_of_thm_global thy th)) ccls
+ val _ = app (fn th => Res_Axioms.trace_msg (fn _ => Display.string_of_thm_global thy th)) ccls
val ccltms = map prop_of ccls
and axtms = map (prop_of o #1) extra_cls
val subs = tfree_classes_of_terms ccltms
and supers = tvar_classes_of_terms axtms
and tycons = type_consts_of_terms thy (ccltms@axtms)
(*TFrees in conjecture clauses; TVars in axiom clauses*)
- val conjectures = ResHolClause.make_conjecture_clauses dfg thy ccls
- val (_, extra_clauses) = ListPair.unzip (ResHolClause.make_axiom_clauses dfg thy extra_cls)
- val (clnames,axiom_clauses) = ListPair.unzip (ResHolClause.make_axiom_clauses dfg thy axcls)
- val helper_clauses = ResHolClause.get_helper_clauses dfg thy isFO (conjectures, extra_cls, [])
- val (supers',arity_clauses) = ResClause.make_arity_clauses_dfg dfg thy tycons supers
- val classrel_clauses = ResClause.make_classrel_clauses thy subs supers'
+ val conjectures = Res_HOL_Clause.make_conjecture_clauses dfg thy ccls
+ val (_, extra_clauses) = ListPair.unzip (Res_HOL_Clause.make_axiom_clauses dfg thy extra_cls)
+ val (clnames,axiom_clauses) = ListPair.unzip (Res_HOL_Clause.make_axiom_clauses dfg thy axcls)
+ val helper_clauses = Res_HOL_Clause.get_helper_clauses dfg thy isFO (conjectures, extra_cls, [])
+ val (supers',arity_clauses) = Res_Clause.make_arity_clauses_dfg dfg thy tycons supers
+ val classrel_clauses = Res_Clause.make_classrel_clauses thy subs supers'
in
(Vector.fromList clnames,
(conjectures, axiom_clauses, extra_clauses, helper_clauses, classrel_clauses, arity_clauses))
--- a/src/HOL/Tools/res_axioms.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/res_axioms.ML Thu Oct 29 18:53:58 2009 +0100
@@ -1,4 +1,5 @@
-(* Author: Jia Meng, Cambridge University Computer Laboratory
+(* Title: HOL/Tools/res_axioms.ML
+ Author: Jia Meng, Cambridge University Computer Laboratory
Transformation of axiom rules (elim/intro/etc) into CNF forms.
*)
@@ -22,7 +23,7 @@
val setup: theory -> theory
end;
-structure ResAxioms: RES_AXIOMS =
+structure Res_Axioms: RES_AXIOMS =
struct
val trace = Unsynchronized.ref false;
@@ -285,7 +286,7 @@
map (skolem_of_def o assume o (cterm_of (theory_of_thm th))) (assume_skofuns s th);
-(*** Blacklisting (duplicated in ResAtp?) ***)
+(*** Blacklisting (duplicated in Res_ATP?) ***)
val max_lambda_nesting = 3;
@@ -316,18 +317,17 @@
fun is_strange_thm th =
case head_of (concl_of th) of
- Const (a,_) => (a <> "Trueprop" andalso a <> "==")
+ Const (a, _) => (a <> "Trueprop" andalso a <> "==")
| _ => false;
fun bad_for_atp th =
- Thm.is_internal th
- orelse too_complex (prop_of th)
+ too_complex (prop_of th)
orelse exists_type type_has_empty_sort (prop_of th)
orelse is_strange_thm th;
val multi_base_blacklist =
["defs","select_defs","update_defs","induct","inducts","split","splits","split_asm",
- "cases","ext_cases"]; (*FIXME: put other record thms here, or use the "Internal" marker*)
+ "cases","ext_cases"]; (* FIXME put other record thms here, or declare as "noatp" *)
(*Keep the full complexity of the original name*)
fun flatten_name s = space_implode "_X" (Long_Name.explode s);
@@ -387,14 +387,14 @@
fun cnf_rules_pairs_aux _ pairs [] = pairs
| cnf_rules_pairs_aux thy pairs ((name,th)::ths) =
let val pairs' = (pair_name_cls 0 (name, cnf_axiom thy th)) @ pairs
- handle THM _ => pairs | ResClause.CLAUSE _ => pairs
+ handle THM _ => pairs | Res_Clause.CLAUSE _ => pairs
in cnf_rules_pairs_aux thy pairs' ths end;
(*The combination of rev and tail recursion preserves the original order*)
fun cnf_rules_pairs thy l = cnf_rules_pairs_aux thy [] (rev l);
-(**** Convert all facts of the theory into clauses (ResClause.clause, or ResHolClause.clause) ****)
+(**** Convert all facts of the theory into clauses (Res_Clause.clause, or Res_HOL_Clause.clause) ****)
local
@@ -421,8 +421,10 @@
fun saturate_skolem_cache thy =
let
- val new_facts = (PureThy.facts_of thy, []) |-> Facts.fold_static (fn (name, ths) =>
- if already_seen thy name then I else cons (name, ths));
+ val facts = PureThy.facts_of thy;
+ val new_facts = (facts, []) |-> Facts.fold_static (fn (name, ths) =>
+ if Facts.is_concealed facts name orelse already_seen thy name then I
+ else cons (name, ths));
val new_thms = (new_facts, []) |-> fold (fn (name, ths) =>
if member (op =) multi_base_blacklist (Long_Name.base_name name) then I
else fold_index (fn (i, th) =>
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/res_blacklist.ML Thu Oct 29 18:53:58 2009 +0100
@@ -0,0 +1,43 @@
+(* Title: HOL/Tools/res_blacklist.ML
+ Author: Lawrence C Paulson, Cambridge University Computer Laboratory
+ Author: Makarius
+
+Theorems blacklisted to sledgehammer. These theorems typically
+produce clauses that are prolific (match too many equality or
+membership literals) and relate to seldom-used facts. Some duplicate
+other rules.
+*)
+
+signature RES_BLACKLIST =
+sig
+ val setup: theory -> theory
+ val blacklisted: Proof.context -> thm -> bool
+ val add: attribute
+ val del: attribute
+end;
+
+structure Res_Blacklist: RES_BLACKLIST =
+struct
+
+structure Data = GenericDataFun
+(
+ type T = thm Termtab.table;
+ val empty = Termtab.empty;
+ val extend = I;
+ fun merge _ tabs = Termtab.merge (K true) tabs;
+);
+
+fun blacklisted ctxt th =
+ Termtab.defined (Data.get (Context.Proof ctxt)) (Thm.full_prop_of th);
+
+fun add_thm th = Data.map (Termtab.update (Thm.full_prop_of th, th));
+fun del_thm th = Data.map (Termtab.delete_safe (Thm.full_prop_of th));
+
+val add = Thm.declaration_attribute add_thm;
+val del = Thm.declaration_attribute del_thm;
+
+val setup =
+ Attrib.setup @{binding noatp} (Attrib.add_del add del) "sledgehammer blacklisting" #>
+ PureThy.add_thms_dynamic (@{binding noatp}, map #2 o Termtab.dest o Data.get);
+
+end;
--- a/src/HOL/Tools/res_clause.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/res_clause.ML Thu Oct 29 18:53:58 2009 +0100
@@ -1,11 +1,12 @@
-(* Author: Jia Meng, Cambridge University Computer Laboratory
- Copyright 2004 University of Cambridge
+(* Title: HOL/Tools/res_clause.ML
+ Author: Jia Meng, Cambridge University Computer Laboratory
-Storing/printing FOL clauses and arity clauses.
-Typed equality is treated differently.
+Storing/printing FOL clauses and arity clauses. Typed equality is
+treated differently.
+
+FIXME: combine with res_hol_clause!
*)
-(*FIXME: combine with res_hol_clause!*)
signature RES_CLAUSE =
sig
val schematic_var_prefix: string
@@ -76,7 +77,7 @@
val tptp_classrelClause : classrelClause -> string
end
-structure ResClause: RES_CLAUSE =
+structure Res_Clause: RES_CLAUSE =
struct
val schematic_var_prefix = "V_";
--- a/src/HOL/Tools/res_hol_clause.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/res_hol_clause.ML Thu Oct 29 18:53:58 2009 +0100
@@ -1,5 +1,5 @@
-(*
- Author: Jia Meng, NICTA
+(* Title: HOL/Tools/res_hol_clause.ML
+ Author: Jia Meng, NICTA
FOL clauses translated from HOL formulae.
*)
@@ -17,13 +17,13 @@
type polarity = bool
type clause_id = int
datatype combterm =
- CombConst of string * ResClause.fol_type * ResClause.fol_type list (*Const and Free*)
- | CombVar of string * ResClause.fol_type
+ CombConst of string * Res_Clause.fol_type * Res_Clause.fol_type list (*Const and Free*)
+ | CombVar of string * Res_Clause.fol_type
| CombApp of combterm * combterm
datatype literal = Literal of polarity * combterm
datatype clause = Clause of {clause_id: clause_id, axiom_name: axiom_name, th: thm,
- kind: ResClause.kind,literals: literal list, ctypes_sorts: typ list}
- val type_of_combterm: combterm -> ResClause.fol_type
+ kind: Res_Clause.kind,literals: literal list, ctypes_sorts: typ list}
+ val type_of_combterm: combterm -> Res_Clause.fol_type
val strip_comb: combterm -> combterm * combterm list
val literals_of_term: theory -> term -> literal list * typ list
exception TOO_TRIVIAL
@@ -38,18 +38,18 @@
clause list
val tptp_write_file: bool -> Path.T ->
clause list * clause list * clause list * clause list *
- ResClause.classrelClause list * ResClause.arityClause list ->
+ Res_Clause.classrelClause list * Res_Clause.arityClause list ->
int * int
val dfg_write_file: bool -> Path.T ->
clause list * clause list * clause list * clause list *
- ResClause.classrelClause list * ResClause.arityClause list ->
+ Res_Clause.classrelClause list * Res_Clause.arityClause list ->
int * int
end
-structure ResHolClause: RES_HOL_CLAUSE =
+structure Res_HOL_Clause: RES_HOL_CLAUSE =
struct
-structure RC = ResClause;
+structure RC = Res_Clause; (* FIXME avoid structure alias *)
(* theorems for combinators and function extensionality *)
val ext = thm "HOL.ext";
@@ -404,7 +404,7 @@
else ct;
fun cnf_helper_thms thy =
- ResAxioms.cnf_rules_pairs thy o map ResAxioms.pairname
+ Res_Axioms.cnf_rules_pairs thy o map Res_Axioms.pairname
fun get_helper_clauses dfg thy isFO (conjectures, axcls, user_lemmas) =
if isFO then [] (*first-order*)
@@ -454,7 +454,7 @@
fold count_constants_lit literals (const_min_arity, const_needs_hBOOL);
fun display_arity const_needs_hBOOL (c,n) =
- ResAxioms.trace_msg (fn () => "Constant: " ^ c ^ " arity:\t" ^ Int.toString n ^
+ Res_Axioms.trace_msg (fn () => "Constant: " ^ c ^ " arity:\t" ^ Int.toString n ^
(if needs_hBOOL const_needs_hBOOL c then " needs hBOOL" else ""));
fun count_constants (conjectures, _, extra_clauses, helper_clauses, _, _) =
@@ -527,4 +527,5 @@
length helper_clauses + 1, length tfree_clss + length dfg_clss)
end;
-end
+end;
+
--- a/src/HOL/Tools/res_reconstruct.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/res_reconstruct.ML Thu Oct 29 18:53:58 2009 +0100
@@ -1,8 +1,9 @@
-(* Author: L C Paulson and Claire Quigley, Cambridge University Computer Laboratory *)
+(* Title: HOL/Tools/res_reconstruct.ML
+ Author: Lawrence C Paulson and Claire Quigley, Cambridge University Computer Laboratory
-(***************************************************************************)
-(* Code to deal with the transfer of proofs from a prover process *)
-(***************************************************************************)
+Transfer of proofs from external provers.
+*)
+
signature RES_RECONSTRUCT =
sig
val chained_hint: string
@@ -23,13 +24,13 @@
string * string vector * (int * int) * Proof.context * thm * int -> string * string list
end;
-structure ResReconstruct : RES_RECONSTRUCT =
+structure Res_Reconstruct : RES_RECONSTRUCT =
struct
val trace_path = Path.basic "atp_trace";
fun trace s =
- if ! ResAxioms.trace then File.append (File.tmp_path trace_path) s
+ if ! Res_Axioms.trace then File.append (File.tmp_path trace_path) s
else ();
fun string_of_thm ctxt = PrintMode.setmp [] (Display.string_of_thm ctxt);
@@ -106,12 +107,12 @@
(*If string s has the prefix s1, return the result of deleting it.*)
fun strip_prefix s1 s =
if String.isPrefix s1 s
- then SOME (ResClause.undo_ascii_of (String.extract (s, size s1, NONE)))
+ then SOME (Res_Clause.undo_ascii_of (String.extract (s, size s1, NONE)))
else NONE;
(*Invert the table of translations between Isabelle and ATPs*)
val type_const_trans_table_inv =
- Symtab.make (map swap (Symtab.dest ResClause.type_const_trans_table));
+ Symtab.make (map swap (Symtab.dest Res_Clause.type_const_trans_table));
fun invert_type_const c =
case Symtab.lookup type_const_trans_table_inv c of
@@ -129,15 +130,15 @@
| Br (a,ts) =>
let val Ts = map type_of_stree ts
in
- case strip_prefix ResClause.tconst_prefix a of
+ case strip_prefix Res_Clause.tconst_prefix a of
SOME b => Type(invert_type_const b, Ts)
| NONE =>
if not (null ts) then raise STREE t (*only tconsts have type arguments*)
else
- case strip_prefix ResClause.tfree_prefix a of
+ case strip_prefix Res_Clause.tfree_prefix a of
SOME b => TFree("'" ^ b, HOLogic.typeS)
| NONE =>
- case strip_prefix ResClause.tvar_prefix a of
+ case strip_prefix Res_Clause.tvar_prefix a of
SOME b => make_tvar b
| NONE => make_tvar a (*Variable from the ATP, say X1*)
end;
@@ -145,7 +146,7 @@
(*Invert the table of translations between Isabelle and ATPs*)
val const_trans_table_inv =
Symtab.update ("fequal", "op =")
- (Symtab.make (map swap (Symtab.dest ResClause.const_trans_table)));
+ (Symtab.make (map swap (Symtab.dest Res_Clause.const_trans_table)));
fun invert_const c =
case Symtab.lookup const_trans_table_inv c of
@@ -166,7 +167,7 @@
| Br ("hBOOL",[t]) => term_of_stree [] thy t (*ignore hBOOL*)
| Br ("hAPP",[t,u]) => term_of_stree (u::args) thy t
| Br (a,ts) =>
- case strip_prefix ResClause.const_prefix a of
+ case strip_prefix Res_Clause.const_prefix a of
SOME "equal" =>
list_comb(Const ("op =", HOLogic.typeT), List.map (term_of_stree [] thy) ts)
| SOME b =>
@@ -180,10 +181,10 @@
| NONE => (*a variable, not a constant*)
let val T = HOLogic.typeT
val opr = (*a Free variable is typically a Skolem function*)
- case strip_prefix ResClause.fixed_var_prefix a of
+ case strip_prefix Res_Clause.fixed_var_prefix a of
SOME b => Free(b,T)
| NONE =>
- case strip_prefix ResClause.schematic_var_prefix a of
+ case strip_prefix Res_Clause.schematic_var_prefix a of
SOME b => make_var (b,T)
| NONE => make_var (a,T) (*Variable from the ATP, say X1*)
in list_comb (opr, List.map (term_of_stree [] thy) (ts@args)) end;
@@ -193,7 +194,7 @@
| constraint_of_stree pol t = case t of
Int _ => raise STREE t
| Br (a,ts) =>
- (case (strip_prefix ResClause.class_prefix a, map type_of_stree ts) of
+ (case (strip_prefix Res_Clause.class_prefix a, map type_of_stree ts) of
(SOME b, [T]) => (pol, b, T)
| _ => raise STREE t);
@@ -443,11 +444,11 @@
val _ = trace (Int.toString (length nonnull_lines) ^ " nonnull_lines extracted\n")
val (_,lines) = List.foldr (add_wanted_prfline ctxt) (0,[]) nonnull_lines
val _ = trace (Int.toString (length lines) ^ " lines extracted\n")
- val (ccls,fixes) = ResAxioms.neg_conjecture_clauses ctxt th sgno
+ val (ccls,fixes) = Res_Axioms.neg_conjecture_clauses ctxt th sgno
val _ = trace (Int.toString (length ccls) ^ " conjecture clauses\n")
val ccls = map forall_intr_vars ccls
val _ =
- if ! ResAxioms.trace then app (fn th => trace ("\nccl: " ^ string_of_thm ctxt th)) ccls
+ if ! Res_Axioms.trace then app (fn th => trace ("\nccl: " ^ string_of_thm ctxt th)) ccls
else ()
val ilines = isar_lines ctxt (map prop_of ccls) (stringify_deps thm_names [] lines)
val _ = trace "\ndecode_tstp_file: finishing\n"
@@ -456,124 +457,128 @@
end;
- (*=== EXTRACTING PROOF-TEXT === *)
+(*=== EXTRACTING PROOF-TEXT === *)
- val begin_proof_strings = ["# SZS output start CNFRefutation.",
- "=========== Refutation ==========",
+val begin_proof_strings = ["# SZS output start CNFRefutation.",
+ "=========== Refutation ==========",
"Here is a proof"];
- val end_proof_strings = ["# SZS output end CNFRefutation",
- "======= End of refutation =======",
+
+val end_proof_strings = ["# SZS output end CNFRefutation",
+ "======= End of refutation =======",
"Formulae used in the proof"];
- fun get_proof_extract proof =
- let
+
+fun get_proof_extract proof =
+ let
(*splits to_split by the first possible of a list of splitters*)
val (begin_string, end_string) =
(find_first (fn s => String.isSubstring s proof) begin_proof_strings,
find_first (fn s => String.isSubstring s proof) end_proof_strings)
- in
- if is_none begin_string orelse is_none end_string
- then error "Could not extract proof (no substring indicating a proof)"
- else proof |> first_field (the begin_string) |> the |> snd
- |> first_field (the end_string) |> the |> fst end;
+ in
+ if is_none begin_string orelse is_none end_string
+ then error "Could not extract proof (no substring indicating a proof)"
+ else proof |> first_field (the begin_string) |> the |> snd
+ |> first_field (the end_string) |> the |> fst
+ end;
(* ==== CHECK IF PROOF OF E OR VAMPIRE WAS SUCCESSFUL === *)
- val failure_strings_E = ["SZS status: Satisfiable","SZS status Satisfiable",
- "SZS status: ResourceOut","SZS status ResourceOut","# Cannot determine problem status"];
- val failure_strings_vampire = ["Satisfiability detected", "Refutation not found", "CANNOT PROVE"];
- val failure_strings_SPASS = ["SPASS beiseite: Completion found.",
- "SPASS beiseite: Ran out of time.", "SPASS beiseite: Maximal number of loops exceeded."];
- val failure_strings_remote = ["Remote-script could not extract proof"];
- fun find_failure proof =
- let val failures =
- map_filter (fn s => if String.isSubstring s proof then SOME s else NONE)
- (failure_strings_E @ failure_strings_vampire @ failure_strings_SPASS @ failure_strings_remote)
- val correct = null failures andalso
- exists (fn s => String.isSubstring s proof) begin_proof_strings andalso
- exists (fn s => String.isSubstring s proof) end_proof_strings
- in
- if correct then NONE
- else if null failures then SOME "Output of ATP not in proper format"
- else SOME (hd failures) end;
+val failure_strings_E = ["SZS status: Satisfiable","SZS status Satisfiable",
+ "SZS status: ResourceOut","SZS status ResourceOut","# Cannot determine problem status"];
+val failure_strings_vampire = ["Satisfiability detected", "Refutation not found", "CANNOT PROVE"];
+val failure_strings_SPASS = ["SPASS beiseite: Completion found.",
+ "SPASS beiseite: Ran out of time.", "SPASS beiseite: Maximal number of loops exceeded."];
+val failure_strings_remote = ["Remote-script could not extract proof"];
+fun find_failure proof =
+ let val failures =
+ map_filter (fn s => if String.isSubstring s proof then SOME s else NONE)
+ (failure_strings_E @ failure_strings_vampire @ failure_strings_SPASS @ failure_strings_remote)
+ val correct = null failures andalso
+ exists (fn s => String.isSubstring s proof) begin_proof_strings andalso
+ exists (fn s => String.isSubstring s proof) end_proof_strings
+ in
+ if correct then NONE
+ else if null failures then SOME "Output of ATP not in proper format"
+ else SOME (hd failures) end;
- (* === EXTRACTING LEMMAS === *)
- (* lines have the form "cnf(108, axiom, ...",
- the number (108) has to be extracted)*)
- fun get_step_nums false proofextract =
- let val toks = String.tokens (not o Char.isAlphaNum)
- fun inputno ("cnf"::ntok::"axiom"::_) = Int.fromString ntok
- | inputno ("cnf"::ntok::"negated"::"conjecture"::_) = Int.fromString ntok
- | inputno _ = NONE
- val lines = split_lines proofextract
- in map_filter (inputno o toks) lines end
- (*String contains multiple lines. We want those of the form
- "253[0:Inp] et cetera..."
- A list consisting of the first number in each line is returned. *)
- | get_step_nums true proofextract =
- let val toks = String.tokens (not o Char.isAlphaNum)
- fun inputno (ntok::"0"::"Inp"::_) = Int.fromString ntok
- | inputno _ = NONE
- val lines = split_lines proofextract
- in map_filter (inputno o toks) lines end
-
- (*extracting lemmas from tstp-output between the lines from above*)
- fun extract_lemmas get_step_nums (proof, thm_names, conj_count, _, _, _) =
+(* === EXTRACTING LEMMAS === *)
+(* lines have the form "cnf(108, axiom, ...",
+the number (108) has to be extracted)*)
+fun get_step_nums false proofextract =
+ let val toks = String.tokens (not o Char.isAlphaNum)
+ fun inputno ("cnf"::ntok::"axiom"::_) = Int.fromString ntok
+ | inputno ("cnf"::ntok::"negated"::"conjecture"::_) = Int.fromString ntok
+ | inputno _ = NONE
+ val lines = split_lines proofextract
+ in map_filter (inputno o toks) lines end
+(*String contains multiple lines. We want those of the form
+ "253[0:Inp] et cetera..."
+ A list consisting of the first number in each line is returned. *)
+| get_step_nums true proofextract =
+ let val toks = String.tokens (not o Char.isAlphaNum)
+ fun inputno (ntok::"0"::"Inp"::_) = Int.fromString ntok
+ | inputno _ = NONE
+ val lines = split_lines proofextract
+ in map_filter (inputno o toks) lines end
+
+(*extracting lemmas from tstp-output between the lines from above*)
+fun extract_lemmas get_step_nums (proof, thm_names, conj_count, _, _, _) =
+ let
+ (* get the names of axioms from their numbers*)
+ fun get_axiom_names thm_names step_nums =
let
- (* get the names of axioms from their numbers*)
- fun get_axiom_names thm_names step_nums =
- let
- val last_axiom = Vector.length thm_names
- fun is_axiom n = n <= last_axiom
- fun is_conj n = n >= fst conj_count andalso n < fst conj_count + snd conj_count
- fun getname i = Vector.sub(thm_names, i-1)
- in
- (sort_distinct string_ord (filter (fn x => x <> "??.unknown")
- (map getname (filter is_axiom step_nums))),
- exists is_conj step_nums)
- end
- val proofextract = get_proof_extract proof
+ val last_axiom = Vector.length thm_names
+ fun is_axiom n = n <= last_axiom
+ fun is_conj n = n >= fst conj_count andalso n < fst conj_count + snd conj_count
+ fun getname i = Vector.sub(thm_names, i-1)
in
- get_axiom_names thm_names (get_step_nums proofextract)
- end;
+ (sort_distinct string_ord (filter (fn x => x <> "??.unknown")
+ (map getname (filter is_axiom step_nums))),
+ exists is_conj step_nums)
+ end
+ val proofextract = get_proof_extract proof
+ in
+ get_axiom_names thm_names (get_step_nums proofextract)
+ end;
- (*Used to label theorems chained into the sledgehammer call*)
- val chained_hint = "CHAINED";
- val nochained = filter_out (fn y => y = chained_hint)
-
- (* metis-command *)
- fun metis_line [] = "apply metis"
- | metis_line xs = "apply (metis " ^ space_implode " " xs ^ ")"
+(*Used to label theorems chained into the sledgehammer call*)
+val chained_hint = "CHAINED";
+val nochained = filter_out (fn y => y = chained_hint)
+
+(* metis-command *)
+fun metis_line [] = "apply metis"
+ | metis_line xs = "apply (metis " ^ space_implode " " xs ^ ")"
- (* atp_minimize [atp=<prover>] <lemmas> *)
- fun minimize_line _ [] = ""
- | minimize_line name lemmas = "For minimizing the number of lemmas try this command:\n" ^
- (Markup.markup Markup.sendback) ("atp_minimize [atp=" ^ name ^ "] " ^
- space_implode " " (nochained lemmas))
+(* atp_minimize [atp=<prover>] <lemmas> *)
+fun minimize_line _ [] = ""
+ | minimize_line name lemmas = "For minimizing the number of lemmas try this command:\n" ^
+ (Markup.markup Markup.sendback) ("atp_minimize [atp=" ^ name ^ "] " ^
+ space_implode " " (nochained lemmas))
- fun sendback_metis_nochained lemmas =
- (Markup.markup Markup.sendback o metis_line) (nochained lemmas)
+fun sendback_metis_nochained lemmas =
+ (Markup.markup Markup.sendback o metis_line) (nochained lemmas)
- fun lemma_list dfg name result =
- let val (lemmas, used_conj) = extract_lemmas (get_step_nums dfg) result
- in (sendback_metis_nochained lemmas ^ "\n" ^ minimize_line name lemmas ^
- (if used_conj then ""
- else "\nWarning: Goal is provable because context is inconsistent."),
- nochained lemmas)
- end;
+fun lemma_list dfg name result =
+ let val (lemmas, used_conj) = extract_lemmas (get_step_nums dfg) result
+ in (sendback_metis_nochained lemmas ^ "\n" ^ minimize_line name lemmas ^
+ (if used_conj then ""
+ else "\nWarning: Goal is provable because context is inconsistent."),
+ nochained lemmas)
+ end;
- (* === Extracting structured Isar-proof === *)
- fun structured_proof name (result as (proof, thm_names, conj_count, ctxt, goal, subgoalno)) =
- let
- (*Could use split_lines, but it can return blank lines...*)
- val lines = String.tokens (equal #"\n");
- val nospaces = String.translate (fn c => if Char.isSpace c then "" else str c)
- val proofextract = get_proof_extract proof
- val cnfs = filter (String.isPrefix "cnf(") (map nospaces (lines proofextract))
- val (one_line_proof, lemma_names) = lemma_list false name result
- val structured = if chained_hint mem_string (String.tokens (fn c => c = #" ") one_line_proof) then ""
- else decode_tstp_file cnfs ctxt goal subgoalno thm_names
- in
- (one_line_proof ^ "\n\n" ^ (Markup.markup Markup.sendback) structured, lemma_names)
- end
+(* === Extracting structured Isar-proof === *)
+fun structured_proof name (result as (proof, thm_names, conj_count, ctxt, goal, subgoalno)) =
+ let
+ (*Could use split_lines, but it can return blank lines...*)
+ val lines = String.tokens (equal #"\n");
+ val nospaces = String.translate (fn c => if Char.isSpace c then "" else str c)
+ val proofextract = get_proof_extract proof
+ val cnfs = filter (String.isPrefix "cnf(") (map nospaces (lines proofextract))
+ val (one_line_proof, lemma_names) = lemma_list false name result
+ val structured =
+ if chained_hint mem_string (String.tokens (fn c => c = #" ") one_line_proof) then ""
+ else decode_tstp_file cnfs ctxt goal subgoalno thm_names
+ in
+ (one_line_proof ^ "\n\n" ^ Markup.markup Markup.sendback structured, lemma_names)
+end
end;
--- a/src/HOL/Tools/sat_solver.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/sat_solver.ML Thu Oct 29 18:53:58 2009 +0100
@@ -272,7 +272,7 @@
else
parse_lines lines
in
- (parse_lines o (List.filter (fn l => l <> "")) o split_lines o File.read) path
+ (parse_lines o filter (fn l => l <> "") o split_lines o File.read) path
end;
(* ------------------------------------------------------------------------- *)
@@ -352,7 +352,7 @@
o (map int_from_string)
o (maps (String.fields (fn c => c mem [#" ", #"\t", #"\n"])))
o filter_preamble
- o (List.filter (fn l => l <> ""))
+ o filter (fn l => l <> "")
o split_lines
o File.read)
path
--- a/src/HOL/Tools/typecopy.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/typecopy.ML Thu Oct 29 18:53:58 2009 +0100
@@ -44,8 +44,8 @@
(* interpretation of type copies *)
-structure TypecopyInterpretation = InterpretationFun(type T = string val eq = op =);
-val interpretation = TypecopyInterpretation.interpretation;
+structure Typecopy_Interpretation = Interpretation(type T = string val eq = op =);
+val interpretation = Typecopy_Interpretation.interpretation;
(* introducing typecopies *)
@@ -76,7 +76,7 @@
in
thy
|> (TypecopyData.map o Symtab.update_new) (tyco, info)
- |> TypecopyInterpretation.data tyco
+ |> Typecopy_Interpretation.data tyco
|> pair (tyco, info)
end
in
@@ -126,7 +126,7 @@
end;
val setup =
- TypecopyInterpretation.init
+ Typecopy_Interpretation.init
#> interpretation add_default_code
end;
--- a/src/HOL/Tools/typedef.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/Tools/typedef.ML Thu Oct 29 18:53:58 2009 +0100
@@ -53,8 +53,8 @@
fun declare_type_name a = Variable.declare_constraints (Logic.mk_type (TFree (a, dummyS)));
-structure TypedefInterpretation = InterpretationFun(type T = string val eq = op =);
-val interpretation = TypedefInterpretation.interpretation;
+structure Typedef_Interpretation = Interpretation(type T = string val eq = op =);
+val interpretation = Typedef_Interpretation.interpretation;
fun prepare_typedef prep_term def name (t, vs, mx) raw_set opt_morphs thy =
let
@@ -169,7 +169,7 @@
in
thy2
|> put_info full_tname info
- |> TypedefInterpretation.data full_tname
+ |> Typedef_Interpretation.data full_tname
|> pair (full_tname, info)
end);
@@ -264,6 +264,6 @@
end;
-val setup = TypedefInterpretation.init;
+val setup = Typedef_Interpretation.init;
end;
--- a/src/HOL/ex/Predicate_Compile_Alternative_Defs.thy Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/ex/Predicate_Compile_Alternative_Defs.thy Thu Oct 29 18:53:58 2009 +0100
@@ -45,7 +45,7 @@
unfolding mem_def[symmetric, of _ a2]
apply auto
unfolding mem_def
- apply auto
+ apply fastsimp
done
qed
--- a/src/HOL/ex/Predicate_Compile_ex.thy Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOL/ex/Predicate_Compile_ex.thy Thu Oct 29 18:53:58 2009 +0100
@@ -1,12 +1,12 @@
theory Predicate_Compile_ex
-imports Main Predicate_Compile_Alternative_Defs
+imports "../Main" Predicate_Compile_Alternative_Defs
begin
subsection {* Basic predicates *}
inductive False' :: "bool"
-code_pred (mode: []) False' .
+code_pred (mode : []) False' .
code_pred [depth_limited] False' .
code_pred [rpred] False' .
@@ -17,7 +17,7 @@
definition EmptySet' :: "'a \<Rightarrow> bool"
where "EmptySet' = {}"
-code_pred (mode: [], [1]) [inductify, show_intermediate_results] EmptySet' .
+code_pred (mode: [], [1]) [inductify] EmptySet' .
inductive EmptyRel :: "'a \<Rightarrow> 'b \<Rightarrow> bool"
@@ -26,7 +26,13 @@
inductive EmptyClosure :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
-code_pred (mode: [], [1], [2], [1, 2])EmptyClosure .
+code_pred
+ (mode: [] ==> [], [] ==> [1], [] ==> [2], [] ==> [1, 2],
+ [1] ==> [], [1] ==> [1], [1] ==> [2], [1] ==> [1, 2],
+ [2] ==> [], [2] ==> [1], [2] ==> [2], [2] ==> [1, 2],
+ [1, 2] ==> [], [1, 2] ==> [1], [1, 2] ==> [2], [1, 2] ==> [1, 2])
+ EmptyClosure .
+
thm EmptyClosure.equation
(* TODO: inductive package is broken!
inductive False'' :: "bool"
@@ -60,8 +66,88 @@
where
"zerozero (0, 0)"
-code_pred zerozero .
-code_pred [rpred, show_compilation] zerozero .
+code_pred (mode: [i], [(i, o)], [(o, i)], [o]) zerozero .
+code_pred [rpred] zerozero .
+
+subsection {* Alternative Rules *}
+
+datatype char = C | D | E | F | G | H
+
+inductive is_C_or_D
+where
+ "(x = C) \<or> (x = D) ==> is_C_or_D x"
+
+code_pred (mode: [1]) is_C_or_D .
+thm is_C_or_D.equation
+
+inductive is_D_or_E
+where
+ "(x = D) \<or> (x = E) ==> is_D_or_E x"
+
+lemma [code_pred_intros]:
+ "is_D_or_E D"
+by (auto intro: is_D_or_E.intros)
+
+lemma [code_pred_intros]:
+ "is_D_or_E E"
+by (auto intro: is_D_or_E.intros)
+
+code_pred (mode: [], [1]) is_D_or_E
+proof -
+ case is_D_or_E
+ from this(1) show thesis
+ proof
+ fix x
+ assume x: "a1 = x"
+ assume "x = D \<or> x = E"
+ from this show thesis
+ proof
+ assume "x = D" from this x is_D_or_E(2) show thesis by simp
+ next
+ assume "x = E" from this x is_D_or_E(3) show thesis by simp
+ qed
+ qed
+qed
+
+thm is_D_or_E.equation
+
+inductive is_F_or_G
+where
+ "x = F \<or> x = G ==> is_F_or_G x"
+
+lemma [code_pred_intros]:
+ "is_F_or_G F"
+by (auto intro: is_F_or_G.intros)
+
+lemma [code_pred_intros]:
+ "is_F_or_G G"
+by (auto intro: is_F_or_G.intros)
+
+inductive is_FGH
+where
+ "is_F_or_G x ==> is_FGH x"
+| "is_FGH H"
+
+text {* Compilation of is_FGH requires elimination rule for is_F_or_G *}
+
+code_pred (mode: [], [1]) is_FGH
+proof -
+ case is_F_or_G
+ from this(1) show thesis
+ proof
+ fix x
+ assume x: "a1 = x"
+ assume "x = F \<or> x = G"
+ from this show thesis
+ proof
+ assume "x = F"
+ from this x is_F_or_G(2) show thesis by simp
+ next
+ assume "x = G"
+ from this x is_F_or_G(3) show thesis by simp
+ qed
+ qed
+qed
subsection {* Preprocessor Inlining *}
@@ -123,7 +209,7 @@
definition odd' where "odd' x == \<not> even x"
-code_pred [inductify] odd' .
+code_pred (mode: [1]) [inductify] odd' .
code_pred [inductify, depth_limited] odd' .
code_pred [inductify, rpred] odd' .
@@ -135,7 +221,7 @@
where
"n mod 2 = 0 \<Longrightarrow> is_even n"
-code_pred is_even .
+code_pred (mode: [1]) is_even .
subsection {* append predicate *}
@@ -172,10 +258,19 @@
lemmas [code_pred_intros] = append2_Nil append2.intros(2)
-code_pred append2
+code_pred (mode: [1, 2], [3], [2, 3], [1, 3], [1, 2, 3]) append2
proof -
case append2
- from append2.cases[OF append2(1)] append2(2-3) show thesis by blast
+ from append2(1) show thesis
+ proof
+ fix xs
+ assume "a1 = []" "a2 = xs" "a3 = xs"
+ from this append2(2) show thesis by simp
+ next
+ fix xs ys zs x
+ assume "a1 = x # xs" "a2 = ys" "a3 = x # zs" "append2 xs ys zs"
+ from this append2(3) show thesis by fastsimp
+ qed
qed
inductive tupled_append :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
@@ -183,7 +278,7 @@
"tupled_append ([], xs, xs)"
| "tupled_append (xs, ys, zs) \<Longrightarrow> tupled_append (x # xs, ys, x # zs)"
-code_pred tupled_append .
+code_pred (mode: [(i,i,o)], [(i,o,i)], [(o,i,i)], [(o,o,i)], [i]) tupled_append .
code_pred [rpred] tupled_append .
thm tupled_append.equation
(*
@@ -197,7 +292,7 @@
| "[| ys = fst (xa, y); x # zs = snd (xa, y);
tupled_append' (xs, ys, zs) |] ==> tupled_append' (x # xs, xa, y)"
-code_pred tupled_append' .
+code_pred (mode: [(i,i,o)], [(i,o,i)], [(o,i,i)], [(o,o,i)], [i]) tupled_append' .
thm tupled_append'.equation
inductive tupled_append'' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
@@ -205,9 +300,7 @@
"tupled_append'' ([], xs, xs)"
| "ys = fst yszs ==> x # zs = snd yszs ==> tupled_append'' (xs, ys, zs) \<Longrightarrow> tupled_append'' (x # xs, yszs)"
-thm tupled_append''.cases
-
-code_pred [inductify] tupled_append'' .
+code_pred (mode: [(i,i,o)], [(i,o,i)], [(o,i,i)], [(o,o,i)], [i]) [inductify] tupled_append'' .
thm tupled_append''.equation
inductive tupled_append''' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
@@ -215,7 +308,7 @@
"tupled_append''' ([], xs, xs)"
| "yszs = (ys, zs) ==> tupled_append''' (xs, yszs) \<Longrightarrow> tupled_append''' (x # xs, ys, x # zs)"
-code_pred [inductify] tupled_append''' .
+code_pred (mode: [(i,i,o)], [(i,o,i)], [(o,i,i)], [(o,o,i)], [i]) [inductify] tupled_append''' .
thm tupled_append'''.equation
subsection {* map_ofP predicate *}
@@ -237,7 +330,7 @@
| "P x ==> filter1 P xs ys ==> filter1 P (x#xs) (x#ys)"
| "\<not> P x ==> filter1 P xs ys ==> filter1 P (x#xs) ys"
-code_pred (mode: [1], [1, 2]) filter1 .
+code_pred (mode: [1] ==> [1], [1] ==> [1, 2]) filter1 .
code_pred [depth_limited] filter1 .
code_pred [rpred] filter1 .
@@ -260,7 +353,7 @@
where
"List.filter P xs = ys ==> filter3 P xs ys"
-code_pred filter3 .
+code_pred (mode: [] ==> [1], [] ==> [1, 2], [1] ==> [1], [1] ==> [1, 2]) filter3 .
code_pred [depth_limited] filter3 .
thm filter3.depth_limited_equation
@@ -268,7 +361,7 @@
where
"List.filter P xs = ys ==> filter4 P xs ys"
-code_pred filter4 .
+code_pred (mode: [1, 2], [1, 2, 3]) filter4 .
code_pred [depth_limited] filter4 .
code_pred [rpred] filter4 .
@@ -288,7 +381,7 @@
"tupled_rev ([], [])"
| "tupled_rev (xs, xs') \<Longrightarrow> tupled_append (xs', [x], ys) \<Longrightarrow> tupled_rev (x#xs, ys)"
-code_pred tupled_rev .
+code_pred (mode: [(i, o)], [(o, i)], [i]) tupled_rev .
thm tupled_rev.equation
subsection {* partition predicate *}
@@ -299,7 +392,7 @@
| "f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) (x # ys) zs"
| "\<not> f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) ys (x # zs)"
-code_pred (mode: [1], [2, 3], [1, 2], [1, 3], [1, 2, 3]) partition .
+code_pred (mode: [1] ==> [1], [1] ==> [2, 3], [1] ==> [1, 2], [1] ==> [1, 3], [1] ==> [1, 2, 3]) partition .
code_pred [depth_limited] partition .
code_pred [rpred] partition .
@@ -314,7 +407,7 @@
| "f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, x # ys, zs)"
| "\<not> f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, ys, x # zs)"
-code_pred tupled_partition .
+code_pred (mode: [i] ==> [i], [i] ==> [(i, i, o)], [i] ==> [(i, o, i)], [i] ==> [(o, i, i)], [i] ==> [(i, o, o)]) tupled_partition .
thm tupled_partition.equation
@@ -325,7 +418,7 @@
subsection {* transitive predicate *}
-code_pred tranclp
+code_pred (mode: [1] ==> [1, 2], [1] ==> [1], [2] ==> [1, 2], [2] ==> [2], [] ==> [1, 2], [] ==> [1], [] ==> [2], [] ==> []) tranclp
proof -
case tranclp
from this converse_tranclpE[OF this(1)] show thesis by metis
@@ -658,6 +751,8 @@
| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
+code_pred (mode: [], [1]) S\<^isub>4p .
+
subsection {* Lambda *}
datatype type =
@@ -708,4 +803,10 @@
| appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
| abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"
+code_pred (mode: [1, 2], [1, 2, 3]) typing .
+thm typing.equation
+
+code_pred (mode: [1], [1, 2]) beta .
+thm beta.equation
+
end
\ No newline at end of file
--- a/src/HOL/ex/predicate_compile.ML Thu Oct 29 16:22:14 2009 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,2182 +0,0 @@
-(* Author: Lukas Bulwahn, TU Muenchen
-
-(Prototype of) A compiler from predicates specified by intro/elim rules
-to equations.
-*)
-
-signature PREDICATE_COMPILE =
-sig
- type mode = int list option list * int list
- (*val add_equations_of: bool -> string list -> theory -> theory *)
- val register_predicate : (thm list * thm * int) -> theory -> theory
- val is_registered : theory -> string -> bool
- (* val fetch_pred_data : theory -> string -> (thm list * thm * int) *)
- val predfun_intro_of: theory -> string -> mode -> thm
- val predfun_elim_of: theory -> string -> mode -> thm
- val strip_intro_concl: int -> term -> term * (term list * term list)
- val predfun_name_of: theory -> string -> mode -> string
- val all_preds_of : theory -> string list
- val modes_of: theory -> string -> mode list
- val string_of_mode : mode -> string
- val intros_of: theory -> string -> thm list
- val nparams_of: theory -> string -> int
- val add_intro: thm -> theory -> theory
- val set_elim: thm -> theory -> theory
- val setup: theory -> theory
- val code_pred: string -> Proof.context -> Proof.state
- val code_pred_cmd: string -> Proof.context -> Proof.state
- val print_stored_rules: theory -> unit
- val print_all_modes: theory -> unit
- val do_proofs: bool Unsynchronized.ref
- val mk_casesrule : Proof.context -> int -> thm list -> term
- val analyze_compr: theory -> term -> term
- val eval_ref: (unit -> term Predicate.pred) option Unsynchronized.ref
- val add_equations : string list -> theory -> theory
- val code_pred_intros_attrib : attribute
- (* used by Quickcheck_Generator *)
- (*val funT_of : mode -> typ -> typ
- val mk_if_pred : term -> term
- val mk_Eval : term * term -> term*)
- val mk_tupleT : typ list -> typ
-(* val mk_predT : typ -> typ *)
- (* temporary for testing of the compilation *)
- datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term |
- GeneratorPrem of term list * term | Generator of (string * typ);
- val prepare_intrs: theory -> string list ->
- (string * typ) list * int * string list * string list * (string * mode list) list *
- (string * (term list * indprem list) list) list * (string * (int option list * int)) list
- datatype compilation_funs = CompilationFuns of {
- mk_predT : typ -> typ,
- dest_predT : typ -> typ,
- mk_bot : typ -> term,
- mk_single : term -> term,
- mk_bind : term * term -> term,
- mk_sup : term * term -> term,
- mk_if : term -> term,
- mk_not : term -> term,
- mk_map : typ -> typ -> term -> term -> term,
- lift_pred : term -> term
- };
- datatype tmode = Mode of mode * int list * tmode option list;
- type moded_clause = term list * (indprem * tmode) list
- type 'a pred_mode_table = (string * (mode * 'a) list) list
- val infer_modes : bool -> theory -> (string * (int list option list * int list) list) list
- -> (string * (int option list * int)) list -> string list
- -> (string * (term list * indprem list) list) list
- -> (moded_clause list) pred_mode_table
- val infer_modes_with_generator : theory -> (string * (int list option list * int list) list) list
- -> (string * (int option list * int)) list -> string list
- -> (string * (term list * indprem list) list) list
- -> (moded_clause list) pred_mode_table
- (*val compile_preds : theory -> compilation_funs -> string list -> string list
- -> (string * typ) list -> (moded_clause list) pred_mode_table -> term pred_mode_table
- val rpred_create_definitions :(string * typ) list -> string * mode list
- -> theory -> theory
- val split_smode : int list -> term list -> (term list * term list) *)
- val print_moded_clauses :
- theory -> (moded_clause list) pred_mode_table -> unit
- val print_compiled_terms : theory -> term pred_mode_table -> unit
- (*val rpred_prove_preds : theory -> term pred_mode_table -> thm pred_mode_table*)
- val rpred_compfuns : compilation_funs
- val dest_funT : typ -> typ * typ
- (* val depending_preds_of : theory -> thm list -> string list *)
- val add_quickcheck_equations : string list -> theory -> theory
- val add_sizelim_equations : string list -> theory -> theory
- val is_inductive_predicate : theory -> string -> bool
- val terms_vs : term list -> string list
- val subsets : int -> int -> int list list
- val check_mode_clause : bool -> theory -> string list ->
- (string * mode list) list -> (string * mode list) list -> mode -> (term list * indprem list)
- -> (term list * (indprem * tmode) list) option
- val string_of_moded_prem : theory -> (indprem * tmode) -> string
- val all_modes_of : theory -> (string * mode list) list
- val all_generator_modes_of : theory -> (string * mode list) list
- val compile_clause : compilation_funs -> term option -> (term list -> term) ->
- theory -> string list -> string list -> mode -> term -> moded_clause -> term
- val preprocess_intro : theory -> thm -> thm
- val is_constrt : theory -> term -> bool
- val is_predT : typ -> bool
- val guess_nparams : typ -> int
-end;
-
-structure Predicate_Compile : PREDICATE_COMPILE =
-struct
-
-(** auxiliary **)
-
-(* debug stuff *)
-
-fun tracing s = (if ! Toplevel.debug then tracing s else ());
-
-fun print_tac s = Seq.single; (* (if ! Toplevel.debug then Tactical.print_tac s else Seq.single); *)
-fun debug_tac msg = Seq.single; (* (fn st => (tracing msg; Seq.single st)); *)
-
-val do_proofs = Unsynchronized.ref true;
-
-fun mycheat_tac thy i st =
- (Tactic.rtac (Skip_Proof.make_thm thy (Var (("A", 0), propT))) i) st
-
-fun remove_last_goal thy st =
- (Tactic.rtac (Skip_Proof.make_thm thy (Var (("A", 0), propT))) (nprems_of st)) st
-
-(* reference to preprocessing of InductiveSet package *)
-
-val ind_set_codegen_preproc = Inductive_Set.codegen_preproc;
-
-(** fundamentals **)
-
-(* syntactic operations *)
-
-fun mk_eq (x, xs) =
- let fun mk_eqs _ [] = []
- | mk_eqs a (b::cs) =
- HOLogic.mk_eq (Free (a, fastype_of b), b) :: mk_eqs a cs
- in mk_eqs x xs end;
-
-fun mk_tupleT [] = HOLogic.unitT
- | mk_tupleT Ts = foldr1 HOLogic.mk_prodT Ts;
-
-fun dest_tupleT (Type (@{type_name Product_Type.unit}, [])) = []
- | dest_tupleT (Type (@{type_name "*"}, [T1, T2])) = T1 :: (dest_tupleT T2)
- | dest_tupleT t = [t]
-
-fun mk_tuple [] = HOLogic.unit
- | mk_tuple ts = foldr1 HOLogic.mk_prod ts;
-
-fun dest_tuple (Const (@{const_name Product_Type.Unity}, _)) = []
- | dest_tuple (Const (@{const_name Pair}, _) $ t1 $ t2) = t1 :: (dest_tuple t2)
- | dest_tuple t = [t]
-
-fun mk_scomp (t, u) =
- let
- val T = fastype_of t
- val U = fastype_of u
- val [A] = binder_types T
- val D = body_type U
- in
- Const (@{const_name "scomp"}, T --> U --> A --> D) $ t $ u
- end;
-
-fun dest_funT (Type ("fun",[S, T])) = (S, T)
- | dest_funT T = raise TYPE ("dest_funT", [T], [])
-
-fun mk_fun_comp (t, u) =
- let
- val (_, B) = dest_funT (fastype_of t)
- val (C, A) = dest_funT (fastype_of u)
- in
- Const(@{const_name "Fun.comp"}, (A --> B) --> (C --> A) --> C --> B) $ t $ u
- end;
-
-fun dest_randomT (Type ("fun", [@{typ Random.seed},
- Type ("*", [Type ("*", [T, @{typ "unit => Code_Evaluation.term"}]) ,@{typ Random.seed}])])) = T
- | dest_randomT T = raise TYPE ("dest_randomT", [T], [])
-
-(* destruction of intro rules *)
-
-(* FIXME: look for other place where this functionality was used before *)
-fun strip_intro_concl nparams intro = let
- val _ $ u = Logic.strip_imp_concl intro
- val (pred, all_args) = strip_comb u
- val (params, args) = chop nparams all_args
-in (pred, (params, args)) end
-
-(** data structures **)
-
-type smode = int list;
-type mode = smode option list * smode;
-datatype tmode = Mode of mode * int list * tmode option list;
-
-fun split_smode is ts =
- let
- fun split_smode' _ _ [] = ([], [])
- | split_smode' is i (t::ts) = (if i mem is then apfst else apsnd) (cons t)
- (split_smode' is (i+1) ts)
- in split_smode' is 1 ts end
-
-fun split_mode (iss, is) ts =
- let
- val (t1, t2) = chop (length iss) ts
- in (t1, split_smode is t2) end
-
-fun string_of_mode (iss, is) = space_implode " -> " (map
- (fn NONE => "X"
- | SOME js => enclose "[" "]" (commas (map string_of_int js)))
- (iss @ [SOME is]));
-
-fun string_of_tmode (Mode (predmode, termmode, param_modes)) =
- "predmode: " ^ (string_of_mode predmode) ^
- (if null param_modes then "" else
- "; " ^ "params: " ^ commas (map (the_default "NONE" o Option.map string_of_tmode) param_modes))
-
-datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term |
- GeneratorPrem of term list * term | Generator of (string * typ);
-
-type moded_clause = term list * (indprem * tmode) list
-type 'a pred_mode_table = (string * (mode * 'a) list) list
-
-datatype predfun_data = PredfunData of {
- name : string,
- definition : thm,
- intro : thm,
- elim : thm
-};
-
-fun rep_predfun_data (PredfunData data) = data;
-fun mk_predfun_data (name, definition, intro, elim) =
- PredfunData {name = name, definition = definition, intro = intro, elim = elim}
-
-datatype function_data = FunctionData of {
- name : string,
- equation : thm option (* is not used at all? *)
-};
-
-fun rep_function_data (FunctionData data) = data;
-fun mk_function_data (name, equation) =
- FunctionData {name = name, equation = equation}
-
-datatype pred_data = PredData of {
- intros : thm list,
- elim : thm option,
- nparams : int,
- functions : (mode * predfun_data) list,
- generators : (mode * function_data) list,
- sizelim_functions : (mode * function_data) list
-};
-
-fun rep_pred_data (PredData data) = data;
-fun mk_pred_data ((intros, elim, nparams), (functions, generators, sizelim_functions)) =
- PredData {intros = intros, elim = elim, nparams = nparams,
- functions = functions, generators = generators, sizelim_functions = sizelim_functions}
-fun map_pred_data f (PredData {intros, elim, nparams, functions, generators, sizelim_functions}) =
- mk_pred_data (f ((intros, elim, nparams), (functions, generators, sizelim_functions)))
-
-fun eq_option eq (NONE, NONE) = true
- | eq_option eq (SOME x, SOME y) = eq (x, y)
- | eq_option eq _ = false
-
-fun eq_pred_data (PredData d1, PredData d2) =
- eq_list (Thm.eq_thm) (#intros d1, #intros d2) andalso
- eq_option (Thm.eq_thm) (#elim d1, #elim d2) andalso
- #nparams d1 = #nparams d2
-
-structure PredData = TheoryDataFun
-(
- type T = pred_data Graph.T;
- val empty = Graph.empty;
- val copy = I;
- val extend = I;
- fun merge _ = Graph.merge eq_pred_data;
-);
-
-(* queries *)
-
-fun lookup_pred_data thy name =
- Option.map rep_pred_data (try (Graph.get_node (PredData.get thy)) name)
-
-fun the_pred_data thy name = case lookup_pred_data thy name
- of NONE => error ("No such predicate " ^ quote name)
- | SOME data => data;
-
-val is_registered = is_some oo lookup_pred_data
-
-val all_preds_of = Graph.keys o PredData.get
-
-val intros_of = #intros oo the_pred_data
-
-fun the_elim_of thy name = case #elim (the_pred_data thy name)
- of NONE => error ("No elimination rule for predicate " ^ quote name)
- | SOME thm => thm
-
-val has_elim = is_some o #elim oo the_pred_data;
-
-val nparams_of = #nparams oo the_pred_data
-
-val modes_of = (map fst) o #functions oo the_pred_data
-
-fun all_modes_of thy = map (fn name => (name, modes_of thy name)) (all_preds_of thy)
-
-val is_compiled = not o null o #functions oo the_pred_data
-
-fun lookup_predfun_data thy name mode =
- Option.map rep_predfun_data (AList.lookup (op =)
- (#functions (the_pred_data thy name)) mode)
-
-fun the_predfun_data thy name mode = case lookup_predfun_data thy name mode
- of NONE => error ("No function defined for mode " ^ string_of_mode mode ^ " of predicate " ^ name)
- | SOME data => data;
-
-val predfun_name_of = #name ooo the_predfun_data
-
-val predfun_definition_of = #definition ooo the_predfun_data
-
-val predfun_intro_of = #intro ooo the_predfun_data
-
-val predfun_elim_of = #elim ooo the_predfun_data
-
-fun lookup_generator_data thy name mode =
- Option.map rep_function_data (AList.lookup (op =)
- (#generators (the_pred_data thy name)) mode)
-
-fun the_generator_data thy name mode = case lookup_generator_data thy name mode
- of NONE => error ("No generator defined for mode " ^ string_of_mode mode ^ " of predicate " ^ name)
- | SOME data => data
-
-val generator_name_of = #name ooo the_generator_data
-
-val generator_modes_of = (map fst) o #generators oo the_pred_data
-
-fun all_generator_modes_of thy =
- map (fn name => (name, generator_modes_of thy name)) (all_preds_of thy)
-
-fun lookup_sizelim_function_data thy name mode =
- Option.map rep_function_data (AList.lookup (op =)
- (#sizelim_functions (the_pred_data thy name)) mode)
-
-fun the_sizelim_function_data thy name mode = case lookup_sizelim_function_data thy name mode
- of NONE => error ("No size-limited function defined for mode " ^ string_of_mode mode
- ^ " of predicate " ^ name)
- | SOME data => data
-
-val sizelim_function_name_of = #name ooo the_sizelim_function_data
-
-(*val generator_modes_of = (map fst) o #generators oo the_pred_data*)
-
-(* diagnostic display functions *)
-
-fun print_modes modes = tracing ("Inferred modes:\n" ^
- cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
- string_of_mode ms)) modes));
-
-fun print_pred_mode_table string_of_entry thy pred_mode_table =
- let
- fun print_mode pred (mode, entry) = "mode : " ^ (string_of_mode mode)
- ^ (string_of_entry pred mode entry)
- fun print_pred (pred, modes) =
- "predicate " ^ pred ^ ": " ^ cat_lines (map (print_mode pred) modes)
- val _ = tracing (cat_lines (map print_pred pred_mode_table))
- in () end;
-
-fun string_of_moded_prem thy (Prem (ts, p), tmode) =
- (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
- "(" ^ (string_of_tmode tmode) ^ ")"
- | string_of_moded_prem thy (GeneratorPrem (ts, p), Mode (predmode, is, _)) =
- (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
- "(generator_mode: " ^ (string_of_mode predmode) ^ ")"
- | string_of_moded_prem thy (Generator (v, T), _) =
- "Generator for " ^ v ^ " of Type " ^ (Syntax.string_of_typ_global thy T)
- | string_of_moded_prem thy (Negprem (ts, p), Mode (_, is, _)) =
- (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
- "(negative mode: " ^ (space_implode ", " (map string_of_int is)) ^ ")"
- | string_of_moded_prem thy (Sidecond t, Mode (_, is, _)) =
- (Syntax.string_of_term_global thy t) ^
- "(sidecond mode: " ^ (space_implode ", " (map string_of_int is)) ^ ")"
- | string_of_moded_prem _ _ = error "string_of_moded_prem: unimplemented"
-
-fun print_moded_clauses thy =
- let
- fun string_of_clause pred mode clauses =
- cat_lines (map (fn (ts, prems) => (space_implode " --> "
- (map (string_of_moded_prem thy) prems)) ^ " --> " ^ pred ^ " "
- ^ (space_implode " " (map (Syntax.string_of_term_global thy) ts))) clauses)
- in print_pred_mode_table string_of_clause thy end;
-
-fun print_compiled_terms thy =
- print_pred_mode_table (fn _ => fn _ => Syntax.string_of_term_global thy) thy
-
-fun print_stored_rules thy =
- let
- val preds = (Graph.keys o PredData.get) thy
- fun print pred () = let
- val _ = writeln ("predicate: " ^ pred)
- val _ = writeln ("number of parameters: " ^ string_of_int (nparams_of thy pred))
- val _ = writeln ("introrules: ")
- val _ = fold (fn thm => fn u => writeln (Display.string_of_thm_global thy thm))
- (rev (intros_of thy pred)) ()
- in
- if (has_elim thy pred) then
- writeln ("elimrule: " ^ Display.string_of_thm_global thy (the_elim_of thy pred))
- else
- writeln ("no elimrule defined")
- end
- in
- fold print preds ()
- end;
-
-fun print_all_modes thy =
- let
- val _ = writeln ("Inferred modes:")
- fun print (pred, modes) u =
- let
- val _ = writeln ("predicate: " ^ pred)
- val _ = writeln ("modes: " ^ (commas (map string_of_mode modes)))
- in u end
- in
- fold print (all_modes_of thy) ()
- end
-
-(** preprocessing rules **)
-
-fun imp_prems_conv cv ct =
- case Thm.term_of ct of
- Const ("==>", _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv) (imp_prems_conv cv) ct
- | _ => Conv.all_conv ct
-
-fun Trueprop_conv cv ct =
- case Thm.term_of ct of
- Const ("Trueprop", _) $ _ => Conv.arg_conv cv ct
- | _ => error "Trueprop_conv"
-
-fun preprocess_intro thy rule =
- Conv.fconv_rule
- (imp_prems_conv
- (Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq})))))
- (Thm.transfer thy rule)
-
-fun preprocess_elim thy nparams elimrule =
- let
- fun replace_eqs (Const ("Trueprop", _) $ (Const ("op =", T) $ lhs $ rhs)) =
- HOLogic.mk_Trueprop (Const (@{const_name Predicate.eq}, T) $ lhs $ rhs)
- | replace_eqs t = t
- val prems = Thm.prems_of elimrule
- val nargs = length (snd (strip_comb (HOLogic.dest_Trueprop (hd prems)))) - nparams
- fun preprocess_case t =
- let
- val params = Logic.strip_params t
- val (assums1, assums2) = chop nargs (Logic.strip_assums_hyp t)
- val assums_hyp' = assums1 @ (map replace_eqs assums2)
- in
- list_all (params, Logic.list_implies (assums_hyp', Logic.strip_assums_concl t))
- end
- val cases' = map preprocess_case (tl prems)
- val elimrule' = Logic.list_implies ((hd prems) :: cases', Thm.concl_of elimrule)
- in
- Thm.equal_elim
- (Thm.symmetric (Conv.implies_concl_conv (MetaSimplifier.rewrite true [@{thm eq_is_eq}])
- (cterm_of thy elimrule')))
- elimrule
- end;
-
-(* special case: predicate with no introduction rule *)
-fun noclause thy predname elim = let
- val T = (Logic.unvarifyT o Sign.the_const_type thy) predname
- val Ts = binder_types T
- val names = Name.variant_list []
- (map (fn i => "x" ^ (string_of_int i)) (1 upto (length Ts)))
- val vs = map2 (curry Free) names Ts
- val clausehd = HOLogic.mk_Trueprop (list_comb (Const (predname, T), vs))
- val intro_t = Logic.mk_implies (@{prop False}, clausehd)
- val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT))
- val elim_t = Logic.list_implies ([clausehd, Logic.mk_implies (@{prop False}, P)], P)
- val intro = Goal.prove (ProofContext.init thy) names [] intro_t
- (fn {...} => etac @{thm FalseE} 1)
- val elim = Goal.prove (ProofContext.init thy) ("P" :: names) [] elim_t
- (fn {...} => etac elim 1)
-in
- ([intro], elim)
-end
-
-fun fetch_pred_data thy name =
- case try (Inductive.the_inductive (ProofContext.init thy)) name of
- SOME (info as (_, result)) =>
- let
- fun is_intro_of intro =
- let
- val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro))
- in (fst (dest_Const const) = name) end;
- val intros = ind_set_codegen_preproc thy ((map (preprocess_intro thy))
- (filter is_intro_of (#intrs result)))
- val pre_elim = nth (#elims result) (find_index (fn s => s = name) (#names (fst info)))
- val nparams = length (Inductive.params_of (#raw_induct result))
- val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams pre_elim)
- val (intros, elim) = if null intros then noclause thy name elim else (intros, elim)
- in
- mk_pred_data ((intros, SOME elim, nparams), ([], [], []))
- end
- | NONE => error ("No such predicate: " ^ quote name)
-
-(* updaters *)
-
-fun apfst3 f (x, y, z) = (f x, y, z)
-fun apsnd3 f (x, y, z) = (x, f y, z)
-fun aptrd3 f (x, y, z) = (x, y, f z)
-
-fun add_predfun name mode data =
- let
- val add = (apsnd o apfst3 o cons) (mode, mk_predfun_data data)
- in PredData.map (Graph.map_node name (map_pred_data add)) end
-
-fun is_inductive_predicate thy name =
- is_some (try (Inductive.the_inductive (ProofContext.init thy)) name)
-
-fun depending_preds_of thy (key, value) =
- let
- val intros = (#intros o rep_pred_data) value
- in
- fold Term.add_const_names (map Thm.prop_of intros) []
- |> filter (fn c => (not (c = key)) andalso (is_inductive_predicate thy c orelse is_registered thy c))
- end;
-
-
-(* code dependency graph *)
-(*
-fun dependencies_of thy name =
- let
- val (intros, elim, nparams) = fetch_pred_data thy name
- val data = mk_pred_data ((intros, SOME elim, nparams), ([], [], []))
- val keys = depending_preds_of thy intros
- in
- (data, keys)
- end;
-*)
-(* guessing number of parameters *)
-fun find_indexes pred xs =
- let
- fun find is n [] = is
- | find is n (x :: xs) = find (if pred x then (n :: is) else is) (n + 1) xs;
- in rev (find [] 0 xs) end;
-
-fun is_predT (T as Type("fun", [_, _])) = (snd (strip_type T) = HOLogic.boolT)
- | is_predT _ = false
-
-fun guess_nparams T =
- let
- val argTs = binder_types T
- val nparams = fold Integer.max
- (map (fn x => x + 1) (find_indexes is_predT argTs)) 0
- in nparams end;
-
-fun add_intro thm thy = let
- val (name, T) = dest_Const (fst (strip_intro_concl 0 (prop_of thm)))
- fun cons_intro gr =
- case try (Graph.get_node gr) name of
- SOME pred_data => Graph.map_node name (map_pred_data
- (apfst (fn (intro, elim, nparams) => (thm::intro, elim, nparams)))) gr
- | NONE =>
- let
- val nparams = the_default (guess_nparams T) (try (#nparams o rep_pred_data o (fetch_pred_data thy)) name)
- in Graph.new_node (name, mk_pred_data (([thm], NONE, nparams), ([], [], []))) gr end;
- in PredData.map cons_intro thy end
-
-fun set_elim thm = let
- val (name, _) = dest_Const (fst
- (strip_comb (HOLogic.dest_Trueprop (hd (prems_of thm)))))
- fun set (intros, _, nparams) = (intros, SOME thm, nparams)
- in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end
-
-fun set_nparams name nparams = let
- fun set (intros, elim, _ ) = (intros, elim, nparams)
- in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end
-
-fun register_predicate (pre_intros, pre_elim, nparams) thy = let
- val (name, _) = dest_Const (fst (strip_intro_concl nparams (prop_of (hd pre_intros))))
- (* preprocessing *)
- val intros = ind_set_codegen_preproc thy (map (preprocess_intro thy) pre_intros)
- val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams pre_elim)
- in
- PredData.map
- (Graph.new_node (name, mk_pred_data ((intros, SOME elim, nparams), ([], [], [])))) thy
- end
-
-fun set_generator_name pred mode name =
- let
- val set = (apsnd o apsnd3 o cons) (mode, mk_function_data (name, NONE))
- in
- PredData.map (Graph.map_node pred (map_pred_data set))
- end
-
-fun set_sizelim_function_name pred mode name =
- let
- val set = (apsnd o aptrd3 o cons) (mode, mk_function_data (name, NONE))
- in
- PredData.map (Graph.map_node pred (map_pred_data set))
- end
-
-(** data structures for generic compilation for different monads **)
-
-(* maybe rename functions more generic:
- mk_predT -> mk_monadT; dest_predT -> dest_monadT
- mk_single -> mk_return (?)
-*)
-datatype compilation_funs = CompilationFuns of {
- mk_predT : typ -> typ,
- dest_predT : typ -> typ,
- mk_bot : typ -> term,
- mk_single : term -> term,
- mk_bind : term * term -> term,
- mk_sup : term * term -> term,
- mk_if : term -> term,
- mk_not : term -> term,
-(* funT_of : mode -> typ -> typ, *)
-(* mk_fun_of : theory -> (string * typ) -> mode -> term, *)
- mk_map : typ -> typ -> term -> term -> term,
- lift_pred : term -> term
-};
-
-fun mk_predT (CompilationFuns funs) = #mk_predT funs
-fun dest_predT (CompilationFuns funs) = #dest_predT funs
-fun mk_bot (CompilationFuns funs) = #mk_bot funs
-fun mk_single (CompilationFuns funs) = #mk_single funs
-fun mk_bind (CompilationFuns funs) = #mk_bind funs
-fun mk_sup (CompilationFuns funs) = #mk_sup funs
-fun mk_if (CompilationFuns funs) = #mk_if funs
-fun mk_not (CompilationFuns funs) = #mk_not funs
-(*fun funT_of (CompilationFuns funs) = #funT_of funs*)
-(*fun mk_fun_of (CompilationFuns funs) = #mk_fun_of funs*)
-fun mk_map (CompilationFuns funs) = #mk_map funs
-fun lift_pred (CompilationFuns funs) = #lift_pred funs
-
-fun funT_of compfuns (iss, is) T =
- let
- val Ts = binder_types T
- val (paramTs, (inargTs, outargTs)) = split_mode (iss, is) Ts
- val paramTs' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) iss paramTs
- in
- (paramTs' @ inargTs) ---> (mk_predT compfuns (mk_tupleT outargTs))
- end;
-
-fun sizelim_funT_of compfuns (iss, is) T =
- let
- val Ts = binder_types T
- val (paramTs, (inargTs, outargTs)) = split_mode (iss, is) Ts
- val paramTs' = map2 (fn SOME is => sizelim_funT_of compfuns ([], is) | NONE => I) iss paramTs
- in
- (paramTs' @ inargTs @ [@{typ "code_numeral"}]) ---> (mk_predT compfuns (mk_tupleT outargTs))
- end;
-
-fun mk_fun_of compfuns thy (name, T) mode =
- Const (predfun_name_of thy name mode, funT_of compfuns mode T)
-
-fun mk_sizelim_fun_of compfuns thy (name, T) mode =
- Const (sizelim_function_name_of thy name mode, sizelim_funT_of compfuns mode T)
-
-fun mk_generator_of compfuns thy (name, T) mode =
- Const (generator_name_of thy name mode, sizelim_funT_of compfuns mode T)
-
-
-structure PredicateCompFuns =
-struct
-
-fun mk_predT T = Type (@{type_name "Predicate.pred"}, [T])
-
-fun dest_predT (Type (@{type_name "Predicate.pred"}, [T])) = T
- | dest_predT T = raise TYPE ("dest_predT", [T], []);
-
-fun mk_bot T = Const (@{const_name Orderings.bot}, mk_predT T);
-
-fun mk_single t =
- let val T = fastype_of t
- in Const(@{const_name Predicate.single}, T --> mk_predT T) $ t end;
-
-fun mk_bind (x, f) =
- let val T as Type ("fun", [_, U]) = fastype_of f
- in
- Const (@{const_name Predicate.bind}, fastype_of x --> T --> U) $ x $ f
- end;
-
-val mk_sup = HOLogic.mk_binop @{const_name sup};
-
-fun mk_if cond = Const (@{const_name Predicate.if_pred},
- HOLogic.boolT --> mk_predT HOLogic.unitT) $ cond;
-
-fun mk_not t = let val T = mk_predT HOLogic.unitT
- in Const (@{const_name Predicate.not_pred}, T --> T) $ t end
-
-fun mk_Enum f =
- let val T as Type ("fun", [T', _]) = fastype_of f
- in
- Const (@{const_name Predicate.Pred}, T --> mk_predT T') $ f
- end;
-
-fun mk_Eval (f, x) =
- let
- val T = fastype_of x
- in
- Const (@{const_name Predicate.eval}, mk_predT T --> T --> HOLogic.boolT) $ f $ x
- end;
-
-fun mk_map T1 T2 tf tp = Const (@{const_name Predicate.map},
- (T1 --> T2) --> mk_predT T1 --> mk_predT T2) $ tf $ tp;
-
-val lift_pred = I
-
-val compfuns = CompilationFuns {mk_predT = mk_predT, dest_predT = dest_predT, mk_bot = mk_bot,
- mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not,
- mk_map = mk_map, lift_pred = lift_pred};
-
-end;
-
-(* termify_code:
-val termT = Type ("Code_Evaluation.term", []);
-fun termifyT T = HOLogic.mk_prodT (T, HOLogic.unitT --> termT)
-*)
-(*
-fun lift_random random =
- let
- val T = dest_randomT (fastype_of random)
- in
- mk_scomp (random,
- mk_fun_comp (HOLogic.pair_const (PredicateCompFuns.mk_predT T) @{typ Random.seed},
- mk_fun_comp (Const (@{const_name Predicate.single}, T --> (PredicateCompFuns.mk_predT T)),
- Const (@{const_name "fst"}, HOLogic.mk_prodT (T, @{typ "unit => term"}) --> T))))
- end;
-*)
-
-structure RPredCompFuns =
-struct
-
-fun mk_rpredT T =
- @{typ "Random.seed"} --> HOLogic.mk_prodT (PredicateCompFuns.mk_predT T, @{typ "Random.seed"})
-
-fun dest_rpredT (Type ("fun", [_,
- Type (@{type_name "*"}, [Type (@{type_name "Predicate.pred"}, [T]), _])])) = T
- | dest_rpredT T = raise TYPE ("dest_rpredT", [T], []);
-
-fun mk_bot T = Const(@{const_name RPred.bot}, mk_rpredT T)
-
-fun mk_single t =
- let
- val T = fastype_of t
- in
- Const (@{const_name RPred.return}, T --> mk_rpredT T) $ t
- end;
-
-fun mk_bind (x, f) =
- let
- val T as (Type ("fun", [_, U])) = fastype_of f
- in
- Const (@{const_name RPred.bind}, fastype_of x --> T --> U) $ x $ f
- end
-
-val mk_sup = HOLogic.mk_binop @{const_name RPred.supp}
-
-fun mk_if cond = Const (@{const_name RPred.if_rpred},
- HOLogic.boolT --> mk_rpredT HOLogic.unitT) $ cond;
-
-fun mk_not t = error "Negation is not defined for RPred"
-
-fun mk_map t = error "FIXME" (*FIXME*)
-
-fun lift_pred t =
- let
- val T = PredicateCompFuns.dest_predT (fastype_of t)
- val lift_predT = PredicateCompFuns.mk_predT T --> mk_rpredT T
- in
- Const (@{const_name "RPred.lift_pred"}, lift_predT) $ t
- end;
-
-val compfuns = CompilationFuns {mk_predT = mk_rpredT, dest_predT = dest_rpredT, mk_bot = mk_bot,
- mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not,
- mk_map = mk_map, lift_pred = lift_pred};
-
-end;
-(* for external use with interactive mode *)
-val rpred_compfuns = RPredCompFuns.compfuns;
-
-fun lift_random random =
- let
- val T = dest_randomT (fastype_of random)
- in
- Const (@{const_name lift_random}, (@{typ Random.seed} -->
- HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) -->
- RPredCompFuns.mk_rpredT T) $ random
- end;
-
-(* Mode analysis *)
-
-(*** check if a term contains only constructor functions ***)
-fun is_constrt thy =
- let
- val cnstrs = flat (maps
- (map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd)
- (Symtab.dest (Datatype.get_all thy)));
- fun check t = (case strip_comb t of
- (Free _, []) => true
- | (Const (s, T), ts) => (case (AList.lookup (op =) cnstrs s, body_type T) of
- (SOME (i, Tname), Type (Tname', _)) => length ts = i andalso Tname = Tname' andalso forall check ts
- | _ => false)
- | _ => false)
- in check end;
-
-(*** check if a type is an equality type (i.e. doesn't contain fun)
- FIXME this is only an approximation ***)
-fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts
- | is_eqT _ = true;
-
-fun term_vs tm = fold_aterms (fn Free (x, T) => cons x | _ => I) tm [];
-val terms_vs = distinct (op =) o maps term_vs;
-
-(** collect all Frees in a term (with duplicates!) **)
-fun term_vTs tm =
- fold_aterms (fn Free xT => cons xT | _ => I) tm [];
-
-(*FIXME this function should not be named merge... make it local instead*)
-fun merge xs [] = xs
- | merge [] ys = ys
- | merge (x::xs) (y::ys) = if length x >= length y then x::merge xs (y::ys)
- else y::merge (x::xs) ys;
-
-fun subsets i j = if i <= j then
- let val is = subsets (i+1) j
- in merge (map (fn ks => i::ks) is) is end
- else [[]];
-
-(* FIXME: should be in library - map_prod *)
-fun cprod ([], ys) = []
- | cprod (x :: xs, ys) = map (pair x) ys @ cprod (xs, ys);
-
-fun cprods xss = List.foldr (map op :: o cprod) [[]] xss;
-
-
-
-(*TODO: cleanup function and put together with modes_of_term *)
-(*
-fun modes_of_param default modes t = let
- val (vs, t') = strip_abs t
- val b = length vs
- fun mk_modes name args = Option.map (maps (fn (m as (iss, is)) =>
- let
- val (args1, args2) =
- if length args < length iss then
- error ("Too few arguments for inductive predicate " ^ name)
- else chop (length iss) args;
- val k = length args2;
- val perm = map (fn i => (find_index_eq (Bound (b - i)) args2) + 1)
- (1 upto b)
- val partial_mode = (1 upto k) \\ perm
- in
- if not (partial_mode subset is) then [] else
- let
- val is' =
- (fold_index (fn (i, j) => if j mem is then cons (i + 1) else I) perm [])
- |> fold (fn i => if i > k then cons (i - k + b) else I) is
-
- val res = map (fn x => Mode (m, is', x)) (cprods (map
- (fn (NONE, _) => [NONE]
- | (SOME js, arg) => map SOME (filter
- (fn Mode (_, js', _) => js=js') (modes_of_term modes arg)))
- (iss ~~ args1)))
- in res end
- end)) (AList.lookup op = modes name)
- in case strip_comb t' of
- (Const (name, _), args) => the_default default (mk_modes name args)
- | (Var ((name, _), _), args) => the (mk_modes name args)
- | (Free (name, _), args) => the (mk_modes name args)
- | _ => default end
-
-and
-*)
-fun modes_of_term modes t =
- let
- val ks = 1 upto length (binder_types (fastype_of t));
- val default = [Mode (([], ks), ks, [])];
- fun mk_modes name args = Option.map (maps (fn (m as (iss, is)) =>
- let
- val (args1, args2) =
- if length args < length iss then
- error ("Too few arguments for inductive predicate " ^ name)
- else chop (length iss) args;
- val k = length args2;
- val prfx = 1 upto k
- in
- if not (is_prefix op = prfx is) then [] else
- let val is' = map (fn i => i - k) (List.drop (is, k))
- in map (fn x => Mode (m, is', x)) (cprods (map
- (fn (NONE, _) => [NONE]
- | (SOME js, arg) => map SOME (filter
- (fn Mode (_, js', _) => js=js') (modes_of_term modes arg)))
- (iss ~~ args1)))
- end
- end)) (AList.lookup op = modes name)
-
- in
- case strip_comb (Envir.eta_contract t) of
- (Const (name, _), args) => the_default default (mk_modes name args)
- | (Var ((name, _), _), args) => the (mk_modes name args)
- | (Free (name, _), args) => the (mk_modes name args)
- | (Abs _, []) => error "Abs at param position" (* modes_of_param default modes t *)
- | _ => default
- end
-
-fun select_mode_prem thy modes vs ps =
- find_first (is_some o snd) (ps ~~ map
- (fn Prem (us, t) => find_first (fn Mode (_, is, _) =>
- let
- val (in_ts, out_ts) = split_smode is us;
- val (out_ts', in_ts') = List.partition (is_constrt thy) out_ts;
- val vTs = maps term_vTs out_ts';
- val dupTs = map snd (duplicates (op =) vTs) @
- map_filter (AList.lookup (op =) vTs) vs;
- in
- subset (op =) (terms_vs (in_ts @ in_ts'), vs) andalso
- forall (is_eqT o fastype_of) in_ts' andalso
- subset (op =) (term_vs t, vs) andalso
- forall is_eqT dupTs
- end)
- (modes_of_term modes t handle Option =>
- error ("Bad predicate: " ^ Syntax.string_of_term_global thy t))
- | Negprem (us, t) => find_first (fn Mode (_, is, _) =>
- length us = length is andalso
- subset (op =) (terms_vs us, vs) andalso
- subset (op =) (term_vs t, vs)
- (modes_of_term modes t handle Option =>
- error ("Bad predicate: " ^ Syntax.string_of_term_global thy t))
- | Sidecond t => if subset (op =) (term_vs t, vs) then SOME (Mode (([], []), [], []))
- else NONE
- ) ps);
-
-fun fold_prem f (Prem (args, _)) = fold f args
- | fold_prem f (Negprem (args, _)) = fold f args
- | fold_prem f (Sidecond t) = f t
-
-fun all_subsets [] = [[]]
- | all_subsets (x::xs) = let val xss' = all_subsets xs in xss' @ (map (cons x) xss') end
-
-fun generator vTs v =
- let
- val T = the (AList.lookup (op =) vTs v)
- in
- (Generator (v, T), Mode (([], []), [], []))
- end;
-
-fun gen_prem (Prem (us, t)) = GeneratorPrem (us, t)
- | gen_prem _ = error "gen_prem : invalid input for gen_prem"
-
-fun param_gen_prem param_vs (p as Prem (us, t as Free (v, _))) =
- if member (op =) param_vs v then
- GeneratorPrem (us, t)
- else p
- | param_gen_prem param_vs p = p
-
-fun check_mode_clause with_generator thy param_vs modes gen_modes (iss, is) (ts, ps) =
- let
- val modes' = modes @ map_filter
- (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
- (param_vs ~~ iss);
- val gen_modes' = gen_modes @ map_filter
- (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
- (param_vs ~~ iss);
- val vTs = distinct (op =) ((fold o fold_prem) Term.add_frees ps (fold Term.add_frees ts []))
- val prem_vs = distinct (op =) ((fold o fold_prem) Term.add_free_names ps [])
- fun check_mode_prems acc_ps vs [] = SOME (acc_ps, vs)
- | check_mode_prems acc_ps vs ps = (case select_mode_prem thy modes' vs ps of
- NONE =>
- (if with_generator then
- (case select_mode_prem thy gen_modes' vs ps of
- SOME (p, SOME mode) => check_mode_prems ((gen_prem p, mode) :: acc_ps)
- (case p of Prem (us, _) => union (op =) vs (terms_vs us) | _ => vs)
- (filter_out (equal p) ps)
- | NONE =>
- let
- val all_generator_vs = all_subsets (prem_vs \\ vs) |> sort (int_ord o (pairself length))
- in
- case (find_first (fn generator_vs => is_some
- (select_mode_prem thy modes' (union (op =) vs generator_vs) ps)) all_generator_vs) of
- SOME generator_vs => check_mode_prems ((map (generator vTs) generator_vs) @ acc_ps)
- (union (op =) vs generator_vs) ps
- | NONE => NONE
- end)
- else
- NONE)
- | SOME (p, SOME mode) => check_mode_prems ((if with_generator then param_gen_prem param_vs p else p, mode) :: acc_ps)
- (case p of Prem (us, _) => union (op =) vs (terms_vs us) | _ => vs)
- (filter_out (equal p) ps))
- val (in_ts, in_ts') = List.partition (is_constrt thy) (fst (split_smode is ts));
- val in_vs = terms_vs in_ts;
- val concl_vs = terms_vs ts
- in
- if forall is_eqT (map snd (duplicates (op =) (maps term_vTs in_ts))) andalso
- forall (is_eqT o fastype_of) in_ts' then
- case check_mode_prems [] (union (op =) param_vs in_vs) ps of
- NONE => NONE
- | SOME (acc_ps, vs) =>
- if with_generator then
- SOME (ts, (rev acc_ps) @ (map (generator vTs) (concl_vs \\ vs)))
- else
- if subset (op =) (concl_vs, vs) then SOME (ts, rev acc_ps) else NONE
- else NONE
- end;
-
-fun check_modes_pred with_generator thy param_vs preds modes gen_modes (p, ms) =
- let val SOME rs = AList.lookup (op =) preds p
- in (p, List.filter (fn m => case find_index
- (is_none o check_mode_clause with_generator thy param_vs modes gen_modes m) rs of
- ~1 => true
- | i => (tracing ("Clause " ^ string_of_int (i + 1) ^ " of " ^
- p ^ " violates mode " ^ string_of_mode m); false)) ms)
- end;
-
-fun get_modes_pred with_generator thy param_vs preds modes gen_modes (p, ms) =
- let
- val SOME rs = AList.lookup (op =) preds p
- in
- (p, map (fn m =>
- (m, map (the o check_mode_clause with_generator thy param_vs modes gen_modes m) rs)) ms)
- end;
-
-fun fixp f (x : (string * mode list) list) =
- let val y = f x
- in if x = y then x else fixp f y end;
-
-fun modes_of_arities arities =
- (map (fn (s, (ks, k)) => (s, cprod (cprods (map
- (fn NONE => [NONE]
- | SOME k' => map SOME (subsets 1 k')) ks),
- subsets 1 k))) arities)
-
-fun infer_modes with_generator thy extra_modes arities param_vs preds =
- let
- val modes =
- fixp (fn modes =>
- map (check_modes_pred with_generator thy param_vs preds (modes @ extra_modes) []) modes)
- (modes_of_arities arities)
- in
- map (get_modes_pred with_generator thy param_vs preds (modes @ extra_modes) []) modes
- end;
-
-fun remove_from rem [] = []
- | remove_from rem ((k, vs) :: xs) =
- (case AList.lookup (op =) rem k of
- NONE => (k, vs)
- | SOME vs' => (k, vs \\ vs'))
- :: remove_from rem xs
-
-fun infer_modes_with_generator thy extra_modes arities param_vs preds =
- let
- val prednames = map fst preds
- val extra_modes = all_modes_of thy
- val gen_modes = all_generator_modes_of thy
- |> filter_out (fn (name, _) => member (op =) prednames name)
- val starting_modes = remove_from extra_modes (modes_of_arities arities)
- val modes =
- fixp (fn modes =>
- map (check_modes_pred true thy param_vs preds extra_modes (gen_modes @ modes)) modes)
- starting_modes
- in
- map (get_modes_pred true thy param_vs preds extra_modes (gen_modes @ modes)) modes
- end;
-
-(* term construction *)
-
-fun mk_v (names, vs) s T = (case AList.lookup (op =) vs s of
- NONE => (Free (s, T), (names, (s, [])::vs))
- | SOME xs =>
- let
- val s' = Name.variant names s;
- val v = Free (s', T)
- in
- (v, (s'::names, AList.update (op =) (s, v::xs) vs))
- end);
-
-fun distinct_v (Free (s, T)) nvs = mk_v nvs s T
- | distinct_v (t $ u) nvs =
- let
- val (t', nvs') = distinct_v t nvs;
- val (u', nvs'') = distinct_v u nvs';
- in (t' $ u', nvs'') end
- | distinct_v x nvs = (x, nvs);
-
-fun compile_match thy compfuns eqs eqs' out_ts success_t =
- let
- val eqs'' = maps mk_eq eqs @ eqs'
- val names = fold Term.add_free_names (success_t :: eqs'' @ out_ts) [];
- val name = Name.variant names "x";
- val name' = Name.variant (name :: names) "y";
- val T = mk_tupleT (map fastype_of out_ts);
- val U = fastype_of success_t;
- val U' = dest_predT compfuns U;
- val v = Free (name, T);
- val v' = Free (name', T);
- in
- lambda v (fst (Datatype.make_case
- (ProofContext.init thy) false [] v
- [(mk_tuple out_ts,
- if null eqs'' then success_t
- else Const (@{const_name HOL.If}, HOLogic.boolT --> U --> U --> U) $
- foldr1 HOLogic.mk_conj eqs'' $ success_t $
- mk_bot compfuns U'),
- (v', mk_bot compfuns U')]))
- end;
-
-(*FIXME function can be removed*)
-fun mk_funcomp f t =
- let
- val names = Term.add_free_names t [];
- val Ts = binder_types (fastype_of t);
- val vs = map Free
- (Name.variant_list names (replicate (length Ts) "x") ~~ Ts)
- in
- fold_rev lambda vs (f (list_comb (t, vs)))
- end;
-(*
-fun compile_param_ext thy compfuns modes (NONE, t) = t
- | compile_param_ext thy compfuns modes (m as SOME (Mode ((iss, is'), is, ms)), t) =
- let
- val (vs, u) = strip_abs t
- val (ivs, ovs) = split_mode is vs
- val (f, args) = strip_comb u
- val (params, args') = chop (length ms) args
- val (inargs, outargs) = split_mode is' args'
- val b = length vs
- val perm = map (fn i => (find_index_eq (Bound (b - i)) args') + 1) (1 upto b)
- val outp_perm =
- snd (split_mode is perm)
- |> map (fn i => i - length (filter (fn x => x < i) is'))
- val names = [] -- TODO
- val out_names = Name.variant_list names (replicate (length outargs) "x")
- val f' = case f of
- Const (name, T) =>
- if AList.defined op = modes name then
- mk_predfun_of thy compfuns (name, T) (iss, is')
- else error "compile param: Not an inductive predicate with correct mode"
- | Free (name, T) => Free (name, param_funT_of compfuns T (SOME is'))
- val outTs = dest_tupleT (dest_predT compfuns (body_type (fastype_of f')))
- val out_vs = map Free (out_names ~~ outTs)
- val params' = map (compile_param thy modes) (ms ~~ params)
- val f_app = list_comb (f', params' @ inargs)
- val single_t = (mk_single compfuns (mk_tuple (map (fn i => nth out_vs (i - 1)) outp_perm)))
- val match_t = compile_match thy compfuns [] [] out_vs single_t
- in list_abs (ivs,
- mk_bind compfuns (f_app, match_t))
- end
- | compile_param_ext _ _ _ _ = error "compile params"
-*)
-
-fun compile_param size thy compfuns (NONE, t) = t
- | compile_param size thy compfuns (m as SOME (Mode ((iss, is'), is, ms)), t) =
- let
- val (f, args) = strip_comb (Envir.eta_contract t)
- val (params, args') = chop (length ms) args
- val params' = map (compile_param size thy compfuns) (ms ~~ params)
- val mk_fun_of = case size of NONE => mk_fun_of | SOME _ => mk_sizelim_fun_of
- val funT_of = case size of NONE => funT_of | SOME _ => sizelim_funT_of
- val f' =
- case f of
- Const (name, T) =>
- mk_fun_of compfuns thy (name, T) (iss, is')
- | Free (name, T) => Free (name, funT_of compfuns (iss, is') T)
- | _ => error ("PredicateCompiler: illegal parameter term")
- in list_comb (f', params' @ args') end
-
-fun compile_expr size thy ((Mode (mode, is, ms)), t) =
- case strip_comb t of
- (Const (name, T), params) =>
- let
- val params' = map (compile_param size thy PredicateCompFuns.compfuns) (ms ~~ params)
- val mk_fun_of = case size of NONE => mk_fun_of | SOME _ => mk_sizelim_fun_of
- in
- list_comb (mk_fun_of PredicateCompFuns.compfuns thy (name, T) mode, params')
- end
- | (Free (name, T), args) =>
- let
- val funT_of = case size of NONE => funT_of | SOME _ => sizelim_funT_of
- in
- list_comb (Free (name, funT_of PredicateCompFuns.compfuns ([], is) T), args)
- end;
-
-fun compile_gen_expr size thy compfuns ((Mode (mode, is, ms)), t) =
- case strip_comb t of
- (Const (name, T), params) =>
- let
- val params' = map (compile_param size thy compfuns) (ms ~~ params)
- in
- list_comb (mk_generator_of compfuns thy (name, T) mode, params')
- end
- | (Free (name, T), args) =>
- list_comb (Free (name, sizelim_funT_of RPredCompFuns.compfuns ([], is) T), args)
-
-(** specific rpred functions -- move them to the correct place in this file *)
-
-(* uncommented termify code; causes more trouble than expected at first *)
-(*
-fun mk_valtermify_term (t as Const (c, T)) = HOLogic.mk_prod (t, Abs ("u", HOLogic.unitT, HOLogic.reflect_term t))
- | mk_valtermify_term (Free (x, T)) = Free (x, termifyT T)
- | mk_valtermify_term (t1 $ t2) =
- let
- val T = fastype_of t1
- val (T1, T2) = dest_funT T
- val t1' = mk_valtermify_term t1
- val t2' = mk_valtermify_term t2
- in
- Const ("Code_Evaluation.valapp", termifyT T --> termifyT T1 --> termifyT T2) $ t1' $ t2'
- end
- | mk_valtermify_term _ = error "Not a valid term for mk_valtermify_term"
-*)
-
-fun compile_clause compfuns size final_term thy all_vs param_vs (iss, is) inp (ts, moded_ps) =
- let
- fun check_constrt t (names, eqs) =
- if is_constrt thy t then (t, (names, eqs)) else
- let
- val s = Name.variant names "x";
- val v = Free (s, fastype_of t)
- in (v, (s::names, HOLogic.mk_eq (v, t)::eqs)) end;
-
- val (in_ts, out_ts) = split_smode is ts;
- val (in_ts', (all_vs', eqs)) =
- fold_map check_constrt in_ts (all_vs, []);
-
- fun compile_prems out_ts' vs names [] =
- let
- val (out_ts'', (names', eqs')) =
- fold_map check_constrt out_ts' (names, []);
- val (out_ts''', (names'', constr_vs)) = fold_map distinct_v
- out_ts'' (names', map (rpair []) vs);
- in
- (* termify code:
- compile_match thy compfuns constr_vs (eqs @ eqs') out_ts'''
- (mk_single compfuns (mk_tuple (map mk_valtermify_term out_ts)))
- *)
- compile_match thy compfuns constr_vs (eqs @ eqs') out_ts'''
- (final_term out_ts)
- end
- | compile_prems out_ts vs names ((p, mode as Mode ((_, is), _, _)) :: ps) =
- let
- val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
- val (out_ts', (names', eqs)) =
- fold_map check_constrt out_ts (names, [])
- val (out_ts'', (names'', constr_vs')) = fold_map distinct_v
- out_ts' ((names', map (rpair []) vs))
- val (compiled_clause, rest) = case p of
- Prem (us, t) =>
- let
- val (in_ts, out_ts''') = split_smode is us;
- val args = case size of
- NONE => in_ts
- | SOME size_t => in_ts @ [size_t]
- val u = lift_pred compfuns
- (list_comb (compile_expr size thy (mode, t), args))
- val rest = compile_prems out_ts''' vs' names'' ps
- in
- (u, rest)
- end
- | Negprem (us, t) =>
- let
- val (in_ts, out_ts''') = split_smode is us
- val u = lift_pred compfuns
- (mk_not PredicateCompFuns.compfuns (list_comb (compile_expr NONE thy (mode, t), in_ts)))
- val rest = compile_prems out_ts''' vs' names'' ps
- in
- (u, rest)
- end
- | Sidecond t =>
- let
- val rest = compile_prems [] vs' names'' ps;
- in
- (mk_if compfuns t, rest)
- end
- | GeneratorPrem (us, t) =>
- let
- val (in_ts, out_ts''') = split_smode is us;
- val args = case size of
- NONE => in_ts
- | SOME size_t => in_ts @ [size_t]
- val u = list_comb (compile_gen_expr size thy compfuns (mode, t), args)
- val rest = compile_prems out_ts''' vs' names'' ps
- in
- (u, rest)
- end
- | Generator (v, T) =>
- let
- val u = lift_random (HOLogic.mk_random T @{term "1::code_numeral"})
- val rest = compile_prems [Free (v, T)] vs' names'' ps;
- in
- (u, rest)
- end
- in
- compile_match thy compfuns constr_vs' eqs out_ts''
- (mk_bind compfuns (compiled_clause, rest))
- end
- val prem_t = compile_prems in_ts' param_vs all_vs' moded_ps;
- in
- mk_bind compfuns (mk_single compfuns inp, prem_t)
- end
-
-fun compile_pred compfuns mk_fun_of use_size thy all_vs param_vs s T mode moded_cls =
- let
- val (Ts1, (Us1, Us2)) = split_mode mode (binder_types T)
- val funT_of = if use_size then sizelim_funT_of else funT_of
- val Ts1' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) (fst mode) Ts1
- val xnames = Name.variant_list (all_vs @ param_vs)
- (map (fn i => "x" ^ string_of_int i) (snd mode));
- val size_name = Name.variant (all_vs @ param_vs @ xnames) "size"
- (* termify code: val xs = map2 (fn s => fn T => Free (s, termifyT T)) xnames Us1; *)
- val xs = map2 (fn s => fn T => Free (s, T)) xnames Us1;
- val params = map2 (fn s => fn T => Free (s, T)) param_vs Ts1'
- val size = Free (size_name, @{typ "code_numeral"})
- val decr_size =
- if use_size then
- SOME (Const ("HOL.minus_class.minus", @{typ "code_numeral => code_numeral => code_numeral"})
- $ size $ Const ("HOL.one_class.one", @{typ "Code_Numeral.code_numeral"}))
- else
- NONE
- val cl_ts =
- map (compile_clause compfuns decr_size (fn out_ts => mk_single compfuns (mk_tuple out_ts))
- thy all_vs param_vs mode (mk_tuple xs)) moded_cls;
- val t = foldr1 (mk_sup compfuns) cl_ts
- val T' = mk_predT compfuns (mk_tupleT Us2)
- val size_t = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T')
- $ HOLogic.mk_eq (size, @{term "0 :: code_numeral"})
- $ mk_bot compfuns (dest_predT compfuns T') $ t
- val fun_const = mk_fun_of compfuns thy (s, T) mode
- val eq = if use_size then
- (list_comb (fun_const, params @ xs @ [size]), size_t)
- else
- (list_comb (fun_const, params @ xs), t)
- in
- HOLogic.mk_Trueprop (HOLogic.mk_eq eq)
- end;
-
-(* special setup for simpset *)
-val HOL_basic_ss' = HOL_basic_ss setSolver
- (mk_solver "all_tac_solver" (fn _ => fn _ => all_tac))
-
-(* Definition of executable functions and their intro and elim rules *)
-
-fun print_arities arities = tracing ("Arities:\n" ^
- cat_lines (map (fn (s, (ks, k)) => s ^ ": " ^
- space_implode " -> " (map
- (fn NONE => "X" | SOME k' => string_of_int k')
- (ks @ [SOME k]))) arities));
-
-fun mk_Eval_of ((x, T), NONE) names = (x, names)
- | mk_Eval_of ((x, T), SOME mode) names = let
- val Ts = binder_types T
- val argnames = Name.variant_list names
- (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
- val args = map Free (argnames ~~ Ts)
- val (inargs, outargs) = split_smode mode args
- val r = PredicateCompFuns.mk_Eval (list_comb (x, inargs), mk_tuple outargs)
- val t = fold_rev lambda args r
-in
- (t, argnames @ names)
-end;
-
-fun create_intro_elim_rule (mode as (iss, is)) defthm mode_id funT pred thy =
-let
- val Ts = binder_types (fastype_of pred)
- val funtrm = Const (mode_id, funT)
- val argnames = Name.variant_list []
- (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
- val (Ts1, Ts2) = chop (length iss) Ts;
- val Ts1' = map2 (fn NONE => I | SOME is => funT_of (PredicateCompFuns.compfuns) ([], is)) iss Ts1
- val args = map Free (argnames ~~ (Ts1' @ Ts2))
- val (params, ioargs) = chop (length iss) args
- val (inargs, outargs) = split_smode is ioargs
- val param_names = Name.variant_list argnames
- (map (fn i => "p" ^ string_of_int i) (1 upto (length iss)))
- val param_vs = map Free (param_names ~~ Ts1)
- val (params', names) = fold_map mk_Eval_of ((params ~~ Ts1) ~~ iss) []
- val predpropI = HOLogic.mk_Trueprop (list_comb (pred, param_vs @ ioargs))
- val predpropE = HOLogic.mk_Trueprop (list_comb (pred, params' @ ioargs))
- val param_eqs = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (param_vs ~~ params')
- val funargs = params @ inargs
- val funpropE = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs),
- if null outargs then Free("y", HOLogic.unitT) else mk_tuple outargs))
- val funpropI = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs),
- mk_tuple outargs))
- val introtrm = Logic.list_implies (predpropI :: param_eqs, funpropI)
- val simprules = [defthm, @{thm eval_pred},
- @{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}]
- val unfolddef_tac = Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps simprules) 1
- val introthm = Goal.prove (ProofContext.init thy) (argnames @ param_names @ ["y"]) [] introtrm (fn {...} => unfolddef_tac)
- val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT));
- val elimtrm = Logic.list_implies ([funpropE, Logic.mk_implies (predpropE, P)], P)
- val elimthm = Goal.prove (ProofContext.init thy) (argnames @ param_names @ ["y", "P"]) [] elimtrm (fn {...} => unfolddef_tac)
-in
- (introthm, elimthm)
-end;
-
-fun create_constname_of_mode thy prefix name mode =
- let
- fun string_of_mode mode = if null mode then "0"
- else space_implode "_" (map string_of_int mode)
- val HOmode = space_implode "_and_"
- (fold (fn NONE => I | SOME mode => cons (string_of_mode mode)) (fst mode) [])
- in
- (Sign.full_bname thy (prefix ^ (Long_Name.base_name name))) ^
- (if HOmode = "" then "_" else "_for_" ^ HOmode ^ "_yields_") ^ (string_of_mode (snd mode))
- end;
-
-fun create_definitions preds (name, modes) thy =
- let
- val compfuns = PredicateCompFuns.compfuns
- val T = AList.lookup (op =) preds name |> the
- fun create_definition (mode as (iss, is)) thy = let
- val mode_cname = create_constname_of_mode thy "" name mode
- val mode_cbasename = Long_Name.base_name mode_cname
- val Ts = binder_types T
- val (Ts1, Ts2) = chop (length iss) Ts
- val (Us1, Us2) = split_smode is Ts2
- val Ts1' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) iss Ts1
- val funT = (Ts1' @ Us1) ---> (mk_predT compfuns (mk_tupleT Us2))
- val names = Name.variant_list []
- (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
- val xs = map Free (names ~~ (Ts1' @ Ts2));
- val (xparams, xargs) = chop (length iss) xs;
- val (xins, xouts) = split_smode is xargs
- val (xparams', names') = fold_map mk_Eval_of ((xparams ~~ Ts1) ~~ iss) names
- fun mk_split_lambda [] t = lambda (Free (Name.variant names' "x", HOLogic.unitT)) t
- | mk_split_lambda [x] t = lambda x t
- | mk_split_lambda xs t =
- let
- fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t))
- | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t))
- in
- mk_split_lambda' xs t
- end;
- val predterm = PredicateCompFuns.mk_Enum (mk_split_lambda xouts
- (list_comb (Const (name, T), xparams' @ xargs)))
- val lhs = list_comb (Const (mode_cname, funT), xparams @ xins)
- val def = Logic.mk_equals (lhs, predterm)
- val ([definition], thy') = thy |>
- Sign.add_consts_i [(Binding.name mode_cbasename, funT, NoSyn)] |>
- PureThy.add_defs false [((Binding.name (mode_cbasename ^ "_def"), def), [])]
- val (intro, elim) =
- create_intro_elim_rule mode definition mode_cname funT (Const (name, T)) thy'
- in thy' |> add_predfun name mode (mode_cname, definition, intro, elim)
- |> PureThy.store_thm (Binding.name (mode_cbasename ^ "I"), intro) |> snd
- |> PureThy.store_thm (Binding.name (mode_cbasename ^ "E"), elim) |> snd
- |> Theory.checkpoint
- end;
- in
- fold create_definition modes thy
- end;
-
-fun sizelim_create_definitions preds (name, modes) thy =
- let
- val T = AList.lookup (op =) preds name |> the
- fun create_definition mode thy =
- let
- val mode_cname = create_constname_of_mode thy "sizelim_" name mode
- val funT = sizelim_funT_of PredicateCompFuns.compfuns mode T
- in
- thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)]
- |> set_sizelim_function_name name mode mode_cname
- end;
- in
- fold create_definition modes thy
- end;
-
-fun rpred_create_definitions preds (name, modes) thy =
- let
- val T = AList.lookup (op =) preds name |> the
- fun create_definition mode thy =
- let
- val mode_cname = create_constname_of_mode thy "gen_" name mode
- val funT = sizelim_funT_of RPredCompFuns.compfuns mode T
- in
- thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)]
- |> set_generator_name name mode mode_cname
- end;
- in
- fold create_definition modes thy
- end;
-
-(* Proving equivalence of term *)
-
-fun is_Type (Type _) = true
- | is_Type _ = false
-
-(* returns true if t is an application of an datatype constructor *)
-(* which then consequently would be splitted *)
-(* else false *)
-fun is_constructor thy t =
- if (is_Type (fastype_of t)) then
- (case Datatype.get_info thy ((fst o dest_Type o fastype_of) t) of
- NONE => false
- | SOME info => (let
- val constr_consts = maps (fn (_, (_, _, constrs)) => map fst constrs) (#descr info)
- val (c, _) = strip_comb t
- in (case c of
- Const (name, _) => name mem_string constr_consts
- | _ => false) end))
- else false
-
-(* MAJOR FIXME: prove_params should be simple
- - different form of introrule for parameters ? *)
-fun prove_param thy (NONE, t) = TRY (rtac @{thm refl} 1)
- | prove_param thy (m as SOME (Mode (mode, is, ms)), t) =
- let
- val (f, args) = strip_comb (Envir.eta_contract t)
- val (params, _) = chop (length ms) args
- val f_tac = case f of
- Const (name, T) => simp_tac (HOL_basic_ss addsimps
- (@{thm eval_pred}::(predfun_definition_of thy name mode)::
- @{thm "Product_Type.split_conv"}::[])) 1
- | Free _ => TRY (rtac @{thm refl} 1)
- | Abs _ => error "prove_param: No valid parameter term"
- in
- REPEAT_DETERM (etac @{thm thin_rl} 1)
- THEN REPEAT_DETERM (rtac @{thm ext} 1)
- THEN print_tac "prove_param"
- THEN f_tac
- THEN print_tac "after simplification in prove_args"
- THEN (EVERY (map (prove_param thy) (ms ~~ params)))
- THEN (REPEAT_DETERM (atac 1))
- end
-
-fun prove_expr thy (Mode (mode, is, ms), t, us) (premposition : int) =
- case strip_comb t of
- (Const (name, T), args) =>
- let
- val introrule = predfun_intro_of thy name mode
- val (args1, args2) = chop (length ms) args
- in
- rtac @{thm bindI} 1
- THEN print_tac "before intro rule:"
- (* for the right assumption in first position *)
- THEN rotate_tac premposition 1
- THEN debug_tac (Display.string_of_thm (ProofContext.init thy) introrule)
- THEN rtac introrule 1
- THEN print_tac "after intro rule"
- (* work with parameter arguments *)
- THEN (atac 1)
- THEN (print_tac "parameter goal")
- THEN (EVERY (map (prove_param thy) (ms ~~ args1)))
- THEN (REPEAT_DETERM (atac 1))
- end
- | _ => rtac @{thm bindI} 1 THEN atac 1
-
-fun SOLVED tac st = FILTER (fn st' => nprems_of st' = nprems_of st - 1) tac st;
-
-fun SOLVEDALL tac st = FILTER (fn st' => nprems_of st' = 0) tac st
-
-fun prove_match thy (out_ts : term list) = let
- fun get_case_rewrite t =
- if (is_constructor thy t) then let
- val case_rewrites = (#case_rewrites (Datatype.the_info thy
- ((fst o dest_Type o fastype_of) t)))
- in case_rewrites @ maps get_case_rewrite (snd (strip_comb t)) end
- else []
- val simprules = @{thm "unit.cases"} :: @{thm "prod.cases"} :: maps get_case_rewrite out_ts
-(* replace TRY by determining if it necessary - are there equations when calling compile match? *)
-in
- (* make this simpset better! *)
- asm_simp_tac (HOL_basic_ss' addsimps simprules) 1
- THEN print_tac "after prove_match:"
- THEN (DETERM (TRY (EqSubst.eqsubst_tac (ProofContext.init thy) [0] [@{thm "HOL.if_P"}] 1
- THEN (REPEAT_DETERM (rtac @{thm conjI} 1 THEN (SOLVED (asm_simp_tac HOL_basic_ss 1))))
- THEN (SOLVED (asm_simp_tac HOL_basic_ss 1)))))
- THEN print_tac "after if simplification"
-end;
-
-(* corresponds to compile_fun -- maybe call that also compile_sidecond? *)
-
-fun prove_sidecond thy modes t =
- let
- fun preds_of t nameTs = case strip_comb t of
- (f as Const (name, T), args) =>
- if AList.defined (op =) modes name then (name, T) :: nameTs
- else fold preds_of args nameTs
- | _ => nameTs
- val preds = preds_of t []
- val defs = map
- (fn (pred, T) => predfun_definition_of thy pred ([], (1 upto (length (binder_types T)))))
- preds
- in
- (* remove not_False_eq_True when simpset in prove_match is better *)
- simp_tac (HOL_basic_ss addsimps @{thm not_False_eq_True} :: @{thm eval_pred} :: defs) 1
- (* need better control here! *)
- end
-
-fun prove_clause thy nargs modes (iss, is) (_, clauses) (ts, moded_ps) =
- let
- val (in_ts, clause_out_ts) = split_smode is ts;
- fun prove_prems out_ts [] =
- (prove_match thy out_ts)
- THEN asm_simp_tac HOL_basic_ss' 1
- THEN (rtac (if null clause_out_ts then @{thm singleI_unit} else @{thm singleI}) 1)
- | prove_prems out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) =
- let
- val premposition = (find_index (equal p) clauses) + nargs
- val rest_tac = (case p of Prem (us, t) =>
- let
- val (_, out_ts''') = split_smode is us
- val rec_tac = prove_prems out_ts''' ps
- in
- print_tac "before clause:"
- THEN asm_simp_tac HOL_basic_ss 1
- THEN print_tac "before prove_expr:"
- THEN prove_expr thy (mode, t, us) premposition
- THEN print_tac "after prove_expr:"
- THEN rec_tac
- end
- | Negprem (us, t) =>
- let
- val (_, out_ts''') = split_smode is us
- val rec_tac = prove_prems out_ts''' ps
- val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
- val (_, params) = strip_comb t
- in
- rtac @{thm bindI} 1
- THEN (if (is_some name) then
- simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, is)]) 1
- THEN rtac @{thm not_predI} 1
- THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
- THEN (REPEAT_DETERM (atac 1))
- (* FIXME: work with parameter arguments *)
- THEN (EVERY (map (prove_param thy) (param_modes ~~ params)))
- else
- rtac @{thm not_predI'} 1)
- THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
- THEN rec_tac
- end
- | Sidecond t =>
- rtac @{thm bindI} 1
- THEN rtac @{thm if_predI} 1
- THEN print_tac "before sidecond:"
- THEN prove_sidecond thy modes t
- THEN print_tac "after sidecond:"
- THEN prove_prems [] ps)
- in (prove_match thy out_ts)
- THEN rest_tac
- end;
- val prems_tac = prove_prems in_ts moded_ps
- in
- rtac @{thm bindI} 1
- THEN rtac @{thm singleI} 1
- THEN prems_tac
- end;
-
-fun select_sup 1 1 = []
- | select_sup _ 1 = [rtac @{thm supI1}]
- | select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1));
-
-fun prove_one_direction thy clauses preds modes pred mode moded_clauses =
- let
- val T = the (AList.lookup (op =) preds pred)
- val nargs = length (binder_types T) - nparams_of thy pred
- val pred_case_rule = the_elim_of thy pred
- in
- REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"}))
- THEN etac (predfun_elim_of thy pred mode) 1
- THEN etac pred_case_rule 1
- THEN (EVERY (map
- (fn i => EVERY' (select_sup (length moded_clauses) i) i)
- (1 upto (length moded_clauses))))
- THEN (EVERY (map2 (prove_clause thy nargs modes mode) clauses moded_clauses))
- THEN print_tac "proved one direction"
- end;
-
-(** Proof in the other direction **)
-
-fun prove_match2 thy out_ts = let
- fun split_term_tac (Free _) = all_tac
- | split_term_tac t =
- if (is_constructor thy t) then let
- val info = Datatype.the_info thy ((fst o dest_Type o fastype_of) t)
- val num_of_constrs = length (#case_rewrites info)
- (* special treatment of pairs -- because of fishing *)
- val split_rules = case (fst o dest_Type o fastype_of) t of
- "*" => [@{thm prod.split_asm}]
- | _ => PureThy.get_thms thy (((fst o dest_Type o fastype_of) t) ^ ".split_asm")
- val (_, ts) = strip_comb t
- in
- (Splitter.split_asm_tac split_rules 1)
-(* THEN (Simplifier.asm_full_simp_tac HOL_basic_ss 1)
- THEN (DETERM (TRY (etac @{thm Pair_inject} 1))) *)
- THEN (REPEAT_DETERM_N (num_of_constrs - 1) (etac @{thm botE} 1 ORELSE etac @{thm botE} 2))
- THEN (EVERY (map split_term_tac ts))
- end
- else all_tac
- in
- split_term_tac (mk_tuple out_ts)
- THEN (DETERM (TRY ((Splitter.split_asm_tac [@{thm "split_if_asm"}] 1) THEN (etac @{thm botE} 2))))
- end
-
-(* VERY LARGE SIMILIRATIY to function prove_param
--- join both functions
-*)
-(* TODO: remove function *)
-
-fun prove_param2 thy (NONE, t) = all_tac
- | prove_param2 thy (m as SOME (Mode (mode, is, ms)), t) = let
- val (f, args) = strip_comb (Envir.eta_contract t)
- val (params, _) = chop (length ms) args
- val f_tac = case f of
- Const (name, T) => full_simp_tac (HOL_basic_ss addsimps
- (@{thm eval_pred}::(predfun_definition_of thy name mode)
- :: @{thm "Product_Type.split_conv"}::[])) 1
- | Free _ => all_tac
- | _ => error "prove_param2: illegal parameter term"
- in
- print_tac "before simplification in prove_args:"
- THEN f_tac
- THEN print_tac "after simplification in prove_args"
- THEN (EVERY (map (prove_param2 thy) (ms ~~ params)))
- end
-
-
-fun prove_expr2 thy (Mode (mode, is, ms), t) =
- (case strip_comb t of
- (Const (name, T), args) =>
- etac @{thm bindE} 1
- THEN (REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"})))
- THEN print_tac "prove_expr2-before"
- THEN (debug_tac (Syntax.string_of_term_global thy
- (prop_of (predfun_elim_of thy name mode))))
- THEN (etac (predfun_elim_of thy name mode) 1)
- THEN print_tac "prove_expr2"
- THEN (EVERY (map (prove_param2 thy) (ms ~~ args)))
- THEN print_tac "finished prove_expr2"
- | _ => etac @{thm bindE} 1)
-
-(* FIXME: what is this for? *)
-(* replace defined by has_mode thy pred *)
-(* TODO: rewrite function *)
-fun prove_sidecond2 thy modes t = let
- fun preds_of t nameTs = case strip_comb t of
- (f as Const (name, T), args) =>
- if AList.defined (op =) modes name then (name, T) :: nameTs
- else fold preds_of args nameTs
- | _ => nameTs
- val preds = preds_of t []
- val defs = map
- (fn (pred, T) => predfun_definition_of thy pred ([], (1 upto (length (binder_types T)))))
- preds
- in
- (* only simplify the one assumption *)
- full_simp_tac (HOL_basic_ss' addsimps @{thm eval_pred} :: defs) 1
- (* need better control here! *)
- THEN print_tac "after sidecond2 simplification"
- end
-
-fun prove_clause2 thy modes pred (iss, is) (ts, ps) i =
- let
- val pred_intro_rule = nth (intros_of thy pred) (i - 1)
- val (in_ts, clause_out_ts) = split_smode is ts;
- fun prove_prems2 out_ts [] =
- print_tac "before prove_match2 - last call:"
- THEN prove_match2 thy out_ts
- THEN print_tac "after prove_match2 - last call:"
- THEN (etac @{thm singleE} 1)
- THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
- THEN (asm_full_simp_tac HOL_basic_ss' 1)
- THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
- THEN (asm_full_simp_tac HOL_basic_ss' 1)
- THEN SOLVED (print_tac "state before applying intro rule:"
- THEN (rtac pred_intro_rule 1)
- (* How to handle equality correctly? *)
- THEN (print_tac "state before assumption matching")
- THEN (REPEAT (atac 1 ORELSE
- (CHANGED (asm_full_simp_tac HOL_basic_ss' 1)
- THEN print_tac "state after simp_tac:"))))
- | prove_prems2 out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) =
- let
- val rest_tac = (case p of
- Prem (us, t) =>
- let
- val (_, out_ts''') = split_smode is us
- val rec_tac = prove_prems2 out_ts''' ps
- in
- (prove_expr2 thy (mode, t)) THEN rec_tac
- end
- | Negprem (us, t) =>
- let
- val (_, out_ts''') = split_smode is us
- val rec_tac = prove_prems2 out_ts''' ps
- val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
- val (_, params) = strip_comb t
- in
- print_tac "before neg prem 2"
- THEN etac @{thm bindE} 1
- THEN (if is_some name then
- full_simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, is)]) 1
- THEN etac @{thm not_predE} 1
- THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
- THEN (EVERY (map (prove_param2 thy) (param_modes ~~ params)))
- else
- etac @{thm not_predE'} 1)
- THEN rec_tac
- end
- | Sidecond t =>
- etac @{thm bindE} 1
- THEN etac @{thm if_predE} 1
- THEN prove_sidecond2 thy modes t
- THEN prove_prems2 [] ps)
- in print_tac "before prove_match2:"
- THEN prove_match2 thy out_ts
- THEN print_tac "after prove_match2:"
- THEN rest_tac
- end;
- val prems_tac = prove_prems2 in_ts ps
- in
- print_tac "starting prove_clause2"
- THEN etac @{thm bindE} 1
- THEN (etac @{thm singleE'} 1)
- THEN (TRY (etac @{thm Pair_inject} 1))
- THEN print_tac "after singleE':"
- THEN prems_tac
- end;
-
-fun prove_other_direction thy modes pred mode moded_clauses =
- let
- fun prove_clause clause i =
- (if i < length moded_clauses then etac @{thm supE} 1 else all_tac)
- THEN (prove_clause2 thy modes pred mode clause i)
- in
- (DETERM (TRY (rtac @{thm unit.induct} 1)))
- THEN (REPEAT_DETERM (CHANGED (rewtac @{thm split_paired_all})))
- THEN (rtac (predfun_intro_of thy pred mode) 1)
- THEN (REPEAT_DETERM (rtac @{thm refl} 2))
- THEN (EVERY (map2 prove_clause moded_clauses (1 upto (length moded_clauses))))
- end;
-
-(** proof procedure **)
-
-fun prove_pred thy clauses preds modes pred mode (moded_clauses, compiled_term) =
- let
- val ctxt = ProofContext.init thy
- val clauses = the (AList.lookup (op =) clauses pred)
- in
- Goal.prove ctxt (Term.add_free_names compiled_term []) [] compiled_term
- (if !do_proofs then
- (fn _ =>
- rtac @{thm pred_iffI} 1
- THEN prove_one_direction thy clauses preds modes pred mode moded_clauses
- THEN print_tac "proved one direction"
- THEN prove_other_direction thy modes pred mode moded_clauses
- THEN print_tac "proved other direction")
- else (fn _ => mycheat_tac thy 1))
- end;
-
-(* composition of mode inference, definition, compilation and proof *)
-
-(** auxillary combinators for table of preds and modes **)
-
-fun map_preds_modes f preds_modes_table =
- map (fn (pred, modes) =>
- (pred, map (fn (mode, value) => (mode, f pred mode value)) modes)) preds_modes_table
-
-fun join_preds_modes table1 table2 =
- map_preds_modes (fn pred => fn mode => fn value =>
- (value, the (AList.lookup (op =) (the (AList.lookup (op =) table2 pred)) mode))) table1
-
-fun maps_modes preds_modes_table =
- map (fn (pred, modes) =>
- (pred, map (fn (mode, value) => value) modes)) preds_modes_table
-
-fun compile_preds compfuns mk_fun_of use_size thy all_vs param_vs preds moded_clauses =
- map_preds_modes (fn pred => compile_pred compfuns mk_fun_of use_size thy all_vs param_vs pred
- (the (AList.lookup (op =) preds pred))) moded_clauses
-
-fun prove thy clauses preds modes moded_clauses compiled_terms =
- map_preds_modes (prove_pred thy clauses preds modes)
- (join_preds_modes moded_clauses compiled_terms)
-
-fun prove_by_skip thy _ _ _ _ compiled_terms =
- map_preds_modes (fn pred => fn mode => fn t => Drule.standard (Skip_Proof.make_thm thy t))
- compiled_terms
-
-fun prepare_intrs thy prednames =
- let
- val intrs = maps (intros_of thy) prednames
- |> map (Logic.unvarify o prop_of)
- val nparams = nparams_of thy (hd prednames)
- val extra_modes = all_modes_of thy |> filter_out (fn (name, _) => member (op =) prednames name)
- val preds = distinct (op =) (map (dest_Const o fst o (strip_intro_concl nparams)) intrs)
- val _ $ u = Logic.strip_imp_concl (hd intrs);
- val params = List.take (snd (strip_comb u), nparams);
- val param_vs = maps term_vs params
- val all_vs = terms_vs intrs
- fun dest_prem t =
- (case strip_comb t of
- (v as Free _, ts) => if v mem params then Prem (ts, v) else Sidecond t
- | (c as Const (@{const_name Not}, _), [t]) => (case dest_prem t of
- Prem (ts, t) => Negprem (ts, t)
- | Negprem _ => error ("Double negation not allowed in premise: " ^ (Syntax.string_of_term_global thy (c $ t)))
- | Sidecond t => Sidecond (c $ t))
- | (c as Const (s, _), ts) =>
- if is_registered thy s then
- let val (ts1, ts2) = chop (nparams_of thy s) ts
- in Prem (ts2, list_comb (c, ts1)) end
- else Sidecond t
- | _ => Sidecond t)
- fun add_clause intr (clauses, arities) =
- let
- val _ $ t = Logic.strip_imp_concl intr;
- val (Const (name, T), ts) = strip_comb t;
- val (ts1, ts2) = chop nparams ts;
- val prems = map (dest_prem o HOLogic.dest_Trueprop) (Logic.strip_imp_prems intr);
- val (Ts, Us) = chop nparams (binder_types T)
- in
- (AList.update op = (name, these (AList.lookup op = clauses name) @
- [(ts2, prems)]) clauses,
- AList.update op = (name, (map (fn U => (case strip_type U of
- (Rs as _ :: _, Type ("bool", [])) => SOME (length Rs)
- | _ => NONE)) Ts,
- length Us)) arities)
- end;
- val (clauses, arities) = fold add_clause intrs ([], []);
- in (preds, nparams, all_vs, param_vs, extra_modes, clauses, arities) end;
-
-(** main function of predicate compiler **)
-
-fun add_equations_of steps prednames thy =
- let
- val _ = tracing ("Starting predicate compiler for predicates " ^ commas prednames ^ "...")
- val (preds, nparams, all_vs, param_vs, extra_modes, clauses, arities) =
- prepare_intrs thy prednames
- val _ = tracing "Infering modes..."
- val moded_clauses = #infer_modes steps thy extra_modes arities param_vs clauses
- val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
- val _ = print_modes modes
- val _ = print_moded_clauses thy moded_clauses
- val _ = tracing "Defining executable functions..."
- val thy' = fold (#create_definitions steps preds) modes thy
- |> Theory.checkpoint
- val _ = tracing "Compiling equations..."
- val compiled_terms =
- (#compile_preds steps) thy' all_vs param_vs preds moded_clauses
- val _ = print_compiled_terms thy' compiled_terms
- val _ = tracing "Proving equations..."
- val result_thms = #prove steps thy' clauses preds (extra_modes @ modes)
- moded_clauses compiled_terms
- val qname = #qname steps
- (* val attrib = gn thy => Attrib.attribute_i thy Code.add_eqn_attrib *)
- val attrib = fn thy => Attrib.attribute_i thy (Attrib.internal (K (Thm.declaration_attribute
- (fn thm => Context.mapping (Code.add_eqn thm) I))))
- val thy'' = fold (fn (name, result_thms) => fn thy => snd (PureThy.add_thmss
- [((Binding.qualify true (Long_Name.base_name name) (Binding.name qname), result_thms),
- [attrib thy ])] thy))
- (maps_modes result_thms) thy'
- |> Theory.checkpoint
- in
- thy''
- end
-
-fun extend' value_of edges_of key (G, visited) =
- let
- val (G', v) = case try (Graph.get_node G) key of
- SOME v => (G, v)
- | NONE => (Graph.new_node (key, value_of key) G, value_of key)
- val (G'', visited') = fold (extend' value_of edges_of) (edges_of (key, v) \\ visited)
- (G', key :: visited)
- in
- (fold (Graph.add_edge o (pair key)) (edges_of (key, v)) G'', visited')
- end;
-
-fun extend value_of edges_of key G = fst (extend' value_of edges_of key (G, []))
-
-fun gen_add_equations steps names thy =
- let
- val thy' = PredData.map (fold (extend (fetch_pred_data thy) (depending_preds_of thy)) names) thy
- |> Theory.checkpoint;
- fun strong_conn_of gr keys =
- Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
- val scc = strong_conn_of (PredData.get thy') names
- val thy'' = fold_rev
- (fn preds => fn thy =>
- if #are_not_defined steps thy preds then add_equations_of steps preds thy else thy)
- scc thy' |> Theory.checkpoint
- in thy'' end
-
-(* different instantiantions of the predicate compiler *)
-
-val add_equations = gen_add_equations
- {infer_modes = infer_modes false,
- create_definitions = create_definitions,
- compile_preds = compile_preds PredicateCompFuns.compfuns mk_fun_of false,
- prove = prove,
- are_not_defined = (fn thy => forall (null o modes_of thy)),
- qname = "equation"}
-
-val add_sizelim_equations = gen_add_equations
- {infer_modes = infer_modes false,
- create_definitions = sizelim_create_definitions,
- compile_preds = compile_preds PredicateCompFuns.compfuns mk_sizelim_fun_of true,
- prove = prove_by_skip,
- are_not_defined = (fn thy => fn preds => true), (* TODO *)
- qname = "sizelim_equation"
- }
-
-val add_quickcheck_equations = gen_add_equations
- {infer_modes = infer_modes_with_generator,
- create_definitions = rpred_create_definitions,
- compile_preds = compile_preds RPredCompFuns.compfuns mk_generator_of true,
- prove = prove_by_skip,
- are_not_defined = (fn thy => fn preds => true), (* TODO *)
- qname = "rpred_equation"}
-
-(** user interface **)
-
-(* generation of case rules from user-given introduction rules *)
-
-fun mk_casesrule ctxt nparams introrules =
- let
- val intros = map (Logic.unvarify o prop_of) introrules
- val (pred, (params, args)) = strip_intro_concl nparams (hd intros)
- val ([propname], ctxt1) = Variable.variant_fixes ["thesis"] ctxt
- val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT))
- val (argnames, ctxt2) = Variable.variant_fixes
- (map (fn i => "a" ^ string_of_int i) (1 upto (length args))) ctxt1
- val argvs = map2 (curry Free) argnames (map fastype_of args)
- fun mk_case intro =
- let
- val (_, (_, args)) = strip_intro_concl nparams intro
- val prems = Logic.strip_imp_prems intro
- val eqprems = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (argvs ~~ args)
- val frees = (fold o fold_aterms)
- (fn t as Free _ =>
- if member (op aconv) params t then I else insert (op aconv) t
- | _ => I) (args @ prems) []
- in fold Logic.all frees (Logic.list_implies (eqprems @ prems, prop)) end
- val assm = HOLogic.mk_Trueprop (list_comb (pred, params @ argvs))
- val cases = map mk_case intros
- in Logic.list_implies (assm :: cases, prop) end;
-
-(* code_pred_intro attribute *)
-
-fun attrib f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
-
-val code_pred_intros_attrib = attrib add_intro;
-
-local
-
-(* TODO: make TheoryDataFun to GenericDataFun & remove duplication of local theory and theory *)
-(* TODO: must create state to prove multiple cases *)
-fun generic_code_pred prep_const raw_const lthy =
- let
- val thy = ProofContext.theory_of lthy
- val const = prep_const thy raw_const
- val lthy' = LocalTheory.theory (PredData.map
- (extend (fetch_pred_data thy) (depending_preds_of thy) const)) lthy
- |> LocalTheory.checkpoint
- val thy' = ProofContext.theory_of lthy'
- val preds = Graph.all_preds (PredData.get thy') [const] |> filter_out (has_elim thy')
- fun mk_cases const =
- let
- val nparams = nparams_of thy' const
- val intros = intros_of thy' const
- in mk_casesrule lthy' nparams intros end
- val cases_rules = map mk_cases preds
- val cases =
- map (fn case_rule => RuleCases.Case {fixes = [],
- assumes = [("", Logic.strip_imp_prems case_rule)],
- binds = [], cases = []}) cases_rules
- val case_env = map2 (fn p => fn c => (Long_Name.base_name p, SOME c)) preds cases
- val lthy'' = lthy'
- |> fold Variable.auto_fixes cases_rules
- |> ProofContext.add_cases true case_env
- fun after_qed thms goal_ctxt =
- let
- val global_thms = ProofContext.export goal_ctxt
- (ProofContext.init (ProofContext.theory_of goal_ctxt)) (map the_single thms)
- in
- goal_ctxt |> LocalTheory.theory (fold set_elim global_thms #> add_equations [const])
- end
- in
- Proof.theorem_i NONE after_qed (map (single o (rpair [])) cases_rules) lthy''
- end;
-
-structure P = OuterParse
-
-in
-
-val code_pred = generic_code_pred (K I);
-val code_pred_cmd = generic_code_pred Code.read_const
-
-val setup = PredData.put (Graph.empty) #>
- Attrib.setup @{binding code_pred_intros} (Scan.succeed (attrib add_intro))
- "adding alternative introduction rules for code generation of inductive predicates"
-(* Attrib.setup @{binding code_ind_cases} (Scan.succeed add_elim_attrib)
- "adding alternative elimination rules for code generation of inductive predicates";
- *)
- (*FIXME name discrepancy in attribs and ML code*)
- (*FIXME intros should be better named intro*)
- (*FIXME why distinguished attribute for cases?*)
-
-val _ = OuterSyntax.local_theory_to_proof "code_pred"
- "prove equations for predicate specified by intro/elim rules"
- OuterKeyword.thy_goal (P.term_group >> code_pred_cmd)
-
-end
-
-(*FIXME
-- Naming of auxiliary rules necessary?
-- add default code equations P x y z = P_i_i_i x y z
-*)
-
-(* transformation for code generation *)
-
-val eval_ref = Unsynchronized.ref (NONE : (unit -> term Predicate.pred) option);
-
-(*FIXME turn this into an LCF-guarded preprocessor for comprehensions*)
-fun analyze_compr thy t_compr =
- let
- val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t
- | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term_global thy t_compr);
- val (body, Ts, fp) = HOLogic.strip_psplits split;
- val (pred as Const (name, T), all_args) = strip_comb body;
- val (params, args) = chop (nparams_of thy name) all_args;
- val user_mode = map_filter I (map_index
- (fn (i, t) => case t of Bound j => if j < length Ts then NONE
- else SOME (i+1) | _ => SOME (i+1)) args); (*FIXME dangling bounds should not occur*)
- val modes = filter (fn Mode (_, is, _) => is = user_mode)
- (modes_of_term (all_modes_of thy) (list_comb (pred, params)));
- val m = case modes
- of [] => error ("No mode possible for comprehension "
- ^ Syntax.string_of_term_global thy t_compr)
- | [m] => m
- | m :: _ :: _ => (warning ("Multiple modes possible for comprehension "
- ^ Syntax.string_of_term_global thy t_compr); m);
- val (inargs, outargs) = split_smode user_mode args;
- val t_pred = list_comb (compile_expr NONE thy (m, list_comb (pred, params)), inargs);
- val t_eval = if null outargs then t_pred else let
- val outargs_bounds = map (fn Bound i => i) outargs;
- val outargsTs = map (nth Ts) outargs_bounds;
- val T_pred = HOLogic.mk_tupleT outargsTs;
- val T_compr = HOLogic.mk_ptupleT fp Ts;
- val arrange_bounds = map_index I outargs_bounds
- |> sort (prod_ord (K EQUAL) int_ord)
- |> map fst;
- val arrange = funpow (length outargs_bounds - 1) HOLogic.mk_split
- (Term.list_abs (map (pair "") outargsTs,
- HOLogic.mk_ptuple fp T_compr (map Bound arrange_bounds)))
- in mk_map PredicateCompFuns.compfuns T_pred T_compr arrange t_pred end
- in t_eval end;
-
-fun eval thy t_compr =
- let
- val t = analyze_compr thy t_compr;
- val T = dest_predT PredicateCompFuns.compfuns (fastype_of t);
- val t' = mk_map PredicateCompFuns.compfuns T HOLogic.termT (HOLogic.term_of_const T) t;
- in (T, Code_ML.eval NONE ("Predicate_Compile.eval_ref", eval_ref) Predicate.map thy t' []) end;
-
-fun values ctxt k t_compr =
- let
- val thy = ProofContext.theory_of ctxt;
- val (T, t) = eval thy t_compr;
- val setT = HOLogic.mk_setT T;
- val (ts, _) = Predicate.yieldn k t;
- val elemsT = HOLogic.mk_set T ts;
- in if k = ~1 orelse length ts < k then elemsT
- else Const (@{const_name Lattices.sup}, setT --> setT --> setT) $ elemsT $ t_compr
- end;
-
-fun values_cmd modes k raw_t state =
- let
- val ctxt = Toplevel.context_of state;
- val t = Syntax.read_term ctxt raw_t;
- val t' = values ctxt k t;
- val ty' = Term.type_of t';
- val ctxt' = Variable.auto_fixes t' ctxt;
- val p = PrintMode.with_modes modes (fn () =>
- Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk,
- Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) ();
- in Pretty.writeln p end;
-
-local structure P = OuterParse in
-
-val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) [];
-
-val _ = OuterSyntax.improper_command "values" "enumerate and print comprehensions" OuterKeyword.diag
- (opt_modes -- Scan.optional P.nat ~1 -- P.term
- >> (fn ((modes, k), t) => Toplevel.no_timing o Toplevel.keep
- (values_cmd modes k t)));
-
-end;
-
-end;
-
--- a/src/HOLCF/Tools/Domain/domain_axioms.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOLCF/Tools/Domain/domain_axioms.ML Thu Oct 29 18:53:58 2009 +0100
@@ -184,7 +184,7 @@
fun one_con (con,args) = let
val nonrec_args = filter_out is_rec args;
- val rec_args = List.filter is_rec args;
+ val rec_args = filter is_rec args;
val recs_cnt = length rec_args;
val allargs = nonrec_args @ rec_args
@ map (upd_vname (fn s=> s^"'")) rec_args;
--- a/src/HOLCF/Tools/Domain/domain_library.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOLCF/Tools/Domain/domain_library.ML Thu Oct 29 18:53:58 2009 +0100
@@ -241,8 +241,8 @@
val upd_vname = upd_third;
fun is_rec arg = rec_of arg >=0;
fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg);
-fun nonlazy args = map vname (filter_out is_lazy args);
-fun nonlazy_rec args = map vname (List.filter is_nonlazy_rec args);
+fun nonlazy args = map vname (filter_out is_lazy args);
+fun nonlazy_rec args = map vname (filter is_nonlazy_rec args);
(* ----- combinators for making dtyps ----- *)
--- a/src/HOLCF/Tools/Domain/domain_theorems.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/HOLCF/Tools/Domain/domain_theorems.ML Thu Oct 29 18:53:58 2009 +0100
@@ -446,7 +446,7 @@
val nlas = nonlazy args;
val vns = map vname args;
val vnn = List.nth (vns, n);
- val nlas' = List.filter (fn v => v <> vnn) nlas;
+ val nlas' = filter (fn v => v <> vnn) nlas;
val lhs = (%%:sel)`(con_app con args);
val goal = lift_defined %: (nlas', mk_trp (lhs === %:vnn));
fun tacs1 ctxt =
@@ -555,7 +555,7 @@
val sargs = case largs of [_] => [] | _ => nonlazy args;
val prop = lift_defined %: (sargs, mk_trp (prem === concl));
in pg con_appls prop end;
- val cons' = List.filter (fn (_,args) => args<>[]) cons;
+ val cons' = filter (fn (_,args) => args<>[]) cons;
in
val _ = trace " Proving inverts...";
val inverts =
@@ -593,7 +593,7 @@
else (%# arg);
val rhs = con_app2 con one_rhs args;
val goal = lift_defined %: (nonlazy_rec args, mk_trp (lhs === rhs));
- val args' = List.filter (fn a => not (is_rec a orelse is_lazy a)) args;
+ val args' = filter_out (fn a => is_rec a orelse is_lazy a) args;
val stricts = abs_strict :: rep_strict :: @{thms domain_fun_stricts};
fun tacs1 ctxt = map (case_UU_tac ctxt stricts 1 o vname) args';
val rules = [ax_abs_iso] @ @{thms domain_fun_simps};
@@ -616,7 +616,7 @@
fun has_nonlazy_rec (_, args) = exists is_nonlazy_rec args;
in
val _ = trace " Proving copy_stricts...";
- val copy_stricts = map one_strict (List.filter has_nonlazy_rec cons);
+ val copy_stricts = map one_strict (filter has_nonlazy_rec cons);
end;
val copy_rews = copy_strict :: copy_apps @ copy_stricts;
@@ -722,7 +722,7 @@
in Library.foldr mk_all (map vname args, lhs === rhs) end;
fun mk_eqns ((dn, _), cons) = map (mk_eqn dn) cons;
val goal = mk_trp (foldr1 mk_conj (maps mk_eqns eqs));
- val simps = List.filter (has_fewer_prems 1) copy_rews;
+ val simps = filter (has_fewer_prems 1) copy_rews;
fun con_tac ctxt (con, args) =
if nonlazy_rec args = []
then all_tac
@@ -747,7 +747,7 @@
let
fun ind_hyp arg = %:(P_name (1 + rec_of arg)) $ bound_arg args arg;
val t1 = mk_trp (%:p $ con_app2 con (bound_arg args) args);
- val t2 = lift ind_hyp (List.filter is_rec args, t1);
+ val t2 = lift ind_hyp (filter is_rec args, t1);
val t3 = lift_defined (bound_arg (map vname args)) (nonlazy args, t2);
in Library.foldr mk_All (map vname args, t3) end;
@@ -767,7 +767,7 @@
maps (fn (_,args) =>
resolve_tac prems 1 ::
map (K(atac 1)) (nonlazy args) @
- map (K(atac 1)) (List.filter is_rec args))
+ map (K(atac 1)) (filter is_rec args))
cons))
conss);
local
@@ -812,10 +812,10 @@
(List.nth (dnames, rec_of arg) ^ "_take n$" ^ vname arg);
fun con_tacs (con, args) =
asm_simp_tac take_ss 1 ::
- map arg_tac (List.filter is_nonlazy_rec args) @
+ map arg_tac (filter is_nonlazy_rec args) @
[resolve_tac prems 1] @
- map (K (atac 1)) (nonlazy args) @
- map (K (etac spec 1)) (List.filter is_rec args);
+ map (K (atac 1)) (nonlazy args) @
+ map (K (etac spec 1)) (filter is_rec args);
fun cases_tacs (cons, cases) =
res_inst_tac context [(("x", 0), "x")] cases 1 ::
asm_simp_tac (take_ss addsimps prems) 1 ::
--- a/src/Provers/Arith/fast_lin_arith.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Provers/Arith/fast_lin_arith.ML Thu Oct 29 18:53:58 2009 +0100
@@ -340,7 +340,7 @@
case ty of Eq => k <> 0 | Le => k > 0 | Lt => k >= 0;
fun calc_blowup l =
- let val (p,n) = List.partition (curry (op <) 0) (List.filter (curry (op <>) 0) l)
+ let val (p,n) = List.partition (curry (op <) 0) (filter (curry (op <>) 0) l)
in length p * length n end;
(* ------------------------------------------------------------------------- *)
--- a/src/Provers/classical.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Provers/classical.ML Thu Oct 29 18:53:58 2009 +0100
@@ -670,7 +670,7 @@
(*version of bimatch_from_nets_tac that only applies rules that
create precisely n subgoals.*)
fun n_bimatch_from_nets_tac n =
- biresolution_from_nets_tac (order_list o List.filter (nsubgoalsP n)) true;
+ biresolution_from_nets_tac (order_list o filter (nsubgoalsP n)) true;
fun eq_contr_tac i = ematch_tac [not_elim] i THEN eq_assume_tac i;
val eq_assume_contr_tac = eq_assume_tac ORELSE' eq_contr_tac;
--- a/src/Provers/splitter.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Provers/splitter.ML Thu Oct 29 18:53:58 2009 +0100
@@ -192,7 +192,7 @@
if n > length ts then []
else let val lev = length apsns
val lbnos = fold add_lbnos (Library.take (n, ts)) []
- val flbnos = List.filter (fn i => i < lev) lbnos
+ val flbnos = filter (fn i => i < lev) lbnos
val tt = incr_boundvars (~lev) t
in if null flbnos then
if T = T' then [(thm,[],pos,TB,tt)] else []
--- a/src/Pure/Isar/code.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Pure/Isar/code.ML Thu Oct 29 18:53:58 2009 +0100
@@ -639,7 +639,8 @@
(* datatypes *)
-structure Type_Interpretation = InterpretationFun(type T = string * serial val eq = eq_snd (op =) : T * T -> bool);
+structure Type_Interpretation =
+ Interpretation(type T = string * serial val eq = eq_snd (op =) : T * T -> bool);
fun add_datatype raw_cs thy =
let
--- a/src/Pure/Isar/expression.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Pure/Isar/expression.ML Thu Oct 29 18:53:58 2009 +0100
@@ -641,8 +641,8 @@
|> bodyT = propT ? Sign.add_advanced_trfuns ([], [], [aprop_tr' (length args) name], [])
|> Sign.declare_const ((bname, predT), NoSyn) |> snd
|> PureThy.add_defs false
- [((Binding.conceal (Binding.map_name Thm.def_name bname), Logic.mk_equals (head, body)),
- [Thm.kind_internal])];
+ [((Binding.conceal (Binding.map_name Thm.def_name bname),
+ Logic.mk_equals (head, body)), [])];
val defs_ctxt = ProofContext.init defs_thy |> Variable.declare_term head;
val cert = Thm.cterm_of defs_thy;
--- a/src/Pure/Isar/rule_cases.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Pure/Isar/rule_cases.ML Thu Oct 29 18:53:58 2009 +0100
@@ -37,6 +37,8 @@
val name: string list -> thm -> thm
val case_names: string list -> attribute
val case_conclusion: string * string list -> attribute
+ val is_inner_rule: thm -> bool
+ val inner_rule: attribute
val save: thm -> thm -> thm
val get: thm -> (string * string list) list * int
val rename_params: string list list -> thm -> thm
@@ -90,7 +92,7 @@
fun extract_case is_open thy (case_outline, raw_prop) name concls =
let
- val rename = if is_open then I else (apfst (Name.internal o Name.clean));
+ val rename = if is_open then I else apfst (Name.internal o Name.clean);
val props = Logic.dest_conjunctions (Drule.norm_hhf thy raw_prop);
val len = length props;
@@ -212,7 +214,7 @@
val consumes_tagN = "consumes";
fun lookup_consumes th =
- (case AList.lookup (op =) (Thm.get_tags th) (consumes_tagN) of
+ (case AList.lookup (op =) (Thm.get_tags th) consumes_tagN of
NONE => NONE
| SOME s =>
(case Lexicon.read_nat s of SOME n => SOME n
@@ -223,14 +225,13 @@
fun put_consumes NONE th = th
| put_consumes (SOME n) th = th
|> Thm.untag_rule consumes_tagN
- |> Thm.tag_rule
- (consumes_tagN, Library.string_of_int (if n < 0 then Thm.nprems_of th + n else n));
+ |> Thm.tag_rule (consumes_tagN, string_of_int (if n < 0 then Thm.nprems_of th + n else n));
fun add_consumes k th = put_consumes (SOME (k + get_consumes th)) th;
val save_consumes = put_consumes o lookup_consumes;
-fun consumes n x = Thm.rule_attribute (K (put_consumes (SOME n))) x;
+fun consumes n = Thm.rule_attribute (K (put_consumes (SOME n)));
fun consumes_default n x =
if is_some (lookup_consumes (#2 x)) then x else consumes n x;
@@ -282,7 +283,24 @@
else NONE)
in fold add_case_concl concls end;
-fun case_conclusion concl = Thm.rule_attribute (fn _ => add_case_concl concl);
+fun case_conclusion concl = Thm.rule_attribute (K (add_case_concl concl));
+
+
+
+(** inner rule **)
+
+val inner_rule_tagN = "inner_rule";
+
+fun is_inner_rule th =
+ AList.defined (op =) (Thm.get_tags th) inner_rule_tagN;
+
+fun put_inner_rule inner =
+ Thm.untag_rule inner_rule_tagN
+ #> inner ? Thm.tag_rule (inner_rule_tagN, "");
+
+val save_inner_rule = put_inner_rule o is_inner_rule;
+
+val inner_rule = Thm.rule_attribute (K (put_inner_rule true));
@@ -290,7 +308,11 @@
(* access hints *)
-fun save th = save_consumes th #> save_case_names th #> save_case_concls th;
+fun save th =
+ save_consumes th #>
+ save_case_names th #>
+ save_case_concls th #>
+ save_inner_rule th;
fun get th =
let
@@ -357,5 +379,5 @@
end;
-structure BasicRuleCases: BASIC_RULE_CASES = RuleCases;
-open BasicRuleCases;
+structure Basic_Rule_Cases: BASIC_RULE_CASES = RuleCases;
+open Basic_Rule_Cases;
--- a/src/Pure/Isar/rule_insts.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Pure/Isar/rule_insts.ML Thu Oct 29 18:53:58 2009 +0100
@@ -268,7 +268,7 @@
(* Separate type and term insts *)
fun has_type_var ((x, _), _) =
(case Symbol.explode x of "'" :: _ => true | _ => false);
- val Tinsts = List.filter has_type_var insts;
+ val Tinsts = filter has_type_var insts;
val tinsts = filter_out has_type_var insts;
(* Tactic *)
--- a/src/Pure/Proof/extraction.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Pure/Proof/extraction.ML Thu Oct 29 18:53:58 2009 +0100
@@ -651,7 +651,7 @@
val nt = Envir.beta_norm t;
val args = filter_out (fn v => member (op =) rtypes
(tname_of (body_type (fastype_of v)))) (vfs_of prop);
- val args' = List.filter (fn v => Logic.occs (v, nt)) args;
+ val args' = filter (fn v => Logic.occs (v, nt)) args;
val t' = mkabs nt args';
val T = fastype_of t';
val cname = extr_name s vs';
--- a/src/Pure/Syntax/parser.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Pure/Syntax/parser.ML Thu Oct 29 18:53:58 2009 +0100
@@ -615,11 +615,11 @@
(*Get all rhss with precedence >= minPrec*)
-fun getRHS minPrec = List.filter (fn (_, _, prec:int) => prec >= minPrec);
+fun getRHS minPrec = filter (fn (_, _, prec:int) => prec >= minPrec);
(*Get all rhss with precedence >= minPrec and < maxPrec*)
fun getRHS' minPrec maxPrec =
- List.filter (fn (_, _, prec:int) => prec >= minPrec andalso prec < maxPrec);
+ filter (fn (_, _, prec:int) => prec >= minPrec andalso prec < maxPrec);
(*Make states using a list of rhss*)
fun mkStates i minPrec lhsID rhss =
@@ -655,19 +655,19 @@
in update (used, []) end;
fun getS A maxPrec Si =
- List.filter
+ filter
(fn (_, _, _, Nonterminal (B, prec) :: _, _, _)
=> A = B andalso prec <= maxPrec
| _ => false) Si;
fun getS' A maxPrec minPrec Si =
- List.filter
+ filter
(fn (_, _, _, Nonterminal (B, prec) :: _, _, _)
=> A = B andalso prec > minPrec andalso prec <= maxPrec
| _ => false) Si;
fun getStates Estate i ii A maxPrec =
- List.filter
+ filter
(fn (_, _, _, Nonterminal (B, prec) :: _, _, _)
=> A = B andalso prec <= maxPrec
| _ => false)
--- a/src/Pure/Syntax/syn_ext.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Pure/Syntax/syn_ext.ML Thu Oct 29 18:53:58 2009 +0100
@@ -218,7 +218,7 @@
val read_symbs = map_filter I o the o Scan.read Symbol.stopper scan_symbs;
fun unique_index xsymbs =
- if length (List.filter is_index xsymbs) <= 1 then xsymbs
+ if length (filter is_index xsymbs) <= 1 then xsymbs
else error "Duplicate index arguments (\\<index>)";
in
@@ -226,7 +226,7 @@
val read_mfix = unique_index o read_symbs o Symbol.explode;
fun mfix_delims sy = fold_rev (fn Delim s => cons s | _ => I) (read_mfix sy) [];
-val mfix_args = length o List.filter is_argument o read_mfix;
+val mfix_args = length o filter is_argument o read_mfix;
val escape_mfix = implode o map (fn s => if is_meta s then "'" ^ s else s) o Symbol.explode;
@@ -276,7 +276,7 @@
val raw_symbs = read_mfix sy handle ERROR msg => err_in_mfix msg mfix;
- val args = List.filter (fn Argument _ => true | _ => false) raw_symbs;
+ val args = filter (fn Argument _ => true | _ => false) raw_symbs;
val (const', typ', parse_rules) =
if not (exists is_index args) then (const, typ, [])
else
@@ -312,7 +312,7 @@
val xprod' =
if Lexicon.is_terminal lhs' then err_in_mfix ("Illegal lhs: " ^ lhs') mfix
else if const <> "" then xprod
- else if length (List.filter is_argument symbs') <> 1 then
+ else if length (filter is_argument symbs') <> 1 then
err_in_mfix "Copy production must have exactly one argument" mfix
else if exists is_terminal symbs' then xprod
else XProd (lhs', map rem_pri symbs', "", chain_pri);
--- a/src/Pure/Tools/find_consts.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Pure/Tools/find_consts.ML Thu Oct 29 18:53:58 2009 +0100
@@ -11,11 +11,10 @@
Strict of string
| Loose of string
| Name of string
-
val find_consts : Proof.context -> (bool * criterion) list -> unit
end;
-structure FindConsts : FIND_CONSTS =
+structure Find_Consts : FIND_CONSTS =
struct
(* search criteria *)
@@ -162,7 +161,7 @@
val _ =
OuterSyntax.improper_command "find_consts" "search constants by type pattern" K.diag
- (Scan.repeat (((Scan.option P.minus >> is_none) -- criterion))
+ (Scan.repeat ((Scan.option P.minus >> is_none) -- criterion)
>> (Toplevel.no_timing oo find_consts_cmd));
end;
--- a/src/Pure/Tools/find_theorems.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Pure/Tools/find_theorems.ML Thu Oct 29 18:53:58 2009 +0100
@@ -18,7 +18,7 @@
(bool * string criterion) list -> unit
end;
-structure FindTheorems: FIND_THEOREMS =
+structure Find_Theorems: FIND_THEOREMS =
struct
(** search criteria **)
@@ -28,24 +28,22 @@
Pattern of 'term;
fun apply_dummies tm =
- strip_abs tm
- |> fst
- |> map (Term.dummy_pattern o snd)
- |> betapplys o pair tm
- |> (fn x => Term.replace_dummy_patterns x 1)
- |> fst;
+ let
+ val (xs, _) = Term.strip_abs tm;
+ val tm' = Term.betapplys (tm, map (Term.dummy_pattern o #2) xs);
+ in #1 (Term.replace_dummy_patterns tm' 1) end;
fun parse_pattern ctxt nm =
let
- val nm' = case Syntax.parse_term ctxt nm of Const (n, _) => n | _ => nm;
val consts = ProofContext.consts_of ctxt;
+ val nm' =
+ (case Syntax.parse_term ctxt nm of
+ Const (c, _) => c
+ | _ => Consts.intern consts nm);
in
- nm'
- |> Consts.intern consts
- |> Consts.the_abbreviation consts
- |> snd
- |> apply_dummies
- handle TYPE _ => ProofContext.read_term_pattern ctxt nm
+ (case try (Consts.the_abbreviation consts) nm' of
+ SOME (_, rhs) => apply_dummies rhs
+ | NONE => ProofContext.read_term_pattern ctxt nm)
end;
fun read_criterion _ (Name name) = Name name
--- a/src/Pure/Tools/named_thms.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Pure/Tools/named_thms.ML Thu Oct 29 18:53:58 2009 +0100
@@ -1,7 +1,8 @@
(* Title: Pure/Tools/named_thms.ML
Author: Makarius
-Named collections of theorems in canonical order.
+Named collections of theorems in canonical order. Based on naive data
+structures -- not scalable!
*)
signature NAMED_THMS =
--- a/src/Pure/axclass.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Pure/axclass.ML Thu Oct 29 18:53:58 2009 +0100
@@ -311,7 +311,7 @@
(Binding.name (Thm.def_name c'), Logic.mk_equals (Const (c, T'), const'))
#>> Thm.varifyT
#-> (fn thm => add_inst_param (c, tyco) (c'', thm)
- #> PureThy.add_thm ((Binding.conceal (Binding.name c'), thm), [Thm.kind_internal])
+ #> PureThy.add_thm ((Binding.conceal (Binding.name c'), thm), [])
#> snd
#> Sign.restore_naming thy
#> pair (Const (c, T))))
--- a/src/Pure/codegen.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Pure/codegen.ML Thu Oct 29 18:53:58 2009 +0100
@@ -137,7 +137,7 @@
| args_of (Ignore :: ms) (_ :: xs) = args_of ms xs
| args_of (_ :: ms) xs = args_of ms xs;
-fun num_args_of x = length (List.filter is_arg x);
+fun num_args_of x = length (filter is_arg x);
(**** theory data ****)
--- a/src/Pure/interpretation.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Pure/interpretation.ML Thu Oct 29 18:53:58 2009 +0100
@@ -13,7 +13,7 @@
val init: theory -> theory
end;
-functor InterpretationFun(type T val eq: T * T -> bool): INTERPRETATION =
+functor Interpretation(type T val eq: T * T -> bool): INTERPRETATION =
struct
type T = T;
--- a/src/Pure/more_thm.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Pure/more_thm.ML Thu Oct 29 18:53:58 2009 +0100
@@ -91,13 +91,9 @@
val lemmaK: string
val corollaryK: string
val internalK: string
- val has_kind: thm -> bool
val get_kind: thm -> string
val kind_rule: string -> thm -> thm
val kind: string -> attribute
- val kind_internal: attribute
- val has_internal: Properties.property list -> bool
- val is_internal: thm -> bool
end;
structure Thm: THM =
@@ -425,16 +421,10 @@
val corollaryK = "corollary";
val internalK = Markup.internalK;
-fun the_kind thm = the (Properties.get (Thm.get_tags thm) Markup.kindN);
-
-val has_kind = can the_kind;
-val get_kind = the_default "" o try the_kind;
+fun get_kind thm = the_default "" (Properties.get (Thm.get_tags thm) Markup.kindN);
fun kind_rule k = tag_rule (Markup.kindN, k) o untag_rule Markup.kindN;
fun kind k x = if k = "" then x else rule_attribute (K (kind_rule k)) x;
-fun kind_internal x = kind internalK x;
-fun has_internal tags = exists (fn tg => tg = (Markup.kindN, internalK)) tags;
-val is_internal = has_internal o Thm.get_tags;
open Thm;
--- a/src/Pure/proofterm.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Pure/proofterm.ML Thu Oct 29 18:53:58 2009 +0100
@@ -1024,7 +1024,7 @@
(** see pattern.ML **)
-fun flt (i: int) = List.filter (fn n => n < i);
+fun flt (i: int) = filter (fn n => n < i);
fun fomatch Ts tymatch j =
let
--- a/src/Pure/unify.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Pure/unify.ML Thu Oct 29 18:53:58 2009 +0100
@@ -473,7 +473,7 @@
if i<lev then Bound i
else if member (op =) banned (i-lev)
then raise CHANGE_FAIL (**flexible occurrence: give up**)
- else Bound (i - length (List.filter (fn j => j < i-lev) banned))
+ else Bound (i - length (filter (fn j => j < i-lev) banned))
| change lev (Abs (a,T,t)) = Abs (a, T, change(lev+1) t)
| change lev (t$u) = change lev t $ change lev u
| change lev t = t
--- a/src/Tools/IsaPlanner/rw_inst.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Tools/IsaPlanner/rw_inst.ML Thu Oct 29 18:53:58 2009 +0100
@@ -169,7 +169,7 @@
OldTerm.add_term_tfrees (t,tfrees)))
([],[]) ts;
val unfixed_tvars =
- List.filter (fn (ix,s) => not (member (op =) ignore_ixs ix)) tvars;
+ filter (fn (ix,s) => not (member (op =) ignore_ixs ix)) tvars;
val (fixtyinsts, _) = List.foldr new_tfree ([], map fst tfrees) unfixed_tvars
in (fixtyinsts, tfrees) end;
--- a/src/Tools/auto_solve.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Tools/auto_solve.ML Thu Oct 29 18:53:58 2009 +0100
@@ -5,7 +5,7 @@
existing theorem. Duplicate lemmas can be detected in this way.
The implementation is based in part on Berghofer and Haftmann's
-quickcheck.ML. It relies critically on the FindTheorems solves
+quickcheck.ML. It relies critically on the Find_Theorems solves
feature.
*)
@@ -45,8 +45,8 @@
let
val ctxt = Proof.context_of state;
- val crits = [(true, FindTheorems.Solves)];
- fun find g = snd (FindTheorems.find_theorems ctxt (SOME g) (SOME (! limit)) false crits);
+ val crits = [(true, Find_Theorems.Solves)];
+ fun find g = snd (Find_Theorems.find_theorems ctxt (SOME g) (SOME (! limit)) false crits);
fun prt_result (goal, results) =
let
@@ -57,7 +57,7 @@
in
Pretty.big_list
(msg ^ " could be solved directly with:")
- (map (FindTheorems.pretty_thm ctxt) results)
+ (map (Find_Theorems.pretty_thm ctxt) results)
end;
fun seek_against_goal () =
--- a/src/Tools/induct.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/Tools/induct.ML Thu Oct 29 18:53:58 2009 +0100
@@ -570,7 +570,7 @@
((fn [] => NONE | ts => List.last ts) #>
(fn NONE => TVar (("'a", 0), []) | SOME t => Term.fastype_of t) #>
find_inductT ctxt)) [[]]
- |> filter_out (forall Thm.is_internal);
+ |> filter_out (forall RuleCases.is_inner_rule);
fun get_inductP ctxt (fact :: _) = map single (find_inductP ctxt (Thm.concl_of fact))
| get_inductP _ _ = [];
--- a/src/ZF/arith_data.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/ZF/arith_data.ML Thu Oct 29 18:53:58 2009 +0100
@@ -68,7 +68,7 @@
fun prove_conv name tacs ctxt prems (t,u) =
if t aconv u then NONE
else
- let val prems' = List.filter (not o is_eq_thm) prems
+ let val prems' = filter_out is_eq_thm prems
val goal = Logic.list_implies (map (#prop o Thm.rep_thm) prems',
FOLogic.mk_Trueprop (mk_eq_iff (t, u)));
in SOME (prems' MRS Goal.prove ctxt [] [] goal (K (EVERY tacs)))
--- a/src/ZF/ind_syntax.ML Thu Oct 29 16:22:14 2009 +0000
+++ b/src/ZF/ind_syntax.ML Thu Oct 29 18:53:58 2009 +0100
@@ -96,8 +96,8 @@
fun union_params (rec_tm, cs) =
let val (_,args) = strip_comb rec_tm
fun is_ind arg = (type_of arg = iT)
- in case List.filter is_ind (args @ cs) of
- [] => @{const 0}
+ in case filter is_ind (args @ cs) of
+ [] => @{const 0}
| u_args => Balanced_Tree.make mk_Un u_args
end;