--- a/Admin/isatest/isatest-makedist Sun Oct 25 08:57:36 2009 +0100
+++ b/Admin/isatest/isatest-makedist Sun Oct 25 08:57:55 2009 +0100
@@ -110,8 +110,8 @@
sleep 15
$SSH macbroy5 "$MAKEALL $HOME/settings/mac-poly"
sleep 15
-#$SSH macbroy6 "$MAKEALL $HOME/settings/at-mac-poly-5.1-para"
-#sleep 15
+$SSH macbroy6 "sleep 10800; $MAKEALL $HOME/settings/at-mac-poly-5.1-para"
+sleep 15
$SSH atbroy51 "$HOME/admin/isatest/isatest-annomaly"
echo ------------------- spawned tests successfully --- `date` --- $HOSTNAME >> $DISTLOG 2>&1
--- a/NEWS Sun Oct 25 08:57:36 2009 +0100
+++ b/NEWS Sun Oct 25 08:57:55 2009 +0100
@@ -165,7 +165,8 @@
* New implementation of quickcheck uses generic code generator;
default generators are provided for all suitable HOL types, records
-and datatypes.
+and datatypes. Old quickcheck can be re-activated importing
+theory Library/SML_Quickcheck.
* Renamed theorems:
Suc_eq_add_numeral_1 -> Suc_eq_plus1
--- a/src/HOL/FunDef.thy Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/FunDef.thy Sun Oct 25 08:57:55 2009 +0100
@@ -9,19 +9,20 @@
uses
"Tools/prop_logic.ML"
"Tools/sat_solver.ML"
- ("Tools/Function/fundef_lib.ML")
- ("Tools/Function/fundef_common.ML")
+ ("Tools/Function/function_lib.ML")
+ ("Tools/Function/function_common.ML")
("Tools/Function/inductive_wrap.ML")
("Tools/Function/context_tree.ML")
- ("Tools/Function/fundef_core.ML")
+ ("Tools/Function/function_core.ML")
("Tools/Function/sum_tree.ML")
("Tools/Function/mutual.ML")
("Tools/Function/pattern_split.ML")
- ("Tools/Function/fundef.ML")
- ("Tools/Function/auto_term.ML")
+ ("Tools/Function/function.ML")
+ ("Tools/Function/relation.ML")
("Tools/Function/measure_functions.ML")
("Tools/Function/lexicographic_order.ML")
- ("Tools/Function/fundef_datatype.ML")
+ ("Tools/Function/pat_completeness.ML")
+ ("Tools/Function/fun.ML")
("Tools/Function/induction_scheme.ML")
("Tools/Function/termination.ML")
("Tools/Function/decompose.ML")
@@ -103,23 +104,25 @@
"wf R \<Longrightarrow> wfP (in_rel R)"
by (simp add: wfP_def)
-use "Tools/Function/fundef_lib.ML"
-use "Tools/Function/fundef_common.ML"
+use "Tools/Function/function_lib.ML"
+use "Tools/Function/function_common.ML"
use "Tools/Function/inductive_wrap.ML"
use "Tools/Function/context_tree.ML"
-use "Tools/Function/fundef_core.ML"
+use "Tools/Function/function_core.ML"
use "Tools/Function/sum_tree.ML"
use "Tools/Function/mutual.ML"
use "Tools/Function/pattern_split.ML"
-use "Tools/Function/auto_term.ML"
-use "Tools/Function/fundef.ML"
-use "Tools/Function/fundef_datatype.ML"
+use "Tools/Function/relation.ML"
+use "Tools/Function/function.ML"
+use "Tools/Function/pat_completeness.ML"
+use "Tools/Function/fun.ML"
use "Tools/Function/induction_scheme.ML"
setup {*
- Fundef.setup
- #> FundefDatatype.setup
- #> InductionScheme.setup
+ Function.setup
+ #> Pat_Completeness.setup
+ #> Function_Fun.setup
+ #> Induction_Scheme.setup
*}
subsection {* Measure Functions *}
@@ -139,7 +142,7 @@
by (rule is_measure_trivial)
use "Tools/Function/lexicographic_order.ML"
-setup LexicographicOrder.setup
+setup Lexicographic_Order.setup
subsection {* Congruence Rules *}
@@ -317,7 +320,7 @@
ML_val -- "setup inactive"
{*
- Context.theory_map (FundefCommon.set_termination_prover (ScnpReconstruct.decomp_scnp
+ Context.theory_map (Function_Common.set_termination_prover (ScnpReconstruct.decomp_scnp
[ScnpSolve.MAX, ScnpSolve.MIN, ScnpSolve.MS]))
*}
--- a/src/HOL/HOL.thy Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/HOL.thy Sun Oct 25 08:57:55 2009 +0100
@@ -2002,8 +2002,12 @@
*} "solve goal by normalization"
+subsection {* Counterexample Search Units *}
+
subsubsection {* Quickcheck *}
+quickcheck_params [size = 5, iterations = 50]
+
ML {*
structure Quickcheck_RecFun_Simps = Named_Thms
(
@@ -2014,37 +2018,8 @@
setup Quickcheck_RecFun_Simps.setup
-setup {*
- Quickcheck.add_generator ("SML", Codegen.test_term)
-*}
-quickcheck_params [size = 5, iterations = 50]
-
-subsection {* Preprocessing for the predicate compiler *}
-
-ML {*
-structure Predicate_Compile_Alternative_Defs = Named_Thms
-(
- val name = "code_pred_def"
- val description = "alternative definitions of constants for the Predicate Compiler"
-)
-*}
-
-ML {*
-structure Predicate_Compile_Inline_Defs = Named_Thms
-(
- val name = "code_pred_inline"
- val description = "inlining definitions for the Predicate Compiler"
-)
-*}
-
-setup {*
- Predicate_Compile_Alternative_Defs.setup
- #> Predicate_Compile_Inline_Defs.setup
- #> Predicate_Compile_Preproc_Const_Defs.setup
-*}
-
-subsection {* Nitpick setup *}
+subsubsection {* Nitpick setup *}
text {* This will be relocated once Nitpick is moved to HOL. *}
@@ -2079,6 +2054,31 @@
*}
+subsection {* Preprocessing for the predicate compiler *}
+
+ML {*
+structure Predicate_Compile_Alternative_Defs = Named_Thms
+(
+ val name = "code_pred_def"
+ val description = "alternative definitions of constants for the Predicate Compiler"
+)
+*}
+
+ML {*
+structure Predicate_Compile_Inline_Defs = Named_Thms
+(
+ val name = "code_pred_inline"
+ val description = "inlining definitions for the Predicate Compiler"
+)
+*}
+
+setup {*
+ Predicate_Compile_Alternative_Defs.setup
+ #> Predicate_Compile_Inline_Defs.setup
+ #> Predicate_Compile_Preproc_Const_Defs.setup
+*}
+
+
subsection {* Legacy tactics and ML bindings *}
ML {*
--- a/src/HOL/IsaMakefile Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/IsaMakefile Sun Oct 25 08:57:55 2009 +0100
@@ -156,21 +156,22 @@
Tools/Datatype/datatype_realizer.ML \
Tools/Datatype/datatype_rep_proofs.ML \
Tools/dseq.ML \
- Tools/Function/auto_term.ML \
Tools/Function/context_tree.ML \
Tools/Function/decompose.ML \
Tools/Function/descent.ML \
- Tools/Function/fundef_common.ML \
- Tools/Function/fundef_core.ML \
- Tools/Function/fundef_datatype.ML \
- Tools/Function/fundef_lib.ML \
- Tools/Function/fundef.ML \
+ Tools/Function/function_common.ML \
+ Tools/Function/function_core.ML \
+ Tools/Function/function_lib.ML \
+ Tools/Function/function.ML \
+ Tools/Function/fun.ML \
Tools/Function/induction_scheme.ML \
Tools/Function/inductive_wrap.ML \
Tools/Function/lexicographic_order.ML \
Tools/Function/measure_functions.ML \
Tools/Function/mutual.ML \
+ Tools/Function/pat_completeness.ML \
Tools/Function/pattern_split.ML \
+ Tools/Function/relation.ML \
Tools/Function/scnp_reconstruct.ML \
Tools/Function/scnp_solve.ML \
Tools/Function/size.ML \
@@ -358,7 +359,7 @@
Library/Enum.thy Library/Float.thy $(SRC)/Tools/float.ML \
$(SRC)/HOL/Tools/float_arith.ML Library/positivstellensatz.ML \
Library/reify_data.ML Library/reflection.ML Library/LaTeXsugar.thy \
- Library/OptionalSugar.thy
+ Library/OptionalSugar.thy Library/SML_Quickcheck.thy
@cd Library; $(ISABELLE_TOOL) usedir $(OUT)/HOL Library
--- a/src/HOL/Library/Library.thy Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Library/Library.thy Sun Oct 25 08:57:55 2009 +0100
@@ -52,6 +52,7 @@
Ramsey
Reflection
RBT
+ SML_Quickcheck
State_Monad
Sum_Of_Squares
Topology_Euclidean_Space
--- a/src/HOL/Library/Multiset.thy Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Library/Multiset.thy Sun Oct 25 08:57:55 2009 +0100
@@ -1607,7 +1607,7 @@
rtac @{thm nonempty_plus} ORELSE' rtac @{thm nonempty_single}
val regroup_munion_conv =
- FundefLib.regroup_conv @{const_name Multiset.Mempty} @{const_name plus}
+ Function_Lib.regroup_conv @{const_name Multiset.Mempty} @{const_name plus}
(map (fn t => t RS eq_reflection) (@{thms union_ac} @ @{thms empty_idemp}))
fun unfold_pwleq_tac i =
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/SML_Quickcheck.thy Sun Oct 25 08:57:55 2009 +0100
@@ -0,0 +1,12 @@
+
+header {* Install quickcheck of SML code generator *}
+
+theory SML_Quickcheck
+imports Main
+begin
+
+setup {*
+ Quickcheck.add_generator ("SML", Codegen.test_term)
+*}
+
+end
--- a/src/HOL/Predicate.thy Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Predicate.thy Sun Oct 25 08:57:55 2009 +0100
@@ -471,49 +471,49 @@
"is_empty (A \<squnion> B) \<longleftrightarrow> is_empty A \<and> is_empty B"
by (auto simp add: is_empty_def intro: sup_eq_bot_eq1 sup_eq_bot_eq2)
-definition singleton :: "'a pred \<Rightarrow> 'a" where
- "singleton A = (if \<exists>!x. eval A x then THE x. eval A x else undefined)"
+definition singleton :: "(unit => 'a) \<Rightarrow> 'a pred \<Rightarrow> 'a" where
+ "singleton dfault A = (if \<exists>!x. eval A x then THE x. eval A x else dfault ())"
lemma singleton_eqI:
- "\<exists>!x. eval A x \<Longrightarrow> eval A x \<Longrightarrow> singleton A = x"
+ "\<exists>!x. eval A x \<Longrightarrow> eval A x \<Longrightarrow> singleton dfault A = x"
by (auto simp add: singleton_def)
lemma eval_singletonI:
- "\<exists>!x. eval A x \<Longrightarrow> eval A (singleton A)"
+ "\<exists>!x. eval A x \<Longrightarrow> eval A (singleton dfault A)"
proof -
assume assm: "\<exists>!x. eval A x"
then obtain x where "eval A x" ..
- moreover with assm have "singleton A = x" by (rule singleton_eqI)
+ moreover with assm have "singleton dfault A = x" by (rule singleton_eqI)
ultimately show ?thesis by simp
qed
lemma single_singleton:
- "\<exists>!x. eval A x \<Longrightarrow> single (singleton A) = A"
+ "\<exists>!x. eval A x \<Longrightarrow> single (singleton dfault A) = A"
proof -
assume assm: "\<exists>!x. eval A x"
- then have "eval A (singleton A)"
+ then have "eval A (singleton dfault A)"
by (rule eval_singletonI)
- moreover from assm have "\<And>x. eval A x \<Longrightarrow> singleton A = x"
+ moreover from assm have "\<And>x. eval A x \<Longrightarrow> singleton dfault A = x"
by (rule singleton_eqI)
- ultimately have "eval (single (singleton A)) = eval A"
+ ultimately have "eval (single (singleton dfault A)) = eval A"
by (simp (no_asm_use) add: single_def expand_fun_eq) blast
then show ?thesis by (simp add: eval_inject)
qed
lemma singleton_undefinedI:
- "\<not> (\<exists>!x. eval A x) \<Longrightarrow> singleton A = undefined"
+ "\<not> (\<exists>!x. eval A x) \<Longrightarrow> singleton dfault A = dfault ()"
by (simp add: singleton_def)
lemma singleton_bot:
- "singleton \<bottom> = undefined"
+ "singleton dfault \<bottom> = dfault ()"
by (auto simp add: bot_pred_def intro: singleton_undefinedI)
lemma singleton_single:
- "singleton (single x) = x"
+ "singleton dfault (single x) = x"
by (auto simp add: intro: singleton_eqI singleI elim: singleE)
lemma singleton_sup_single_single:
- "singleton (single x \<squnion> single y) = (if x = y then x else undefined)"
+ "singleton dfault (single x \<squnion> single y) = (if x = y then x else dfault ())"
proof (cases "x = y")
case True then show ?thesis by (simp add: singleton_single)
next
@@ -523,25 +523,25 @@
by (auto intro: supI1 supI2 singleI)
with False have "\<not> (\<exists>!z. eval (single x \<squnion> single y) z)"
by blast
- then have "singleton (single x \<squnion> single y) = undefined"
+ then have "singleton dfault (single x \<squnion> single y) = dfault ()"
by (rule singleton_undefinedI)
with False show ?thesis by simp
qed
lemma singleton_sup_aux:
- "singleton (A \<squnion> B) = (if A = \<bottom> then singleton B
- else if B = \<bottom> then singleton A
- else singleton
- (single (singleton A) \<squnion> single (singleton B)))"
+ "singleton dfault (A \<squnion> B) = (if A = \<bottom> then singleton dfault B
+ else if B = \<bottom> then singleton dfault A
+ else singleton dfault
+ (single (singleton dfault A) \<squnion> single (singleton dfault B)))"
proof (cases "(\<exists>!x. eval A x) \<and> (\<exists>!y. eval B y)")
case True then show ?thesis by (simp add: single_singleton)
next
case False
from False have A_or_B:
- "singleton A = undefined \<or> singleton B = undefined"
+ "singleton dfault A = dfault () \<or> singleton dfault B = dfault ()"
by (auto intro!: singleton_undefinedI)
- then have rhs: "singleton
- (single (singleton A) \<squnion> single (singleton B)) = undefined"
+ then have rhs: "singleton dfault
+ (single (singleton dfault A) \<squnion> single (singleton dfault B)) = dfault ()"
by (auto simp add: singleton_sup_single_single singleton_single)
from False have not_unique:
"\<not> (\<exists>!x. eval A x) \<or> \<not> (\<exists>!y. eval B y)" by simp
@@ -551,7 +551,7 @@
by (blast elim: not_bot)
with True not_unique have "\<not> (\<exists>!x. eval (A \<squnion> B) x)"
by (auto simp add: sup_pred_def bot_pred_def)
- then have "singleton (A \<squnion> B) = undefined" by (rule singleton_undefinedI)
+ then have "singleton dfault (A \<squnion> B) = dfault ()" by (rule singleton_undefinedI)
with True rhs show ?thesis by simp
next
case False then show ?thesis by auto
@@ -559,10 +559,10 @@
qed
lemma singleton_sup:
- "singleton (A \<squnion> B) = (if A = \<bottom> then singleton B
- else if B = \<bottom> then singleton A
- else if singleton A = singleton B then singleton A else undefined)"
-using singleton_sup_aux [of A B] by (simp only: singleton_sup_single_single)
+ "singleton dfault (A \<squnion> B) = (if A = \<bottom> then singleton dfault B
+ else if B = \<bottom> then singleton dfault A
+ else if singleton dfault A = singleton dfault B then singleton dfault A else dfault ())"
+using singleton_sup_aux [of dfault A B] by (simp only: singleton_sup_single_single)
subsubsection {* Derived operations *}
@@ -743,36 +743,43 @@
"is_empty (Seq f) \<longleftrightarrow> null (f ())"
by (simp add: null_is_empty Seq_def)
-primrec the_only :: "'a seq \<Rightarrow> 'a" where
- [code del]: "the_only Empty = undefined"
- | "the_only (Insert x P) = (if is_empty P then x else let y = singleton P in if x = y then x else undefined)"
- | "the_only (Join P xq) = (if is_empty P then the_only xq else if null xq then singleton P
- else let x = singleton P; y = the_only xq in
- if x = y then x else undefined)"
+primrec the_only :: "(unit \<Rightarrow> 'a) \<Rightarrow> 'a seq \<Rightarrow> 'a" where
+ [code del]: "the_only dfault Empty = dfault ()"
+ | "the_only dfault (Insert x P) = (if is_empty P then x else let y = singleton dfault P in if x = y then x else dfault ())"
+ | "the_only dfault (Join P xq) = (if is_empty P then the_only dfault xq else if null xq then singleton dfault P
+ else let x = singleton dfault P; y = the_only dfault xq in
+ if x = y then x else dfault ())"
lemma the_only_singleton:
- "the_only xq = singleton (pred_of_seq xq)"
+ "the_only dfault xq = singleton dfault (pred_of_seq xq)"
by (induct xq)
(auto simp add: singleton_bot singleton_single is_empty_def
null_is_empty Let_def singleton_sup)
lemma singleton_code [code]:
- "singleton (Seq f) = (case f ()
- of Empty \<Rightarrow> undefined
+ "singleton dfault (Seq f) = (case f ()
+ of Empty \<Rightarrow> dfault ()
| Insert x P \<Rightarrow> if is_empty P then x
- else let y = singleton P in
- if x = y then x else undefined
- | Join P xq \<Rightarrow> if is_empty P then the_only xq
- else if null xq then singleton P
- else let x = singleton P; y = the_only xq in
- if x = y then x else undefined)"
+ else let y = singleton dfault P in
+ if x = y then x else dfault ()
+ | Join P xq \<Rightarrow> if is_empty P then the_only dfault xq
+ else if null xq then singleton dfault P
+ else let x = singleton dfault P; y = the_only dfault xq in
+ if x = y then x else dfault ())"
by (cases "f ()")
(auto simp add: Seq_def the_only_singleton is_empty_def
null_is_empty singleton_bot singleton_single singleton_sup Let_def)
-lemma meta_fun_cong:
-"f == g ==> f x == g x"
-by simp
+definition not_unique :: "'a pred => 'a"
+where
+ [code del]: "not_unique A = (THE x. eval A x)"
+
+definition the :: "'a pred => 'a"
+where
+ [code del]: "the A = (THE x. eval A x)"
+
+lemma the_eq[code]: "the A = singleton (\<lambda>x. not_unique A) A"
+by (auto simp add: the_def singleton_def not_unique_def)
ML {*
signature PREDICATE =
@@ -819,6 +826,8 @@
code_const Seq and Empty and Insert and Join
(Eval "Predicate.Seq" and "Predicate.Empty" and "Predicate.Insert/ (_,/ _)" and "Predicate.Join/ (_,/ _)")
+code_abort not_unique
+
text {* dummy setup for @{text code_pred} and @{text values} keywords *}
ML {*
@@ -852,6 +861,6 @@
hide (open) type pred seq
hide (open) const Pred eval single bind is_empty singleton if_pred not_pred
- Empty Insert Join Seq member pred_of_seq "apply" adjunct null the_only eq map
+ Empty Insert Join Seq member pred_of_seq "apply" adjunct null the_only eq map not_unique the
end
--- a/src/HOL/Product_Type.thy Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Product_Type.thy Sun Oct 25 08:57:55 2009 +0100
@@ -6,7 +6,7 @@
header {* Cartesian products *}
theory Product_Type
-imports Inductive
+imports Inductive Nat
uses
("Tools/split_rule.ML")
("Tools/inductive_set.ML")
--- a/src/HOL/SizeChange/sct.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/SizeChange/sct.ML Sun Oct 25 08:57:55 2009 +0100
@@ -112,7 +112,7 @@
end
fun bind_many [] = I
- | bind_many vs = FundefLib.tupled_lambda (foldr1 HOLogic.mk_prod vs)
+ | bind_many vs = Function_Lib.tupled_lambda (foldr1 HOLogic.mk_prod vs)
(* Builds relation descriptions from a relation definition *)
fun mk_reldescs (Abs a) =
--- a/src/HOL/Tools/ATP_Manager/atp_wrapper.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/ATP_Manager/atp_wrapper.ML Sun Oct 25 08:57:55 2009 +0100
@@ -141,7 +141,8 @@
(* write out problem file and call prover *)
fun cmd_line probfile = "TIMEFORMAT='%3U'; { time " ^ space_implode " "
- [File.shell_path cmd, args, File.shell_path probfile] ^ " ; } 2>&1"
+ [File.shell_path cmd, args, File.shell_path probfile] ^ " 2> /dev/null" ^
+ " ; } 2>&1"
fun split_time s =
let
val split = String.tokens (fn c => str c = "\n");
--- a/src/HOL/Tools/Function/auto_term.ML Sun Oct 25 08:57:36 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,36 +0,0 @@
-(* Title: HOL/Tools/Function/auto_term.ML
- Author: Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions.
-Method "relation" to commence a termination proof using a user-specified relation.
-*)
-
-signature FUNDEF_RELATION =
-sig
- val relation_tac: Proof.context -> term -> int -> tactic
- val setup: theory -> theory
-end
-
-structure FundefRelation : FUNDEF_RELATION =
-struct
-
-fun inst_thm ctxt rel st =
- let
- val cert = Thm.cterm_of (ProofContext.theory_of ctxt)
- val rel' = cert (singleton (Variable.polymorphic ctxt) rel)
- val st' = Thm.incr_indexes (#maxidx (Thm.rep_cterm rel') + 1) st
- val Rvar = cert (Var (the_single (Term.add_vars (prop_of st') [])))
- in
- Drule.cterm_instantiate [(Rvar, rel')] st'
- end
-
-fun relation_tac ctxt rel i =
- TRY (FundefCommon.apply_termination_rule ctxt i)
- THEN PRIMITIVE (inst_thm ctxt rel)
-
-val setup =
- Method.setup @{binding relation}
- (Args.term >> (fn rel => fn ctxt => SIMPLE_METHOD' (relation_tac ctxt rel)))
- "proves termination using a user-specified wellfounded relation"
-
-end
--- a/src/HOL/Tools/Function/context_tree.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Function/context_tree.ML Sun Oct 25 08:57:55 2009 +0100
@@ -5,15 +5,15 @@
Builds and traverses trees of nested contexts along a term.
*)
-signature FUNDEF_CTXTREE =
+signature FUNCTION_CTXTREE =
sig
type ctxt = (string * typ) list * thm list (* poor man's contexts: fixes + assumes *)
type ctx_tree
(* FIXME: This interface is a mess and needs to be cleaned up! *)
- val get_fundef_congs : Proof.context -> thm list
- val add_fundef_cong : thm -> Context.generic -> Context.generic
- val map_fundef_congs : (thm list -> thm list) -> Context.generic -> Context.generic
+ val get_function_congs : Proof.context -> thm list
+ val add_function_cong : thm -> Context.generic -> Context.generic
+ val map_function_congs : (thm list -> thm list) -> Context.generic -> Context.generic
val cong_add: attribute
val cong_del: attribute
@@ -36,15 +36,15 @@
val rewrite_by_tree : theory -> term -> thm -> (thm * thm) list -> ctx_tree -> thm * (thm * thm) list
end
-structure FundefCtxTree : FUNDEF_CTXTREE =
+structure Function_Ctx_Tree : FUNCTION_CTXTREE =
struct
type ctxt = (string * typ) list * thm list
-open FundefCommon
-open FundefLib
+open Function_Common
+open Function_Lib
-structure FundefCongs = GenericDataFun
+structure FunctionCongs = GenericDataFun
(
type T = thm list
val empty = []
@@ -52,14 +52,14 @@
fun merge _ = Thm.merge_thms
);
-val get_fundef_congs = FundefCongs.get o Context.Proof
-val map_fundef_congs = FundefCongs.map
-val add_fundef_cong = FundefCongs.map o Thm.add_thm
+val get_function_congs = FunctionCongs.get o Context.Proof
+val map_function_congs = FunctionCongs.map
+val add_function_cong = FunctionCongs.map o Thm.add_thm
(* congruence rules *)
-val cong_add = Thm.declaration_attribute (map_fundef_congs o Thm.add_thm o safe_mk_meta_eq);
-val cong_del = Thm.declaration_attribute (map_fundef_congs o Thm.del_thm o safe_mk_meta_eq);
+val cong_add = Thm.declaration_attribute (map_function_congs o Thm.add_thm o safe_mk_meta_eq);
+val cong_del = Thm.declaration_attribute (map_function_congs o Thm.del_thm o safe_mk_meta_eq);
type depgraph = int IntGraph.T
@@ -128,7 +128,7 @@
fun mk_tree fvar h ctxt t =
let
- val congs = get_fundef_congs ctxt
+ val congs = get_function_congs ctxt
val congs_deps = map (fn c => (c, cong_deps c)) (congs @ default_congs) (* FIXME: Save in theory *)
fun matchcall (a $ b) = if a = Free fvar then SOME b else NONE
--- a/src/HOL/Tools/Function/decompose.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Function/decompose.ML Sun Oct 25 08:57:55 2009 +0100
@@ -33,8 +33,8 @@
Const (@{const_name Set.empty}, fastype_of c1))
|> HOLogic.mk_Trueprop (* "C1 O C2 = {}" *)
- val chain = case FundefLib.try_proof (cterm_of thy goal) chain_tac of
- FundefLib.Solved thm => SOME thm
+ val chain = case Function_Lib.try_proof (cterm_of thy goal) chain_tac of
+ Function_Lib.Solved thm => SOME thm
| _ => NONE
in
Termination.note_chain c1 c2 chain D
@@ -62,7 +62,7 @@
let
val is = map (fn c => find_index (curry op aconv c) cs') cs
in
- CONVERSION (Conv.arg_conv (Conv.arg_conv (FundefLib.regroup_union_conv is))) i
+ CONVERSION (Conv.arg_conv (Conv.arg_conv (Function_Lib.regroup_union_conv is))) i
end)
--- a/src/HOL/Tools/Function/descent.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Function/descent.ML Sun Oct 25 08:57:55 2009 +0100
@@ -35,7 +35,7 @@
(measures_of p) (measures_of q) D
end
in
- cont (FundefCommon.PROFILE "deriving descents" (fold derive cs) D) i
+ cont (Function_Common.PROFILE "deriving descents" (fold derive cs) D) i
end)
val derive_diag = gen_descent true
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/fun.ML Sun Oct 25 08:57:55 2009 +0100
@@ -0,0 +1,178 @@
+(* Title: HOL/Tools/Function/fun.ML
+ Author: Alexander Krauss, TU Muenchen
+
+Sequential mode for function definitions
+Command "fun" for fully automated function definitions
+*)
+
+signature FUNCTION_FUN =
+sig
+ val add_fun : Function_Common.function_config ->
+ (binding * typ option * mixfix) list -> (Attrib.binding * term) list ->
+ bool -> local_theory -> Proof.context
+ val add_fun_cmd : Function_Common.function_config ->
+ (binding * string option * mixfix) list -> (Attrib.binding * string) list ->
+ bool -> local_theory -> Proof.context
+
+ val setup : theory -> theory
+end
+
+structure Function_Fun : FUNCTION_FUN =
+struct
+
+open Function_Lib
+open Function_Common
+
+
+fun check_pats ctxt geq =
+ let
+ fun err str = error (cat_lines ["Malformed definition:",
+ str ^ " not allowed in sequential mode.",
+ Syntax.string_of_term ctxt geq])
+ val thy = ProofContext.theory_of ctxt
+
+ fun check_constr_pattern (Bound _) = ()
+ | check_constr_pattern t =
+ let
+ val (hd, args) = strip_comb t
+ in
+ (((case Datatype.info_of_constr thy (dest_Const hd) of
+ SOME _ => ()
+ | NONE => err "Non-constructor pattern")
+ handle TERM ("dest_Const", _) => err "Non-constructor patterns");
+ map check_constr_pattern args;
+ ())
+ end
+
+ val (fname, qs, gs, args, rhs) = split_def ctxt geq
+
+ val _ = if not (null gs) then err "Conditional equations" else ()
+ val _ = map check_constr_pattern args
+
+ (* just count occurrences to check linearity *)
+ val _ = if fold (fold_aterms (fn Bound _ => Integer.add 1 | _ => I)) args 0 > length qs
+ then err "Nonlinear patterns" else ()
+ in
+ ()
+ end
+
+val by_pat_completeness_auto =
+ Proof.global_future_terminal_proof
+ (Method.Basic Pat_Completeness.pat_completeness,
+ SOME (Method.Source_i (Args.src (("HOL.auto", []), Position.none))))
+
+fun termination_by method int =
+ Function.termination_proof NONE
+ #> Proof.global_future_terminal_proof (Method.Basic method, NONE) int
+
+fun mk_catchall fixes arity_of =
+ let
+ fun mk_eqn ((fname, fT), _) =
+ let
+ val n = arity_of fname
+ val (argTs, rT) = chop n (binder_types fT)
+ |> apsnd (fn Ts => Ts ---> body_type fT)
+
+ val qs = map Free (Name.invent_list [] "a" n ~~ argTs)
+ in
+ HOLogic.mk_eq(list_comb (Free (fname, fT), qs),
+ Const ("HOL.undefined", rT))
+ |> HOLogic.mk_Trueprop
+ |> fold_rev Logic.all qs
+ end
+ in
+ map mk_eqn fixes
+ end
+
+fun add_catchall ctxt fixes spec =
+ let val fqgars = map (split_def ctxt) spec
+ val arity_of = map (fn (fname,_,_,args,_) => (fname, length args)) fqgars
+ |> AList.lookup (op =) #> the
+ in
+ spec @ mk_catchall fixes arity_of
+ end
+
+fun warn_if_redundant ctxt origs tss =
+ let
+ fun msg t = "Ignoring redundant equation: " ^ quote (Syntax.string_of_term ctxt t)
+
+ val (tss', _) = chop (length origs) tss
+ fun check (t, []) = (warning (msg t); [])
+ | check (t, s) = s
+ in
+ (map check (origs ~~ tss'); tss)
+ end
+
+
+fun sequential_preproc (config as FunctionConfig {sequential, ...}) ctxt fixes spec =
+ if sequential then
+ let
+ val (bnds, eqss) = split_list spec
+
+ val eqs = map the_single eqss
+
+ val feqs = eqs
+ |> tap (check_defs ctxt fixes) (* Standard checks *)
+ |> tap (map (check_pats ctxt)) (* More checks for sequential mode *)
+
+ val compleqs = add_catchall ctxt fixes feqs (* Completion *)
+
+ val spliteqs = warn_if_redundant ctxt feqs
+ (Function_Split.split_all_equations ctxt compleqs)
+
+ fun restore_spec thms =
+ bnds ~~ Library.take (length bnds, Library.unflat spliteqs thms)
+
+ val spliteqs' = flat (Library.take (length bnds, spliteqs))
+ val fnames = map (fst o fst) fixes
+ val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) spliteqs'
+
+ fun sort xs = partition_list (fn i => fn (j,_) => i = j) 0 (length fnames - 1) (indices ~~ xs)
+ |> map (map snd)
+
+
+ val bnds' = bnds @ replicate (length spliteqs - length bnds) Attrib.empty_binding
+
+ (* using theorem names for case name currently disabled *)
+ val case_names = map_index (fn (i, (_, es)) => mk_case_names i "" (length es))
+ (bnds' ~~ spliteqs)
+ |> flat
+ in
+ (flat spliteqs, restore_spec, sort, case_names)
+ end
+ else
+ Function_Common.empty_preproc check_defs config ctxt fixes spec
+
+val setup =
+ Context.theory_map (Function_Common.set_preproc sequential_preproc)
+
+
+val fun_config = FunctionConfig { sequential=true, default="%x. undefined" (*FIXME dynamic scoping*),
+ domintros=false, partials=false, tailrec=false }
+
+fun gen_fun add config fixes statements int lthy =
+ let val group = serial_string () in
+ lthy
+ |> LocalTheory.set_group group
+ |> add fixes statements config
+ |> by_pat_completeness_auto int
+ |> LocalTheory.restore
+ |> LocalTheory.set_group group
+ |> termination_by (Function_Common.get_termination_prover lthy) int
+ end;
+
+val add_fun = gen_fun Function.add_function
+val add_fun_cmd = gen_fun Function.add_function_cmd
+
+
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val _ =
+ OuterSyntax.local_theory' "fun" "define general recursive functions (short version)" K.thy_decl
+ (function_parser fun_config
+ >> (fn ((config, fixes), statements) => add_fun_cmd config fixes statements));
+
+end
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/function.ML Sun Oct 25 08:57:55 2009 +0100
@@ -0,0 +1,227 @@
+(* Title: HOL/Tools/Function/fundef.ML
+ Author: Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions.
+Isar commands.
+*)
+
+signature FUNCTION =
+sig
+ val add_function : (binding * typ option * mixfix) list
+ -> (Attrib.binding * term) list
+ -> Function_Common.function_config
+ -> local_theory
+ -> Proof.state
+ val add_function_cmd : (binding * string option * mixfix) list
+ -> (Attrib.binding * string) list
+ -> Function_Common.function_config
+ -> local_theory
+ -> Proof.state
+
+ val termination_proof : term option -> local_theory -> Proof.state
+ val termination_proof_cmd : string option -> local_theory -> Proof.state
+ val termination : term option -> local_theory -> Proof.state
+ val termination_cmd : string option -> local_theory -> Proof.state
+
+ val setup : theory -> theory
+ val get_congs : Proof.context -> thm list
+end
+
+
+structure Function : FUNCTION =
+struct
+
+open Function_Lib
+open Function_Common
+
+val simp_attribs = map (Attrib.internal o K)
+ [Simplifier.simp_add,
+ Code.add_default_eqn_attribute,
+ Nitpick_Simps.add,
+ Quickcheck_RecFun_Simps.add]
+
+val psimp_attribs = map (Attrib.internal o K)
+ [Simplifier.simp_add,
+ Nitpick_Psimps.add]
+
+fun note_theorem ((name, atts), ths) =
+ LocalTheory.note Thm.generatedK ((Binding.qualified_name name, atts), ths)
+
+fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_"
+
+fun add_simps fnames post sort extra_qualify label moreatts simps lthy =
+ let
+ val spec = post simps
+ |> map (apfst (apsnd (fn ats => moreatts @ ats)))
+ |> map (apfst (apfst extra_qualify))
+
+ val (saved_spec_simps, lthy) =
+ fold_map (LocalTheory.note Thm.generatedK) spec lthy
+
+ val saved_simps = maps snd saved_spec_simps
+ val simps_by_f = sort saved_simps
+
+ fun add_for_f fname simps =
+ note_theorem ((Long_Name.qualify fname label, []), simps) #> snd
+ in
+ (saved_simps,
+ fold2 add_for_f fnames simps_by_f lthy)
+ end
+
+fun gen_add_function is_external prep default_constraint fixspec eqns config lthy =
+ let
+ val constrn_fxs = map (fn (b, T, mx) => (b, SOME (the_default default_constraint T), mx))
+ val ((fixes0, spec0), ctxt') = prep (constrn_fxs fixspec) eqns lthy
+ val fixes = map (apfst (apfst Binding.name_of)) fixes0;
+ val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0;
+ val (eqs, post, sort_cont, cnames) = get_preproc lthy config ctxt' fixes spec
+
+ val defname = mk_defname fixes
+ val FunctionConfig {partials, ...} = config
+
+ val ((goalstate, cont), lthy) =
+ Function_Mutual.prepare_function_mutual config defname fixes eqs lthy
+
+ fun afterqed [[proof]] lthy =
+ let
+ val FunctionResult {fs, R, psimps, trsimps, simple_pinducts, termination,
+ domintros, cases, ...} =
+ cont (Thm.close_derivation proof)
+
+ val fnames = map (fst o fst) fixes
+ val qualify = Long_Name.qualify defname
+ val addsmps = add_simps fnames post sort_cont
+
+ val (((psimps', pinducts'), (_, [termination'])), lthy) =
+ lthy
+ |> addsmps (Binding.qualify false "partial") "psimps"
+ psimp_attribs psimps
+ ||> fold_option (snd oo addsmps I "simps" simp_attribs) trsimps
+ ||>> note_theorem ((qualify "pinduct",
+ [Attrib.internal (K (RuleCases.case_names cnames)),
+ Attrib.internal (K (RuleCases.consumes 1)),
+ Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts)
+ ||>> note_theorem ((qualify "termination", []), [termination])
+ ||> (snd o note_theorem ((qualify "cases",
+ [Attrib.internal (K (RuleCases.case_names cnames))]), [cases]))
+ ||> fold_option (snd oo curry note_theorem (qualify "domintros", [])) domintros
+
+ val cdata = FunctionCtxData { add_simps=addsmps, case_names=cnames, psimps=psimps',
+ pinducts=snd pinducts', termination=termination',
+ fs=fs, R=R, defname=defname }
+ val _ =
+ if not is_external then ()
+ else Specification.print_consts lthy (K false) (map fst fixes)
+ in
+ lthy
+ |> LocalTheory.declaration (add_function_data o morph_function_data cdata)
+ end
+ in
+ lthy
+ |> is_external ? LocalTheory.set_group (serial_string ())
+ |> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
+ |> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd
+ end
+
+val add_function = gen_add_function false Specification.check_spec (TypeInfer.anyT HOLogic.typeS)
+val add_function_cmd = gen_add_function true Specification.read_spec "_::type"
+
+fun gen_termination_proof prep_term raw_term_opt lthy =
+ let
+ val term_opt = Option.map (prep_term lthy) raw_term_opt
+ val data = the (case term_opt of
+ SOME t => (import_function_data t lthy
+ handle Option.Option =>
+ error ("Not a function: " ^ quote (Syntax.string_of_term lthy t)))
+ | NONE => (import_last_function lthy handle Option.Option => error "Not a function"))
+
+ val FunctionCtxData { termination, R, add_simps, case_names, psimps,
+ pinducts, defname, ...} = data
+ val domT = domain_type (fastype_of R)
+ val goal = HOLogic.mk_Trueprop
+ (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT)))
+ fun afterqed [[totality]] lthy =
+ let
+ val totality = Thm.close_derivation totality
+ val remove_domain_condition =
+ full_simplify (HOL_basic_ss addsimps [totality, True_implies_equals])
+ val tsimps = map remove_domain_condition psimps
+ val tinduct = map remove_domain_condition pinducts
+ val qualify = Long_Name.qualify defname;
+ in
+ lthy
+ |> add_simps I "simps" simp_attribs tsimps |> snd
+ |> note_theorem
+ ((qualify "induct",
+ [Attrib.internal (K (RuleCases.case_names case_names))]),
+ tinduct) |> snd
+ end
+ in
+ lthy
+ |> ProofContext.note_thmss ""
+ [((Binding.empty, [ContextRules.rule_del]), [([allI], [])])] |> snd
+ |> ProofContext.note_thmss ""
+ [((Binding.empty, [ContextRules.intro_bang (SOME 1)]), [([allI], [])])] |> snd
+ |> ProofContext.note_thmss ""
+ [((Binding.name "termination", [ContextRules.intro_bang (SOME 0)]),
+ [([Goal.norm_result termination], [])])] |> snd
+ |> Proof.theorem_i NONE afterqed [[(goal, [])]]
+ end
+
+val termination_proof = gen_termination_proof Syntax.check_term;
+val termination_proof_cmd = gen_termination_proof Syntax.read_term;
+
+fun termination term_opt lthy =
+ lthy
+ |> LocalTheory.set_group (serial_string ())
+ |> termination_proof term_opt;
+
+fun termination_cmd term_opt lthy =
+ lthy
+ |> LocalTheory.set_group (serial_string ())
+ |> termination_proof_cmd term_opt;
+
+
+(* Datatype hook to declare datatype congs as "function_congs" *)
+
+
+fun add_case_cong n thy =
+ Context.theory_map (Function_Ctx_Tree.map_function_congs (Thm.add_thm
+ (Datatype.the_info thy n
+ |> #case_cong
+ |> safe_mk_meta_eq)))
+ thy
+
+val setup_case_cong = Datatype.interpretation (K (fold add_case_cong))
+
+
+(* setup *)
+
+val setup =
+ Attrib.setup @{binding fundef_cong}
+ (Attrib.add_del Function_Ctx_Tree.cong_add Function_Ctx_Tree.cong_del)
+ "declaration of congruence rule for function definitions"
+ #> setup_case_cong
+ #> Function_Relation.setup
+ #> Function_Common.Termination_Simps.setup
+
+val get_congs = Function_Ctx_Tree.get_function_congs
+
+
+(* outer syntax *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val _ =
+ OuterSyntax.local_theory_to_proof "function" "define general recursive functions" K.thy_goal
+ (function_parser default_config
+ >> (fn ((config, fixes), statements) => add_function_cmd fixes statements config));
+
+val _ =
+ OuterSyntax.local_theory_to_proof "termination" "prove termination of a recursive function" K.thy_goal
+ (Scan.option P.term >> termination_cmd);
+
+end;
+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/function_common.ML Sun Oct 25 08:57:55 2009 +0100
@@ -0,0 +1,348 @@
+(* Title: HOL/Tools/Function/fundef_common.ML
+ Author: Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions.
+Common definitions and other infrastructure.
+*)
+
+structure Function_Common =
+struct
+
+local open Function_Lib in
+
+(* Profiling *)
+val profile = Unsynchronized.ref false;
+
+fun PROFILE msg = if !profile then timeap_msg msg else I
+
+
+val acc_const_name = @{const_name accp}
+fun mk_acc domT R =
+ Const (acc_const_name, (domT --> domT --> HOLogic.boolT) --> domT --> HOLogic.boolT) $ R
+
+val function_name = suffix "C"
+val graph_name = suffix "_graph"
+val rel_name = suffix "_rel"
+val dom_name = suffix "_dom"
+
+(* Termination rules *)
+
+structure TerminationRule = GenericDataFun
+(
+ type T = thm list
+ val empty = []
+ val extend = I
+ fun merge _ = Thm.merge_thms
+);
+
+val get_termination_rules = TerminationRule.get
+val store_termination_rule = TerminationRule.map o cons
+val apply_termination_rule = resolve_tac o get_termination_rules o Context.Proof
+
+
+(* Function definition result data *)
+
+datatype function_result =
+ FunctionResult of
+ {
+ fs: term list,
+ G: term,
+ R: term,
+
+ psimps : thm list,
+ trsimps : thm list option,
+
+ simple_pinducts : thm list,
+ cases : thm,
+ termination : thm,
+ domintros : thm list option
+ }
+
+
+datatype function_context_data =
+ FunctionCtxData of
+ {
+ defname : string,
+
+ (* contains no logical entities: invariant under morphisms *)
+ add_simps : (binding -> binding) -> string -> Attrib.src list -> thm list
+ -> local_theory -> thm list * local_theory,
+ case_names : string list,
+
+ fs : term list,
+ R : term,
+
+ psimps: thm list,
+ pinducts: thm list,
+ termination: thm
+ }
+
+fun morph_function_data (FunctionCtxData {add_simps, case_names, fs, R,
+ psimps, pinducts, termination, defname}) phi =
+ let
+ val term = Morphism.term phi val thm = Morphism.thm phi val fact = Morphism.fact phi
+ val name = Binding.name_of o Morphism.binding phi o Binding.name
+ in
+ FunctionCtxData { add_simps = add_simps, case_names = case_names,
+ fs = map term fs, R = term R, psimps = fact psimps,
+ pinducts = fact pinducts, termination = thm termination,
+ defname = name defname }
+ end
+
+structure FunctionData = GenericDataFun
+(
+ type T = (term * function_context_data) Item_Net.T;
+ val empty = Item_Net.init
+ (op aconv o pairself fst : (term * function_context_data) * (term * function_context_data) -> bool)
+ fst;
+ val copy = I;
+ val extend = I;
+ fun merge _ (tab1, tab2) = Item_Net.merge (tab1, tab2)
+);
+
+val get_function = FunctionData.get o Context.Proof;
+
+
+(* Generally useful?? *)
+fun lift_morphism thy f =
+ let
+ val term = Drule.term_rule thy f
+ in
+ Morphism.thm_morphism f $> Morphism.term_morphism term
+ $> Morphism.typ_morphism (Logic.type_map term)
+ end
+
+fun import_function_data t ctxt =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val ct = cterm_of thy t
+ val inst_morph = lift_morphism thy o Thm.instantiate
+
+ fun match (trm, data) =
+ SOME (morph_function_data data (inst_morph (Thm.match (cterm_of thy trm, ct))))
+ handle Pattern.MATCH => NONE
+ in
+ get_first match (Item_Net.retrieve (get_function ctxt) t)
+ end
+
+fun import_last_function ctxt =
+ case Item_Net.content (get_function ctxt) of
+ [] => NONE
+ | (t, data) :: _ =>
+ let
+ val ([t'], ctxt') = Variable.import_terms true [t] ctxt
+ in
+ import_function_data t' ctxt'
+ end
+
+val all_function_data = Item_Net.content o get_function
+
+fun add_function_data (data as FunctionCtxData {fs, termination, ...}) =
+ FunctionData.map (fold (fn f => Item_Net.insert (f, data)) fs)
+ #> store_termination_rule termination
+
+
+(* Simp rules for termination proofs *)
+
+structure Termination_Simps = Named_Thms
+(
+ val name = "termination_simp"
+ val description = "Simplification rule for termination proofs"
+);
+
+
+(* Default Termination Prover *)
+
+structure TerminationProver = GenericDataFun
+(
+ type T = Proof.context -> Proof.method
+ val empty = (fn _ => error "Termination prover not configured")
+ val extend = I
+ fun merge _ (a,b) = b (* FIXME *)
+);
+
+val set_termination_prover = TerminationProver.put
+val get_termination_prover = TerminationProver.get o Context.Proof
+
+
+(* Configuration management *)
+datatype function_opt
+ = Sequential
+ | Default of string
+ | DomIntros
+ | No_Partials
+ | Tailrec
+
+datatype function_config
+ = FunctionConfig of
+ {
+ sequential: bool,
+ default: string,
+ domintros: bool,
+ partials: bool,
+ tailrec: bool
+ }
+
+fun apply_opt Sequential (FunctionConfig {sequential, default, domintros, partials, tailrec}) =
+ FunctionConfig {sequential=true, default=default, domintros=domintros, partials=partials, tailrec=tailrec}
+ | apply_opt (Default d) (FunctionConfig {sequential, default, domintros, partials, tailrec}) =
+ FunctionConfig {sequential=sequential, default=d, domintros=domintros, partials=partials, tailrec=tailrec}
+ | apply_opt DomIntros (FunctionConfig {sequential, default, domintros, partials, tailrec}) =
+ FunctionConfig {sequential=sequential, default=default, domintros=true, partials=partials, tailrec=tailrec}
+ | apply_opt Tailrec (FunctionConfig {sequential, default, domintros, partials, tailrec}) =
+ FunctionConfig {sequential=sequential, default=default, domintros=domintros, partials=partials, tailrec=true}
+ | apply_opt No_Partials (FunctionConfig {sequential, default, domintros, partials, tailrec}) =
+ FunctionConfig {sequential=sequential, default=default, domintros=domintros, partials=false, tailrec=true}
+
+val default_config =
+ FunctionConfig { sequential=false, default="%x. undefined" (*FIXME dynamic scoping*),
+ domintros=false, partials=true, tailrec=false }
+
+
+(* Analyzing function equations *)
+
+fun split_def ctxt geq =
+ let
+ fun input_error msg = cat_lines [msg, Syntax.string_of_term ctxt geq]
+ val qs = Term.strip_qnt_vars "all" geq
+ val imp = Term.strip_qnt_body "all" geq
+ val (gs, eq) = Logic.strip_horn imp
+
+ val (f_args, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eq)
+ handle TERM _ => error (input_error "Not an equation")
+
+ val (head, args) = strip_comb f_args
+
+ val fname = fst (dest_Free head)
+ handle TERM _ => error (input_error "Head symbol must not be a bound variable")
+ in
+ (fname, qs, gs, args, rhs)
+ end
+
+(* Check for all sorts of errors in the input *)
+fun check_defs ctxt fixes eqs =
+ let
+ val fnames = map (fst o fst) fixes
+
+ fun check geq =
+ let
+ fun input_error msg = error (cat_lines [msg, Syntax.string_of_term ctxt geq])
+
+ val fqgar as (fname, qs, gs, args, rhs) = split_def ctxt geq
+
+ val _ = fname mem fnames
+ orelse input_error
+ ("Head symbol of left hand side must be "
+ ^ plural "" "one out of " fnames ^ commas_quote fnames)
+
+ val _ = length args > 0 orelse input_error "Function has no arguments:"
+
+ fun add_bvs t is = add_loose_bnos (t, 0, is)
+ val rvs = (subtract (op =) (fold add_bvs args []) (add_bvs rhs []))
+ |> map (fst o nth (rev qs))
+
+ val _ = null rvs orelse input_error
+ ("Variable" ^ plural " " "s " rvs ^ commas_quote rvs
+ ^ " occur" ^ plural "s" "" rvs ^ " on right hand side only:")
+
+ val _ = forall (not o Term.exists_subterm
+ (fn Free (n, _) => n mem fnames | _ => false)) (gs @ args)
+ orelse input_error "Defined function may not occur in premises or arguments"
+
+ val freeargs = map (fn t => subst_bounds (rev (map Free qs), t)) args
+ val funvars = filter (fn q => exists (exists_subterm (fn (Free q') $ _ => q = q' | _ => false)) freeargs) qs
+ val _ = null funvars
+ orelse (warning (cat_lines
+ ["Bound variable" ^ plural " " "s " funvars
+ ^ commas_quote (map fst funvars) ^
+ " occur" ^ plural "s" "" funvars ^ " in function position.",
+ "Misspelled constructor???"]); true)
+ in
+ (fname, length args)
+ end
+
+ val _ = AList.group (op =) (map check eqs)
+ |> map (fn (fname, ars) =>
+ length (distinct (op =) ars) = 1
+ orelse error ("Function " ^ quote fname ^
+ " has different numbers of arguments in different equations"))
+
+ fun check_sorts ((fname, fT), _) =
+ Sorts.of_sort (Sign.classes_of (ProofContext.theory_of ctxt)) (fT, HOLogic.typeS)
+ orelse error (cat_lines
+ ["Type of " ^ quote fname ^ " is not of sort " ^ quote "type" ^ ":",
+ setmp_CRITICAL show_sorts true (Syntax.string_of_typ ctxt) fT])
+
+ val _ = map check_sorts fixes
+ in
+ ()
+ end
+
+(* Preprocessors *)
+
+type fixes = ((string * typ) * mixfix) list
+type 'a spec = (Attrib.binding * 'a list) list
+type preproc = function_config -> Proof.context -> fixes -> term spec
+ -> (term list * (thm list -> thm spec) * (thm list -> thm list list) * string list)
+
+val fname_of = fst o dest_Free o fst o strip_comb o fst
+ o HOLogic.dest_eq o HOLogic.dest_Trueprop o Logic.strip_imp_concl o snd o dest_all_all
+
+fun mk_case_names i "" k = mk_case_names i (string_of_int (i + 1)) k
+ | mk_case_names _ n 0 = []
+ | mk_case_names _ n 1 = [n]
+ | mk_case_names _ n k = map (fn i => n ^ "_" ^ string_of_int i) (1 upto k)
+
+fun empty_preproc check _ ctxt fixes spec =
+ let
+ val (bnds, tss) = split_list spec
+ val ts = flat tss
+ val _ = check ctxt fixes ts
+ val fnames = map (fst o fst) fixes
+ val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) ts
+
+ fun sort xs = partition_list (fn i => fn (j,_) => i = j) 0 (length fnames - 1)
+ (indices ~~ xs)
+ |> map (map snd)
+
+ (* using theorem names for case name currently disabled *)
+ val cnames = map_index (fn (i, _) => mk_case_names i "" 1) bnds |> flat
+ in
+ (ts, curry op ~~ bnds o Library.unflat tss, sort, cnames)
+ end
+
+structure Preprocessor = GenericDataFun
+(
+ type T = preproc
+ val empty : T = empty_preproc check_defs
+ val extend = I
+ fun merge _ (a, _) = a
+);
+
+val get_preproc = Preprocessor.get o Context.Proof
+val set_preproc = Preprocessor.map o K
+
+
+
+local
+ structure P = OuterParse and K = OuterKeyword
+
+ val option_parser =
+ P.group "option" ((P.reserved "sequential" >> K Sequential)
+ || ((P.reserved "default" |-- P.term) >> Default)
+ || (P.reserved "domintros" >> K DomIntros)
+ || (P.reserved "no_partials" >> K No_Partials)
+ || (P.reserved "tailrec" >> K Tailrec))
+
+ fun config_parser default =
+ (Scan.optional (P.$$$ "(" |-- P.!!! (P.list1 option_parser) --| P.$$$ ")") [])
+ >> (fn opts => fold apply_opt opts default)
+in
+ fun function_parser default_cfg =
+ config_parser default_cfg -- P.fixes -- SpecParse.where_alt_specs
+end
+
+
+end
+end
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/function_core.ML Sun Oct 25 08:57:55 2009 +0100
@@ -0,0 +1,956 @@
+(* Title: HOL/Tools/Function/function_core.ML
+ Author: Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions:
+Main functionality.
+*)
+
+signature FUNCTION_CORE =
+sig
+ val trace: bool Unsynchronized.ref
+
+ val prepare_function : Function_Common.function_config
+ -> string (* defname *)
+ -> ((bstring * typ) * mixfix) list (* defined symbol *)
+ -> ((bstring * typ) list * term list * term * term) list (* specification *)
+ -> local_theory
+
+ -> (term (* f *)
+ * thm (* goalstate *)
+ * (thm -> Function_Common.function_result) (* continuation *)
+ ) * local_theory
+
+end
+
+structure Function_Core : FUNCTION_CORE =
+struct
+
+val trace = Unsynchronized.ref false;
+fun trace_msg msg = if ! trace then tracing (msg ()) else ();
+
+val boolT = HOLogic.boolT
+val mk_eq = HOLogic.mk_eq
+
+open Function_Lib
+open Function_Common
+
+datatype globals =
+ Globals of {
+ fvar: term,
+ domT: typ,
+ ranT: typ,
+ h: term,
+ y: term,
+ x: term,
+ z: term,
+ a: term,
+ P: term,
+ D: term,
+ Pbool:term
+}
+
+
+datatype rec_call_info =
+ RCInfo of
+ {
+ RIvs: (string * typ) list, (* Call context: fixes and assumes *)
+ CCas: thm list,
+ rcarg: term, (* The recursive argument *)
+
+ llRI: thm,
+ h_assum: term
+ }
+
+
+datatype clause_context =
+ ClauseContext of
+ {
+ ctxt : Proof.context,
+
+ qs : term list,
+ gs : term list,
+ lhs: term,
+ rhs: term,
+
+ cqs: cterm list,
+ ags: thm list,
+ case_hyp : thm
+ }
+
+
+fun transfer_clause_ctx thy (ClauseContext { ctxt, qs, gs, lhs, rhs, cqs, ags, case_hyp }) =
+ ClauseContext { ctxt = ProofContext.transfer thy ctxt,
+ qs = qs, gs = gs, lhs = lhs, rhs = rhs, cqs = cqs, ags = ags, case_hyp = case_hyp }
+
+
+datatype clause_info =
+ ClauseInfo of
+ {
+ no: int,
+ qglr : ((string * typ) list * term list * term * term),
+ cdata : clause_context,
+
+ tree: Function_Ctx_Tree.ctx_tree,
+ lGI: thm,
+ RCs: rec_call_info list
+ }
+
+
+(* Theory dependencies. *)
+val Pair_inject = @{thm Product_Type.Pair_inject};
+
+val acc_induct_rule = @{thm accp_induct_rule};
+
+val ex1_implies_ex = @{thm FunDef.fundef_ex1_existence};
+val ex1_implies_un = @{thm FunDef.fundef_ex1_uniqueness};
+val ex1_implies_iff = @{thm FunDef.fundef_ex1_iff};
+
+val acc_downward = @{thm accp_downward};
+val accI = @{thm accp.accI};
+val case_split = @{thm HOL.case_split};
+val fundef_default_value = @{thm FunDef.fundef_default_value};
+val not_acc_down = @{thm not_accp_down};
+
+
+
+fun find_calls tree =
+ let
+ fun add_Ri (fixes,assumes) (_ $ arg) _ (_, xs) = ([], (fixes, assumes, arg) :: xs)
+ | add_Ri _ _ _ _ = raise Match
+ in
+ rev (Function_Ctx_Tree.traverse_tree add_Ri tree [])
+ end
+
+
+(** building proof obligations *)
+
+fun mk_compat_proof_obligations domT ranT fvar f glrs =
+ let
+ fun mk_impl ((qs, gs, lhs, rhs),(qs', gs', lhs', rhs')) =
+ let
+ val shift = incr_boundvars (length qs')
+ in
+ Logic.mk_implies
+ (HOLogic.mk_Trueprop (HOLogic.eq_const domT $ shift lhs $ lhs'),
+ HOLogic.mk_Trueprop (HOLogic.eq_const ranT $ shift rhs $ rhs'))
+ |> fold_rev (curry Logic.mk_implies) (map shift gs @ gs')
+ |> fold_rev (fn (n,T) => fn b => Term.all T $ Abs(n,T,b)) (qs @ qs')
+ |> curry abstract_over fvar
+ |> curry subst_bound f
+ end
+ in
+ map mk_impl (unordered_pairs glrs)
+ end
+
+
+fun mk_completeness (Globals {x, Pbool, ...}) clauses qglrs =
+ let
+ fun mk_case (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) =
+ HOLogic.mk_Trueprop Pbool
+ |> curry Logic.mk_implies (HOLogic.mk_Trueprop (mk_eq (x, lhs)))
+ |> fold_rev (curry Logic.mk_implies) gs
+ |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
+ in
+ HOLogic.mk_Trueprop Pbool
+ |> fold_rev (curry Logic.mk_implies o mk_case) (clauses ~~ qglrs)
+ |> mk_forall_rename ("x", x)
+ |> mk_forall_rename ("P", Pbool)
+ end
+
+(** making a context with it's own local bindings **)
+
+fun mk_clause_context x ctxt (pre_qs,pre_gs,pre_lhs,pre_rhs) =
+ let
+ val (qs, ctxt') = Variable.variant_fixes (map fst pre_qs) ctxt
+ |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs
+
+ val thy = ProofContext.theory_of ctxt'
+
+ fun inst t = subst_bounds (rev qs, t)
+ val gs = map inst pre_gs
+ val lhs = inst pre_lhs
+ val rhs = inst pre_rhs
+
+ val cqs = map (cterm_of thy) qs
+ val ags = map (assume o cterm_of thy) gs
+
+ val case_hyp = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (x, lhs))))
+ in
+ ClauseContext { ctxt = ctxt', qs = qs, gs = gs, lhs = lhs, rhs = rhs,
+ cqs = cqs, ags = ags, case_hyp = case_hyp }
+ end
+
+
+(* lowlevel term function. FIXME: remove *)
+fun abstract_over_list vs body =
+ let
+ fun abs lev v tm =
+ if v aconv tm then Bound lev
+ else
+ (case tm of
+ Abs (a, T, t) => Abs (a, T, abs (lev + 1) v t)
+ | t $ u => abs lev v t $ abs lev v u
+ | t => t);
+ in
+ fold_index (fn (i, v) => fn t => abs i v t) vs body
+ end
+
+
+
+fun mk_clause_info globals G f no cdata qglr tree RCs GIntro_thm RIntro_thms =
+ let
+ val Globals {h, fvar, x, ...} = globals
+
+ val ClauseContext { ctxt, qs, cqs, ags, ... } = cdata
+ val cert = Thm.cterm_of (ProofContext.theory_of ctxt)
+
+ (* Instantiate the GIntro thm with "f" and import into the clause context. *)
+ val lGI = GIntro_thm
+ |> forall_elim (cert f)
+ |> fold forall_elim cqs
+ |> fold Thm.elim_implies ags
+
+ fun mk_call_info (rcfix, rcassm, rcarg) RI =
+ let
+ val llRI = RI
+ |> fold forall_elim cqs
+ |> fold (forall_elim o cert o Free) rcfix
+ |> fold Thm.elim_implies ags
+ |> fold Thm.elim_implies rcassm
+
+ val h_assum =
+ HOLogic.mk_Trueprop (G $ rcarg $ (h $ rcarg))
+ |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
+ |> fold_rev (Logic.all o Free) rcfix
+ |> Pattern.rewrite_term (ProofContext.theory_of ctxt) [(f, h)] []
+ |> abstract_over_list (rev qs)
+ in
+ RCInfo {RIvs=rcfix, rcarg=rcarg, CCas=rcassm, llRI=llRI, h_assum=h_assum}
+ end
+
+ val RC_infos = map2 mk_call_info RCs RIntro_thms
+ in
+ ClauseInfo
+ {
+ no=no,
+ cdata=cdata,
+ qglr=qglr,
+
+ lGI=lGI,
+ RCs=RC_infos,
+ tree=tree
+ }
+ end
+
+
+
+
+
+
+
+(* replace this by a table later*)
+fun store_compat_thms 0 thms = []
+ | store_compat_thms n thms =
+ let
+ val (thms1, thms2) = chop n thms
+ in
+ (thms1 :: store_compat_thms (n - 1) thms2)
+ end
+
+(* expects i <= j *)
+fun lookup_compat_thm i j cts =
+ nth (nth cts (i - 1)) (j - i)
+
+(* Returns "Gsi, Gsj, lhs_i = lhs_j |-- rhs_j_f = rhs_i_f" *)
+(* if j < i, then turn around *)
+fun get_compat_thm thy cts i j ctxi ctxj =
+ let
+ val ClauseContext {cqs=cqsi,ags=agsi,lhs=lhsi,...} = ctxi
+ val ClauseContext {cqs=cqsj,ags=agsj,lhs=lhsj,...} = ctxj
+
+ val lhsi_eq_lhsj = cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj)))
+ in if j < i then
+ let
+ val compat = lookup_compat_thm j i cts
+ in
+ compat (* "!!qj qi. Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
+ |> fold forall_elim (cqsj @ cqsi) (* "Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
+ |> fold Thm.elim_implies agsj
+ |> fold Thm.elim_implies agsi
+ |> Thm.elim_implies ((assume lhsi_eq_lhsj) RS sym) (* "Gsj, Gsi, lhsi = lhsj |-- rhsj = rhsi" *)
+ end
+ else
+ let
+ val compat = lookup_compat_thm i j cts
+ in
+ compat (* "!!qi qj. Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
+ |> fold forall_elim (cqsi @ cqsj) (* "Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
+ |> fold Thm.elim_implies agsi
+ |> fold Thm.elim_implies agsj
+ |> Thm.elim_implies (assume lhsi_eq_lhsj)
+ |> (fn thm => thm RS sym) (* "Gsi, Gsj, lhsi = lhsj |-- rhsj = rhsi" *)
+ end
+ end
+
+
+
+
+(* Generates the replacement lemma in fully quantified form. *)
+fun mk_replacement_lemma thy h ih_elim clause =
+ let
+ val ClauseInfo {cdata=ClauseContext {qs, lhs, rhs, cqs, ags, case_hyp, ...}, RCs, tree, ...} = clause
+ local open Conv in
+ val ih_conv = arg1_conv o arg_conv o arg_conv
+ end
+
+ val ih_elim_case = Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_elim
+
+ val Ris = map (fn RCInfo {llRI, ...} => llRI) RCs
+ val h_assums = map (fn RCInfo {h_assum, ...} => assume (cterm_of thy (subst_bounds (rev qs, h_assum)))) RCs
+
+ val (eql, _) = Function_Ctx_Tree.rewrite_by_tree thy h ih_elim_case (Ris ~~ h_assums) tree
+
+ val replace_lemma = (eql RS meta_eq_to_obj_eq)
+ |> implies_intr (cprop_of case_hyp)
+ |> fold_rev (implies_intr o cprop_of) h_assums
+ |> fold_rev (implies_intr o cprop_of) ags
+ |> fold_rev forall_intr cqs
+ |> Thm.close_derivation
+ in
+ replace_lemma
+ end
+
+
+fun mk_uniqueness_clause thy globals f compat_store clausei clausej RLj =
+ let
+ val Globals {h, y, x, fvar, ...} = globals
+ val ClauseInfo {no=i, cdata=cctxi as ClauseContext {ctxt=ctxti, lhs=lhsi, case_hyp, ...}, ...} = clausei
+ val ClauseInfo {no=j, qglr=cdescj, RCs=RCsj, ...} = clausej
+
+ val cctxj as ClauseContext {ags = agsj', lhs = lhsj', rhs = rhsj', qs = qsj', cqs = cqsj', ...}
+ = mk_clause_context x ctxti cdescj
+
+ val rhsj'h = Pattern.rewrite_term thy [(fvar,h)] [] rhsj'
+ val compat = get_compat_thm thy compat_store i j cctxi cctxj
+ val Ghsj' = map (fn RCInfo {h_assum, ...} => assume (cterm_of thy (subst_bounds (rev qsj', h_assum)))) RCsj
+
+ val RLj_import =
+ RLj |> fold forall_elim cqsj'
+ |> fold Thm.elim_implies agsj'
+ |> fold Thm.elim_implies Ghsj'
+
+ val y_eq_rhsj'h = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h))))
+ val lhsi_eq_lhsj' = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj')))) (* lhs_i = lhs_j' |-- lhs_i = lhs_j' *)
+ in
+ (trans OF [case_hyp, lhsi_eq_lhsj']) (* lhs_i = lhs_j' |-- x = lhs_j' *)
+ |> implies_elim RLj_import (* Rj1' ... Rjk', lhs_i = lhs_j' |-- rhs_j'_h = rhs_j'_f *)
+ |> (fn it => trans OF [it, compat]) (* lhs_i = lhs_j', Gj', Rj1' ... Rjk' |-- rhs_j'_h = rhs_i_f *)
+ |> (fn it => trans OF [y_eq_rhsj'h, it]) (* lhs_i = lhs_j', Gj', Rj1' ... Rjk', y = rhs_j_h' |-- y = rhs_i_f *)
+ |> fold_rev (implies_intr o cprop_of) Ghsj'
+ |> fold_rev (implies_intr o cprop_of) agsj' (* lhs_i = lhs_j' , y = rhs_j_h' |-- Gj', Rj1'...Rjk' ==> y = rhs_i_f *)
+ |> implies_intr (cprop_of y_eq_rhsj'h)
+ |> implies_intr (cprop_of lhsi_eq_lhsj')
+ |> fold_rev forall_intr (cterm_of thy h :: cqsj')
+ end
+
+
+
+fun mk_uniqueness_case ctxt thy globals G f ihyp ih_intro G_cases compat_store clauses rep_lemmas clausei =
+ let
+ val Globals {x, y, ranT, fvar, ...} = globals
+ val ClauseInfo {cdata = ClauseContext {lhs, rhs, qs, cqs, ags, case_hyp, ...}, lGI, RCs, ...} = clausei
+ val rhsC = Pattern.rewrite_term thy [(fvar, f)] [] rhs
+
+ val ih_intro_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ih_intro
+
+ fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) = (llRI RS ih_intro_case)
+ |> fold_rev (implies_intr o cprop_of) CCas
+ |> fold_rev (forall_intr o cterm_of thy o Free) RIvs
+
+ val existence = fold (curry op COMP o prep_RC) RCs lGI
+
+ val P = cterm_of thy (mk_eq (y, rhsC))
+ val G_lhs_y = assume (cterm_of thy (HOLogic.mk_Trueprop (G $ lhs $ y)))
+
+ val unique_clauses = map2 (mk_uniqueness_clause thy globals f compat_store clausei) clauses rep_lemmas
+
+ val uniqueness = G_cases
+ |> forall_elim (cterm_of thy lhs)
+ |> forall_elim (cterm_of thy y)
+ |> forall_elim P
+ |> Thm.elim_implies G_lhs_y
+ |> fold Thm.elim_implies unique_clauses
+ |> implies_intr (cprop_of G_lhs_y)
+ |> forall_intr (cterm_of thy y)
+
+ val P2 = cterm_of thy (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *)
+
+ val exactly_one =
+ ex1I |> instantiate' [SOME (ctyp_of thy ranT)] [SOME P2, SOME (cterm_of thy rhsC)]
+ |> curry (op COMP) existence
+ |> curry (op COMP) uniqueness
+ |> simplify (HOL_basic_ss addsimps [case_hyp RS sym])
+ |> implies_intr (cprop_of case_hyp)
+ |> fold_rev (implies_intr o cprop_of) ags
+ |> fold_rev forall_intr cqs
+
+ val function_value =
+ existence
+ |> implies_intr ihyp
+ |> implies_intr (cprop_of case_hyp)
+ |> forall_intr (cterm_of thy x)
+ |> forall_elim (cterm_of thy lhs)
+ |> curry (op RS) refl
+ in
+ (exactly_one, function_value)
+ end
+
+
+
+
+fun prove_stuff ctxt globals G f R clauses complete compat compat_store G_elim f_def =
+ let
+ val Globals {h, domT, ranT, x, ...} = globals
+ val thy = ProofContext.theory_of ctxt
+
+ (* Inductive Hypothesis: !!z. (z,x):R ==> EX!y. (z,y):G *)
+ val ihyp = Term.all domT $ Abs ("z", domT,
+ Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
+ HOLogic.mk_Trueprop (Const ("Ex1", (ranT --> boolT) --> boolT) $
+ Abs ("y", ranT, G $ Bound 1 $ Bound 0))))
+ |> cterm_of thy
+
+ val ihyp_thm = assume ihyp |> Thm.forall_elim_vars 0
+ val ih_intro = ihyp_thm RS (f_def RS ex1_implies_ex)
+ val ih_elim = ihyp_thm RS (f_def RS ex1_implies_un)
+ |> instantiate' [] [NONE, SOME (cterm_of thy h)]
+
+ val _ = trace_msg (K "Proving Replacement lemmas...")
+ val repLemmas = map (mk_replacement_lemma thy h ih_elim) clauses
+
+ val _ = trace_msg (K "Proving cases for unique existence...")
+ val (ex1s, values) =
+ split_list (map (mk_uniqueness_case ctxt thy globals G f ihyp ih_intro G_elim compat_store clauses repLemmas) clauses)
+
+ val _ = trace_msg (K "Proving: Graph is a function")
+ val graph_is_function = complete
+ |> Thm.forall_elim_vars 0
+ |> fold (curry op COMP) ex1s
+ |> implies_intr (ihyp)
+ |> implies_intr (cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ x)))
+ |> forall_intr (cterm_of thy x)
+ |> (fn it => Drule.compose_single (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)
+ |> (fn it => fold (forall_intr o cterm_of thy o Var) (Term.add_vars (prop_of it) []) it)
+
+ val goalstate = Conjunction.intr graph_is_function complete
+ |> Thm.close_derivation
+ |> Goal.protect
+ |> fold_rev (implies_intr o cprop_of) compat
+ |> implies_intr (cprop_of complete)
+ in
+ (goalstate, values)
+ end
+
+
+fun define_graph Gname fvar domT ranT clauses RCss lthy =
+ let
+ val GT = domT --> ranT --> boolT
+ val Gvar = Free (the_single (Variable.variant_frees lthy [] [(Gname, GT)]))
+
+ fun mk_GIntro (ClauseContext {qs, gs, lhs, rhs, ...}) RCs =
+ let
+ fun mk_h_assm (rcfix, rcassm, rcarg) =
+ HOLogic.mk_Trueprop (Gvar $ rcarg $ (fvar $ rcarg))
+ |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
+ |> fold_rev (Logic.all o Free) rcfix
+ in
+ HOLogic.mk_Trueprop (Gvar $ lhs $ rhs)
+ |> fold_rev (curry Logic.mk_implies o mk_h_assm) RCs
+ |> fold_rev (curry Logic.mk_implies) gs
+ |> fold_rev Logic.all (fvar :: qs)
+ end
+
+ val G_intros = map2 mk_GIntro clauses RCss
+
+ val (GIntro_thms, (G, G_elim, G_induct, lthy)) =
+ Function_Inductive_Wrap.inductive_def G_intros ((dest_Free Gvar, NoSyn), lthy)
+ in
+ ((G, GIntro_thms, G_elim, G_induct), lthy)
+ end
+
+
+
+fun define_function fdefname (fname, mixfix) domT ranT G default lthy =
+ let
+ val f_def =
+ Abs ("x", domT, Const (@{const_name FunDef.THE_default}, ranT --> (ranT --> boolT) --> ranT) $ (default $ Bound 0) $
+ Abs ("y", ranT, G $ Bound 1 $ Bound 0))
+ |> Syntax.check_term lthy
+
+ val ((f, (_, f_defthm)), lthy) =
+ LocalTheory.define Thm.internalK ((Binding.name (function_name fname), mixfix), ((Binding.name fdefname, []), f_def)) lthy
+ in
+ ((f, f_defthm), lthy)
+ end
+
+
+fun define_recursion_relation Rname domT ranT fvar f qglrs clauses RCss lthy =
+ let
+
+ val RT = domT --> domT --> boolT
+ val Rvar = Free (the_single (Variable.variant_frees lthy [] [(Rname, RT)]))
+
+ fun mk_RIntro (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) (rcfix, rcassm, rcarg) =
+ HOLogic.mk_Trueprop (Rvar $ rcarg $ lhs)
+ |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
+ |> fold_rev (curry Logic.mk_implies) gs
+ |> fold_rev (Logic.all o Free) rcfix
+ |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
+ (* "!!qs xs. CS ==> G => (r, lhs) : R" *)
+
+ val R_intross = map2 (map o mk_RIntro) (clauses ~~ qglrs) RCss
+
+ val (RIntro_thmss, (R, R_elim, _, lthy)) =
+ fold_burrow Function_Inductive_Wrap.inductive_def R_intross ((dest_Free Rvar, NoSyn), lthy)
+ in
+ ((R, RIntro_thmss, R_elim), lthy)
+ end
+
+
+fun fix_globals domT ranT fvar ctxt =
+ let
+ val ([h, y, x, z, a, D, P, Pbool],ctxt') =
+ Variable.variant_fixes ["h_fd", "y_fd", "x_fd", "z_fd", "a_fd", "D_fd", "P_fd", "Pb_fd"] ctxt
+ in
+ (Globals {h = Free (h, domT --> ranT),
+ y = Free (y, ranT),
+ x = Free (x, domT),
+ z = Free (z, domT),
+ a = Free (a, domT),
+ D = Free (D, domT --> boolT),
+ P = Free (P, domT --> boolT),
+ Pbool = Free (Pbool, boolT),
+ fvar = fvar,
+ domT = domT,
+ ranT = ranT
+ },
+ ctxt')
+ end
+
+
+
+fun inst_RC thy fvar f (rcfix, rcassm, rcarg) =
+ let
+ fun inst_term t = subst_bound(f, abstract_over (fvar, t))
+ in
+ (rcfix, map (assume o cterm_of thy o inst_term o prop_of) rcassm, inst_term rcarg)
+ end
+
+
+
+(**********************************************************
+ * PROVING THE RULES
+ **********************************************************)
+
+fun mk_psimps thy globals R clauses valthms f_iff graph_is_function =
+ let
+ val Globals {domT, z, ...} = globals
+
+ fun mk_psimp (ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {cqs, lhs, ags, ...}, ...}) valthm =
+ let
+ val lhs_acc = cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *)
+ val z_smaller = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ lhs)) (* "R z lhs" *)
+ in
+ ((assume z_smaller) RS ((assume lhs_acc) RS acc_downward))
+ |> (fn it => it COMP graph_is_function)
+ |> implies_intr z_smaller
+ |> forall_intr (cterm_of thy z)
+ |> (fn it => it COMP valthm)
+ |> implies_intr lhs_acc
+ |> asm_simplify (HOL_basic_ss addsimps [f_iff])
+ |> fold_rev (implies_intr o cprop_of) ags
+ |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
+ end
+ in
+ map2 mk_psimp clauses valthms
+ end
+
+
+(** Induction rule **)
+
+
+val acc_subset_induct = @{thm Orderings.predicate1I} RS @{thm accp_subset_induct}
+
+
+fun mk_partial_induct_rule thy globals R complete_thm clauses =
+ let
+ val Globals {domT, x, z, a, P, D, ...} = globals
+ val acc_R = mk_acc domT R
+
+ val x_D = assume (cterm_of thy (HOLogic.mk_Trueprop (D $ x)))
+ val a_D = cterm_of thy (HOLogic.mk_Trueprop (D $ a))
+
+ val D_subset = cterm_of thy (Logic.all x
+ (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x), HOLogic.mk_Trueprop (acc_R $ x))))
+
+ val D_dcl = (* "!!x z. [| x: D; (z,x):R |] ==> z:D" *)
+ Logic.all x
+ (Logic.all z (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x),
+ Logic.mk_implies (HOLogic.mk_Trueprop (R $ z $ x),
+ HOLogic.mk_Trueprop (D $ z)))))
+ |> cterm_of thy
+
+
+ (* Inductive Hypothesis: !!z. (z,x):R ==> P z *)
+ val ihyp = Term.all domT $ Abs ("z", domT,
+ Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
+ HOLogic.mk_Trueprop (P $ Bound 0)))
+ |> cterm_of thy
+
+ val aihyp = assume ihyp
+
+ fun prove_case clause =
+ let
+ val ClauseInfo {cdata = ClauseContext {ctxt, qs, cqs, ags, gs, lhs, case_hyp, ...}, RCs,
+ qglr = (oqs, _, _, _), ...} = clause
+
+ val case_hyp_conv = K (case_hyp RS eq_reflection)
+ local open Conv in
+ val lhs_D = fconv_rule (arg_conv (arg_conv (case_hyp_conv))) x_D
+ val sih = fconv_rule (More_Conv.binder_conv (K (arg1_conv (arg_conv (arg_conv case_hyp_conv)))) ctxt) aihyp
+ end
+
+ fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) =
+ sih |> forall_elim (cterm_of thy rcarg)
+ |> Thm.elim_implies llRI
+ |> fold_rev (implies_intr o cprop_of) CCas
+ |> fold_rev (forall_intr o cterm_of thy o Free) RIvs
+
+ val P_recs = map mk_Prec RCs (* [P rec1, P rec2, ... ] *)
+
+ val step = HOLogic.mk_Trueprop (P $ lhs)
+ |> fold_rev (curry Logic.mk_implies o prop_of) P_recs
+ |> fold_rev (curry Logic.mk_implies) gs
+ |> curry Logic.mk_implies (HOLogic.mk_Trueprop (D $ lhs))
+ |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
+ |> cterm_of thy
+
+ val P_lhs = assume step
+ |> fold forall_elim cqs
+ |> Thm.elim_implies lhs_D
+ |> fold Thm.elim_implies ags
+ |> fold Thm.elim_implies P_recs
+
+ val res = cterm_of thy (HOLogic.mk_Trueprop (P $ x))
+ |> Conv.arg_conv (Conv.arg_conv case_hyp_conv)
+ |> symmetric (* P lhs == P x *)
+ |> (fn eql => equal_elim eql P_lhs) (* "P x" *)
+ |> implies_intr (cprop_of case_hyp)
+ |> fold_rev (implies_intr o cprop_of) ags
+ |> fold_rev forall_intr cqs
+ in
+ (res, step)
+ end
+
+ val (cases, steps) = split_list (map prove_case clauses)
+
+ val istep = complete_thm
+ |> Thm.forall_elim_vars 0
+ |> fold (curry op COMP) cases (* P x *)
+ |> implies_intr ihyp
+ |> implies_intr (cprop_of x_D)
+ |> forall_intr (cterm_of thy x)
+
+ val subset_induct_rule =
+ acc_subset_induct
+ |> (curry op COMP) (assume D_subset)
+ |> (curry op COMP) (assume D_dcl)
+ |> (curry op COMP) (assume a_D)
+ |> (curry op COMP) istep
+ |> fold_rev implies_intr steps
+ |> implies_intr a_D
+ |> implies_intr D_dcl
+ |> implies_intr D_subset
+
+ val subset_induct_all = fold_rev (forall_intr o cterm_of thy) [P, a, D] subset_induct_rule
+
+ val simple_induct_rule =
+ subset_induct_rule
+ |> forall_intr (cterm_of thy D)
+ |> forall_elim (cterm_of thy acc_R)
+ |> assume_tac 1 |> Seq.hd
+ |> (curry op COMP) (acc_downward
+ |> (instantiate' [SOME (ctyp_of thy domT)]
+ (map (SOME o cterm_of thy) [R, x, z]))
+ |> forall_intr (cterm_of thy z)
+ |> forall_intr (cterm_of thy x))
+ |> forall_intr (cterm_of thy a)
+ |> forall_intr (cterm_of thy P)
+ in
+ simple_induct_rule
+ end
+
+
+
+(* FIXME: This should probably use fixed goals, to be more reliable and faster *)
+fun mk_domain_intro ctxt (Globals {domT, ...}) R R_cases clause =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val ClauseInfo {cdata = ClauseContext {qs, gs, lhs, rhs, cqs, ...},
+ qglr = (oqs, _, _, _), ...} = clause
+ val goal = HOLogic.mk_Trueprop (mk_acc domT R $ lhs)
+ |> fold_rev (curry Logic.mk_implies) gs
+ |> cterm_of thy
+ in
+ Goal.init goal
+ |> (SINGLE (resolve_tac [accI] 1)) |> the
+ |> (SINGLE (eresolve_tac [Thm.forall_elim_vars 0 R_cases] 1)) |> the
+ |> (SINGLE (auto_tac (clasimpset_of ctxt))) |> the
+ |> Goal.conclude
+ |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
+ end
+
+
+
+(** Termination rule **)
+
+val wf_induct_rule = @{thm Wellfounded.wfP_induct_rule};
+val wf_in_rel = @{thm FunDef.wf_in_rel};
+val in_rel_def = @{thm FunDef.in_rel_def};
+
+fun mk_nest_term_case thy globals R' ihyp clause =
+ let
+ val Globals {x, z, ...} = globals
+ val ClauseInfo {cdata = ClauseContext {qs,cqs,ags,lhs,rhs,case_hyp,...},tree,
+ qglr=(oqs, _, _, _), ...} = clause
+
+ val ih_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ihyp
+
+ fun step (fixes, assumes) (_ $ arg) u (sub,(hyps,thms)) =
+ let
+ val used = map (fn (ctx,thm) => Function_Ctx_Tree.export_thm thy ctx thm) (u @ sub)
+
+ val hyp = HOLogic.mk_Trueprop (R' $ arg $ lhs)
+ |> fold_rev (curry Logic.mk_implies o prop_of) used (* additional hyps *)
+ |> Function_Ctx_Tree.export_term (fixes, assumes)
+ |> fold_rev (curry Logic.mk_implies o prop_of) ags
+ |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
+ |> cterm_of thy
+
+ val thm = assume hyp
+ |> fold forall_elim cqs
+ |> fold Thm.elim_implies ags
+ |> Function_Ctx_Tree.import_thm thy (fixes, assumes)
+ |> fold Thm.elim_implies used (* "(arg, lhs) : R'" *)
+
+ val z_eq_arg = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (z, arg))))
+
+ val acc = thm COMP ih_case
+ val z_acc_local = acc
+ |> Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (K (symmetric (z_eq_arg RS eq_reflection)))))
+
+ val ethm = z_acc_local
+ |> Function_Ctx_Tree.export_thm thy (fixes,
+ z_eq_arg :: case_hyp :: ags @ assumes)
+ |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
+
+ val sub' = sub @ [(([],[]), acc)]
+ in
+ (sub', (hyp :: hyps, ethm :: thms))
+ end
+ | step _ _ _ _ = raise Match
+ in
+ Function_Ctx_Tree.traverse_tree step tree
+ end
+
+
+fun mk_nest_term_rule thy globals R R_cases clauses =
+ let
+ val Globals { domT, x, z, ... } = globals
+ val acc_R = mk_acc domT R
+
+ val R' = Free ("R", fastype_of R)
+
+ val Rrel = Free ("R", HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)))
+ val inrel_R = Const (@{const_name FunDef.in_rel}, HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) --> fastype_of R) $ Rrel
+
+ val wfR' = cterm_of thy (HOLogic.mk_Trueprop (Const (@{const_name Wellfounded.wfP}, (domT --> domT --> boolT) --> boolT) $ R')) (* "wf R'" *)
+
+ (* Inductive Hypothesis: !!z. (z,x):R' ==> z : acc R *)
+ val ihyp = Term.all domT $ Abs ("z", domT,
+ Logic.mk_implies (HOLogic.mk_Trueprop (R' $ Bound 0 $ x),
+ HOLogic.mk_Trueprop (acc_R $ Bound 0)))
+ |> cterm_of thy
+
+ val ihyp_a = assume ihyp |> Thm.forall_elim_vars 0
+
+ val R_z_x = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ x))
+
+ val (hyps,cases) = fold (mk_nest_term_case thy globals R' ihyp_a) clauses ([],[])
+ in
+ R_cases
+ |> forall_elim (cterm_of thy z)
+ |> forall_elim (cterm_of thy x)
+ |> forall_elim (cterm_of thy (acc_R $ z))
+ |> curry op COMP (assume R_z_x)
+ |> fold_rev (curry op COMP) cases
+ |> implies_intr R_z_x
+ |> forall_intr (cterm_of thy z)
+ |> (fn it => it COMP accI)
+ |> implies_intr ihyp
+ |> forall_intr (cterm_of thy x)
+ |> (fn it => Drule.compose_single(it,2,wf_induct_rule))
+ |> curry op RS (assume wfR')
+ |> forall_intr_vars
+ |> (fn it => it COMP allI)
+ |> fold implies_intr hyps
+ |> implies_intr wfR'
+ |> forall_intr (cterm_of thy R')
+ |> forall_elim (cterm_of thy (inrel_R))
+ |> curry op RS wf_in_rel
+ |> full_simplify (HOL_basic_ss addsimps [in_rel_def])
+ |> forall_intr (cterm_of thy Rrel)
+ end
+
+
+
+(* Tail recursion (probably very fragile)
+ *
+ * FIXME:
+ * - Need to do forall_elim_vars on psimps: Unneccesary, if psimps would be taken from the same context.
+ * - Must we really replace the fvar by f here?
+ * - Splitting is not configured automatically: Problems with case?
+ *)
+fun mk_trsimps octxt globals f G R f_def R_cases G_induct clauses psimps =
+ let
+ val Globals {domT, ranT, fvar, ...} = globals
+
+ val R_cases = Thm.forall_elim_vars 0 R_cases (* FIXME: Should be already in standard form. *)
+
+ val graph_implies_dom = (* "G ?x ?y ==> dom ?x" *)
+ Goal.prove octxt ["x", "y"] [HOLogic.mk_Trueprop (G $ Free ("x", domT) $ Free ("y", ranT))]
+ (HOLogic.mk_Trueprop (mk_acc domT R $ Free ("x", domT)))
+ (fn {prems=[a], ...} =>
+ ((rtac (G_induct OF [a]))
+ THEN_ALL_NEW (rtac accI)
+ THEN_ALL_NEW (etac R_cases)
+ THEN_ALL_NEW (asm_full_simp_tac (simpset_of octxt))) 1)
+
+ val default_thm = (forall_intr_vars graph_implies_dom) COMP (f_def COMP fundef_default_value)
+
+ fun mk_trsimp clause psimp =
+ let
+ val ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {ctxt, cqs, qs, gs, lhs, rhs, ...}, ...} = clause
+ val thy = ProofContext.theory_of ctxt
+ val rhs_f = Pattern.rewrite_term thy [(fvar, f)] [] rhs
+
+ val trsimp = Logic.list_implies(gs, HOLogic.mk_Trueprop (HOLogic.mk_eq(f $ lhs, rhs_f))) (* "f lhs = rhs" *)
+ val lhs_acc = (mk_acc domT R $ lhs) (* "acc R lhs" *)
+ fun simp_default_tac ss = asm_full_simp_tac (ss addsimps [default_thm, Let_def])
+ in
+ Goal.prove ctxt [] [] trsimp
+ (fn _ =>
+ rtac (instantiate' [] [SOME (cterm_of thy lhs_acc)] case_split) 1
+ THEN (rtac (Thm.forall_elim_vars 0 psimp) THEN_ALL_NEW assume_tac) 1
+ THEN (simp_default_tac (simpset_of ctxt) 1)
+ THEN (etac not_acc_down 1)
+ THEN ((etac R_cases) THEN_ALL_NEW (simp_default_tac (simpset_of ctxt))) 1)
+ |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
+ end
+ in
+ map2 mk_trsimp clauses psimps
+ end
+
+
+fun prepare_function config defname [((fname, fT), mixfix)] abstract_qglrs lthy =
+ let
+ val FunctionConfig {domintros, tailrec, default=default_str, ...} = config
+
+ val fvar = Free (fname, fT)
+ val domT = domain_type fT
+ val ranT = range_type fT
+
+ val default = Syntax.parse_term lthy default_str
+ |> TypeInfer.constrain fT |> Syntax.check_term lthy
+
+ val (globals, ctxt') = fix_globals domT ranT fvar lthy
+
+ val Globals { x, h, ... } = globals
+
+ val clauses = map (mk_clause_context x ctxt') abstract_qglrs
+
+ val n = length abstract_qglrs
+
+ fun build_tree (ClauseContext { ctxt, rhs, ...}) =
+ Function_Ctx_Tree.mk_tree (fname, fT) h ctxt rhs
+
+ val trees = map build_tree clauses
+ val RCss = map find_calls trees
+
+ val ((G, GIntro_thms, G_elim, G_induct), lthy) =
+ PROFILE "def_graph" (define_graph (graph_name defname) fvar domT ranT clauses RCss) lthy
+
+ val ((f, f_defthm), lthy) =
+ PROFILE "def_fun" (define_function (defname ^ "_sumC_def") (fname, mixfix) domT ranT G default) lthy
+
+ val RCss = map (map (inst_RC (ProofContext.theory_of lthy) fvar f)) RCss
+ val trees = map (Function_Ctx_Tree.inst_tree (ProofContext.theory_of lthy) fvar f) trees
+
+ val ((R, RIntro_thmss, R_elim), lthy) =
+ PROFILE "def_rel" (define_recursion_relation (rel_name defname) domT ranT fvar f abstract_qglrs clauses RCss) lthy
+
+ val (_, lthy) =
+ LocalTheory.abbrev Syntax.mode_default ((Binding.name (dom_name defname), NoSyn), mk_acc domT R) lthy
+
+ val newthy = ProofContext.theory_of lthy
+ val clauses = map (transfer_clause_ctx newthy) clauses
+
+ val cert = cterm_of (ProofContext.theory_of lthy)
+
+ val xclauses = PROFILE "xclauses" (map7 (mk_clause_info globals G f) (1 upto n) clauses abstract_qglrs trees RCss GIntro_thms) RIntro_thmss
+
+ val complete = mk_completeness globals clauses abstract_qglrs |> cert |> assume
+ val compat = mk_compat_proof_obligations domT ranT fvar f abstract_qglrs |> map (cert #> assume)
+
+ val compat_store = store_compat_thms n compat
+
+ val (goalstate, values) = PROFILE "prove_stuff" (prove_stuff lthy globals G f R xclauses complete compat compat_store G_elim) f_defthm
+
+ val mk_trsimps = mk_trsimps lthy globals f G R f_defthm R_elim G_induct xclauses
+
+ fun mk_partial_rules provedgoal =
+ let
+ val newthy = theory_of_thm provedgoal (*FIXME*)
+
+ val (graph_is_function, complete_thm) =
+ provedgoal
+ |> Conjunction.elim
+ |> apfst (Thm.forall_elim_vars 0)
+
+ val f_iff = graph_is_function RS (f_defthm RS ex1_implies_iff)
+
+ val psimps = PROFILE "Proving simplification rules" (mk_psimps newthy globals R xclauses values f_iff) graph_is_function
+
+ val simple_pinduct = PROFILE "Proving partial induction rule"
+ (mk_partial_induct_rule newthy globals R complete_thm) xclauses
+
+
+ val total_intro = PROFILE "Proving nested termination rule" (mk_nest_term_rule newthy globals R R_elim) xclauses
+
+ val dom_intros = if domintros
+ then SOME (PROFILE "Proving domain introduction rules" (map (mk_domain_intro lthy globals R R_elim)) xclauses)
+ else NONE
+ val trsimps = if tailrec then SOME (mk_trsimps psimps) else NONE
+
+ in
+ FunctionResult {fs=[f], G=G, R=R, cases=complete_thm,
+ psimps=psimps, simple_pinducts=[simple_pinduct],
+ termination=total_intro, trsimps=trsimps,
+ domintros=dom_intros}
+ end
+ in
+ ((f, goalstate, mk_partial_rules), lthy)
+ end
+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/function_lib.ML Sun Oct 25 08:57:55 2009 +0100
@@ -0,0 +1,180 @@
+(* Title: HOL/Tools/Function/fundef_lib.ML
+ Author: Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions.
+Some fairly general functions that should probably go somewhere else...
+*)
+
+structure Function_Lib =
+struct
+
+fun map_option f NONE = NONE
+ | map_option f (SOME x) = SOME (f x);
+
+fun fold_option f NONE y = y
+ | fold_option f (SOME x) y = f x y;
+
+fun fold_map_option f NONE y = (NONE, y)
+ | fold_map_option f (SOME x) y = apfst SOME (f x y);
+
+(* Ex: "The variable" ^ plural " is" "s are" vs *)
+fun plural sg pl [x] = sg
+ | plural sg pl _ = pl
+
+(* lambda-abstracts over an arbitrarily nested tuple
+ ==> hologic.ML? *)
+fun tupled_lambda vars t =
+ case vars of
+ (Free v) => lambda (Free v) t
+ | (Var v) => lambda (Var v) t
+ | (Const ("Pair", Type ("fun", [Ta, Type ("fun", [Tb, _])]))) $ us $ vs =>
+ (HOLogic.split_const (Ta,Tb, fastype_of t)) $ (tupled_lambda us (tupled_lambda vs t))
+ | _ => raise Match
+
+
+fun dest_all (Const ("all", _) $ Abs (a as (_,T,_))) =
+ let
+ val (n, body) = Term.dest_abs a
+ in
+ (Free (n, T), body)
+ end
+ | dest_all _ = raise Match
+
+
+(* Removes all quantifiers from a term, replacing bound variables by frees. *)
+fun dest_all_all (t as (Const ("all",_) $ _)) =
+ let
+ val (v,b) = dest_all t
+ val (vs, b') = dest_all_all b
+ in
+ (v :: vs, b')
+ end
+ | dest_all_all t = ([],t)
+
+
+(* FIXME: similar to Variable.focus *)
+fun dest_all_all_ctx ctx (Const ("all", _) $ Abs (a as (n,T,b))) =
+ let
+ val [(n', _)] = Variable.variant_frees ctx [] [(n,T)]
+ val (_, ctx') = ProofContext.add_fixes [(Binding.name n', SOME T, NoSyn)] ctx
+
+ val (n'', body) = Term.dest_abs (n', T, b)
+ val _ = (n' = n'') orelse error "dest_all_ctx"
+ (* Note: We assume that n' does not occur in the body. Otherwise it would be fixed. *)
+
+ val (ctx'', vs, bd) = dest_all_all_ctx ctx' body
+ in
+ (ctx'', (n', T) :: vs, bd)
+ end
+ | dest_all_all_ctx ctx t =
+ (ctx, [], t)
+
+
+fun map3 _ [] [] [] = []
+ | map3 f (x :: xs) (y :: ys) (z :: zs) = f x y z :: map3 f xs ys zs
+ | map3 _ _ _ _ = raise Library.UnequalLengths;
+
+fun map4 _ [] [] [] [] = []
+ | map4 f (x :: xs) (y :: ys) (z :: zs) (u :: us) = f x y z u :: map4 f xs ys zs us
+ | map4 _ _ _ _ _ = raise Library.UnequalLengths;
+
+fun map6 _ [] [] [] [] [] [] = []
+ | map6 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) = f x y z u v w :: map6 f xs ys zs us vs ws
+ | map6 _ _ _ _ _ _ _ = raise Library.UnequalLengths;
+
+fun map7 _ [] [] [] [] [] [] [] = []
+ | map7 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) (b :: bs) = f x y z u v w b :: map7 f xs ys zs us vs ws bs
+ | map7 _ _ _ _ _ _ _ _ = raise Library.UnequalLengths;
+
+
+
+(* forms all "unordered pairs": [1, 2, 3] ==> [(1, 1), (1, 2), (1, 3), (2, 2), (2, 3), (3, 3)] *)
+(* ==> library *)
+fun unordered_pairs [] = []
+ | unordered_pairs (x::xs) = map (pair x) (x::xs) @ unordered_pairs xs
+
+
+(* Replaces Frees by name. Works with loose Bounds. *)
+fun replace_frees assoc =
+ map_aterms (fn c as Free (n, _) => the_default c (AList.lookup (op =) assoc n)
+ | t => t)
+
+
+fun rename_bound n (Q $ Abs(_, T, b)) = (Q $ Abs(n, T, b))
+ | rename_bound n _ = raise Match
+
+fun mk_forall_rename (n, v) =
+ rename_bound n o Logic.all v
+
+fun forall_intr_rename (n, cv) thm =
+ let
+ val allthm = forall_intr cv thm
+ val (_ $ abs) = prop_of allthm
+ in
+ Thm.rename_boundvars abs (Abs (n, dummyT, Term.dummy_pattern dummyT)) allthm
+ end
+
+
+(* Returns the frees in a term in canonical order, excluding the fixes from the context *)
+fun frees_in_term ctxt t =
+ Term.add_frees t []
+ |> filter_out (Variable.is_fixed ctxt o fst)
+ |> rev
+
+
+datatype proof_attempt = Solved of thm | Stuck of thm | Fail
+
+fun try_proof cgoal tac =
+ case SINGLE tac (Goal.init cgoal) of
+ NONE => Fail
+ | SOME st =>
+ if Thm.no_prems st
+ then Solved (Goal.finish (Syntax.init_pretty_global (Thm.theory_of_cterm cgoal)) st)
+ else Stuck st
+
+
+fun dest_binop_list cn (t as (Const (n, _) $ a $ b)) =
+ if cn = n then dest_binop_list cn a @ dest_binop_list cn b else [ t ]
+ | dest_binop_list _ t = [ t ]
+
+
+(* separate two parts in a +-expression:
+ "a + b + c + d + e" --> "(a + b + d) + (c + e)"
+
+ Here, + can be any binary operation that is AC.
+
+ cn - The name of the binop-constructor (e.g. @{const_name Un})
+ ac - the AC rewrite rules for cn
+ is - the list of indices of the expressions that should become the first part
+ (e.g. [0,1,3] in the above example)
+*)
+
+fun regroup_conv neu cn ac is ct =
+ let
+ val mk = HOLogic.mk_binop cn
+ val t = term_of ct
+ val xs = dest_binop_list cn t
+ val js = subtract (op =) is (0 upto (length xs) - 1)
+ val ty = fastype_of t
+ val thy = theory_of_cterm ct
+ in
+ Goal.prove_internal []
+ (cterm_of thy
+ (Logic.mk_equals (t,
+ if is = []
+ then mk (Const (neu, ty), foldr1 mk (map (nth xs) js))
+ else if js = []
+ then mk (foldr1 mk (map (nth xs) is), Const (neu, ty))
+ else mk (foldr1 mk (map (nth xs) is), foldr1 mk (map (nth xs) js)))))
+ (K (rewrite_goals_tac ac
+ THEN rtac Drule.reflexive_thm 1))
+ end
+
+(* instance for unions *)
+fun regroup_union_conv t = regroup_conv @{const_name Set.empty} @{const_name Lattices.sup}
+ (map (fn t => t RS eq_reflection) (@{thms Un_ac} @
+ @{thms Un_empty_right} @
+ @{thms Un_empty_left})) t
+
+
+end
--- a/src/HOL/Tools/Function/fundef.ML Sun Oct 25 08:57:36 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,226 +0,0 @@
-(* Title: HOL/Tools/Function/fundef.ML
- Author: Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions.
-Isar commands.
-*)
-
-signature FUNDEF =
-sig
- val add_fundef : (binding * typ option * mixfix) list
- -> (Attrib.binding * term) list
- -> FundefCommon.fundef_config
- -> local_theory
- -> Proof.state
- val add_fundef_cmd : (binding * string option * mixfix) list
- -> (Attrib.binding * string) list
- -> FundefCommon.fundef_config
- -> local_theory
- -> Proof.state
-
- val termination_proof : term option -> local_theory -> Proof.state
- val termination_proof_cmd : string option -> local_theory -> Proof.state
- val termination : term option -> local_theory -> Proof.state
- val termination_cmd : string option -> local_theory -> Proof.state
-
- val setup : theory -> theory
- val get_congs : Proof.context -> thm list
-end
-
-
-structure Fundef : FUNDEF =
-struct
-
-open FundefLib
-open FundefCommon
-
-val simp_attribs = map (Attrib.internal o K)
- [Simplifier.simp_add,
- Code.add_default_eqn_attribute,
- Nitpick_Simps.add,
- Quickcheck_RecFun_Simps.add]
-
-val psimp_attribs = map (Attrib.internal o K)
- [Simplifier.simp_add,
- Nitpick_Psimps.add]
-
-fun note_theorem ((name, atts), ths) =
- LocalTheory.note Thm.generatedK ((Binding.qualified_name name, atts), ths)
-
-fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_"
-
-fun add_simps fnames post sort extra_qualify label moreatts simps lthy =
- let
- val spec = post simps
- |> map (apfst (apsnd (fn ats => moreatts @ ats)))
- |> map (apfst (apfst extra_qualify))
-
- val (saved_spec_simps, lthy) =
- fold_map (LocalTheory.note Thm.generatedK) spec lthy
-
- val saved_simps = maps snd saved_spec_simps
- val simps_by_f = sort saved_simps
-
- fun add_for_f fname simps =
- note_theorem ((Long_Name.qualify fname label, []), simps) #> snd
- in
- (saved_simps,
- fold2 add_for_f fnames simps_by_f lthy)
- end
-
-fun gen_add_fundef is_external prep default_constraint fixspec eqns config lthy =
- let
- val constrn_fxs = map (fn (b, T, mx) => (b, SOME (the_default default_constraint T), mx))
- val ((fixes0, spec0), ctxt') = prep (constrn_fxs fixspec) eqns lthy
- val fixes = map (apfst (apfst Binding.name_of)) fixes0;
- val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0;
- val (eqs, post, sort_cont, cnames) = FundefCommon.get_preproc lthy config ctxt' fixes spec
-
- val defname = mk_defname fixes
-
- val ((goalstate, cont), lthy) =
- FundefMutual.prepare_fundef_mutual config defname fixes eqs lthy
-
- fun afterqed [[proof]] lthy =
- let
- val FundefResult {fs, R, psimps, trsimps, simple_pinducts, termination,
- domintros, cases, ...} =
- cont (Thm.close_derivation proof)
-
- val fnames = map (fst o fst) fixes
- val qualify = Long_Name.qualify defname
- val addsmps = add_simps fnames post sort_cont
-
- val (((psimps', pinducts'), (_, [termination'])), lthy) =
- lthy
- |> addsmps (Binding.qualify false "partial") "psimps"
- psimp_attribs psimps
- ||> fold_option (snd oo addsmps I "simps" simp_attribs) trsimps
- ||>> note_theorem ((qualify "pinduct",
- [Attrib.internal (K (RuleCases.case_names cnames)),
- Attrib.internal (K (RuleCases.consumes 1)),
- Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts)
- ||>> note_theorem ((qualify "termination", []), [termination])
- ||> (snd o note_theorem ((qualify "cases",
- [Attrib.internal (K (RuleCases.case_names cnames))]), [cases]))
- ||> fold_option (snd oo curry note_theorem (qualify "domintros", [])) domintros
-
- val cdata = FundefCtxData { add_simps=addsmps, case_names=cnames, psimps=psimps',
- pinducts=snd pinducts', termination=termination',
- fs=fs, R=R, defname=defname }
- val _ =
- if not is_external then ()
- else Specification.print_consts lthy (K false) (map fst fixes)
- in
- lthy
- |> LocalTheory.declaration (add_fundef_data o morph_fundef_data cdata)
- end
- in
- lthy
- |> is_external ? LocalTheory.set_group (serial_string ())
- |> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
- |> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd
- end
-
-val add_fundef = gen_add_fundef false Specification.check_spec (TypeInfer.anyT HOLogic.typeS)
-val add_fundef_cmd = gen_add_fundef true Specification.read_spec "_::type"
-
-fun gen_termination_proof prep_term raw_term_opt lthy =
- let
- val term_opt = Option.map (prep_term lthy) raw_term_opt
- val data = the (case term_opt of
- SOME t => (import_fundef_data t lthy
- handle Option.Option =>
- error ("Not a function: " ^ quote (Syntax.string_of_term lthy t)))
- | NONE => (import_last_fundef lthy handle Option.Option => error "Not a function"))
-
- val FundefCtxData { termination, R, add_simps, case_names, psimps,
- pinducts, defname, ...} = data
- val domT = domain_type (fastype_of R)
- val goal = HOLogic.mk_Trueprop
- (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT)))
- fun afterqed [[totality]] lthy =
- let
- val totality = Thm.close_derivation totality
- val remove_domain_condition =
- full_simplify (HOL_basic_ss addsimps [totality, True_implies_equals])
- val tsimps = map remove_domain_condition psimps
- val tinduct = map remove_domain_condition pinducts
- val qualify = Long_Name.qualify defname;
- in
- lthy
- |> add_simps I "simps" simp_attribs tsimps |> snd
- |> note_theorem
- ((qualify "induct",
- [Attrib.internal (K (RuleCases.case_names case_names))]),
- tinduct) |> snd
- end
- in
- lthy
- |> ProofContext.note_thmss ""
- [((Binding.empty, [ContextRules.rule_del]), [([allI], [])])] |> snd
- |> ProofContext.note_thmss ""
- [((Binding.empty, [ContextRules.intro_bang (SOME 1)]), [([allI], [])])] |> snd
- |> ProofContext.note_thmss ""
- [((Binding.name "termination", [ContextRules.intro_bang (SOME 0)]),
- [([Goal.norm_result termination], [])])] |> snd
- |> Proof.theorem_i NONE afterqed [[(goal, [])]]
- end
-
-val termination_proof = gen_termination_proof Syntax.check_term;
-val termination_proof_cmd = gen_termination_proof Syntax.read_term;
-
-fun termination term_opt lthy =
- lthy
- |> LocalTheory.set_group (serial_string ())
- |> termination_proof term_opt;
-
-fun termination_cmd term_opt lthy =
- lthy
- |> LocalTheory.set_group (serial_string ())
- |> termination_proof_cmd term_opt;
-
-
-(* Datatype hook to declare datatype congs as "fundef_congs" *)
-
-
-fun add_case_cong n thy =
- Context.theory_map (FundefCtxTree.map_fundef_congs (Thm.add_thm
- (Datatype.the_info thy n
- |> #case_cong
- |> safe_mk_meta_eq)))
- thy
-
-val setup_case_cong = Datatype.interpretation (K (fold add_case_cong))
-
-
-(* setup *)
-
-val setup =
- Attrib.setup @{binding fundef_cong}
- (Attrib.add_del FundefCtxTree.cong_add FundefCtxTree.cong_del)
- "declaration of congruence rule for function definitions"
- #> setup_case_cong
- #> FundefRelation.setup
- #> FundefCommon.Termination_Simps.setup
-
-val get_congs = FundefCtxTree.get_fundef_congs
-
-
-(* outer syntax *)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ =
- OuterSyntax.local_theory_to_proof "function" "define general recursive functions" K.thy_goal
- (fundef_parser default_config
- >> (fn ((config, fixes), statements) => add_fundef_cmd fixes statements config));
-
-val _ =
- OuterSyntax.local_theory_to_proof "termination" "prove termination of a recursive function" K.thy_goal
- (Scan.option P.term >> termination_cmd);
-
-end;
-
-
-end
--- a/src/HOL/Tools/Function/fundef_common.ML Sun Oct 25 08:57:36 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,343 +0,0 @@
-(* Title: HOL/Tools/Function/fundef_common.ML
- Author: Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions.
-Common definitions and other infrastructure.
-*)
-
-structure FundefCommon =
-struct
-
-local open FundefLib in
-
-(* Profiling *)
-val profile = Unsynchronized.ref false;
-
-fun PROFILE msg = if !profile then timeap_msg msg else I
-
-
-val acc_const_name = @{const_name accp}
-fun mk_acc domT R =
- Const (acc_const_name, (domT --> domT --> HOLogic.boolT) --> domT --> HOLogic.boolT) $ R
-
-val function_name = suffix "C"
-val graph_name = suffix "_graph"
-val rel_name = suffix "_rel"
-val dom_name = suffix "_dom"
-
-(* Termination rules *)
-
-structure TerminationRule = GenericDataFun
-(
- type T = thm list
- val empty = []
- val extend = I
- fun merge _ = Thm.merge_thms
-);
-
-val get_termination_rules = TerminationRule.get
-val store_termination_rule = TerminationRule.map o cons
-val apply_termination_rule = resolve_tac o get_termination_rules o Context.Proof
-
-
-(* Function definition result data *)
-
-datatype fundef_result =
- FundefResult of
- {
- fs: term list,
- G: term,
- R: term,
-
- psimps : thm list,
- trsimps : thm list option,
-
- simple_pinducts : thm list,
- cases : thm,
- termination : thm,
- domintros : thm list option
- }
-
-
-datatype fundef_context_data =
- FundefCtxData of
- {
- defname : string,
-
- (* contains no logical entities: invariant under morphisms *)
- add_simps : (binding -> binding) -> string -> Attrib.src list -> thm list
- -> local_theory -> thm list * local_theory,
- case_names : string list,
-
- fs : term list,
- R : term,
-
- psimps: thm list,
- pinducts: thm list,
- termination: thm
- }
-
-fun morph_fundef_data (FundefCtxData {add_simps, case_names, fs, R,
- psimps, pinducts, termination, defname}) phi =
- let
- val term = Morphism.term phi val thm = Morphism.thm phi val fact = Morphism.fact phi
- val name = Binding.name_of o Morphism.binding phi o Binding.name
- in
- FundefCtxData { add_simps = add_simps, case_names = case_names,
- fs = map term fs, R = term R, psimps = fact psimps,
- pinducts = fact pinducts, termination = thm termination,
- defname = name defname }
- end
-
-structure FundefData = GenericDataFun
-(
- type T = (term * fundef_context_data) Item_Net.T;
- val empty = Item_Net.init
- (op aconv o pairself fst : (term * fundef_context_data) * (term * fundef_context_data) -> bool)
- fst;
- val copy = I;
- val extend = I;
- fun merge _ (tab1, tab2) = Item_Net.merge (tab1, tab2)
-);
-
-val get_fundef = FundefData.get o Context.Proof;
-
-
-(* Generally useful?? *)
-fun lift_morphism thy f =
- let
- val term = Drule.term_rule thy f
- in
- Morphism.thm_morphism f $> Morphism.term_morphism term
- $> Morphism.typ_morphism (Logic.type_map term)
- end
-
-fun import_fundef_data t ctxt =
- let
- val thy = ProofContext.theory_of ctxt
- val ct = cterm_of thy t
- val inst_morph = lift_morphism thy o Thm.instantiate
-
- fun match (trm, data) =
- SOME (morph_fundef_data data (inst_morph (Thm.match (cterm_of thy trm, ct))))
- handle Pattern.MATCH => NONE
- in
- get_first match (Item_Net.retrieve (get_fundef ctxt) t)
- end
-
-fun import_last_fundef ctxt =
- case Item_Net.content (get_fundef ctxt) of
- [] => NONE
- | (t, data) :: _ =>
- let
- val ([t'], ctxt') = Variable.import_terms true [t] ctxt
- in
- import_fundef_data t' ctxt'
- end
-
-val all_fundef_data = Item_Net.content o get_fundef
-
-fun add_fundef_data (data as FundefCtxData {fs, termination, ...}) =
- FundefData.map (fold (fn f => Item_Net.insert (f, data)) fs)
- #> store_termination_rule termination
-
-
-(* Simp rules for termination proofs *)
-
-structure Termination_Simps = Named_Thms
-(
- val name = "termination_simp"
- val description = "Simplification rule for termination proofs"
-);
-
-
-(* Default Termination Prover *)
-
-structure TerminationProver = GenericDataFun
-(
- type T = Proof.context -> Proof.method
- val empty = (fn _ => error "Termination prover not configured")
- val extend = I
- fun merge _ (a,b) = b (* FIXME *)
-);
-
-val set_termination_prover = TerminationProver.put
-val get_termination_prover = TerminationProver.get o Context.Proof
-
-
-(* Configuration management *)
-datatype fundef_opt
- = Sequential
- | Default of string
- | DomIntros
- | Tailrec
-
-datatype fundef_config
- = FundefConfig of
- {
- sequential: bool,
- default: string,
- domintros: bool,
- tailrec: bool
- }
-
-fun apply_opt Sequential (FundefConfig {sequential, default, domintros,tailrec}) =
- FundefConfig {sequential=true, default=default, domintros=domintros, tailrec=tailrec}
- | apply_opt (Default d) (FundefConfig {sequential, default, domintros,tailrec}) =
- FundefConfig {sequential=sequential, default=d, domintros=domintros, tailrec=tailrec}
- | apply_opt DomIntros (FundefConfig {sequential, default, domintros,tailrec}) =
- FundefConfig {sequential=sequential, default=default, domintros=true,tailrec=tailrec}
- | apply_opt Tailrec (FundefConfig {sequential, default, domintros,tailrec}) =
- FundefConfig {sequential=sequential, default=default, domintros=domintros,tailrec=true}
-
-val default_config =
- FundefConfig { sequential=false, default="%x. undefined" (*FIXME dynamic scoping*),
- domintros=false, tailrec=false }
-
-
-(* Analyzing function equations *)
-
-fun split_def ctxt geq =
- let
- fun input_error msg = cat_lines [msg, Syntax.string_of_term ctxt geq]
- val qs = Term.strip_qnt_vars "all" geq
- val imp = Term.strip_qnt_body "all" geq
- val (gs, eq) = Logic.strip_horn imp
-
- val (f_args, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eq)
- handle TERM _ => error (input_error "Not an equation")
-
- val (head, args) = strip_comb f_args
-
- val fname = fst (dest_Free head)
- handle TERM _ => error (input_error "Head symbol must not be a bound variable")
- in
- (fname, qs, gs, args, rhs)
- end
-
-(* Check for all sorts of errors in the input *)
-fun check_defs ctxt fixes eqs =
- let
- val fnames = map (fst o fst) fixes
-
- fun check geq =
- let
- fun input_error msg = error (cat_lines [msg, Syntax.string_of_term ctxt geq])
-
- val fqgar as (fname, qs, gs, args, rhs) = split_def ctxt geq
-
- val _ = fname mem fnames
- orelse input_error
- ("Head symbol of left hand side must be "
- ^ plural "" "one out of " fnames ^ commas_quote fnames)
-
- val _ = length args > 0 orelse input_error "Function has no arguments:"
-
- fun add_bvs t is = add_loose_bnos (t, 0, is)
- val rvs = (subtract (op =) (fold add_bvs args []) (add_bvs rhs []))
- |> map (fst o nth (rev qs))
-
- val _ = null rvs orelse input_error
- ("Variable" ^ plural " " "s " rvs ^ commas_quote rvs
- ^ " occur" ^ plural "s" "" rvs ^ " on right hand side only:")
-
- val _ = forall (not o Term.exists_subterm
- (fn Free (n, _) => n mem fnames | _ => false)) (gs @ args)
- orelse input_error "Defined function may not occur in premises or arguments"
-
- val freeargs = map (fn t => subst_bounds (rev (map Free qs), t)) args
- val funvars = filter (fn q => exists (exists_subterm (fn (Free q') $ _ => q = q' | _ => false)) freeargs) qs
- val _ = null funvars
- orelse (warning (cat_lines
- ["Bound variable" ^ plural " " "s " funvars
- ^ commas_quote (map fst funvars) ^
- " occur" ^ plural "s" "" funvars ^ " in function position.",
- "Misspelled constructor???"]); true)
- in
- (fname, length args)
- end
-
- val _ = AList.group (op =) (map check eqs)
- |> map (fn (fname, ars) =>
- length (distinct (op =) ars) = 1
- orelse error ("Function " ^ quote fname ^
- " has different numbers of arguments in different equations"))
-
- fun check_sorts ((fname, fT), _) =
- Sorts.of_sort (Sign.classes_of (ProofContext.theory_of ctxt)) (fT, HOLogic.typeS)
- orelse error (cat_lines
- ["Type of " ^ quote fname ^ " is not of sort " ^ quote "type" ^ ":",
- setmp_CRITICAL show_sorts true (Syntax.string_of_typ ctxt) fT])
-
- val _ = map check_sorts fixes
- in
- ()
- end
-
-(* Preprocessors *)
-
-type fixes = ((string * typ) * mixfix) list
-type 'a spec = (Attrib.binding * 'a list) list
-type preproc = fundef_config -> Proof.context -> fixes -> term spec
- -> (term list * (thm list -> thm spec) * (thm list -> thm list list) * string list)
-
-val fname_of = fst o dest_Free o fst o strip_comb o fst
- o HOLogic.dest_eq o HOLogic.dest_Trueprop o Logic.strip_imp_concl o snd o dest_all_all
-
-fun mk_case_names i "" k = mk_case_names i (string_of_int (i + 1)) k
- | mk_case_names _ n 0 = []
- | mk_case_names _ n 1 = [n]
- | mk_case_names _ n k = map (fn i => n ^ "_" ^ string_of_int i) (1 upto k)
-
-fun empty_preproc check _ ctxt fixes spec =
- let
- val (bnds, tss) = split_list spec
- val ts = flat tss
- val _ = check ctxt fixes ts
- val fnames = map (fst o fst) fixes
- val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) ts
-
- fun sort xs = partition_list (fn i => fn (j,_) => i = j) 0 (length fnames - 1)
- (indices ~~ xs)
- |> map (map snd)
-
- (* using theorem names for case name currently disabled *)
- val cnames = map_index (fn (i, _) => mk_case_names i "" 1) bnds |> flat
- in
- (ts, curry op ~~ bnds o Library.unflat tss, sort, cnames)
- end
-
-structure Preprocessor = GenericDataFun
-(
- type T = preproc
- val empty : T = empty_preproc check_defs
- val extend = I
- fun merge _ (a, _) = a
-);
-
-val get_preproc = Preprocessor.get o Context.Proof
-val set_preproc = Preprocessor.map o K
-
-
-
-local
- structure P = OuterParse and K = OuterKeyword
-
- val option_parser =
- P.group "option" ((P.reserved "sequential" >> K Sequential)
- || ((P.reserved "default" |-- P.term) >> Default)
- || (P.reserved "domintros" >> K DomIntros)
- || (P.reserved "tailrec" >> K Tailrec))
-
- fun config_parser default =
- (Scan.optional (P.$$$ "(" |-- P.!!! (P.list1 option_parser) --| P.$$$ ")") [])
- >> (fn opts => fold apply_opt opts default)
-in
- fun fundef_parser default_cfg =
- config_parser default_cfg -- P.fixes -- SpecParse.where_alt_specs
-end
-
-
-end
-end
-
--- a/src/HOL/Tools/Function/fundef_core.ML Sun Oct 25 08:57:36 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,956 +0,0 @@
-(* Title: HOL/Tools/Function/fundef_core.ML
- Author: Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions:
-Main functionality.
-*)
-
-signature FUNDEF_CORE =
-sig
- val trace: bool Unsynchronized.ref
-
- val prepare_fundef : FundefCommon.fundef_config
- -> string (* defname *)
- -> ((bstring * typ) * mixfix) list (* defined symbol *)
- -> ((bstring * typ) list * term list * term * term) list (* specification *)
- -> local_theory
-
- -> (term (* f *)
- * thm (* goalstate *)
- * (thm -> FundefCommon.fundef_result) (* continuation *)
- ) * local_theory
-
-end
-
-structure FundefCore : FUNDEF_CORE =
-struct
-
-val trace = Unsynchronized.ref false;
-fun trace_msg msg = if ! trace then tracing (msg ()) else ();
-
-val boolT = HOLogic.boolT
-val mk_eq = HOLogic.mk_eq
-
-open FundefLib
-open FundefCommon
-
-datatype globals =
- Globals of {
- fvar: term,
- domT: typ,
- ranT: typ,
- h: term,
- y: term,
- x: term,
- z: term,
- a: term,
- P: term,
- D: term,
- Pbool:term
-}
-
-
-datatype rec_call_info =
- RCInfo of
- {
- RIvs: (string * typ) list, (* Call context: fixes and assumes *)
- CCas: thm list,
- rcarg: term, (* The recursive argument *)
-
- llRI: thm,
- h_assum: term
- }
-
-
-datatype clause_context =
- ClauseContext of
- {
- ctxt : Proof.context,
-
- qs : term list,
- gs : term list,
- lhs: term,
- rhs: term,
-
- cqs: cterm list,
- ags: thm list,
- case_hyp : thm
- }
-
-
-fun transfer_clause_ctx thy (ClauseContext { ctxt, qs, gs, lhs, rhs, cqs, ags, case_hyp }) =
- ClauseContext { ctxt = ProofContext.transfer thy ctxt,
- qs = qs, gs = gs, lhs = lhs, rhs = rhs, cqs = cqs, ags = ags, case_hyp = case_hyp }
-
-
-datatype clause_info =
- ClauseInfo of
- {
- no: int,
- qglr : ((string * typ) list * term list * term * term),
- cdata : clause_context,
-
- tree: FundefCtxTree.ctx_tree,
- lGI: thm,
- RCs: rec_call_info list
- }
-
-
-(* Theory dependencies. *)
-val Pair_inject = @{thm Product_Type.Pair_inject};
-
-val acc_induct_rule = @{thm accp_induct_rule};
-
-val ex1_implies_ex = @{thm FunDef.fundef_ex1_existence};
-val ex1_implies_un = @{thm FunDef.fundef_ex1_uniqueness};
-val ex1_implies_iff = @{thm FunDef.fundef_ex1_iff};
-
-val acc_downward = @{thm accp_downward};
-val accI = @{thm accp.accI};
-val case_split = @{thm HOL.case_split};
-val fundef_default_value = @{thm FunDef.fundef_default_value};
-val not_acc_down = @{thm not_accp_down};
-
-
-
-fun find_calls tree =
- let
- fun add_Ri (fixes,assumes) (_ $ arg) _ (_, xs) = ([], (fixes, assumes, arg) :: xs)
- | add_Ri _ _ _ _ = raise Match
- in
- rev (FundefCtxTree.traverse_tree add_Ri tree [])
- end
-
-
-(** building proof obligations *)
-
-fun mk_compat_proof_obligations domT ranT fvar f glrs =
- let
- fun mk_impl ((qs, gs, lhs, rhs),(qs', gs', lhs', rhs')) =
- let
- val shift = incr_boundvars (length qs')
- in
- Logic.mk_implies
- (HOLogic.mk_Trueprop (HOLogic.eq_const domT $ shift lhs $ lhs'),
- HOLogic.mk_Trueprop (HOLogic.eq_const ranT $ shift rhs $ rhs'))
- |> fold_rev (curry Logic.mk_implies) (map shift gs @ gs')
- |> fold_rev (fn (n,T) => fn b => Term.all T $ Abs(n,T,b)) (qs @ qs')
- |> curry abstract_over fvar
- |> curry subst_bound f
- end
- in
- map mk_impl (unordered_pairs glrs)
- end
-
-
-fun mk_completeness (Globals {x, Pbool, ...}) clauses qglrs =
- let
- fun mk_case (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) =
- HOLogic.mk_Trueprop Pbool
- |> curry Logic.mk_implies (HOLogic.mk_Trueprop (mk_eq (x, lhs)))
- |> fold_rev (curry Logic.mk_implies) gs
- |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
- in
- HOLogic.mk_Trueprop Pbool
- |> fold_rev (curry Logic.mk_implies o mk_case) (clauses ~~ qglrs)
- |> mk_forall_rename ("x", x)
- |> mk_forall_rename ("P", Pbool)
- end
-
-(** making a context with it's own local bindings **)
-
-fun mk_clause_context x ctxt (pre_qs,pre_gs,pre_lhs,pre_rhs) =
- let
- val (qs, ctxt') = Variable.variant_fixes (map fst pre_qs) ctxt
- |>> map2 (fn (_, T) => fn n => Free (n, T)) pre_qs
-
- val thy = ProofContext.theory_of ctxt'
-
- fun inst t = subst_bounds (rev qs, t)
- val gs = map inst pre_gs
- val lhs = inst pre_lhs
- val rhs = inst pre_rhs
-
- val cqs = map (cterm_of thy) qs
- val ags = map (assume o cterm_of thy) gs
-
- val case_hyp = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (x, lhs))))
- in
- ClauseContext { ctxt = ctxt', qs = qs, gs = gs, lhs = lhs, rhs = rhs,
- cqs = cqs, ags = ags, case_hyp = case_hyp }
- end
-
-
-(* lowlevel term function. FIXME: remove *)
-fun abstract_over_list vs body =
- let
- fun abs lev v tm =
- if v aconv tm then Bound lev
- else
- (case tm of
- Abs (a, T, t) => Abs (a, T, abs (lev + 1) v t)
- | t $ u => abs lev v t $ abs lev v u
- | t => t);
- in
- fold_index (fn (i, v) => fn t => abs i v t) vs body
- end
-
-
-
-fun mk_clause_info globals G f no cdata qglr tree RCs GIntro_thm RIntro_thms =
- let
- val Globals {h, fvar, x, ...} = globals
-
- val ClauseContext { ctxt, qs, cqs, ags, ... } = cdata
- val cert = Thm.cterm_of (ProofContext.theory_of ctxt)
-
- (* Instantiate the GIntro thm with "f" and import into the clause context. *)
- val lGI = GIntro_thm
- |> forall_elim (cert f)
- |> fold forall_elim cqs
- |> fold Thm.elim_implies ags
-
- fun mk_call_info (rcfix, rcassm, rcarg) RI =
- let
- val llRI = RI
- |> fold forall_elim cqs
- |> fold (forall_elim o cert o Free) rcfix
- |> fold Thm.elim_implies ags
- |> fold Thm.elim_implies rcassm
-
- val h_assum =
- HOLogic.mk_Trueprop (G $ rcarg $ (h $ rcarg))
- |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
- |> fold_rev (Logic.all o Free) rcfix
- |> Pattern.rewrite_term (ProofContext.theory_of ctxt) [(f, h)] []
- |> abstract_over_list (rev qs)
- in
- RCInfo {RIvs=rcfix, rcarg=rcarg, CCas=rcassm, llRI=llRI, h_assum=h_assum}
- end
-
- val RC_infos = map2 mk_call_info RCs RIntro_thms
- in
- ClauseInfo
- {
- no=no,
- cdata=cdata,
- qglr=qglr,
-
- lGI=lGI,
- RCs=RC_infos,
- tree=tree
- }
- end
-
-
-
-
-
-
-
-(* replace this by a table later*)
-fun store_compat_thms 0 thms = []
- | store_compat_thms n thms =
- let
- val (thms1, thms2) = chop n thms
- in
- (thms1 :: store_compat_thms (n - 1) thms2)
- end
-
-(* expects i <= j *)
-fun lookup_compat_thm i j cts =
- nth (nth cts (i - 1)) (j - i)
-
-(* Returns "Gsi, Gsj, lhs_i = lhs_j |-- rhs_j_f = rhs_i_f" *)
-(* if j < i, then turn around *)
-fun get_compat_thm thy cts i j ctxi ctxj =
- let
- val ClauseContext {cqs=cqsi,ags=agsi,lhs=lhsi,...} = ctxi
- val ClauseContext {cqs=cqsj,ags=agsj,lhs=lhsj,...} = ctxj
-
- val lhsi_eq_lhsj = cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj)))
- in if j < i then
- let
- val compat = lookup_compat_thm j i cts
- in
- compat (* "!!qj qi. Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
- |> fold forall_elim (cqsj @ cqsi) (* "Gsj => Gsi => lhsj = lhsi ==> rhsj = rhsi" *)
- |> fold Thm.elim_implies agsj
- |> fold Thm.elim_implies agsi
- |> Thm.elim_implies ((assume lhsi_eq_lhsj) RS sym) (* "Gsj, Gsi, lhsi = lhsj |-- rhsj = rhsi" *)
- end
- else
- let
- val compat = lookup_compat_thm i j cts
- in
- compat (* "!!qi qj. Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
- |> fold forall_elim (cqsi @ cqsj) (* "Gsi => Gsj => lhsi = lhsj ==> rhsi = rhsj" *)
- |> fold Thm.elim_implies agsi
- |> fold Thm.elim_implies agsj
- |> Thm.elim_implies (assume lhsi_eq_lhsj)
- |> (fn thm => thm RS sym) (* "Gsi, Gsj, lhsi = lhsj |-- rhsj = rhsi" *)
- end
- end
-
-
-
-
-(* Generates the replacement lemma in fully quantified form. *)
-fun mk_replacement_lemma thy h ih_elim clause =
- let
- val ClauseInfo {cdata=ClauseContext {qs, lhs, rhs, cqs, ags, case_hyp, ...}, RCs, tree, ...} = clause
- local open Conv in
- val ih_conv = arg1_conv o arg_conv o arg_conv
- end
-
- val ih_elim_case = Conv.fconv_rule (ih_conv (K (case_hyp RS eq_reflection))) ih_elim
-
- val Ris = map (fn RCInfo {llRI, ...} => llRI) RCs
- val h_assums = map (fn RCInfo {h_assum, ...} => assume (cterm_of thy (subst_bounds (rev qs, h_assum)))) RCs
-
- val (eql, _) = FundefCtxTree.rewrite_by_tree thy h ih_elim_case (Ris ~~ h_assums) tree
-
- val replace_lemma = (eql RS meta_eq_to_obj_eq)
- |> implies_intr (cprop_of case_hyp)
- |> fold_rev (implies_intr o cprop_of) h_assums
- |> fold_rev (implies_intr o cprop_of) ags
- |> fold_rev forall_intr cqs
- |> Thm.close_derivation
- in
- replace_lemma
- end
-
-
-fun mk_uniqueness_clause thy globals f compat_store clausei clausej RLj =
- let
- val Globals {h, y, x, fvar, ...} = globals
- val ClauseInfo {no=i, cdata=cctxi as ClauseContext {ctxt=ctxti, lhs=lhsi, case_hyp, ...}, ...} = clausei
- val ClauseInfo {no=j, qglr=cdescj, RCs=RCsj, ...} = clausej
-
- val cctxj as ClauseContext {ags = agsj', lhs = lhsj', rhs = rhsj', qs = qsj', cqs = cqsj', ...}
- = mk_clause_context x ctxti cdescj
-
- val rhsj'h = Pattern.rewrite_term thy [(fvar,h)] [] rhsj'
- val compat = get_compat_thm thy compat_store i j cctxi cctxj
- val Ghsj' = map (fn RCInfo {h_assum, ...} => assume (cterm_of thy (subst_bounds (rev qsj', h_assum)))) RCsj
-
- val RLj_import =
- RLj |> fold forall_elim cqsj'
- |> fold Thm.elim_implies agsj'
- |> fold Thm.elim_implies Ghsj'
-
- val y_eq_rhsj'h = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (y, rhsj'h))))
- val lhsi_eq_lhsj' = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (lhsi, lhsj')))) (* lhs_i = lhs_j' |-- lhs_i = lhs_j' *)
- in
- (trans OF [case_hyp, lhsi_eq_lhsj']) (* lhs_i = lhs_j' |-- x = lhs_j' *)
- |> implies_elim RLj_import (* Rj1' ... Rjk', lhs_i = lhs_j' |-- rhs_j'_h = rhs_j'_f *)
- |> (fn it => trans OF [it, compat]) (* lhs_i = lhs_j', Gj', Rj1' ... Rjk' |-- rhs_j'_h = rhs_i_f *)
- |> (fn it => trans OF [y_eq_rhsj'h, it]) (* lhs_i = lhs_j', Gj', Rj1' ... Rjk', y = rhs_j_h' |-- y = rhs_i_f *)
- |> fold_rev (implies_intr o cprop_of) Ghsj'
- |> fold_rev (implies_intr o cprop_of) agsj' (* lhs_i = lhs_j' , y = rhs_j_h' |-- Gj', Rj1'...Rjk' ==> y = rhs_i_f *)
- |> implies_intr (cprop_of y_eq_rhsj'h)
- |> implies_intr (cprop_of lhsi_eq_lhsj')
- |> fold_rev forall_intr (cterm_of thy h :: cqsj')
- end
-
-
-
-fun mk_uniqueness_case ctxt thy globals G f ihyp ih_intro G_cases compat_store clauses rep_lemmas clausei =
- let
- val Globals {x, y, ranT, fvar, ...} = globals
- val ClauseInfo {cdata = ClauseContext {lhs, rhs, qs, cqs, ags, case_hyp, ...}, lGI, RCs, ...} = clausei
- val rhsC = Pattern.rewrite_term thy [(fvar, f)] [] rhs
-
- val ih_intro_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ih_intro
-
- fun prep_RC (RCInfo {llRI, RIvs, CCas, ...}) = (llRI RS ih_intro_case)
- |> fold_rev (implies_intr o cprop_of) CCas
- |> fold_rev (forall_intr o cterm_of thy o Free) RIvs
-
- val existence = fold (curry op COMP o prep_RC) RCs lGI
-
- val P = cterm_of thy (mk_eq (y, rhsC))
- val G_lhs_y = assume (cterm_of thy (HOLogic.mk_Trueprop (G $ lhs $ y)))
-
- val unique_clauses = map2 (mk_uniqueness_clause thy globals f compat_store clausei) clauses rep_lemmas
-
- val uniqueness = G_cases
- |> forall_elim (cterm_of thy lhs)
- |> forall_elim (cterm_of thy y)
- |> forall_elim P
- |> Thm.elim_implies G_lhs_y
- |> fold Thm.elim_implies unique_clauses
- |> implies_intr (cprop_of G_lhs_y)
- |> forall_intr (cterm_of thy y)
-
- val P2 = cterm_of thy (lambda y (G $ lhs $ y)) (* P2 y := (lhs, y): G *)
-
- val exactly_one =
- ex1I |> instantiate' [SOME (ctyp_of thy ranT)] [SOME P2, SOME (cterm_of thy rhsC)]
- |> curry (op COMP) existence
- |> curry (op COMP) uniqueness
- |> simplify (HOL_basic_ss addsimps [case_hyp RS sym])
- |> implies_intr (cprop_of case_hyp)
- |> fold_rev (implies_intr o cprop_of) ags
- |> fold_rev forall_intr cqs
-
- val function_value =
- existence
- |> implies_intr ihyp
- |> implies_intr (cprop_of case_hyp)
- |> forall_intr (cterm_of thy x)
- |> forall_elim (cterm_of thy lhs)
- |> curry (op RS) refl
- in
- (exactly_one, function_value)
- end
-
-
-
-
-fun prove_stuff ctxt globals G f R clauses complete compat compat_store G_elim f_def =
- let
- val Globals {h, domT, ranT, x, ...} = globals
- val thy = ProofContext.theory_of ctxt
-
- (* Inductive Hypothesis: !!z. (z,x):R ==> EX!y. (z,y):G *)
- val ihyp = Term.all domT $ Abs ("z", domT,
- Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
- HOLogic.mk_Trueprop (Const ("Ex1", (ranT --> boolT) --> boolT) $
- Abs ("y", ranT, G $ Bound 1 $ Bound 0))))
- |> cterm_of thy
-
- val ihyp_thm = assume ihyp |> Thm.forall_elim_vars 0
- val ih_intro = ihyp_thm RS (f_def RS ex1_implies_ex)
- val ih_elim = ihyp_thm RS (f_def RS ex1_implies_un)
- |> instantiate' [] [NONE, SOME (cterm_of thy h)]
-
- val _ = trace_msg (K "Proving Replacement lemmas...")
- val repLemmas = map (mk_replacement_lemma thy h ih_elim) clauses
-
- val _ = trace_msg (K "Proving cases for unique existence...")
- val (ex1s, values) =
- split_list (map (mk_uniqueness_case ctxt thy globals G f ihyp ih_intro G_elim compat_store clauses repLemmas) clauses)
-
- val _ = trace_msg (K "Proving: Graph is a function")
- val graph_is_function = complete
- |> Thm.forall_elim_vars 0
- |> fold (curry op COMP) ex1s
- |> implies_intr (ihyp)
- |> implies_intr (cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ x)))
- |> forall_intr (cterm_of thy x)
- |> (fn it => Drule.compose_single (it, 2, acc_induct_rule)) (* "EX! y. (?x,y):G" *)
- |> (fn it => fold (forall_intr o cterm_of thy o Var) (Term.add_vars (prop_of it) []) it)
-
- val goalstate = Conjunction.intr graph_is_function complete
- |> Thm.close_derivation
- |> Goal.protect
- |> fold_rev (implies_intr o cprop_of) compat
- |> implies_intr (cprop_of complete)
- in
- (goalstate, values)
- end
-
-
-fun define_graph Gname fvar domT ranT clauses RCss lthy =
- let
- val GT = domT --> ranT --> boolT
- val Gvar = Free (the_single (Variable.variant_frees lthy [] [(Gname, GT)]))
-
- fun mk_GIntro (ClauseContext {qs, gs, lhs, rhs, ...}) RCs =
- let
- fun mk_h_assm (rcfix, rcassm, rcarg) =
- HOLogic.mk_Trueprop (Gvar $ rcarg $ (fvar $ rcarg))
- |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
- |> fold_rev (Logic.all o Free) rcfix
- in
- HOLogic.mk_Trueprop (Gvar $ lhs $ rhs)
- |> fold_rev (curry Logic.mk_implies o mk_h_assm) RCs
- |> fold_rev (curry Logic.mk_implies) gs
- |> fold_rev Logic.all (fvar :: qs)
- end
-
- val G_intros = map2 mk_GIntro clauses RCss
-
- val (GIntro_thms, (G, G_elim, G_induct, lthy)) =
- FundefInductiveWrap.inductive_def G_intros ((dest_Free Gvar, NoSyn), lthy)
- in
- ((G, GIntro_thms, G_elim, G_induct), lthy)
- end
-
-
-
-fun define_function fdefname (fname, mixfix) domT ranT G default lthy =
- let
- val f_def =
- Abs ("x", domT, Const (@{const_name FunDef.THE_default}, ranT --> (ranT --> boolT) --> ranT) $ (default $ Bound 0) $
- Abs ("y", ranT, G $ Bound 1 $ Bound 0))
- |> Syntax.check_term lthy
-
- val ((f, (_, f_defthm)), lthy) =
- LocalTheory.define Thm.internalK ((Binding.name (function_name fname), mixfix), ((Binding.name fdefname, []), f_def)) lthy
- in
- ((f, f_defthm), lthy)
- end
-
-
-fun define_recursion_relation Rname domT ranT fvar f qglrs clauses RCss lthy =
- let
-
- val RT = domT --> domT --> boolT
- val Rvar = Free (the_single (Variable.variant_frees lthy [] [(Rname, RT)]))
-
- fun mk_RIntro (ClauseContext {qs, gs, lhs, ...}, (oqs, _, _, _)) (rcfix, rcassm, rcarg) =
- HOLogic.mk_Trueprop (Rvar $ rcarg $ lhs)
- |> fold_rev (curry Logic.mk_implies o prop_of) rcassm
- |> fold_rev (curry Logic.mk_implies) gs
- |> fold_rev (Logic.all o Free) rcfix
- |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
- (* "!!qs xs. CS ==> G => (r, lhs) : R" *)
-
- val R_intross = map2 (map o mk_RIntro) (clauses ~~ qglrs) RCss
-
- val (RIntro_thmss, (R, R_elim, _, lthy)) =
- fold_burrow FundefInductiveWrap.inductive_def R_intross ((dest_Free Rvar, NoSyn), lthy)
- in
- ((R, RIntro_thmss, R_elim), lthy)
- end
-
-
-fun fix_globals domT ranT fvar ctxt =
- let
- val ([h, y, x, z, a, D, P, Pbool],ctxt') =
- Variable.variant_fixes ["h_fd", "y_fd", "x_fd", "z_fd", "a_fd", "D_fd", "P_fd", "Pb_fd"] ctxt
- in
- (Globals {h = Free (h, domT --> ranT),
- y = Free (y, ranT),
- x = Free (x, domT),
- z = Free (z, domT),
- a = Free (a, domT),
- D = Free (D, domT --> boolT),
- P = Free (P, domT --> boolT),
- Pbool = Free (Pbool, boolT),
- fvar = fvar,
- domT = domT,
- ranT = ranT
- },
- ctxt')
- end
-
-
-
-fun inst_RC thy fvar f (rcfix, rcassm, rcarg) =
- let
- fun inst_term t = subst_bound(f, abstract_over (fvar, t))
- in
- (rcfix, map (assume o cterm_of thy o inst_term o prop_of) rcassm, inst_term rcarg)
- end
-
-
-
-(**********************************************************
- * PROVING THE RULES
- **********************************************************)
-
-fun mk_psimps thy globals R clauses valthms f_iff graph_is_function =
- let
- val Globals {domT, z, ...} = globals
-
- fun mk_psimp (ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {cqs, lhs, ags, ...}, ...}) valthm =
- let
- val lhs_acc = cterm_of thy (HOLogic.mk_Trueprop (mk_acc domT R $ lhs)) (* "acc R lhs" *)
- val z_smaller = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ lhs)) (* "R z lhs" *)
- in
- ((assume z_smaller) RS ((assume lhs_acc) RS acc_downward))
- |> (fn it => it COMP graph_is_function)
- |> implies_intr z_smaller
- |> forall_intr (cterm_of thy z)
- |> (fn it => it COMP valthm)
- |> implies_intr lhs_acc
- |> asm_simplify (HOL_basic_ss addsimps [f_iff])
- |> fold_rev (implies_intr o cprop_of) ags
- |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
- end
- in
- map2 mk_psimp clauses valthms
- end
-
-
-(** Induction rule **)
-
-
-val acc_subset_induct = @{thm Orderings.predicate1I} RS @{thm accp_subset_induct}
-
-
-fun mk_partial_induct_rule thy globals R complete_thm clauses =
- let
- val Globals {domT, x, z, a, P, D, ...} = globals
- val acc_R = mk_acc domT R
-
- val x_D = assume (cterm_of thy (HOLogic.mk_Trueprop (D $ x)))
- val a_D = cterm_of thy (HOLogic.mk_Trueprop (D $ a))
-
- val D_subset = cterm_of thy (Logic.all x
- (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x), HOLogic.mk_Trueprop (acc_R $ x))))
-
- val D_dcl = (* "!!x z. [| x: D; (z,x):R |] ==> z:D" *)
- Logic.all x
- (Logic.all z (Logic.mk_implies (HOLogic.mk_Trueprop (D $ x),
- Logic.mk_implies (HOLogic.mk_Trueprop (R $ z $ x),
- HOLogic.mk_Trueprop (D $ z)))))
- |> cterm_of thy
-
-
- (* Inductive Hypothesis: !!z. (z,x):R ==> P z *)
- val ihyp = Term.all domT $ Abs ("z", domT,
- Logic.mk_implies (HOLogic.mk_Trueprop (R $ Bound 0 $ x),
- HOLogic.mk_Trueprop (P $ Bound 0)))
- |> cterm_of thy
-
- val aihyp = assume ihyp
-
- fun prove_case clause =
- let
- val ClauseInfo {cdata = ClauseContext {ctxt, qs, cqs, ags, gs, lhs, case_hyp, ...}, RCs,
- qglr = (oqs, _, _, _), ...} = clause
-
- val case_hyp_conv = K (case_hyp RS eq_reflection)
- local open Conv in
- val lhs_D = fconv_rule (arg_conv (arg_conv (case_hyp_conv))) x_D
- val sih = fconv_rule (More_Conv.binder_conv (K (arg1_conv (arg_conv (arg_conv case_hyp_conv)))) ctxt) aihyp
- end
-
- fun mk_Prec (RCInfo {llRI, RIvs, CCas, rcarg, ...}) =
- sih |> forall_elim (cterm_of thy rcarg)
- |> Thm.elim_implies llRI
- |> fold_rev (implies_intr o cprop_of) CCas
- |> fold_rev (forall_intr o cterm_of thy o Free) RIvs
-
- val P_recs = map mk_Prec RCs (* [P rec1, P rec2, ... ] *)
-
- val step = HOLogic.mk_Trueprop (P $ lhs)
- |> fold_rev (curry Logic.mk_implies o prop_of) P_recs
- |> fold_rev (curry Logic.mk_implies) gs
- |> curry Logic.mk_implies (HOLogic.mk_Trueprop (D $ lhs))
- |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
- |> cterm_of thy
-
- val P_lhs = assume step
- |> fold forall_elim cqs
- |> Thm.elim_implies lhs_D
- |> fold Thm.elim_implies ags
- |> fold Thm.elim_implies P_recs
-
- val res = cterm_of thy (HOLogic.mk_Trueprop (P $ x))
- |> Conv.arg_conv (Conv.arg_conv case_hyp_conv)
- |> symmetric (* P lhs == P x *)
- |> (fn eql => equal_elim eql P_lhs) (* "P x" *)
- |> implies_intr (cprop_of case_hyp)
- |> fold_rev (implies_intr o cprop_of) ags
- |> fold_rev forall_intr cqs
- in
- (res, step)
- end
-
- val (cases, steps) = split_list (map prove_case clauses)
-
- val istep = complete_thm
- |> Thm.forall_elim_vars 0
- |> fold (curry op COMP) cases (* P x *)
- |> implies_intr ihyp
- |> implies_intr (cprop_of x_D)
- |> forall_intr (cterm_of thy x)
-
- val subset_induct_rule =
- acc_subset_induct
- |> (curry op COMP) (assume D_subset)
- |> (curry op COMP) (assume D_dcl)
- |> (curry op COMP) (assume a_D)
- |> (curry op COMP) istep
- |> fold_rev implies_intr steps
- |> implies_intr a_D
- |> implies_intr D_dcl
- |> implies_intr D_subset
-
- val subset_induct_all = fold_rev (forall_intr o cterm_of thy) [P, a, D] subset_induct_rule
-
- val simple_induct_rule =
- subset_induct_rule
- |> forall_intr (cterm_of thy D)
- |> forall_elim (cterm_of thy acc_R)
- |> assume_tac 1 |> Seq.hd
- |> (curry op COMP) (acc_downward
- |> (instantiate' [SOME (ctyp_of thy domT)]
- (map (SOME o cterm_of thy) [R, x, z]))
- |> forall_intr (cterm_of thy z)
- |> forall_intr (cterm_of thy x))
- |> forall_intr (cterm_of thy a)
- |> forall_intr (cterm_of thy P)
- in
- simple_induct_rule
- end
-
-
-
-(* FIXME: This should probably use fixed goals, to be more reliable and faster *)
-fun mk_domain_intro ctxt (Globals {domT, ...}) R R_cases clause =
- let
- val thy = ProofContext.theory_of ctxt
- val ClauseInfo {cdata = ClauseContext {qs, gs, lhs, rhs, cqs, ...},
- qglr = (oqs, _, _, _), ...} = clause
- val goal = HOLogic.mk_Trueprop (mk_acc domT R $ lhs)
- |> fold_rev (curry Logic.mk_implies) gs
- |> cterm_of thy
- in
- Goal.init goal
- |> (SINGLE (resolve_tac [accI] 1)) |> the
- |> (SINGLE (eresolve_tac [Thm.forall_elim_vars 0 R_cases] 1)) |> the
- |> (SINGLE (auto_tac (clasimpset_of ctxt))) |> the
- |> Goal.conclude
- |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
- end
-
-
-
-(** Termination rule **)
-
-val wf_induct_rule = @{thm Wellfounded.wfP_induct_rule};
-val wf_in_rel = @{thm FunDef.wf_in_rel};
-val in_rel_def = @{thm FunDef.in_rel_def};
-
-fun mk_nest_term_case thy globals R' ihyp clause =
- let
- val Globals {x, z, ...} = globals
- val ClauseInfo {cdata = ClauseContext {qs,cqs,ags,lhs,rhs,case_hyp,...},tree,
- qglr=(oqs, _, _, _), ...} = clause
-
- val ih_case = full_simplify (HOL_basic_ss addsimps [case_hyp]) ihyp
-
- fun step (fixes, assumes) (_ $ arg) u (sub,(hyps,thms)) =
- let
- val used = map (fn (ctx,thm) => FundefCtxTree.export_thm thy ctx thm) (u @ sub)
-
- val hyp = HOLogic.mk_Trueprop (R' $ arg $ lhs)
- |> fold_rev (curry Logic.mk_implies o prop_of) used (* additional hyps *)
- |> FundefCtxTree.export_term (fixes, assumes)
- |> fold_rev (curry Logic.mk_implies o prop_of) ags
- |> fold_rev mk_forall_rename (map fst oqs ~~ qs)
- |> cterm_of thy
-
- val thm = assume hyp
- |> fold forall_elim cqs
- |> fold Thm.elim_implies ags
- |> FundefCtxTree.import_thm thy (fixes, assumes)
- |> fold Thm.elim_implies used (* "(arg, lhs) : R'" *)
-
- val z_eq_arg = assume (cterm_of thy (HOLogic.mk_Trueprop (mk_eq (z, arg))))
-
- val acc = thm COMP ih_case
- val z_acc_local = acc
- |> Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (K (symmetric (z_eq_arg RS eq_reflection)))))
-
- val ethm = z_acc_local
- |> FundefCtxTree.export_thm thy (fixes,
- z_eq_arg :: case_hyp :: ags @ assumes)
- |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
-
- val sub' = sub @ [(([],[]), acc)]
- in
- (sub', (hyp :: hyps, ethm :: thms))
- end
- | step _ _ _ _ = raise Match
- in
- FundefCtxTree.traverse_tree step tree
- end
-
-
-fun mk_nest_term_rule thy globals R R_cases clauses =
- let
- val Globals { domT, x, z, ... } = globals
- val acc_R = mk_acc domT R
-
- val R' = Free ("R", fastype_of R)
-
- val Rrel = Free ("R", HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)))
- val inrel_R = Const (@{const_name FunDef.in_rel}, HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) --> fastype_of R) $ Rrel
-
- val wfR' = cterm_of thy (HOLogic.mk_Trueprop (Const (@{const_name Wellfounded.wfP}, (domT --> domT --> boolT) --> boolT) $ R')) (* "wf R'" *)
-
- (* Inductive Hypothesis: !!z. (z,x):R' ==> z : acc R *)
- val ihyp = Term.all domT $ Abs ("z", domT,
- Logic.mk_implies (HOLogic.mk_Trueprop (R' $ Bound 0 $ x),
- HOLogic.mk_Trueprop (acc_R $ Bound 0)))
- |> cterm_of thy
-
- val ihyp_a = assume ihyp |> Thm.forall_elim_vars 0
-
- val R_z_x = cterm_of thy (HOLogic.mk_Trueprop (R $ z $ x))
-
- val (hyps,cases) = fold (mk_nest_term_case thy globals R' ihyp_a) clauses ([],[])
- in
- R_cases
- |> forall_elim (cterm_of thy z)
- |> forall_elim (cterm_of thy x)
- |> forall_elim (cterm_of thy (acc_R $ z))
- |> curry op COMP (assume R_z_x)
- |> fold_rev (curry op COMP) cases
- |> implies_intr R_z_x
- |> forall_intr (cterm_of thy z)
- |> (fn it => it COMP accI)
- |> implies_intr ihyp
- |> forall_intr (cterm_of thy x)
- |> (fn it => Drule.compose_single(it,2,wf_induct_rule))
- |> curry op RS (assume wfR')
- |> forall_intr_vars
- |> (fn it => it COMP allI)
- |> fold implies_intr hyps
- |> implies_intr wfR'
- |> forall_intr (cterm_of thy R')
- |> forall_elim (cterm_of thy (inrel_R))
- |> curry op RS wf_in_rel
- |> full_simplify (HOL_basic_ss addsimps [in_rel_def])
- |> forall_intr (cterm_of thy Rrel)
- end
-
-
-
-(* Tail recursion (probably very fragile)
- *
- * FIXME:
- * - Need to do forall_elim_vars on psimps: Unneccesary, if psimps would be taken from the same context.
- * - Must we really replace the fvar by f here?
- * - Splitting is not configured automatically: Problems with case?
- *)
-fun mk_trsimps octxt globals f G R f_def R_cases G_induct clauses psimps =
- let
- val Globals {domT, ranT, fvar, ...} = globals
-
- val R_cases = Thm.forall_elim_vars 0 R_cases (* FIXME: Should be already in standard form. *)
-
- val graph_implies_dom = (* "G ?x ?y ==> dom ?x" *)
- Goal.prove octxt ["x", "y"] [HOLogic.mk_Trueprop (G $ Free ("x", domT) $ Free ("y", ranT))]
- (HOLogic.mk_Trueprop (mk_acc domT R $ Free ("x", domT)))
- (fn {prems=[a], ...} =>
- ((rtac (G_induct OF [a]))
- THEN_ALL_NEW (rtac accI)
- THEN_ALL_NEW (etac R_cases)
- THEN_ALL_NEW (asm_full_simp_tac (simpset_of octxt))) 1)
-
- val default_thm = (forall_intr_vars graph_implies_dom) COMP (f_def COMP fundef_default_value)
-
- fun mk_trsimp clause psimp =
- let
- val ClauseInfo {qglr = (oqs, _, _, _), cdata = ClauseContext {ctxt, cqs, qs, gs, lhs, rhs, ...}, ...} = clause
- val thy = ProofContext.theory_of ctxt
- val rhs_f = Pattern.rewrite_term thy [(fvar, f)] [] rhs
-
- val trsimp = Logic.list_implies(gs, HOLogic.mk_Trueprop (HOLogic.mk_eq(f $ lhs, rhs_f))) (* "f lhs = rhs" *)
- val lhs_acc = (mk_acc domT R $ lhs) (* "acc R lhs" *)
- fun simp_default_tac ss = asm_full_simp_tac (ss addsimps [default_thm, Let_def])
- in
- Goal.prove ctxt [] [] trsimp
- (fn _ =>
- rtac (instantiate' [] [SOME (cterm_of thy lhs_acc)] case_split) 1
- THEN (rtac (Thm.forall_elim_vars 0 psimp) THEN_ALL_NEW assume_tac) 1
- THEN (simp_default_tac (simpset_of ctxt) 1)
- THEN (etac not_acc_down 1)
- THEN ((etac R_cases) THEN_ALL_NEW (simp_default_tac (simpset_of ctxt))) 1)
- |> fold_rev forall_intr_rename (map fst oqs ~~ cqs)
- end
- in
- map2 mk_trsimp clauses psimps
- end
-
-
-fun prepare_fundef config defname [((fname, fT), mixfix)] abstract_qglrs lthy =
- let
- val FundefConfig {domintros, tailrec, default=default_str, ...} = config
-
- val fvar = Free (fname, fT)
- val domT = domain_type fT
- val ranT = range_type fT
-
- val default = Syntax.parse_term lthy default_str
- |> TypeInfer.constrain fT |> Syntax.check_term lthy
-
- val (globals, ctxt') = fix_globals domT ranT fvar lthy
-
- val Globals { x, h, ... } = globals
-
- val clauses = map (mk_clause_context x ctxt') abstract_qglrs
-
- val n = length abstract_qglrs
-
- fun build_tree (ClauseContext { ctxt, rhs, ...}) =
- FundefCtxTree.mk_tree (fname, fT) h ctxt rhs
-
- val trees = map build_tree clauses
- val RCss = map find_calls trees
-
- val ((G, GIntro_thms, G_elim, G_induct), lthy) =
- PROFILE "def_graph" (define_graph (graph_name defname) fvar domT ranT clauses RCss) lthy
-
- val ((f, f_defthm), lthy) =
- PROFILE "def_fun" (define_function (defname ^ "_sumC_def") (fname, mixfix) domT ranT G default) lthy
-
- val RCss = map (map (inst_RC (ProofContext.theory_of lthy) fvar f)) RCss
- val trees = map (FundefCtxTree.inst_tree (ProofContext.theory_of lthy) fvar f) trees
-
- val ((R, RIntro_thmss, R_elim), lthy) =
- PROFILE "def_rel" (define_recursion_relation (rel_name defname) domT ranT fvar f abstract_qglrs clauses RCss) lthy
-
- val (_, lthy) =
- LocalTheory.abbrev Syntax.mode_default ((Binding.name (dom_name defname), NoSyn), mk_acc domT R) lthy
-
- val newthy = ProofContext.theory_of lthy
- val clauses = map (transfer_clause_ctx newthy) clauses
-
- val cert = cterm_of (ProofContext.theory_of lthy)
-
- val xclauses = PROFILE "xclauses" (map7 (mk_clause_info globals G f) (1 upto n) clauses abstract_qglrs trees RCss GIntro_thms) RIntro_thmss
-
- val complete = mk_completeness globals clauses abstract_qglrs |> cert |> assume
- val compat = mk_compat_proof_obligations domT ranT fvar f abstract_qglrs |> map (cert #> assume)
-
- val compat_store = store_compat_thms n compat
-
- val (goalstate, values) = PROFILE "prove_stuff" (prove_stuff lthy globals G f R xclauses complete compat compat_store G_elim) f_defthm
-
- val mk_trsimps = mk_trsimps lthy globals f G R f_defthm R_elim G_induct xclauses
-
- fun mk_partial_rules provedgoal =
- let
- val newthy = theory_of_thm provedgoal (*FIXME*)
-
- val (graph_is_function, complete_thm) =
- provedgoal
- |> Conjunction.elim
- |> apfst (Thm.forall_elim_vars 0)
-
- val f_iff = graph_is_function RS (f_defthm RS ex1_implies_iff)
-
- val psimps = PROFILE "Proving simplification rules" (mk_psimps newthy globals R xclauses values f_iff) graph_is_function
-
- val simple_pinduct = PROFILE "Proving partial induction rule"
- (mk_partial_induct_rule newthy globals R complete_thm) xclauses
-
-
- val total_intro = PROFILE "Proving nested termination rule" (mk_nest_term_rule newthy globals R R_elim) xclauses
-
- val dom_intros = if domintros
- then SOME (PROFILE "Proving domain introduction rules" (map (mk_domain_intro lthy globals R R_elim)) xclauses)
- else NONE
- val trsimps = if tailrec then SOME (mk_trsimps psimps) else NONE
-
- in
- FundefResult {fs=[f], G=G, R=R, cases=complete_thm,
- psimps=psimps, simple_pinducts=[simple_pinduct],
- termination=total_intro, trsimps=trsimps,
- domintros=dom_intros}
- end
- in
- ((f, goalstate, mk_partial_rules), lthy)
- end
-
-
-end
--- a/src/HOL/Tools/Function/fundef_datatype.ML Sun Oct 25 08:57:36 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,330 +0,0 @@
-(* Title: HOL/Tools/Function/fundef_datatype.ML
- Author: Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions.
-A tactic to prove completeness of datatype patterns.
-*)
-
-signature FUNDEF_DATATYPE =
-sig
- val pat_completeness_tac: Proof.context -> int -> tactic
- val pat_completeness: Proof.context -> Proof.method
- val prove_completeness : theory -> term list -> term -> term list list -> term list list -> thm
-
- val setup : theory -> theory
-
- val add_fun : FundefCommon.fundef_config ->
- (binding * typ option * mixfix) list -> (Attrib.binding * term) list ->
- bool -> local_theory -> Proof.context
- val add_fun_cmd : FundefCommon.fundef_config ->
- (binding * string option * mixfix) list -> (Attrib.binding * string) list ->
- bool -> local_theory -> Proof.context
-end
-
-structure FundefDatatype : FUNDEF_DATATYPE =
-struct
-
-open FundefLib
-open FundefCommon
-
-
-fun check_pats ctxt geq =
- let
- fun err str = error (cat_lines ["Malformed definition:",
- str ^ " not allowed in sequential mode.",
- Syntax.string_of_term ctxt geq])
- val thy = ProofContext.theory_of ctxt
-
- fun check_constr_pattern (Bound _) = ()
- | check_constr_pattern t =
- let
- val (hd, args) = strip_comb t
- in
- (((case Datatype.info_of_constr thy (dest_Const hd) of
- SOME _ => ()
- | NONE => err "Non-constructor pattern")
- handle TERM ("dest_Const", _) => err "Non-constructor patterns");
- map check_constr_pattern args;
- ())
- end
-
- val (fname, qs, gs, args, rhs) = split_def ctxt geq
-
- val _ = if not (null gs) then err "Conditional equations" else ()
- val _ = map check_constr_pattern args
-
- (* just count occurrences to check linearity *)
- val _ = if fold (fold_aterms (fn Bound _ => Integer.add 1 | _ => I)) args 0 > length qs
- then err "Nonlinear patterns" else ()
- in
- ()
- end
-
-
-fun mk_argvar i T = Free ("_av" ^ (string_of_int i), T)
-fun mk_patvar i T = Free ("_pv" ^ (string_of_int i), T)
-
-fun inst_free var inst thm =
- forall_elim inst (forall_intr var thm)
-
-
-fun inst_case_thm thy x P thm =
- let
- val [Pv, xv] = Term.add_vars (prop_of thm) []
- in
- cterm_instantiate [(cterm_of thy (Var xv), cterm_of thy x),
- (cterm_of thy (Var Pv), cterm_of thy P)] thm
- end
-
-
-fun invent_vars constr i =
- let
- val Ts = binder_types (fastype_of constr)
- val j = i + length Ts
- val is = i upto (j - 1)
- val avs = map2 mk_argvar is Ts
- val pvs = map2 mk_patvar is Ts
- in
- (avs, pvs, j)
- end
-
-
-fun filter_pats thy cons pvars [] = []
- | filter_pats thy cons pvars (([], thm) :: pts) = raise Match
- | filter_pats thy cons pvars ((pat :: pats, thm) :: pts) =
- case pat of
- Free _ => let val inst = list_comb (cons, pvars)
- in (inst :: pats, inst_free (cterm_of thy pat) (cterm_of thy inst) thm)
- :: (filter_pats thy cons pvars pts) end
- | _ => if fst (strip_comb pat) = cons
- then (pat :: pats, thm) :: (filter_pats thy cons pvars pts)
- else filter_pats thy cons pvars pts
-
-
-fun inst_constrs_of thy (T as Type (name, _)) =
- map (fn (Cn,CT) =>
- Envir.subst_term_types (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT)))
- (the (Datatype.get_constrs thy name))
- | inst_constrs_of thy _ = raise Match
-
-
-fun transform_pat thy avars c_assum ([] , thm) = raise Match
- | transform_pat thy avars c_assum (pat :: pats, thm) =
- let
- val (_, subps) = strip_comb pat
- val eqs = map (cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (avars ~~ subps)
- val a_eqs = map assume eqs
- val c_eq_pat = simplify (HOL_basic_ss addsimps a_eqs) c_assum
- in
- (subps @ pats, fold_rev implies_intr eqs
- (implies_elim thm c_eq_pat))
- end
-
-
-exception COMPLETENESS
-
-fun constr_case thy P idx (v :: vs) pats cons =
- let
- val (avars, pvars, newidx) = invent_vars cons idx
- val c_hyp = cterm_of thy (HOLogic.mk_Trueprop (HOLogic.mk_eq (v, list_comb (cons, avars))))
- val c_assum = assume c_hyp
- val newpats = map (transform_pat thy avars c_assum) (filter_pats thy cons pvars pats)
- in
- o_alg thy P newidx (avars @ vs) newpats
- |> implies_intr c_hyp
- |> fold_rev (forall_intr o cterm_of thy) avars
- end
- | constr_case _ _ _ _ _ _ = raise Match
-and o_alg thy P idx [] (([], Pthm) :: _) = Pthm
- | o_alg thy P idx (v :: vs) [] = raise COMPLETENESS
- | o_alg thy P idx (v :: vs) pts =
- if forall (is_Free o hd o fst) pts (* Var case *)
- then o_alg thy P idx vs (map (fn (pv :: pats, thm) =>
- (pats, refl RS (inst_free (cterm_of thy pv) (cterm_of thy v) thm))) pts)
- else (* Cons case *)
- let
- val T = fastype_of v
- val (tname, _) = dest_Type T
- val {exhaust=case_thm, ...} = Datatype.the_info thy tname
- val constrs = inst_constrs_of thy T
- val c_cases = map (constr_case thy P idx (v :: vs) pts) constrs
- in
- inst_case_thm thy v P case_thm
- |> fold (curry op COMP) c_cases
- end
- | o_alg _ _ _ _ _ = raise Match
-
-
-fun prove_completeness thy xs P qss patss =
- let
- fun mk_assum qs pats =
- HOLogic.mk_Trueprop P
- |> fold_rev (curry Logic.mk_implies o HOLogic.mk_Trueprop o HOLogic.mk_eq) (xs ~~ pats)
- |> fold_rev Logic.all qs
- |> cterm_of thy
-
- val hyps = map2 mk_assum qss patss
-
- fun inst_hyps hyp qs = fold (forall_elim o cterm_of thy) qs (assume hyp)
-
- val assums = map2 inst_hyps hyps qss
- in
- o_alg thy P 2 xs (patss ~~ assums)
- |> fold_rev implies_intr hyps
- end
-
-
-
-fun pat_completeness_tac ctxt = SUBGOAL (fn (subgoal, i) =>
- let
- val thy = ProofContext.theory_of ctxt
- val (vs, subgf) = dest_all_all subgoal
- val (cases, _ $ thesis) = Logic.strip_horn subgf
- handle Bind => raise COMPLETENESS
-
- fun pat_of assum =
- let
- val (qs, imp) = dest_all_all assum
- val prems = Logic.strip_imp_prems imp
- in
- (qs, map (HOLogic.dest_eq o HOLogic.dest_Trueprop) prems)
- end
-
- val (qss, x_pats) = split_list (map pat_of cases)
- val xs = map fst (hd x_pats)
- handle Empty => raise COMPLETENESS
-
- val patss = map (map snd) x_pats
-
- val complete_thm = prove_completeness thy xs thesis qss patss
- |> fold_rev (forall_intr o cterm_of thy) vs
- in
- PRIMITIVE (fn st => Drule.compose_single(complete_thm, i, st))
- end
- handle COMPLETENESS => no_tac)
-
-
-fun pat_completeness ctxt = SIMPLE_METHOD' (pat_completeness_tac ctxt)
-
-val by_pat_completeness_auto =
- Proof.global_future_terminal_proof
- (Method.Basic pat_completeness,
- SOME (Method.Source_i (Args.src (("HOL.auto", []), Position.none))))
-
-fun termination_by method int =
- Fundef.termination_proof NONE
- #> Proof.global_future_terminal_proof (Method.Basic method, NONE) int
-
-fun mk_catchall fixes arity_of =
- let
- fun mk_eqn ((fname, fT), _) =
- let
- val n = arity_of fname
- val (argTs, rT) = chop n (binder_types fT)
- |> apsnd (fn Ts => Ts ---> body_type fT)
-
- val qs = map Free (Name.invent_list [] "a" n ~~ argTs)
- in
- HOLogic.mk_eq(list_comb (Free (fname, fT), qs),
- Const ("HOL.undefined", rT))
- |> HOLogic.mk_Trueprop
- |> fold_rev Logic.all qs
- end
- in
- map mk_eqn fixes
- end
-
-fun add_catchall ctxt fixes spec =
- let val fqgars = map (split_def ctxt) spec
- val arity_of = map (fn (fname,_,_,args,_) => (fname, length args)) fqgars
- |> AList.lookup (op =) #> the
- in
- spec @ mk_catchall fixes arity_of
- end
-
-fun warn_if_redundant ctxt origs tss =
- let
- fun msg t = "Ignoring redundant equation: " ^ quote (Syntax.string_of_term ctxt t)
-
- val (tss', _) = chop (length origs) tss
- fun check (t, []) = (warning (msg t); [])
- | check (t, s) = s
- in
- (map check (origs ~~ tss'); tss)
- end
-
-
-fun sequential_preproc (config as FundefConfig {sequential, ...}) ctxt fixes spec =
- if sequential then
- let
- val (bnds, eqss) = split_list spec
-
- val eqs = map the_single eqss
-
- val feqs = eqs
- |> tap (check_defs ctxt fixes) (* Standard checks *)
- |> tap (map (check_pats ctxt)) (* More checks for sequential mode *)
-
- val compleqs = add_catchall ctxt fixes feqs (* Completion *)
-
- val spliteqs = warn_if_redundant ctxt feqs
- (FundefSplit.split_all_equations ctxt compleqs)
-
- fun restore_spec thms =
- bnds ~~ Library.take (length bnds, Library.unflat spliteqs thms)
-
- val spliteqs' = flat (Library.take (length bnds, spliteqs))
- val fnames = map (fst o fst) fixes
- val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) spliteqs'
-
- fun sort xs = partition_list (fn i => fn (j,_) => i = j) 0 (length fnames - 1) (indices ~~ xs)
- |> map (map snd)
-
-
- val bnds' = bnds @ replicate (length spliteqs - length bnds) Attrib.empty_binding
-
- (* using theorem names for case name currently disabled *)
- val case_names = map_index (fn (i, (_, es)) => mk_case_names i "" (length es))
- (bnds' ~~ spliteqs)
- |> flat
- in
- (flat spliteqs, restore_spec, sort, case_names)
- end
- else
- FundefCommon.empty_preproc check_defs config ctxt fixes spec
-
-val setup =
- Method.setup @{binding pat_completeness} (Scan.succeed pat_completeness)
- "Completeness prover for datatype patterns"
- #> Context.theory_map (FundefCommon.set_preproc sequential_preproc)
-
-
-val fun_config = FundefConfig { sequential=true, default="%x. undefined" (*FIXME dynamic scoping*),
- domintros=false, tailrec=false }
-
-fun gen_fun add config fixes statements int lthy =
- let val group = serial_string () in
- lthy
- |> LocalTheory.set_group group
- |> add fixes statements config
- |> by_pat_completeness_auto int
- |> LocalTheory.restore
- |> LocalTheory.set_group group
- |> termination_by (FundefCommon.get_termination_prover lthy) int
- end;
-
-val add_fun = gen_fun Fundef.add_fundef
-val add_fun_cmd = gen_fun Fundef.add_fundef_cmd
-
-
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ =
- OuterSyntax.local_theory' "fun" "define general recursive functions (short version)" K.thy_decl
- (fundef_parser fun_config
- >> (fn ((config, fixes), statements) => add_fun_cmd config fixes statements));
-
-end
-
-end
--- a/src/HOL/Tools/Function/fundef_lib.ML Sun Oct 25 08:57:36 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,179 +0,0 @@
-(* Title: HOL/Tools/Function/fundef_lib.ML
- Author: Alexander Krauss, TU Muenchen
-
-A package for general recursive function definitions.
-Some fairly general functions that should probably go somewhere else...
-*)
-
-structure FundefLib = struct
-
-fun map_option f NONE = NONE
- | map_option f (SOME x) = SOME (f x);
-
-fun fold_option f NONE y = y
- | fold_option f (SOME x) y = f x y;
-
-fun fold_map_option f NONE y = (NONE, y)
- | fold_map_option f (SOME x) y = apfst SOME (f x y);
-
-(* Ex: "The variable" ^ plural " is" "s are" vs *)
-fun plural sg pl [x] = sg
- | plural sg pl _ = pl
-
-(* lambda-abstracts over an arbitrarily nested tuple
- ==> hologic.ML? *)
-fun tupled_lambda vars t =
- case vars of
- (Free v) => lambda (Free v) t
- | (Var v) => lambda (Var v) t
- | (Const ("Pair", Type ("fun", [Ta, Type ("fun", [Tb, _])]))) $ us $ vs =>
- (HOLogic.split_const (Ta,Tb, fastype_of t)) $ (tupled_lambda us (tupled_lambda vs t))
- | _ => raise Match
-
-
-fun dest_all (Const ("all", _) $ Abs (a as (_,T,_))) =
- let
- val (n, body) = Term.dest_abs a
- in
- (Free (n, T), body)
- end
- | dest_all _ = raise Match
-
-
-(* Removes all quantifiers from a term, replacing bound variables by frees. *)
-fun dest_all_all (t as (Const ("all",_) $ _)) =
- let
- val (v,b) = dest_all t
- val (vs, b') = dest_all_all b
- in
- (v :: vs, b')
- end
- | dest_all_all t = ([],t)
-
-
-(* FIXME: similar to Variable.focus *)
-fun dest_all_all_ctx ctx (Const ("all", _) $ Abs (a as (n,T,b))) =
- let
- val [(n', _)] = Variable.variant_frees ctx [] [(n,T)]
- val (_, ctx') = ProofContext.add_fixes [(Binding.name n', SOME T, NoSyn)] ctx
-
- val (n'', body) = Term.dest_abs (n', T, b)
- val _ = (n' = n'') orelse error "dest_all_ctx"
- (* Note: We assume that n' does not occur in the body. Otherwise it would be fixed. *)
-
- val (ctx'', vs, bd) = dest_all_all_ctx ctx' body
- in
- (ctx'', (n', T) :: vs, bd)
- end
- | dest_all_all_ctx ctx t =
- (ctx, [], t)
-
-
-fun map3 _ [] [] [] = []
- | map3 f (x :: xs) (y :: ys) (z :: zs) = f x y z :: map3 f xs ys zs
- | map3 _ _ _ _ = raise Library.UnequalLengths;
-
-fun map4 _ [] [] [] [] = []
- | map4 f (x :: xs) (y :: ys) (z :: zs) (u :: us) = f x y z u :: map4 f xs ys zs us
- | map4 _ _ _ _ _ = raise Library.UnequalLengths;
-
-fun map6 _ [] [] [] [] [] [] = []
- | map6 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) = f x y z u v w :: map6 f xs ys zs us vs ws
- | map6 _ _ _ _ _ _ _ = raise Library.UnequalLengths;
-
-fun map7 _ [] [] [] [] [] [] [] = []
- | map7 f (x :: xs) (y :: ys) (z :: zs) (u :: us) (v :: vs) (w :: ws) (b :: bs) = f x y z u v w b :: map7 f xs ys zs us vs ws bs
- | map7 _ _ _ _ _ _ _ _ = raise Library.UnequalLengths;
-
-
-
-(* forms all "unordered pairs": [1, 2, 3] ==> [(1, 1), (1, 2), (1, 3), (2, 2), (2, 3), (3, 3)] *)
-(* ==> library *)
-fun unordered_pairs [] = []
- | unordered_pairs (x::xs) = map (pair x) (x::xs) @ unordered_pairs xs
-
-
-(* Replaces Frees by name. Works with loose Bounds. *)
-fun replace_frees assoc =
- map_aterms (fn c as Free (n, _) => the_default c (AList.lookup (op =) assoc n)
- | t => t)
-
-
-fun rename_bound n (Q $ Abs(_, T, b)) = (Q $ Abs(n, T, b))
- | rename_bound n _ = raise Match
-
-fun mk_forall_rename (n, v) =
- rename_bound n o Logic.all v
-
-fun forall_intr_rename (n, cv) thm =
- let
- val allthm = forall_intr cv thm
- val (_ $ abs) = prop_of allthm
- in
- Thm.rename_boundvars abs (Abs (n, dummyT, Term.dummy_pattern dummyT)) allthm
- end
-
-
-(* Returns the frees in a term in canonical order, excluding the fixes from the context *)
-fun frees_in_term ctxt t =
- Term.add_frees t []
- |> filter_out (Variable.is_fixed ctxt o fst)
- |> rev
-
-
-datatype proof_attempt = Solved of thm | Stuck of thm | Fail
-
-fun try_proof cgoal tac =
- case SINGLE tac (Goal.init cgoal) of
- NONE => Fail
- | SOME st =>
- if Thm.no_prems st
- then Solved (Goal.finish (Syntax.init_pretty_global (Thm.theory_of_cterm cgoal)) st)
- else Stuck st
-
-
-fun dest_binop_list cn (t as (Const (n, _) $ a $ b)) =
- if cn = n then dest_binop_list cn a @ dest_binop_list cn b else [ t ]
- | dest_binop_list _ t = [ t ]
-
-
-(* separate two parts in a +-expression:
- "a + b + c + d + e" --> "(a + b + d) + (c + e)"
-
- Here, + can be any binary operation that is AC.
-
- cn - The name of the binop-constructor (e.g. @{const_name Un})
- ac - the AC rewrite rules for cn
- is - the list of indices of the expressions that should become the first part
- (e.g. [0,1,3] in the above example)
-*)
-
-fun regroup_conv neu cn ac is ct =
- let
- val mk = HOLogic.mk_binop cn
- val t = term_of ct
- val xs = dest_binop_list cn t
- val js = subtract (op =) is (0 upto (length xs) - 1)
- val ty = fastype_of t
- val thy = theory_of_cterm ct
- in
- Goal.prove_internal []
- (cterm_of thy
- (Logic.mk_equals (t,
- if is = []
- then mk (Const (neu, ty), foldr1 mk (map (nth xs) js))
- else if js = []
- then mk (foldr1 mk (map (nth xs) is), Const (neu, ty))
- else mk (foldr1 mk (map (nth xs) is), foldr1 mk (map (nth xs) js)))))
- (K (rewrite_goals_tac ac
- THEN rtac Drule.reflexive_thm 1))
- end
-
-(* instance for unions *)
-fun regroup_union_conv t = regroup_conv @{const_name Set.empty} @{const_name Lattices.sup}
- (map (fn t => t RS eq_reflection) (@{thms Un_ac} @
- @{thms Un_empty_right} @
- @{thms Un_empty_left})) t
-
-
-end
--- a/src/HOL/Tools/Function/induction_scheme.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Function/induction_scheme.ML Sun Oct 25 08:57:55 2009 +0100
@@ -13,10 +13,10 @@
end
-structure InductionScheme : INDUCTION_SCHEME =
+structure Induction_Scheme : INDUCTION_SCHEME =
struct
-open FundefLib
+open Function_Lib
type rec_call_info = int * (string * typ) list * term list * term list
@@ -244,7 +244,7 @@
(* Rule for case splitting along the sum types *)
val xss = map (fn (SchemeBranch { xs, ... }) => map Free xs) branches
val pats = map_index (uncurry inject) xss
- val sum_split_rule = FundefDatatype.prove_completeness thy [x] (P_comp $ x) xss (map single pats)
+ val sum_split_rule = Pat_Completeness.prove_completeness thy [x] (P_comp $ x) xss (map single pats)
fun prove_branch (bidx, (SchemeBranch { P, xs, ws, Cs, ... }, (complete_thm, pat))) =
let
--- a/src/HOL/Tools/Function/inductive_wrap.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Function/inductive_wrap.ML Sun Oct 25 08:57:55 2009 +0100
@@ -6,17 +6,17 @@
the introduction and elimination rules.
*)
-signature FUNDEF_INDUCTIVE_WRAP =
+signature FUNCTION_INDUCTIVE_WRAP =
sig
val inductive_def : term list
-> ((bstring * typ) * mixfix) * local_theory
-> thm list * (term * thm * thm * local_theory)
end
-structure FundefInductiveWrap: FUNDEF_INDUCTIVE_WRAP =
+structure Function_Inductive_Wrap: FUNCTION_INDUCTIVE_WRAP =
struct
-open FundefLib
+open Function_Lib
fun requantify ctxt lfix orig_def thm =
let
--- a/src/HOL/Tools/Function/lexicographic_order.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Function/lexicographic_order.ML Sun Oct 25 08:57:55 2009 +0100
@@ -13,10 +13,10 @@
val setup: theory -> theory
end
-structure LexicographicOrder : LEXICOGRAPHIC_ORDER =
+structure Lexicographic_Order : LEXICOGRAPHIC_ORDER =
struct
-open FundefLib
+open Function_Lib
(** General stuff **)
@@ -58,7 +58,7 @@
fun dest_term (t : term) =
let
- val (vars, prop) = FundefLib.dest_all_all t
+ val (vars, prop) = Function_Lib.dest_all_all t
val (prems, concl) = Logic.strip_horn prop
val (lhs, rhs) = concl
|> HOLogic.dest_Trueprop
@@ -215,9 +215,9 @@
end
fun lexicographic_order_tac ctxt =
- TRY (FundefCommon.apply_termination_rule ctxt 1)
+ TRY (Function_Common.apply_termination_rule ctxt 1)
THEN lex_order_tac ctxt
- (auto_tac (clasimpset_of ctxt addsimps2 FundefCommon.Termination_Simps.get ctxt))
+ (auto_tac (clasimpset_of ctxt addsimps2 Function_Common.Termination_Simps.get ctxt))
val lexicographic_order = SIMPLE_METHOD o lexicographic_order_tac
@@ -225,7 +225,7 @@
Method.setup @{binding lexicographic_order}
(Method.sections clasimp_modifiers >> (K lexicographic_order))
"termination prover for lexicographic orderings"
- #> Context.theory_map (FundefCommon.set_termination_prover lexicographic_order)
+ #> Context.theory_map (Function_Common.set_termination_prover lexicographic_order)
end;
--- a/src/HOL/Tools/Function/mutual.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Function/mutual.ML Sun Oct 25 08:57:55 2009 +0100
@@ -5,29 +5,26 @@
Tools for mutual recursive definitions.
*)
-signature FUNDEF_MUTUAL =
+signature FUNCTION_MUTUAL =
sig
- val prepare_fundef_mutual : FundefCommon.fundef_config
+ val prepare_function_mutual : Function_Common.function_config
-> string (* defname *)
-> ((string * typ) * mixfix) list
-> term list
-> local_theory
-> ((thm (* goalstate *)
- * (thm -> FundefCommon.fundef_result) (* proof continuation *)
+ * (thm -> Function_Common.function_result) (* proof continuation *)
) * local_theory)
end
-structure FundefMutual: FUNDEF_MUTUAL =
+structure Function_Mutual: FUNCTION_MUTUAL =
struct
-open FundefLib
-open FundefCommon
-
-
-
+open Function_Lib
+open Function_Common
type qgar = string * (string * typ) list * term list * term list * term
@@ -268,7 +265,7 @@
fun mk_partial_rules_mutual lthy inner_cont (m as Mutual {parts, fqgars, ...}) proof =
let
val result = inner_cont proof
- val FundefResult {fs=[f], G, R, cases, psimps, trsimps, simple_pinducts=[simple_pinduct],
+ val FunctionResult {fs=[f], G, R, cases, psimps, trsimps, simple_pinducts=[simple_pinduct],
termination,domintros} = result
val (all_f_defs, fs) = map (fn MutualPart {f_defthm = SOME f_def, f = SOME f, cargTs, ...} =>
@@ -288,20 +285,20 @@
val mtermination = full_simplify rew_ss termination
val mdomintros = map_option (map (full_simplify rew_ss)) domintros
in
- FundefResult { fs=fs, G=G, R=R,
+ FunctionResult { fs=fs, G=G, R=R,
psimps=mpsimps, simple_pinducts=minducts,
cases=cases, termination=mtermination,
domintros=mdomintros,
trsimps=mtrsimps}
end
-fun prepare_fundef_mutual config defname fixes eqss lthy =
+fun prepare_function_mutual config defname fixes eqss lthy =
let
val mutual = analyze_eqs lthy defname (map fst fixes) (map Envir.beta_eta_contract eqss)
val Mutual {fsum_var=(n, T), qglrs, ...} = mutual
val ((fsum, goalstate, cont), lthy') =
- FundefCore.prepare_fundef config defname [((n, T), NoSyn)] qglrs lthy
+ Function_Core.prepare_function config defname [((n, T), NoSyn)] qglrs lthy
val (mutual', lthy'') = define_projections fixes mutual fsum lthy'
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/pat_completeness.ML Sun Oct 25 08:57:55 2009 +0100
@@ -0,0 +1,163 @@
+(* Title: HOL/Tools/Function/fundef_datatype.ML
+ Author: Alexander Krauss, TU Muenchen
+
+Method "pat_completeness" to prove completeness of datatype patterns.
+*)
+
+signature PAT_COMPLETENESS =
+sig
+ val pat_completeness_tac: Proof.context -> int -> tactic
+ val pat_completeness: Proof.context -> Proof.method
+ val prove_completeness : theory -> term list -> term -> term list list ->
+ term list list -> thm
+
+ val setup : theory -> theory
+end
+
+structure Pat_Completeness : PAT_COMPLETENESS =
+struct
+
+open Function_Lib
+open Function_Common
+
+
+fun mk_argvar i T = Free ("_av" ^ (string_of_int i), T)
+fun mk_patvar i T = Free ("_pv" ^ (string_of_int i), T)
+
+fun inst_free var inst = forall_elim inst o forall_intr var
+
+fun inst_case_thm thy x P thm =
+ let val [Pv, xv] = Term.add_vars (prop_of thm) []
+ in
+ thm |> cterm_instantiate (map (pairself (cterm_of thy))
+ [(Var xv, x), (Var Pv, P)])
+ end
+
+fun invent_vars constr i =
+ let
+ val Ts = binder_types (fastype_of constr)
+ val j = i + length Ts
+ val is = i upto (j - 1)
+ val avs = map2 mk_argvar is Ts
+ val pvs = map2 mk_patvar is Ts
+ in
+ (avs, pvs, j)
+ end
+
+fun filter_pats thy cons pvars [] = []
+ | filter_pats thy cons pvars (([], thm) :: pts) = raise Match
+ | filter_pats thy cons pvars (((pat as Free _) :: pats, thm) :: pts) =
+ let val inst = list_comb (cons, pvars)
+ in (inst :: pats, inst_free (cterm_of thy pat) (cterm_of thy inst) thm)
+ :: (filter_pats thy cons pvars pts)
+ end
+ | filter_pats thy cons pvars ((pat :: pats, thm) :: pts) =
+ if fst (strip_comb pat) = cons
+ then (pat :: pats, thm) :: (filter_pats thy cons pvars pts)
+ else filter_pats thy cons pvars pts
+
+
+fun inst_constrs_of thy (T as Type (name, _)) =
+ map (fn (Cn,CT) =>
+ Envir.subst_term_types (Sign.typ_match thy (body_type CT, T) Vartab.empty) (Const (Cn, CT)))
+ (the (Datatype.get_constrs thy name))
+ | inst_constrs_of thy _ = raise Match
+
+
+fun transform_pat thy avars c_assum ([] , thm) = raise Match
+ | transform_pat thy avars c_assum (pat :: pats, thm) =
+ let
+ val (_, subps) = strip_comb pat
+ val eqs = map (cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (avars ~~ subps)
+ val c_eq_pat = simplify (HOL_basic_ss addsimps (map assume eqs)) c_assum
+ in
+ (subps @ pats,
+ fold_rev implies_intr eqs (implies_elim thm c_eq_pat))
+ end
+
+
+exception COMPLETENESS
+
+fun constr_case thy P idx (v :: vs) pats cons =
+ let
+ val (avars, pvars, newidx) = invent_vars cons idx
+ val c_hyp = cterm_of thy (HOLogic.mk_Trueprop (HOLogic.mk_eq (v, list_comb (cons, avars))))
+ val c_assum = assume c_hyp
+ val newpats = map (transform_pat thy avars c_assum) (filter_pats thy cons pvars pats)
+ in
+ o_alg thy P newidx (avars @ vs) newpats
+ |> implies_intr c_hyp
+ |> fold_rev (forall_intr o cterm_of thy) avars
+ end
+ | constr_case _ _ _ _ _ _ = raise Match
+and o_alg thy P idx [] (([], Pthm) :: _) = Pthm
+ | o_alg thy P idx (v :: vs) [] = raise COMPLETENESS
+ | o_alg thy P idx (v :: vs) pts =
+ if forall (is_Free o hd o fst) pts (* Var case *)
+ then o_alg thy P idx vs
+ (map (fn (pv :: pats, thm) =>
+ (pats, refl RS (inst_free (cterm_of thy pv) (cterm_of thy v) thm))) pts)
+ else (* Cons case *)
+ let
+ val T = fastype_of v
+ val (tname, _) = dest_Type T
+ val {exhaust=case_thm, ...} = Datatype.the_info thy tname
+ val constrs = inst_constrs_of thy T
+ val c_cases = map (constr_case thy P idx (v :: vs) pts) constrs
+ in
+ inst_case_thm thy v P case_thm
+ |> fold (curry op COMP) c_cases
+ end
+ | o_alg _ _ _ _ _ = raise Match
+
+fun prove_completeness thy xs P qss patss =
+ let
+ fun mk_assum qs pats =
+ HOLogic.mk_Trueprop P
+ |> fold_rev (curry Logic.mk_implies o HOLogic.mk_Trueprop o HOLogic.mk_eq) (xs ~~ pats)
+ |> fold_rev Logic.all qs
+ |> cterm_of thy
+
+ val hyps = map2 mk_assum qss patss
+ fun inst_hyps hyp qs = fold (forall_elim o cterm_of thy) qs (assume hyp)
+ val assums = map2 inst_hyps hyps qss
+ in
+ o_alg thy P 2 xs (patss ~~ assums)
+ |> fold_rev implies_intr hyps
+ end
+
+fun pat_completeness_tac ctxt = SUBGOAL (fn (subgoal, i) =>
+ let
+ val thy = ProofContext.theory_of ctxt
+ val (vs, subgf) = dest_all_all subgoal
+ val (cases, _ $ thesis) = Logic.strip_horn subgf
+ handle Bind => raise COMPLETENESS
+
+ fun pat_of assum =
+ let
+ val (qs, imp) = dest_all_all assum
+ val prems = Logic.strip_imp_prems imp
+ in
+ (qs, map (HOLogic.dest_eq o HOLogic.dest_Trueprop) prems)
+ end
+
+ val (qss, x_pats) = split_list (map pat_of cases)
+ val xs = map fst (hd x_pats)
+ handle Empty => raise COMPLETENESS
+
+ val patss = map (map snd) x_pats
+ val complete_thm = prove_completeness thy xs thesis qss patss
+ |> fold_rev (forall_intr o cterm_of thy) vs
+ in
+ PRIMITIVE (fn st => Drule.compose_single(complete_thm, i, st))
+ end
+ handle COMPLETENESS => no_tac)
+
+
+val pat_completeness = SIMPLE_METHOD' o pat_completeness_tac
+
+val setup =
+ Method.setup @{binding pat_completeness} (Scan.succeed pat_completeness)
+ "Completeness prover for datatype patterns"
+
+end
--- a/src/HOL/Tools/Function/pattern_split.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Function/pattern_split.ML Sun Oct 25 08:57:55 2009 +0100
@@ -8,7 +8,7 @@
*)
-signature FUNDEF_SPLIT =
+signature FUNCTION_SPLIT =
sig
val split_some_equations :
Proof.context -> (bool * term) list -> term list list
@@ -17,10 +17,10 @@
Proof.context -> term list -> term list list
end
-structure FundefSplit : FUNDEF_SPLIT =
+structure Function_Split : FUNCTION_SPLIT =
struct
-open FundefLib
+open Function_Lib
(* We use proof context for the variable management *)
(* FIXME: no __ *)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/relation.ML Sun Oct 25 08:57:55 2009 +0100
@@ -0,0 +1,36 @@
+(* Title: HOL/Tools/Function/relation.ML
+ Author: Alexander Krauss, TU Muenchen
+
+A package for general recursive function definitions.
+Method "relation" to commence a termination proof using a user-specified relation.
+*)
+
+signature FUNCTION_RELATION =
+sig
+ val relation_tac: Proof.context -> term -> int -> tactic
+ val setup: theory -> theory
+end
+
+structure Function_Relation : FUNCTION_RELATION =
+struct
+
+fun inst_thm ctxt rel st =
+ let
+ val cert = Thm.cterm_of (ProofContext.theory_of ctxt)
+ val rel' = cert (singleton (Variable.polymorphic ctxt) rel)
+ val st' = Thm.incr_indexes (#maxidx (Thm.rep_cterm rel') + 1) st
+ val Rvar = cert (Var (the_single (Term.add_vars (prop_of st') [])))
+ in
+ Drule.cterm_instantiate [(Rvar, rel')] st'
+ end
+
+fun relation_tac ctxt rel i =
+ TRY (Function_Common.apply_termination_rule ctxt i)
+ THEN PRIMITIVE (inst_thm ctxt rel)
+
+val setup =
+ Method.setup @{binding relation}
+ (Args.term >> (fn rel => fn ctxt => SIMPLE_METHOD' (relation_tac ctxt rel)))
+ "proves termination using a user-specified wellfounded relation"
+
+end
--- a/src/HOL/Tools/Function/scnp_reconstruct.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Function/scnp_reconstruct.ML Sun Oct 25 08:57:55 2009 +0100
@@ -38,8 +38,8 @@
structure ScnpReconstruct : SCNP_RECONSTRUCT =
struct
-val PROFILE = FundefCommon.PROFILE
-fun TRACE x = if ! FundefCommon.profile then tracing x else ()
+val PROFILE = Function_Common.PROFILE
+fun TRACE x = if ! Function_Common.profile then tracing x else ()
open ScnpSolve
@@ -64,7 +64,7 @@
reduction_pair : thm
}
-structure MultisetSetup = TheoryDataFun
+structure Multiset_Setup = TheoryDataFun
(
type T = multiset_setup option
val empty = NONE
@@ -73,10 +73,10 @@
fun merge _ (v1, v2) = if is_some v2 then v2 else v1
)
-val multiset_setup = MultisetSetup.put o SOME
+val multiset_setup = Multiset_Setup.put o SOME
fun undef x = error "undef"
-fun get_multiset_setup thy = MultisetSetup.get thy
+fun get_multiset_setup thy = Multiset_Setup.get thy
|> the_default (Multiset
{ msetT = undef, mk_mset=undef,
mset_regroup_conv=undef, mset_member_tac = undef,
@@ -287,7 +287,7 @@
|> cterm_of thy
in
PROFILE "Proof Reconstruction"
- (CONVERSION (Conv.arg_conv (Conv.arg_conv (FundefLib.regroup_union_conv sl))) 1
+ (CONVERSION (Conv.arg_conv (Conv.arg_conv (Function_Lib.regroup_union_conv sl))) 1
THEN (rtac @{thm reduction_pair_lemma} 1)
THEN (rtac @{thm rp_inv_image_rp} 1)
THEN (rtac (order_rpair ms_rp label) 1)
@@ -350,7 +350,7 @@
fun single_scnp_tac use_tags orders ctxt cont err_cont D = Termination.CALLS (fn (cs, i) =>
let
- val ms_configured = is_some (MultisetSetup.get (ProofContext.theory_of ctxt))
+ val ms_configured = is_some (Multiset_Setup.get (ProofContext.theory_of ctxt))
val orders' = if ms_configured then orders
else filter_out (curry op = MS) orders
val gp = gen_probl D cs
@@ -395,7 +395,7 @@
end
fun gen_sizechange_tac orders autom_tac ctxt err_cont =
- TRY (FundefCommon.apply_termination_rule ctxt 1)
+ TRY (Function_Common.apply_termination_rule ctxt 1)
THEN TRY (Termination.wf_union_tac ctxt)
THEN
(rtac @{thm wf_empty} 1
@@ -406,7 +406,7 @@
fun decomp_scnp orders ctxt =
let
- val extra_simps = FundefCommon.Termination_Simps.get ctxt
+ val extra_simps = Function_Common.Termination_Simps.get ctxt
val autom_tac = auto_tac (clasimpset_of ctxt addsimps2 extra_simps)
in
SIMPLE_METHOD
--- a/src/HOL/Tools/Function/scnp_solve.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Function/scnp_solve.ML Sun Oct 25 08:57:55 2009 +0100
@@ -73,7 +73,7 @@
(* SAT solving *)
val solver = Unsynchronized.ref "auto";
fun sat_solver x =
- FundefCommon.PROFILE "sat_solving..." (SatSolver.invoke_solver (!solver)) x
+ Function_Common.PROFILE "sat_solving..." (SatSolver.invoke_solver (!solver)) x
(* "Virtual constructors" for various propositional variables *)
fun var_constrs (gp as GP (arities, gl)) =
--- a/src/HOL/Tools/Function/termination.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Function/termination.ML Sun Oct 25 08:57:55 2009 +0100
@@ -48,7 +48,7 @@
structure Termination : TERMINATION =
struct
-open FundefLib
+open Function_Lib
val term2_ord = prod_ord TermOrd.fast_term_ord TermOrd.fast_term_ord
structure Term2tab = Table(type key = term * term val ord = term2_ord);
@@ -145,7 +145,7 @@
fun mk_sum_skel rel =
let
- val cs = FundefLib.dest_binop_list @{const_name Lattices.sup} rel
+ val cs = Function_Lib.dest_binop_list @{const_name Lattices.sup} rel
fun collect_pats (Const (@{const_name Collect}, _) $ Abs (_, _, c)) =
let
val (Const ("op &", _) $ (Const ("op =", _) $ _ $ (Const ("Pair", _) $ r $ l)) $ Gam)
@@ -233,7 +233,7 @@
fun CALLS tac i st =
if Thm.no_prems st then all_tac st
else case Thm.term_of (Thm.cprem_of st i) of
- (_ $ (_ $ rel)) => tac (FundefLib.dest_binop_list @{const_name Lattices.sup} rel, i) st
+ (_ $ (_ $ rel)) => tac (Function_Lib.dest_binop_list @{const_name Lattices.sup} rel, i) st
|_ => no_tac st
type ttac = (data -> int -> tactic) -> (data -> int -> tactic) -> data -> int -> tactic
@@ -251,7 +251,7 @@
local
fun dest_term (t : term) = (* FIXME, cf. Lexicographic order *)
let
- val (vars, prop) = FundefLib.dest_all_all t
+ val (vars, prop) = Function_Lib.dest_all_all t
val (prems, concl) = Logic.strip_horn prop
val (lhs, rhs) = concl
|> HOLogic.dest_Trueprop
--- a/src/HOL/Tools/Predicate_Compile/pred_compile_aux.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Predicate_Compile/pred_compile_aux.ML Sun Oct 25 08:57:55 2009 +0100
@@ -6,8 +6,16 @@
structure Predicate_Compile_Aux =
struct
+(* general syntactic functions *)
+
+(*Like dest_conj, but flattens conjunctions however nested*)
+fun conjuncts_aux (Const ("op &", _) $ t $ t') conjs = conjuncts_aux t (conjuncts_aux t' conjs)
+ | conjuncts_aux t conjs = t::conjs;
+
+fun conjuncts t = conjuncts_aux t [];
+
(* syntactic functions *)
-
+
fun is_equationlike_term (Const ("==", _) $ _ $ _) = true
| is_equationlike_term (Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ _)) = true
| is_equationlike_term _ = false
@@ -69,26 +77,52 @@
in ((ps', t''), nctxt') end;
+(* introduction rule combinators *)
+(* combinators to apply a function to all literals of an introduction rules *)
-(*
fun map_atoms f intro =
-fun fold_atoms f intro =
-*)
+ let
+ val (literals, head) = Logic.strip_horn intro
+ fun appl t = (case t of
+ (@{term "Not"} $ t') => HOLogic.mk_not (f t')
+ | _ => f t)
+ in
+ Logic.list_implies
+ (map (HOLogic.mk_Trueprop o appl o HOLogic.dest_Trueprop) literals, head)
+ end
+
+fun fold_atoms f intro s =
+ let
+ val (literals, head) = Logic.strip_horn intro
+ fun appl t s = (case t of
+ (@{term "Not"} $ t') => f t' s
+ | _ => f t s)
+ in fold appl (map HOLogic.dest_Trueprop literals) s end
+
fun fold_map_atoms f intro s =
let
val (literals, head) = Logic.strip_horn intro
fun appl t s = (case t of
- (@{term "Not"} $ t') =>
- let
- val (t'', s') = f t' s
- in (@{term "Not"} $ t'', s') end
+ (@{term "Not"} $ t') => apfst HOLogic.mk_not (f t' s)
| _ => f t s)
val (literals', s') = fold_map appl (map HOLogic.dest_Trueprop literals) s
in
(Logic.list_implies (map HOLogic.mk_Trueprop literals', head), s')
end;
+
+fun maps_premises f intro =
+ let
+ val (premises, head) = Logic.strip_horn intro
+ in
+ Logic.list_implies (maps f premises, head)
+ end
+(* lifting term operations to theorems *)
+
+fun map_term thy f th =
+ Skip_Proof.make_thm thy (f (prop_of th))
+
(*
fun equals_conv lhs_cv rhs_cv ct =
case Thm.term_of ct of
@@ -96,5 +130,107 @@
| _ => error "equals_conv"
*)
+(* Different options for compiler *)
+
+datatype options = Options of {
+ expected_modes : (string * int list list) option,
+ show_steps : bool,
+ show_mode_inference : bool,
+ show_proof_trace : bool,
+ show_intermediate_results : bool,
+ show_compilation : bool,
+ skip_proof : bool,
+
+ inductify : bool,
+ rpred : bool,
+ depth_limited : bool
+};
+
+fun expected_modes (Options opt) = #expected_modes opt
+fun show_steps (Options opt) = #show_steps opt
+fun show_mode_inference (Options opt) = #show_mode_inference opt
+fun show_intermediate_results (Options opt) = #show_intermediate_results opt
+fun show_proof_trace (Options opt) = #show_proof_trace opt
+fun show_compilation (Options opt) = #show_compilation opt
+fun skip_proof (Options opt) = #skip_proof opt
+
+fun is_inductify (Options opt) = #inductify opt
+fun is_rpred (Options opt) = #rpred opt
+fun is_depth_limited (Options opt) = #depth_limited opt
+
+val default_options = Options {
+ expected_modes = NONE,
+ show_steps = false,
+ show_intermediate_results = false,
+ show_proof_trace = false,
+ show_mode_inference = false,
+ show_compilation = false,
+ skip_proof = false,
+
+ inductify = false,
+ rpred = false,
+ depth_limited = false
+}
+
+
+fun print_step options s =
+ if show_steps options then tracing s else ()
+
+(* tuple processing *)
+
+fun expand_tuples thy intro =
+ let
+ fun rewrite_args [] (pats, intro_t, ctxt) = (pats, intro_t, ctxt)
+ | rewrite_args (arg::args) (pats, intro_t, ctxt) =
+ (case HOLogic.strip_tupleT (fastype_of arg) of
+ (Ts as _ :: _ :: _) =>
+ let
+ fun rewrite_arg' (Const ("Pair", _) $ _ $ t2, Type ("*", [_, T2]))
+ (args, (pats, intro_t, ctxt)) = rewrite_arg' (t2, T2) (args, (pats, intro_t, ctxt))
+ | rewrite_arg' (t, Type ("*", [T1, T2])) (args, (pats, intro_t, ctxt)) =
+ let
+ val ([x, y], ctxt') = Variable.variant_fixes ["x", "y"] ctxt
+ val pat = (t, HOLogic.mk_prod (Free (x, T1), Free (y, T2)))
+ val intro_t' = Pattern.rewrite_term thy [pat] [] intro_t
+ val args' = map (Pattern.rewrite_term thy [pat] []) args
+ in
+ rewrite_arg' (Free (y, T2), T2) (args', (pat::pats, intro_t', ctxt'))
+ end
+ | rewrite_arg' _ (args, (pats, intro_t, ctxt)) = (args, (pats, intro_t, ctxt))
+ val (args', (pats, intro_t', ctxt')) = rewrite_arg' (arg, fastype_of arg)
+ (args, (pats, intro_t, ctxt))
+ in
+ rewrite_args args' (pats, intro_t', ctxt')
+ end
+ | _ => rewrite_args args (pats, intro_t, ctxt))
+ fun rewrite_prem atom =
+ let
+ val (_, args) = strip_comb atom
+ in rewrite_args args end
+ val ctxt = ProofContext.init thy
+ val (((T_insts, t_insts), [intro']), ctxt1) = Variable.import false [intro] ctxt
+ val intro_t = prop_of intro'
+ val concl = Logic.strip_imp_concl intro_t
+ val (p, args) = strip_comb (HOLogic.dest_Trueprop concl)
+ val (pats', intro_t', ctxt2) = rewrite_args args ([], intro_t, ctxt1)
+ val (pats', intro_t', ctxt3) =
+ fold_atoms rewrite_prem intro_t' (pats', intro_t', ctxt2)
+ fun rewrite_pat (ct1, ct2) =
+ (ct1, cterm_of thy (Pattern.rewrite_term thy pats' [] (term_of ct2)))
+ val t_insts' = map rewrite_pat t_insts
+ val intro'' = Thm.instantiate (T_insts, t_insts') intro
+ val [intro'''] = Variable.export ctxt3 ctxt [intro'']
+ val intro'''' = Simplifier.full_simplify
+ (HOL_basic_ss addsimps [@{thm fst_conv}, @{thm snd_conv}, @{thm Pair_eq}])
+ intro'''
+ (* splitting conjunctions introduced by Pair_eq*)
+ fun split_conj prem =
+ map HOLogic.mk_Trueprop (conjuncts (HOLogic.dest_Trueprop prem))
+ val intro''''' = map_term thy (maps_premises split_conj) intro''''
+ in
+ intro'''''
+ end
+
+
end;
--- a/src/HOL/Tools/Predicate_Compile/pred_compile_data.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Predicate_Compile/pred_compile_data.ML Sun Oct 25 08:57:55 2009 +0100
@@ -3,16 +3,16 @@
Book-keeping datastructure for the predicate compiler
*)
-signature PRED_COMPILE_DATA =
+signature PREDICATE_COMPILE_DATA =
sig
type specification_table;
val make_const_spec_table : theory -> specification_table
val get_specification : specification_table -> string -> thm list
- val obtain_specification_graph : specification_table -> string -> thm list Graph.T
+ val obtain_specification_graph : theory -> specification_table -> string -> thm list Graph.T
val normalize_equation : theory -> thm -> thm
end;
-structure Pred_Compile_Data : PRED_COMPILE_DATA =
+structure Predicate_Compile_Data : PREDICATE_COMPILE_DATA =
struct
open Predicate_Compile_Aux;
@@ -81,11 +81,13 @@
Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ _) => th RS @{thm eq_reflection}
| _ => th
+val meta_fun_cong = @{lemma "f == g ==> f x == g x" by simp}
+
fun full_fun_cong_expand th =
let
val (f, args) = strip_comb (fst (Logic.dest_equals (prop_of th)))
val i = length (binder_types (fastype_of f)) - length args
- in funpow i (fn th => th RS @{thm meta_fun_cong}) th end;
+ in funpow i (fn th => th RS meta_fun_cong) th end;
fun declare_names s xs ctxt =
let
@@ -117,7 +119,7 @@
fun normalize_equation thy th =
mk_meta_equation th
- |> Pred_Compile_Set.unfold_set_notation
+ |> Predicate_Compile_Set.unfold_set_notation
|> full_fun_cong_expand
|> split_all_pairs thy
|> tap check_equation_format
@@ -127,19 +129,18 @@
fun store_thm_in_table ignore_consts thy th=
let
- val th = AxClass.unoverload thy th
+ val th = th
|> inline_equations thy
+ |> Predicate_Compile_Set.unfold_set_notation
+ |> AxClass.unoverload thy
val (const, th) =
if is_equationlike th then
let
- val _ = priority "Normalizing definition..."
val eq = normalize_equation thy th
in
(defining_const_of_equation eq, eq)
end
- else if (is_introlike th) then
- let val th = Pred_Compile_Set.unfold_set_notation th
- in (defining_const_of_introrule th, th) end
+ else if (is_introlike th) then (defining_const_of_introrule th, th)
else error "store_thm: unexpected definition format"
in
if not (member (op =) ignore_consts const) then
@@ -147,18 +148,6 @@
else I
end
-(*
-fun make_const_spec_table_warning thy =
- fold
- (fn th => fn thy => case try (store_thm th) thy of
- SOME thy => thy
- | NONE => (warning ("store_thm fails for " ^ Display.string_of_thm_global thy th) ; thy))
- (Predicate_Compile_Preproc_Const_Defs.get (ProofContext.init thy)) thy
-
-fun make_const_spec_table thy =
- fold store_thm (Predicate_Compile_Preproc_Const_Defs.get (ProofContext.init thy)) thy
- |> (fn thy => fold store_thm (Nitpick_Simps.get (ProofContext.init thy)) thy)
-*)
fun make_const_spec_table thy =
let
fun store ignore_const f = fold (store_thm_in_table ignore_const thy) (map (Thm.transfer thy) (f (ProofContext.init thy)))
@@ -166,35 +155,27 @@
|> store [] Predicate_Compile_Alternative_Defs.get
val ignore_consts = Symtab.keys table
in
- table
+ table
|> store ignore_consts Predicate_Compile_Preproc_Const_Defs.get
|> store ignore_consts Nitpick_Simps.get
|> store ignore_consts Nitpick_Intros.get
end
- (*
-fun get_specification thy constname =
- case Symtab.lookup (#const_spec_table (Data.get thy)) constname of
- SOME thms => thms
- | NONE => error ("get_specification: lookup of constant " ^ quote constname ^ " failed")
- *)
+
fun get_specification table constname =
case Symtab.lookup table constname of
- SOME thms =>
- let
- val _ = tracing ("Looking up specification of " ^ constname ^ ": "
- ^ (commas (map Display.string_of_thm_without_context thms)))
- in thms end
+ SOME thms => thms
| NONE => error ("get_specification: lookup of constant " ^ quote constname ^ " failed")
val logic_operator_names =
- [@{const_name "=="}, @{const_name "op ="}, @{const_name "op -->"}, @{const_name "All"}, @{const_name "op &"}]
+ [@{const_name "=="}, @{const_name "op ="}, @{const_name "op -->"}, @{const_name "All"}, @{const_name "Ex"},
+ @{const_name "op &"}]
-val special_cases = member (op =) [@{const_name "Suc"}, @{const_name Nat.zero_nat_inst.zero_nat},
+val special_cases = member (op =) [
+ @{const_name "False"},
+ @{const_name "Suc"}, @{const_name Nat.zero_nat_inst.zero_nat},
@{const_name Nat.one_nat_inst.one_nat},
@{const_name "HOL.ord_class.less"}, @{const_name "HOL.ord_class.less_eq"}, @{const_name "HOL.zero_class.zero"},
@{const_name "HOL.one_class.one"}, @{const_name HOL.plus_class.plus},
-@{const_name "Nat.nat.nat_case"}, @{const_name "List.list.list_case"},
-@{const_name "Option.option.option_case"},
@{const_name Nat.ord_nat_inst.less_eq_nat},
@{const_name number_nat_inst.number_of_nat},
@{const_name Int.Bit0},
@@ -203,13 +184,19 @@
@{const_name "Int.zero_int_inst.zero_int"},
@{const_name "List.filter"}]
-fun obtain_specification_graph table constname =
+fun case_consts thy s = is_some (Datatype.info_of_case thy s)
+
+fun obtain_specification_graph thy table constname =
let
fun is_nondefining_constname c = member (op =) logic_operator_names c
val is_defining_constname = member (op =) (Symtab.keys table)
+ fun has_code_pred_intros c = is_some (try (Predicate_Compile_Core.intros_of thy) c)
fun defiants_of specs =
fold (Term.add_const_names o prop_of) specs []
|> filter is_defining_constname
+ |> filter_out is_nondefining_constname
+ |> filter_out has_code_pred_intros
+ |> filter_out (case_consts thy)
|> filter_out special_cases
fun extend constname =
let
--- a/src/HOL/Tools/Predicate_Compile/pred_compile_fun.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Predicate_Compile/pred_compile_fun.ML Sun Oct 25 08:57:55 2009 +0100
@@ -5,9 +5,10 @@
signature PREDICATE_COMPILE_FUN =
sig
- val define_predicates : (string * thm list) list -> theory -> theory
+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 =
@@ -63,6 +64,8 @@
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)
@@ -99,23 +102,29 @@
(Free (Long_Name.base_name name ^ "P", pred_type T))
end
-fun mk_param lookup_pred (t as Free (v, _)) = lookup_pred t
- | mk_param lookup_pred t =
+fun mk_param thy lookup_pred (t as Free (v, _)) = lookup_pred t
+ | mk_param thy lookup_pred t =
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 = lookup_pred f
- val pred_body = list_comb (P, args @ [resvar])
- val param = fold_rev lambda (vs' @ [resvar]) pred_body
- in param end;
-
-
+ 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) *)
@@ -210,10 +219,14 @@
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))]
+ [(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))]
@@ -243,8 +256,10 @@
maps mk_prems_of_assm assms
end
else
- let
+ 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
@@ -261,8 +276,7 @@
val pred = lookup_pred t
val nparams = get_nparams pred
val (params, args) = chop nparams args
- val _ = tracing ("mk_prems'': " ^ (Syntax.string_of_term_global thy t) ^ " has " ^ string_of_int nparams ^ " parameters.")
- val params' = map (mk_param lookup_pred) params
+ val params' = map (mk_param thy lookup_pred) params
in
folds_map mk_prems' args (names', prems)
|> map (fn (argvs, (names'', prems')) =>
@@ -281,7 +295,8 @@
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)
+ | mk_prems'' t =
+ error ("Invalid term: " ^ Syntax.string_of_term_global thy t)
in
map (pair resvar) (mk_prems'' f)
end
@@ -292,7 +307,7 @@
(* 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
+ ([], thy)
else
let
val consts = map fst specs
@@ -307,16 +322,14 @@
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 (@{const_name Cons}, T)) =
- Const ("Preprocessing.ConsP", pred_type T) (* FIXME: temporary - Cons lookup *)
- | lookup_pred (Const (name, T)) =
+ 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)) =
+ | lookup_pred (Free (name, T)) =
if member op = (map fst pnames) name then
Free (name, transform_ho_typ T)
else
@@ -347,12 +360,10 @@
Logic.list_implies (prems, HOLogic.mk_Trueprop (list_comb (pred, args @ [resultt]))))
end
fun mk_rewr_thm (func, pred) = @{thm refl}
- in
+ in
case try (maps mk_intros) ((funs ~~ preds) ~~ (argss' ~~ rhss)) of
- NONE => thy
- | SOME intr_ts => let
- val _ = map (tracing o (Syntax.string_of_term_global thy)) intr_ts
- in
+ NONE => ([], thy)
+ | SOME intr_ts =>
if is_some (try (map (cterm_of thy)) intr_ts) then
let
val (ind_result, thy') =
@@ -367,10 +378,15 @@
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 *)
- in Pred_Compile_Preproc.map (fold Symtab.update_new (consts ~~ prednames)) thy' end
+ 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
- thy
- end
+ let
+ val _ = tracing "Introduction rules of function_predicate are not welltyped"
+ in ([], thy) end
end
(* preprocessing intro rules - uses oracle *)
@@ -391,7 +407,6 @@
| get_nparams t = error ("No parameters for " ^ (Syntax.string_of_term_global thy t))
val intro_t = (Logic.unvarify o prop_of) intro
- val _ = tracing (Syntax.string_of_term_global thy intro_t)
val (prems, concl) = Logic.strip_horn intro_t
val frees = map fst (Term.add_frees intro_t [])
fun rewrite prem names =
@@ -415,8 +430,6 @@
rewrite concl frees'
|> map (fn (concl'::conclprems, _) =>
Logic.list_implies ((flat prems') @ conclprems, concl')))
- val _ = tracing ("intro_ts': " ^
- commas (map (Syntax.string_of_term_global thy) intro_ts'))
in
map (Drule.standard o the_oracle () o cterm_of thy) intro_ts'
end;
--- a/src/HOL/Tools/Predicate_Compile/pred_compile_pred.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Predicate_Compile/pred_compile_pred.ML Sun Oct 25 08:57:55 2009 +0100
@@ -123,16 +123,67 @@
else
error ("unexpected specification for constant " ^ quote constname ^ ":\n"
^ commas (map (quote o Display.string_of_thm_global thy) specs))
- val _ = tracing ("Introduction rules of definitions before flattening: "
- ^ commas (map (Display.string_of_thm ctxt) intros))
- val _ = tracing "Defining local predicates and their intro rules..."
val (intros', (local_defs, thy')) = flatten_intros constname intros thy
val (intross, thy'') = fold_map preprocess local_defs thy'
in
- (intros' :: flat intross,thy'')
+ ((constname, intros') :: flat intross,thy'')
end;
fun preprocess_term t thy = error "preprocess_pred_term: to implement"
-
-
+
+fun is_Abs (Abs _) = true
+ | is_Abs _ = false
+
+fun flat_higher_order_arguments (intross, thy) =
+ let
+ fun process constname atom (new_defs, thy) =
+ let
+ val (pred, args) = strip_comb atom
+ val abs_args = filter is_Abs args
+ fun replace_abs_arg (abs_arg as Abs _ ) (new_defs, thy) =
+ let
+ val _ = tracing ("Introduce new constant for " ^
+ Syntax.string_of_term_global thy abs_arg)
+ val vars = map Var (Term.add_vars abs_arg [])
+ val abs_arg' = Logic.unvarify abs_arg
+ val frees = map Free (Term.add_frees abs_arg' [])
+ val constname = Name.variant (map (Long_Name.base_name o fst) new_defs)
+ ((Long_Name.base_name constname) ^ "_hoaux")
+ val full_constname = Sign.full_bname thy constname
+ val constT = map fastype_of frees ---> (fastype_of abs_arg')
+ val const = Const (full_constname, constT)
+ val lhs = list_comb (const, frees)
+ val def = Logic.mk_equals (lhs, abs_arg')
+ val _ = tracing (Syntax.string_of_term_global thy def)
+ val ([definition], thy') = thy
+ |> Sign.add_consts_i [(Binding.name constname, constT, NoSyn)]
+ |> PureThy.add_defs false [((Binding.name (constname ^ "_def"), def), [])]
+ in
+ (list_comb (Logic.varify const, vars), ((full_constname, [definition])::new_defs, thy'))
+ end
+ | replace_abs_arg arg (new_defs, thy) = (arg, (new_defs, thy))
+ val (args', (new_defs', thy')) = fold_map replace_abs_arg args (new_defs, thy)
+ in
+ (list_comb (pred, args'), (new_defs', thy'))
+ end
+ fun flat_intro intro (new_defs, thy) =
+ let
+ val constname = fst (dest_Const (fst (strip_comb
+ (HOLogic.dest_Trueprop (Logic.strip_imp_concl (prop_of intro))))))
+ val (intro_ts, (new_defs, thy)) = fold_map_atoms (process constname) (prop_of intro) (new_defs, thy)
+ val th = Skip_Proof.make_thm thy intro_ts
+ in
+ (th, (new_defs, thy))
+ end
+ fun fold_map_spec f [] s = ([], s)
+ | fold_map_spec f ((c, ths) :: specs) s =
+ let
+ val (ths', s') = f ths s
+ val (specs', s'') = fold_map_spec f specs s'
+ in ((c, ths') :: specs', s'') end
+ val (intross', (new_defs, thy')) = fold_map_spec (fold_map flat_intro) intross ([], thy)
+ in
+ (intross', (new_defs, thy'))
+ end
+
end;
--- a/src/HOL/Tools/Predicate_Compile/pred_compile_quickcheck.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Predicate_Compile/pred_compile_quickcheck.ML Sun Oct 25 08:57:55 2009 +0100
@@ -3,18 +3,19 @@
A quickcheck generator based on the predicate compiler
*)
-signature PRED_COMPILE_QUICKCHECK =
+signature PREDICATE_COMPILE_QUICKCHECK =
sig
val quickcheck : Proof.context -> term -> int -> term list option
val test_ref :
((unit -> int -> int * int -> term list Predicate.pred * (int * int)) option) Unsynchronized.ref
end;
-structure Pred_Compile_Quickcheck : PRED_COMPILE_QUICKCHECK =
+structure Predicate_Compile_Quickcheck : PREDICATE_COMPILE_QUICKCHECK =
struct
val test_ref =
Unsynchronized.ref (NONE : (unit -> int -> int * int -> term list Predicate.pred * (int * int)) option)
+
val target = "Quickcheck"
fun dest_compfuns (Predicate_Compile_Core.CompilationFuns funs) = funs
@@ -63,21 +64,21 @@
val _ = tracing (Display.string_of_thm ctxt' intro)
val thy'' = thy'
|> Context.theory_map (Predicate_Compile_Preproc_Const_Defs.add_thm intro)
- |> Predicate_Compile.preprocess full_constname
- |> Predicate_Compile_Core.add_equations [full_constname]
- |> Predicate_Compile_Core.add_sizelim_equations [full_constname]
- |> Predicate_Compile_Core.add_quickcheck_equations [full_constname]
- val sizelim_modes = Predicate_Compile_Core.sizelim_modes_of thy'' full_constname
+ |> Predicate_Compile.preprocess Predicate_Compile_Aux.default_options full_constname
+ (* |> Predicate_Compile_Core.add_equations Predicate_Compile_Aux.default_options [full_constname]*)
+ (* |> Predicate_Compile_Core.add_depth_limited_equations Predicate_Compile_Aux.default_options [full_constname]*)
+ (* |> Predicate_Compile_Core.add_quickcheck_equations Predicate_Compile_Aux.default_options [full_constname] *)
+ val depth_limited_modes = Predicate_Compile_Core.depth_limited_modes_of thy'' full_constname
val modes = Predicate_Compile_Core.generator_modes_of thy'' full_constname
val prog =
if member (op =) modes ([], []) then
let
val name = Predicate_Compile_Core.generator_name_of thy'' full_constname ([], [])
- val T = @{typ code_numeral} --> (mk_rpredT (HOLogic.mk_tupleT (map snd vs')))
- in Const (name, T) $ Bound 0 end
- else if member (op =) sizelim_modes ([], []) then
+ val T = [@{typ bool}, @{typ code_numeral}] ---> (mk_rpredT (HOLogic.mk_tupleT (map snd vs')))
+ in Const (name, T) $ @{term True} $ Bound 0 end
+ else if member (op =) depth_limited_modes ([], []) then
let
- val name = Predicate_Compile_Core.sizelim_function_name_of thy'' full_constname ([], [])
+ val name = Predicate_Compile_Core.depth_limited_function_name_of thy'' full_constname ([], [])
val T = @{typ code_numeral} --> (mk_predT (HOLogic.mk_tupleT (map snd vs')))
in lift_pred (Const (name, T) $ Bound 0) end
else error "Predicate Compile Quickcheck failed"
@@ -85,7 +86,7 @@
mk_split_lambda (map Free vs') (mk_return (HOLogic.mk_list @{typ term}
(map2 HOLogic.mk_term_of (map snd vs') (map Free vs'))))))
val _ = tracing (Syntax.string_of_term ctxt' qc_term)
- val compile = Code_ML.eval (SOME target) ("Pred_Compile_Quickcheck.test_ref", test_ref)
+ val compile = Code_ML.eval (SOME target) ("Predicate_Compile_Quickcheck.test_ref", test_ref)
(fn proc => fn g => fn s => g s #>> (Predicate.map o map) proc)
thy'' qc_term []
in
--- a/src/HOL/Tools/Predicate_Compile/pred_compile_set.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Predicate_Compile/pred_compile_set.ML Sun Oct 25 08:57:55 2009 +0100
@@ -3,7 +3,7 @@
Preprocessing sets to predicates
*)
-signature PRED_COMPILE_SET =
+signature PREDICATE_COMPILE_SET =
sig
(*
val preprocess_intro : thm -> theory -> thm * theory
@@ -12,10 +12,11 @@
val unfold_set_notation : thm -> thm;
end;
-structure Pred_Compile_Set : PRED_COMPILE_SET =
+structure Predicate_Compile_Set : PREDICATE_COMPILE_SET =
struct
(*FIXME: unfolding Ball in pretty adhoc here *)
-val unfold_set_lemmas = [@{thm Collect_def}, @{thm mem_def}, @{thm Ball_def}]
+val unfold_set_lemmas = [@{thm Collect_def}, @{thm mem_def},
+@{thm Ball_def}, @{thm Bex_def}]
val unfold_set_notation = Simplifier.rewrite_rule unfold_set_lemmas
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile.ML Sun Oct 25 08:57:55 2009 +0100
@@ -4,87 +4,170 @@
signature PREDICATE_COMPILE =
sig
val setup : theory -> theory
- val preprocess : string -> theory -> theory
+ val preprocess : Predicate_Compile_Aux.options -> string -> theory -> theory
end;
structure Predicate_Compile : PREDICATE_COMPILE =
struct
+(* options *)
+val fail_safe_mode = true
+
open Predicate_Compile_Aux;
val priority = tracing;
(* Some last processing *)
+
fun remove_pointless_clauses intro =
if Logic.strip_imp_prems (prop_of intro) = [@{prop "False"}] then
[]
else [intro]
-fun preprocess_strong_conn_constnames gr constnames thy =
+fun tracing s = ()
+
+fun print_intross options thy msg intross =
+ if show_intermediate_results options then
+ Output.tracing (msg ^
+ (space_implode "\n" (map
+ (fn (c, intros) => "Introduction rule(s) of " ^ c ^ ":\n" ^
+ commas (map (Display.string_of_thm_global thy) intros)) intross)))
+ else ()
+
+fun print_specs thy specs =
+ map (fn (c, thms) => "Constant " ^ c ^ " has specification:\n"
+ ^ (space_implode "\n" (map (Display.string_of_thm_global thy) thms)) ^ "\n") specs
+
+fun map_specs f specs =
+ map (fn (s, ths) => (s, f ths)) specs
+
+fun process_specification options specs thy' =
+ let
+ val _ = print_step options "Compiling predicates to flat introrules..."
+ val specs = map (apsnd (map
+ (fn th => if is_equationlike th then Predicate_Compile_Data.normalize_equation thy' th else th))) specs
+ val (intross1, thy'') = apfst flat (fold_map Predicate_Compile_Pred.preprocess specs thy')
+ val _ = print_intross options thy'' "Flattened introduction rules: " intross1
+ val _ = print_step options "Replacing functions in introrules..."
+ val intross2 =
+ if fail_safe_mode then
+ case try (map_specs (maps (Predicate_Compile_Fun.rewrite_intro thy''))) intross1 of
+ SOME intross => intross
+ | NONE => let val _ = warning "Function replacement failed!" in intross1 end
+ else map_specs (maps (Predicate_Compile_Fun.rewrite_intro thy'')) intross1
+ val _ = print_intross options thy'' "Introduction rules with replaced functions: " intross2
+ val _ = print_step options "Introducing new constants for abstractions at higher-order argument positions..."
+ val (intross3, (new_defs, thy''')) = Predicate_Compile_Pred.flat_higher_order_arguments (intross2, thy'')
+ val (new_intross, thy'''') =
+ if not (null new_defs) then
+ let
+ val _ = print_step options "Recursively obtaining introduction rules for new definitions..."
+ in process_specification options new_defs thy''' end
+ else ([], thy''')
+ in
+ (intross3 @ new_intross, thy'''')
+ end
+
+
+fun preprocess_strong_conn_constnames options gr constnames thy =
let
val get_specs = map (fn k => (k, Graph.get_node gr k))
- val _ = priority ("Preprocessing scc of " ^ commas constnames)
+ val _ = print_step options ("Preprocessing scc of " ^ commas constnames)
val (prednames, funnames) = List.partition (is_pred thy) constnames
(* untangle recursion by defining predicates for all functions *)
- val _ = priority "Compiling functions to predicates..."
- val _ = tracing ("funnames: " ^ commas funnames)
- val thy' =
- thy |> not (null funnames) ? Predicate_Compile_Fun.define_predicates
- (get_specs funnames)
- val _ = priority "Compiling predicates to flat introrules..."
- val (intross, thy'') = apfst flat (fold_map Predicate_Compile_Pred.preprocess
- (get_specs prednames) thy')
- val _ = tracing ("Flattened introduction rules: " ^
- commas (map (Display.string_of_thm_global thy'') (flat intross)))
- val _ = priority "Replacing functions in introrules..."
- (* val _ = burrow (maps (Predicate_Compile_Fun.rewrite_intro thy'')) intross *)
- val intross' =
- case try (burrow (maps (Predicate_Compile_Fun.rewrite_intro thy''))) intross of
- SOME intross' => intross'
- | NONE => let val _ = warning "Function replacement failed!" in intross end
- val _ = tracing ("Introduction rules with replaced functions: " ^
- commas (map (Display.string_of_thm_global thy'') (flat intross')))
- val intross'' = burrow (maps remove_pointless_clauses) intross'
- val intross'' = burrow (map (AxClass.overload thy'')) intross''
- val _ = priority "Registering intro rules..."
- val thy''' = fold Predicate_Compile_Core.register_intros intross'' thy''
+ val _ = print_step options
+ ("Compiling functions (" ^ commas funnames ^ ") to predicates...")
+ val (fun_pred_specs, thy') =
+ if not (null funnames) then Predicate_Compile_Fun.define_predicates
+ (get_specs funnames) thy else ([], thy)
+ val _ = print_specs thy' fun_pred_specs
+ val specs = (get_specs prednames) @ fun_pred_specs
+ val (intross3, thy''') = process_specification options specs thy'
+ val _ = print_intross options thy''' "Introduction rules with new constants: " intross3
+ val intross4 = map_specs (maps remove_pointless_clauses) intross3
+ val _ = print_intross options thy''' "After removing pointless clauses: " intross4
+ (*val intross5 = map (fn s, ths) => ( s, map (AxClass.overload thy''') ths)) intross4*)
+ val intross6 = map_specs (map (expand_tuples thy''')) intross4
+ val _ = print_intross options thy''' "introduction rules before registering: " intross6
+ val _ = print_step options "Registering introduction rules..."
+ val thy'''' = fold Predicate_Compile_Core.register_intros intross6 thy'''
in
- thy'''
+ thy''''
end;
-fun preprocess const thy =
+fun preprocess options const thy =
let
- val _ = tracing ("Fetching definitions from theory...")
- val table = Pred_Compile_Data.make_const_spec_table thy
- val gr = Pred_Compile_Data.obtain_specification_graph table const
- val _ = tracing (commas (Graph.all_succs gr [const]))
+ val _ = print_step options "Fetching definitions from theory..."
+ val table = Predicate_Compile_Data.make_const_spec_table thy
+ val gr = Predicate_Compile_Data.obtain_specification_graph thy table const
val gr = Graph.subgraph (member (op =) (Graph.all_succs gr [const])) gr
- in fold_rev (preprocess_strong_conn_constnames gr)
+ in fold_rev (preprocess_strong_conn_constnames options gr)
(Graph.strong_conn gr) thy
end
-fun code_pred_cmd ((inductify_all, rpred), raw_const) lthy =
- if inductify_all then
- let
- val thy = ProofContext.theory_of lthy
- val const = Code.read_const thy raw_const
- val lthy' = LocalTheory.theory (preprocess const) lthy
- |> LocalTheory.checkpoint
- val _ = tracing "Starting Predicate Compile Core..."
- in Predicate_Compile_Core.code_pred_cmd rpred raw_const lthy' end
- else
- Predicate_Compile_Core.code_pred_cmd rpred raw_const lthy
+fun extract_options ((modes, raw_options), const) =
+ let
+ fun chk s = member (op =) raw_options s
+ in
+ Options {
+ expected_modes = Option.map (pair const) modes,
+ show_steps = chk "show_steps",
+ show_intermediate_results = chk "show_intermediate_results",
+ show_proof_trace = chk "show_proof_trace",
+ show_mode_inference = chk "show_mode_inference",
+ show_compilation = chk "show_compilation",
+ skip_proof = chk "skip_proof",
+ inductify = chk "inductify",
+ rpred = chk "rpred",
+ depth_limited = chk "depth_limited"
+ }
+ end
+
+fun code_pred_cmd ((modes, raw_options), raw_const) lthy =
+ let
+ val thy = ProofContext.theory_of lthy
+ val const = Code.read_const thy raw_const
+ val options = extract_options ((modes, raw_options), const)
+ in
+ if (is_inductify options) then
+ let
+ val lthy' = LocalTheory.theory (preprocess options const) lthy
+ |> LocalTheory.checkpoint
+ val const = case Predicate_Compile_Fun.pred_of_function (ProofContext.theory_of lthy') const of
+ SOME c => c
+ | NONE => const
+ val _ = print_step options "Starting Predicate Compile Core..."
+ in
+ Predicate_Compile_Core.code_pred options const lthy'
+ end
+ else
+ Predicate_Compile_Core.code_pred_cmd options raw_const lthy
+ end
val setup = Predicate_Compile_Fun.setup_oracle #> Predicate_Compile_Core.setup
-val _ = List.app OuterKeyword.keyword ["inductify_all", "rpred"]
+val bool_options = ["show_steps", "show_intermediate_results", "show_proof_trace",
+ "show_mode_inference", "show_compilation", "skip_proof", "inductify", "rpred", "depth_limited"]
local structure P = OuterParse
in
+val opt_modes =
+ Scan.optional (P.$$$ "(" |-- Args.$$$ "mode" |-- P.$$$ ":" |--
+ P.enum1 "," (P.$$$ "[" |-- P.enum "," P.nat --| P.$$$ "]")
+ --| P.$$$ ")" >> SOME) NONE
+
+val scan_params =
+ let
+ val scan_bool_param = foldl1 (op ||) (map Args.$$$ bool_options)
+ in
+ Scan.optional (P.$$$ "[" |-- P.enum1 "," scan_bool_param --| P.$$$ "]") []
+ end
+
val _ = OuterSyntax.local_theory_to_proof "code_pred"
"prove equations for predicate specified by intro/elim rules"
- OuterKeyword.thy_goal (P.opt_keyword "inductify_all" -- P.opt_keyword "rpred" -- 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_core.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Sun Oct 25 08:57:55 2009 +0100
@@ -7,26 +7,25 @@
signature PREDICATE_COMPILE_CORE =
sig
val setup: theory -> theory
- val code_pred: bool -> string -> Proof.context -> Proof.state
- val code_pred_cmd: bool -> string -> Proof.context -> Proof.state
+ 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 add_equations_of: bool -> string list -> theory -> theory *)
- val register_predicate : (thm list * thm * int) -> theory -> theory
- val register_intros : thm list -> theory -> theory
+ 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 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 sizelim_modes_of: theory -> string -> mode list
- val sizelim_function_name_of : theory -> string -> mode -> string
+ 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 intros_of: theory -> string -> thm list
val nparams_of: theory -> string -> int
@@ -35,24 +34,12 @@
val set_nparams : string -> int -> theory -> theory
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 mk_casesrule : Proof.context -> term -> int -> thm list -> term
+ val eval_ref : (unit -> term Predicate.pred) option Unsynchronized.ref
+ val random_eval_ref : (unit -> int * int -> term Predicate.pred * (int * int)) option Unsynchronized.ref
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,
@@ -64,66 +51,39 @@
mk_not : term -> term,
mk_map : typ -> typ -> term -> term -> term,
lift_pred : term -> term
- };
- type moded_clause = term list * (indprem * tmode) list
- type 'a pred_mode_table = (string * (mode * 'a) list) list
- val infer_modes : theory -> (string * mode list) list
- -> (string * mode list) list
- -> string list
- -> (string * (term list * indprem list) list) list
- -> (moded_clause list) pred_mode_table
- val infer_modes_with_generator : theory -> (string * mode list) list
- -> (string * mode list) 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 pred_compfuns : compilation_funs
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
- val cprods_subset : 'a list list -> 'a list list
+ (* val add_equations : Predicate_Compile_Aux.options -> string list -> theory -> theory
+ val add_quickcheck_equations : Predicate_Compile_Aux.options -> string list -> theory -> theory
+ val add_depth_limited_equations : Predicate_Compile_Aux.options -> string list -> theory -> theory
+ *)
end;
structure Predicate_Compile_Core : PREDICATE_COMPILE_CORE =
struct
open Predicate_Compile_Aux;
+
(** auxiliary **)
(* debug stuff *)
-fun tracing s = (if ! Toplevel.debug then tracing s else ());
+fun tracing s = (if ! Toplevel.debug then Output.tracing s else ());
+
+fun print_tac s = Seq.single;
+
+fun print_tac' options s =
+ if show_proof_trace options then Tactical.print_tac s else Seq.single;
-fun print_tac s = Seq.single; (*Tactical.print_tac s;*) (* (if ! Toplevel.debug then Tactical.print_tac s else Seq.single); *)
-fun debug_tac msg = Seq.single; (* (fn st => (tracing msg; Seq.single st)); *)
+fun debug_tac msg = Seq.single; (* (fn st => (Output.tracing msg; Seq.single st)); *)
-val do_proofs = Unsynchronized.ref true;
+datatype assertion = Max_number_of_subgoals of int
+fun assert_tac (Max_number_of_subgoals i) st =
+ if (nprems_of st <= i) then Seq.single st
+ else error ("assert_tac: Numbers of subgoals mismatch at goal state :"
+ ^ "\n" ^ Pretty.string_of (Pretty.chunks
+ (Goal_Display.pretty_goals_without_context (! Goal_Display.goals_limit) st)));
(* reference to preprocessing of InductiveSet package *)
@@ -139,20 +99,6 @@
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
@@ -189,6 +135,13 @@
(** data structures **)
+(* new datatype for modes: *)
+(*
+datatype instantiation = Input | Output
+type arg_mode = Tuple of instantiation list | Atom of instantiation | HigherOrderMode of mode
+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;
@@ -241,35 +194,8 @@
(if null param_modes then "" else
"; " ^ "params: " ^ commas (map (the_default "NONE" o Option.map string_of_tmode) param_modes))
-(* generation of case rules from user-given introduction rules *)
-
-fun mk_casesrule ctxt nparams introrules =
- let
- val ((_, intros_th), ctxt1) = Variable.import false introrules ctxt
- val intros = map prop_of intros_th
- val (pred, (params, args)) = strip_intro_concl nparams (hd intros)
- val ([propname], ctxt2) = Variable.variant_fixes ["thesis"] ctxt1
- val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT))
- val (argnames, ctxt3) = Variable.variant_fixes
- (map (fn i => "a" ^ string_of_int i) (1 upto (length args))) ctxt2
- 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;
-
-
-datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term |
- GeneratorPrem of term list * term | Generator of (string * typ);
+datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term
+ | Generator of (string * typ);
type moded_clause = term list * (indprem * tmode) list
type 'a pred_mode_table = (string * (mode * 'a) list) list
@@ -300,25 +226,25 @@
nparams : int,
functions : (mode * predfun_data) list,
generators : (mode * function_data) list,
- sizelim_functions : (mode * function_data) list
+ depth_limited_functions : (mode * function_data) list
};
fun rep_pred_data (PredData data) = data;
-fun mk_pred_data ((intros, elim, nparams), (functions, generators, sizelim_functions)) =
+fun mk_pred_data ((intros, elim, nparams), (functions, generators, depth_limited_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)))
-
+ functions = functions, generators = generators, depth_limited_functions = depth_limited_functions}
+fun map_pred_data f (PredData {intros, elim, nparams, functions, generators, depth_limited_functions}) =
+ mk_pred_data (f ((intros, elim, nparams), (functions, generators, depth_limited_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;
@@ -353,7 +279,7 @@
val modes_of = (map fst) o #functions oo the_pred_data
-val sizelim_modes_of = (map fst) o #sizelim_functions oo the_pred_data
+val depth_limited_modes_of = (map fst) o #depth_limited_functions oo the_pred_data
val rpred_modes_of = (map fst) o #generators oo the_pred_data
@@ -380,7 +306,7 @@
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
@@ -392,24 +318,25 @@
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 =
+fun lookup_depth_limited_function_data thy name mode =
Option.map rep_function_data (AList.lookup (op =)
- (#sizelim_functions (the_pred_data thy name)) mode)
+ (#depth_limited_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
+fun the_depth_limited_function_data thy name mode = case lookup_depth_limited_function_data thy name mode
+ of NONE => error ("No depth-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 depth_limited_function_name_of = #name ooo the_depth_limited_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_modes modes =
+ Output.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
@@ -417,15 +344,20 @@
^ (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))
+ val _ = Output.tracing (cat_lines (map print_pred pred_mode_table))
in () end;
+fun string_of_prem thy (Prem (ts, p)) =
+ (Syntax.string_of_term_global thy (list_comb (p, ts))) ^ "(premise)"
+ | string_of_prem thy (Negprem (ts, p)) =
+ (Syntax.string_of_term_global thy (HOLogic.mk_not (list_comb (p, ts)))) ^ "(negative premise)"
+ | string_of_prem thy (Sidecond t) =
+ (Syntax.string_of_term_global thy t) ^ "(sidecondition)"
+ | string_of_prem thy _ = error "string_of_prem: unexpected input"
+
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, _)) =
@@ -435,18 +367,25 @@
(Syntax.string_of_term_global thy t) ^
"(sidecond mode: " ^ string_of_smode is ^ ")"
| string_of_moded_prem _ _ = error "string_of_moded_prem: unimplemented"
-
+
fun print_moded_clauses thy =
- let
+ 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 string_of_clause thy pred (ts, prems) =
+ (space_implode " --> "
+ (map (string_of_prem thy) prems)) ^ " --> " ^ pred ^ " "
+ ^ (space_implode " " (map (Syntax.string_of_term_global thy) ts))
+
+fun print_compiled_terms options thy =
+ if show_compilation options then
+ print_pred_mode_table (fn _ => fn _ => Syntax.string_of_term_global thy) thy
+ else K ()
+
fun print_stored_rules thy =
let
val preds = (Graph.keys o PredData.get) thy
@@ -477,7 +416,105 @@
in
fold print (all_modes_of thy) ()
end
-
+
+(* validity checks *)
+
+fun check_expected_modes (options : Predicate_Compile_Aux.options) modes =
+ 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
+ error ("expected modes were not inferred:\n"
+ ^ "inferred modes for " ^ s ^ ": "
+ ^ commas (map ((enclose "[" "]") o string_of_smode o snd) modes))
+ else ()
+ | NONE => ())
+ | NONE => ()
+
+(* importing introduction rules *)
+
+fun unify_consts thy cs intr_ts =
+ (let
+ val add_term_consts_2 = fold_aterms (fn Const c => insert (op =) c | _ => I);
+ fun varify (t, (i, ts)) =
+ let val t' = map_types (Logic.incr_tvar (i + 1)) (#2 (Type.varify [] t))
+ in (maxidx_of_term t', t'::ts) end;
+ val (i, cs') = List.foldr varify (~1, []) cs;
+ val (i', intr_ts') = List.foldr varify (i, []) intr_ts;
+ 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)
+ 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)
+ in (map subst cs', map subst intr_ts')
+ end) handle Type.TUNIFY =>
+ (warning "Occurrences of recursive constant have non-unifiable types"; (cs, intr_ts));
+
+fun import_intros inp_pred nparams [] ctxt =
+ let
+ val ([outp_pred], ctxt') = Variable.import_terms false [inp_pred] ctxt
+ val (paramTs, _) = chop nparams (binder_types (fastype_of outp_pred))
+ val (param_names, ctxt'') = Variable.variant_fixes (map (fn i => "p" ^ (string_of_int i))
+ (1 upto nparams)) ctxt'
+ val params = map Free (param_names ~~ paramTs)
+ in (((outp_pred, params), []), ctxt') end
+ | import_intros inp_pred nparams (th :: ths) ctxt =
+ let
+ val ((_, [th']), ctxt') = Variable.import false [th] ctxt
+ val thy = ProofContext.theory_of ctxt'
+ val (pred, (params, args)) = strip_intro_concl nparams (prop_of th')
+ val ho_args = filter (is_predT o fastype_of) args
+ fun subst_of (pred', pred) =
+ let
+ val subst = Sign.typ_match thy (fastype_of pred', fastype_of pred) Vartab.empty
+ in map (fn (indexname, (s, T)) => ((indexname, s), T)) (Vartab.dest subst) end
+ fun instantiate_typ th =
+ let
+ val (pred', _) = strip_intro_concl 0 (prop_of th)
+ val _ = if not (fst (dest_Const pred) = fst (dest_Const pred')) then
+ error "Trying to instantiate another predicate" else ()
+ in Thm.certify_instantiate (subst_of (pred', pred), []) th end;
+ fun instantiate_ho_args th =
+ let
+ val (_, (params', args')) = strip_intro_concl nparams (prop_of th)
+ val ho_args' = map dest_Var (filter (is_predT o fastype_of) args')
+ in Thm.certify_instantiate ([], map dest_Var params' ~~ params) th end
+ val outp_pred =
+ Term_Subst.instantiate (subst_of (inp_pred, pred), []) inp_pred
+ val ((_, ths'), ctxt1) =
+ Variable.import false (map (instantiate_typ #> instantiate_ho_args) ths) ctxt'
+ in
+ (((outp_pred, params), th' :: ths'), ctxt1)
+ end
+
+(* generation of case rules from user-given introduction rules *)
+
+fun mk_casesrule ctxt pred nparams introrules =
+ let
+ val (((pred, params), intros_th), ctxt1) = import_intros pred nparams introrules ctxt
+ val intros = map prop_of intros_th
+ val ([propname], ctxt2) = Variable.variant_fixes ["thesis"] ctxt1
+ val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT))
+ val (_, argsT) = chop nparams (binder_types (fastype_of pred))
+ val (argnames, ctxt3) = Variable.variant_fixes
+ (map (fn i => "a" ^ string_of_int i) (1 upto length argsT)) ctxt2
+ val argvs = map2 (curry Free) argnames argsT
+ 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;
+
(** preprocessing rules **)
fun imp_prems_conv cv ct =
@@ -498,39 +535,30 @@
fun preprocess_elim thy nparams elimrule =
let
- val _ = tracing ("Preprocessing elimination rule "
- ^ (Display.string_of_thm_global thy elimrule))
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 ctxt = ProofContext.init thy
+ val ((_, [elimrule]), ctxt') = Variable.import false [elimrule] ctxt
val prems = Thm.prems_of elimrule
val nargs = length (snd (strip_comb (HOLogic.dest_Trueprop (hd prems)))) - nparams
fun preprocess_case t =
- let
+ 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
+ in
list_all (params, Logic.list_implies (assums_hyp', Logic.strip_assums_concl t))
- end
+ end
val cases' = map preprocess_case (tl prems)
val elimrule' = Logic.list_implies ((hd prems) :: cases', Thm.concl_of elimrule)
- (*val _ = tracing ("elimrule': "^ (Syntax.string_of_term_global thy elimrule'))*)
val bigeq = (Thm.symmetric (Conv.implies_concl_conv
(MetaSimplifier.rewrite true [@{thm Predicate.eq_is_eq}])
(cterm_of thy elimrule')))
- (*
- val _ = tracing ("bigeq:" ^ (Display.string_of_thm_global thy bigeq))
- val res =
- Thm.equal_elim bigeq elimrule
- *)
- (*
- val t = (fn {...} => mycheat_tac thy 1)
- val eq = Goal.prove (ProofContext.init thy) [] [] (Logic.mk_equals ((Thm.prop_of elimrule), elimrule')) t
- *)
- val _ = tracing "Preprocessed elimination rule"
+ val tac = (fn _ => Skip_Proof.cheat_tac thy)
+ val eq = Goal.prove ctxt' [] [] (Logic.mk_equals ((Thm.prop_of elimrule), elimrule')) tac
in
- Thm.equal_elim bigeq elimrule
+ Thm.equal_elim eq elimrule |> singleton (Variable.export ctxt' ctxt)
end;
(* special case: predicate with no introduction rule *)
@@ -552,6 +580,8 @@
([intro], elim)
end
+fun expand_tuples_elim th = th
+
fun fetch_pred_data thy name =
case try (Inductive.the_inductive (ProofContext.init thy)) name of
SOME (info as (_, result)) =>
@@ -560,17 +590,23 @@
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 intros = ind_set_codegen_preproc thy
+ (map (expand_tuples thy #> preprocess_intro thy) (filter is_intro_of (#intrs result)))
+ val index = find_index (fn s => s = name) (#names (fst info))
+ val pre_elim = nth (#elims result) index
+ val pred = nth (#preds result) index
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)
+ (*val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams
+ (expand_tuples_elim pre_elim))*)
+ val elim =
+ (Drule.standard o Skip_Proof.make_thm thy)
+ (mk_casesrule (ProofContext.init thy) pred nparams intros)
+ 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)
@@ -605,6 +641,7 @@
(data, keys)
end;
*)
+
(* guessing number of parameters *)
fun find_indexes pred xs =
let
@@ -624,7 +661,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 (intro, elim, nparams) => (thm::intro, elim, nparams)))) gr
+ (apfst (fn (intros, elim, nparams) => (thm::intros, elim, nparams)))) gr
| NONE =>
let
val nparams = the_default (guess_nparams T) (try (#nparams o rep_pred_data o (fetch_pred_data thy)) name)
@@ -640,30 +677,34 @@
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 =
+
+fun register_predicate (constname, 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
- if not (member (op =) (Graph.keys (PredData.get thy)) name) then
+ if not (member (op =) (Graph.keys (PredData.get thy)) constname) then
PredData.map
- (Graph.new_node (name, mk_pred_data ((intros, SOME elim, nparams), ([], [], [])))) thy
+ (Graph.new_node (constname, mk_pred_data ((intros, SOME elim, nparams), ([], [], [])))) thy
else thy
end
-fun register_intros pre_intros thy =
+fun register_intros (constname, pre_intros) thy =
let
- val (c, T) = dest_Const (fst (strip_intro_concl 0 (prop_of (hd pre_intros))))
- val _ = tracing ("Registering introduction rules of " ^ c)
- val _ = tracing (commas (map (Display.string_of_thm_global thy) pre_intros))
+ val T = Sign.the_const_type thy constname
+ fun constname_of_intro intr = fst (dest_Const (fst (strip_intro_concl 0 (prop_of intr))))
+ val _ = if not (forall (fn intr => constname_of_intro intr = constname) pre_intros) then
+ error ("register_intros: Introduction rules of different constants are used\n" ^
+ "expected rules for " ^ constname ^ ", but received rules for " ^
+ commas (map constname_of_intro pre_intros))
+ else ()
+ val pred = Const (constname, T)
val nparams = guess_nparams T
val pre_elim =
(Drule.standard o Skip_Proof.make_thm thy)
- (mk_casesrule (ProofContext.init thy) nparams pre_intros)
- in register_predicate (pre_intros, pre_elim, nparams) thy end
+ (mk_casesrule (ProofContext.init thy) pred nparams pre_intros)
+ in register_predicate (constname, pre_intros, pre_elim, nparams) thy end
fun set_generator_name pred mode name =
let
@@ -672,7 +713,7 @@
PredData.map (Graph.map_node pred (map_pred_data set))
end
-fun set_sizelim_function_name pred mode name =
+fun set_depth_limited_function_name pred mode name =
let
val set = (apsnd o aptrd3 o cons) (mode, mk_function_data (name, NONE))
in
@@ -694,8 +735,6 @@
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
};
@@ -708,8 +747,6 @@
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
@@ -719,7 +756,7 @@
val (paramTs, (inargTs, outargTs)) = split_modeT (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))
+ (paramTs' @ inargTs) ---> (mk_predT compfuns (HOLogic.mk_tupleT outargTs))
end;
fun mk_fun_of compfuns thy (name, T) mode =
@@ -809,9 +846,11 @@
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_not t = let val T = mk_rpredT HOLogic.unitT
+ in Const (@{const_name RPred.not_rpred}, T --> T) $ t end
-fun mk_map t = error "FIXME" (*FIXME*)
+fun mk_map T1 T2 tf tp = Const (@{const_name RPred.map},
+ (T1 --> T2) --> mk_rpredT T1 --> mk_rpredT T2) $ tf $ tp
fun lift_pred t =
let
@@ -839,20 +878,21 @@
RPredCompFuns.mk_rpredT T) $ random
end;
-fun sizelim_funT_of compfuns (iss, is) T =
+fun depth_limited_funT_of compfuns (iss, is) T =
let
val Ts = binder_types T
val (paramTs, (inargTs, outargTs)) = split_modeT (iss, is) Ts
- val paramTs' = map2 (fn SOME is => sizelim_funT_of PredicateCompFuns.compfuns ([], is) | NONE => I) iss paramTs
+ val paramTs' = map2 (fn SOME is => depth_limited_funT_of compfuns ([], is) | NONE => I) iss paramTs
in
- (paramTs' @ inargTs @ [@{typ "code_numeral"}]) ---> (mk_predT compfuns (mk_tupleT outargTs))
+ (paramTs' @ inargTs @ [@{typ bool}, @{typ "code_numeral"}])
+ ---> (mk_predT compfuns (HOLogic.mk_tupleT outargTs))
end;
-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_depth_limited_fun_of compfuns thy (name, T) mode =
+ Const (depth_limited_function_name_of thy name mode, depth_limited_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)
+ Const (generator_name_of thy name mode, depth_limited_funT_of compfuns mode T)
(* Mode analysis *)
@@ -882,16 +922,16 @@
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 [[]];
+fun subsets i j =
+ if i <= j then
+ let
+ 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;
+ val is = subsets (i+1) j
+ in merge (map (fn ks => i::ks) is) is end
+ else [[]];
(* FIXME: should be in library - cprod = map_prod I *)
fun cprod ([], ys) = []
@@ -905,46 +945,9 @@
val yss = (cprods_subset xss)
in maps (fn ys => map (fn x => cons x ys) xs) yss @ yss end
-(*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 = map_index (fn (i, T) => (i, NONE)) (binder_types (fastype_of t));
+ val ks = map_index (fn (i, T) => (i + 1, NONE)) (binder_types (fastype_of t));
val default = [Mode (([], ks), ks, [])];
fun mk_modes name args = Option.map (maps (fn (m as (iss, is)) =>
let
@@ -956,7 +959,7 @@
val prfx = map (rpair NONE) (1 upto k)
in
if not (is_prefix op = prfx is) then [] else
- let val is' = List.drop (is, k)
+ let val is' = map (fn (i, t) => (i - k, t)) (List.drop (is, k))
in map (fn x => Mode (m, is', x)) (cprods (map
(fn (NONE, _) => [NONE]
| (SOME js, arg) => map SOME (filter
@@ -992,7 +995,7 @@
(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
+ is = map (rpair NONE) (1 upto length us) andalso
subset (op =) (terms_vs us, vs) andalso
subset (op =) (term_vs t, vs))
(modes_of_term modes t handle Option =>
@@ -1015,24 +1018,8 @@
(Generator (v, T), Mode (([], []), [], []))
end;
-fun gen_prem (Prem (us, t)) = GeneratorPrem (us, t)
- | gen_prem (Negprem (us, t)) = error "it is a negated prem"
- | gen_prem (Sidecond t) = error "it is a sidecond"
- | 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 _ = tracing ("param_vs:" ^ commas param_vs)
- val _ = tracing ("iss:" ^
- commas (map (fn is => case is of SOME is => string_of_smode is | NONE => "NONE") iss))
- *)
val modes' = modes @ map_filter
(fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
(param_vs ~~ iss);
@@ -1046,7 +1033,7 @@
NONE =>
(if with_generator then
(case select_mode_prem thy gen_modes' vs ps of
- SOME (p as Prem _, SOME mode) => check_mode_prems ((gen_prem p, mode) :: acc_ps)
+ SOME (p as Prem _, SOME mode) => check_mode_prems ((p, mode) :: acc_ps)
(case p of Prem (us, _) => union (op =) vs (terms_vs us) | _ => vs)
(filter_out (equal p) ps)
| _ =>
@@ -1057,14 +1044,11 @@
(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 => let
- val _ = tracing ("ps:" ^ (commas
- (map (fn p => string_of_moded_prem thy (p, Mode (([], []), [], []))) ps)))
- in (*error "mode analysis failed"*)NONE end
+ | 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)
+ | SOME (p, SOME mode) => check_mode_prems ((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));
@@ -1083,33 +1067,41 @@
else NONE
end;
-fun check_modes_pred with_generator thy param_vs clauses modes gen_modes (p, ms) =
- let val SOME rs = AList.lookup (op =) clauses p
+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 " ^
+ p ^ " violates mode " ^ string_of_mode m)
+ val _ = Output.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
(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);
- tracing (commas (map (Syntax.string_of_term_global thy) (fst (nth rs i)))); false)) ms)
+ | i => (print_failed_mode options thy modes p m rs i; false)) ms)
end;
fun get_modes_pred with_generator thy param_vs clauses modes gen_modes (p, ms) =
let
- val SOME rs = AList.lookup (op =) clauses p
+ val rs = case AList.lookup (op =) clauses p of SOME rs => rs | NONE => []
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 infer_modes thy extra_modes all_modes param_vs clauses =
+fun infer_modes options thy extra_modes all_modes param_vs clauses =
let
val modes =
fixp (fn modes =>
- map (check_modes_pred false thy param_vs clauses (modes @ extra_modes) []) modes)
+ map (check_modes_pred options false thy param_vs clauses (modes @ extra_modes) []) modes)
all_modes
in
map (get_modes_pred false thy param_vs clauses (modes @ extra_modes) []) modes
@@ -1122,19 +1114,24 @@
| SOME vs' => (k, subtract (op =) vs' vs))
:: remove_from rem xs
-fun infer_modes_with_generator thy extra_modes all_modes param_vs clauses =
+fun infer_modes_with_generator options thy extra_modes all_modes param_vs clauses =
let
val prednames = map fst clauses
- val extra_modes = all_modes_of thy
+ 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 all_modes
+ val starting_modes = remove_from extra_modes' all_modes
+ fun eq_mode (m1, m2) = (m1 = m2)
val modes =
fixp (fn modes =>
- map (check_modes_pred true thy param_vs clauses extra_modes (gen_modes @ modes)) modes)
- starting_modes
+ map (check_modes_pred options true thy param_vs clauses extra_modes' (gen_modes @ modes)) modes)
+ starting_modes
in
- map (get_modes_pred true thy param_vs clauses extra_modes (gen_modes @ modes)) modes
+ AList.join (op =)
+ (fn _ => fn ((mps1, mps2)) =>
+ merge (fn ((m1, _), (m2, _)) => eq_mode (m1, m2)) (mps1, mps2))
+ (infer_modes options thy extra_modes all_modes param_vs clauses,
+ map (get_modes_pred true thy param_vs clauses extra_modes (gen_modes @ modes)) modes)
end;
(* term construction *)
@@ -1157,136 +1154,10 @@
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) DatatypeCase.Quiet [] 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 neg_in_sizelim size thy compfuns (NONE, t) = t
- | compile_param neg_in_sizelim 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 neg_in_sizelim 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) =>
- case neg_in_sizelim of
- SOME _ => Free (name, sizelim_funT_of compfuns (iss, is') T)
- | NONE => Free (name, funT_of compfuns (iss, is') T)
-
- | _ => error ("PredicateCompiler: illegal parameter term")
- in
- (case neg_in_sizelim of SOME size_t =>
- (fn t =>
- let
- val Ts = fst (split_last (binder_types (fastype_of t)))
- val names = map (fn i => "x" ^ string_of_int i) (1 upto length Ts)
- in
- list_abs (names ~~ Ts, list_comb (t, (map Bound ((length Ts) - 1 downto 0)) @ [size_t]))
- end)
- | NONE => I)
- (list_comb (f', params' @ args'))
- end
-
-fun compile_expr neg_in_sizelim size thy ((Mode (mode, is, ms)), t) =
- case strip_comb t of
- (Const (name, T), params) =>
- let
- val params' = map (compile_param neg_in_sizelim 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) inargs =
- case strip_comb t of
- (Const (name, T), params) =>
- let
- val params' = map (compile_param NONE size thy PredicateCompFuns.compfuns) (ms ~~ params)
- in
- list_comb (mk_generator_of compfuns thy (name, T) mode, params' @ inargs)
- end
- | (Free (name, T), params) =>
- lift_pred compfuns
- (list_comb (Free (name, sizelim_funT_of PredicateCompFuns.compfuns ([], is) T), params @ inargs))
-
-
(** specific rpred functions -- move them to the correct place in this file *)
-fun mk_Eval_of size ((x, T), NONE) names = (x, names)
- | mk_Eval_of size ((x, T), SOME mode) names =
+fun mk_Eval_of additional_arguments ((x, T), NONE) names = (x, names)
+ | mk_Eval_of additional_arguments ((x, T), SOME mode) names =
let
val Ts = binder_types T
(*val argnames = Name.variant_list names
@@ -1331,32 +1202,99 @@
end
val (inoutargs, args) = split_list (map mk_arg (1 upto (length Ts) ~~ Ts))
val (inargs, outargs) = pairself flat (split_list inoutargs)
- val size_t = case size of NONE => [] | SOME size_t => [size_t]
- val r = PredicateCompFuns.mk_Eval (list_comb (x, inargs @ size_t), mk_tuple outargs)
+ val r = PredicateCompFuns.mk_Eval
+ (list_comb (x, inargs @ additional_arguments), HOLogic.mk_tuple outargs)
val t = fold_rev mk_split_lambda args r
in
(t, names)
end;
-fun compile_arg size thy param_vs iss arg =
+fun compile_arg compilation_modifiers compfuns additional_arguments thy param_vs iss arg =
let
- val funT_of = case size of NONE => funT_of | SOME _ => sizelim_funT_of
fun map_params (t as Free (f, T)) =
if member (op =) param_vs f then
case (the (AList.lookup (op =) (param_vs ~~ iss) f)) of
- SOME is => let val T' = funT_of PredicateCompFuns.compfuns ([], is) T
- in fst (mk_Eval_of size ((Free (f, T'), T), SOME is) []) end
+ SOME is =>
+ let
+ val T' = #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
| map_params t = t
in map_aterms map_params arg end
-
-fun compile_clause compfuns size final_term thy all_vs param_vs (iss, is) inp (ts, moded_ps) =
+
+fun compile_match compilation_modifiers compfuns additional_arguments param_vs iss thy eqs eqs' out_ts success_t =
+ let
+ val eqs'' = maps mk_eq eqs @ eqs'
+ val eqs'' =
+ map (compile_arg compilation_modifiers compfuns additional_arguments thy param_vs iss) 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 = HOLogic.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) DatatypeCase.Quiet [] v
+ [(HOLogic.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 compilation_modifiers compfuns thy (NONE, t) = t
+ | compile_param compilation_modifiers compfuns thy (m as SOME (Mode (mode, _, ms)), t) =
+ let
+ val (f, args) = strip_comb (Envir.eta_contract t)
+ val (params, args') = chop (length ms) args
+ 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)
+ | _ => error ("PredicateCompiler: illegal parameter term")
+ in
+ list_comb (f', params' @ args')
+ end
+
+fun compile_expr compilation_modifiers compfuns thy ((Mode (mode, _, ms)), t) inargs additional_arguments =
+ case strip_comb t of
+ (Const (name, T), params) =>
+ 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
+ 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)
+
+fun compile_clause compilation_modifiers compfuns thy all_vs param_vs additional_arguments (iss, is) inp (ts, moded_ps) =
+ let
+ val compile_match = compile_match compilation_modifiers compfuns additional_arguments param_vs iss thy
fun check_constrt t (names, eqs) =
if is_constrt thy t then (t, (names, eqs)) else
let
- val s = Name.variant names "x";
+ val s = Name.variant names "x"
val v = Free (s, fastype_of t)
in (v, (s::names, HOLogic.mk_eq (v, t)::eqs)) end;
@@ -1371,12 +1309,8 @@
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)
+ compile_match constr_vs (eqs @ eqs') out_ts'''
+ (mk_single compfuns (HOLogic.mk_tuple out_ts))
end
| compile_prems out_ts vs names ((p, mode as Mode ((_, is), _, _)) :: ps) =
let
@@ -1385,16 +1319,16 @@
fold_map check_constrt out_ts (names, [])
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
val (compiled_clause, rest) = case p of
Prem (us, t) =>
let
val (in_ts, out_ts''') = split_smode is us;
- val in_ts = map (compile_arg size thy param_vs iss) in_ts
- val args = case size of
- NONE => in_ts
- | SOME size_t => in_ts @ [size_t]
- val u = lift_pred compfuns
- (list_comb (compile_expr NONE size thy (mode, t), args))
+ val in_ts = map (compile_arg compilation_modifiers compfuns additional_arguments
+ thy param_vs iss) in_ts
+ val u =
+ compile_expr compilation_modifiers compfuns thy (mode, t) in_ts additional_arguments'
val rest = compile_prems out_ts''' vs' names'' ps
in
(u, rest)
@@ -1402,38 +1336,32 @@
| 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 size NONE thy (mode, t), in_ts)))
+ val in_ts = map (compile_arg compilation_modifiers compfuns additional_arguments
+ thy param_vs iss) in_ts
+ val u = mk_not compfuns
+ (compile_expr compilation_modifiers compfuns thy (mode, t) in_ts additional_arguments')
val rest = compile_prems out_ts''' vs' names'' ps
in
(u, rest)
end
| Sidecond t =>
let
+ val t = compile_arg compilation_modifiers compfuns additional_arguments
+ thy param_vs iss t
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 = 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 (the size))
+ val [size] = additional_arguments
+ val u = lift_random (HOLogic.mk_random T size)
val rest = compile_prems [Free (v, T)] vs' names'' ps;
in
(u, rest)
end
in
- compile_match thy compfuns constr_vs' eqs out_ts''
+ compile_match 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;
@@ -1441,14 +1369,13 @@
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 =
+fun compile_pred compilation_modifiers compfuns thy all_vs param_vs s T mode moded_cls =
let
val (Ts1, Ts2) = chop (length (fst mode)) (binder_types T)
val (Us1, Us2) = split_smodeT (snd mode) Ts2
- val funT_of = if use_size then sizelim_funT_of else funT_of
- val Ts1' = map2 (fn NONE => I | SOME is => funT_of PredicateCompFuns.compfuns ([], is)) (fst mode) Ts1
- val size_name = Name.variant (all_vs @ param_vs) "size"
- fun mk_input_term (i, NONE) =
+ val Ts1' =
+ map2 (fn NONE => I | SOME is => #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
[] => error "strange unit input"
@@ -1461,30 +1388,22 @@
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 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 additional_arguments = #additional_arguments compilation_modifiers (all_vs @ param_vs)
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 in_ts)) 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 @ in_ts @ [size]), size_t)
- else
- (list_comb (fun_const, params @ in_ts), t)
+ 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
+ (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)
in
- HOLogic.mk_Trueprop (HOLogic.mk_eq eq)
+ HOLogic.mk_Trueprop
+ (HOLogic.mk_eq (list_comb (fun_const, params @ in_ts @ additional_arguments), compilation))
end;
-
+
(* special setup for simpset *)
val HOL_basic_ss' = HOL_basic_ss addsimps (@{thms "HOL.simp_thms"} @ [@{thm Pair_eq}])
setSolver (mk_solver "all_tac_solver" (fn _ => fn _ => all_tac))
@@ -1531,15 +1450,15 @@
val param_names' = Name.variant_list (param_names @ 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 NONE) ((params ~~ Ts1) ~~ iss) []
+ val (params', names) = fold_map (mk_Eval_of []) ((params ~~ Ts1) ~~ iss) []
val predpropI = HOLogic.mk_Trueprop (list_comb (pred, param_vs @ args))
val predpropE = HOLogic.mk_Trueprop (list_comb (pred, params' @ args))
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))
+ if null outargs then Free("y", HOLogic.unitT) else HOLogic.mk_tuple outargs))
val funpropI = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs),
- mk_tuple outargs))
+ HOLogic.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"}, @{thm pair_collapse}]
@@ -1601,7 +1520,7 @@
val (Ts1, Ts2) = chop (length iss) Ts
val (Us1, Us2) = split_smodeT 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 funT = (Ts1' @ Us1) ---> (mk_predT compfuns (HOLogic.mk_tupleT Us2))
val names = Name.variant_list []
(map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
(* old *)
@@ -1634,7 +1553,7 @@
val (xinoutargs, names) = fold_map mk_vars ((1 upto (length Ts2)) ~~ Ts2) param_names
val (xinout, xargs) = split_list xinoutargs
val (xins, xouts) = pairself flat (split_list xinout)
- val (xparams', names') = fold_map (mk_Eval_of NONE) ((xparams ~~ Ts1) ~~ iss) names
+ 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 =
@@ -1663,16 +1582,16 @@
fold create_definition modes thy
end;
-fun sizelim_create_definitions preds (name, modes) thy =
+fun create_definitions_of_depth_limited_functions 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
+ val mode_cname = create_constname_of_mode thy "depth_limited_" name mode
+ val funT = depth_limited_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
+ |> set_depth_limited_function_name name mode mode_cname
end;
in
fold create_definition modes thy
@@ -1682,9 +1601,10 @@
let
val Ts = binder_types T
val (paramTs, (inargTs, outargTs)) = split_modeT (iss, is) Ts
- val paramTs' = map2 (fn SOME is => sizelim_funT_of PredicateCompFuns.compfuns ([], is) | NONE => I) iss paramTs
+ val paramTs' = map2 (fn SOME is => generator_funT_of ([], is) | NONE => I) iss paramTs
in
- (paramTs' @ inargTs @ [@{typ "code_numeral"}]) ---> (mk_predT RPredCompFuns.compfuns (mk_tupleT outargTs))
+ (paramTs' @ inargTs @ [@{typ code_numeral}]) --->
+ (mk_predT RPredCompFuns.compfuns (HOLogic.mk_tupleT outargTs))
end
fun rpred_create_definitions preds (name, modes) thy =
@@ -1696,7 +1616,7 @@
val funT = generator_funT_of 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
+ |> set_generator_name name mode mode_cname
end;
in
fold create_definition modes thy
@@ -1818,7 +1738,7 @@
(* need better control here! *)
end
-fun prove_clause thy nargs modes (iss, is) (_, clauses) (ts, moded_ps) =
+fun prove_clause options 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 [] =
@@ -1874,7 +1794,8 @@
end;
val prems_tac = prove_prems in_ts moded_ps
in
- rtac @{thm bindI} 1
+ print_tac' options "Proving clause..."
+ THEN rtac @{thm bindI} 1
THEN rtac @{thm singleI} 1
THEN prems_tac
end;
@@ -1883,20 +1804,20 @@
| 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 =
+fun prove_one_direction options 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 print_tac "before applying elim rule"
+ THEN print_tac' options "before applying elim rule"
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 (EVERY (map2 (prove_clause options thy nargs modes mode) clauses moded_clauses))
THEN print_tac "proved one direction"
end;
@@ -1914,15 +1835,20 @@
| _ => 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))
+ (print_tac ("Term " ^ (Syntax.string_of_term_global thy t) ^
+ "splitting with rules \n" ^
+ commas (map (Display.string_of_thm_global thy) split_rules)))
+ THEN TRY ((Splitter.split_asm_tac split_rules 1)
+ THEN (print_tac "after splitting with split_asm rules")
+ (* 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 (assert_tac (Max_number_of_subgoals 2))
THEN (EVERY (map split_term_tac ts))
end
else all_tac
in
- split_term_tac (mk_tuple out_ts)
+ split_term_tac (HOLogic.mk_tuple out_ts)
THEN (DETERM (TRY ((Splitter.split_asm_tac [@{thm "split_if_asm"}] 1) THEN (etac @{thm botE} 2))))
end
@@ -2053,7 +1979,7 @@
THEN prems_tac
end;
-fun prove_other_direction thy modes pred mode moded_clauses =
+fun prove_other_direction options 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)
@@ -2063,26 +1989,28 @@
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))))
+ THEN (if null moded_clauses then
+ etac @{thm botE} 1
+ else 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) =
+fun prove_pred options thy clauses preds modes pred mode (moded_clauses, compiled_term) =
let
val ctxt = ProofContext.init thy
- val clauses = the (AList.lookup (op =) clauses pred)
+ val clauses = case AList.lookup (op =) clauses pred of SOME rs => rs | NONE => []
in
Goal.prove ctxt (Term.add_free_names compiled_term []) [] compiled_term
- (if !do_proofs then
+ (if not (skip_proof options) then
(fn _ =>
rtac @{thm pred_iffI} 1
- THEN print_tac "after pred_iffI"
- 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 _ => Skip_Proof.cheat_tac thy)
+ THEN print_tac' options "after pred_iffI"
+ THEN prove_one_direction options thy clauses preds modes pred mode moded_clauses
+ THEN print_tac' options "proved one direction"
+ THEN prove_other_direction options thy modes pred mode moded_clauses
+ THEN print_tac' options "proved other direction")
+ else (fn _ => Skip_Proof.cheat_tac thy))
end;
(* composition of mode inference, definition, compilation and proof *)
@@ -2101,48 +2029,57 @@
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)
+fun compile_preds comp_modifiers compfuns thy all_vs param_vs preds moded_clauses =
+ map_preds_modes (fn pred => compile_pred comp_modifiers compfuns thy all_vs param_vs pred
+ (the (AList.lookup (op =) preds pred))) moded_clauses
+
+fun prove options thy clauses preds modes moded_clauses compiled_terms =
+ map_preds_modes (prove_pred options thy clauses preds modes)
(join_preds_modes moded_clauses compiled_terms)
-fun prove_by_skip thy _ _ _ _ compiled_terms =
+fun prove_by_skip options thy _ _ _ _ compiled_terms =
map_preds_modes (fn pred => fn mode => fn t => Drule.standard (Skip_Proof.make_thm thy t))
compiled_terms
+
+fun dest_prem thy params 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 thy params 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 prepare_intrs thy prednames =
+fun prepare_intrs thy prednames intros =
let
- val intrs = maps (intros_of thy) prednames
- |> map (Logic.unvarify o prop_of)
+ val intrs = map prop_of intros
val nparams = nparams_of thy (hd prednames)
+ val preds = map (fn c => Const (c, Sign.the_const_type thy c)) prednames
+ val (preds, intrs) = unify_consts thy preds intrs
+ val ([preds, intrs], _) = fold_burrow (Variable.import_terms false) [preds, intrs] (ProofContext.init thy)
+ val preds = map dest_Const preds
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 params = case intrs of
+ [] =>
+ let
+ val (paramTs, _) = chop nparams (binder_types (snd (hd preds)))
+ val param_names = Name.variant_list [] (map (fn i => "p" ^ string_of_int i) (1 upto length paramTs))
+ in map Free (param_names ~~ paramTs) end
+ | intr :: _ => fst (chop nparams
+ (snd (strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl intr)))))
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 prems = map (dest_prem thy params 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) @
@@ -2177,31 +2114,74 @@
val all_modes = map (fn (s, T) => (s, modes_of_typ T)) preds
in (preds, nparams, all_vs, param_vs, extra_modes, clauses, all_modes) end;
+fun check_format_of_intro_rule thy intro =
+ let
+ val concl = Logic.strip_imp_concl (prop_of intro)
+ val (p, args) = strip_comb (HOLogic.dest_Trueprop concl)
+ val params = List.take (args, nparams_of thy (fst (dest_Const p)))
+ fun check_arg arg = case HOLogic.strip_tupleT (fastype_of arg) of
+ (Ts as _ :: _ :: _) =>
+ if (length (HOLogic.strip_tuple arg) = length Ts) then true
+ else
+ error ("Format of introduction rule is invalid: tuples must be expanded:"
+ ^ (Syntax.string_of_term_global thy arg) ^ " in " ^
+ (Display.string_of_thm_global thy intro))
+ | _ => true
+ val prems = Logic.strip_imp_prems (prop_of intro)
+ fun check_prem (Prem (args, _)) = forall check_arg args
+ | check_prem (Negprem (args, _)) = forall check_arg args
+ | check_prem _ = true
+ in
+ forall check_arg args andalso
+ forall (check_prem o dest_prem thy params o HOLogic.dest_Trueprop) prems
+ end
+
+(*
+fun check_intros_elim_match thy prednames =
+ let
+ fun check predname =
+ let
+ val intros = intros_of thy predname
+ val elim = the_elim_of thy predname
+ val nparams = nparams_of thy predname
+ val elim' =
+ (Drule.standard o (Skip_Proof.make_thm thy))
+ (mk_casesrule (ProofContext.init thy) nparams intros)
+ in
+ if not (Thm.equiv_thm (elim, elim')) then
+ error "Introduction and elimination rules do not match!"
+ else true
+ end
+ in forall check prednames end
+*)
+
(** main function of predicate compiler **)
-fun add_equations_of steps prednames thy =
+fun add_equations_of steps options prednames thy =
let
- val _ = tracing ("Starting predicate compiler for predicates " ^ commas prednames ^ "...")
+ 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
- val _ = tracing "Infering modes..."
- val moded_clauses = #infer_modes steps thy extra_modes all_modes param_vs clauses
+ 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 modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
+ val _ = check_expected_modes options modes
val _ = print_modes modes
- val _ = print_moded_clauses thy moded_clauses
- val _ = tracing "Defining executable functions..."
+ (*val _ = print_moded_clauses thy moded_clauses*)
+ val _ = print_step options "Defining executable functions..."
val thy' = fold (#create_definitions steps preds) modes thy
|> Theory.checkpoint
- val _ = tracing "Compiling equations..."
+ val _ = print_step options "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)
+ 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)
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
@@ -2226,7 +2206,7 @@
fun extend value_of edges_of key G = fst (extend' value_of edges_of key (G, []))
-fun gen_add_equations steps names thy =
+fun gen_add_equations steps options names thy =
let
val thy' = PredData.map (fold (extend (fetch_pred_data thy) (depending_preds_of thy)) names) thy
|> Theory.checkpoint;
@@ -2235,33 +2215,83 @@
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)
+ if #are_not_defined 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,
+ additional_arguments = K [],
+ wrap_compilation = K (K (K (K (K I)))),
+ transform_additional_arguments = K I
+ }
+
+val depth_limited_comp_modifiers =
+ {const_name_of = depth_limited_function_name_of,
+ funT_of = depth_limited_funT_of,
+ additional_arguments = fn names =>
+ let
+ val [depth_name, polarity_name] = Name.variant_list names ["depth", "polarity"]
+ in [Free (polarity_name, @{typ "bool"}), Free (depth_name, @{typ "code_numeral"})] end,
+ wrap_compilation =
+ fn compfuns => fn s => fn T => fn mode => fn additional_arguments => fn compilation =>
+ let
+ val [polarity, depth] = additional_arguments
+ val (_, Ts2) = chop (length (fst mode)) (binder_types T)
+ val (_, Us2) = split_smodeT (snd mode) Ts2
+ val T' = mk_predT compfuns (HOLogic.mk_tupleT Us2)
+ val if_const = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T')
+ val full_mode = null Us2
+ in
+ if_const $ HOLogic.mk_eq (depth, @{term "0 :: code_numeral"})
+ $ (if_const $ polarity $ mk_bot compfuns (dest_predT compfuns T')
+ $ (if full_mode then mk_single compfuns HOLogic.unit else Const (@{const_name undefined}, T')))
+ $ compilation
+ end,
+ transform_additional_arguments =
+ fn prem => fn additional_arguments =>
+ let
+ val [polarity, depth] = additional_arguments
+ val polarity' = (case prem of Prem _ => I | Negprem _ => HOLogic.mk_not | _ => I) polarity
+ val depth' =
+ Const ("HOL.minus_class.minus", @{typ "code_numeral => code_numeral => code_numeral"})
+ $ depth $ Const ("HOL.one_class.one", @{typ "Code_Numeral.code_numeral"})
+ in [polarity', depth'] end
+ }
+
+val rpred_comp_modifiers =
+ {const_name_of = generator_name_of,
+ funT_of = K generator_funT_of,
+ 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
+ }
+
+
val add_equations = gen_add_equations
{infer_modes = infer_modes,
create_definitions = create_definitions,
- compile_preds = compile_preds PredicateCompFuns.compfuns mk_fun_of false,
+ compile_preds = compile_preds predicate_comp_modifiers PredicateCompFuns.compfuns,
prove = prove,
are_not_defined = fn thy => forall (null o modes_of thy),
qname = "equation"}
-val add_sizelim_equations = gen_add_equations
+val add_depth_limited_equations = gen_add_equations
{infer_modes = infer_modes,
- create_definitions = sizelim_create_definitions,
- compile_preds = compile_preds PredicateCompFuns.compfuns mk_sizelim_fun_of true,
+ 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 sizelim_modes_of thy),
- qname = "sizelim_equation"
- }
+ are_not_defined = fn thy => forall (null o depth_limited_modes_of thy),
+ qname = "depth_limited_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,
+ compile_preds = compile_preds rpred_comp_modifiers RPredCompFuns.compfuns,
prove = prove_by_skip,
are_not_defined = fn thy => forall (null o rpred_modes_of thy),
qname = "rpred_equation"}
@@ -2283,15 +2313,11 @@
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?*)
(* TODO: make TheoryDataFun to GenericDataFun & remove duplication of local theory and theory *)
-fun generic_code_pred prep_const rpred raw_const lthy =
+fun generic_code_pred prep_const options raw_const lthy =
let
val thy = ProofContext.theory_of lthy
val const = prep_const thy raw_const
@@ -2302,9 +2328,11 @@
val preds = Graph.all_preds (PredData.get thy') [const] |> filter_out (has_elim thy')
fun mk_cases const =
let
+ val T = Sign.the_const_type thy const
+ val pred = Const (const, T)
val nparams = nparams_of thy' const
val intros = intros_of thy' const
- in mk_casesrule lthy' nparams intros end
+ in mk_casesrule lthy' pred nparams intros end
val cases_rules = map mk_cases preds
val cases =
map (fn case_rule => RuleCases.Case {fixes = [],
@@ -2320,11 +2348,14 @@
(ProofContext.init (ProofContext.theory_of goal_ctxt)) (map the_single thms)
in
goal_ctxt |> LocalTheory.theory (fold set_elim global_thms #>
- (if rpred then
- (add_equations [const] #>
- add_sizelim_equations [const] #> add_quickcheck_equations [const])
- else add_equations [const]))
- end
+ (if is_rpred options then
+ (add_equations options [const] #>
+ add_quickcheck_equations options [const])
+ else if is_depth_limited options then
+ add_depth_limited_equations options [const]
+ else
+ add_equations options [const]))
+ end
in
Proof.theorem_i NONE after_qed (map (single o (rpair [])) cases_rules) lthy''
end;
@@ -2335,9 +2366,11 @@
(* transformation for code generation *)
val eval_ref = Unsynchronized.ref (NONE : (unit -> term Predicate.pred) option);
+val random_eval_ref = Unsynchronized.ref (NONE : (unit -> int * int -> term Predicate.pred * (int * int)) option);
(*FIXME turn this into an LCF-guarded preprocessor for comprehensions*)
-fun analyze_compr thy t_compr =
+(* TODO: *)
+fun analyze_compr thy compfuns (depth_limit, random) 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);
@@ -2348,6 +2381,8 @@
(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 user_mode' = map (rpair NONE) user_mode
+ val all_modes_of = if random then all_generator_modes_of else all_modes_of
+ (*val compile_expr = if random then compile_gen_expr else compile_expr*)
val modes = filter (fn Mode (_, is, _) => is = user_mode')
(modes_of_term (all_modes_of thy) (list_comb (pred, params)));
val m = case modes
@@ -2357,7 +2392,17 @@
| 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 NONE thy (m, list_comb (pred, params)), inargs);
+ val additional_arguments =
+ case depth_limit of
+ NONE => (if random then [@{term "5 :: code_numeral"}] else [])
+ | SOME d => [@{term "True"}, HOLogic.mk_number @{typ "code_numeral"} d]
+ val comp_modifiers =
+ case depth_limit of NONE =>
+ (if random then rpred_comp_modifiers else predicate_comp_modifiers) | SOME _ => depth_limited_comp_modifiers
+ val mk_fun_of = if random then mk_generator_of else
+ if (is_some depth_limit) then mk_depth_limited_fun_of else mk_fun_of
+ val t_pred = compile_expr comp_modifiers compfuns thy
+ (m, list_comb (pred, params)) inargs additional_arguments;
val t_eval = if null outargs then t_pred else
let
val outargs_bounds = map (fn Bound i => i) outargs;
@@ -2370,22 +2415,30 @@
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 mk_map compfuns T_pred T_compr arrange t_pred end
in t_eval end;
-fun eval thy t_compr =
+fun eval thy (options as (depth_limit, random)) 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_Core.eval_ref", eval_ref) Predicate.map thy t' []) end;
+ val compfuns = if random then RPredCompFuns.compfuns else PredicateCompFuns.compfuns
+ val t = analyze_compr thy compfuns options t_compr;
+ val T = dest_predT compfuns (fastype_of t);
+ val t' = mk_map compfuns T HOLogic.termT (HOLogic.term_of_const T) t;
+ val eval =
+ if random then
+ Code_ML.eval NONE ("Predicate_Compile_Core.random_eval_ref", random_eval_ref)
+ (fn proc => fn g => fn s => g s |>> Predicate.map proc) thy t' []
+ |> Random_Engine.run
+ else
+ Code_ML.eval NONE ("Predicate_Compile_Core.eval_ref", eval_ref) Predicate.map thy t' []
+ in (T, eval) end;
-fun values ctxt k t_compr =
+fun values ctxt options k t_compr =
let
val thy = ProofContext.theory_of ctxt;
- val (T, t) = eval thy t_compr;
+ val (T, ts) = eval thy options t_compr;
+ val (ts, _) = Predicate.yieldn k ts;
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 Set.union}, setT --> setT --> setT) $ elemsT $ t_compr
@@ -2398,11 +2451,11 @@
in
end;
*)
-fun values_cmd modes k raw_t state =
+fun values_cmd modes options 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 t' = values ctxt options k t;
val ty' = Term.type_of t';
val ctxt' = Variable.auto_fixes t' ctxt;
val p = PrintMode.with_modes modes (fn () =>
@@ -2410,15 +2463,24 @@
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 _ = List.app OuterKeyword.keyword ["depth_limit", "random"]
+
+val options =
+ let
+ val depth_limit = Scan.optional (P.$$$ "depth_limit" |-- P.$$$ "=" |-- P.nat >> SOME) NONE
+ val random = Scan.optional (P.$$$ "random" >> K true) false
+ in
+ Scan.optional (P.$$$ "[" |-- depth_limit -- random --| P.$$$ "]") (NONE, false)
+ end
+
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)));
+ (opt_modes -- options -- Scan.optional P.nat ~1 -- P.term
+ >> (fn (((modes, options), k), t) => Toplevel.no_timing o Toplevel.keep
+ (values_cmd modes options k t)));
end;
--- a/src/HOL/Tools/inductive.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/inductive.ML Sun Oct 25 08:57:55 2009 +0100
@@ -144,7 +144,7 @@
val (tab, monos) = get_inductives ctxt;
val space = Consts.space_of (ProofContext.consts_of ctxt);
in
- [Pretty.strs ("(co)inductives:" :: map #1 (NameSpace.extern_table (space, tab))),
+ [Pretty.strs ("(co)inductives:" :: map #1 (Name_Space.extern_table (space, tab))),
Pretty.big_list "monotonicity rules:" (map (Display.pretty_thm ctxt) monos)]
|> Pretty.chunks |> Pretty.writeln
end;
--- a/src/HOL/Tools/quickcheck_generators.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/quickcheck_generators.ML Sun Oct 25 08:57:55 2009 +0100
@@ -317,7 +317,7 @@
|> fold (fn (v, sort) => Vartab.update ((v, 0), sort)) raw_vs
|> fold meet_random insts;
in SOME (fn (v, _) => (v, (the o Vartab.lookup vtab) (v, 0)))
- end handle CLASS_ERROR => NONE;
+ end handle Sorts.CLASS_ERROR _ => NONE;
fun ensure_random_datatype config raw_tycos thy =
let
--- a/src/HOL/Tools/record.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/record.ML Sun Oct 25 08:57:55 2009 +0100
@@ -1810,7 +1810,7 @@
fun record_definition (args, bname) parent (parents: parent_info list) raw_fields thy =
let
- val external_names = NameSpace.external_names (Sign.naming_of thy);
+ val external_names = Name_Space.external_names (Sign.naming_of thy);
val alphas = map fst args;
val name = Sign.full_bname thy bname;
--- a/src/HOL/Tools/res_atp.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/Tools/res_atp.ML Sun Oct 25 08:57:55 2009 +0100
@@ -355,7 +355,7 @@
if run_blacklist_filter andalso is_package_def name then I
else
let val xname = Facts.extern facts name in
- if NameSpace.is_hidden xname then I
+ if Name_Space.is_hidden xname then I
else cons (xname, filter_out ResAxioms.bad_for_atp ths)
end) facts [];
--- a/src/HOL/ex/Predicate_Compile.thy Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/ex/Predicate_Compile.thy Sun Oct 25 08:57:55 2009 +0100
@@ -12,6 +12,6 @@
begin
setup {* Predicate_Compile.setup *}
-setup {* Quickcheck.add_generator ("pred_compile", Pred_Compile_Quickcheck.quickcheck) *}
+setup {* Quickcheck.add_generator ("pred_compile", Predicate_Compile_Quickcheck.quickcheck) *}
end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Predicate_Compile_Alternative_Defs.thy Sun Oct 25 08:57:55 2009 +0100
@@ -0,0 +1,76 @@
+theory Predicate_Compile_Alternative_Defs
+imports Predicate_Compile
+begin
+
+section {* Set operations *}
+(*
+definition Empty where "Empty == {}"
+declare empty_def[symmetric, code_pred_inline]
+*)
+declare eq_reflection[OF empty_def, code_pred_inline]
+(*
+definition Union where "Union A B == A Un B"
+
+lemma [code_pred_intros]: "A x ==> Union A B x"
+and [code_pred_intros] : "B x ==> Union A B x"
+unfolding Union_def Un_def Collect_def mem_def by auto
+
+code_pred Union
+unfolding Union_def Un_def Collect_def mem_def by auto
+
+declare Union_def[symmetric, code_pred_inline]
+*)
+declare eq_reflection[OF Un_def, code_pred_inline]
+
+section {* Alternative list definitions *}
+
+subsection {* Alternative rules for set *}
+
+lemma set_ConsI1 [code_pred_intros]:
+ "set (x # xs) x"
+unfolding mem_def[symmetric, of _ x]
+by auto
+
+lemma set_ConsI2 [code_pred_intros]:
+ "set xs x ==> set (x' # xs) x"
+unfolding mem_def[symmetric, of _ x]
+by auto
+
+code_pred set
+proof -
+ case set
+ from this show thesis
+ apply (case_tac a1)
+ apply auto
+ unfolding mem_def[symmetric, of _ a2]
+ apply auto
+ unfolding mem_def
+ apply auto
+ done
+qed
+
+
+subsection {* Alternative rules for list_all2 *}
+
+lemma list_all2_NilI [code_pred_intros]: "list_all2 P [] []"
+by auto
+
+lemma list_all2_ConsI [code_pred_intros]: "list_all2 P xs ys ==> P x y ==> list_all2 P (x#xs) (y#ys)"
+by auto
+
+code_pred list_all2
+proof -
+ case list_all2
+ from this show thesis
+ apply -
+ apply (case_tac a1)
+ apply (case_tac a2)
+ apply auto
+ apply (case_tac a2)
+ apply auto
+ done
+qed
+
+
+
+end
\ No newline at end of file
--- a/src/HOL/ex/Predicate_Compile_ex.thy Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/ex/Predicate_Compile_ex.thy Sun Oct 25 08:57:55 2009 +0100
@@ -1,5 +1,5 @@
theory Predicate_Compile_ex
-imports Main Predicate_Compile
+imports Main Predicate_Compile_Alternative_Defs
begin
inductive even :: "nat \<Rightarrow> bool" and odd :: "nat \<Rightarrow> bool" where
@@ -7,66 +7,216 @@
| "even n \<Longrightarrow> odd (Suc n)"
| "odd n \<Longrightarrow> even (Suc n)"
-code_pred even .
+code_pred (mode: [], [1]) even .
+code_pred [depth_limited] even .
+code_pred [rpred] even .
thm odd.equation
thm even.equation
+thm odd.depth_limited_equation
+thm even.depth_limited_equation
+thm even.rpred_equation
+thm odd.rpred_equation
values "{x. even 2}"
values "{x. odd 2}"
values 10 "{n. even n}"
values 10 "{n. odd n}"
+values [depth_limit = 2] "{x. even 6}"
+values [depth_limit = 7] "{x. even 6}"
+values [depth_limit = 2] "{x. odd 7}"
+values [depth_limit = 8] "{x. odd 7}"
+
+values [depth_limit = 7] 10 "{n. even n}"
+
+definition odd' where "odd' x == \<not> even x"
+
+code_pred [inductify] odd' .
+code_pred [inductify, depth_limited] odd' .
+code_pred [inductify, rpred] odd' .
+
+thm odd'.depth_limited_equation
+values [depth_limit = 2] "{x. odd' 7}"
+values [depth_limit = 9] "{x. odd' 7}"
inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
"append [] xs xs"
| "append xs ys zs \<Longrightarrow> append (x # xs) ys (x # zs)"
-code_pred append .
-code_pred (inductify_all) (rpred) append .
+code_pred (mode: [1, 2], [3], [2, 3], [1, 3], [1, 2, 3]) append .
+code_pred [depth_limited] append .
+code_pred [rpred] append .
thm append.equation
+thm append.depth_limited_equation
thm append.rpred_equation
values "{(ys, xs). append xs ys [0, Suc 0, 2]}"
values "{zs. append [0, Suc 0, 2] [17, 8] zs}"
-values "{ys. append [0, Suc 0, 2] ys [0, Suc 0, 2, 17, 0,5]}"
+values "{ys. append [0, Suc 0, 2] ys [0, Suc 0, 2, 17, 0, 5]}"
+values [depth_limit = 3] "{(xs, ys). append xs ys [1, 2, 3, 4, (5::nat)]}"
+values [random] 15 "{(ys, zs). append [1::nat, 2] ys zs}"
+
+value [code] "Predicate.the (append_1_2 [0::int, 1, 2] [3, 4, 5])"
+value [code] "Predicate.the (append_3 ([]::int list))"
+
+subsection {* Tricky case with alternative rules *}
+
+inductive append2
+where
+ "append2 [] xs xs"
+| "append2 xs ys zs \<Longrightarrow> append2 (x # xs) ys (x # zs)"
+
+lemma append2_Nil: "append2 [] (xs::'b list) xs"
+ by (simp add: append2.intros(1))
+
+lemmas [code_pred_intros] = append2_Nil append2.intros(2)
+
+code_pred append2
+proof -
+ case append2
+ from append2.cases[OF append2(1)] append2(2-3) show thesis by blast
+qed
+
+subsection {* Tricky cases with tuples *}
+
+inductive zerozero :: "nat * nat => bool"
+where
+ "zerozero (0, 0)"
+
+code_pred zerozero .
+code_pred [rpred] zerozero .
+
+inductive tupled_append :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
+where
+ "tupled_append ([], xs, xs)"
+| "tupled_append (xs, ys, zs) \<Longrightarrow> tupled_append (x # xs, ys, x # zs)"
+
+code_pred tupled_append .
+code_pred [rpred] tupled_append .
+thm tupled_append.equation
+(*
+TODO: values with tupled modes
+values "{xs. tupled_append ([1,2,3], [4,5], xs)}"
+*)
+
+inductive tupled_append'
+where
+"tupled_append' ([], xs, xs)"
+| "[| ys = fst (xa, y); x # zs = snd (xa, y);
+ tupled_append' (xs, ys, zs) |] ==> tupled_append' (x # xs, xa, y)"
+
+code_pred tupled_append' .
+thm tupled_append'.equation
+
+inductive tupled_append'' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
+where
+ "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'' .
+thm tupled_append''.equation
+
+inductive tupled_append''' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
+where
+ "tupled_append''' ([], xs, xs)"
+| "yszs = (ys, zs) ==> tupled_append''' (xs, yszs) \<Longrightarrow> tupled_append''' (x # xs, ys, x # zs)"
+
+code_pred [inductify] tupled_append''' .
+thm tupled_append'''.equation
+
+inductive map_ofP :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
+where
+ "map_ofP ((a, b)#xs) a b"
+| "map_ofP xs a b \<Longrightarrow> map_ofP (x#xs) a b"
+
+code_pred (mode: [1], [1, 2], [1, 2, 3], [1, 3]) map_ofP .
+thm map_ofP.equation
+
+inductive filter1
+for P
+where
+ "filter1 P [] []"
+| "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 [depth_limited] filter1 .
+code_pred [rpred] filter1 .
+
+thm filter1.equation
+
+inductive filter2
+where
+ "filter2 P [] []"
+| "P x ==> filter2 P xs ys ==> filter2 P (x#xs) (x#ys)"
+| "\<not> P x ==> filter2 P xs ys ==> filter2 P (x#xs) ys"
+
+code_pred (mode: [1, 2, 3], [1, 2]) filter2 .
+code_pred [depth_limited] filter2 .
+code_pred [rpred] filter2 .
+thm filter2.equation
+thm filter2.rpred_equation
+
+inductive filter3
+for P
+where
+ "List.filter P xs = ys ==> filter3 P xs ys"
+
+code_pred filter3 .
+code_pred [depth_limited] filter3 .
+thm filter3.depth_limited_equation
+(*code_pred [rpred] filter3 .*)
+inductive filter4
+where
+ "List.filter P xs = ys ==> filter4 P xs ys"
+
+code_pred filter4 .
+code_pred [depth_limited] filter4 .
+code_pred [rpred] filter4 .
+
+section {* reverse *}
inductive rev where
"rev [] []"
| "rev xs xs' ==> append xs' [x] ys ==> rev (x#xs) ys"
-code_pred rev .
+code_pred (mode: [1], [2], [1, 2]) rev .
thm rev.equation
values "{xs. rev [0, 1, 2, 3::nat] xs}"
+inductive tupled_rev where
+ "tupled_rev ([], [])"
+| "tupled_rev (xs, xs') \<Longrightarrow> tupled_append (xs', [x], ys) \<Longrightarrow> tupled_rev (x#xs, ys)"
+
+code_pred tupled_rev .
+thm tupled_rev.equation
+
inductive partition :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
for f where
"partition f [] [] []"
| "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 partition .
+code_pred (mode: [1], [2, 3], [1, 2], [1, 3], [1, 2, 3]) partition .
+code_pred [depth_limited] partition .
+code_pred [rpred] partition .
-thm partition.equation
+inductive tupled_partition :: "('a \<Rightarrow> bool) \<Rightarrow> ('a list \<times> 'a list \<times> 'a list) \<Rightarrow> bool"
+ for f where
+ "tupled_partition f ([], [], [])"
+ | "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 .
+
+thm tupled_partition.equation
-inductive member
-for xs
-where "x \<in> set xs ==> member xs x"
-
-lemma [code_pred_intros]:
- "member (x#xs') x"
-by (auto intro: member.intros)
-
-lemma [code_pred_intros]:
-"member xs x ==> member (x'#xs) x"
-by (auto intro: member.intros elim!: member.cases)
-(* strange bug must be repaired! *)
-(*
-code_pred member sorry
-*)
inductive is_even :: "nat \<Rightarrow> bool"
where
"n mod 2 = 0 \<Longrightarrow> is_even n"
@@ -88,18 +238,20 @@
case tranclp
from this converse_tranclpE[OF this(1)] show thesis by metis
qed
-(*
-code_pred (inductify_all) (rpred) tranclp .
+
+code_pred [depth_limited] tranclp .
+code_pred [rpred] tranclp .
thm tranclp.equation
thm tranclp.rpred_equation
-*)
+
inductive succ :: "nat \<Rightarrow> nat \<Rightarrow> bool" where
"succ 0 1"
| "succ m n \<Longrightarrow> succ (Suc m) (Suc n)"
code_pred succ .
-
+code_pred [rpred] succ .
thm succ.equation
+thm succ.rpred_equation
values 10 "{(m, n). succ n m}"
values "{m. succ 0 m}"
@@ -141,6 +293,16 @@
(While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1))))
[3,5] t}"
+inductive tupled_exec :: "(com \<times> state \<times> state) \<Rightarrow> bool" where
+"tupled_exec (Skip, s, s)" |
+"tupled_exec (Ass x e, s, s[x := e(s)])" |
+"tupled_exec (c1, s1, s2) ==> tupled_exec (c2, s2, s3) ==> tupled_exec (Seq c1 c2, s1, s3)" |
+"b s ==> tupled_exec (c1, s, t) ==> tupled_exec (IF b c1 c2, s, t)" |
+"~b s ==> tupled_exec (c2, s, t) ==> tupled_exec (IF b c1 c2, s, t)" |
+"~b s ==> tupled_exec (While b c, s, s)" |
+"b s1 ==> tupled_exec (c, s1, s2) ==> tupled_exec (While b c, s2, s3) ==> tupled_exec (While b c, s1, s3)"
+
+code_pred tupled_exec .
subsection{* CCS *}
@@ -171,6 +333,17 @@
values 3 "{(a,q). step (par nil nil) a q}"
*)
+inductive tupled_step :: "(proc \<times> nat \<times> proc) \<Rightarrow> bool"
+where
+"tupled_step (pre n p, n, p)" |
+"tupled_step (p1, a, q) \<Longrightarrow> tupled_step (or p1 p2, a, q)" |
+"tupled_step (p2, a, q) \<Longrightarrow> tupled_step (or p1 p2, a, q)" |
+"tupled_step (p1, a, q) \<Longrightarrow> tupled_step (par p1 p2, a, par q p2)" |
+"tupled_step (p2, a, q) \<Longrightarrow> tupled_step (par p1 p2, a, par p1 q)"
+
+code_pred tupled_step .
+thm tupled_step.equation
+
subsection {* divmod *}
inductive divmod_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
@@ -179,52 +352,75 @@
code_pred divmod_rel ..
-value [code] "Predicate.singleton (divmod_rel_1_2 1705 42)"
+value [code] "Predicate.the (divmod_rel_1_2 1705 42)"
section {* Executing definitions *}
definition Min
where "Min s r x \<equiv> s x \<and> (\<forall>y. r x y \<longrightarrow> x = y)"
-code_pred (inductify_all) Min .
+code_pred [inductify] Min .
subsection {* Examples with lists *}
-inductive filterP for Pa where
-"(filterP::('a => bool) => 'a list => 'a list => bool) (Pa::'a => bool) [] []"
-| "[| (res::'a list) = (y::'a) # (resa::'a list); (filterP::('a => bool) => 'a list => 'a list => bool) (Pa::'a => bool) (xt::'a list) resa; Pa y |]
-==> filterP Pa (y # xt) res"
-| "[| (filterP::('a => bool) => 'a list => 'a list => bool) (Pa::'a => bool) (xt::'a list) (res::'a list); ~ Pa (y::'a) |] ==> filterP Pa (y # xt) res"
+subsubsection {* Lexicographic order *}
+
+thm lexord_def
+code_pred [inductify] lexord .
+code_pred [inductify, rpred] lexord .
+thm lexord.equation
+thm lexord.rpred_equation
+
+inductive less_than_nat :: "nat * nat => bool"
+where
+ "less_than_nat (0, x)"
+| "less_than_nat (x, y) ==> less_than_nat (Suc x, Suc y)"
+
+code_pred less_than_nat .
+
+code_pred [depth_limited] less_than_nat .
+code_pred [rpred] less_than_nat .
+inductive test_lexord :: "nat list * nat list => bool"
+where
+ "lexord less_than_nat (xs, ys) ==> test_lexord (xs, ys)"
+
+code_pred [rpred] test_lexord .
+code_pred [depth_limited] test_lexord .
+thm test_lexord.depth_limited_equation
+thm test_lexord.rpred_equation
+
+values "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"
+values [depth_limit = 5] "{x. test_lexord ([1, 2, 3], [1, 2, 5])}"
+(*values [random] "{xys. test_lexord xys}"*)
+(*values [depth_limit = 5 random] "{xy. lexord less_than_nat xy}"*)
(*
-code_pred (inductify_all) (rpred) filterP .
-thm filterP.rpred_equation
+lemma "(u, v) : lexord less_than_nat ==> (x @ u, y @ v) : lexord less_than_nat"
+quickcheck[generator=pred_compile]
+oops
*)
-
-code_pred (inductify_all) lexord .
-
-thm lexord.equation
-
-lemma "(u, v) : lexord r ==> (x @ u, y @ v) : lexord r"
-(*quickcheck[generator=pred_compile]*)
-oops
-
lemmas [code_pred_def] = lexn_conv lex_conv lenlex_conv
-code_pred (inductify_all) lexn .
+code_pred [inductify] lexn .
thm lexn.equation
-code_pred (inductify_all) lenlex .
+code_pred [rpred] lexn .
+
+thm lexn.rpred_equation
+
+code_pred [inductify, show_steps] lenlex .
thm lenlex.equation
-(*
-code_pred (inductify_all) (rpred) lenlex .
+
+code_pred [inductify, rpred] lenlex .
thm lenlex.rpred_equation
-*)
+
thm lists.intros
-code_pred (inductify_all) lists .
+code_pred [inductify] lists .
thm lists.equation
+section {* AVL Tree Example *}
+
datatype 'a tree = ET | MKT 'a "'a tree" "'a tree" nat
fun height :: "'a tree => nat" where
"height ET = 0"
@@ -236,40 +432,107 @@
"avl (MKT x l r h) = ((height l = height r \<or> height l = 1 + height r \<or> height r = 1+height l) \<and>
h = max (height l) (height r) + 1 \<and> avl l \<and> avl r)"
-code_pred (inductify_all) avl .
+code_pred [inductify] avl .
thm avl.equation
-lemma [code_pred_inline]: "bot_fun_inst.bot_fun == (\<lambda>(y::'a::type). False)"
-unfolding bot_fun_inst.bot_fun[symmetric] bot_bool_eq[symmetric] bot_fun_eq by simp
+code_pred [rpred] avl .
+thm avl.rpred_equation
+(*values [random] 10 "{t. avl (t::int tree)}"*)
fun set_of
where
"set_of ET = {}"
| "set_of (MKT n l r h) = insert n (set_of l \<union> set_of r)"
-fun is_ord
+fun is_ord :: "nat tree => bool"
where
"is_ord ET = True"
| "is_ord (MKT n l r h) =
((\<forall>n' \<in> set_of l. n' < n) \<and> (\<forall>n' \<in> set_of r. n < n') \<and> is_ord l \<and> is_ord r)"
-declare Un_def[code_pred_def]
-
-code_pred (inductify_all) set_of .
+code_pred (mode: [1], [1, 2]) [inductify] set_of .
thm set_of.equation
-(* FIXME *)
-(*
-code_pred (inductify_all) is_ord .
+
+code_pred [inductify] is_ord .
thm is_ord.equation
-*)
+code_pred [rpred] is_ord .
+thm is_ord.rpred_equation
+
section {* Definitions about Relations *}
-code_pred (inductify_all) converse .
+code_pred [inductify] converse .
thm converse.equation
+code_pred [inductify] rel_comp .
+thm rel_comp.equation
+code_pred [inductify] Image .
+thm Image.equation
+(*TODO: *)
+ML {* Toplevel.debug := true *}
+declare Id_on_def[unfolded UNION_def, code_pred_def]
+
+code_pred [inductify] Id_on .
+thm Id_on.equation
+code_pred [inductify] Domain .
+thm Domain.equation
+code_pred [inductify] Range .
+thm sym_def
+code_pred [inductify] Field .
+declare Sigma_def[unfolded UNION_def, code_pred_def]
+declare refl_on_def[unfolded UNION_def, code_pred_def]
+code_pred [inductify] refl_on .
+thm refl_on.equation
+code_pred [inductify] total_on .
+thm total_on.equation
+(*
+code_pred [inductify] sym .
+thm sym.equation
+*)
+code_pred [inductify] antisym .
+thm antisym.equation
+code_pred [inductify] trans .
+thm trans.equation
+code_pred [inductify] single_valued .
+thm single_valued.equation
+code_pred [inductify] inv_image .
+thm inv_image.equation
-code_pred (inductify_all) Domain .
-thm Domain.equation
+section {* List functions *}
+
+code_pred [inductify] length .
+thm size_listP.equation
+code_pred [inductify, rpred] length .
+thm size_listP.rpred_equation
+values [random] 20 "{xs. size_listP (xs::nat list) (5::nat)}"
+code_pred [inductify] concat .
+code_pred [inductify] hd .
+code_pred [inductify] tl .
+code_pred [inductify] last .
+code_pred [inductify] butlast .
+(*code_pred [inductify] listsum .*)
+code_pred [inductify] take .
+code_pred [inductify] drop .
+code_pred [inductify] zip .
+code_pred [inductify] upt .
+code_pred [inductify] remdups .
+code_pred [inductify] remove1 .
+code_pred [inductify] removeAll .
+code_pred [inductify] distinct .
+code_pred [inductify] replicate .
+code_pred [inductify] splice .
+code_pred [inductify] List.rev .
+code_pred [inductify] map .
+code_pred [inductify] foldr .
+code_pred [inductify] foldl .
+code_pred [inductify] filter .
+code_pred [inductify, rpred] filter .
+thm filterP.rpred_equation
+
+definition test where "test xs = filter (\<lambda>x. x = (1::nat)) xs"
+code_pred [inductify] test .
+thm testP.equation
+code_pred [inductify, rpred] test .
+thm testP.rpred_equation
section {* Context Free Grammar *}
@@ -283,15 +546,86 @@
| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
-code_pred (inductify_all) S\<^isub>1p .
+code_pred [inductify] S\<^isub>1p .
+code_pred [inductify, rpred] S\<^isub>1p .
+thm S\<^isub>1p.equation
+thm S\<^isub>1p.rpred_equation
-thm S\<^isub>1p.equation
+values [random] 5 "{x. S\<^isub>1p x}"
+
+inductive is_a where
+ "is_a a"
+
+inductive is_b where
+ "is_b b"
-theorem S\<^isub>1_sound:
-"w \<in> S\<^isub>1 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-quickcheck[generator=pred_compile]
+code_pred is_a .
+code_pred [depth_limited] is_a .
+code_pred [rpred] is_a .
+
+values [random] "{x. is_a x}"
+code_pred [depth_limited] is_b .
+code_pred [rpred] is_b .
+
+code_pred [inductify, depth_limited] filter .
+
+values [depth_limit=5] "{x. filterP is_a [a, b] x}"
+values [depth_limit=3] "{x. filterP is_b [a, b] x}"
+
+lemma "w \<in> S\<^isub>1 \<Longrightarrow> length (filter (\<lambda>x. x = a) w) = 1"
+(*quickcheck[generator=pred_compile, size=10]*)
oops
+inductive test_lemma where
+ "S\<^isub>1p w ==> filterP is_a w r1 ==> size_listP r1 r2 ==> filterP is_b w r3 ==> size_listP r3 r4 ==> r2 \<noteq> r4 ==> test_lemma w"
+(*
+code_pred [rpred] test_lemma .
+*)
+(*
+definition test_lemma' where
+ "test_lemma' w == (w \<in> S\<^isub>1 \<and> (\<not> length [x <- w. x = a] = length [x <- w. x = b]))"
+
+code_pred [inductify, rpred] test_lemma' .
+thm test_lemma'.rpred_equation
+*)
+(*thm test_lemma'.rpred_equation*)
+(*
+values [depth_limit=3 random] "{x. S\<^isub>1 x}"
+*)
+code_pred [depth_limited] is_b .
+(*
+code_pred [inductify, depth_limited] filter .
+*)
+thm filterP.intros
+thm filterP.equation
+(*
+values [depth_limit=3] "{x. filterP is_b [a, b] x}"
+find_theorems "test_lemma'_hoaux"
+code_pred [depth_limited] test_lemma'_hoaux .
+thm test_lemma'_hoaux.depth_limited_equation
+values [depth_limit=2] "{x. test_lemma'_hoaux b}"
+inductive test1 where
+ "\<not> test_lemma'_hoaux x ==> test1 x"
+code_pred test1 .
+code_pred [depth_limited] test1 .
+thm test1.depth_limited_equation
+thm test_lemma'_hoaux.depth_limited_equation
+thm test1.intros
+
+values [depth_limit=2] "{x. test1 b}"
+
+thm filterP.intros
+thm filterP.depth_limited_equation
+values [depth_limit=3] "{x. filterP test_lemma'_hoaux [a, b] x}"
+values [depth_limit=4 random] "{w. test_lemma w}"
+values [depth_limit=4 random] "{w. test_lemma' w}"
+*)
+(*
+theorem S\<^isub>1_sound:
+"w \<in> S\<^isub>1p \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+quickcheck[generator=pred_compile, size=15]
+oops
+*)
inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
"[] \<in> S\<^isub>2"
| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
@@ -299,14 +633,18 @@
| "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
| "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
-(*
-code_pred (inductify_all) (rpred) S\<^isub>2 .
-ML {* Predicate_Compile_Core.intros_of @{theory} @{const_name "B\<^isub>2"} *}
-*)
+
+code_pred [inductify, rpred] S\<^isub>2 .
+thm S\<^isub>2.rpred_equation
+thm A\<^isub>2.rpred_equation
+thm B\<^isub>2.rpred_equation
+
+values [random] 10 "{x. S\<^isub>2 x}"
+
theorem S\<^isub>2_sound:
"w \<in> S\<^isub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
(*quickcheck[generator=SML]*)
-quickcheck[generator=pred_compile, size=15, iterations=100]
+(*quickcheck[generator=pred_compile, size=15, iterations=1]*)
oops
inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
@@ -317,23 +655,35 @@
| "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3"
| "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3"
+code_pred [inductify] S\<^isub>3 .
+thm S\<^isub>3.equation
+
+values 10 "{x. S\<^isub>3 x}"
(*
-code_pred (inductify_all) S\<^isub>3 .
-*)
theorem S\<^isub>3_sound:
"w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
quickcheck[generator=pred_compile, size=10, iterations=1]
oops
-
+*)
lemma "\<not> (length w > 2) \<or> \<not> (length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b])"
-quickcheck[size=10, generator = pred_compile]
+(*quickcheck[size=10, generator = pred_compile]*)
oops
+(*
+inductive test
+where
+ "length [x \<leftarrow> w. a = x] = length [x \<leftarrow> w. x = b] ==> test w"
+ML {* @{term "[x \<leftarrow> w. x = a]"} *}
+code_pred (inductify_all) test .
+thm test.equation
+*)
+(*
theorem S\<^isub>3_complete:
-"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>3"
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. b = x] \<longrightarrow> w \<in> S\<^isub>3"
(*quickcheck[generator=SML]*)
quickcheck[generator=pred_compile, size=10, iterations=100]
oops
+*)
inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
"[] \<in> S\<^isub>4"
@@ -343,15 +693,15 @@
| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
| "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"
-
+(*
theorem S\<^isub>4_sound:
"w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
quickcheck[generator = pred_compile, size=2, iterations=1]
oops
-
+*)
theorem S\<^isub>4_complete:
"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4"
-quickcheck[generator = pred_compile, size=5, iterations=1]
+(*quickcheck[generator = pred_compile, size=5, iterations=1]*)
oops
theorem S\<^isub>4_A\<^isub>4_B\<^isub>4_sound_and_complete:
@@ -361,8 +711,8 @@
(*quickcheck[generator = pred_compile, size=5, iterations=1]*)
oops
+section {* Lambda *}
-section {* Lambda *}
datatype type =
Atom nat
| Fun type type (infixr "\<Rightarrow>" 200)
@@ -378,15 +728,15 @@
"[]\<langle>i\<rangle> = None"
| "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
-(*
+
inductive nth_el' :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool"
where
"nth_el' (x # xs) 0 x"
| "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y"
-*)
+
inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" ("_ \<turnstile> _ : _" [50, 50, 50] 50)
where
- Var [intro!]: "nth_el env x = Some T \<Longrightarrow> env \<turnstile> Var x : T"
+ Var [intro!]: "nth_el' env x T \<Longrightarrow> env \<turnstile> Var x : T"
| Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)"
(* | App [intro!]: "env \<turnstile> s : T \<Rightarrow> U \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U" *)
| App [intro!]: "env \<turnstile> s : U \<Rightarrow> T \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U"
@@ -414,22 +764,8 @@
| abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"
lemma "Gamma \<turnstile> t : T \<Longrightarrow> t \<rightarrow>\<^sub>\<beta> t' \<Longrightarrow> Gamma \<turnstile> t' : T"
-quickcheck[generator = pred_compile, size = 10, iterations = 1000]
+(*quickcheck[generator = pred_compile, size = 10, iterations = 1]*)
oops
-(* FIXME *)
-(*
-inductive test for P where
-"[| filter P vs = res |]
-==> test P vs res"
-
-code_pred test .
-*)
-(*
-export_code test_for_1_yields_1_2 in SML file -
-code_pred (inductify_all) (rpred) test .
-
-thm test.equation
-*)
lemma filter_eq_ConsD:
"filter P ys = x#xs \<Longrightarrow>
--- a/src/HOL/ex/RPred.thy Sun Oct 25 08:57:36 2009 +0100
+++ b/src/HOL/ex/RPred.thy Sun Oct 25 08:57:55 2009 +0100
@@ -2,7 +2,7 @@
imports Quickcheck Random Predicate
begin
-types 'a rpred = "Random.seed \<Rightarrow> ('a Predicate.pred \<times> Random.seed)"
+types 'a "rpred" = "Random.seed \<Rightarrow> ('a Predicate.pred \<times> Random.seed)"
section {* The RandomPred Monad *}
@@ -33,9 +33,9 @@
(* Missing a good definition for negation: not_rpred *)
-definition not_rpred :: "unit Predicate.pred \<Rightarrow> unit rpred"
+definition not_rpred :: "unit rpred \<Rightarrow> unit rpred"
where
- "not_rpred = Pair o Predicate.not_pred"
+ "not_rpred P = (\<lambda>s. let (P', s') = P s in if Predicate.eval P' () then (Orderings.bot, s') else (Predicate.single (), s'))"
definition lift_pred :: "'a Predicate.pred \<Rightarrow> 'a rpred"
where
@@ -44,9 +44,9 @@
definition lift_random :: "(Random.seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed) \<Rightarrow> 'a rpred"
where "lift_random g = scomp g (Pair o (Predicate.single o fst))"
-definition map_rpred :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a rpred \<Rightarrow> 'b rpred)"
-where "map_rpred f P = bind P (return o f)"
+definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a rpred \<Rightarrow> 'b rpred)"
+ where "map f P = bind P (return o f)"
-hide (open) const return bind supp
+hide (open) const return bind supp map
end
\ No newline at end of file
--- a/src/Pure/General/markup.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/General/markup.ML Sun Oct 25 08:57:55 2009 +0100
@@ -18,6 +18,9 @@
val kindN: string
val internalK: string
val property_internal: Properties.property
+ val entityN: string val entity: string -> T
+ val defN: string
+ val refN: string
val lineN: string
val columnN: string
val offsetN: string
@@ -161,6 +164,14 @@
val property_internal = (kindN, internalK);
+(* formal entities *)
+
+val (entityN, entity) = markup_string "entity" nameN;
+
+val defN = "def";
+val refN = "ref";
+
+
(* position *)
val lineN = "line";
--- a/src/Pure/General/markup.scala Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/General/markup.scala Sun Oct 25 08:57:55 2009 +0100
@@ -15,6 +15,13 @@
val KIND = "kind"
+ /* formal entities */
+
+ val ENTITY = "entity"
+ val DEF = "def"
+ val REF = "ref"
+
+
/* position */
val LINE = "line"
--- a/src/Pure/General/name_space.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/General/name_space.ML Sun Oct 25 08:57:55 2009 +0100
@@ -20,7 +20,9 @@
val hidden: string -> string
val is_hidden: string -> bool
type T
- val empty: T
+ val empty: string -> T
+ val kind_of: T -> string
+ val the_entry: T -> string -> {is_system: bool, pos: Position.T, id: serial}
val intern: T -> xstring -> string
val extern: T -> string -> xstring
val extern_flags: {long_names: bool, short_names: bool, unique_names: bool} ->
@@ -29,7 +31,7 @@
val merge: T * T -> T
type naming
val default_naming: naming
- val declare: naming -> binding -> T -> string * T
+ val declare: bool -> naming -> binding -> T -> string * T
val full_name: naming -> binding -> string
val external_names: naming -> string -> string list
val add_path: string -> naming -> naming
@@ -37,16 +39,16 @@
val parent_path: naming -> naming
val mandatory_path: string -> naming -> naming
type 'a table = T * 'a Symtab.table
- val define: naming -> binding * 'a -> 'a table -> string * 'a table (*exception Symtab.DUP*)
- val empty_table: 'a table
- val merge_tables: ('a * 'a -> bool) -> 'a table * 'a table -> 'a table (*exception Symtab.DUP*)
- val join_tables: (string -> 'a * 'a -> 'a) (*exception Symtab.DUP/Symtab.SAME*) ->
- 'a table * 'a table -> 'a table (*exception Symtab.DUP*)
+ val define: bool -> naming -> binding * 'a -> 'a table -> string * 'a table
+ val empty_table: string -> 'a table
+ val merge_tables: 'a table * 'a table -> 'a table
+ val join_tables: (string -> 'a * 'a -> 'a) (*Symtab.SAME*) ->
+ 'a table * 'a table -> 'a table
val dest_table: 'a table -> (string * 'a) list
val extern_table: 'a table -> (xstring * 'a) list
end;
-structure NameSpace: NAME_SPACE =
+structure Name_Space: NAME_SPACE =
struct
@@ -58,33 +60,71 @@
val is_hidden = String.isPrefix "??.";
+(* datatype entry *)
+
+type entry =
+ {externals: xstring list,
+ is_system: bool,
+ pos: Position.T,
+ id: serial};
+
+fun str_of_entry def (name, {pos, id, ...}: entry) =
+ let
+ val occurrence = (if def then Markup.defN else Markup.refN, string_of_int id);
+ val props = occurrence :: Position.properties_of pos;
+ in Markup.markup (Markup.properties props (Markup.entity name)) name end;
+
+fun err_dup kind entry1 entry2 =
+ error ("Duplicate " ^ kind ^ " declaration " ^
+ quote (str_of_entry true entry1) ^ " vs. " ^ quote (str_of_entry true entry2));
+
+
(* datatype T *)
datatype T =
- NameSpace of
- (string list * string list) Symtab.table * (*internals, hidden internals*)
- xstring list Symtab.table; (*externals*)
+ Name_Space of
+ {kind: string,
+ internals: (string list * string list) Symtab.table, (*visible, hidden*)
+ entries: entry Symtab.table};
+
+fun make_name_space (kind, internals, entries) =
+ Name_Space {kind = kind, internals = internals, entries = entries};
+
+fun map_name_space f (Name_Space {kind = kind, internals = internals, entries = entries}) =
+ make_name_space (f (kind, internals, entries));
+
+fun map_internals f xname = map_name_space (fn (kind, internals, entries) =>
+ (kind, Symtab.map_default (xname, ([], [])) f internals, entries));
+
-val empty = NameSpace (Symtab.empty, Symtab.empty);
+fun empty kind = make_name_space (kind, Symtab.empty, Symtab.empty);
+
+fun kind_of (Name_Space {kind, ...}) = kind;
-fun lookup (NameSpace (tab, _)) xname =
- (case Symtab.lookup tab xname of
+fun the_entry (Name_Space {kind, entries, ...}) name =
+ (case Symtab.lookup entries name of
+ NONE => error ("Unknown " ^ kind ^ " " ^ quote name)
+ | SOME {is_system, pos, id, ...} => {is_system = is_system, pos = pos, id = id});
+
+
+(* name accesses *)
+
+fun lookup (Name_Space {internals, ...}) xname =
+ (case Symtab.lookup internals xname of
NONE => (xname, true)
| SOME ([], []) => (xname, true)
| SOME ([name], _) => (name, true)
| SOME (name :: _, _) => (name, false)
| SOME ([], name' :: _) => (hidden name', true));
-fun get_accesses (NameSpace (_, xtab)) name =
- (case Symtab.lookup xtab name of
+fun get_accesses (Name_Space {entries, ...}) name =
+ (case Symtab.lookup entries name of
NONE => [name]
- | SOME xnames => xnames);
+ | SOME {externals, ...} => externals);
-fun put_accesses name xnames (NameSpace (tab, xtab)) =
- NameSpace (tab, Symtab.update (name, xnames) xtab);
-
-fun valid_accesses (NameSpace (tab, _)) name = Symtab.fold (fn (xname, (names, _)) =>
- if not (null names) andalso hd names = name then cons xname else I) tab [];
+fun valid_accesses (Name_Space {internals, ...}) name =
+ Symtab.fold (fn (xname, (names, _)) =>
+ if not (null names) andalso hd names = name then cons xname else I) internals [];
(* intern and extern *)
@@ -116,21 +156,13 @@
unique_names = ! unique_names} space name;
-(* basic operations *)
-
-local
-
-fun map_space f xname (NameSpace (tab, xtab)) =
- NameSpace (Symtab.map_default (xname, ([], [])) f tab, xtab);
+(* modify internals *)
-in
-
-val del_name = map_space o apfst o remove (op =);
-fun del_name_extra name = map_space (apfst (fn [] => [] | x :: xs => x :: remove (op =) name xs));
-val add_name = map_space o apfst o update (op =);
-val add_name' = map_space o apsnd o update (op =);
-
-end;
+val del_name = map_internals o apfst o remove (op =);
+fun del_name_extra name =
+ map_internals (apfst (fn [] => [] | x :: xs => x :: remove (op =) name xs));
+val add_name = map_internals o apfst o update (op =);
+val add_name' = map_internals o apsnd o update (op =);
(* hide *)
@@ -152,17 +184,24 @@
(* merge *)
-fun merge (NameSpace (tab1, xtab1), NameSpace (tab2, xtab2)) =
+fun merge
+ (Name_Space {kind = kind1, internals = internals1, entries = entries1},
+ Name_Space {kind = kind2, internals = internals2, entries = entries2}) =
let
- val tab' = (tab1, tab2) |> Symtab.join
+ val kind' =
+ if kind1 = kind2 then kind1
+ else error ("Attempt to merge different kinds of name spaces " ^
+ quote kind1 ^ " vs. " ^ quote kind2);
+ val internals' = (internals1, internals2) |> Symtab.join
(K (fn ((names1, names1'), (names2, names2')) =>
- if pointer_eq (names1, names2) andalso pointer_eq (names1', names2') then raise Symtab.SAME
+ if pointer_eq (names1, names2) andalso pointer_eq (names1', names2')
+ then raise Symtab.SAME
else (Library.merge (op =) (names1, names2), Library.merge (op =) (names1', names2'))));
- val xtab' = (xtab1, xtab2) |> Symtab.join
- (K (fn xnames =>
- if pointer_eq xnames then raise Symtab.SAME
- else (Library.merge (op =) xnames)));
- in NameSpace (tab', xtab') end;
+ val entries' = (entries1, entries2) |> Symtab.join
+ (fn name => fn (entry1, entry2) =>
+ if #id entry1 = #id entry2 then raise Symtab.SAME
+ else err_dup kind' (name, entry1) (name, entry2));
+ in make_name_space (kind', internals', entries') end;
@@ -225,13 +264,27 @@
(* declaration *)
-fun declare naming binding space =
+fun new_entry strict entry =
+ map_name_space (fn (kind, internals, entries) =>
+ let
+ val entries' =
+ (if strict then Symtab.update_new else Symtab.update) entry entries
+ handle Symtab.DUP dup =>
+ err_dup kind (dup, the (Symtab.lookup entries dup)) entry;
+ in (kind, internals, entries') end);
+
+fun declare strict naming binding space =
let
val names = full naming binding;
val name = Long_Name.implode names;
val _ = name = "" andalso err_bad binding;
val (accs, accs') = accesses naming binding;
- val space' = space |> fold (add_name name) accs |> put_accesses name accs';
+ val entry =
+ {externals = accs',
+ is_system = false,
+ pos = Position.default (Binding.pos_of binding),
+ id = serial ()};
+ val space' = space |> fold (add_name name) accs |> new_entry strict (name, entry);
in (name, space') end;
@@ -240,14 +293,14 @@
type 'a table = T * 'a Symtab.table;
-fun define naming (binding, x) (space, tab) =
- let val (name, space') = declare naming binding space
- in (name, (space', Symtab.update_new (name, x) tab)) end;
+fun define strict naming (binding, x) (space, tab) =
+ let val (name, space') = declare strict naming binding space
+ in (name, (space', Symtab.update (name, x) tab)) end;
-val empty_table = (empty, Symtab.empty);
+fun empty_table kind = (empty kind, Symtab.empty);
-fun merge_tables eq ((space1, tab1), (space2, tab2)) =
- (merge (space1, space2), Symtab.merge eq (tab1, tab2));
+fun merge_tables ((space1, tab1), (space2, tab2)) =
+ (merge (space1, space2), Symtab.merge (K true) (tab1, tab2));
fun join_tables f ((space1, tab1), (space2, tab2)) =
(merge (space1, space2), Symtab.join f (tab1, tab2));
@@ -261,6 +314,6 @@
end;
-structure Basic_Name_Space: BASIC_NAME_SPACE = NameSpace;
+structure Basic_Name_Space: BASIC_NAME_SPACE = Name_Space;
open Basic_Name_Space;
--- a/src/Pure/General/position.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/General/position.ML Sun Oct 25 08:57:55 2009 +0100
@@ -37,6 +37,7 @@
val range: T -> T -> range
val thread_data: unit -> T
val setmp_thread_data: T -> ('a -> 'b) -> 'a -> 'b
+ val default: T -> T
end;
structure Position: POSITION =
@@ -178,4 +179,8 @@
end;
+fun default pos =
+ if pos = none then thread_data ()
+ else pos;
+
end;
--- a/src/Pure/Isar/attrib.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/Isar/attrib.ML Sun Oct 25 08:57:55 2009 +0100
@@ -67,35 +67,33 @@
structure Attributes = TheoryDataFun
(
- type T = (((src -> attribute) * string) * stamp) NameSpace.table;
- val empty = NameSpace.empty_table;
+ type T = ((src -> attribute) * string) Name_Space.table;
+ val empty = Name_Space.empty_table "attribute";
val copy = I;
val extend = I;
- fun merge _ tables : T = NameSpace.merge_tables (eq_snd (op =)) tables handle Symtab.DUP dup =>
- error ("Attempt to merge different versions of attribute " ^ quote dup);
+ fun merge _ tables : T = Name_Space.merge_tables tables;
);
fun print_attributes thy =
let
val attribs = Attributes.get thy;
- fun prt_attr (name, ((_, comment), _)) = Pretty.block
+ fun prt_attr (name, (_, comment)) = Pretty.block
[Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
in
- [Pretty.big_list "attributes:" (map prt_attr (NameSpace.extern_table attribs))]
+ [Pretty.big_list "attributes:" (map prt_attr (Name_Space.extern_table attribs))]
|> Pretty.chunks |> Pretty.writeln
end;
-fun add_attribute name att comment thy = thy |> Attributes.map (fn atts =>
- #2 (NameSpace.define (Sign.naming_of thy) (name, ((att, comment), stamp ())) atts)
- handle Symtab.DUP dup => error ("Duplicate declaration of attribute " ^ quote dup));
+fun add_attribute name att comment thy = thy
+ |> Attributes.map (#2 o Name_Space.define true (Sign.naming_of thy) (name, (att, comment)));
(* name space *)
-val intern = NameSpace.intern o #1 o Attributes.get;
+val intern = Name_Space.intern o #1 o Attributes.get;
val intern_src = Args.map_name o intern;
-val extern = NameSpace.extern o #1 o Attributes.get o ProofContext.theory_of;
+val extern = Name_Space.extern o #1 o Attributes.get o ProofContext.theory_of;
(* pretty printing *)
@@ -117,7 +115,7 @@
let val ((name, _), pos) = Args.dest_src src in
(case Symtab.lookup attrs name of
NONE => error ("Unknown attribute: " ^ quote name ^ Position.str_of pos)
- | SOME ((att, _), _) => (Position.report (Markup.attribute name) pos; att src))
+ | SOME (att, _) => (Position.report (Markup.attribute name) pos; att src))
end;
in attr end;
@@ -340,7 +338,7 @@
Pretty.block [Pretty.str (name ^ ": " ^ Config.print_type value ^ " ="), Pretty.brk 1,
Pretty.str (Config.print_value value)]
end;
- val configs = NameSpace.extern_table (#1 (Attributes.get thy), Configs.get thy);
+ val configs = Name_Space.extern_table (#1 (Attributes.get thy), Configs.get thy);
in Pretty.writeln (Pretty.big_list "configuration options" (map prt configs)) end;
--- a/src/Pure/Isar/isar_cmd.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/Isar/isar_cmd.ML Sun Oct 25 08:57:55 2009 +0100
@@ -400,7 +400,7 @@
val {classes = (space, algebra), ...} = Type.rep_tsig (Sign.tsig_of thy);
val {classes, ...} = Sorts.rep_algebra algebra;
fun entry (c, (i, (_, cs))) =
- (i, {name = NameSpace.extern space c, ID = c, parents = cs,
+ (i, {name = Name_Space.extern space c, ID = c, parents = cs,
dir = "", unfold = true, path = ""});
val gr =
Graph.fold (cons o entry) classes []
--- a/src/Pure/Isar/local_theory.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/Isar/local_theory.ML Sun Oct 25 08:57:55 2009 +0100
@@ -17,7 +17,7 @@
val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
val raw_theory: (theory -> theory) -> local_theory -> local_theory
val checkpoint: local_theory -> local_theory
- val full_naming: local_theory -> NameSpace.naming
+ val full_naming: local_theory -> Name_Space.naming
val full_name: local_theory -> binding -> string
val theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
val theory: (theory -> theory) -> local_theory -> local_theory
@@ -139,9 +139,9 @@
fun full_naming lthy =
Sign.naming_of (ProofContext.theory_of lthy)
- |> NameSpace.mandatory_path (#theory_prefix (get_lthy lthy));
+ |> Name_Space.mandatory_path (#theory_prefix (get_lthy lthy));
-fun full_name naming = NameSpace.full_name (full_naming naming);
+fun full_name naming = Name_Space.full_name (full_naming naming);
fun theory_result f lthy = lthy |> raw_theory_result (fn thy => thy
|> Sign.mandatory_path (#theory_prefix (get_lthy lthy))
--- a/src/Pure/Isar/locale.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/Isar/locale.ML Sun Oct 25 08:57:55 2009 +0100
@@ -124,15 +124,15 @@
structure Locales = TheoryDataFun
(
- type T = locale NameSpace.table;
- val empty = NameSpace.empty_table;
+ type T = locale Name_Space.table;
+ val empty = Name_Space.empty_table "locale";
val copy = I;
val extend = I;
- fun merge _ = NameSpace.join_tables (K merge_locale);
+ fun merge _ = Name_Space.join_tables (K merge_locale);
);
-val intern = NameSpace.intern o #1 o Locales.get;
-val extern = NameSpace.extern o #1 o Locales.get;
+val intern = Name_Space.intern o #1 o Locales.get;
+val extern = Name_Space.extern o #1 o Locales.get;
val get_locale = Symtab.lookup o #2 o Locales.get;
val defined = Symtab.defined o #2 o Locales.get;
@@ -143,7 +143,7 @@
| NONE => error ("Unknown locale " ^ quote name));
fun register_locale binding parameters spec intros axioms decls notes dependencies thy =
- thy |> Locales.map (NameSpace.define (Sign.naming_of thy)
+ thy |> Locales.map (Name_Space.define true (Sign.naming_of thy)
(binding,
mk_locale ((parameters, spec, intros, axioms),
((pairself (map (fn decl => (decl, stamp ()))) decls, map (fn n => (n, stamp ())) notes),
@@ -153,7 +153,7 @@
Locales.map o apsnd o Symtab.map_entry name o map_locale o apsnd;
fun print_locales thy =
- Pretty.strs ("locales:" :: map #1 (NameSpace.extern_table (Locales.get thy)))
+ Pretty.strs ("locales:" :: map #1 (Name_Space.extern_table (Locales.get thy)))
|> Pretty.writeln;
--- a/src/Pure/Isar/method.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/Isar/method.ML Sun Oct 25 08:57:55 2009 +0100
@@ -322,32 +322,30 @@
structure Methods = TheoryDataFun
(
- type T = (((src -> Proof.context -> method) * string) * stamp) NameSpace.table;
- val empty = NameSpace.empty_table;
+ type T = ((src -> Proof.context -> method) * string) Name_Space.table;
+ val empty = Name_Space.empty_table "method";
val copy = I;
val extend = I;
- fun merge _ tables : T = NameSpace.merge_tables (eq_snd (op =)) tables handle Symtab.DUP dup =>
- error ("Attempt to merge different versions of method " ^ quote dup);
+ fun merge _ tables : T = Name_Space.merge_tables tables;
);
fun print_methods thy =
let
val meths = Methods.get thy;
- fun prt_meth (name, ((_, comment), _)) = Pretty.block
+ fun prt_meth (name, (_, comment)) = Pretty.block
[Pretty.str (name ^ ":"), Pretty.brk 2, Pretty.str comment];
in
- [Pretty.big_list "methods:" (map prt_meth (NameSpace.extern_table meths))]
+ [Pretty.big_list "methods:" (map prt_meth (Name_Space.extern_table meths))]
|> Pretty.chunks |> Pretty.writeln
end;
-fun add_method name meth comment thy = thy |> Methods.map (fn meths =>
- #2 (NameSpace.define (Sign.naming_of thy) (name, ((meth, comment), stamp ())) meths)
- handle Symtab.DUP dup => error ("Duplicate declaration of method " ^ quote dup));
+fun add_method name meth comment thy = thy
+ |> Methods.map (#2 o Name_Space.define true (Sign.naming_of thy) (name, (meth, comment)));
(* get methods *)
-val intern = NameSpace.intern o #1 o Methods.get;
+val intern = Name_Space.intern o #1 o Methods.get;
val defined = Symtab.defined o #2 o Methods.get;
fun method_i thy =
@@ -357,11 +355,11 @@
let val ((name, _), pos) = Args.dest_src src in
(case Symtab.lookup meths name of
NONE => error ("Unknown proof method: " ^ quote name ^ Position.str_of pos)
- | SOME ((mth, _), _) => (Position.report (Markup.method name) pos; mth src))
+ | SOME (mth, _) => (Position.report (Markup.method name) pos; mth src))
end;
in meth end;
-fun method thy = method_i thy o Args.map_name (NameSpace.intern (#1 (Methods.get thy)));
+fun method thy = method_i thy o Args.map_name (Name_Space.intern (#1 (Methods.get thy)));
(* method setup *)
--- a/src/Pure/Isar/object_logic.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/Isar/object_logic.ML Sun Oct 25 08:57:55 2009 +0100
@@ -90,7 +90,7 @@
val base_sort = get_base_sort thy;
val b = Binding.map_name (Syntax.type_name mx) a;
val _ = has_duplicates (op =) vs andalso
- error ("Duplicate parameters in type declaration: " ^ quote (Binding.str_of b));
+ error ("Duplicate parameters in type declaration " ^ quote (Binding.str_of b));
val name = Sign.full_name thy b;
val n = length vs;
val T = Type (name, map (fn v => TFree (v, [])) vs);
--- a/src/Pure/Isar/proof_context.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/Isar/proof_context.ML Sun Oct 25 08:57:55 2009 +0100
@@ -21,7 +21,7 @@
val restore_mode: Proof.context -> Proof.context -> Proof.context
val abbrev_mode: Proof.context -> bool
val set_stmt: bool -> Proof.context -> Proof.context
- val naming_of: Proof.context -> NameSpace.naming
+ val naming_of: Proof.context -> Name_Space.naming
val full_name: Proof.context -> binding -> string
val consts_of: Proof.context -> Consts.T
val const_syntax_name: Proof.context -> string -> string
@@ -170,7 +170,7 @@
datatype ctxt =
Ctxt of
{mode: mode, (*inner syntax mode*)
- naming: NameSpace.naming, (*local naming conventions*)
+ naming: Name_Space.naming, (*local naming conventions*)
syntax: LocalSyntax.T, (*local syntax*)
consts: Consts.T * Consts.T, (*local/global consts*)
facts: Facts.T, (*local facts*)
@@ -180,7 +180,7 @@
Ctxt {mode = mode, naming = naming, syntax = syntax,
consts = consts, facts = facts, cases = cases};
-val local_naming = NameSpace.default_naming |> NameSpace.add_path "local";
+val local_naming = Name_Space.default_naming |> Name_Space.add_path "local";
structure ContextData = ProofDataFun
(
@@ -231,7 +231,7 @@
map_mode (fn (_, pattern, schematic, abbrev) => (stmt, pattern, schematic, abbrev));
val naming_of = #naming o rep_context;
-val full_name = NameSpace.full_name o naming_of;
+val full_name = Name_Space.full_name o naming_of;
val syntax_of = #syntax o rep_context;
val syn_of = LocalSyntax.syn_of o syntax_of;
@@ -924,10 +924,10 @@
(* name space operations *)
-val add_path = map_naming o NameSpace.add_path;
-val mandatory_path = map_naming o NameSpace.mandatory_path;
-val restore_naming = map_naming o K o naming_of;
-val reset_naming = map_naming (K local_naming);
+val add_path = map_naming o Name_Space.add_path;
+val mandatory_path = map_naming o Name_Space.mandatory_path;
+val restore_naming = map_naming o K o naming_of;
+val reset_naming = map_naming (K local_naming);
(* facts *)
@@ -1230,7 +1230,7 @@
| add_abbr (c, (T, SOME t)) =
if not show_globals andalso Symtab.defined globals c then I
else cons (c, Logic.mk_equals (Const (c, T), t));
- val abbrevs = NameSpace.extern_table (space, Symtab.make (Symtab.fold add_abbr consts []));
+ val abbrevs = Name_Space.extern_table (space, Symtab.make (Symtab.fold add_abbr consts []));
in
if null abbrevs andalso not (! verbose) then []
else [Pretty.big_list "abbreviations:" (map (pretty_term_abbrev ctxt o #2) abbrevs)]
--- a/src/Pure/Thy/thy_output.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/Thy/thy_output.ML Sun Oct 25 08:57:55 2009 +0100
@@ -420,9 +420,9 @@
("show_sorts", setmp_CRITICAL Syntax.show_sorts o boolean),
("show_structs", setmp_CRITICAL show_structs o boolean),
("show_question_marks", setmp_CRITICAL show_question_marks o boolean),
- ("long_names", setmp_CRITICAL NameSpace.long_names o boolean),
- ("short_names", setmp_CRITICAL NameSpace.short_names o boolean),
- ("unique_names", setmp_CRITICAL NameSpace.unique_names o boolean),
+ ("long_names", setmp_CRITICAL Name_Space.long_names o boolean),
+ ("short_names", setmp_CRITICAL Name_Space.short_names o boolean),
+ ("unique_names", setmp_CRITICAL Name_Space.unique_names o boolean),
("eta_contract", setmp_CRITICAL Syntax.eta_contract o boolean),
("display", setmp_CRITICAL display o boolean),
("break", setmp_CRITICAL break o boolean),
--- a/src/Pure/Tools/find_theorems.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/Tools/find_theorems.ML Sun Oct 25 08:57:55 2009 +0100
@@ -326,7 +326,7 @@
local
val index_ord = option_ord (K EQUAL);
-val hidden_ord = bool_ord o pairself NameSpace.is_hidden;
+val hidden_ord = bool_ord o pairself Name_Space.is_hidden;
val qual_ord = int_ord o pairself (length o Long_Name.explode);
val txt_ord = int_ord o pairself size;
@@ -355,7 +355,8 @@
val space = Facts.space_of (PureThy.facts_of (ProofContext.theory_of ctxt));
val shorten =
- NameSpace.extern_flags {long_names = false, short_names = false, unique_names = false} space;
+ Name_Space.extern_flags
+ {long_names = false, short_names = false, unique_names = false} space;
fun nicer (Facts.Named ((x, _), i)) (Facts.Named ((y, _), j)) =
nicer_name (shorten x, i) (shorten y, j)
--- a/src/Pure/codegen.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/codegen.ML Sun Oct 25 08:57:55 2009 +0100
@@ -337,15 +337,16 @@
val tc = Sign.intern_type thy s;
in
case Symtab.lookup (snd (#types (Type.rep_tsig (Sign.tsig_of thy)))) tc of
- SOME ((Type.LogicalType i, _), _) =>
+ SOME (Type.LogicalType i, _) =>
if num_args_of (fst syn) > i then
error ("More arguments than corresponding type constructor " ^ s)
- else (case AList.lookup (op =) types tc of
- NONE => CodegenData.put {codegens = codegens,
- tycodegens = tycodegens, consts = consts,
- types = (tc, syn) :: types,
- preprocs = preprocs, modules = modules} thy
- | SOME _ => error ("Type " ^ tc ^ " already associated with code"))
+ else
+ (case AList.lookup (op =) types tc of
+ NONE => CodegenData.put {codegens = codegens,
+ tycodegens = tycodegens, consts = consts,
+ types = (tc, syn) :: types,
+ preprocs = preprocs, modules = modules} thy
+ | SOME _ => error ("Type " ^ tc ^ " already associated with code"))
| _ => error ("Not a type constructor: " ^ s)
end;
--- a/src/Pure/consts.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/consts.ML Sun Oct 25 08:57:55 2009 +0100
@@ -11,15 +11,15 @@
val eq_consts: T * T -> bool
val retrieve_abbrevs: T -> string list -> term -> (term * term) list
val dest: T ->
- {constants: (typ * term option) NameSpace.table,
- constraints: typ NameSpace.table}
+ {constants: (typ * term option) Name_Space.table,
+ constraints: typ Name_Space.table}
val the_type: T -> string -> typ (*exception TYPE*)
val the_abbreviation: T -> string -> typ * term (*exception TYPE*)
val type_scheme: T -> string -> typ (*exception TYPE*)
val the_tags: T -> string -> Properties.T (*exception TYPE*)
val is_monomorphic: T -> string -> bool (*exception TYPE*)
val the_constraint: T -> string -> typ (*exception TYPE*)
- val space_of: T -> NameSpace.T
+ val space_of: T -> Name_Space.T
val intern: T -> xstring -> string
val extern: T -> string -> xstring
val extern_early: T -> string -> xstring
@@ -29,9 +29,9 @@
val certify: Pretty.pp -> Type.tsig -> bool -> T -> term -> term (*exception TYPE*)
val typargs: T -> string * typ -> typ list
val instance: T -> string * typ list -> typ
- val declare: bool -> NameSpace.naming -> Properties.T -> (binding * typ) -> T -> T
+ val declare: bool -> Name_Space.naming -> Properties.T -> binding * typ -> T -> T
val constrain: string * typ option -> T -> T
- val abbreviate: Pretty.pp -> Type.tsig -> NameSpace.naming -> string -> Properties.T ->
+ val abbreviate: Pretty.pp -> Type.tsig -> Name_Space.naming -> string -> Properties.T ->
binding * term -> T -> (term * term) * T
val revert_abbrev: string -> string -> T -> T
val hide: bool -> string -> T -> T
@@ -50,7 +50,7 @@
type abbrev = {rhs: term, normal_rhs: term, force_expand: bool};
datatype T = Consts of
- {decls: ((decl * abbrev option) * serial) NameSpace.table,
+ {decls: (decl * abbrev option) Name_Space.table,
constraints: typ Symtab.table,
rev_abbrevs: (term * term) Item_Net.T Symtab.table};
@@ -84,7 +84,7 @@
fun dest (Consts {decls = (space, decls), constraints, ...}) =
{constants = (space,
- Symtab.fold (fn (c, (({T, ...}, abbr), _)) =>
+ Symtab.fold (fn (c, ({T, ...}, abbr)) =>
Symtab.update (c, (T, Option.map #rhs abbr))) decls Symtab.empty),
constraints = (space, constraints)};
@@ -93,7 +93,7 @@
fun the_const (Consts {decls = (_, tab), ...}) c =
(case Symtab.lookup tab c of
- SOME (decl, _) => decl
+ SOME decl => decl
| NONE => raise TYPE ("Unknown constant: " ^ quote c, [], []));
fun the_type consts c =
@@ -123,8 +123,8 @@
fun space_of (Consts {decls = (space, _), ...}) = space;
-val intern = NameSpace.intern o space_of;
-val extern = NameSpace.extern o space_of;
+val intern = Name_Space.intern o space_of;
+val extern = Name_Space.extern o space_of;
fun extern_early consts c =
(case try (the_const consts) c of
@@ -221,27 +221,20 @@
(** build consts **)
-fun err_dup_const c =
- error ("Duplicate declaration of constant " ^ quote c);
-
-fun extend_decls naming decl tab = NameSpace.define naming decl tab
- handle Symtab.DUP c => err_dup_const c;
-
-
(* name space *)
fun hide fully c = map_consts (fn (decls, constraints, rev_abbrevs) =>
- (apfst (NameSpace.hide fully c) decls, constraints, rev_abbrevs));
+ (apfst (Name_Space.hide fully c) decls, constraints, rev_abbrevs));
(* declarations *)
-fun declare authentic naming tags (b, declT) = map_consts (fn (decls, constraints, rev_abbrevs) =>
- let
- val tags' = tags |> Position.default_properties (Position.thread_data ());
- val decl = {T = declT, typargs = typargs_of declT, tags = tags', authentic = authentic};
- val (_, decls') = decls |> extend_decls naming (b, ((decl, NONE), serial ()));
- in (decls', constraints, rev_abbrevs) end);
+fun declare authentic naming tags (b, declT) =
+ map_consts (fn (decls, constraints, rev_abbrevs) =>
+ let
+ val decl = {T = declT, typargs = typargs_of declT, tags = tags, authentic = authentic};
+ val (_, decls') = decls |> Name_Space.define true naming (b, (decl, NONE));
+ in (decls', constraints, rev_abbrevs) end);
(* constraints *)
@@ -278,7 +271,7 @@
val force_expand = mode = PrintMode.internal;
val _ = Term.exists_subterm Term.is_Var raw_rhs andalso
- error ("Illegal schematic variables on rhs of abbreviation: " ^ Binding.str_of b);
+ error ("Illegal schematic variables on rhs of abbreviation " ^ quote (Binding.str_of b));
val rhs = raw_rhs
|> Term.map_types (Type.cert_typ tsig)
@@ -286,15 +279,14 @@
|> Term.close_schematic_term;
val normal_rhs = expand_term rhs;
val T = Term.fastype_of rhs;
- val lhs = Const (NameSpace.full_name naming b, T);
+ val lhs = Const (Name_Space.full_name naming b, T);
in
consts |> map_consts (fn (decls, constraints, rev_abbrevs) =>
let
- val tags' = tags |> Position.default_properties (Position.thread_data ());
- val decl = {T = T, typargs = typargs_of T, tags = tags', authentic = true};
+ val decl = {T = T, typargs = typargs_of T, tags = tags, authentic = true};
val abbr = {rhs = rhs, normal_rhs = normal_rhs, force_expand = force_expand};
val (_, decls') = decls
- |> extend_decls naming (b, ((decl, SOME abbr), serial ()));
+ |> Name_Space.define true naming (b, (decl, SOME abbr));
val rev_abbrevs' = rev_abbrevs
|> insert_abbrevs mode (rev_abbrev lhs rhs);
in (decls', constraints, rev_abbrevs') end)
@@ -313,14 +305,13 @@
(* empty and merge *)
-val empty = make_consts (NameSpace.empty_table, Symtab.empty, Symtab.empty);
+val empty = make_consts (Name_Space.empty_table "constant", Symtab.empty, Symtab.empty);
fun merge
(Consts {decls = decls1, constraints = constraints1, rev_abbrevs = rev_abbrevs1},
Consts {decls = decls2, constraints = constraints2, rev_abbrevs = rev_abbrevs2}) =
let
- val decls' = NameSpace.merge_tables (eq_snd (op =)) (decls1, decls2)
- handle Symtab.DUP c => err_dup_const c;
+ val decls' = Name_Space.merge_tables (decls1, decls2);
val constraints' = Symtab.join (K fst) (constraints1, constraints2);
val rev_abbrevs' = Symtab.join (K Item_Net.merge) (rev_abbrevs1, rev_abbrevs2);
in make_consts (decls', constraints', rev_abbrevs') end;
--- a/src/Pure/display.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/display.ML Sun Oct 25 08:57:55 2009 +0100
@@ -146,14 +146,14 @@
[Pretty.str "default sort:", Pretty.brk 1, prt_sort S];
val tfrees = map (fn v => TFree (v, []));
- fun pretty_type syn (t, ((Type.LogicalType n, _), _)) =
+ fun pretty_type syn (t, (Type.LogicalType n, _)) =
if syn then NONE
else SOME (prt_typ (Type (t, tfrees (Name.invents Name.context Name.aT n))))
- | pretty_type syn (t, ((Type.Abbreviation (vs, U, syn'), _), _)) =
+ | pretty_type syn (t, (Type.Abbreviation (vs, U, syn'), _)) =
if syn <> syn' then NONE
else SOME (Pretty.block
[prt_typ (Type (t, tfrees vs)), Pretty.str " =", Pretty.brk 1, prt_typ U])
- | pretty_type syn (t, ((Type.Nonterminal, _), _)) =
+ | pretty_type syn (t, (Type.Nonterminal, _)) =
if not syn then NONE
else SOME (prt_typ (Type (t, [])));
@@ -179,25 +179,25 @@
val {restricts, reducts} = Defs.dest defs;
val {naming = _, syn = _, tsig, consts} = Sign.rep_sg thy;
val {constants, constraints} = Consts.dest consts;
- val extern_const = NameSpace.extern (#1 constants);
+ val extern_const = Name_Space.extern (#1 constants);
val {classes, default, types, ...} = Type.rep_tsig tsig;
val (class_space, class_algebra) = classes;
val {classes, arities} = Sorts.rep_algebra class_algebra;
- val clsses = NameSpace.dest_table (class_space, Symtab.make (Graph.dest classes));
- val tdecls = NameSpace.dest_table types;
- val arties = NameSpace.dest_table (Sign.type_space thy, arities);
+ val clsses = Name_Space.dest_table (class_space, Symtab.make (Graph.dest classes));
+ val tdecls = Name_Space.dest_table types;
+ val arties = Name_Space.dest_table (Sign.type_space thy, arities);
fun prune_const c = not verbose andalso
member (op =) (Consts.the_tags consts c) Markup.property_internal;
- val cnsts = NameSpace.extern_table (#1 constants,
+ val cnsts = Name_Space.extern_table (#1 constants,
Symtab.make (filter_out (prune_const o fst) (Symtab.dest (#2 constants))));
val log_cnsts = map_filter (fn (c, (ty, NONE)) => SOME (c, ty) | _ => NONE) cnsts;
val abbrevs = map_filter (fn (c, (ty, SOME t)) => SOME (c, (ty, t)) | _ => NONE) cnsts;
- val cnstrs = NameSpace.extern_table constraints;
+ val cnstrs = Name_Space.extern_table constraints;
- val axms = NameSpace.extern_table axioms;
+ val axms = Name_Space.extern_table axioms;
val (reds0, (reds1, reds2)) = filter_out (prune_const o fst o fst) reducts
|> map (fn (lhs, rhs) =>
--- a/src/Pure/drule.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/drule.ML Sun Oct 25 08:57:55 2009 +0100
@@ -452,7 +452,7 @@
val read_prop = certify o SimpleSyntax.read_prop;
-fun get_def thy = Thm.axiom thy o NameSpace.intern (Theory.axiom_space thy) o Thm.def_name;
+fun get_def thy = Thm.axiom thy o Name_Space.intern (Theory.axiom_space thy) o Thm.def_name;
fun store_thm name th =
Context.>>> (Context.map_theory_result (PureThy.store_thm (Binding.name name, th)));
--- a/src/Pure/facts.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/facts.ML Sun Oct 25 08:57:55 2009 +0100
@@ -20,7 +20,7 @@
val selections: string * thm list -> (ref * thm) list
type T
val empty: T
- val space_of: T -> NameSpace.T
+ val space_of: T -> Name_Space.T
val intern: T -> xstring -> string
val extern: T -> string -> xstring
val lookup: Context.generic -> T -> string -> (bool * thm list) option
@@ -31,9 +31,9 @@
val props: T -> thm list
val could_unify: T -> term -> thm list
val merge: T * T -> T
- val add_global: NameSpace.naming -> binding * thm list -> T -> string * T
- val add_local: bool -> NameSpace.naming -> binding * thm list -> T -> string * T
- val add_dynamic: NameSpace.naming -> binding * (Context.generic -> thm list) -> T -> string * T
+ val add_global: Name_Space.naming -> binding * thm list -> T -> string * T
+ val add_local: bool -> Name_Space.naming -> binding * thm list -> T -> string * T
+ val add_dynamic: Name_Space.naming -> binding * (Context.generic -> thm list) -> T -> string * T
val del: string -> T -> T
val hide: bool -> string -> T -> T
end;
@@ -122,11 +122,11 @@
datatype fact = Static of thm list | Dynamic of Context.generic -> thm list;
datatype T = Facts of
- {facts: (fact * serial) NameSpace.table,
+ {facts: fact Name_Space.table,
props: thm Net.net};
fun make_facts facts props = Facts {facts = facts, props = props};
-val empty = make_facts NameSpace.empty_table Net.empty;
+val empty = make_facts (Name_Space.empty_table "fact") Net.empty;
(* named facts *)
@@ -136,18 +136,19 @@
val space_of = #1 o facts_of;
val table_of = #2 o facts_of;
-val intern = NameSpace.intern o space_of;
-val extern = NameSpace.extern o space_of;
+val intern = Name_Space.intern o space_of;
+val extern = Name_Space.extern o space_of;
val defined = Symtab.defined o table_of;
fun lookup context facts name =
(case Symtab.lookup (table_of facts) name of
NONE => NONE
- | SOME (Static ths, _) => SOME (true, ths)
- | SOME (Dynamic f, _) => SOME (false, f context));
+ | SOME (Static ths) => SOME (true, ths)
+ | SOME (Dynamic f) => SOME (false, f context));
-fun fold_static f = Symtab.fold (fn (name, (Static ths, _)) => f (name, ths) | _ => I) o table_of;
+fun fold_static f =
+ Symtab.fold (fn (name, Static ths) => f (name, ths) | _ => I) o table_of;
(* content difference *)
@@ -174,61 +175,52 @@
(* merge facts *)
-fun err_dup_fact name = error ("Duplicate fact " ^ quote name);
-
-(* FIXME stamp identity! *)
-fun eq_entry ((Static ths1, _), (Static ths2, _)) = is_equal (list_ord Thm.thm_ord (ths1, ths2))
- | eq_entry ((_, id1: serial), (_, id2)) = id1 = id2;
-
fun merge (Facts {facts = facts1, props = props1}, Facts {facts = facts2, props = props2}) =
let
- val facts' = NameSpace.merge_tables eq_entry (facts1, facts2)
- handle Symtab.DUP dup => err_dup_fact dup;
+ val facts' = Name_Space.merge_tables (facts1, facts2);
val props' = Net.merge (is_equal o prop_ord) (props1, props2);
in make_facts facts' props' end;
(* add static entries *)
-fun add permissive do_props naming (b, ths) (Facts {facts, props}) =
+local
+
+fun add strict do_props naming (b, ths) (Facts {facts, props}) =
let
- val (name, facts') = if Binding.is_empty b then ("", facts)
- else let
- val (space, tab) = facts;
- val (name, space') = NameSpace.declare naming b space;
- val entry = (name, (Static ths, serial ()));
- val tab' = (if permissive then Symtab.update else Symtab.update_new) entry tab
- handle Symtab.DUP dup => err_dup_fact dup;
- in (name, (space', tab')) end;
+ val (name, facts') =
+ if Binding.is_empty b then ("", facts)
+ else Name_Space.define strict naming (b, Static ths) facts;
val props' =
- if do_props then fold (fn th => Net.insert_term (K false) (Thm.full_prop_of th, th)) ths props
+ if do_props
+ then fold (fn th => Net.insert_term (K false) (Thm.full_prop_of th, th)) ths props
else props;
in (name, make_facts facts' props') end;
-val add_global = add false false;
-val add_local = add true;
+in
+
+val add_global = add true false;
+val add_local = add false;
+
+end;
(* add dynamic entries *)
-fun add_dynamic naming (b, f) (Facts {facts = (space, tab), props}) =
- let
- val (name, space') = NameSpace.declare naming b space;
- val entry = (name, (Dynamic f, serial ()));
- val tab' = Symtab.update_new entry tab
- handle Symtab.DUP dup => err_dup_fact dup;
- in (name, make_facts (space', tab') props) end;
+fun add_dynamic naming (b, f) (Facts {facts, props}) =
+ let val (name, facts') = Name_Space.define true naming (b, Dynamic f) facts;
+ in (name, make_facts facts' props) end;
(* remove entries *)
fun del name (Facts {facts = (space, tab), props}) =
let
- val space' = NameSpace.hide true name space handle ERROR _ => space;
+ val space' = Name_Space.hide true name space handle ERROR _ => space; (* FIXME handle?? *)
val tab' = Symtab.delete_safe name tab;
in make_facts (space', tab') props end;
fun hide fully name (Facts {facts = (space, tab), props}) =
- make_facts (NameSpace.hide fully name space, tab) props;
+ make_facts (Name_Space.hide fully name space, tab) props;
end;
--- a/src/Pure/name.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/name.ML Sun Oct 25 08:57:55 2009 +0100
@@ -45,7 +45,7 @@
(* checked binding *)
-val of_binding = Long_Name.base_name o NameSpace.full_name NameSpace.default_naming;
+val of_binding = Long_Name.base_name o Name_Space.full_name Name_Space.default_naming;
(* encoded bounds *)
--- a/src/Pure/pure_thy.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/pure_thy.ML Sun Oct 25 08:57:55 2009 +0100
@@ -146,7 +146,7 @@
else
let
val naming = Sign.naming_of thy;
- val name = NameSpace.full_name naming b;
+ val name = Name_Space.full_name naming b;
val (thy', thms') =
register_proofs (apsnd (post_name name) (app_att (thy, pre_name name thms)));
val thms'' = map (Thm.transfer thy') thms';
--- a/src/Pure/sign.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/sign.ML Sun Oct 25 08:57:55 2009 +0100
@@ -8,11 +8,11 @@
signature SIGN =
sig
val rep_sg: theory ->
- {naming: NameSpace.naming,
+ {naming: Name_Space.naming,
syn: Syntax.syntax,
tsig: Type.tsig,
consts: Consts.T}
- val naming_of: theory -> NameSpace.naming
+ val naming_of: theory -> Name_Space.naming
val full_name: theory -> binding -> string
val full_name_path: theory -> string -> binding -> string
val full_bname: theory -> bstring -> string
@@ -44,9 +44,9 @@
val const_typargs: theory -> string * typ -> typ list
val const_instance: theory -> string * typ list -> typ
val mk_const: theory -> string * typ list -> term
- val class_space: theory -> NameSpace.T
- val type_space: theory -> NameSpace.T
- val const_space: theory -> NameSpace.T
+ val class_space: theory -> Name_Space.T
+ val type_space: theory -> Name_Space.T
+ val const_space: theory -> Name_Space.T
val intern_class: theory -> xstring -> string
val extern_class: theory -> string -> xstring
val intern_type: theory -> xstring -> string
@@ -137,7 +137,7 @@
(** datatype sign **)
datatype sign = Sign of
- {naming: NameSpace.naming, (*common naming conventions*)
+ {naming: Name_Space.naming, (*common naming conventions*)
syn: Syntax.syntax, (*concrete syntax for terms, types, sorts*)
tsig: Type.tsig, (*order-sorted signature of types*)
consts: Consts.T}; (*polymorphic constants*)
@@ -150,17 +150,17 @@
type T = sign;
val copy = I;
fun extend (Sign {syn, tsig, consts, ...}) =
- make_sign (NameSpace.default_naming, syn, tsig, consts);
+ make_sign (Name_Space.default_naming, syn, tsig, consts);
val empty =
- make_sign (NameSpace.default_naming, Syntax.basic_syn, Type.empty_tsig, Consts.empty);
+ make_sign (Name_Space.default_naming, Syntax.basic_syn, Type.empty_tsig, Consts.empty);
fun merge pp (sign1, sign2) =
let
val Sign {naming = _, syn = syn1, tsig = tsig1, consts = consts1} = sign1;
val Sign {naming = _, syn = syn2, tsig = tsig2, consts = consts2} = sign2;
- val naming = NameSpace.default_naming;
+ val naming = Name_Space.default_naming;
val syn = Syntax.merge_syntaxes syn1 syn2;
val tsig = Type.merge_tsigs pp (tsig1, tsig2);
val consts = Consts.merge (consts1, consts2);
@@ -182,10 +182,10 @@
val naming_of = #naming o rep_sg;
-val full_name = NameSpace.full_name o naming_of;
-fun full_name_path thy path = NameSpace.full_name (NameSpace.add_path path (naming_of thy));
+val full_name = Name_Space.full_name o naming_of;
+fun full_name_path thy path = Name_Space.full_name (Name_Space.add_path path (naming_of thy));
-fun full_bname thy = NameSpace.full_name (naming_of thy) o Binding.name;
+fun full_bname thy = Name_Space.full_name (naming_of thy) o Binding.name;
fun full_bname_path thy path = full_name_path thy path o Binding.name;
@@ -240,12 +240,12 @@
val type_space = #1 o #types o Type.rep_tsig o tsig_of;
val const_space = Consts.space_of o consts_of;
-val intern_class = NameSpace.intern o class_space;
-val extern_class = NameSpace.extern o class_space;
-val intern_type = NameSpace.intern o type_space;
-val extern_type = NameSpace.extern o type_space;
-val intern_const = NameSpace.intern o const_space;
-val extern_const = NameSpace.extern o const_space;
+val intern_class = Name_Space.intern o class_space;
+val extern_class = Name_Space.extern o class_space;
+val intern_type = Name_Space.intern o type_space;
+val extern_type = Name_Space.extern o type_space;
+val intern_const = Name_Space.intern o const_space;
+val extern_const = Name_Space.extern o const_space;
val intern_sort = map o intern_class;
val extern_sort = map o extern_class;
@@ -526,8 +526,7 @@
fun declare_const tags ((b, T), mx) thy =
let
val pos = Binding.pos_of b;
- val tags' = Position.default_properties pos tags;
- val ([const as Const (c, _)], thy') = gen_add_consts (K I) true tags' [(b, T, mx)] thy;
+ val ([const as Const (c, _)], thy') = gen_add_consts (K I) true tags [(b, T, mx)] thy;
val _ = Position.report (Markup.const_decl c) pos;
in (const, thy') end;
@@ -612,10 +611,10 @@
(* naming *)
-val add_path = map_naming o NameSpace.add_path;
-val root_path = map_naming NameSpace.root_path;
-val parent_path = map_naming NameSpace.parent_path;
-val mandatory_path = map_naming o NameSpace.mandatory_path;
+val add_path = map_naming o Name_Space.add_path;
+val root_path = map_naming Name_Space.root_path;
+val parent_path = map_naming Name_Space.parent_path;
+val mandatory_path = map_naming o Name_Space.mandatory_path;
fun local_path thy = thy |> root_path |> add_path (Context.theory_name thy);
--- a/src/Pure/simplifier.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/simplifier.ML Sun Oct 25 08:57:55 2009 +0100
@@ -143,18 +143,14 @@
(** named simprocs **)
-fun err_dup_simproc name = error ("Duplicate simproc: " ^ quote name);
-
-
(* data *)
structure Simprocs = GenericDataFun
(
- type T = simproc NameSpace.table;
- val empty = NameSpace.empty_table;
+ type T = simproc Name_Space.table;
+ val empty = Name_Space.empty_table "simproc";
val extend = I;
- fun merge _ simprocs = NameSpace.merge_tables eq_simproc simprocs
- handle Symtab.DUP dup => err_dup_simproc dup;
+ fun merge _ simprocs = Name_Space.merge_tables simprocs;
);
@@ -163,7 +159,7 @@
fun get_simproc context xname =
let
val (space, tab) = Simprocs.get context;
- val name = NameSpace.intern space xname;
+ val name = Name_Space.intern space xname;
in
(case Symtab.lookup tab name of
SOME proc => proc
@@ -201,9 +197,7 @@
val b' = Morphism.binding phi b;
val simproc' = morph_simproc phi simproc;
in
- Simprocs.map (fn simprocs =>
- NameSpace.define naming (b', simproc') simprocs |> snd
- handle Symtab.DUP dup => err_dup_simproc dup)
+ Simprocs.map (#2 o Name_Space.define true naming (b', simproc'))
#> map_ss (fn ss => ss addsimprocs [simproc'])
end)
end;
--- a/src/Pure/theory.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/theory.ML Sun Oct 25 08:57:55 2009 +0100
@@ -19,7 +19,7 @@
val checkpoint: theory -> theory
val copy: theory -> theory
val requires: theory -> string -> string -> unit
- val axiom_space: theory -> NameSpace.T
+ val axiom_space: theory -> Name_Space.T
val axiom_table: theory -> term Symtab.table
val axioms_of: theory -> (string * term) list
val all_axioms_of: theory -> (string * term) list
@@ -80,29 +80,27 @@
perhaps (perhaps_loop (perhaps_apply (map fst wrappers)));
datatype thy = Thy of
- {axioms: term NameSpace.table,
+ {axioms: term Name_Space.table,
defs: Defs.T,
wrappers: wrapper list * wrapper list};
fun make_thy (axioms, defs, wrappers) = Thy {axioms = axioms, defs = defs, wrappers = wrappers};
-fun err_dup_axm dup = error ("Duplicate axiom: " ^ quote dup);
-
structure ThyData = TheoryDataFun
(
type T = thy;
- val empty = make_thy (NameSpace.empty_table, Defs.empty, ([], []));
+ val empty_axioms = Name_Space.empty_table "axiom";
+ val empty = make_thy (empty_axioms, Defs.empty, ([], []));
val copy = I;
- fun extend (Thy {axioms = _, defs, wrappers}) =
- make_thy (NameSpace.empty_table, defs, wrappers);
+ fun extend (Thy {axioms = _, defs, wrappers}) = make_thy (empty_axioms, defs, wrappers);
fun merge pp (thy1, thy2) =
let
val Thy {axioms = _, defs = defs1, wrappers = (bgs1, ens1)} = thy1;
val Thy {axioms = _, defs = defs2, wrappers = (bgs2, ens2)} = thy2;
- val axioms' = NameSpace.empty_table;
+ val axioms' = empty_axioms;
val defs' = Defs.merge pp (defs1, defs2);
val bgs' = Library.merge (eq_snd op =) (bgs1, bgs2);
val ens' = Library.merge (eq_snd op =) (ens1, ens2);
@@ -166,7 +164,7 @@
fun read_axm thy (b, str) =
cert_axm thy (b, Syntax.read_prop_global thy str) handle ERROR msg =>
- cat_error msg ("The error(s) above occurred in axiom: " ^ quote (Binding.str_of b));
+ cat_error msg ("The error(s) above occurred in axiom " ^ quote (Binding.str_of b));
(* add_axioms(_i) *)
@@ -176,8 +174,7 @@
fun gen_add_axioms prep_axm raw_axms thy = thy |> map_axioms (fn axioms =>
let
val axms = map (apsnd Logic.varify o prep_axm thy) raw_axms;
- val axioms' = fold (snd oo NameSpace.define (Sign.naming_of thy)) axms axioms
- handle Symtab.DUP dup => err_dup_axm dup;
+ val axioms' = fold (#2 oo Name_Space.define true (Sign.naming_of thy)) axms axioms;
in axioms' end);
in
--- a/src/Pure/thm.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/thm.ML Sun Oct 25 08:57:55 2009 +0100
@@ -1724,25 +1724,21 @@
(* authentic derivation names *)
-fun err_dup_ora dup = error ("Duplicate oracle: " ^ quote dup);
-
structure Oracles = TheoryDataFun
(
- type T = serial NameSpace.table;
- val empty = NameSpace.empty_table;
+ type T = unit Name_Space.table;
+ val empty = Name_Space.empty_table "oracle";
val copy = I;
val extend = I;
- fun merge _ oracles : T = NameSpace.merge_tables (op =) oracles
- handle Symtab.DUP dup => err_dup_ora dup;
+ fun merge _ oracles : T = Name_Space.merge_tables oracles;
);
-val extern_oracles = map #1 o NameSpace.extern_table o Oracles.get;
+val extern_oracles = map #1 o Name_Space.extern_table o Oracles.get;
fun add_oracle (b, oracle) thy =
let
val naming = Sign.naming_of thy;
- val (name, tab') = NameSpace.define naming (b, serial ()) (Oracles.get thy)
- handle Symtab.DUP _ => err_dup_ora (Binding.str_of b);
+ val (name, tab') = Name_Space.define true naming (b, ()) (Oracles.get thy);
val thy' = Oracles.put tab' thy;
in ((name, invoke_oracle (Theory.check_thy thy') name oracle), thy') end;
--- a/src/Pure/type.ML Sun Oct 25 08:57:36 2009 +0100
+++ b/src/Pure/type.ML Sun Oct 25 08:57:55 2009 +0100
@@ -14,9 +14,9 @@
Nonterminal
type tsig
val rep_tsig: tsig ->
- {classes: NameSpace.T * Sorts.algebra,
+ {classes: Name_Space.T * Sorts.algebra,
default: sort,
- types: ((decl * Properties.T) * serial) NameSpace.table,
+ types: (decl * Properties.T) Name_Space.table,
log_types: string list}
val empty_tsig: tsig
val defaultS: tsig -> sort
@@ -70,12 +70,12 @@
val eq_type: tyenv -> typ * typ -> bool
(*extend and merge type signatures*)
- val add_class: Pretty.pp -> NameSpace.naming -> binding * class list -> tsig -> tsig
+ val add_class: Pretty.pp -> Name_Space.naming -> binding * class list -> tsig -> tsig
val hide_class: bool -> string -> tsig -> tsig
val set_defsort: sort -> tsig -> tsig
- val add_type: NameSpace.naming -> Properties.T -> binding * int -> tsig -> tsig
- val add_abbrev: NameSpace.naming -> Properties.T -> binding * string list * typ -> tsig -> tsig
- val add_nonterminal: NameSpace.naming -> Properties.T -> binding -> tsig -> tsig
+ val add_type: Name_Space.naming -> Properties.T -> binding * int -> tsig -> tsig
+ val add_abbrev: Name_Space.naming -> Properties.T -> binding * string list * typ -> tsig -> tsig
+ val add_nonterminal: Name_Space.naming -> Properties.T -> binding -> tsig -> tsig
val hide_type: bool -> string -> tsig -> tsig
val add_arity: Pretty.pp -> arity -> tsig -> tsig
val add_classrel: Pretty.pp -> class * class -> tsig -> tsig
@@ -94,18 +94,14 @@
Abbreviation of string list * typ * bool |
Nonterminal;
-fun str_of_decl (LogicalType _) = "logical type constructor"
- | str_of_decl (Abbreviation _) = "type abbreviation"
- | str_of_decl Nonterminal = "syntactic type";
-
(* type tsig *)
datatype tsig =
TSig of {
- classes: NameSpace.T * Sorts.algebra, (*order-sorted algebra of type classes*)
+ classes: Name_Space.T * Sorts.algebra, (*order-sorted algebra of type classes*)
default: sort, (*default sort on input*)
- types: ((decl * Properties.T) * serial) NameSpace.table, (*declared types*)
+ types: (decl * Properties.T) Name_Space.table, (*declared types*)
log_types: string list}; (*logical types sorted by number of arguments*)
fun rep_tsig (TSig comps) = comps;
@@ -113,18 +109,18 @@
fun make_tsig (classes, default, types, log_types) =
TSig {classes = classes, default = default, types = types, log_types = log_types};
-fun build_tsig ((space, classes), default, types) =
+fun build_tsig (classes, default, types) =
let
val log_types =
- Symtab.fold (fn (c, ((LogicalType n, _), _)) => cons (c, n) | _ => I) (snd types) []
- |> Library.sort (Library.int_ord o pairself snd) |> map fst;
- in make_tsig ((space, classes), default, types, log_types) end;
+ Symtab.fold (fn (c, (LogicalType n, _)) => cons (c, n) | _ => I) (snd types) []
+ |> Library.sort (int_ord o pairself snd) |> map fst;
+ in make_tsig (classes, default, types, log_types) end;
fun map_tsig f (TSig {classes, default, types, log_types = _}) =
build_tsig (f (classes, default, types));
val empty_tsig =
- build_tsig ((NameSpace.empty, Sorts.empty_algebra), [], NameSpace.empty_table);
+ build_tsig ((Name_Space.empty "class", Sorts.empty_algebra), [], Name_Space.empty_table "type");
(* classes and sorts *)
@@ -167,7 +163,7 @@
fun undecl_type c = "Undeclared type constructor: " ^ quote c;
-fun lookup_type (TSig {types, ...}) = Option.map fst o Symtab.lookup (snd types);
+fun lookup_type (TSig {types = (_, types), ...}) = Symtab.lookup types;
fun the_tags tsig c =
(case lookup_type tsig c of
@@ -515,12 +511,12 @@
let
val cs' = map (cert_class tsig) cs
handle TYPE (msg, _, _) => error msg;
- val (c', space') = space |> NameSpace.declare naming c;
+ val (c', space') = space |> Name_Space.declare true naming c;
val classes' = classes |> Sorts.add_class pp (c', cs');
in ((space', classes'), default, types) end);
fun hide_class fully c = map_tsig (fn ((space, classes), default, types) =>
- ((NameSpace.hide fully c space, classes), default, types));
+ ((Name_Space.hide fully c space, classes), default, types));
(* arities *)
@@ -530,7 +526,7 @@
val _ =
(case lookup_type tsig t of
SOME (LogicalType n, _) => if length Ss <> n then error (bad_nargs t) else ()
- | SOME (decl, _) => error ("Illegal " ^ str_of_decl decl ^ ": " ^ quote t)
+ | SOME _ => error ("Logical type constructor expected: " ^ quote t)
| NONE => error (undecl_type t));
val (Ss', S') = (map (cert_sort tsig) Ss, cert_sort tsig S)
handle TYPE (msg, _, _) => error msg;
@@ -559,68 +555,49 @@
local
-fun err_in_decls c decl decl' =
- let val s = str_of_decl decl and s' = str_of_decl decl' in
- if s = s' then error ("Duplicate declaration of " ^ s ^ ": " ^ quote c)
- else error ("Conflict of " ^ s ^ " with " ^ s' ^ ": " ^ quote c)
- end;
-
-fun new_decl naming tags (c, decl) (space, types) =
- let
- val tags' = tags |> Position.default_properties (Position.thread_data ());
- val (c', space') = NameSpace.declare naming c space;
- val types' =
- (case Symtab.lookup types c' of
- SOME ((decl', _), _) => err_in_decls c' decl decl'
- | NONE => Symtab.update (c', ((decl, tags'), serial ())) types);
- in (space', types') end;
-
-fun the_decl (_, types) = fst o fst o the o Symtab.lookup types;
+fun new_decl naming tags (c, decl) types =
+ #2 (Name_Space.define true naming (c, (decl, tags)) types);
fun map_types f = map_tsig (fn (classes, default, types) =>
let
val (space', tab') = f types;
- val _ = NameSpace.intern space' "dummy" = "dummy" orelse
+ val _ = Name_Space.intern space' "dummy" = "dummy" orelse
error "Illegal declaration of dummy type";
in (classes, default, (space', tab')) end);
fun syntactic types (Type (c, Ts)) =
- (case Symtab.lookup types c of SOME ((Nonterminal, _), _) => true | _ => false)
+ (case Symtab.lookup types c of SOME (Nonterminal, _) => true | _ => false)
orelse exists (syntactic types) Ts
| syntactic _ _ = false;
in
fun add_type naming tags (c, n) =
- if n < 0 then error ("Bad type constructor declaration: " ^ quote (Binding.str_of c))
+ if n < 0 then error ("Bad type constructor declaration " ^ quote (Binding.str_of c))
else map_types (new_decl naming tags (c, LogicalType n));
fun add_abbrev naming tags (a, vs, rhs) tsig = tsig |> map_types (fn types =>
let
fun err msg =
- cat_error msg ("The error(s) above occurred in type abbreviation: " ^ quote (Binding.str_of a));
+ cat_error msg ("The error(s) above occurred in type abbreviation " ^ quote (Binding.str_of a));
val rhs' = strip_sorts (no_tvars (cert_typ_mode mode_syntax tsig rhs))
handle TYPE (msg, _, _) => err msg;
- in
- (case duplicates (op =) vs of
- [] => []
- | dups => err ("Duplicate variables on lhs: " ^ commas_quote dups));
- (case subtract (op =) vs (map #1 (Term.add_tfreesT rhs' [])) of
- [] => []
- | extras => err ("Extra variables on rhs: " ^ commas_quote extras));
- types |> new_decl naming tags (a, Abbreviation (vs, rhs', syntactic (#2 types) rhs'))
- end);
+ val _ =
+ (case duplicates (op =) vs of
+ [] => []
+ | dups => err ("Duplicate variables on lhs: " ^ commas_quote dups));
+ val _ =
+ (case subtract (op =) vs (map #1 (Term.add_tfreesT rhs' [])) of
+ [] => []
+ | extras => err ("Extra variables on rhs: " ^ commas_quote extras));
+ in types |> new_decl naming tags (a, Abbreviation (vs, rhs', syntactic (#2 types) rhs')) end);
fun add_nonterminal naming tags = map_types o new_decl naming tags o rpair Nonterminal;
-fun merge_types (types1, types2) =
- NameSpace.merge_tables (Library.eq_snd (op = : serial * serial -> bool)) (types1, types2)
- handle Symtab.DUP d => err_in_decls d (the_decl types1 d) (the_decl types2 d);
-
end;
fun hide_type fully c = map_tsig (fn (classes, default, (space, types)) =>
- (classes, default, (NameSpace.hide fully c space, types)));
+ (classes, default, (Name_Space.hide fully c space, types)));
(* merge type signatures *)
@@ -632,10 +609,10 @@
val (TSig {classes = (space2, classes2), default = default2, types = types2,
log_types = _}) = tsig2;
- val space' = NameSpace.merge (space1, space2);
+ val space' = Name_Space.merge (space1, space2);
val classes' = Sorts.merge_algebra pp (classes1, classes2);
val default' = Sorts.inter_sort classes' (default1, default2);
- val types' = merge_types (types1, types2);
+ val types' = Name_Space.merge_tables (types1, types2);
in build_tsig ((space', classes'), default', types') end;
end;