--- a/src/HOL/Code_Eval.thy Wed Jun 10 21:04:36 2009 +0200
+++ b/src/HOL/Code_Eval.thy Thu Jun 11 08:04:26 2009 +0200
@@ -129,7 +129,7 @@
(Eval "Term.term")
code_const Const and App
- (Eval "Term.Const/ (_, _)" and "Term.$/ (_, _)")
+ (Eval "Term.Const/ ((_), (_))" and "Term.$/ ((_), (_))")
code_const "term_of \<Colon> String.literal \<Rightarrow> term"
(Eval "HOLogic.mk'_message'_string")
--- a/src/HOL/Fun.thy Wed Jun 10 21:04:36 2009 +0200
+++ b/src/HOL/Fun.thy Thu Jun 11 08:04:26 2009 +0200
@@ -133,7 +133,7 @@
shows "inj f"
using assms unfolding inj_on_def by auto
-text{*For Proofs in @{text "Tools/datatype_rep_proofs"}*}
+text{*For Proofs in @{text "Tools/datatype_package/datatype_rep_proofs"}*}
lemma datatype_injI:
"(!! x. ALL y. f(x) = f(y) --> x=y) ==> inj(f)"
by (simp add: inj_on_def)
--- a/src/HOL/Induct/Tree.thy Wed Jun 10 21:04:36 2009 +0200
+++ b/src/HOL/Induct/Tree.thy Thu Jun 11 08:04:26 2009 +0200
@@ -1,12 +1,13 @@
(* Title: HOL/Induct/Tree.thy
- ID: $Id$
Author: Stefan Berghofer, TU Muenchen
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
*)
header {* Infinitely branching trees *}
-theory Tree imports Main begin
+theory Tree
+imports Main
+begin
datatype 'a tree =
Atom 'a
--- a/src/HOL/Inductive.thy Wed Jun 10 21:04:36 2009 +0200
+++ b/src/HOL/Inductive.thy Thu Jun 11 08:04:26 2009 +0200
@@ -10,15 +10,15 @@
("Tools/inductive_package.ML")
"Tools/dseq.ML"
("Tools/inductive_codegen.ML")
- ("Tools/datatype_aux.ML")
- ("Tools/datatype_prop.ML")
- ("Tools/datatype_rep_proofs.ML")
- ("Tools/datatype_abs_proofs.ML")
- ("Tools/datatype_case.ML")
- ("Tools/datatype_package.ML")
+ ("Tools/datatype_package/datatype_aux.ML")
+ ("Tools/datatype_package/datatype_prop.ML")
+ ("Tools/datatype_package/datatype_rep_proofs.ML")
+ ("Tools/datatype_package/datatype_abs_proofs.ML")
+ ("Tools/datatype_package/datatype_case.ML")
+ ("Tools/datatype_package/datatype_package.ML")
("Tools/old_primrec_package.ML")
("Tools/primrec_package.ML")
- ("Tools/datatype_codegen.ML")
+ ("Tools/datatype_package/datatype_codegen.ML")
begin
subsection {* Least and greatest fixed points *}
@@ -335,17 +335,18 @@
text {* Package setup. *}
-use "Tools/datatype_aux.ML"
-use "Tools/datatype_prop.ML"
-use "Tools/datatype_rep_proofs.ML"
-use "Tools/datatype_abs_proofs.ML"
-use "Tools/datatype_case.ML"
-use "Tools/datatype_package.ML"
+use "Tools/datatype_package/datatype_aux.ML"
+use "Tools/datatype_package/datatype_prop.ML"
+use "Tools/datatype_package/datatype_rep_proofs.ML"
+use "Tools/datatype_package/datatype_abs_proofs.ML"
+use "Tools/datatype_package/datatype_case.ML"
+use "Tools/datatype_package/datatype_package.ML"
setup DatatypePackage.setup
+
use "Tools/old_primrec_package.ML"
use "Tools/primrec_package.ML"
-use "Tools/datatype_codegen.ML"
+use "Tools/datatype_package/datatype_codegen.ML"
setup DatatypeCodegen.setup
use "Tools/inductive_codegen.ML"
--- a/src/HOL/IsaMakefile Wed Jun 10 21:04:36 2009 +0200
+++ b/src/HOL/IsaMakefile Thu Jun 11 08:04:26 2009 +0200
@@ -141,14 +141,14 @@
Sum_Type.thy \
Tools/arith_data.ML \
Tools/cnf_funcs.ML \
- Tools/datatype_abs_proofs.ML \
- Tools/datatype_aux.ML \
- Tools/datatype_case.ML \
- Tools/datatype_codegen.ML \
- Tools/datatype_package.ML \
- Tools/datatype_prop.ML \
- Tools/datatype_realizer.ML \
- Tools/datatype_rep_proofs.ML \
+ Tools/datatype_package/datatype_abs_proofs.ML \
+ Tools/datatype_package/datatype_aux.ML \
+ Tools/datatype_package/datatype_case.ML \
+ Tools/datatype_package/datatype_codegen.ML \
+ Tools/datatype_package/datatype_package.ML \
+ Tools/datatype_package/datatype_prop.ML \
+ Tools/datatype_package/datatype_realizer.ML \
+ Tools/datatype_package/datatype_rep_proofs.ML \
Tools/dseq.ML \
Tools/function_package/auto_term.ML \
Tools/function_package/context_tree.ML \
@@ -854,7 +854,7 @@
ex/Lagrange.thy ex/LocaleTest2.thy ex/MT.thy ex/MergeSort.thy \
ex/Meson_Test.thy ex/MonoidGroup.thy ex/Multiquote.thy ex/NatSum.thy \
ex/Numeral.thy ex/PER.thy ex/PresburgerEx.thy ex/Primrec.thy \
- ex/Quickcheck_Examples.thy ex/Quickcheck_Generators.thy ex/ROOT.ML \
+ ex/Quickcheck_Examples.thy ex/ROOT.ML \
ex/Recdefs.thy ex/Records.thy ex/ReflectionEx.thy \
ex/Refute_Examples.thy ex/SAT_Examples.thy ex/SVC_Oracle.thy \
ex/Serbian.thy ex/Sqrt.thy ex/Sqrt_Script.thy \
--- a/src/HOL/Library/Enum.thy Wed Jun 10 21:04:36 2009 +0200
+++ b/src/HOL/Library/Enum.thy Thu Jun 11 08:04:26 2009 +0200
@@ -1,5 +1,4 @@
-(* Author: Florian Haftmann, TU Muenchen
-*)
+(* Author: Florian Haftmann, TU Muenchen *)
header {* Finite types as explicit enumerations *}
--- a/src/HOL/Product_Type.thy Wed Jun 10 21:04:36 2009 +0200
+++ b/src/HOL/Product_Type.thy Thu Jun 11 08:04:26 2009 +0200
@@ -12,7 +12,7 @@
("Tools/split_rule.ML")
("Tools/inductive_set_package.ML")
("Tools/inductive_realizer.ML")
- ("Tools/datatype_realizer.ML")
+ ("Tools/datatype_package/datatype_realizer.ML")
begin
subsection {* @{typ bool} is a datatype *}
@@ -1155,7 +1155,7 @@
use "Tools/inductive_set_package.ML"
setup InductiveSetPackage.setup
-use "Tools/datatype_realizer.ML"
+use "Tools/datatype_package/datatype_realizer.ML"
setup DatatypeRealizer.setup
end
--- a/src/HOL/Quickcheck.thy Wed Jun 10 21:04:36 2009 +0200
+++ b/src/HOL/Quickcheck.thy Thu Jun 11 08:04:26 2009 +0200
@@ -87,6 +87,28 @@
subsection {* Complex generators *}
+text {* Towards @{typ "'a \<Rightarrow> 'b"} *}
+
+axiomatization random_fun_aux :: "typerep \<Rightarrow> typerep \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> term)
+ \<Rightarrow> (Random.seed \<Rightarrow> ('b \<times> (unit \<Rightarrow> term)) \<times> Random.seed) \<Rightarrow> (Random.seed \<Rightarrow> Random.seed \<times> Random.seed)
+ \<Rightarrow> Random.seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
+
+definition random_fun_lift :: "(code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('b \<times> (unit \<Rightarrow> term)) \<times> Random.seed)
+ \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> (('a\<Colon>term_of \<Rightarrow> 'b\<Colon>typerep) \<times> (unit \<Rightarrow> term)) \<times> Random.seed" where
+ "random_fun_lift f i = random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Eval.term_of (f i) Random.split_seed"
+
+instantiation "fun" :: ("{eq, term_of}", "{type, random}") random
+begin
+
+definition random_fun :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> Random.seed" where
+ "random = random_fun_lift random"
+
+instance ..
+
+end
+
+text {* Towards type copies and datatypes *}
+
definition collapse :: "('a \<Rightarrow> ('a \<Rightarrow> 'b \<times> 'a) \<times> 'a) \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a" where
"collapse f = (f o\<rightarrow> id)"
@@ -107,31 +129,15 @@
use "Tools/quickcheck_generators.ML"
setup {* Quickcheck_Generators.setup *}
-code_reserved Quickcheck Quickcheck_Generators
-
-text {* Type @{typ "'a \<Rightarrow> 'b"} *}
-
-axiomatization random_fun_aux :: "typerep \<Rightarrow> typerep \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> term)
- \<Rightarrow> (Random.seed \<Rightarrow> ('b \<times> (unit \<Rightarrow> term)) \<times> Random.seed) \<Rightarrow> (Random.seed \<Rightarrow> Random.seed \<times> Random.seed)
- \<Rightarrow> Random.seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
-
code_const random_fun_aux (Quickcheck "Quickcheck'_Generators.random'_fun")
-- {* With enough criminal energy this can be abused to derive @{prop False};
for this reason we use a distinguished target @{text Quickcheck}
not spoiling the regular trusted code generation *}
-instantiation "fun" :: ("{eq, term_of}", "{type, random}") random
-begin
-
-definition random_fun :: "code_numeral \<Rightarrow> Random.seed \<Rightarrow> (('a \<Rightarrow> 'b) \<times> (unit \<Rightarrow> term)) \<times> Random.seed" where
- "random n = random_fun_aux TYPEREP('a) TYPEREP('b) (op =) Code_Eval.term_of (random n) Random.split_seed"
-
-instance ..
-
-end
+code_reserved Quickcheck Quickcheck_Generators
-hide (open) const collapse beyond
+hide (open) const collapse beyond random_fun_aux random_fun_lift
no_notation fcomp (infixl "o>" 60)
no_notation scomp (infixl "o\<rightarrow>" 60)
--- a/src/HOL/Tools/datatype_abs_proofs.ML Wed Jun 10 21:04:36 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,446 +0,0 @@
-(* Title: HOL/Tools/datatype_abs_proofs.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Proofs and defintions independent of concrete representation
-of datatypes (i.e. requiring only abstract properties such as
-injectivity / distinctness of constructors and induction)
-
- - case distinction (exhaustion) theorems
- - characteristic equations for primrec combinators
- - characteristic equations for case combinators
- - equations for splitting "P (case ...)" expressions
- - "nchotomy" and "case_cong" theorems for TFL
-*)
-
-signature DATATYPE_ABS_PROOFS =
-sig
- val prove_casedist_thms : string list ->
- DatatypeAux.descr list -> (string * sort) list -> thm ->
- attribute list -> theory -> thm list * theory
- val prove_primrec_thms : bool -> string list ->
- DatatypeAux.descr list -> (string * sort) list ->
- DatatypeAux.datatype_info Symtab.table -> thm list list -> thm list list ->
- simpset -> thm -> theory -> (string list * thm list) * theory
- val prove_case_thms : bool -> string list ->
- DatatypeAux.descr list -> (string * sort) list ->
- string list -> thm list -> theory -> (thm list list * string list) * theory
- val prove_split_thms : string list ->
- DatatypeAux.descr list -> (string * sort) list ->
- thm list list -> thm list list -> thm list -> thm list list -> theory ->
- (thm * thm) list * theory
- val prove_nchotomys : string list -> DatatypeAux.descr list ->
- (string * sort) list -> thm list -> theory -> thm list * theory
- val prove_weak_case_congs : string list -> DatatypeAux.descr list ->
- (string * sort) list -> theory -> thm list * theory
- val prove_case_congs : string list ->
- DatatypeAux.descr list -> (string * sort) list ->
- thm list -> thm list list -> theory -> thm list * theory
-end;
-
-structure DatatypeAbsProofs: DATATYPE_ABS_PROOFS =
-struct
-
-open DatatypeAux;
-
-(************************ case distinction theorems ***************************)
-
-fun prove_casedist_thms new_type_names descr sorts induct case_names_exhausts thy =
- let
- val _ = message "Proving case distinction theorems ...";
-
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val newTs = Library.take (length (hd descr), recTs);
-
- val {maxidx, ...} = rep_thm induct;
- val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
-
- fun prove_casedist_thm ((i, t), T) =
- let
- val dummyPs = map (fn (Var (_, Type (_, [T', T'']))) =>
- Abs ("z", T', Const ("True", T''))) induct_Ps;
- val P = Abs ("z", T, HOLogic.imp $ HOLogic.mk_eq (Var (("a", maxidx+1), T), Bound 0) $
- Var (("P", 0), HOLogic.boolT))
- val insts = Library.take (i, dummyPs) @ (P::(Library.drop (i + 1, dummyPs)));
- val cert = cterm_of thy;
- val insts' = (map cert induct_Ps) ~~ (map cert insts);
- val induct' = refl RS ((List.nth
- (split_conj_thm (cterm_instantiate insts' induct), i)) RSN (2, rev_mp))
-
- in
- SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
- (fn {prems, ...} => EVERY
- [rtac induct' 1,
- REPEAT (rtac TrueI 1),
- REPEAT ((rtac impI 1) THEN (eresolve_tac prems 1)),
- REPEAT (rtac TrueI 1)])
- end;
-
- val casedist_thms = map prove_casedist_thm ((0 upto (length newTs - 1)) ~~
- (DatatypeProp.make_casedists descr sorts) ~~ newTs)
- in
- thy
- |> store_thms_atts "exhaust" new_type_names (map single case_names_exhausts) casedist_thms
- end;
-
-
-(*************************** primrec combinators ******************************)
-
-fun prove_primrec_thms flat_names new_type_names descr sorts
- (dt_info : datatype_info Symtab.table) constr_inject dist_rewrites dist_ss induct thy =
- let
- val _ = message "Constructing primrec combinators ...";
-
- val big_name = space_implode "_" new_type_names;
- val thy0 = add_path flat_names big_name thy;
-
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
- val newTs = Library.take (length (hd descr), recTs);
-
- val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
-
- val big_rec_name' = big_name ^ "_rec_set";
- val rec_set_names' =
- if length descr' = 1 then [big_rec_name'] else
- map ((curry (op ^) (big_rec_name' ^ "_")) o string_of_int)
- (1 upto (length descr'));
- val rec_set_names = map (Sign.full_bname thy0) rec_set_names';
-
- val (rec_result_Ts, reccomb_fn_Ts) = DatatypeProp.make_primrec_Ts descr sorts used;
-
- val rec_set_Ts = map (fn (T1, T2) =>
- reccomb_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts);
-
- val rec_fns = map (uncurry (mk_Free "f"))
- (reccomb_fn_Ts ~~ (1 upto (length reccomb_fn_Ts)));
- val rec_sets' = map (fn c => list_comb (Free c, rec_fns))
- (rec_set_names' ~~ rec_set_Ts);
- val rec_sets = map (fn c => list_comb (Const c, rec_fns))
- (rec_set_names ~~ rec_set_Ts);
-
- (* introduction rules for graph of primrec function *)
-
- fun make_rec_intr T rec_set ((rec_intr_ts, l), (cname, cargs)) =
- let
- fun mk_prem ((dt, U), (j, k, prems, t1s, t2s)) =
- let val free1 = mk_Free "x" U j
- in (case (strip_dtyp dt, strip_type U) of
- ((_, DtRec m), (Us, _)) =>
- let
- val free2 = mk_Free "y" (Us ---> List.nth (rec_result_Ts, m)) k;
- val i = length Us
- in (j + 1, k + 1, HOLogic.mk_Trueprop (HOLogic.list_all
- (map (pair "x") Us, List.nth (rec_sets', m) $
- app_bnds free1 i $ app_bnds free2 i)) :: prems,
- free1::t1s, free2::t2s)
- end
- | _ => (j + 1, k, prems, free1::t1s, t2s))
- end;
-
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val (_, _, prems, t1s, t2s) = List.foldr mk_prem (1, 1, [], [], []) (cargs ~~ Ts)
-
- in (rec_intr_ts @ [Logic.list_implies (prems, HOLogic.mk_Trueprop
- (rec_set $ list_comb (Const (cname, Ts ---> T), t1s) $
- list_comb (List.nth (rec_fns, l), t1s @ t2s)))], l + 1)
- end;
-
- val (rec_intr_ts, _) = Library.foldl (fn (x, ((d, T), set_name)) =>
- Library.foldl (make_rec_intr T set_name) (x, #3 (snd d)))
- (([], 0), descr' ~~ recTs ~~ rec_sets');
-
- val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
- InductivePackage.add_inductive_global (serial_string ())
- {quiet_mode = ! quiet_mode, verbose = false, kind = Thm.internalK,
- alt_name = Binding.name big_rec_name', coind = false, no_elim = false, no_ind = true,
- skip_mono = true, fork_mono = false}
- (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
- (map dest_Free rec_fns)
- (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] thy0;
-
- (* prove uniqueness and termination of primrec combinators *)
-
- val _ = message "Proving termination and uniqueness of primrec functions ...";
-
- fun mk_unique_tac ((tac, intrs), ((((i, (tname, _, constrs)), elim), T), T')) =
- let
- val distinct_tac =
- (if i < length newTs then
- full_simp_tac (HOL_ss addsimps (List.nth (dist_rewrites, i))) 1
- else full_simp_tac dist_ss 1);
-
- val inject = map (fn r => r RS iffD1)
- (if i < length newTs then List.nth (constr_inject, i)
- else #inject (the (Symtab.lookup dt_info tname)));
-
- fun mk_unique_constr_tac n ((tac, intr::intrs, j), (cname, cargs)) =
- let
- val k = length (List.filter is_rec_type cargs)
-
- in (EVERY [DETERM tac,
- REPEAT (etac ex1E 1), rtac ex1I 1,
- DEPTH_SOLVE_1 (ares_tac [intr] 1),
- REPEAT_DETERM_N k (etac thin_rl 1 THEN rotate_tac 1 1),
- etac elim 1,
- REPEAT_DETERM_N j distinct_tac,
- TRY (dresolve_tac inject 1),
- REPEAT (etac conjE 1), hyp_subst_tac 1,
- REPEAT (EVERY [etac allE 1, dtac mp 1, atac 1]),
- TRY (hyp_subst_tac 1),
- rtac refl 1,
- REPEAT_DETERM_N (n - j - 1) distinct_tac],
- intrs, j + 1)
- end;
-
- val (tac', intrs', _) = Library.foldl (mk_unique_constr_tac (length constrs))
- ((tac, intrs, 0), constrs);
-
- in (tac', intrs') end;
-
- val rec_unique_thms =
- let
- val rec_unique_ts = map (fn (((set_t, T1), T2), i) =>
- Const ("Ex1", (T2 --> HOLogic.boolT) --> HOLogic.boolT) $
- absfree ("y", T2, set_t $ mk_Free "x" T1 i $ Free ("y", T2)))
- (rec_sets ~~ recTs ~~ rec_result_Ts ~~ (1 upto length recTs));
- val cert = cterm_of thy1
- val insts = map (fn ((i, T), t) => absfree ("x" ^ (string_of_int i), T, t))
- ((1 upto length recTs) ~~ recTs ~~ rec_unique_ts);
- val induct' = cterm_instantiate ((map cert induct_Ps) ~~
- (map cert insts)) induct;
- val (tac, _) = Library.foldl mk_unique_tac
- (((rtac induct' THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1
- THEN rewrite_goals_tac [mk_meta_eq choice_eq], rec_intrs),
- descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts);
-
- in split_conj_thm (SkipProof.prove_global thy1 [] []
- (HOLogic.mk_Trueprop (mk_conj rec_unique_ts)) (K tac))
- end;
-
- val rec_total_thms = map (fn r => r RS theI') rec_unique_thms;
-
- (* define primrec combinators *)
-
- val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
- val reccomb_names = map (Sign.full_bname thy1)
- (if length descr' = 1 then [big_reccomb_name] else
- (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
- (1 upto (length descr'))));
- val reccombs = map (fn ((name, T), T') => list_comb
- (Const (name, reccomb_fn_Ts @ [T] ---> T'), rec_fns))
- (reccomb_names ~~ recTs ~~ rec_result_Ts);
-
- val (reccomb_defs, thy2) =
- thy1
- |> Sign.add_consts_i (map (fn ((name, T), T') =>
- (Binding.name (Long_Name.base_name name), reccomb_fn_Ts @ [T] ---> T', NoSyn))
- (reccomb_names ~~ recTs ~~ rec_result_Ts))
- |> (PureThy.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') =>
- (Binding.name (Long_Name.base_name name ^ "_def"), Logic.mk_equals (comb, absfree ("x", T,
- Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
- set $ Free ("x", T) $ Free ("y", T'))))))
- (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts))
- ||> parent_path flat_names
- ||> Theory.checkpoint;
-
-
- (* prove characteristic equations for primrec combinators *)
-
- val _ = message "Proving characteristic theorems for primrec combinators ..."
-
- val rec_thms = map (fn t => SkipProof.prove_global thy2 [] [] t
- (fn _ => EVERY
- [rewrite_goals_tac reccomb_defs,
- rtac the1_equality 1,
- resolve_tac rec_unique_thms 1,
- resolve_tac rec_intrs 1,
- REPEAT (rtac allI 1 ORELSE resolve_tac rec_total_thms 1)]))
- (DatatypeProp.make_primrecs new_type_names descr sorts thy2)
-
- in
- thy2
- |> Sign.add_path (space_implode "_" new_type_names)
- |> PureThy.add_thmss [((Binding.name "recs", rec_thms),
- [Nitpick_Const_Simp_Thms.add])]
- ||> Sign.parent_path
- ||> Theory.checkpoint
- |-> (fn thms => pair (reccomb_names, Library.flat thms))
- end;
-
-
-(***************************** case combinators *******************************)
-
-fun prove_case_thms flat_names new_type_names descr sorts reccomb_names primrec_thms thy =
- let
- val _ = message "Proving characteristic theorems for case combinators ...";
-
- val thy1 = add_path flat_names (space_implode "_" new_type_names) thy;
-
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
- val newTs = Library.take (length (hd descr), recTs);
- val T' = TFree (Name.variant used "'t", HOLogic.typeS);
-
- fun mk_dummyT dt = binder_types (typ_of_dtyp descr' sorts dt) ---> T';
-
- val case_dummy_fns = map (fn (_, (_, _, constrs)) => map (fn (_, cargs) =>
- let
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val Ts' = map mk_dummyT (List.filter is_rec_type cargs)
- in Const (@{const_name undefined}, Ts @ Ts' ---> T')
- end) constrs) descr';
-
- val case_names = map (fn s => Sign.full_bname thy1 (s ^ "_case")) new_type_names;
-
- (* define case combinators via primrec combinators *)
-
- val (case_defs, thy2) = Library.foldl (fn ((defs, thy),
- ((((i, (_, _, constrs)), T), name), recname)) =>
- let
- val (fns1, fns2) = ListPair.unzip (map (fn ((_, cargs), j) =>
- let
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val Ts' = Ts @ map mk_dummyT (List.filter is_rec_type cargs);
- val frees' = map (uncurry (mk_Free "x")) (Ts' ~~ (1 upto length Ts'));
- val frees = Library.take (length cargs, frees');
- val free = mk_Free "f" (Ts ---> T') j
- in
- (free, list_abs_free (map dest_Free frees',
- list_comb (free, frees)))
- end) (constrs ~~ (1 upto length constrs)));
-
- val caseT = (map (snd o dest_Free) fns1) @ [T] ---> T';
- val fns = (List.concat (Library.take (i, case_dummy_fns))) @
- fns2 @ (List.concat (Library.drop (i + 1, case_dummy_fns)));
- val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T');
- val decl = ((Binding.name (Long_Name.base_name name), caseT), NoSyn);
- val def = (Binding.name (Long_Name.base_name name ^ "_def"),
- Logic.mk_equals (list_comb (Const (name, caseT), fns1),
- list_comb (reccomb, (List.concat (Library.take (i, case_dummy_fns))) @
- fns2 @ (List.concat (Library.drop (i + 1, case_dummy_fns))) )));
- val ([def_thm], thy') =
- thy
- |> Sign.declare_const [] decl |> snd
- |> (PureThy.add_defs false o map Thm.no_attributes) [def];
-
- in (defs @ [def_thm], thy')
- end) (([], thy1), (hd descr) ~~ newTs ~~ case_names ~~
- (Library.take (length newTs, reccomb_names)))
- ||> Theory.checkpoint;
-
- val case_thms = map (map (fn t => SkipProof.prove_global thy2 [] [] t
- (fn _ => EVERY [rewrite_goals_tac (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1])))
- (DatatypeProp.make_cases new_type_names descr sorts thy2)
- in
- thy2
- |> Context.the_theory o fold (fold Nitpick_Const_Simp_Thms.add_thm) case_thms
- o Context.Theory
- |> parent_path flat_names
- |> store_thmss "cases" new_type_names case_thms
- |-> (fn thmss => pair (thmss, case_names))
- end;
-
-
-(******************************* case splitting *******************************)
-
-fun prove_split_thms new_type_names descr sorts constr_inject dist_rewrites
- casedist_thms case_thms thy =
- let
- val _ = message "Proving equations for case splitting ...";
-
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val newTs = Library.take (length (hd descr), recTs);
-
- fun prove_split_thms ((((((t1, t2), inject), dist_rewrites'),
- exhaustion), case_thms'), T) =
- let
- val cert = cterm_of thy;
- val _ $ (_ $ lhs $ _) = hd (Logic.strip_assums_hyp (hd (prems_of exhaustion)));
- val exhaustion' = cterm_instantiate
- [(cert lhs, cert (Free ("x", T)))] exhaustion;
- val tacf = K (EVERY [rtac exhaustion' 1, ALLGOALS (asm_simp_tac
- (HOL_ss addsimps (dist_rewrites' @ inject @ case_thms')))])
- in
- (SkipProof.prove_global thy [] [] t1 tacf,
- SkipProof.prove_global thy [] [] t2 tacf)
- end;
-
- val split_thm_pairs = map prove_split_thms
- ((DatatypeProp.make_splits new_type_names descr sorts thy) ~~ constr_inject ~~
- dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs);
-
- val (split_thms, split_asm_thms) = ListPair.unzip split_thm_pairs
-
- in
- thy
- |> store_thms "split" new_type_names split_thms
- ||>> store_thms "split_asm" new_type_names split_asm_thms
- |-> (fn (thms1, thms2) => pair (thms1 ~~ thms2))
- end;
-
-fun prove_weak_case_congs new_type_names descr sorts thy =
- let
- fun prove_weak_case_cong t =
- SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
- (fn {prems, ...} => EVERY [rtac ((hd prems) RS arg_cong) 1])
-
- val weak_case_congs = map prove_weak_case_cong (DatatypeProp.make_weak_case_congs
- new_type_names descr sorts thy)
-
- in thy |> store_thms "weak_case_cong" new_type_names weak_case_congs end;
-
-(************************* additional theorems for TFL ************************)
-
-fun prove_nchotomys new_type_names descr sorts casedist_thms thy =
- let
- val _ = message "Proving additional theorems for TFL ...";
-
- fun prove_nchotomy (t, exhaustion) =
- let
- (* For goal i, select the correct disjunct to attack, then prove it *)
- fun tac i 0 = EVERY [TRY (rtac disjI1 i),
- hyp_subst_tac i, REPEAT (rtac exI i), rtac refl i]
- | tac i n = rtac disjI2 i THEN tac i (n - 1)
- in
- SkipProof.prove_global thy [] [] t (fn _ =>
- EVERY [rtac allI 1,
- exh_tac (K exhaustion) 1,
- ALLGOALS (fn i => tac i (i-1))])
- end;
-
- val nchotomys =
- map prove_nchotomy (DatatypeProp.make_nchotomys descr sorts ~~ casedist_thms)
-
- in thy |> store_thms "nchotomy" new_type_names nchotomys end;
-
-fun prove_case_congs new_type_names descr sorts nchotomys case_thms thy =
- let
- fun prove_case_cong ((t, nchotomy), case_rewrites) =
- let
- val (Const ("==>", _) $ tm $ _) = t;
- val (Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ Ma)) = tm;
- val cert = cterm_of thy;
- val nchotomy' = nchotomy RS spec;
- val [v] = Term.add_vars (concl_of nchotomy') [];
- val nchotomy'' = cterm_instantiate [(cert (Var v), cert Ma)] nchotomy'
- in
- SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
- (fn {prems, ...} =>
- let val simplify = asm_simp_tac (HOL_ss addsimps (prems @ case_rewrites))
- in EVERY [simp_tac (HOL_ss addsimps [hd prems]) 1,
- cut_facts_tac [nchotomy''] 1,
- REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1),
- REPEAT (etac exE 1) THEN simplify 1 (* Get last disjunct *)]
- end)
- end;
-
- val case_congs = map prove_case_cong (DatatypeProp.make_case_congs
- new_type_names descr sorts thy ~~ nchotomys ~~ case_thms)
-
- in thy |> store_thms "case_cong" new_type_names case_congs end;
-
-end;
--- a/src/HOL/Tools/datatype_aux.ML Wed Jun 10 21:04:36 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,348 +0,0 @@
-(* Title: HOL/Tools/datatype_aux.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Auxiliary functions for defining datatypes.
-*)
-
-signature DATATYPE_AUX =
-sig
- val quiet_mode : bool ref
- val message : string -> unit
-
- val add_path : bool -> string -> theory -> theory
- val parent_path : bool -> theory -> theory
-
- val store_thmss_atts : string -> string list -> attribute list list -> thm list list
- -> theory -> thm list list * theory
- val store_thmss : string -> string list -> thm list list -> theory -> thm list list * theory
- val store_thms_atts : string -> string list -> attribute list list -> thm list
- -> theory -> thm list * theory
- val store_thms : string -> string list -> thm list -> theory -> thm list * theory
-
- val split_conj_thm : thm -> thm list
- val mk_conj : term list -> term
- val mk_disj : term list -> term
-
- val app_bnds : term -> int -> term
-
- val cong_tac : int -> tactic
- val indtac : thm -> string list -> int -> tactic
- val exh_tac : (string -> thm) -> int -> tactic
-
- datatype simproc_dist = FewConstrs of thm list
- | ManyConstrs of thm * simpset;
-
- datatype dtyp =
- DtTFree of string
- | DtType of string * (dtyp list)
- | DtRec of int;
- type descr
- type datatype_info
-
- exception Datatype
- exception Datatype_Empty of string
- val name_of_typ : typ -> string
- val dtyp_of_typ : (string * string list) list -> typ -> dtyp
- val mk_Free : string -> typ -> int -> term
- val is_rec_type : dtyp -> bool
- val typ_of_dtyp : descr -> (string * sort) list -> dtyp -> typ
- val dest_DtTFree : dtyp -> string
- val dest_DtRec : dtyp -> int
- val strip_dtyp : dtyp -> dtyp list * dtyp
- val body_index : dtyp -> int
- val mk_fun_dtyp : dtyp list -> dtyp -> dtyp
- val get_nonrec_types : descr -> (string * sort) list -> typ list
- val get_branching_types : descr -> (string * sort) list -> typ list
- val get_arities : descr -> int list
- val get_rec_types : descr -> (string * sort) list -> typ list
- val check_nonempty : descr list -> unit
- val unfold_datatypes :
- theory -> descr -> (string * sort) list -> datatype_info Symtab.table ->
- descr -> int -> descr list * int
-end;
-
-structure DatatypeAux : DATATYPE_AUX =
-struct
-
-val quiet_mode = ref false;
-fun message s = if !quiet_mode then () else writeln s;
-
-fun add_path flat_names s = if flat_names then I else Sign.add_path s;
-fun parent_path flat_names = if flat_names then I else Sign.parent_path;
-
-
-(* store theorems in theory *)
-
-fun store_thmss_atts label tnames attss thmss =
- fold_map (fn ((tname, atts), thms) =>
- Sign.add_path tname
- #> PureThy.add_thmss [((Binding.name label, thms), atts)]
- #-> (fn thm::_ => Sign.parent_path #> pair thm)) (tnames ~~ attss ~~ thmss)
- ##> Theory.checkpoint;
-
-fun store_thmss label tnames = store_thmss_atts label tnames (replicate (length tnames) []);
-
-fun store_thms_atts label tnames attss thmss =
- fold_map (fn ((tname, atts), thms) =>
- Sign.add_path tname
- #> PureThy.add_thms [((Binding.name label, thms), atts)]
- #-> (fn thm::_ => Sign.parent_path #> pair thm)) (tnames ~~ attss ~~ thmss)
- ##> Theory.checkpoint;
-
-fun store_thms label tnames = store_thms_atts label tnames (replicate (length tnames) []);
-
-
-(* split theorem thm_1 & ... & thm_n into n theorems *)
-
-fun split_conj_thm th =
- ((th RS conjunct1)::(split_conj_thm (th RS conjunct2))) handle THM _ => [th];
-
-val mk_conj = foldr1 (HOLogic.mk_binop "op &");
-val mk_disj = foldr1 (HOLogic.mk_binop "op |");
-
-fun app_bnds t i = list_comb (t, map Bound (i - 1 downto 0));
-
-
-fun cong_tac i st = (case Logic.strip_assums_concl
- (List.nth (prems_of st, i - 1)) of
- _ $ (_ $ (f $ x) $ (g $ y)) =>
- let
- val cong' = Thm.lift_rule (Thm.cprem_of st i) cong;
- val _ $ (_ $ (f' $ x') $ (g' $ y')) =
- Logic.strip_assums_concl (prop_of cong');
- val insts = map (pairself (cterm_of (Thm.theory_of_thm st)) o
- apsnd (curry list_abs (Logic.strip_params (concl_of cong'))) o
- apfst head_of) [(f', f), (g', g), (x', x), (y', y)]
- in compose_tac (false, cterm_instantiate insts cong', 2) i st
- handle THM _ => no_tac st
- end
- | _ => no_tac st);
-
-(* instantiate induction rule *)
-
-fun indtac indrule indnames i st =
- let
- val ts = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule));
- val ts' = HOLogic.dest_conj (HOLogic.dest_Trueprop
- (Logic.strip_imp_concl (List.nth (prems_of st, i - 1))));
- val getP = if can HOLogic.dest_imp (hd ts) then
- (apfst SOME) o HOLogic.dest_imp else pair NONE;
- val flt = if null indnames then I else
- filter (fn Free (s, _) => s mem indnames | _ => false);
- fun abstr (t1, t2) = (case t1 of
- NONE => (case flt (OldTerm.term_frees t2) of
- [Free (s, T)] => SOME (absfree (s, T, t2))
- | _ => NONE)
- | SOME (_ $ t') => SOME (Abs ("x", fastype_of t', abstract_over (t', t2))))
- val cert = cterm_of (Thm.theory_of_thm st);
- val insts = List.mapPartial (fn (t, u) => case abstr (getP u) of
- NONE => NONE
- | SOME u' => SOME (t |> getP |> snd |> head_of |> cert, cert u')) (ts ~~ ts');
- val indrule' = cterm_instantiate insts indrule
- in
- rtac indrule' i st
- end;
-
-(* perform exhaustive case analysis on last parameter of subgoal i *)
-
-fun exh_tac exh_thm_of i state =
- let
- val thy = Thm.theory_of_thm state;
- val prem = nth (prems_of state) (i - 1);
- val params = Logic.strip_params prem;
- val (_, Type (tname, _)) = hd (rev params);
- val exhaustion = Thm.lift_rule (Thm.cprem_of state i) (exh_thm_of tname);
- val prem' = hd (prems_of exhaustion);
- val _ $ (_ $ lhs $ _) = hd (rev (Logic.strip_assums_hyp prem'));
- val exhaustion' = cterm_instantiate [(cterm_of thy (head_of lhs),
- cterm_of thy (List.foldr (fn ((_, T), t) => Abs ("z", T, t))
- (Bound 0) params))] exhaustion
- in compose_tac (false, exhaustion', nprems_of exhaustion) i state
- end;
-
-(* handling of distinctness theorems *)
-
-datatype simproc_dist = FewConstrs of thm list
- | ManyConstrs of thm * simpset;
-
-(********************** Internal description of datatypes *********************)
-
-datatype dtyp =
- DtTFree of string
- | DtType of string * (dtyp list)
- | DtRec of int;
-
-(* information about datatypes *)
-
-(* index, datatype name, type arguments, constructor name, types of constructor's arguments *)
-type descr = (int * (string * dtyp list * (string * dtyp list) list)) list;
-
-type datatype_info =
- {index : int,
- alt_names : string list option,
- descr : descr,
- sorts : (string * sort) list,
- rec_names : string list,
- rec_rewrites : thm list,
- case_name : string,
- case_rewrites : thm list,
- induction : thm,
- exhaustion : thm,
- distinct : simproc_dist,
- inject : thm list,
- nchotomy : thm,
- case_cong : thm,
- weak_case_cong : thm};
-
-fun mk_Free s T i = Free (s ^ (string_of_int i), T);
-
-fun subst_DtTFree _ substs (T as (DtTFree name)) =
- AList.lookup (op =) substs name |> the_default T
- | subst_DtTFree i substs (DtType (name, ts)) =
- DtType (name, map (subst_DtTFree i substs) ts)
- | subst_DtTFree i _ (DtRec j) = DtRec (i + j);
-
-exception Datatype;
-exception Datatype_Empty of string;
-
-fun dest_DtTFree (DtTFree a) = a
- | dest_DtTFree _ = raise Datatype;
-
-fun dest_DtRec (DtRec i) = i
- | dest_DtRec _ = raise Datatype;
-
-fun is_rec_type (DtType (_, dts)) = exists is_rec_type dts
- | is_rec_type (DtRec _) = true
- | is_rec_type _ = false;
-
-fun strip_dtyp (DtType ("fun", [T, U])) = apfst (cons T) (strip_dtyp U)
- | strip_dtyp T = ([], T);
-
-val body_index = dest_DtRec o snd o strip_dtyp;
-
-fun mk_fun_dtyp [] U = U
- | mk_fun_dtyp (T :: Ts) U = DtType ("fun", [T, mk_fun_dtyp Ts U]);
-
-fun name_of_typ (Type (s, Ts)) =
- let val s' = Long_Name.base_name s
- in space_implode "_" (List.filter (not o equal "") (map name_of_typ Ts) @
- [if Syntax.is_identifier s' then s' else "x"])
- end
- | name_of_typ _ = "";
-
-fun dtyp_of_typ _ (TFree (n, _)) = DtTFree n
- | dtyp_of_typ _ (TVar _) = error "Illegal schematic type variable(s)"
- | dtyp_of_typ new_dts (Type (tname, Ts)) =
- (case AList.lookup (op =) new_dts tname of
- NONE => DtType (tname, map (dtyp_of_typ new_dts) Ts)
- | SOME vs => if map (try (fst o dest_TFree)) Ts = map SOME vs then
- DtRec (find_index (curry op = tname o fst) new_dts)
- else error ("Illegal occurrence of recursive type " ^ tname));
-
-fun typ_of_dtyp descr sorts (DtTFree a) = TFree (a, (the o AList.lookup (op =) sorts) a)
- | typ_of_dtyp descr sorts (DtRec i) =
- let val (s, ds, _) = (the o AList.lookup (op =) descr) i
- in Type (s, map (typ_of_dtyp descr sorts) ds) end
- | typ_of_dtyp descr sorts (DtType (s, ds)) =
- Type (s, map (typ_of_dtyp descr sorts) ds);
-
-(* find all non-recursive types in datatype description *)
-
-fun get_nonrec_types descr sorts =
- map (typ_of_dtyp descr sorts) (Library.foldl (fn (Ts, (_, (_, _, constrs))) =>
- Library.foldl (fn (Ts', (_, cargs)) =>
- filter_out is_rec_type cargs union Ts') (Ts, constrs)) ([], descr));
-
-(* get all recursive types in datatype description *)
-
-fun get_rec_types descr sorts = map (fn (_ , (s, ds, _)) =>
- Type (s, map (typ_of_dtyp descr sorts) ds)) descr;
-
-(* get all branching types *)
-
-fun get_branching_types descr sorts =
- map (typ_of_dtyp descr sorts) (fold (fn (_, (_, _, constrs)) =>
- fold (fn (_, cargs) => fold (strip_dtyp #> fst #> fold (insert op =)) cargs)
- constrs) descr []);
-
-fun get_arities descr = fold (fn (_, (_, _, constrs)) =>
- fold (fn (_, cargs) => fold (insert op =) (map (length o fst o strip_dtyp)
- (List.filter is_rec_type cargs))) constrs) descr [];
-
-(* nonemptiness check for datatypes *)
-
-fun check_nonempty descr =
- let
- val descr' = List.concat descr;
- fun is_nonempty_dt is i =
- let
- val (_, _, constrs) = (the o AList.lookup (op =) descr') i;
- fun arg_nonempty (_, DtRec i) = if i mem is then false
- else is_nonempty_dt (i::is) i
- | arg_nonempty _ = true;
- in exists ((forall (arg_nonempty o strip_dtyp)) o snd) constrs
- end
- in assert_all (fn (i, _) => is_nonempty_dt [i] i) (hd descr)
- (fn (_, (s, _, _)) => raise Datatype_Empty s)
- end;
-
-(* unfold a list of mutually recursive datatype specifications *)
-(* all types of the form DtType (dt_name, [..., DtRec _, ...]) *)
-(* need to be unfolded *)
-
-fun unfold_datatypes sign orig_descr sorts (dt_info : datatype_info Symtab.table) descr i =
- let
- fun typ_error T msg = error ("Non-admissible type expression\n" ^
- Syntax.string_of_typ_global sign (typ_of_dtyp (orig_descr @ descr) sorts T) ^ "\n" ^ msg);
-
- fun get_dt_descr T i tname dts =
- (case Symtab.lookup dt_info tname of
- NONE => typ_error T (tname ^ " is not a datatype - can't use it in\
- \ nested recursion")
- | (SOME {index, descr, ...}) =>
- let val (_, vars, _) = (the o AList.lookup (op =) descr) index;
- val subst = ((map dest_DtTFree vars) ~~ dts) handle Library.UnequalLengths =>
- typ_error T ("Type constructor " ^ tname ^ " used with wrong\
- \ number of arguments")
- in (i + index, map (fn (j, (tn, args, cs)) => (i + j,
- (tn, map (subst_DtTFree i subst) args,
- map (apsnd (map (subst_DtTFree i subst))) cs))) descr)
- end);
-
- (* unfold a single constructor argument *)
-
- fun unfold_arg ((i, Ts, descrs), T) =
- if is_rec_type T then
- let val (Us, U) = strip_dtyp T
- in if exists is_rec_type Us then
- typ_error T "Non-strictly positive recursive occurrence of type"
- else (case U of
- DtType (tname, dts) =>
- let
- val (index, descr) = get_dt_descr T i tname dts;
- val (descr', i') = unfold_datatypes sign orig_descr sorts
- dt_info descr (i + length descr)
- in (i', Ts @ [mk_fun_dtyp Us (DtRec index)], descrs @ descr') end
- | _ => (i, Ts @ [T], descrs))
- end
- else (i, Ts @ [T], descrs);
-
- (* unfold a constructor *)
-
- fun unfold_constr ((i, constrs, descrs), (cname, cargs)) =
- let val (i', cargs', descrs') = Library.foldl unfold_arg ((i, [], descrs), cargs)
- in (i', constrs @ [(cname, cargs')], descrs') end;
-
- (* unfold a single datatype *)
-
- fun unfold_datatype ((i, dtypes, descrs), (j, (tname, tvars, constrs))) =
- let val (i', constrs', descrs') =
- Library.foldl unfold_constr ((i, [], descrs), constrs)
- in (i', dtypes @ [(j, (tname, tvars, constrs'))], descrs')
- end;
-
- val (i', descr', descrs) = Library.foldl unfold_datatype ((i, [],[]), descr);
-
- in (descr' :: descrs, i') end;
-
-end;
--- a/src/HOL/Tools/datatype_case.ML Wed Jun 10 21:04:36 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,469 +0,0 @@
-(* Title: HOL/Tools/datatype_case.ML
- Author: Konrad Slind, Cambridge University Computer Laboratory
- Author: Stefan Berghofer, TU Muenchen
-
-Nested case expressions on datatypes.
-*)
-
-signature DATATYPE_CASE =
-sig
- val make_case: (string -> DatatypeAux.datatype_info option) ->
- Proof.context -> bool -> string list -> term -> (term * term) list ->
- term * (term * (int * bool)) list
- val dest_case: (string -> DatatypeAux.datatype_info option) -> bool ->
- string list -> term -> (term * (term * term) list) option
- val strip_case: (string -> DatatypeAux.datatype_info option) -> bool ->
- term -> (term * (term * term) list) option
- val case_tr: bool -> (theory -> string -> DatatypeAux.datatype_info option)
- -> Proof.context -> term list -> term
- val case_tr': (theory -> string -> DatatypeAux.datatype_info option) ->
- string -> Proof.context -> term list -> term
-end;
-
-structure DatatypeCase : DATATYPE_CASE =
-struct
-
-exception CASE_ERROR of string * int;
-
-fun match_type thy pat ob = Sign.typ_match thy (pat, ob) Vartab.empty;
-
-(*---------------------------------------------------------------------------
- * Get information about datatypes
- *---------------------------------------------------------------------------*)
-
-fun ty_info (tab : string -> DatatypeAux.datatype_info option) s =
- case tab s of
- SOME {descr, case_name, index, sorts, ...} =>
- let
- val (_, (tname, dts, constrs)) = nth descr index;
- val mk_ty = DatatypeAux.typ_of_dtyp descr sorts;
- val T = Type (tname, map mk_ty dts)
- in
- SOME {case_name = case_name,
- constructors = map (fn (cname, dts') =>
- Const (cname, Logic.varifyT (map mk_ty dts' ---> T))) constrs}
- end
- | NONE => NONE;
-
-
-(*---------------------------------------------------------------------------
- * Each pattern carries with it a tag (i,b) where
- * i is the clause it came from and
- * b=true indicates that clause was given by the user
- * (or is an instantiation of a user supplied pattern)
- * b=false --> i = ~1
- *---------------------------------------------------------------------------*)
-
-fun pattern_subst theta (tm, x) = (subst_free theta tm, x);
-
-fun row_of_pat x = fst (snd x);
-
-fun add_row_used ((prfx, pats), (tm, tag)) =
- fold Term.add_free_names (tm :: pats @ prfx);
-
-(* try to preserve names given by user *)
-fun default_names names ts =
- map (fn ("", Free (name', _)) => name' | (name, _) => name) (names ~~ ts);
-
-fun strip_constraints (Const ("_constrain", _) $ t $ tT) =
- strip_constraints t ||> cons tT
- | strip_constraints t = (t, []);
-
-fun mk_fun_constrain tT t = Syntax.const "_constrain" $ t $
- (Syntax.free "fun" $ tT $ Syntax.free "dummy");
-
-
-(*---------------------------------------------------------------------------
- * Produce an instance of a constructor, plus genvars for its arguments.
- *---------------------------------------------------------------------------*)
-fun fresh_constr ty_match ty_inst colty used c =
- let
- val (_, Ty) = dest_Const c
- val Ts = binder_types Ty;
- val names = Name.variant_list used
- (DatatypeProp.make_tnames (map Logic.unvarifyT Ts));
- val ty = body_type Ty;
- val ty_theta = ty_match ty colty handle Type.TYPE_MATCH =>
- raise CASE_ERROR ("type mismatch", ~1)
- val c' = ty_inst ty_theta c
- val gvars = map (ty_inst ty_theta o Free) (names ~~ Ts)
- in (c', gvars)
- end;
-
-
-(*---------------------------------------------------------------------------
- * Goes through a list of rows and picks out the ones beginning with a
- * pattern with constructor = name.
- *---------------------------------------------------------------------------*)
-fun mk_group (name, T) rows =
- let val k = length (binder_types T)
- in fold (fn (row as ((prfx, p :: rst), rhs as (_, (i, _)))) =>
- fn ((in_group, not_in_group), (names, cnstrts)) => (case strip_comb p of
- (Const (name', _), args) =>
- if name = name' then
- if length args = k then
- let val (args', cnstrts') = split_list (map strip_constraints args)
- in
- ((((prfx, args' @ rst), rhs) :: in_group, not_in_group),
- (default_names names args', map2 append cnstrts cnstrts'))
- end
- else raise CASE_ERROR
- ("Wrong number of arguments for constructor " ^ name, i)
- else ((in_group, row :: not_in_group), (names, cnstrts))
- | _ => raise CASE_ERROR ("Not a constructor pattern", i)))
- rows (([], []), (replicate k "", replicate k [])) |>> pairself rev
- end;
-
-(*---------------------------------------------------------------------------
- * Partition the rows. Not efficient: we should use hashing.
- *---------------------------------------------------------------------------*)
-fun partition _ _ _ _ _ _ _ [] = raise CASE_ERROR ("partition: no rows", ~1)
- | partition ty_match ty_inst type_of used constructors colty res_ty
- (rows as (((prfx, _ :: rstp), _) :: _)) =
- let
- fun part {constrs = [], rows = [], A} = rev A
- | part {constrs = [], rows = (_, (_, (i, _))) :: _, A} =
- raise CASE_ERROR ("Not a constructor pattern", i)
- | part {constrs = c :: crst, rows, A} =
- let
- val ((in_group, not_in_group), (names, cnstrts)) =
- mk_group (dest_Const c) rows;
- val used' = fold add_row_used in_group used;
- val (c', gvars) = fresh_constr ty_match ty_inst colty used' c;
- val in_group' =
- if null in_group (* Constructor not given *)
- then
- let
- val Ts = map type_of rstp;
- val xs = Name.variant_list
- (fold Term.add_free_names gvars used')
- (replicate (length rstp) "x")
- in
- [((prfx, gvars @ map Free (xs ~~ Ts)),
- (Const ("HOL.undefined", res_ty), (~1, false)))]
- end
- else in_group
- in
- part{constrs = crst,
- rows = not_in_group,
- A = {constructor = c',
- new_formals = gvars,
- names = names,
- constraints = cnstrts,
- group = in_group'} :: A}
- end
- in part {constrs = constructors, rows = rows, A = []}
- end;
-
-(*---------------------------------------------------------------------------
- * Misc. routines used in mk_case
- *---------------------------------------------------------------------------*)
-
-fun mk_pat ((c, c'), l) =
- let
- val L = length (binder_types (fastype_of c))
- fun build (prfx, tag, plist) =
- let val (args, plist') = chop L plist
- in (prfx, tag, list_comb (c', args) :: plist') end
- in map build l end;
-
-fun v_to_prfx (prfx, v::pats) = (v::prfx,pats)
- | v_to_prfx _ = raise CASE_ERROR ("mk_case: v_to_prfx", ~1);
-
-fun v_to_pats (v::prfx,tag, pats) = (prfx, tag, v::pats)
- | v_to_pats _ = raise CASE_ERROR ("mk_case: v_to_pats", ~1);
-
-
-(*----------------------------------------------------------------------------
- * Translation of pattern terms into nested case expressions.
- *
- * This performs the translation and also builds the full set of patterns.
- * Thus it supports the construction of induction theorems even when an
- * incomplete set of patterns is given.
- *---------------------------------------------------------------------------*)
-
-fun mk_case tab ctxt ty_match ty_inst type_of used range_ty =
- let
- val name = Name.variant used "a";
- fun expand constructors used ty ((_, []), _) =
- raise CASE_ERROR ("mk_case: expand_var_row", ~1)
- | expand constructors used ty (row as ((prfx, p :: rst), rhs)) =
- if is_Free p then
- let
- val used' = add_row_used row used;
- fun expnd c =
- let val capp =
- list_comb (fresh_constr ty_match ty_inst ty used' c)
- in ((prfx, capp :: rst), pattern_subst [(p, capp)] rhs)
- end
- in map expnd constructors end
- else [row]
- fun mk {rows = [], ...} = raise CASE_ERROR ("no rows", ~1)
- | mk {path = [], rows = ((prfx, []), (tm, tag)) :: _} = (* Done *)
- ([(prfx, tag, [])], tm)
- | mk {path, rows as ((row as ((_, [Free _]), _)) :: _ :: _)} =
- mk {path = path, rows = [row]}
- | mk {path = u :: rstp, rows as ((_, _ :: _), _) :: _} =
- let val col0 = map (fn ((_, p :: _), (_, (i, _))) => (p, i)) rows
- in case Option.map (apfst head_of)
- (find_first (not o is_Free o fst) col0) of
- NONE =>
- let
- val rows' = map (fn ((v, _), row) => row ||>
- pattern_subst [(v, u)] |>> v_to_prfx) (col0 ~~ rows);
- val (pref_patl, tm) = mk {path = rstp, rows = rows'}
- in (map v_to_pats pref_patl, tm) end
- | SOME (Const (cname, cT), i) => (case ty_info tab cname of
- NONE => raise CASE_ERROR ("Not a datatype constructor: " ^ cname, i)
- | SOME {case_name, constructors} =>
- let
- val pty = body_type cT;
- val used' = fold Term.add_free_names rstp used;
- val nrows = maps (expand constructors used' pty) rows;
- val subproblems = partition ty_match ty_inst type_of used'
- constructors pty range_ty nrows;
- val new_formals = map #new_formals subproblems
- val constructors' = map #constructor subproblems
- val news = map (fn {new_formals, group, ...} =>
- {path = new_formals @ rstp, rows = group}) subproblems;
- val (pat_rect, dtrees) = split_list (map mk news);
- val case_functions = map2
- (fn {new_formals, names, constraints, ...} =>
- fold_rev (fn ((x as Free (_, T), s), cnstrts) => fn t =>
- Abs (if s = "" then name else s, T,
- abstract_over (x, t)) |>
- fold mk_fun_constrain cnstrts)
- (new_formals ~~ names ~~ constraints))
- subproblems dtrees;
- val types = map type_of (case_functions @ [u]);
- val case_const = Const (case_name, types ---> range_ty)
- val tree = list_comb (case_const, case_functions @ [u])
- val pat_rect1 = flat (map mk_pat
- (constructors ~~ constructors' ~~ pat_rect))
- in (pat_rect1, tree)
- end)
- | SOME (t, i) => raise CASE_ERROR ("Not a datatype constructor: " ^
- Syntax.string_of_term ctxt t, i)
- end
- | mk _ = raise CASE_ERROR ("Malformed row matrix", ~1)
- in mk
- end;
-
-fun case_error s = error ("Error in case expression:\n" ^ s);
-
-(* Repeated variable occurrences in a pattern are not allowed. *)
-fun no_repeat_vars ctxt pat = fold_aterms
- (fn x as Free (s, _) => (fn xs =>
- if member op aconv xs x then
- case_error (quote s ^ " occurs repeatedly in the pattern " ^
- quote (Syntax.string_of_term ctxt pat))
- else x :: xs)
- | _ => I) pat [];
-
-fun gen_make_case ty_match ty_inst type_of tab ctxt err used x clauses =
- let
- fun string_of_clause (pat, rhs) = Syntax.string_of_term ctxt
- (Syntax.const "_case1" $ pat $ rhs);
- val _ = map (no_repeat_vars ctxt o fst) clauses;
- val rows = map_index (fn (i, (pat, rhs)) =>
- (([], [pat]), (rhs, (i, true)))) clauses;
- val rangeT = (case distinct op = (map (type_of o snd) clauses) of
- [] => case_error "no clauses given"
- | [T] => T
- | _ => case_error "all cases must have the same result type");
- val used' = fold add_row_used rows used;
- val (patts, case_tm) = mk_case tab ctxt ty_match ty_inst type_of
- used' rangeT {path = [x], rows = rows}
- handle CASE_ERROR (msg, i) => case_error (msg ^
- (if i < 0 then ""
- else "\nIn clause\n" ^ string_of_clause (nth clauses i)));
- val patts1 = map
- (fn (_, tag, [pat]) => (pat, tag)
- | _ => case_error "error in pattern-match translation") patts;
- val patts2 = Library.sort (Library.int_ord o Library.pairself row_of_pat) patts1
- val finals = map row_of_pat patts2
- val originals = map (row_of_pat o #2) rows
- val _ = case originals \\ finals of
- [] => ()
- | is => (if err then case_error else warning)
- ("The following clauses are redundant (covered by preceding clauses):\n" ^
- cat_lines (map (string_of_clause o nth clauses) is));
- in
- (case_tm, patts2)
- end;
-
-fun make_case tab ctxt = gen_make_case
- (match_type (ProofContext.theory_of ctxt)) Envir.subst_TVars fastype_of tab ctxt;
-val make_case_untyped = gen_make_case (K (K Vartab.empty))
- (K (Term.map_types (K dummyT))) (K dummyT);
-
-
-(* parse translation *)
-
-fun case_tr err tab_of ctxt [t, u] =
- let
- val thy = ProofContext.theory_of ctxt;
- (* replace occurrences of dummy_pattern by distinct variables *)
- (* internalize constant names *)
- fun prep_pat ((c as Const ("_constrain", _)) $ t $ tT) used =
- let val (t', used') = prep_pat t used
- in (c $ t' $ tT, used') end
- | prep_pat (Const ("dummy_pattern", T)) used =
- let val x = Name.variant used "x"
- in (Free (x, T), x :: used) end
- | prep_pat (Const (s, T)) used =
- (case try (unprefix Syntax.constN) s of
- SOME c => (Const (c, T), used)
- | NONE => (Const (Sign.intern_const thy s, T), used))
- | prep_pat (v as Free (s, T)) used =
- let val s' = Sign.intern_const thy s
- in
- if Sign.declared_const thy s' then
- (Const (s', T), used)
- else (v, used)
- end
- | prep_pat (t $ u) used =
- let
- val (t', used') = prep_pat t used;
- val (u', used'') = prep_pat u used'
- in
- (t' $ u', used'')
- end
- | prep_pat t used = case_error ("Bad pattern: " ^ Syntax.string_of_term ctxt t);
- fun dest_case1 (t as Const ("_case1", _) $ l $ r) =
- let val (l', cnstrts) = strip_constraints l
- in ((fst (prep_pat l' (Term.add_free_names t [])), r), cnstrts)
- end
- | dest_case1 t = case_error "dest_case1";
- fun dest_case2 (Const ("_case2", _) $ t $ u) = t :: dest_case2 u
- | dest_case2 t = [t];
- val (cases, cnstrts) = split_list (map dest_case1 (dest_case2 u));
- val (case_tm, _) = make_case_untyped (tab_of thy) ctxt err []
- (fold (fn tT => fn t => Syntax.const "_constrain" $ t $ tT)
- (flat cnstrts) t) cases;
- in case_tm end
- | case_tr _ _ _ ts = case_error "case_tr";
-
-
-(*---------------------------------------------------------------------------
- * Pretty printing of nested case expressions
- *---------------------------------------------------------------------------*)
-
-(* destruct one level of pattern matching *)
-
-fun gen_dest_case name_of type_of tab d used t =
- case apfst name_of (strip_comb t) of
- (SOME cname, ts as _ :: _) =>
- let
- val (fs, x) = split_last ts;
- fun strip_abs i t =
- let
- val zs = strip_abs_vars t;
- val _ = if length zs < i then raise CASE_ERROR ("", 0) else ();
- val (xs, ys) = chop i zs;
- val u = list_abs (ys, strip_abs_body t);
- val xs' = map Free (Name.variant_list (OldTerm.add_term_names (u, used))
- (map fst xs) ~~ map snd xs)
- in (xs', subst_bounds (rev xs', u)) end;
- fun is_dependent i t =
- let val k = length (strip_abs_vars t) - i
- in k < 0 orelse exists (fn j => j >= k)
- (loose_bnos (strip_abs_body t))
- end;
- fun count_cases (_, _, true) = I
- | count_cases (c, (_, body), false) =
- AList.map_default op aconv (body, []) (cons c);
- val is_undefined = name_of #> equal (SOME "HOL.undefined");
- fun mk_case (c, (xs, body), _) = (list_comb (c, xs), body)
- in case ty_info tab cname of
- SOME {constructors, case_name} =>
- if length fs = length constructors then
- let
- val cases = map (fn (Const (s, U), t) =>
- let
- val k = length (binder_types U);
- val p as (xs, _) = strip_abs k t
- in
- (Const (s, map type_of xs ---> type_of x),
- p, is_dependent k t)
- end) (constructors ~~ fs);
- val cases' = sort (int_ord o swap o pairself (length o snd))
- (fold_rev count_cases cases []);
- val R = type_of t;
- val dummy = if d then Const ("dummy_pattern", R)
- else Free (Name.variant used "x", R)
- in
- SOME (x, map mk_case (case find_first (is_undefined o fst) cases' of
- SOME (_, cs) =>
- if length cs = length constructors then [hd cases]
- else filter_out (fn (_, (_, body), _) => is_undefined body) cases
- | NONE => case cases' of
- [] => cases
- | (default, cs) :: _ =>
- if length cs = 1 then cases
- else if length cs = length constructors then
- [hd cases, (dummy, ([], default), false)]
- else
- filter_out (fn (c, _, _) => member op aconv cs c) cases @
- [(dummy, ([], default), false)]))
- end handle CASE_ERROR _ => NONE
- else NONE
- | _ => NONE
- end
- | _ => NONE;
-
-val dest_case = gen_dest_case (try (dest_Const #> fst)) fastype_of;
-val dest_case' = gen_dest_case
- (try (dest_Const #> fst #> unprefix Syntax.constN)) (K dummyT);
-
-
-(* destruct nested patterns *)
-
-fun strip_case'' dest (pat, rhs) =
- case dest (Term.add_free_names pat []) rhs of
- SOME (exp as Free _, clauses) =>
- if member op aconv (OldTerm.term_frees pat) exp andalso
- not (exists (fn (_, rhs') =>
- member op aconv (OldTerm.term_frees rhs') exp) clauses)
- then
- maps (strip_case'' dest) (map (fn (pat', rhs') =>
- (subst_free [(exp, pat')] pat, rhs')) clauses)
- else [(pat, rhs)]
- | _ => [(pat, rhs)];
-
-fun gen_strip_case dest t = case dest [] t of
- SOME (x, clauses) =>
- SOME (x, maps (strip_case'' dest) clauses)
- | NONE => NONE;
-
-val strip_case = gen_strip_case oo dest_case;
-val strip_case' = gen_strip_case oo dest_case';
-
-
-(* print translation *)
-
-fun case_tr' tab_of cname ctxt ts =
- let
- val thy = ProofContext.theory_of ctxt;
- val consts = ProofContext.consts_of ctxt;
- fun mk_clause (pat, rhs) =
- let val xs = Term.add_frees pat []
- in
- Syntax.const "_case1" $
- map_aterms
- (fn Free p => Syntax.mark_boundT p
- | Const (s, _) => Const (Consts.extern_early consts s, dummyT)
- | t => t) pat $
- map_aterms
- (fn x as Free (s, T) =>
- if member (op =) xs (s, T) then Syntax.mark_bound s else x
- | t => t) rhs
- end
- in case strip_case' (tab_of thy) true (list_comb (Syntax.const cname, ts)) of
- SOME (x, clauses) => Syntax.const "_case_syntax" $ x $
- foldr1 (fn (t, u) => Syntax.const "_case2" $ t $ u)
- (map mk_clause clauses)
- | NONE => raise Match
- end;
-
-end;
--- a/src/HOL/Tools/datatype_codegen.ML Wed Jun 10 21:04:36 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,451 +0,0 @@
-(* Title: HOL/Tools/datatype_codegen.ML
- Author: Stefan Berghofer and Florian Haftmann, TU Muenchen
-
-Code generator facilities for inductive datatypes.
-*)
-
-signature DATATYPE_CODEGEN =
-sig
- val find_shortest_path: DatatypeAux.descr -> int -> (string * int) option
- val mk_eq_eqns: theory -> string -> (thm * bool) list
- val mk_case_cert: theory -> string -> thm
- val setup: theory -> theory
-end;
-
-structure DatatypeCodegen : DATATYPE_CODEGEN =
-struct
-
-(** SML code generator **)
-
-open Codegen;
-
-(**** datatype definition ****)
-
-(* find shortest path to constructor with no recursive arguments *)
-
-fun find_nonempty (descr: DatatypeAux.descr) is i =
- let
- val (_, _, constrs) = the (AList.lookup (op =) descr i);
- fun arg_nonempty (_, DatatypeAux.DtRec i) = if member (op =) is i
- then NONE
- else Option.map (curry op + 1 o snd) (find_nonempty descr (i::is) i)
- | arg_nonempty _ = SOME 0;
- fun max xs = Library.foldl
- (fn (NONE, _) => NONE
- | (SOME i, SOME j) => SOME (Int.max (i, j))
- | (_, NONE) => NONE) (SOME 0, xs);
- val xs = sort (int_ord o pairself snd)
- (map_filter (fn (s, dts) => Option.map (pair s)
- (max (map (arg_nonempty o DatatypeAux.strip_dtyp) dts))) constrs)
- in case xs of [] => NONE | x :: _ => SOME x end;
-
-fun find_shortest_path descr i = find_nonempty descr [i] i;
-
-fun add_dt_defs thy defs dep module (descr: DatatypeAux.descr) sorts gr =
- let
- val descr' = List.filter (can (map DatatypeAux.dest_DtTFree o #2 o snd)) descr;
- val rtnames = map (#1 o snd) (List.filter (fn (_, (_, _, cs)) =>
- exists (exists DatatypeAux.is_rec_type o snd) cs) descr');
-
- val (_, (tname, _, _)) :: _ = descr';
- val node_id = tname ^ " (type)";
- val module' = if_library (thyname_of_type thy tname) module;
-
- fun mk_dtdef prfx [] gr = ([], gr)
- | mk_dtdef prfx ((_, (tname, dts, cs))::xs) gr =
- let
- val tvs = map DatatypeAux.dest_DtTFree dts;
- val cs' = map (apsnd (map (DatatypeAux.typ_of_dtyp descr sorts))) cs;
- val ((_, type_id), gr') = mk_type_id module' tname gr;
- val (ps, gr'') = gr' |>
- fold_map (fn (cname, cargs) =>
- fold_map (invoke_tycodegen thy defs node_id module' false)
- cargs ##>>
- mk_const_id module' cname) cs';
- val (rest, gr''') = mk_dtdef "and " xs gr''
- in
- (Pretty.block (str prfx ::
- (if null tvs then [] else
- [mk_tuple (map str tvs), str " "]) @
- [str (type_id ^ " ="), Pretty.brk 1] @
- List.concat (separate [Pretty.brk 1, str "| "]
- (map (fn (ps', (_, cname)) => [Pretty.block
- (str cname ::
- (if null ps' then [] else
- List.concat ([str " of", Pretty.brk 1] ::
- separate [str " *", Pretty.brk 1]
- (map single ps'))))]) ps))) :: rest, gr''')
- end;
-
- fun mk_constr_term cname Ts T ps =
- List.concat (separate [str " $", Pretty.brk 1]
- ([str ("Const (\"" ^ cname ^ "\","), Pretty.brk 1,
- mk_type false (Ts ---> T), str ")"] :: ps));
-
- fun mk_term_of_def gr prfx [] = []
- | mk_term_of_def gr prfx ((_, (tname, dts, cs)) :: xs) =
- let
- val cs' = map (apsnd (map (DatatypeAux.typ_of_dtyp descr sorts))) cs;
- val dts' = map (DatatypeAux.typ_of_dtyp descr sorts) dts;
- val T = Type (tname, dts');
- val rest = mk_term_of_def gr "and " xs;
- val (eqs, _) = fold_map (fn (cname, Ts) => fn prfx =>
- let val args = map (fn i =>
- str ("x" ^ string_of_int i)) (1 upto length Ts)
- in (Pretty.blk (4,
- [str prfx, mk_term_of gr module' false T, Pretty.brk 1,
- if null Ts then str (snd (get_const_id gr cname))
- else parens (Pretty.block
- [str (snd (get_const_id gr cname)),
- Pretty.brk 1, mk_tuple args]),
- str " =", Pretty.brk 1] @
- mk_constr_term cname Ts T
- (map2 (fn x => fn U => [Pretty.block [mk_term_of gr module' false U,
- Pretty.brk 1, x]]) args Ts)), " | ")
- end) cs' prfx
- in eqs @ rest end;
-
- fun mk_gen_of_def gr prfx [] = []
- | mk_gen_of_def gr prfx ((i, (tname, dts, cs)) :: xs) =
- let
- val tvs = map DatatypeAux.dest_DtTFree dts;
- val Us = map (DatatypeAux.typ_of_dtyp descr sorts) dts;
- val T = Type (tname, Us);
- val (cs1, cs2) =
- List.partition (exists DatatypeAux.is_rec_type o snd) cs;
- val SOME (cname, _) = find_shortest_path descr i;
-
- fun mk_delay p = Pretty.block
- [str "fn () =>", Pretty.brk 1, p];
-
- fun mk_force p = Pretty.block [p, Pretty.brk 1, str "()"];
-
- fun mk_constr s b (cname, dts) =
- let
- val gs = map (fn dt => mk_app false (mk_gen gr module' false rtnames s
- (DatatypeAux.typ_of_dtyp descr sorts dt))
- [str (if b andalso DatatypeAux.is_rec_type dt then "0"
- else "j")]) dts;
- val Ts = map (DatatypeAux.typ_of_dtyp descr sorts) dts;
- val xs = map str
- (DatatypeProp.indexify_names (replicate (length dts) "x"));
- val ts = map str
- (DatatypeProp.indexify_names (replicate (length dts) "t"));
- val (_, id) = get_const_id gr cname
- in
- mk_let
- (map2 (fn p => fn q => mk_tuple [p, q]) xs ts ~~ gs)
- (mk_tuple
- [case xs of
- _ :: _ :: _ => Pretty.block
- [str id, Pretty.brk 1, mk_tuple xs]
- | _ => mk_app false (str id) xs,
- mk_delay (Pretty.block (mk_constr_term cname Ts T
- (map (single o mk_force) ts)))])
- end;
-
- fun mk_choice [c] = mk_constr "(i-1)" false c
- | mk_choice cs = Pretty.block [str "one_of",
- Pretty.brk 1, Pretty.blk (1, str "[" ::
- List.concat (separate [str ",", Pretty.fbrk]
- (map (single o mk_delay o mk_constr "(i-1)" false) cs)) @
- [str "]"]), Pretty.brk 1, str "()"];
-
- val gs = maps (fn s =>
- let val s' = strip_tname s
- in [str (s' ^ "G"), str (s' ^ "T")] end) tvs;
- val gen_name = "gen_" ^ snd (get_type_id gr tname)
-
- in
- Pretty.blk (4, separate (Pretty.brk 1)
- (str (prfx ^ gen_name ^
- (if null cs1 then "" else "'")) :: gs @
- (if null cs1 then [] else [str "i"]) @
- [str "j"]) @
- [str " =", Pretty.brk 1] @
- (if not (null cs1) andalso not (null cs2)
- then [str "frequency", Pretty.brk 1,
- Pretty.blk (1, [str "[",
- mk_tuple [str "i", mk_delay (mk_choice cs1)],
- str ",", Pretty.fbrk,
- mk_tuple [str "1", mk_delay (mk_choice cs2)],
- str "]"]), Pretty.brk 1, str "()"]
- else if null cs2 then
- [Pretty.block [str "(case", Pretty.brk 1,
- str "i", Pretty.brk 1, str "of",
- Pretty.brk 1, str "0 =>", Pretty.brk 1,
- mk_constr "0" true (cname, valOf (AList.lookup (op =) cs cname)),
- Pretty.brk 1, str "| _ =>", Pretty.brk 1,
- mk_choice cs1, str ")"]]
- else [mk_choice cs2])) ::
- (if null cs1 then []
- else [Pretty.blk (4, separate (Pretty.brk 1)
- (str ("and " ^ gen_name) :: gs @ [str "i"]) @
- [str " =", Pretty.brk 1] @
- separate (Pretty.brk 1) (str (gen_name ^ "'") :: gs @
- [str "i", str "i"]))]) @
- mk_gen_of_def gr "and " xs
- end
-
- in
- (module', (add_edge_acyclic (node_id, dep) gr
- handle Graph.CYCLES _ => gr) handle Graph.UNDEF _ =>
- let
- val gr1 = add_edge (node_id, dep)
- (new_node (node_id, (NONE, "", "")) gr);
- val (dtdef, gr2) = mk_dtdef "datatype " descr' gr1 ;
- in
- map_node node_id (K (NONE, module',
- string_of (Pretty.blk (0, separate Pretty.fbrk dtdef @
- [str ";"])) ^ "\n\n" ^
- (if "term_of" mem !mode then
- string_of (Pretty.blk (0, separate Pretty.fbrk
- (mk_term_of_def gr2 "fun " descr') @ [str ";"])) ^ "\n\n"
- else "") ^
- (if "test" mem !mode then
- string_of (Pretty.blk (0, separate Pretty.fbrk
- (mk_gen_of_def gr2 "fun " descr') @ [str ";"])) ^ "\n\n"
- else ""))) gr2
- end)
- end;
-
-
-(**** case expressions ****)
-
-fun pretty_case thy defs dep module brack constrs (c as Const (_, T)) ts gr =
- let val i = length constrs
- in if length ts <= i then
- invoke_codegen thy defs dep module brack (eta_expand c ts (i+1)) gr
- else
- let
- val ts1 = Library.take (i, ts);
- val t :: ts2 = Library.drop (i, ts);
- val names = List.foldr OldTerm.add_term_names
- (map (fst o fst o dest_Var) (List.foldr OldTerm.add_term_vars [] ts1)) ts1;
- val (Ts, dT) = split_last (Library.take (i+1, fst (strip_type T)));
-
- fun pcase [] [] [] gr = ([], gr)
- | pcase ((cname, cargs)::cs) (t::ts) (U::Us) gr =
- let
- val j = length cargs;
- val xs = Name.variant_list names (replicate j "x");
- val Us' = Library.take (j, fst (strip_type U));
- val frees = map Free (xs ~~ Us');
- val (cp, gr0) = invoke_codegen thy defs dep module false
- (list_comb (Const (cname, Us' ---> dT), frees)) gr;
- val t' = Envir.beta_norm (list_comb (t, frees));
- val (p, gr1) = invoke_codegen thy defs dep module false t' gr0;
- val (ps, gr2) = pcase cs ts Us gr1;
- in
- ([Pretty.block [cp, str " =>", Pretty.brk 1, p]] :: ps, gr2)
- end;
-
- val (ps1, gr1) = pcase constrs ts1 Ts gr ;
- val ps = List.concat (separate [Pretty.brk 1, str "| "] ps1);
- val (p, gr2) = invoke_codegen thy defs dep module false t gr1;
- val (ps2, gr3) = fold_map (invoke_codegen thy defs dep module true) ts2 gr2;
- in ((if not (null ts2) andalso brack then parens else I)
- (Pretty.block (separate (Pretty.brk 1)
- (Pretty.block ([str "(case ", p, str " of",
- Pretty.brk 1] @ ps @ [str ")"]) :: ps2))), gr3)
- end
- end;
-
-
-(**** constructors ****)
-
-fun pretty_constr thy defs dep module brack args (c as Const (s, T)) ts gr =
- let val i = length args
- in if i > 1 andalso length ts < i then
- invoke_codegen thy defs dep module brack (eta_expand c ts i) gr
- else
- let
- val id = mk_qual_id module (get_const_id gr s);
- val (ps, gr') = fold_map
- (invoke_codegen thy defs dep module (i = 1)) ts gr;
- in (case args of
- _ :: _ :: _ => (if brack then parens else I)
- (Pretty.block [str id, Pretty.brk 1, mk_tuple ps])
- | _ => (mk_app brack (str id) ps), gr')
- end
- end;
-
-
-(**** code generators for terms and types ****)
-
-fun datatype_codegen thy defs dep module brack t gr = (case strip_comb t of
- (c as Const (s, T), ts) =>
- (case DatatypePackage.datatype_of_case thy s of
- SOME {index, descr, ...} =>
- if is_some (get_assoc_code thy (s, T)) then NONE else
- SOME (pretty_case thy defs dep module brack
- (#3 (the (AList.lookup op = descr index))) c ts gr )
- | NONE => case (DatatypePackage.datatype_of_constr thy s, strip_type T) of
- (SOME {index, descr, ...}, (_, U as Type (tyname, _))) =>
- if is_some (get_assoc_code thy (s, T)) then NONE else
- let
- val SOME (tyname', _, constrs) = AList.lookup op = descr index;
- val SOME args = AList.lookup op = constrs s
- in
- if tyname <> tyname' then NONE
- else SOME (pretty_constr thy defs
- dep module brack args c ts (snd (invoke_tycodegen thy defs dep module false U gr)))
- end
- | _ => NONE)
- | _ => NONE);
-
-fun datatype_tycodegen thy defs dep module brack (Type (s, Ts)) gr =
- (case DatatypePackage.get_datatype thy s of
- NONE => NONE
- | SOME {descr, sorts, ...} =>
- if is_some (get_assoc_type thy s) then NONE else
- let
- val (ps, gr') = fold_map
- (invoke_tycodegen thy defs dep module false) Ts gr;
- val (module', gr'') = add_dt_defs thy defs dep module descr sorts gr' ;
- val (tyid, gr''') = mk_type_id module' s gr''
- in SOME (Pretty.block ((if null Ts then [] else
- [mk_tuple ps, str " "]) @
- [str (mk_qual_id module tyid)]), gr''')
- end)
- | datatype_tycodegen _ _ _ _ _ _ _ = NONE;
-
-
-(** generic code generator **)
-
-(* case certificates *)
-
-fun mk_case_cert thy tyco =
- let
- val raw_thms =
- (#case_rewrites o DatatypePackage.the_datatype thy) tyco;
- val thms as hd_thm :: _ = raw_thms
- |> Conjunction.intr_balanced
- |> Thm.unvarify
- |> Conjunction.elim_balanced (length raw_thms)
- |> map Simpdata.mk_meta_eq
- |> map Drule.zero_var_indexes
- val params = fold_aterms (fn (Free (v, _)) => insert (op =) v
- | _ => I) (Thm.prop_of hd_thm) [];
- val rhs = hd_thm
- |> Thm.prop_of
- |> Logic.dest_equals
- |> fst
- |> Term.strip_comb
- |> apsnd (fst o split_last)
- |> list_comb;
- val lhs = Free (Name.variant params "case", Term.fastype_of rhs);
- val asm = (Thm.cterm_of thy o Logic.mk_equals) (lhs, rhs);
- in
- thms
- |> Conjunction.intr_balanced
- |> MetaSimplifier.rewrite_rule [(Thm.symmetric o Thm.assume) asm]
- |> Thm.implies_intr asm
- |> Thm.generalize ([], params) 0
- |> AxClass.unoverload thy
- |> Thm.varifyT
- end;
-
-
-(* equality *)
-
-fun mk_eq_eqns thy dtco =
- let
- val (vs, cos) = DatatypePackage.the_datatype_spec thy dtco;
- val { descr, index, inject = inject_thms, ... } = DatatypePackage.the_datatype thy dtco;
- val ty = Type (dtco, map TFree vs);
- fun mk_eq (t1, t2) = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
- $ t1 $ t2;
- fun true_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.true_const);
- fun false_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.false_const);
- val triv_injects = map_filter
- (fn (c, []) => SOME (HOLogic.mk_Trueprop (true_eq (Const (c, ty), Const (c, ty))))
- | _ => NONE) cos;
- fun prep_inject (trueprop $ (equiv $ (_ $ t1 $ t2) $ rhs)) =
- trueprop $ (equiv $ mk_eq (t1, t2) $ rhs);
- val injects = map prep_inject (nth (DatatypeProp.make_injs [descr] vs) index);
- fun prep_distinct (trueprop $ (not $ (_ $ t1 $ t2))) =
- [trueprop $ false_eq (t1, t2), trueprop $ false_eq (t2, t1)];
- val distincts = maps prep_distinct (snd (nth (DatatypeProp.make_distincts [descr] vs) index));
- val refl = HOLogic.mk_Trueprop (true_eq (Free ("x", ty), Free ("x", ty)));
- val simpset = Simplifier.context (ProofContext.init thy) (HOL_basic_ss
- addsimps (map Simpdata.mk_eq (@{thm eq} :: @{thm eq_True} :: inject_thms))
- addsimprocs [DatatypePackage.distinct_simproc]);
- fun prove prop = Goal.prove_global thy [] [] prop (K (ALLGOALS (simp_tac simpset)))
- |> Simpdata.mk_eq;
- in map (rpair true o prove) (triv_injects @ injects @ distincts) @ [(prove refl, false)] end;
-
-fun add_equality vs dtcos thy =
- let
- fun add_def dtco lthy =
- let
- val ty = Type (dtco, map TFree vs);
- fun mk_side const_name = Const (const_name, ty --> ty --> HOLogic.boolT)
- $ Free ("x", ty) $ Free ("y", ty);
- val def = HOLogic.mk_Trueprop (HOLogic.mk_eq
- (mk_side @{const_name eq_class.eq}, mk_side @{const_name "op ="}));
- val def' = Syntax.check_term lthy def;
- val ((_, (_, thm)), lthy') = Specification.definition
- (NONE, (Attrib.empty_binding, def')) lthy;
- val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy);
- val thm' = singleton (ProofContext.export lthy' ctxt_thy) thm;
- in (thm', lthy') end;
- fun tac thms = Class.intro_classes_tac []
- THEN ALLGOALS (ProofContext.fact_tac thms);
- fun add_eq_thms dtco thy =
- let
- val const = AxClass.param_of_inst thy (@{const_name eq_class.eq}, dtco);
- val thy_ref = Theory.check_thy thy;
- fun mk_thms () = rev ((mk_eq_eqns (Theory.deref thy_ref) dtco));
- in
- Code.add_eqnl (const, Lazy.lazy mk_thms) thy
- end;
- in
- thy
- |> TheoryTarget.instantiation (dtcos, vs, [HOLogic.class_eq])
- |> fold_map add_def dtcos
- |-> (fn def_thms => Class.prove_instantiation_exit_result (map o Morphism.thm)
- (fn _ => fn def_thms => tac def_thms) def_thms)
- |-> (fn def_thms => fold Code.del_eqn def_thms)
- |> fold add_eq_thms dtcos
- end;
-
-
-(* liberal addition of code data for datatypes *)
-
-fun mk_constr_consts thy vs dtco cos =
- let
- val cs = map (fn (c, tys) => (c, tys ---> Type (dtco, map TFree vs))) cos;
- val cs' = map (fn c_ty as (_, ty) => (AxClass.unoverload_const thy c_ty, ty)) cs;
- in if is_some (try (Code.constrset_of_consts thy) cs')
- then SOME cs
- else NONE
- end;
-
-fun add_all_code dtcos thy =
- let
- val (vs :: _, coss) = (split_list o map (DatatypePackage.the_datatype_spec thy)) dtcos;
- val any_css = map2 (mk_constr_consts thy vs) dtcos coss;
- val css = if exists is_none any_css then []
- else map_filter I any_css;
- val case_rewrites = maps (#case_rewrites o DatatypePackage.the_datatype thy) dtcos;
- val certs = map (mk_case_cert thy) dtcos;
- in
- if null css then thy
- else thy
- |> fold Code.add_datatype css
- |> fold_rev Code.add_default_eqn case_rewrites
- |> fold Code.add_case certs
- |> add_equality vs dtcos
- end;
-
-
-
-(** theory setup **)
-
-val setup =
- add_codegen "datatype" datatype_codegen
- #> add_tycodegen "datatype" datatype_tycodegen
- #> DatatypePackage.interpretation add_all_code
-
-end;
--- a/src/HOL/Tools/datatype_package.ML Wed Jun 10 21:04:36 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,709 +0,0 @@
-(* Title: HOL/Tools/datatype_package.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Datatype package for Isabelle/HOL.
-*)
-
-signature DATATYPE_PACKAGE =
-sig
- val quiet_mode : bool ref
- val get_datatypes : theory -> DatatypeAux.datatype_info Symtab.table
- val print_datatypes : theory -> unit
- val get_datatype : theory -> string -> DatatypeAux.datatype_info option
- val the_datatype : theory -> string -> DatatypeAux.datatype_info
- val datatype_of_constr : theory -> string -> DatatypeAux.datatype_info option
- val datatype_of_case : theory -> string -> DatatypeAux.datatype_info option
- val the_datatype_spec : theory -> string -> (string * sort) list * (string * typ list) list
- val get_datatype_constrs : theory -> string -> (string * typ) list option
- val construction_interpretation : theory
- -> {atom : typ -> 'a, dtyp : string -> 'a, rtyp : string * typ list -> 'a list -> 'a}
- -> (string * sort) list -> string list
- -> (string * (string * 'a list) list) list
- * ((string * typ list) * (string * 'a list) list) list
- val distinct_simproc : simproc
- val make_case : Proof.context -> bool -> string list -> term ->
- (term * term) list -> term * (term * (int * bool)) list
- val strip_case : Proof.context -> bool -> term -> (term * (term * term) list) option
- val read_typ: theory ->
- (typ list * (string * sort) list) * string -> typ list * (string * sort) list
- val interpretation : (string list -> theory -> theory) -> theory -> theory
- val rep_datatype : ({distinct : thm list list,
- inject : thm list list,
- exhaustion : thm list,
- rec_thms : thm list,
- case_thms : thm list list,
- split_thms : (thm * thm) list,
- induction : thm,
- simps : thm list} -> Proof.context -> Proof.context) -> string list option -> term list
- -> theory -> Proof.state;
- val rep_datatype_cmd : string list option -> string list -> theory -> Proof.state;
- val add_datatype : bool -> bool -> string list -> (string list * binding * mixfix *
- (binding * typ list * mixfix) list) list -> theory ->
- {distinct : thm list list,
- inject : thm list list,
- exhaustion : thm list,
- rec_thms : thm list,
- case_thms : thm list list,
- split_thms : (thm * thm) list,
- induction : thm,
- simps : thm list} * theory
- val add_datatype_cmd : bool -> string list -> (string list * binding * mixfix *
- (binding * string list * mixfix) list) list -> theory ->
- {distinct : thm list list,
- inject : thm list list,
- exhaustion : thm list,
- rec_thms : thm list,
- case_thms : thm list list,
- split_thms : (thm * thm) list,
- induction : thm,
- simps : thm list} * theory
- val setup: theory -> theory
-end;
-
-structure DatatypePackage : DATATYPE_PACKAGE =
-struct
-
-open DatatypeAux;
-
-val quiet_mode = quiet_mode;
-
-
-(* theory data *)
-
-structure DatatypesData = TheoryDataFun
-(
- type T =
- {types: datatype_info Symtab.table,
- constrs: datatype_info Symtab.table,
- cases: datatype_info Symtab.table};
-
- val empty =
- {types = Symtab.empty, constrs = Symtab.empty, cases = Symtab.empty};
- val copy = I;
- val extend = I;
- fun merge _
- ({types = types1, constrs = constrs1, cases = cases1},
- {types = types2, constrs = constrs2, cases = cases2}) =
- {types = Symtab.merge (K true) (types1, types2),
- constrs = Symtab.merge (K true) (constrs1, constrs2),
- cases = Symtab.merge (K true) (cases1, cases2)};
-);
-
-val get_datatypes = #types o DatatypesData.get;
-val map_datatypes = DatatypesData.map;
-
-fun print_datatypes thy =
- Pretty.writeln (Pretty.strs ("datatypes:" ::
- map #1 (NameSpace.extern_table (Sign.type_space thy, get_datatypes thy))));
-
-
-(** theory information about datatypes **)
-
-fun put_dt_infos (dt_infos : (string * datatype_info) list) =
- map_datatypes (fn {types, constrs, cases} =>
- {types = fold Symtab.update dt_infos types,
- constrs = fold Symtab.default (*conservative wrt. overloaded constructors*)
- (maps (fn (_, info as {descr, index, ...}) => map (rpair info o fst)
- (#3 (the (AList.lookup op = descr index)))) dt_infos) constrs,
- cases = fold Symtab.update
- (map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)
- cases});
-
-val get_datatype = Symtab.lookup o get_datatypes;
-
-fun the_datatype thy name = (case get_datatype thy name of
- SOME info => info
- | NONE => error ("Unknown datatype " ^ quote name));
-
-val datatype_of_constr = Symtab.lookup o #constrs o DatatypesData.get;
-val datatype_of_case = Symtab.lookup o #cases o DatatypesData.get;
-
-fun get_datatype_descr thy dtco =
- get_datatype thy dtco
- |> Option.map (fn info as { descr, index, ... } =>
- (info, (((fn SOME (_, dtys, cos) => (dtys, cos)) o AList.lookup (op =) descr) index)));
-
-fun the_datatype_spec thy dtco =
- let
- val info as { descr, index, sorts = raw_sorts, ... } = the_datatype thy dtco;
- val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr index;
- val sorts = map ((fn v => (v, (the o AList.lookup (op =) raw_sorts) v))
- o DatatypeAux.dest_DtTFree) dtys;
- val cos = map
- (fn (co, tys) => (co, map (DatatypeAux.typ_of_dtyp descr sorts) tys)) raw_cos;
- in (sorts, cos) end;
-
-fun get_datatype_constrs thy dtco =
- case try (the_datatype_spec thy) dtco
- of SOME (sorts, cos) =>
- let
- fun subst (v, sort) = TVar ((v, 0), sort);
- fun subst_ty (TFree v) = subst v
- | subst_ty ty = ty;
- val dty = Type (dtco, map subst sorts);
- fun mk_co (co, tys) = (co, map (Term.map_atyps subst_ty) tys ---> dty);
- in SOME (map mk_co cos) end
- | NONE => NONE;
-
-fun construction_interpretation thy { atom, dtyp, rtyp } sorts tycos =
- let
- val descr = (#descr o the_datatype thy o hd) tycos;
- val k = length tycos;
- val descr_of = the o AList.lookup (op =) descr;
- val typ_of_dtyp = typ_of_dtyp descr sorts;
- fun interpT (dT as DtTFree _) = atom (typ_of_dtyp dT)
- | interpT (dT as DtType (tyco, dTs)) =
- let
- val Ts = map typ_of_dtyp dTs;
- in if is_rec_type dT
- then rtyp (tyco, Ts) (map interpT dTs)
- else atom (Type (tyco, Ts))
- end
- | interpT (DtRec l) = if l < k then (dtyp o #1 o descr_of) l
- else let
- val (tyco, dTs, _) = descr_of l;
- val Ts = map typ_of_dtyp dTs;
- in rtyp (tyco, Ts) (map interpT dTs) end;
- fun interpC (c, dTs) = (c, map interpT dTs);
- fun interpD (_, (tyco, _, cs)) = (tyco, map interpC cs);
- fun interpR (_, (tyco, dTs, cs)) = ((tyco, map typ_of_dtyp dTs), map interpC cs);
- in chop k descr |> (apfst o map) interpD |> (apsnd o map) interpR end;
-
-
-
-(** induct method setup **)
-
-(* case names *)
-
-local
-
-fun dt_recs (DtTFree _) = []
- | dt_recs (DtType (_, dts)) = maps dt_recs dts
- | dt_recs (DtRec i) = [i];
-
-fun dt_cases (descr: descr) (_, args, constrs) =
- let
- fun the_bname i = Long_Name.base_name (#1 (the (AList.lookup (op =) descr i)));
- val bnames = map the_bname (distinct (op =) (maps dt_recs args));
- in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end;
-
-
-fun induct_cases descr =
- DatatypeProp.indexify_names (maps (dt_cases descr) (map #2 descr));
-
-fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i));
-
-in
-
-fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
-
-fun mk_case_names_exhausts descr new =
- map (RuleCases.case_names o exhaust_cases descr o #1)
- (filter (fn ((_, (name, _, _))) => member (op =) new name) descr);
-
-end;
-
-fun add_rules simps case_thms rec_thms inject distinct
- weak_case_congs cong_att =
- PureThy.add_thmss [((Binding.name "simps", simps), []),
- ((Binding.empty, flat case_thms @
- flat distinct @ rec_thms), [Simplifier.simp_add]),
- ((Binding.empty, rec_thms), [Code.add_default_eqn_attribute]),
- ((Binding.empty, flat inject), [iff_add]),
- ((Binding.empty, map (fn th => th RS notE) (flat distinct)), [Classical.safe_elim NONE]),
- ((Binding.empty, weak_case_congs), [cong_att])]
- #> snd;
-
-
-(* add_cases_induct *)
-
-fun add_cases_induct infos induction thy =
- let
- val inducts = ProjectRule.projections (ProofContext.init thy) induction;
-
- fun named_rules (name, {index, exhaustion, ...}: datatype_info) =
- [((Binding.empty, nth inducts index), [Induct.induct_type name]),
- ((Binding.empty, exhaustion), [Induct.cases_type name])];
- fun unnamed_rule i =
- ((Binding.empty, nth inducts i), [Thm.kind_internal, Induct.induct_type ""]);
- in
- thy |> PureThy.add_thms
- (maps named_rules infos @
- map unnamed_rule (length infos upto length inducts - 1)) |> snd
- |> PureThy.add_thmss [((Binding.name "inducts", inducts), [])] |> snd
- end;
-
-
-
-(**** simplification procedure for showing distinctness of constructors ****)
-
-fun stripT (i, Type ("fun", [_, T])) = stripT (i + 1, T)
- | stripT p = p;
-
-fun stripC (i, f $ x) = stripC (i + 1, f)
- | stripC p = p;
-
-val distinctN = "constr_distinct";
-
-fun distinct_rule thy ss tname eq_t = case #distinct (the_datatype thy tname) of
- FewConstrs thms => Goal.prove (Simplifier.the_context ss) [] [] eq_t (K
- (EVERY [rtac eq_reflection 1, rtac iffI 1, rtac notE 1,
- atac 2, resolve_tac thms 1, etac FalseE 1]))
- | ManyConstrs (thm, simpset) =>
- let
- val [In0_inject, In1_inject, In0_not_In1, In1_not_In0] =
- map (PureThy.get_thm (ThyInfo.the_theory "Datatype" thy))
- ["In0_inject", "In1_inject", "In0_not_In1", "In1_not_In0"];
- in
- Goal.prove (Simplifier.the_context ss) [] [] eq_t (K
- (EVERY [rtac eq_reflection 1, rtac iffI 1, dtac thm 1,
- full_simp_tac (Simplifier.inherit_context ss simpset) 1,
- REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
- eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1,
- etac FalseE 1]))
- end;
-
-fun distinct_proc thy ss (t as Const ("op =", _) $ t1 $ t2) =
- (case (stripC (0, t1), stripC (0, t2)) of
- ((i, Const (cname1, T1)), (j, Const (cname2, T2))) =>
- (case (stripT (0, T1), stripT (0, T2)) of
- ((i', Type (tname1, _)), (j', Type (tname2, _))) =>
- if tname1 = tname2 andalso not (cname1 = cname2) andalso i = i' andalso j = j' then
- (case (get_datatype_descr thy) tname1 of
- SOME (_, (_, constrs)) => let val cnames = map fst constrs
- in if cname1 mem cnames andalso cname2 mem cnames then
- SOME (distinct_rule thy ss tname1
- (Logic.mk_equals (t, Const ("False", HOLogic.boolT))))
- else NONE
- end
- | NONE => NONE)
- else NONE
- | _ => NONE)
- | _ => NONE)
- | distinct_proc _ _ _ = NONE;
-
-val distinct_simproc =
- Simplifier.simproc @{theory HOL} distinctN ["s = t"] distinct_proc;
-
-val dist_ss = HOL_ss addsimprocs [distinct_simproc];
-
-val simproc_setup =
- Simplifier.map_simpset (fn ss => ss addsimprocs [distinct_simproc]);
-
-
-(**** translation rules for case ****)
-
-fun make_case ctxt = DatatypeCase.make_case
- (datatype_of_constr (ProofContext.theory_of ctxt)) ctxt;
-
-fun strip_case ctxt = DatatypeCase.strip_case
- (datatype_of_case (ProofContext.theory_of ctxt));
-
-fun add_case_tr' case_names thy =
- Sign.add_advanced_trfuns ([], [],
- map (fn case_name =>
- let val case_name' = Sign.const_syntax_name thy case_name
- in (case_name', DatatypeCase.case_tr' datatype_of_case case_name')
- end) case_names, []) thy;
-
-val trfun_setup =
- Sign.add_advanced_trfuns ([],
- [("_case_syntax", DatatypeCase.case_tr true datatype_of_constr)],
- [], []);
-
-
-(* prepare types *)
-
-fun read_typ thy ((Ts, sorts), str) =
- let
- val ctxt = ProofContext.init thy
- |> fold (Variable.declare_typ o TFree) sorts;
- val T = Syntax.read_typ ctxt str;
- in (Ts @ [T], Term.add_tfreesT T sorts) end;
-
-fun cert_typ sign ((Ts, sorts), raw_T) =
- let
- val T = Type.no_tvars (Sign.certify_typ sign raw_T) handle
- TYPE (msg, _, _) => error msg;
- val sorts' = Term.add_tfreesT T sorts;
- in (Ts @ [T],
- case duplicates (op =) (map fst sorts') of
- [] => sorts'
- | dups => error ("Inconsistent sort constraints for " ^ commas dups))
- end;
-
-
-(**** make datatype info ****)
-
-fun make_dt_info alt_names descr sorts induct reccomb_names rec_thms
- (((((((((i, (_, (tname, _, _))), case_name), case_thms),
- exhaustion_thm), distinct_thm), inject), nchotomy), case_cong), weak_case_cong) =
- (tname,
- {index = i,
- alt_names = alt_names,
- descr = descr,
- sorts = sorts,
- rec_names = reccomb_names,
- rec_rewrites = rec_thms,
- case_name = case_name,
- case_rewrites = case_thms,
- induction = induct,
- exhaustion = exhaustion_thm,
- distinct = distinct_thm,
- inject = inject,
- nchotomy = nchotomy,
- case_cong = case_cong,
- weak_case_cong = weak_case_cong});
-
-structure DatatypeInterpretation = InterpretationFun(type T = string list val eq = op =);
-val interpretation = DatatypeInterpretation.interpretation;
-
-
-(******************* definitional introduction of datatypes *******************)
-
-fun add_datatype_def flat_names new_type_names descr sorts types_syntax constr_syntax dt_info
- case_names_induct case_names_exhausts thy =
- let
- val _ = message ("Proofs for datatype(s) " ^ commas_quote new_type_names);
-
- val ((inject, distinct, dist_rewrites, simproc_dists, induct), thy2) = thy |>
- DatatypeRepProofs.representation_proofs flat_names dt_info new_type_names descr sorts
- types_syntax constr_syntax case_names_induct;
-
- val (casedist_thms, thy3) = DatatypeAbsProofs.prove_casedist_thms new_type_names descr
- sorts induct case_names_exhausts thy2;
- val ((reccomb_names, rec_thms), thy4) = DatatypeAbsProofs.prove_primrec_thms
- flat_names new_type_names descr sorts dt_info inject dist_rewrites
- (Simplifier.theory_context thy3 dist_ss) induct thy3;
- val ((case_thms, case_names), thy6) = DatatypeAbsProofs.prove_case_thms
- flat_names new_type_names descr sorts reccomb_names rec_thms thy4;
- val (split_thms, thy7) = DatatypeAbsProofs.prove_split_thms new_type_names
- descr sorts inject dist_rewrites casedist_thms case_thms thy6;
- val (nchotomys, thy8) = DatatypeAbsProofs.prove_nchotomys new_type_names
- descr sorts casedist_thms thy7;
- val (case_congs, thy9) = DatatypeAbsProofs.prove_case_congs new_type_names
- descr sorts nchotomys case_thms thy8;
- val (weak_case_congs, thy10) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
- descr sorts thy9;
-
- val dt_infos = map
- (make_dt_info (SOME new_type_names) (flat descr) sorts induct reccomb_names rec_thms)
- ((0 upto length (hd descr) - 1) ~~ (hd descr) ~~ case_names ~~ case_thms ~~
- casedist_thms ~~ simproc_dists ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
-
- val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
-
- val thy12 =
- thy10
- |> add_case_tr' case_names
- |> Sign.add_path (space_implode "_" new_type_names)
- |> add_rules simps case_thms rec_thms inject distinct
- weak_case_congs (Simplifier.attrib (op addcongs))
- |> put_dt_infos dt_infos
- |> add_cases_induct dt_infos induct
- |> Sign.parent_path
- |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms) |> snd
- |> DatatypeInterpretation.data (map fst dt_infos);
- in
- ({distinct = distinct,
- inject = inject,
- exhaustion = casedist_thms,
- rec_thms = rec_thms,
- case_thms = case_thms,
- split_thms = split_thms,
- induction = induct,
- simps = simps}, thy12)
- end;
-
-
-(*********************** declare existing type as datatype *********************)
-
-fun prove_rep_datatype alt_names new_type_names descr sorts induct inject half_distinct thy =
- let
- val ((_, [induct']), _) =
- Variable.importT_thms [induct] (Variable.thm_context induct);
-
- fun err t = error ("Ill-formed predicate in induction rule: " ^
- Syntax.string_of_term_global thy t);
-
- fun get_typ (t as _ $ Var (_, Type (tname, Ts))) =
- ((tname, map (fst o dest_TFree) Ts) handle TERM _ => err t)
- | get_typ t = err t;
- val dtnames = map get_typ (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct')));
-
- val dt_info = get_datatypes thy;
-
- val distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct;
- val (case_names_induct, case_names_exhausts) =
- (mk_case_names_induct descr, mk_case_names_exhausts descr (map #1 dtnames));
-
- val _ = message ("Proofs for datatype(s) " ^ commas_quote new_type_names);
-
- val (casedist_thms, thy2) = thy |>
- DatatypeAbsProofs.prove_casedist_thms new_type_names [descr] sorts induct
- case_names_exhausts;
- val ((reccomb_names, rec_thms), thy3) = DatatypeAbsProofs.prove_primrec_thms
- false new_type_names [descr] sorts dt_info inject distinct
- (Simplifier.theory_context thy2 dist_ss) induct thy2;
- val ((case_thms, case_names), thy4) = DatatypeAbsProofs.prove_case_thms false
- new_type_names [descr] sorts reccomb_names rec_thms thy3;
- val (split_thms, thy5) = DatatypeAbsProofs.prove_split_thms
- new_type_names [descr] sorts inject distinct casedist_thms case_thms thy4;
- val (nchotomys, thy6) = DatatypeAbsProofs.prove_nchotomys new_type_names
- [descr] sorts casedist_thms thy5;
- val (case_congs, thy7) = DatatypeAbsProofs.prove_case_congs new_type_names
- [descr] sorts nchotomys case_thms thy6;
- val (weak_case_congs, thy8) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
- [descr] sorts thy7;
-
- val ((_, [induct']), thy10) =
- thy8
- |> store_thmss "inject" new_type_names inject
- ||>> store_thmss "distinct" new_type_names distinct
- ||> Sign.add_path (space_implode "_" new_type_names)
- ||>> PureThy.add_thms [((Binding.name "induct", induct), [case_names_induct])];
-
- val dt_infos = map (make_dt_info alt_names descr sorts induct' reccomb_names rec_thms)
- ((0 upto length descr - 1) ~~ descr ~~ case_names ~~ case_thms ~~ casedist_thms ~~
- map FewConstrs distinct ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
-
- val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
-
- val thy11 =
- thy10
- |> add_case_tr' case_names
- |> add_rules simps case_thms rec_thms inject distinct
- weak_case_congs (Simplifier.attrib (op addcongs))
- |> put_dt_infos dt_infos
- |> add_cases_induct dt_infos induct'
- |> Sign.parent_path
- |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms)
- |> snd
- |> DatatypeInterpretation.data (map fst dt_infos);
- in
- ({distinct = distinct,
- inject = inject,
- exhaustion = casedist_thms,
- rec_thms = rec_thms,
- case_thms = case_thms,
- split_thms = split_thms,
- induction = induct',
- simps = simps}, thy11)
- end;
-
-fun gen_rep_datatype prep_term after_qed alt_names raw_ts thy =
- let
- fun constr_of_term (Const (c, T)) = (c, T)
- | constr_of_term t =
- error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
- fun no_constr (c, T) = error ("Bad constructor: "
- ^ Sign.extern_const thy c ^ "::"
- ^ Syntax.string_of_typ_global thy T);
- fun type_of_constr (cT as (_, T)) =
- let
- val frees = OldTerm.typ_tfrees T;
- val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) T
- handle TYPE _ => no_constr cT
- val _ = if has_duplicates (eq_fst (op =)) vs then no_constr cT else ();
- val _ = if length frees <> length vs then no_constr cT else ();
- in (tyco, (vs, cT)) end;
-
- val raw_cs = AList.group (op =) (map (type_of_constr o constr_of_term o prep_term thy) raw_ts);
- val _ = case map_filter (fn (tyco, _) =>
- if Symtab.defined (get_datatypes thy) tyco then SOME tyco else NONE) raw_cs
- of [] => ()
- | tycos => error ("Type(s) " ^ commas (map quote tycos)
- ^ " already represented inductivly");
- val raw_vss = maps (map (map snd o fst) o snd) raw_cs;
- val ms = case distinct (op =) (map length raw_vss)
- of [n] => 0 upto n - 1
- | _ => error ("Different types in given constructors");
- fun inter_sort m = map (fn xs => nth xs m) raw_vss
- |> Library.foldr1 (Sorts.inter_sort (Sign.classes_of thy))
- val sorts = map inter_sort ms;
- val vs = Name.names Name.context Name.aT sorts;
-
- fun norm_constr (raw_vs, (c, T)) = (c, map_atyps
- (TFree o (the o AList.lookup (op =) (map fst raw_vs ~~ vs)) o fst o dest_TFree) T);
-
- val cs = map (apsnd (map norm_constr)) raw_cs;
- val dtyps_of_typ = map (dtyp_of_typ (map (rpair (map fst vs) o fst) cs))
- o fst o strip_type;
- val new_type_names = map Long_Name.base_name (the_default (map fst cs) alt_names);
-
- fun mk_spec (i, (tyco, constr)) = (i, (tyco,
- map (DtTFree o fst) vs,
- (map o apsnd) dtyps_of_typ constr))
- val descr = map_index mk_spec cs;
- val injs = DatatypeProp.make_injs [descr] vs;
- val half_distincts = map snd (DatatypeProp.make_distincts [descr] vs);
- val ind = DatatypeProp.make_ind [descr] vs;
- val rules = (map o map o map) Logic.close_form [[[ind]], injs, half_distincts];
-
- fun after_qed' raw_thms =
- let
- val [[[induct]], injs, half_distincts] =
- unflat rules (map Drule.zero_var_indexes_list raw_thms);
- (*FIXME somehow dubious*)
- in
- ProofContext.theory_result
- (prove_rep_datatype alt_names new_type_names descr vs induct injs half_distincts)
- #-> after_qed
- end;
- in
- thy
- |> ProofContext.init
- |> Proof.theorem_i NONE after_qed' ((map o map) (rpair []) (flat rules))
- end;
-
-val rep_datatype = gen_rep_datatype Sign.cert_term;
-val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global (K I);
-
-
-
-(******************************** add datatype ********************************)
-
-fun gen_add_datatype prep_typ err flat_names new_type_names dts thy =
- let
- val _ = Theory.requires thy "Datatype" "datatype definitions";
-
- (* this theory is used just for parsing *)
-
- val tmp_thy = thy |>
- Theory.copy |>
- Sign.add_types (map (fn (tvs, tname, mx, _) =>
- (tname, length tvs, mx)) dts);
-
- val (tyvars, _, _, _)::_ = dts;
- val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) =>
- let val full_tname = Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname)
- in (case duplicates (op =) tvs of
- [] => if eq_set (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
- else error ("Mutually recursive datatypes must have same type parameters")
- | dups => error ("Duplicate parameter(s) for datatype " ^ quote (Binding.str_of tname) ^
- " : " ^ commas dups))
- end) dts);
-
- val _ = (case duplicates (op =) (map fst new_dts) @ duplicates (op =) new_type_names of
- [] => () | dups => error ("Duplicate datatypes: " ^ commas dups));
-
- fun prep_dt_spec ((tvs, tname, mx, constrs), tname') (dts', constr_syntax, sorts, i) =
- let
- fun prep_constr (cname, cargs, mx') (constrs, constr_syntax', sorts') =
- let
- val (cargs', sorts'') = Library.foldl (prep_typ tmp_thy) (([], sorts'), cargs);
- val _ = (case fold (curry OldTerm.add_typ_tfree_names) cargs' [] \\ tvs of
- [] => ()
- | vs => error ("Extra type variables on rhs: " ^ commas vs))
- in (constrs @ [((if flat_names then Sign.full_name tmp_thy else
- Sign.full_name_path tmp_thy tname')
- (Binding.map_name (Syntax.const_name mx') cname),
- map (dtyp_of_typ new_dts) cargs')],
- constr_syntax' @ [(cname, mx')], sorts'')
- end handle ERROR msg => cat_error msg
- ("The error above occured in constructor " ^ quote (Binding.str_of cname) ^
- " of datatype " ^ quote (Binding.str_of tname));
-
- val (constrs', constr_syntax', sorts') =
- fold prep_constr constrs ([], [], sorts)
-
- in
- case duplicates (op =) (map fst constrs') of
- [] =>
- (dts' @ [(i, (Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname),
- map DtTFree tvs, constrs'))],
- constr_syntax @ [constr_syntax'], sorts', i + 1)
- | dups => error ("Duplicate constructors " ^ commas dups ^
- " in datatype " ^ quote (Binding.str_of tname))
- end;
-
- val (dts', constr_syntax, sorts', i) =
- fold prep_dt_spec (dts ~~ new_type_names) ([], [], [], 0);
- val sorts = sorts' @ (map (rpair (Sign.defaultS tmp_thy)) (tyvars \\ map fst sorts'));
- val dt_info = get_datatypes thy;
- val (descr, _) = unfold_datatypes tmp_thy dts' sorts dt_info dts' i;
- val _ = check_nonempty descr handle (exn as Datatype_Empty s) =>
- if err then error ("Nonemptiness check failed for datatype " ^ s)
- else raise exn;
-
- val descr' = flat descr;
- val case_names_induct = mk_case_names_induct descr';
- val case_names_exhausts = mk_case_names_exhausts descr' (map #1 new_dts);
- in
- add_datatype_def
- flat_names new_type_names descr sorts types_syntax constr_syntax dt_info
- case_names_induct case_names_exhausts thy
- end;
-
-val add_datatype = gen_add_datatype cert_typ;
-val add_datatype_cmd = gen_add_datatype read_typ true;
-
-
-
-(** package setup **)
-
-(* setup theory *)
-
-val setup =
- DatatypeRepProofs.distinctness_limit_setup #>
- simproc_setup #>
- trfun_setup #>
- DatatypeInterpretation.init;
-
-
-(* outer syntax *)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val datatype_decl =
- Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.binding -- P.opt_infix --
- (P.$$$ "=" |-- P.enum1 "|" (P.binding -- Scan.repeat P.typ -- P.opt_mixfix));
-
-fun mk_datatype args =
- let
- val names = map
- (fn ((((NONE, _), t), _), _) => Binding.name_of t | ((((SOME t, _), _), _), _) => t) args;
- val specs = map (fn ((((_, vs), t), mx), cons) =>
- (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
- in snd o add_datatype_cmd false names specs end;
-
-val _ =
- OuterSyntax.command "datatype" "define inductive datatypes" K.thy_decl
- (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
-
-val _ =
- OuterSyntax.command "rep_datatype" "represent existing types inductively" K.thy_goal
- (Scan.option (P.$$$ "(" |-- Scan.repeat1 P.name --| P.$$$ ")") -- Scan.repeat1 P.term
- >> (fn (alt_names, ts) => Toplevel.print
- o Toplevel.theory_to_proof (rep_datatype_cmd alt_names ts)));
-
-end;
-
-
-(* document antiquotation *)
-
-val _ = ThyOutput.antiquotation "datatype" Args.tyname
- (fn {source = src, context = ctxt, ...} => fn dtco =>
- let
- val thy = ProofContext.theory_of ctxt;
- val (vs, cos) = the_datatype_spec thy dtco;
- val ty = Type (dtco, map TFree vs);
- fun pretty_typ_bracket (ty as Type (_, _ :: _)) =
- Pretty.enclose "(" ")" [Syntax.pretty_typ ctxt ty]
- | pretty_typ_bracket ty =
- Syntax.pretty_typ ctxt ty;
- fun pretty_constr (co, tys) =
- (Pretty.block o Pretty.breaks)
- (Syntax.pretty_term ctxt (Const (co, tys ---> ty)) ::
- map pretty_typ_bracket tys);
- val pretty_datatype =
- Pretty.block
- (Pretty.command "datatype" :: Pretty.brk 1 ::
- Syntax.pretty_typ ctxt ty ::
- Pretty.str " =" :: Pretty.brk 1 ::
- flat (separate [Pretty.brk 1, Pretty.str "| "]
- (map (single o pretty_constr) cos)));
- in ThyOutput.output (ThyOutput.maybe_pretty_source (K pretty_datatype) src [()]) end);
-
-end;
-
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/datatype_package/datatype_abs_proofs.ML Thu Jun 11 08:04:26 2009 +0200
@@ -0,0 +1,446 @@
+(* Title: HOL/Tools/datatype_abs_proofs.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Proofs and defintions independent of concrete representation
+of datatypes (i.e. requiring only abstract properties such as
+injectivity / distinctness of constructors and induction)
+
+ - case distinction (exhaustion) theorems
+ - characteristic equations for primrec combinators
+ - characteristic equations for case combinators
+ - equations for splitting "P (case ...)" expressions
+ - "nchotomy" and "case_cong" theorems for TFL
+*)
+
+signature DATATYPE_ABS_PROOFS =
+sig
+ val prove_casedist_thms : string list ->
+ DatatypeAux.descr list -> (string * sort) list -> thm ->
+ attribute list -> theory -> thm list * theory
+ val prove_primrec_thms : bool -> string list ->
+ DatatypeAux.descr list -> (string * sort) list ->
+ DatatypeAux.datatype_info Symtab.table -> thm list list -> thm list list ->
+ simpset -> thm -> theory -> (string list * thm list) * theory
+ val prove_case_thms : bool -> string list ->
+ DatatypeAux.descr list -> (string * sort) list ->
+ string list -> thm list -> theory -> (thm list list * string list) * theory
+ val prove_split_thms : string list ->
+ DatatypeAux.descr list -> (string * sort) list ->
+ thm list list -> thm list list -> thm list -> thm list list -> theory ->
+ (thm * thm) list * theory
+ val prove_nchotomys : string list -> DatatypeAux.descr list ->
+ (string * sort) list -> thm list -> theory -> thm list * theory
+ val prove_weak_case_congs : string list -> DatatypeAux.descr list ->
+ (string * sort) list -> theory -> thm list * theory
+ val prove_case_congs : string list ->
+ DatatypeAux.descr list -> (string * sort) list ->
+ thm list -> thm list list -> theory -> thm list * theory
+end;
+
+structure DatatypeAbsProofs: DATATYPE_ABS_PROOFS =
+struct
+
+open DatatypeAux;
+
+(************************ case distinction theorems ***************************)
+
+fun prove_casedist_thms new_type_names descr sorts induct case_names_exhausts thy =
+ let
+ val _ = message "Proving case distinction theorems ...";
+
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val newTs = Library.take (length (hd descr), recTs);
+
+ val {maxidx, ...} = rep_thm induct;
+ val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
+
+ fun prove_casedist_thm ((i, t), T) =
+ let
+ val dummyPs = map (fn (Var (_, Type (_, [T', T'']))) =>
+ Abs ("z", T', Const ("True", T''))) induct_Ps;
+ val P = Abs ("z", T, HOLogic.imp $ HOLogic.mk_eq (Var (("a", maxidx+1), T), Bound 0) $
+ Var (("P", 0), HOLogic.boolT))
+ val insts = Library.take (i, dummyPs) @ (P::(Library.drop (i + 1, dummyPs)));
+ val cert = cterm_of thy;
+ val insts' = (map cert induct_Ps) ~~ (map cert insts);
+ val induct' = refl RS ((List.nth
+ (split_conj_thm (cterm_instantiate insts' induct), i)) RSN (2, rev_mp))
+
+ in
+ SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
+ (fn {prems, ...} => EVERY
+ [rtac induct' 1,
+ REPEAT (rtac TrueI 1),
+ REPEAT ((rtac impI 1) THEN (eresolve_tac prems 1)),
+ REPEAT (rtac TrueI 1)])
+ end;
+
+ val casedist_thms = map prove_casedist_thm ((0 upto (length newTs - 1)) ~~
+ (DatatypeProp.make_casedists descr sorts) ~~ newTs)
+ in
+ thy
+ |> store_thms_atts "exhaust" new_type_names (map single case_names_exhausts) casedist_thms
+ end;
+
+
+(*************************** primrec combinators ******************************)
+
+fun prove_primrec_thms flat_names new_type_names descr sorts
+ (dt_info : datatype_info Symtab.table) constr_inject dist_rewrites dist_ss induct thy =
+ let
+ val _ = message "Constructing primrec combinators ...";
+
+ val big_name = space_implode "_" new_type_names;
+ val thy0 = add_path flat_names big_name thy;
+
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
+ val newTs = Library.take (length (hd descr), recTs);
+
+ val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct)));
+
+ val big_rec_name' = big_name ^ "_rec_set";
+ val rec_set_names' =
+ if length descr' = 1 then [big_rec_name'] else
+ map ((curry (op ^) (big_rec_name' ^ "_")) o string_of_int)
+ (1 upto (length descr'));
+ val rec_set_names = map (Sign.full_bname thy0) rec_set_names';
+
+ val (rec_result_Ts, reccomb_fn_Ts) = DatatypeProp.make_primrec_Ts descr sorts used;
+
+ val rec_set_Ts = map (fn (T1, T2) =>
+ reccomb_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts);
+
+ val rec_fns = map (uncurry (mk_Free "f"))
+ (reccomb_fn_Ts ~~ (1 upto (length reccomb_fn_Ts)));
+ val rec_sets' = map (fn c => list_comb (Free c, rec_fns))
+ (rec_set_names' ~~ rec_set_Ts);
+ val rec_sets = map (fn c => list_comb (Const c, rec_fns))
+ (rec_set_names ~~ rec_set_Ts);
+
+ (* introduction rules for graph of primrec function *)
+
+ fun make_rec_intr T rec_set ((rec_intr_ts, l), (cname, cargs)) =
+ let
+ fun mk_prem ((dt, U), (j, k, prems, t1s, t2s)) =
+ let val free1 = mk_Free "x" U j
+ in (case (strip_dtyp dt, strip_type U) of
+ ((_, DtRec m), (Us, _)) =>
+ let
+ val free2 = mk_Free "y" (Us ---> List.nth (rec_result_Ts, m)) k;
+ val i = length Us
+ in (j + 1, k + 1, HOLogic.mk_Trueprop (HOLogic.list_all
+ (map (pair "x") Us, List.nth (rec_sets', m) $
+ app_bnds free1 i $ app_bnds free2 i)) :: prems,
+ free1::t1s, free2::t2s)
+ end
+ | _ => (j + 1, k, prems, free1::t1s, t2s))
+ end;
+
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val (_, _, prems, t1s, t2s) = List.foldr mk_prem (1, 1, [], [], []) (cargs ~~ Ts)
+
+ in (rec_intr_ts @ [Logic.list_implies (prems, HOLogic.mk_Trueprop
+ (rec_set $ list_comb (Const (cname, Ts ---> T), t1s) $
+ list_comb (List.nth (rec_fns, l), t1s @ t2s)))], l + 1)
+ end;
+
+ val (rec_intr_ts, _) = Library.foldl (fn (x, ((d, T), set_name)) =>
+ Library.foldl (make_rec_intr T set_name) (x, #3 (snd d)))
+ (([], 0), descr' ~~ recTs ~~ rec_sets');
+
+ val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
+ InductivePackage.add_inductive_global (serial_string ())
+ {quiet_mode = ! quiet_mode, verbose = false, kind = Thm.internalK,
+ alt_name = Binding.name big_rec_name', coind = false, no_elim = false, no_ind = true,
+ skip_mono = true, fork_mono = false}
+ (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
+ (map dest_Free rec_fns)
+ (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] thy0;
+
+ (* prove uniqueness and termination of primrec combinators *)
+
+ val _ = message "Proving termination and uniqueness of primrec functions ...";
+
+ fun mk_unique_tac ((tac, intrs), ((((i, (tname, _, constrs)), elim), T), T')) =
+ let
+ val distinct_tac =
+ (if i < length newTs then
+ full_simp_tac (HOL_ss addsimps (List.nth (dist_rewrites, i))) 1
+ else full_simp_tac dist_ss 1);
+
+ val inject = map (fn r => r RS iffD1)
+ (if i < length newTs then List.nth (constr_inject, i)
+ else #inject (the (Symtab.lookup dt_info tname)));
+
+ fun mk_unique_constr_tac n ((tac, intr::intrs, j), (cname, cargs)) =
+ let
+ val k = length (List.filter is_rec_type cargs)
+
+ in (EVERY [DETERM tac,
+ REPEAT (etac ex1E 1), rtac ex1I 1,
+ DEPTH_SOLVE_1 (ares_tac [intr] 1),
+ REPEAT_DETERM_N k (etac thin_rl 1 THEN rotate_tac 1 1),
+ etac elim 1,
+ REPEAT_DETERM_N j distinct_tac,
+ TRY (dresolve_tac inject 1),
+ REPEAT (etac conjE 1), hyp_subst_tac 1,
+ REPEAT (EVERY [etac allE 1, dtac mp 1, atac 1]),
+ TRY (hyp_subst_tac 1),
+ rtac refl 1,
+ REPEAT_DETERM_N (n - j - 1) distinct_tac],
+ intrs, j + 1)
+ end;
+
+ val (tac', intrs', _) = Library.foldl (mk_unique_constr_tac (length constrs))
+ ((tac, intrs, 0), constrs);
+
+ in (tac', intrs') end;
+
+ val rec_unique_thms =
+ let
+ val rec_unique_ts = map (fn (((set_t, T1), T2), i) =>
+ Const ("Ex1", (T2 --> HOLogic.boolT) --> HOLogic.boolT) $
+ absfree ("y", T2, set_t $ mk_Free "x" T1 i $ Free ("y", T2)))
+ (rec_sets ~~ recTs ~~ rec_result_Ts ~~ (1 upto length recTs));
+ val cert = cterm_of thy1
+ val insts = map (fn ((i, T), t) => absfree ("x" ^ (string_of_int i), T, t))
+ ((1 upto length recTs) ~~ recTs ~~ rec_unique_ts);
+ val induct' = cterm_instantiate ((map cert induct_Ps) ~~
+ (map cert insts)) induct;
+ val (tac, _) = Library.foldl mk_unique_tac
+ (((rtac induct' THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1
+ THEN rewrite_goals_tac [mk_meta_eq choice_eq], rec_intrs),
+ descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts);
+
+ in split_conj_thm (SkipProof.prove_global thy1 [] []
+ (HOLogic.mk_Trueprop (mk_conj rec_unique_ts)) (K tac))
+ end;
+
+ val rec_total_thms = map (fn r => r RS theI') rec_unique_thms;
+
+ (* define primrec combinators *)
+
+ val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
+ val reccomb_names = map (Sign.full_bname thy1)
+ (if length descr' = 1 then [big_reccomb_name] else
+ (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
+ (1 upto (length descr'))));
+ val reccombs = map (fn ((name, T), T') => list_comb
+ (Const (name, reccomb_fn_Ts @ [T] ---> T'), rec_fns))
+ (reccomb_names ~~ recTs ~~ rec_result_Ts);
+
+ val (reccomb_defs, thy2) =
+ thy1
+ |> Sign.add_consts_i (map (fn ((name, T), T') =>
+ (Binding.name (Long_Name.base_name name), reccomb_fn_Ts @ [T] ---> T', NoSyn))
+ (reccomb_names ~~ recTs ~~ rec_result_Ts))
+ |> (PureThy.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') =>
+ (Binding.name (Long_Name.base_name name ^ "_def"), Logic.mk_equals (comb, absfree ("x", T,
+ Const ("The", (T' --> HOLogic.boolT) --> T') $ absfree ("y", T',
+ set $ Free ("x", T) $ Free ("y", T'))))))
+ (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts))
+ ||> parent_path flat_names
+ ||> Theory.checkpoint;
+
+
+ (* prove characteristic equations for primrec combinators *)
+
+ val _ = message "Proving characteristic theorems for primrec combinators ..."
+
+ val rec_thms = map (fn t => SkipProof.prove_global thy2 [] [] t
+ (fn _ => EVERY
+ [rewrite_goals_tac reccomb_defs,
+ rtac the1_equality 1,
+ resolve_tac rec_unique_thms 1,
+ resolve_tac rec_intrs 1,
+ REPEAT (rtac allI 1 ORELSE resolve_tac rec_total_thms 1)]))
+ (DatatypeProp.make_primrecs new_type_names descr sorts thy2)
+
+ in
+ thy2
+ |> Sign.add_path (space_implode "_" new_type_names)
+ |> PureThy.add_thmss [((Binding.name "recs", rec_thms),
+ [Nitpick_Const_Simp_Thms.add])]
+ ||> Sign.parent_path
+ ||> Theory.checkpoint
+ |-> (fn thms => pair (reccomb_names, Library.flat thms))
+ end;
+
+
+(***************************** case combinators *******************************)
+
+fun prove_case_thms flat_names new_type_names descr sorts reccomb_names primrec_thms thy =
+ let
+ val _ = message "Proving characteristic theorems for case combinators ...";
+
+ val thy1 = add_path flat_names (space_implode "_" new_type_names) thy;
+
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
+ val newTs = Library.take (length (hd descr), recTs);
+ val T' = TFree (Name.variant used "'t", HOLogic.typeS);
+
+ fun mk_dummyT dt = binder_types (typ_of_dtyp descr' sorts dt) ---> T';
+
+ val case_dummy_fns = map (fn (_, (_, _, constrs)) => map (fn (_, cargs) =>
+ let
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val Ts' = map mk_dummyT (List.filter is_rec_type cargs)
+ in Const (@{const_name undefined}, Ts @ Ts' ---> T')
+ end) constrs) descr';
+
+ val case_names = map (fn s => Sign.full_bname thy1 (s ^ "_case")) new_type_names;
+
+ (* define case combinators via primrec combinators *)
+
+ val (case_defs, thy2) = Library.foldl (fn ((defs, thy),
+ ((((i, (_, _, constrs)), T), name), recname)) =>
+ let
+ val (fns1, fns2) = ListPair.unzip (map (fn ((_, cargs), j) =>
+ let
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val Ts' = Ts @ map mk_dummyT (List.filter is_rec_type cargs);
+ val frees' = map (uncurry (mk_Free "x")) (Ts' ~~ (1 upto length Ts'));
+ val frees = Library.take (length cargs, frees');
+ val free = mk_Free "f" (Ts ---> T') j
+ in
+ (free, list_abs_free (map dest_Free frees',
+ list_comb (free, frees)))
+ end) (constrs ~~ (1 upto length constrs)));
+
+ val caseT = (map (snd o dest_Free) fns1) @ [T] ---> T';
+ val fns = (List.concat (Library.take (i, case_dummy_fns))) @
+ fns2 @ (List.concat (Library.drop (i + 1, case_dummy_fns)));
+ val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T');
+ val decl = ((Binding.name (Long_Name.base_name name), caseT), NoSyn);
+ val def = (Binding.name (Long_Name.base_name name ^ "_def"),
+ Logic.mk_equals (list_comb (Const (name, caseT), fns1),
+ list_comb (reccomb, (List.concat (Library.take (i, case_dummy_fns))) @
+ fns2 @ (List.concat (Library.drop (i + 1, case_dummy_fns))) )));
+ val ([def_thm], thy') =
+ thy
+ |> Sign.declare_const [] decl |> snd
+ |> (PureThy.add_defs false o map Thm.no_attributes) [def];
+
+ in (defs @ [def_thm], thy')
+ end) (([], thy1), (hd descr) ~~ newTs ~~ case_names ~~
+ (Library.take (length newTs, reccomb_names)))
+ ||> Theory.checkpoint;
+
+ val case_thms = map (map (fn t => SkipProof.prove_global thy2 [] [] t
+ (fn _ => EVERY [rewrite_goals_tac (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1])))
+ (DatatypeProp.make_cases new_type_names descr sorts thy2)
+ in
+ thy2
+ |> Context.the_theory o fold (fold Nitpick_Const_Simp_Thms.add_thm) case_thms
+ o Context.Theory
+ |> parent_path flat_names
+ |> store_thmss "cases" new_type_names case_thms
+ |-> (fn thmss => pair (thmss, case_names))
+ end;
+
+
+(******************************* case splitting *******************************)
+
+fun prove_split_thms new_type_names descr sorts constr_inject dist_rewrites
+ casedist_thms case_thms thy =
+ let
+ val _ = message "Proving equations for case splitting ...";
+
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val newTs = Library.take (length (hd descr), recTs);
+
+ fun prove_split_thms ((((((t1, t2), inject), dist_rewrites'),
+ exhaustion), case_thms'), T) =
+ let
+ val cert = cterm_of thy;
+ val _ $ (_ $ lhs $ _) = hd (Logic.strip_assums_hyp (hd (prems_of exhaustion)));
+ val exhaustion' = cterm_instantiate
+ [(cert lhs, cert (Free ("x", T)))] exhaustion;
+ val tacf = K (EVERY [rtac exhaustion' 1, ALLGOALS (asm_simp_tac
+ (HOL_ss addsimps (dist_rewrites' @ inject @ case_thms')))])
+ in
+ (SkipProof.prove_global thy [] [] t1 tacf,
+ SkipProof.prove_global thy [] [] t2 tacf)
+ end;
+
+ val split_thm_pairs = map prove_split_thms
+ ((DatatypeProp.make_splits new_type_names descr sorts thy) ~~ constr_inject ~~
+ dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs);
+
+ val (split_thms, split_asm_thms) = ListPair.unzip split_thm_pairs
+
+ in
+ thy
+ |> store_thms "split" new_type_names split_thms
+ ||>> store_thms "split_asm" new_type_names split_asm_thms
+ |-> (fn (thms1, thms2) => pair (thms1 ~~ thms2))
+ end;
+
+fun prove_weak_case_congs new_type_names descr sorts thy =
+ let
+ fun prove_weak_case_cong t =
+ SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
+ (fn {prems, ...} => EVERY [rtac ((hd prems) RS arg_cong) 1])
+
+ val weak_case_congs = map prove_weak_case_cong (DatatypeProp.make_weak_case_congs
+ new_type_names descr sorts thy)
+
+ in thy |> store_thms "weak_case_cong" new_type_names weak_case_congs end;
+
+(************************* additional theorems for TFL ************************)
+
+fun prove_nchotomys new_type_names descr sorts casedist_thms thy =
+ let
+ val _ = message "Proving additional theorems for TFL ...";
+
+ fun prove_nchotomy (t, exhaustion) =
+ let
+ (* For goal i, select the correct disjunct to attack, then prove it *)
+ fun tac i 0 = EVERY [TRY (rtac disjI1 i),
+ hyp_subst_tac i, REPEAT (rtac exI i), rtac refl i]
+ | tac i n = rtac disjI2 i THEN tac i (n - 1)
+ in
+ SkipProof.prove_global thy [] [] t (fn _ =>
+ EVERY [rtac allI 1,
+ exh_tac (K exhaustion) 1,
+ ALLGOALS (fn i => tac i (i-1))])
+ end;
+
+ val nchotomys =
+ map prove_nchotomy (DatatypeProp.make_nchotomys descr sorts ~~ casedist_thms)
+
+ in thy |> store_thms "nchotomy" new_type_names nchotomys end;
+
+fun prove_case_congs new_type_names descr sorts nchotomys case_thms thy =
+ let
+ fun prove_case_cong ((t, nchotomy), case_rewrites) =
+ let
+ val (Const ("==>", _) $ tm $ _) = t;
+ val (Const ("Trueprop", _) $ (Const ("op =", _) $ _ $ Ma)) = tm;
+ val cert = cterm_of thy;
+ val nchotomy' = nchotomy RS spec;
+ val [v] = Term.add_vars (concl_of nchotomy') [];
+ val nchotomy'' = cterm_instantiate [(cert (Var v), cert Ma)] nchotomy'
+ in
+ SkipProof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
+ (fn {prems, ...} =>
+ let val simplify = asm_simp_tac (HOL_ss addsimps (prems @ case_rewrites))
+ in EVERY [simp_tac (HOL_ss addsimps [hd prems]) 1,
+ cut_facts_tac [nchotomy''] 1,
+ REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1),
+ REPEAT (etac exE 1) THEN simplify 1 (* Get last disjunct *)]
+ end)
+ end;
+
+ val case_congs = map prove_case_cong (DatatypeProp.make_case_congs
+ new_type_names descr sorts thy ~~ nchotomys ~~ case_thms)
+
+ in thy |> store_thms "case_cong" new_type_names case_congs end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/datatype_package/datatype_aux.ML Thu Jun 11 08:04:26 2009 +0200
@@ -0,0 +1,369 @@
+(* Title: HOL/Tools/datatype_aux.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Auxiliary functions for defining datatypes.
+*)
+
+signature DATATYPE_AUX =
+sig
+ val quiet_mode : bool ref
+ val message : string -> unit
+
+ val add_path : bool -> string -> theory -> theory
+ val parent_path : bool -> theory -> theory
+
+ val store_thmss_atts : string -> string list -> attribute list list -> thm list list
+ -> theory -> thm list list * theory
+ val store_thmss : string -> string list -> thm list list -> theory -> thm list list * theory
+ val store_thms_atts : string -> string list -> attribute list list -> thm list
+ -> theory -> thm list * theory
+ val store_thms : string -> string list -> thm list -> theory -> thm list * theory
+
+ val split_conj_thm : thm -> thm list
+ val mk_conj : term list -> term
+ val mk_disj : term list -> term
+
+ val app_bnds : term -> int -> term
+
+ val cong_tac : int -> tactic
+ val indtac : thm -> string list -> int -> tactic
+ val exh_tac : (string -> thm) -> int -> tactic
+
+ datatype simproc_dist = FewConstrs of thm list
+ | ManyConstrs of thm * simpset;
+
+ datatype dtyp =
+ DtTFree of string
+ | DtType of string * (dtyp list)
+ | DtRec of int;
+ type descr
+ type datatype_info
+
+ exception Datatype
+ exception Datatype_Empty of string
+ val name_of_typ : typ -> string
+ val dtyp_of_typ : (string * string list) list -> typ -> dtyp
+ val mk_Free : string -> typ -> int -> term
+ val is_rec_type : dtyp -> bool
+ val typ_of_dtyp : descr -> (string * sort) list -> dtyp -> typ
+ val dest_DtTFree : dtyp -> string
+ val dest_DtRec : dtyp -> int
+ val strip_dtyp : dtyp -> dtyp list * dtyp
+ val body_index : dtyp -> int
+ val mk_fun_dtyp : dtyp list -> dtyp -> dtyp
+ val get_nonrec_types : descr -> (string * sort) list -> typ list
+ val get_branching_types : descr -> (string * sort) list -> typ list
+ val get_arities : descr -> int list
+ val get_rec_types : descr -> (string * sort) list -> typ list
+ val interpret_construction : descr -> (string * sort) list
+ -> { atyp: typ -> 'a, dtyp: typ list -> int * bool -> string * typ list -> 'a }
+ -> ((string * Term.typ list) * (string * 'a list) list) list
+ val check_nonempty : descr list -> unit
+ val unfold_datatypes :
+ theory -> descr -> (string * sort) list -> datatype_info Symtab.table ->
+ descr -> int -> descr list * int
+end;
+
+structure DatatypeAux : DATATYPE_AUX =
+struct
+
+val quiet_mode = ref false;
+fun message s = if !quiet_mode then () else writeln s;
+
+fun add_path flat_names s = if flat_names then I else Sign.add_path s;
+fun parent_path flat_names = if flat_names then I else Sign.parent_path;
+
+
+(* store theorems in theory *)
+
+fun store_thmss_atts label tnames attss thmss =
+ fold_map (fn ((tname, atts), thms) =>
+ Sign.add_path tname
+ #> PureThy.add_thmss [((Binding.name label, thms), atts)]
+ #-> (fn thm::_ => Sign.parent_path #> pair thm)) (tnames ~~ attss ~~ thmss)
+ ##> Theory.checkpoint;
+
+fun store_thmss label tnames = store_thmss_atts label tnames (replicate (length tnames) []);
+
+fun store_thms_atts label tnames attss thmss =
+ fold_map (fn ((tname, atts), thms) =>
+ Sign.add_path tname
+ #> PureThy.add_thms [((Binding.name label, thms), atts)]
+ #-> (fn thm::_ => Sign.parent_path #> pair thm)) (tnames ~~ attss ~~ thmss)
+ ##> Theory.checkpoint;
+
+fun store_thms label tnames = store_thms_atts label tnames (replicate (length tnames) []);
+
+
+(* split theorem thm_1 & ... & thm_n into n theorems *)
+
+fun split_conj_thm th =
+ ((th RS conjunct1)::(split_conj_thm (th RS conjunct2))) handle THM _ => [th];
+
+val mk_conj = foldr1 (HOLogic.mk_binop "op &");
+val mk_disj = foldr1 (HOLogic.mk_binop "op |");
+
+fun app_bnds t i = list_comb (t, map Bound (i - 1 downto 0));
+
+
+fun cong_tac i st = (case Logic.strip_assums_concl
+ (List.nth (prems_of st, i - 1)) of
+ _ $ (_ $ (f $ x) $ (g $ y)) =>
+ let
+ val cong' = Thm.lift_rule (Thm.cprem_of st i) cong;
+ val _ $ (_ $ (f' $ x') $ (g' $ y')) =
+ Logic.strip_assums_concl (prop_of cong');
+ val insts = map (pairself (cterm_of (Thm.theory_of_thm st)) o
+ apsnd (curry list_abs (Logic.strip_params (concl_of cong'))) o
+ apfst head_of) [(f', f), (g', g), (x', x), (y', y)]
+ in compose_tac (false, cterm_instantiate insts cong', 2) i st
+ handle THM _ => no_tac st
+ end
+ | _ => no_tac st);
+
+(* instantiate induction rule *)
+
+fun indtac indrule indnames i st =
+ let
+ val ts = HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule));
+ val ts' = HOLogic.dest_conj (HOLogic.dest_Trueprop
+ (Logic.strip_imp_concl (List.nth (prems_of st, i - 1))));
+ val getP = if can HOLogic.dest_imp (hd ts) then
+ (apfst SOME) o HOLogic.dest_imp else pair NONE;
+ val flt = if null indnames then I else
+ filter (fn Free (s, _) => s mem indnames | _ => false);
+ fun abstr (t1, t2) = (case t1 of
+ NONE => (case flt (OldTerm.term_frees t2) of
+ [Free (s, T)] => SOME (absfree (s, T, t2))
+ | _ => NONE)
+ | SOME (_ $ t') => SOME (Abs ("x", fastype_of t', abstract_over (t', t2))))
+ val cert = cterm_of (Thm.theory_of_thm st);
+ val insts = List.mapPartial (fn (t, u) => case abstr (getP u) of
+ NONE => NONE
+ | SOME u' => SOME (t |> getP |> snd |> head_of |> cert, cert u')) (ts ~~ ts');
+ val indrule' = cterm_instantiate insts indrule
+ in
+ rtac indrule' i st
+ end;
+
+(* perform exhaustive case analysis on last parameter of subgoal i *)
+
+fun exh_tac exh_thm_of i state =
+ let
+ val thy = Thm.theory_of_thm state;
+ val prem = nth (prems_of state) (i - 1);
+ val params = Logic.strip_params prem;
+ val (_, Type (tname, _)) = hd (rev params);
+ val exhaustion = Thm.lift_rule (Thm.cprem_of state i) (exh_thm_of tname);
+ val prem' = hd (prems_of exhaustion);
+ val _ $ (_ $ lhs $ _) = hd (rev (Logic.strip_assums_hyp prem'));
+ val exhaustion' = cterm_instantiate [(cterm_of thy (head_of lhs),
+ cterm_of thy (List.foldr (fn ((_, T), t) => Abs ("z", T, t))
+ (Bound 0) params))] exhaustion
+ in compose_tac (false, exhaustion', nprems_of exhaustion) i state
+ end;
+
+(* handling of distinctness theorems *)
+
+datatype simproc_dist = FewConstrs of thm list
+ | ManyConstrs of thm * simpset;
+
+(********************** Internal description of datatypes *********************)
+
+datatype dtyp =
+ DtTFree of string
+ | DtType of string * (dtyp list)
+ | DtRec of int;
+
+(* information about datatypes *)
+
+(* index, datatype name, type arguments, constructor name, types of constructor's arguments *)
+type descr = (int * (string * dtyp list * (string * dtyp list) list)) list;
+
+type datatype_info =
+ {index : int,
+ alt_names : string list option,
+ descr : descr,
+ sorts : (string * sort) list,
+ rec_names : string list,
+ rec_rewrites : thm list,
+ case_name : string,
+ case_rewrites : thm list,
+ induction : thm,
+ exhaustion : thm,
+ distinct : simproc_dist,
+ inject : thm list,
+ nchotomy : thm,
+ case_cong : thm,
+ weak_case_cong : thm};
+
+fun mk_Free s T i = Free (s ^ (string_of_int i), T);
+
+fun subst_DtTFree _ substs (T as (DtTFree name)) =
+ AList.lookup (op =) substs name |> the_default T
+ | subst_DtTFree i substs (DtType (name, ts)) =
+ DtType (name, map (subst_DtTFree i substs) ts)
+ | subst_DtTFree i _ (DtRec j) = DtRec (i + j);
+
+exception Datatype;
+exception Datatype_Empty of string;
+
+fun dest_DtTFree (DtTFree a) = a
+ | dest_DtTFree _ = raise Datatype;
+
+fun dest_DtRec (DtRec i) = i
+ | dest_DtRec _ = raise Datatype;
+
+fun is_rec_type (DtType (_, dts)) = exists is_rec_type dts
+ | is_rec_type (DtRec _) = true
+ | is_rec_type _ = false;
+
+fun strip_dtyp (DtType ("fun", [T, U])) = apfst (cons T) (strip_dtyp U)
+ | strip_dtyp T = ([], T);
+
+val body_index = dest_DtRec o snd o strip_dtyp;
+
+fun mk_fun_dtyp [] U = U
+ | mk_fun_dtyp (T :: Ts) U = DtType ("fun", [T, mk_fun_dtyp Ts U]);
+
+fun name_of_typ (Type (s, Ts)) =
+ let val s' = Long_Name.base_name s
+ in space_implode "_" (List.filter (not o equal "") (map name_of_typ Ts) @
+ [if Syntax.is_identifier s' then s' else "x"])
+ end
+ | name_of_typ _ = "";
+
+fun dtyp_of_typ _ (TFree (n, _)) = DtTFree n
+ | dtyp_of_typ _ (TVar _) = error "Illegal schematic type variable(s)"
+ | dtyp_of_typ new_dts (Type (tname, Ts)) =
+ (case AList.lookup (op =) new_dts tname of
+ NONE => DtType (tname, map (dtyp_of_typ new_dts) Ts)
+ | SOME vs => if map (try (fst o dest_TFree)) Ts = map SOME vs then
+ DtRec (find_index (curry op = tname o fst) new_dts)
+ else error ("Illegal occurrence of recursive type " ^ tname));
+
+fun typ_of_dtyp descr sorts (DtTFree a) = TFree (a, (the o AList.lookup (op =) sorts) a)
+ | typ_of_dtyp descr sorts (DtRec i) =
+ let val (s, ds, _) = (the o AList.lookup (op =) descr) i
+ in Type (s, map (typ_of_dtyp descr sorts) ds) end
+ | typ_of_dtyp descr sorts (DtType (s, ds)) =
+ Type (s, map (typ_of_dtyp descr sorts) ds);
+
+(* find all non-recursive types in datatype description *)
+
+fun get_nonrec_types descr sorts =
+ map (typ_of_dtyp descr sorts) (Library.foldl (fn (Ts, (_, (_, _, constrs))) =>
+ Library.foldl (fn (Ts', (_, cargs)) =>
+ filter_out is_rec_type cargs union Ts') (Ts, constrs)) ([], descr));
+
+(* get all recursive types in datatype description *)
+
+fun get_rec_types descr sorts = map (fn (_ , (s, ds, _)) =>
+ Type (s, map (typ_of_dtyp descr sorts) ds)) descr;
+
+(* get all branching types *)
+
+fun get_branching_types descr sorts =
+ map (typ_of_dtyp descr sorts) (fold (fn (_, (_, _, constrs)) =>
+ fold (fn (_, cargs) => fold (strip_dtyp #> fst #> fold (insert op =)) cargs)
+ constrs) descr []);
+
+fun get_arities descr = fold (fn (_, (_, _, constrs)) =>
+ fold (fn (_, cargs) => fold (insert op =) (map (length o fst o strip_dtyp)
+ (List.filter is_rec_type cargs))) constrs) descr [];
+
+(* interpret construction of datatype *)
+
+fun interpret_construction descr vs { atyp, dtyp } =
+ let
+ val typ_of_dtyp = typ_of_dtyp descr vs;
+ fun interpT dT = case strip_dtyp dT
+ of (dTs, DtRec l) =>
+ let
+ val (tyco, dTs', _) = (the o AList.lookup (op =) descr) l;
+ val Ts = map typ_of_dtyp dTs;
+ val Ts' = map typ_of_dtyp dTs';
+ val is_proper = forall (can dest_TFree) Ts';
+ in dtyp Ts (l, is_proper) (tyco, Ts') end
+ | _ => atyp (typ_of_dtyp dT);
+ fun interpC (c, dTs) = (c, map interpT dTs);
+ fun interpD (_, (tyco, dTs, cs)) = ((tyco, map typ_of_dtyp dTs), map interpC cs);
+ in map interpD descr end;
+
+(* nonemptiness check for datatypes *)
+
+fun check_nonempty descr =
+ let
+ val descr' = List.concat descr;
+ fun is_nonempty_dt is i =
+ let
+ val (_, _, constrs) = (the o AList.lookup (op =) descr') i;
+ fun arg_nonempty (_, DtRec i) = if i mem is then false
+ else is_nonempty_dt (i::is) i
+ | arg_nonempty _ = true;
+ in exists ((forall (arg_nonempty o strip_dtyp)) o snd) constrs
+ end
+ in assert_all (fn (i, _) => is_nonempty_dt [i] i) (hd descr)
+ (fn (_, (s, _, _)) => raise Datatype_Empty s)
+ end;
+
+(* unfold a list of mutually recursive datatype specifications *)
+(* all types of the form DtType (dt_name, [..., DtRec _, ...]) *)
+(* need to be unfolded *)
+
+fun unfold_datatypes sign orig_descr sorts (dt_info : datatype_info Symtab.table) descr i =
+ let
+ fun typ_error T msg = error ("Non-admissible type expression\n" ^
+ Syntax.string_of_typ_global sign (typ_of_dtyp (orig_descr @ descr) sorts T) ^ "\n" ^ msg);
+
+ fun get_dt_descr T i tname dts =
+ (case Symtab.lookup dt_info tname of
+ NONE => typ_error T (tname ^ " is not a datatype - can't use it in\
+ \ nested recursion")
+ | (SOME {index, descr, ...}) =>
+ let val (_, vars, _) = (the o AList.lookup (op =) descr) index;
+ val subst = ((map dest_DtTFree vars) ~~ dts) handle Library.UnequalLengths =>
+ typ_error T ("Type constructor " ^ tname ^ " used with wrong\
+ \ number of arguments")
+ in (i + index, map (fn (j, (tn, args, cs)) => (i + j,
+ (tn, map (subst_DtTFree i subst) args,
+ map (apsnd (map (subst_DtTFree i subst))) cs))) descr)
+ end);
+
+ (* unfold a single constructor argument *)
+
+ fun unfold_arg ((i, Ts, descrs), T) =
+ if is_rec_type T then
+ let val (Us, U) = strip_dtyp T
+ in if exists is_rec_type Us then
+ typ_error T "Non-strictly positive recursive occurrence of type"
+ else (case U of
+ DtType (tname, dts) =>
+ let
+ val (index, descr) = get_dt_descr T i tname dts;
+ val (descr', i') = unfold_datatypes sign orig_descr sorts
+ dt_info descr (i + length descr)
+ in (i', Ts @ [mk_fun_dtyp Us (DtRec index)], descrs @ descr') end
+ | _ => (i, Ts @ [T], descrs))
+ end
+ else (i, Ts @ [T], descrs);
+
+ (* unfold a constructor *)
+
+ fun unfold_constr ((i, constrs, descrs), (cname, cargs)) =
+ let val (i', cargs', descrs') = Library.foldl unfold_arg ((i, [], descrs), cargs)
+ in (i', constrs @ [(cname, cargs')], descrs') end;
+
+ (* unfold a single datatype *)
+
+ fun unfold_datatype ((i, dtypes, descrs), (j, (tname, tvars, constrs))) =
+ let val (i', constrs', descrs') =
+ Library.foldl unfold_constr ((i, [], descrs), constrs)
+ in (i', dtypes @ [(j, (tname, tvars, constrs'))], descrs')
+ end;
+
+ val (i', descr', descrs) = Library.foldl unfold_datatype ((i, [],[]), descr);
+
+ in (descr' :: descrs, i') end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/datatype_package/datatype_case.ML Thu Jun 11 08:04:26 2009 +0200
@@ -0,0 +1,469 @@
+(* Title: HOL/Tools/datatype_case.ML
+ Author: Konrad Slind, Cambridge University Computer Laboratory
+ Author: Stefan Berghofer, TU Muenchen
+
+Nested case expressions on datatypes.
+*)
+
+signature DATATYPE_CASE =
+sig
+ val make_case: (string -> DatatypeAux.datatype_info option) ->
+ Proof.context -> bool -> string list -> term -> (term * term) list ->
+ term * (term * (int * bool)) list
+ val dest_case: (string -> DatatypeAux.datatype_info option) -> bool ->
+ string list -> term -> (term * (term * term) list) option
+ val strip_case: (string -> DatatypeAux.datatype_info option) -> bool ->
+ term -> (term * (term * term) list) option
+ val case_tr: bool -> (theory -> string -> DatatypeAux.datatype_info option)
+ -> Proof.context -> term list -> term
+ val case_tr': (theory -> string -> DatatypeAux.datatype_info option) ->
+ string -> Proof.context -> term list -> term
+end;
+
+structure DatatypeCase : DATATYPE_CASE =
+struct
+
+exception CASE_ERROR of string * int;
+
+fun match_type thy pat ob = Sign.typ_match thy (pat, ob) Vartab.empty;
+
+(*---------------------------------------------------------------------------
+ * Get information about datatypes
+ *---------------------------------------------------------------------------*)
+
+fun ty_info (tab : string -> DatatypeAux.datatype_info option) s =
+ case tab s of
+ SOME {descr, case_name, index, sorts, ...} =>
+ let
+ val (_, (tname, dts, constrs)) = nth descr index;
+ val mk_ty = DatatypeAux.typ_of_dtyp descr sorts;
+ val T = Type (tname, map mk_ty dts)
+ in
+ SOME {case_name = case_name,
+ constructors = map (fn (cname, dts') =>
+ Const (cname, Logic.varifyT (map mk_ty dts' ---> T))) constrs}
+ end
+ | NONE => NONE;
+
+
+(*---------------------------------------------------------------------------
+ * Each pattern carries with it a tag (i,b) where
+ * i is the clause it came from and
+ * b=true indicates that clause was given by the user
+ * (or is an instantiation of a user supplied pattern)
+ * b=false --> i = ~1
+ *---------------------------------------------------------------------------*)
+
+fun pattern_subst theta (tm, x) = (subst_free theta tm, x);
+
+fun row_of_pat x = fst (snd x);
+
+fun add_row_used ((prfx, pats), (tm, tag)) =
+ fold Term.add_free_names (tm :: pats @ prfx);
+
+(* try to preserve names given by user *)
+fun default_names names ts =
+ map (fn ("", Free (name', _)) => name' | (name, _) => name) (names ~~ ts);
+
+fun strip_constraints (Const ("_constrain", _) $ t $ tT) =
+ strip_constraints t ||> cons tT
+ | strip_constraints t = (t, []);
+
+fun mk_fun_constrain tT t = Syntax.const "_constrain" $ t $
+ (Syntax.free "fun" $ tT $ Syntax.free "dummy");
+
+
+(*---------------------------------------------------------------------------
+ * Produce an instance of a constructor, plus genvars for its arguments.
+ *---------------------------------------------------------------------------*)
+fun fresh_constr ty_match ty_inst colty used c =
+ let
+ val (_, Ty) = dest_Const c
+ val Ts = binder_types Ty;
+ val names = Name.variant_list used
+ (DatatypeProp.make_tnames (map Logic.unvarifyT Ts));
+ val ty = body_type Ty;
+ val ty_theta = ty_match ty colty handle Type.TYPE_MATCH =>
+ raise CASE_ERROR ("type mismatch", ~1)
+ val c' = ty_inst ty_theta c
+ val gvars = map (ty_inst ty_theta o Free) (names ~~ Ts)
+ in (c', gvars)
+ end;
+
+
+(*---------------------------------------------------------------------------
+ * Goes through a list of rows and picks out the ones beginning with a
+ * pattern with constructor = name.
+ *---------------------------------------------------------------------------*)
+fun mk_group (name, T) rows =
+ let val k = length (binder_types T)
+ in fold (fn (row as ((prfx, p :: rst), rhs as (_, (i, _)))) =>
+ fn ((in_group, not_in_group), (names, cnstrts)) => (case strip_comb p of
+ (Const (name', _), args) =>
+ if name = name' then
+ if length args = k then
+ let val (args', cnstrts') = split_list (map strip_constraints args)
+ in
+ ((((prfx, args' @ rst), rhs) :: in_group, not_in_group),
+ (default_names names args', map2 append cnstrts cnstrts'))
+ end
+ else raise CASE_ERROR
+ ("Wrong number of arguments for constructor " ^ name, i)
+ else ((in_group, row :: not_in_group), (names, cnstrts))
+ | _ => raise CASE_ERROR ("Not a constructor pattern", i)))
+ rows (([], []), (replicate k "", replicate k [])) |>> pairself rev
+ end;
+
+(*---------------------------------------------------------------------------
+ * Partition the rows. Not efficient: we should use hashing.
+ *---------------------------------------------------------------------------*)
+fun partition _ _ _ _ _ _ _ [] = raise CASE_ERROR ("partition: no rows", ~1)
+ | partition ty_match ty_inst type_of used constructors colty res_ty
+ (rows as (((prfx, _ :: rstp), _) :: _)) =
+ let
+ fun part {constrs = [], rows = [], A} = rev A
+ | part {constrs = [], rows = (_, (_, (i, _))) :: _, A} =
+ raise CASE_ERROR ("Not a constructor pattern", i)
+ | part {constrs = c :: crst, rows, A} =
+ let
+ val ((in_group, not_in_group), (names, cnstrts)) =
+ mk_group (dest_Const c) rows;
+ val used' = fold add_row_used in_group used;
+ val (c', gvars) = fresh_constr ty_match ty_inst colty used' c;
+ val in_group' =
+ if null in_group (* Constructor not given *)
+ then
+ let
+ val Ts = map type_of rstp;
+ val xs = Name.variant_list
+ (fold Term.add_free_names gvars used')
+ (replicate (length rstp) "x")
+ in
+ [((prfx, gvars @ map Free (xs ~~ Ts)),
+ (Const ("HOL.undefined", res_ty), (~1, false)))]
+ end
+ else in_group
+ in
+ part{constrs = crst,
+ rows = not_in_group,
+ A = {constructor = c',
+ new_formals = gvars,
+ names = names,
+ constraints = cnstrts,
+ group = in_group'} :: A}
+ end
+ in part {constrs = constructors, rows = rows, A = []}
+ end;
+
+(*---------------------------------------------------------------------------
+ * Misc. routines used in mk_case
+ *---------------------------------------------------------------------------*)
+
+fun mk_pat ((c, c'), l) =
+ let
+ val L = length (binder_types (fastype_of c))
+ fun build (prfx, tag, plist) =
+ let val (args, plist') = chop L plist
+ in (prfx, tag, list_comb (c', args) :: plist') end
+ in map build l end;
+
+fun v_to_prfx (prfx, v::pats) = (v::prfx,pats)
+ | v_to_prfx _ = raise CASE_ERROR ("mk_case: v_to_prfx", ~1);
+
+fun v_to_pats (v::prfx,tag, pats) = (prfx, tag, v::pats)
+ | v_to_pats _ = raise CASE_ERROR ("mk_case: v_to_pats", ~1);
+
+
+(*----------------------------------------------------------------------------
+ * Translation of pattern terms into nested case expressions.
+ *
+ * This performs the translation and also builds the full set of patterns.
+ * Thus it supports the construction of induction theorems even when an
+ * incomplete set of patterns is given.
+ *---------------------------------------------------------------------------*)
+
+fun mk_case tab ctxt ty_match ty_inst type_of used range_ty =
+ let
+ val name = Name.variant used "a";
+ fun expand constructors used ty ((_, []), _) =
+ raise CASE_ERROR ("mk_case: expand_var_row", ~1)
+ | expand constructors used ty (row as ((prfx, p :: rst), rhs)) =
+ if is_Free p then
+ let
+ val used' = add_row_used row used;
+ fun expnd c =
+ let val capp =
+ list_comb (fresh_constr ty_match ty_inst ty used' c)
+ in ((prfx, capp :: rst), pattern_subst [(p, capp)] rhs)
+ end
+ in map expnd constructors end
+ else [row]
+ fun mk {rows = [], ...} = raise CASE_ERROR ("no rows", ~1)
+ | mk {path = [], rows = ((prfx, []), (tm, tag)) :: _} = (* Done *)
+ ([(prfx, tag, [])], tm)
+ | mk {path, rows as ((row as ((_, [Free _]), _)) :: _ :: _)} =
+ mk {path = path, rows = [row]}
+ | mk {path = u :: rstp, rows as ((_, _ :: _), _) :: _} =
+ let val col0 = map (fn ((_, p :: _), (_, (i, _))) => (p, i)) rows
+ in case Option.map (apfst head_of)
+ (find_first (not o is_Free o fst) col0) of
+ NONE =>
+ let
+ val rows' = map (fn ((v, _), row) => row ||>
+ pattern_subst [(v, u)] |>> v_to_prfx) (col0 ~~ rows);
+ val (pref_patl, tm) = mk {path = rstp, rows = rows'}
+ in (map v_to_pats pref_patl, tm) end
+ | SOME (Const (cname, cT), i) => (case ty_info tab cname of
+ NONE => raise CASE_ERROR ("Not a datatype constructor: " ^ cname, i)
+ | SOME {case_name, constructors} =>
+ let
+ val pty = body_type cT;
+ val used' = fold Term.add_free_names rstp used;
+ val nrows = maps (expand constructors used' pty) rows;
+ val subproblems = partition ty_match ty_inst type_of used'
+ constructors pty range_ty nrows;
+ val new_formals = map #new_formals subproblems
+ val constructors' = map #constructor subproblems
+ val news = map (fn {new_formals, group, ...} =>
+ {path = new_formals @ rstp, rows = group}) subproblems;
+ val (pat_rect, dtrees) = split_list (map mk news);
+ val case_functions = map2
+ (fn {new_formals, names, constraints, ...} =>
+ fold_rev (fn ((x as Free (_, T), s), cnstrts) => fn t =>
+ Abs (if s = "" then name else s, T,
+ abstract_over (x, t)) |>
+ fold mk_fun_constrain cnstrts)
+ (new_formals ~~ names ~~ constraints))
+ subproblems dtrees;
+ val types = map type_of (case_functions @ [u]);
+ val case_const = Const (case_name, types ---> range_ty)
+ val tree = list_comb (case_const, case_functions @ [u])
+ val pat_rect1 = flat (map mk_pat
+ (constructors ~~ constructors' ~~ pat_rect))
+ in (pat_rect1, tree)
+ end)
+ | SOME (t, i) => raise CASE_ERROR ("Not a datatype constructor: " ^
+ Syntax.string_of_term ctxt t, i)
+ end
+ | mk _ = raise CASE_ERROR ("Malformed row matrix", ~1)
+ in mk
+ end;
+
+fun case_error s = error ("Error in case expression:\n" ^ s);
+
+(* Repeated variable occurrences in a pattern are not allowed. *)
+fun no_repeat_vars ctxt pat = fold_aterms
+ (fn x as Free (s, _) => (fn xs =>
+ if member op aconv xs x then
+ case_error (quote s ^ " occurs repeatedly in the pattern " ^
+ quote (Syntax.string_of_term ctxt pat))
+ else x :: xs)
+ | _ => I) pat [];
+
+fun gen_make_case ty_match ty_inst type_of tab ctxt err used x clauses =
+ let
+ fun string_of_clause (pat, rhs) = Syntax.string_of_term ctxt
+ (Syntax.const "_case1" $ pat $ rhs);
+ val _ = map (no_repeat_vars ctxt o fst) clauses;
+ val rows = map_index (fn (i, (pat, rhs)) =>
+ (([], [pat]), (rhs, (i, true)))) clauses;
+ val rangeT = (case distinct op = (map (type_of o snd) clauses) of
+ [] => case_error "no clauses given"
+ | [T] => T
+ | _ => case_error "all cases must have the same result type");
+ val used' = fold add_row_used rows used;
+ val (patts, case_tm) = mk_case tab ctxt ty_match ty_inst type_of
+ used' rangeT {path = [x], rows = rows}
+ handle CASE_ERROR (msg, i) => case_error (msg ^
+ (if i < 0 then ""
+ else "\nIn clause\n" ^ string_of_clause (nth clauses i)));
+ val patts1 = map
+ (fn (_, tag, [pat]) => (pat, tag)
+ | _ => case_error "error in pattern-match translation") patts;
+ val patts2 = Library.sort (Library.int_ord o Library.pairself row_of_pat) patts1
+ val finals = map row_of_pat patts2
+ val originals = map (row_of_pat o #2) rows
+ val _ = case originals \\ finals of
+ [] => ()
+ | is => (if err then case_error else warning)
+ ("The following clauses are redundant (covered by preceding clauses):\n" ^
+ cat_lines (map (string_of_clause o nth clauses) is));
+ in
+ (case_tm, patts2)
+ end;
+
+fun make_case tab ctxt = gen_make_case
+ (match_type (ProofContext.theory_of ctxt)) Envir.subst_TVars fastype_of tab ctxt;
+val make_case_untyped = gen_make_case (K (K Vartab.empty))
+ (K (Term.map_types (K dummyT))) (K dummyT);
+
+
+(* parse translation *)
+
+fun case_tr err tab_of ctxt [t, u] =
+ let
+ val thy = ProofContext.theory_of ctxt;
+ (* replace occurrences of dummy_pattern by distinct variables *)
+ (* internalize constant names *)
+ fun prep_pat ((c as Const ("_constrain", _)) $ t $ tT) used =
+ let val (t', used') = prep_pat t used
+ in (c $ t' $ tT, used') end
+ | prep_pat (Const ("dummy_pattern", T)) used =
+ let val x = Name.variant used "x"
+ in (Free (x, T), x :: used) end
+ | prep_pat (Const (s, T)) used =
+ (case try (unprefix Syntax.constN) s of
+ SOME c => (Const (c, T), used)
+ | NONE => (Const (Sign.intern_const thy s, T), used))
+ | prep_pat (v as Free (s, T)) used =
+ let val s' = Sign.intern_const thy s
+ in
+ if Sign.declared_const thy s' then
+ (Const (s', T), used)
+ else (v, used)
+ end
+ | prep_pat (t $ u) used =
+ let
+ val (t', used') = prep_pat t used;
+ val (u', used'') = prep_pat u used'
+ in
+ (t' $ u', used'')
+ end
+ | prep_pat t used = case_error ("Bad pattern: " ^ Syntax.string_of_term ctxt t);
+ fun dest_case1 (t as Const ("_case1", _) $ l $ r) =
+ let val (l', cnstrts) = strip_constraints l
+ in ((fst (prep_pat l' (Term.add_free_names t [])), r), cnstrts)
+ end
+ | dest_case1 t = case_error "dest_case1";
+ fun dest_case2 (Const ("_case2", _) $ t $ u) = t :: dest_case2 u
+ | dest_case2 t = [t];
+ val (cases, cnstrts) = split_list (map dest_case1 (dest_case2 u));
+ val (case_tm, _) = make_case_untyped (tab_of thy) ctxt err []
+ (fold (fn tT => fn t => Syntax.const "_constrain" $ t $ tT)
+ (flat cnstrts) t) cases;
+ in case_tm end
+ | case_tr _ _ _ ts = case_error "case_tr";
+
+
+(*---------------------------------------------------------------------------
+ * Pretty printing of nested case expressions
+ *---------------------------------------------------------------------------*)
+
+(* destruct one level of pattern matching *)
+
+fun gen_dest_case name_of type_of tab d used t =
+ case apfst name_of (strip_comb t) of
+ (SOME cname, ts as _ :: _) =>
+ let
+ val (fs, x) = split_last ts;
+ fun strip_abs i t =
+ let
+ val zs = strip_abs_vars t;
+ val _ = if length zs < i then raise CASE_ERROR ("", 0) else ();
+ val (xs, ys) = chop i zs;
+ val u = list_abs (ys, strip_abs_body t);
+ val xs' = map Free (Name.variant_list (OldTerm.add_term_names (u, used))
+ (map fst xs) ~~ map snd xs)
+ in (xs', subst_bounds (rev xs', u)) end;
+ fun is_dependent i t =
+ let val k = length (strip_abs_vars t) - i
+ in k < 0 orelse exists (fn j => j >= k)
+ (loose_bnos (strip_abs_body t))
+ end;
+ fun count_cases (_, _, true) = I
+ | count_cases (c, (_, body), false) =
+ AList.map_default op aconv (body, []) (cons c);
+ val is_undefined = name_of #> equal (SOME "HOL.undefined");
+ fun mk_case (c, (xs, body), _) = (list_comb (c, xs), body)
+ in case ty_info tab cname of
+ SOME {constructors, case_name} =>
+ if length fs = length constructors then
+ let
+ val cases = map (fn (Const (s, U), t) =>
+ let
+ val k = length (binder_types U);
+ val p as (xs, _) = strip_abs k t
+ in
+ (Const (s, map type_of xs ---> type_of x),
+ p, is_dependent k t)
+ end) (constructors ~~ fs);
+ val cases' = sort (int_ord o swap o pairself (length o snd))
+ (fold_rev count_cases cases []);
+ val R = type_of t;
+ val dummy = if d then Const ("dummy_pattern", R)
+ else Free (Name.variant used "x", R)
+ in
+ SOME (x, map mk_case (case find_first (is_undefined o fst) cases' of
+ SOME (_, cs) =>
+ if length cs = length constructors then [hd cases]
+ else filter_out (fn (_, (_, body), _) => is_undefined body) cases
+ | NONE => case cases' of
+ [] => cases
+ | (default, cs) :: _ =>
+ if length cs = 1 then cases
+ else if length cs = length constructors then
+ [hd cases, (dummy, ([], default), false)]
+ else
+ filter_out (fn (c, _, _) => member op aconv cs c) cases @
+ [(dummy, ([], default), false)]))
+ end handle CASE_ERROR _ => NONE
+ else NONE
+ | _ => NONE
+ end
+ | _ => NONE;
+
+val dest_case = gen_dest_case (try (dest_Const #> fst)) fastype_of;
+val dest_case' = gen_dest_case
+ (try (dest_Const #> fst #> unprefix Syntax.constN)) (K dummyT);
+
+
+(* destruct nested patterns *)
+
+fun strip_case'' dest (pat, rhs) =
+ case dest (Term.add_free_names pat []) rhs of
+ SOME (exp as Free _, clauses) =>
+ if member op aconv (OldTerm.term_frees pat) exp andalso
+ not (exists (fn (_, rhs') =>
+ member op aconv (OldTerm.term_frees rhs') exp) clauses)
+ then
+ maps (strip_case'' dest) (map (fn (pat', rhs') =>
+ (subst_free [(exp, pat')] pat, rhs')) clauses)
+ else [(pat, rhs)]
+ | _ => [(pat, rhs)];
+
+fun gen_strip_case dest t = case dest [] t of
+ SOME (x, clauses) =>
+ SOME (x, maps (strip_case'' dest) clauses)
+ | NONE => NONE;
+
+val strip_case = gen_strip_case oo dest_case;
+val strip_case' = gen_strip_case oo dest_case';
+
+
+(* print translation *)
+
+fun case_tr' tab_of cname ctxt ts =
+ let
+ val thy = ProofContext.theory_of ctxt;
+ val consts = ProofContext.consts_of ctxt;
+ fun mk_clause (pat, rhs) =
+ let val xs = Term.add_frees pat []
+ in
+ Syntax.const "_case1" $
+ map_aterms
+ (fn Free p => Syntax.mark_boundT p
+ | Const (s, _) => Const (Consts.extern_early consts s, dummyT)
+ | t => t) pat $
+ map_aterms
+ (fn x as Free (s, T) =>
+ if member (op =) xs (s, T) then Syntax.mark_bound s else x
+ | t => t) rhs
+ end
+ in case strip_case' (tab_of thy) true (list_comb (Syntax.const cname, ts)) of
+ SOME (x, clauses) => Syntax.const "_case_syntax" $ x $
+ foldr1 (fn (t, u) => Syntax.const "_case2" $ t $ u)
+ (map mk_clause clauses)
+ | NONE => raise Match
+ end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/datatype_package/datatype_codegen.ML Thu Jun 11 08:04:26 2009 +0200
@@ -0,0 +1,451 @@
+(* Title: HOL/Tools/datatype_codegen.ML
+ Author: Stefan Berghofer and Florian Haftmann, TU Muenchen
+
+Code generator facilities for inductive datatypes.
+*)
+
+signature DATATYPE_CODEGEN =
+sig
+ val find_shortest_path: DatatypeAux.descr -> int -> (string * int) option
+ val mk_eq_eqns: theory -> string -> (thm * bool) list
+ val mk_case_cert: theory -> string -> thm
+ val setup: theory -> theory
+end;
+
+structure DatatypeCodegen : DATATYPE_CODEGEN =
+struct
+
+(** SML code generator **)
+
+open Codegen;
+
+(**** datatype definition ****)
+
+(* find shortest path to constructor with no recursive arguments *)
+
+fun find_nonempty (descr: DatatypeAux.descr) is i =
+ let
+ val (_, _, constrs) = the (AList.lookup (op =) descr i);
+ fun arg_nonempty (_, DatatypeAux.DtRec i) = if member (op =) is i
+ then NONE
+ else Option.map (curry op + 1 o snd) (find_nonempty descr (i::is) i)
+ | arg_nonempty _ = SOME 0;
+ fun max xs = Library.foldl
+ (fn (NONE, _) => NONE
+ | (SOME i, SOME j) => SOME (Int.max (i, j))
+ | (_, NONE) => NONE) (SOME 0, xs);
+ val xs = sort (int_ord o pairself snd)
+ (map_filter (fn (s, dts) => Option.map (pair s)
+ (max (map (arg_nonempty o DatatypeAux.strip_dtyp) dts))) constrs)
+ in case xs of [] => NONE | x :: _ => SOME x end;
+
+fun find_shortest_path descr i = find_nonempty descr [i] i;
+
+fun add_dt_defs thy defs dep module (descr: DatatypeAux.descr) sorts gr =
+ let
+ val descr' = List.filter (can (map DatatypeAux.dest_DtTFree o #2 o snd)) descr;
+ val rtnames = map (#1 o snd) (List.filter (fn (_, (_, _, cs)) =>
+ exists (exists DatatypeAux.is_rec_type o snd) cs) descr');
+
+ val (_, (tname, _, _)) :: _ = descr';
+ val node_id = tname ^ " (type)";
+ val module' = if_library (thyname_of_type thy tname) module;
+
+ fun mk_dtdef prfx [] gr = ([], gr)
+ | mk_dtdef prfx ((_, (tname, dts, cs))::xs) gr =
+ let
+ val tvs = map DatatypeAux.dest_DtTFree dts;
+ val cs' = map (apsnd (map (DatatypeAux.typ_of_dtyp descr sorts))) cs;
+ val ((_, type_id), gr') = mk_type_id module' tname gr;
+ val (ps, gr'') = gr' |>
+ fold_map (fn (cname, cargs) =>
+ fold_map (invoke_tycodegen thy defs node_id module' false)
+ cargs ##>>
+ mk_const_id module' cname) cs';
+ val (rest, gr''') = mk_dtdef "and " xs gr''
+ in
+ (Pretty.block (str prfx ::
+ (if null tvs then [] else
+ [mk_tuple (map str tvs), str " "]) @
+ [str (type_id ^ " ="), Pretty.brk 1] @
+ List.concat (separate [Pretty.brk 1, str "| "]
+ (map (fn (ps', (_, cname)) => [Pretty.block
+ (str cname ::
+ (if null ps' then [] else
+ List.concat ([str " of", Pretty.brk 1] ::
+ separate [str " *", Pretty.brk 1]
+ (map single ps'))))]) ps))) :: rest, gr''')
+ end;
+
+ fun mk_constr_term cname Ts T ps =
+ List.concat (separate [str " $", Pretty.brk 1]
+ ([str ("Const (\"" ^ cname ^ "\","), Pretty.brk 1,
+ mk_type false (Ts ---> T), str ")"] :: ps));
+
+ fun mk_term_of_def gr prfx [] = []
+ | mk_term_of_def gr prfx ((_, (tname, dts, cs)) :: xs) =
+ let
+ val cs' = map (apsnd (map (DatatypeAux.typ_of_dtyp descr sorts))) cs;
+ val dts' = map (DatatypeAux.typ_of_dtyp descr sorts) dts;
+ val T = Type (tname, dts');
+ val rest = mk_term_of_def gr "and " xs;
+ val (eqs, _) = fold_map (fn (cname, Ts) => fn prfx =>
+ let val args = map (fn i =>
+ str ("x" ^ string_of_int i)) (1 upto length Ts)
+ in (Pretty.blk (4,
+ [str prfx, mk_term_of gr module' false T, Pretty.brk 1,
+ if null Ts then str (snd (get_const_id gr cname))
+ else parens (Pretty.block
+ [str (snd (get_const_id gr cname)),
+ Pretty.brk 1, mk_tuple args]),
+ str " =", Pretty.brk 1] @
+ mk_constr_term cname Ts T
+ (map2 (fn x => fn U => [Pretty.block [mk_term_of gr module' false U,
+ Pretty.brk 1, x]]) args Ts)), " | ")
+ end) cs' prfx
+ in eqs @ rest end;
+
+ fun mk_gen_of_def gr prfx [] = []
+ | mk_gen_of_def gr prfx ((i, (tname, dts, cs)) :: xs) =
+ let
+ val tvs = map DatatypeAux.dest_DtTFree dts;
+ val Us = map (DatatypeAux.typ_of_dtyp descr sorts) dts;
+ val T = Type (tname, Us);
+ val (cs1, cs2) =
+ List.partition (exists DatatypeAux.is_rec_type o snd) cs;
+ val SOME (cname, _) = find_shortest_path descr i;
+
+ fun mk_delay p = Pretty.block
+ [str "fn () =>", Pretty.brk 1, p];
+
+ fun mk_force p = Pretty.block [p, Pretty.brk 1, str "()"];
+
+ fun mk_constr s b (cname, dts) =
+ let
+ val gs = map (fn dt => mk_app false (mk_gen gr module' false rtnames s
+ (DatatypeAux.typ_of_dtyp descr sorts dt))
+ [str (if b andalso DatatypeAux.is_rec_type dt then "0"
+ else "j")]) dts;
+ val Ts = map (DatatypeAux.typ_of_dtyp descr sorts) dts;
+ val xs = map str
+ (DatatypeProp.indexify_names (replicate (length dts) "x"));
+ val ts = map str
+ (DatatypeProp.indexify_names (replicate (length dts) "t"));
+ val (_, id) = get_const_id gr cname
+ in
+ mk_let
+ (map2 (fn p => fn q => mk_tuple [p, q]) xs ts ~~ gs)
+ (mk_tuple
+ [case xs of
+ _ :: _ :: _ => Pretty.block
+ [str id, Pretty.brk 1, mk_tuple xs]
+ | _ => mk_app false (str id) xs,
+ mk_delay (Pretty.block (mk_constr_term cname Ts T
+ (map (single o mk_force) ts)))])
+ end;
+
+ fun mk_choice [c] = mk_constr "(i-1)" false c
+ | mk_choice cs = Pretty.block [str "one_of",
+ Pretty.brk 1, Pretty.blk (1, str "[" ::
+ List.concat (separate [str ",", Pretty.fbrk]
+ (map (single o mk_delay o mk_constr "(i-1)" false) cs)) @
+ [str "]"]), Pretty.brk 1, str "()"];
+
+ val gs = maps (fn s =>
+ let val s' = strip_tname s
+ in [str (s' ^ "G"), str (s' ^ "T")] end) tvs;
+ val gen_name = "gen_" ^ snd (get_type_id gr tname)
+
+ in
+ Pretty.blk (4, separate (Pretty.brk 1)
+ (str (prfx ^ gen_name ^
+ (if null cs1 then "" else "'")) :: gs @
+ (if null cs1 then [] else [str "i"]) @
+ [str "j"]) @
+ [str " =", Pretty.brk 1] @
+ (if not (null cs1) andalso not (null cs2)
+ then [str "frequency", Pretty.brk 1,
+ Pretty.blk (1, [str "[",
+ mk_tuple [str "i", mk_delay (mk_choice cs1)],
+ str ",", Pretty.fbrk,
+ mk_tuple [str "1", mk_delay (mk_choice cs2)],
+ str "]"]), Pretty.brk 1, str "()"]
+ else if null cs2 then
+ [Pretty.block [str "(case", Pretty.brk 1,
+ str "i", Pretty.brk 1, str "of",
+ Pretty.brk 1, str "0 =>", Pretty.brk 1,
+ mk_constr "0" true (cname, valOf (AList.lookup (op =) cs cname)),
+ Pretty.brk 1, str "| _ =>", Pretty.brk 1,
+ mk_choice cs1, str ")"]]
+ else [mk_choice cs2])) ::
+ (if null cs1 then []
+ else [Pretty.blk (4, separate (Pretty.brk 1)
+ (str ("and " ^ gen_name) :: gs @ [str "i"]) @
+ [str " =", Pretty.brk 1] @
+ separate (Pretty.brk 1) (str (gen_name ^ "'") :: gs @
+ [str "i", str "i"]))]) @
+ mk_gen_of_def gr "and " xs
+ end
+
+ in
+ (module', (add_edge_acyclic (node_id, dep) gr
+ handle Graph.CYCLES _ => gr) handle Graph.UNDEF _ =>
+ let
+ val gr1 = add_edge (node_id, dep)
+ (new_node (node_id, (NONE, "", "")) gr);
+ val (dtdef, gr2) = mk_dtdef "datatype " descr' gr1 ;
+ in
+ map_node node_id (K (NONE, module',
+ string_of (Pretty.blk (0, separate Pretty.fbrk dtdef @
+ [str ";"])) ^ "\n\n" ^
+ (if "term_of" mem !mode then
+ string_of (Pretty.blk (0, separate Pretty.fbrk
+ (mk_term_of_def gr2 "fun " descr') @ [str ";"])) ^ "\n\n"
+ else "") ^
+ (if "test" mem !mode then
+ string_of (Pretty.blk (0, separate Pretty.fbrk
+ (mk_gen_of_def gr2 "fun " descr') @ [str ";"])) ^ "\n\n"
+ else ""))) gr2
+ end)
+ end;
+
+
+(**** case expressions ****)
+
+fun pretty_case thy defs dep module brack constrs (c as Const (_, T)) ts gr =
+ let val i = length constrs
+ in if length ts <= i then
+ invoke_codegen thy defs dep module brack (eta_expand c ts (i+1)) gr
+ else
+ let
+ val ts1 = Library.take (i, ts);
+ val t :: ts2 = Library.drop (i, ts);
+ val names = List.foldr OldTerm.add_term_names
+ (map (fst o fst o dest_Var) (List.foldr OldTerm.add_term_vars [] ts1)) ts1;
+ val (Ts, dT) = split_last (Library.take (i+1, fst (strip_type T)));
+
+ fun pcase [] [] [] gr = ([], gr)
+ | pcase ((cname, cargs)::cs) (t::ts) (U::Us) gr =
+ let
+ val j = length cargs;
+ val xs = Name.variant_list names (replicate j "x");
+ val Us' = Library.take (j, fst (strip_type U));
+ val frees = map Free (xs ~~ Us');
+ val (cp, gr0) = invoke_codegen thy defs dep module false
+ (list_comb (Const (cname, Us' ---> dT), frees)) gr;
+ val t' = Envir.beta_norm (list_comb (t, frees));
+ val (p, gr1) = invoke_codegen thy defs dep module false t' gr0;
+ val (ps, gr2) = pcase cs ts Us gr1;
+ in
+ ([Pretty.block [cp, str " =>", Pretty.brk 1, p]] :: ps, gr2)
+ end;
+
+ val (ps1, gr1) = pcase constrs ts1 Ts gr ;
+ val ps = List.concat (separate [Pretty.brk 1, str "| "] ps1);
+ val (p, gr2) = invoke_codegen thy defs dep module false t gr1;
+ val (ps2, gr3) = fold_map (invoke_codegen thy defs dep module true) ts2 gr2;
+ in ((if not (null ts2) andalso brack then parens else I)
+ (Pretty.block (separate (Pretty.brk 1)
+ (Pretty.block ([str "(case ", p, str " of",
+ Pretty.brk 1] @ ps @ [str ")"]) :: ps2))), gr3)
+ end
+ end;
+
+
+(**** constructors ****)
+
+fun pretty_constr thy defs dep module brack args (c as Const (s, T)) ts gr =
+ let val i = length args
+ in if i > 1 andalso length ts < i then
+ invoke_codegen thy defs dep module brack (eta_expand c ts i) gr
+ else
+ let
+ val id = mk_qual_id module (get_const_id gr s);
+ val (ps, gr') = fold_map
+ (invoke_codegen thy defs dep module (i = 1)) ts gr;
+ in (case args of
+ _ :: _ :: _ => (if brack then parens else I)
+ (Pretty.block [str id, Pretty.brk 1, mk_tuple ps])
+ | _ => (mk_app brack (str id) ps), gr')
+ end
+ end;
+
+
+(**** code generators for terms and types ****)
+
+fun datatype_codegen thy defs dep module brack t gr = (case strip_comb t of
+ (c as Const (s, T), ts) =>
+ (case DatatypePackage.datatype_of_case thy s of
+ SOME {index, descr, ...} =>
+ if is_some (get_assoc_code thy (s, T)) then NONE else
+ SOME (pretty_case thy defs dep module brack
+ (#3 (the (AList.lookup op = descr index))) c ts gr )
+ | NONE => case (DatatypePackage.datatype_of_constr thy s, strip_type T) of
+ (SOME {index, descr, ...}, (_, U as Type (tyname, _))) =>
+ if is_some (get_assoc_code thy (s, T)) then NONE else
+ let
+ val SOME (tyname', _, constrs) = AList.lookup op = descr index;
+ val SOME args = AList.lookup op = constrs s
+ in
+ if tyname <> tyname' then NONE
+ else SOME (pretty_constr thy defs
+ dep module brack args c ts (snd (invoke_tycodegen thy defs dep module false U gr)))
+ end
+ | _ => NONE)
+ | _ => NONE);
+
+fun datatype_tycodegen thy defs dep module brack (Type (s, Ts)) gr =
+ (case DatatypePackage.get_datatype thy s of
+ NONE => NONE
+ | SOME {descr, sorts, ...} =>
+ if is_some (get_assoc_type thy s) then NONE else
+ let
+ val (ps, gr') = fold_map
+ (invoke_tycodegen thy defs dep module false) Ts gr;
+ val (module', gr'') = add_dt_defs thy defs dep module descr sorts gr' ;
+ val (tyid, gr''') = mk_type_id module' s gr''
+ in SOME (Pretty.block ((if null Ts then [] else
+ [mk_tuple ps, str " "]) @
+ [str (mk_qual_id module tyid)]), gr''')
+ end)
+ | datatype_tycodegen _ _ _ _ _ _ _ = NONE;
+
+
+(** generic code generator **)
+
+(* case certificates *)
+
+fun mk_case_cert thy tyco =
+ let
+ val raw_thms =
+ (#case_rewrites o DatatypePackage.the_datatype thy) tyco;
+ val thms as hd_thm :: _ = raw_thms
+ |> Conjunction.intr_balanced
+ |> Thm.unvarify
+ |> Conjunction.elim_balanced (length raw_thms)
+ |> map Simpdata.mk_meta_eq
+ |> map Drule.zero_var_indexes
+ val params = fold_aterms (fn (Free (v, _)) => insert (op =) v
+ | _ => I) (Thm.prop_of hd_thm) [];
+ val rhs = hd_thm
+ |> Thm.prop_of
+ |> Logic.dest_equals
+ |> fst
+ |> Term.strip_comb
+ |> apsnd (fst o split_last)
+ |> list_comb;
+ val lhs = Free (Name.variant params "case", Term.fastype_of rhs);
+ val asm = (Thm.cterm_of thy o Logic.mk_equals) (lhs, rhs);
+ in
+ thms
+ |> Conjunction.intr_balanced
+ |> MetaSimplifier.rewrite_rule [(Thm.symmetric o Thm.assume) asm]
+ |> Thm.implies_intr asm
+ |> Thm.generalize ([], params) 0
+ |> AxClass.unoverload thy
+ |> Thm.varifyT
+ end;
+
+
+(* equality *)
+
+fun mk_eq_eqns thy dtco =
+ let
+ val (vs, cos) = DatatypePackage.the_datatype_spec thy dtco;
+ val { descr, index, inject = inject_thms, ... } = DatatypePackage.the_datatype thy dtco;
+ val ty = Type (dtco, map TFree vs);
+ fun mk_eq (t1, t2) = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
+ $ t1 $ t2;
+ fun true_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.true_const);
+ fun false_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.false_const);
+ val triv_injects = map_filter
+ (fn (c, []) => SOME (HOLogic.mk_Trueprop (true_eq (Const (c, ty), Const (c, ty))))
+ | _ => NONE) cos;
+ fun prep_inject (trueprop $ (equiv $ (_ $ t1 $ t2) $ rhs)) =
+ trueprop $ (equiv $ mk_eq (t1, t2) $ rhs);
+ val injects = map prep_inject (nth (DatatypeProp.make_injs [descr] vs) index);
+ fun prep_distinct (trueprop $ (not $ (_ $ t1 $ t2))) =
+ [trueprop $ false_eq (t1, t2), trueprop $ false_eq (t2, t1)];
+ val distincts = maps prep_distinct (snd (nth (DatatypeProp.make_distincts [descr] vs) index));
+ val refl = HOLogic.mk_Trueprop (true_eq (Free ("x", ty), Free ("x", ty)));
+ val simpset = Simplifier.context (ProofContext.init thy) (HOL_basic_ss
+ addsimps (map Simpdata.mk_eq (@{thm eq} :: @{thm eq_True} :: inject_thms))
+ addsimprocs [DatatypePackage.distinct_simproc]);
+ fun prove prop = Goal.prove_global thy [] [] prop (K (ALLGOALS (simp_tac simpset)))
+ |> Simpdata.mk_eq;
+ in map (rpair true o prove) (triv_injects @ injects @ distincts) @ [(prove refl, false)] end;
+
+fun add_equality vs dtcos thy =
+ let
+ fun add_def dtco lthy =
+ let
+ val ty = Type (dtco, map TFree vs);
+ fun mk_side const_name = Const (const_name, ty --> ty --> HOLogic.boolT)
+ $ Free ("x", ty) $ Free ("y", ty);
+ val def = HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (mk_side @{const_name eq_class.eq}, mk_side @{const_name "op ="}));
+ val def' = Syntax.check_term lthy def;
+ val ((_, (_, thm)), lthy') = Specification.definition
+ (NONE, (Attrib.empty_binding, def')) lthy;
+ val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy);
+ val thm' = singleton (ProofContext.export lthy' ctxt_thy) thm;
+ in (thm', lthy') end;
+ fun tac thms = Class.intro_classes_tac []
+ THEN ALLGOALS (ProofContext.fact_tac thms);
+ fun add_eq_thms dtco thy =
+ let
+ val const = AxClass.param_of_inst thy (@{const_name eq_class.eq}, dtco);
+ val thy_ref = Theory.check_thy thy;
+ fun mk_thms () = rev ((mk_eq_eqns (Theory.deref thy_ref) dtco));
+ in
+ Code.add_eqnl (const, Lazy.lazy mk_thms) thy
+ end;
+ in
+ thy
+ |> TheoryTarget.instantiation (dtcos, vs, [HOLogic.class_eq])
+ |> fold_map add_def dtcos
+ |-> (fn def_thms => Class.prove_instantiation_exit_result (map o Morphism.thm)
+ (fn _ => fn def_thms => tac def_thms) def_thms)
+ |-> (fn def_thms => fold Code.del_eqn def_thms)
+ |> fold add_eq_thms dtcos
+ end;
+
+
+(* liberal addition of code data for datatypes *)
+
+fun mk_constr_consts thy vs dtco cos =
+ let
+ val cs = map (fn (c, tys) => (c, tys ---> Type (dtco, map TFree vs))) cos;
+ val cs' = map (fn c_ty as (_, ty) => (AxClass.unoverload_const thy c_ty, ty)) cs;
+ in if is_some (try (Code.constrset_of_consts thy) cs')
+ then SOME cs
+ else NONE
+ end;
+
+fun add_all_code dtcos thy =
+ let
+ val (vs :: _, coss) = (split_list o map (DatatypePackage.the_datatype_spec thy)) dtcos;
+ val any_css = map2 (mk_constr_consts thy vs) dtcos coss;
+ val css = if exists is_none any_css then []
+ else map_filter I any_css;
+ val case_rewrites = maps (#case_rewrites o DatatypePackage.the_datatype thy) dtcos;
+ val certs = map (mk_case_cert thy) dtcos;
+ in
+ if null css then thy
+ else thy
+ |> fold Code.add_datatype css
+ |> fold_rev Code.add_default_eqn case_rewrites
+ |> fold Code.add_case certs
+ |> add_equality vs dtcos
+ end;
+
+
+
+(** theory setup **)
+
+val setup =
+ add_codegen "datatype" datatype_codegen
+ #> add_tycodegen "datatype" datatype_tycodegen
+ #> DatatypePackage.interpretation add_all_code
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/datatype_package/datatype_package.ML Thu Jun 11 08:04:26 2009 +0200
@@ -0,0 +1,712 @@
+(* Title: HOL/Tools/datatype_package.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Datatype package for Isabelle/HOL.
+*)
+
+signature DATATYPE_PACKAGE =
+sig
+ type datatype_info = DatatypeAux.datatype_info
+ type descr = DatatypeAux.descr
+ val get_datatypes : theory -> datatype_info Symtab.table
+ val get_datatype : theory -> string -> datatype_info option
+ val the_datatype : theory -> string -> datatype_info
+ val datatype_of_constr : theory -> string -> datatype_info option
+ val datatype_of_case : theory -> string -> datatype_info option
+ val the_datatype_spec : theory -> string -> (string * sort) list * (string * typ list) list
+ val the_datatype_descr : theory -> string list
+ -> descr * (string * sort) list * string list
+ * (string list * string list) * (typ list * typ list)
+ val get_datatype_constrs : theory -> string -> (string * typ) list option
+ val distinct_simproc : simproc
+ val make_case : Proof.context -> bool -> string list -> term ->
+ (term * term) list -> term * (term * (int * bool)) list
+ val strip_case : Proof.context -> bool -> term -> (term * (term * term) list) option
+ val read_typ: theory ->
+ (typ list * (string * sort) list) * string -> typ list * (string * sort) list
+ val interpretation : (string list -> theory -> theory) -> theory -> theory
+ val rep_datatype : ({distinct : thm list list,
+ inject : thm list list,
+ exhaustion : thm list,
+ rec_thms : thm list,
+ case_thms : thm list list,
+ split_thms : (thm * thm) list,
+ induction : thm,
+ simps : thm list} -> Proof.context -> Proof.context) -> string list option -> term list
+ -> theory -> Proof.state;
+ val rep_datatype_cmd : string list option -> string list -> theory -> Proof.state;
+ val add_datatype : bool -> bool -> string list -> (string list * binding * mixfix *
+ (binding * typ list * mixfix) list) list -> theory ->
+ {distinct : thm list list,
+ inject : thm list list,
+ exhaustion : thm list,
+ rec_thms : thm list,
+ case_thms : thm list list,
+ split_thms : (thm * thm) list,
+ induction : thm,
+ simps : thm list} * theory
+ val add_datatype_cmd : bool -> string list -> (string list * binding * mixfix *
+ (binding * string list * mixfix) list) list -> theory ->
+ {distinct : thm list list,
+ inject : thm list list,
+ exhaustion : thm list,
+ rec_thms : thm list,
+ case_thms : thm list list,
+ split_thms : (thm * thm) list,
+ induction : thm,
+ simps : thm list} * theory
+ val setup: theory -> theory
+ val quiet_mode : bool ref
+ val print_datatypes : theory -> unit
+end;
+
+structure DatatypePackage : DATATYPE_PACKAGE =
+struct
+
+open DatatypeAux;
+
+val quiet_mode = quiet_mode;
+
+
+(* theory data *)
+
+structure DatatypesData = TheoryDataFun
+(
+ type T =
+ {types: datatype_info Symtab.table,
+ constrs: datatype_info Symtab.table,
+ cases: datatype_info Symtab.table};
+
+ val empty =
+ {types = Symtab.empty, constrs = Symtab.empty, cases = Symtab.empty};
+ val copy = I;
+ val extend = I;
+ fun merge _
+ ({types = types1, constrs = constrs1, cases = cases1},
+ {types = types2, constrs = constrs2, cases = cases2}) =
+ {types = Symtab.merge (K true) (types1, types2),
+ constrs = Symtab.merge (K true) (constrs1, constrs2),
+ cases = Symtab.merge (K true) (cases1, cases2)};
+);
+
+val get_datatypes = #types o DatatypesData.get;
+val map_datatypes = DatatypesData.map;
+
+fun print_datatypes thy =
+ Pretty.writeln (Pretty.strs ("datatypes:" ::
+ map #1 (NameSpace.extern_table (Sign.type_space thy, get_datatypes thy))));
+
+
+(** theory information about datatypes **)
+
+fun put_dt_infos (dt_infos : (string * datatype_info) list) =
+ map_datatypes (fn {types, constrs, cases} =>
+ {types = fold Symtab.update dt_infos types,
+ constrs = fold Symtab.default (*conservative wrt. overloaded constructors*)
+ (maps (fn (_, info as {descr, index, ...}) => map (rpair info o fst)
+ (#3 (the (AList.lookup op = descr index)))) dt_infos) constrs,
+ cases = fold Symtab.update
+ (map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)
+ cases});
+
+val get_datatype = Symtab.lookup o get_datatypes;
+
+fun the_datatype thy name = (case get_datatype thy name of
+ SOME info => info
+ | NONE => error ("Unknown datatype " ^ quote name));
+
+val datatype_of_constr = Symtab.lookup o #constrs o DatatypesData.get;
+val datatype_of_case = Symtab.lookup o #cases o DatatypesData.get;
+
+fun get_datatype_descr thy dtco =
+ get_datatype thy dtco
+ |> Option.map (fn info as { descr, index, ... } =>
+ (info, (((fn SOME (_, dtys, cos) => (dtys, cos)) o AList.lookup (op =) descr) index)));
+
+fun the_datatype_spec thy dtco =
+ let
+ val info as { descr, index, sorts = raw_sorts, ... } = the_datatype thy dtco;
+ val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr index;
+ val sorts = map ((fn v => (v, (the o AList.lookup (op =) raw_sorts) v))
+ o DatatypeAux.dest_DtTFree) dtys;
+ val cos = map
+ (fn (co, tys) => (co, map (DatatypeAux.typ_of_dtyp descr sorts) tys)) raw_cos;
+ in (sorts, cos) end;
+
+fun the_datatype_descr thy (raw_tycos as raw_tyco :: _) =
+ let
+ val info = the_datatype thy raw_tyco;
+ val descr = #descr info;
+
+ val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr (#index info);
+ val vs = map ((fn v => (v, (the o AList.lookup (op =) (#sorts info)) v))
+ o dest_DtTFree) dtys;
+
+ fun is_DtTFree (DtTFree _) = true
+ | is_DtTFree _ = false
+ val k = find_index (fn (_, (_, dTs, _)) => not (forall is_DtTFree dTs)) descr;
+ val protoTs as (dataTs, _) = chop k descr
+ |> (pairself o map) (fn (_, (tyco, dTs, _)) => (tyco, map (typ_of_dtyp descr vs) dTs));
+
+ val tycos = map fst dataTs;
+ val _ = if gen_eq_set (op =) (tycos, raw_tycos) then ()
+ else error ("Type constructors " ^ commas (map quote raw_tycos)
+ ^ "do not belong exhaustively to one mutual recursive datatype");
+
+ val (Ts, Us) = (pairself o map) Type protoTs;
+
+ val names = map Long_Name.base_name (the_default tycos (#alt_names info));
+ val (auxnames, _) = Name.make_context names
+ |> fold_map (yield_singleton Name.variants o name_of_typ) Us
+
+ in (descr, vs, tycos, (names, auxnames), (Ts, Us)) end;
+
+fun get_datatype_constrs thy dtco =
+ case try (the_datatype_spec thy) dtco
+ of SOME (sorts, cos) =>
+ let
+ fun subst (v, sort) = TVar ((v, 0), sort);
+ fun subst_ty (TFree v) = subst v
+ | subst_ty ty = ty;
+ val dty = Type (dtco, map subst sorts);
+ fun mk_co (co, tys) = (co, map (Term.map_atyps subst_ty) tys ---> dty);
+ in SOME (map mk_co cos) end
+ | NONE => NONE;
+
+
+(** induct method setup **)
+
+(* case names *)
+
+local
+
+fun dt_recs (DtTFree _) = []
+ | dt_recs (DtType (_, dts)) = maps dt_recs dts
+ | dt_recs (DtRec i) = [i];
+
+fun dt_cases (descr: descr) (_, args, constrs) =
+ let
+ fun the_bname i = Long_Name.base_name (#1 (the (AList.lookup (op =) descr i)));
+ val bnames = map the_bname (distinct (op =) (maps dt_recs args));
+ in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end;
+
+
+fun induct_cases descr =
+ DatatypeProp.indexify_names (maps (dt_cases descr) (map #2 descr));
+
+fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i));
+
+in
+
+fun mk_case_names_induct descr = RuleCases.case_names (induct_cases descr);
+
+fun mk_case_names_exhausts descr new =
+ map (RuleCases.case_names o exhaust_cases descr o #1)
+ (filter (fn ((_, (name, _, _))) => member (op =) new name) descr);
+
+end;
+
+fun add_rules simps case_thms rec_thms inject distinct
+ weak_case_congs cong_att =
+ PureThy.add_thmss [((Binding.name "simps", simps), []),
+ ((Binding.empty, flat case_thms @
+ flat distinct @ rec_thms), [Simplifier.simp_add]),
+ ((Binding.empty, rec_thms), [Code.add_default_eqn_attribute]),
+ ((Binding.empty, flat inject), [iff_add]),
+ ((Binding.empty, map (fn th => th RS notE) (flat distinct)), [Classical.safe_elim NONE]),
+ ((Binding.empty, weak_case_congs), [cong_att])]
+ #> snd;
+
+
+(* add_cases_induct *)
+
+fun add_cases_induct infos induction thy =
+ let
+ val inducts = ProjectRule.projections (ProofContext.init thy) induction;
+
+ fun named_rules (name, {index, exhaustion, ...}: datatype_info) =
+ [((Binding.empty, nth inducts index), [Induct.induct_type name]),
+ ((Binding.empty, exhaustion), [Induct.cases_type name])];
+ fun unnamed_rule i =
+ ((Binding.empty, nth inducts i), [Thm.kind_internal, Induct.induct_type ""]);
+ in
+ thy |> PureThy.add_thms
+ (maps named_rules infos @
+ map unnamed_rule (length infos upto length inducts - 1)) |> snd
+ |> PureThy.add_thmss [((Binding.name "inducts", inducts), [])] |> snd
+ end;
+
+
+
+(**** simplification procedure for showing distinctness of constructors ****)
+
+fun stripT (i, Type ("fun", [_, T])) = stripT (i + 1, T)
+ | stripT p = p;
+
+fun stripC (i, f $ x) = stripC (i + 1, f)
+ | stripC p = p;
+
+val distinctN = "constr_distinct";
+
+fun distinct_rule thy ss tname eq_t = case #distinct (the_datatype thy tname) of
+ FewConstrs thms => Goal.prove (Simplifier.the_context ss) [] [] eq_t (K
+ (EVERY [rtac eq_reflection 1, rtac iffI 1, rtac notE 1,
+ atac 2, resolve_tac thms 1, etac FalseE 1]))
+ | ManyConstrs (thm, simpset) =>
+ let
+ val [In0_inject, In1_inject, In0_not_In1, In1_not_In0] =
+ map (PureThy.get_thm (ThyInfo.the_theory "Datatype" thy))
+ ["In0_inject", "In1_inject", "In0_not_In1", "In1_not_In0"];
+ in
+ Goal.prove (Simplifier.the_context ss) [] [] eq_t (K
+ (EVERY [rtac eq_reflection 1, rtac iffI 1, dtac thm 1,
+ full_simp_tac (Simplifier.inherit_context ss simpset) 1,
+ REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
+ eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1,
+ etac FalseE 1]))
+ end;
+
+fun distinct_proc thy ss (t as Const ("op =", _) $ t1 $ t2) =
+ (case (stripC (0, t1), stripC (0, t2)) of
+ ((i, Const (cname1, T1)), (j, Const (cname2, T2))) =>
+ (case (stripT (0, T1), stripT (0, T2)) of
+ ((i', Type (tname1, _)), (j', Type (tname2, _))) =>
+ if tname1 = tname2 andalso not (cname1 = cname2) andalso i = i' andalso j = j' then
+ (case (get_datatype_descr thy) tname1 of
+ SOME (_, (_, constrs)) => let val cnames = map fst constrs
+ in if cname1 mem cnames andalso cname2 mem cnames then
+ SOME (distinct_rule thy ss tname1
+ (Logic.mk_equals (t, Const ("False", HOLogic.boolT))))
+ else NONE
+ end
+ | NONE => NONE)
+ else NONE
+ | _ => NONE)
+ | _ => NONE)
+ | distinct_proc _ _ _ = NONE;
+
+val distinct_simproc =
+ Simplifier.simproc @{theory HOL} distinctN ["s = t"] distinct_proc;
+
+val dist_ss = HOL_ss addsimprocs [distinct_simproc];
+
+val simproc_setup =
+ Simplifier.map_simpset (fn ss => ss addsimprocs [distinct_simproc]);
+
+
+(**** translation rules for case ****)
+
+fun make_case ctxt = DatatypeCase.make_case
+ (datatype_of_constr (ProofContext.theory_of ctxt)) ctxt;
+
+fun strip_case ctxt = DatatypeCase.strip_case
+ (datatype_of_case (ProofContext.theory_of ctxt));
+
+fun add_case_tr' case_names thy =
+ Sign.add_advanced_trfuns ([], [],
+ map (fn case_name =>
+ let val case_name' = Sign.const_syntax_name thy case_name
+ in (case_name', DatatypeCase.case_tr' datatype_of_case case_name')
+ end) case_names, []) thy;
+
+val trfun_setup =
+ Sign.add_advanced_trfuns ([],
+ [("_case_syntax", DatatypeCase.case_tr true datatype_of_constr)],
+ [], []);
+
+
+(* prepare types *)
+
+fun read_typ thy ((Ts, sorts), str) =
+ let
+ val ctxt = ProofContext.init thy
+ |> fold (Variable.declare_typ o TFree) sorts;
+ val T = Syntax.read_typ ctxt str;
+ in (Ts @ [T], Term.add_tfreesT T sorts) end;
+
+fun cert_typ sign ((Ts, sorts), raw_T) =
+ let
+ val T = Type.no_tvars (Sign.certify_typ sign raw_T) handle
+ TYPE (msg, _, _) => error msg;
+ val sorts' = Term.add_tfreesT T sorts;
+ in (Ts @ [T],
+ case duplicates (op =) (map fst sorts') of
+ [] => sorts'
+ | dups => error ("Inconsistent sort constraints for " ^ commas dups))
+ end;
+
+
+(**** make datatype info ****)
+
+fun make_dt_info alt_names descr sorts induct reccomb_names rec_thms
+ (((((((((i, (_, (tname, _, _))), case_name), case_thms),
+ exhaustion_thm), distinct_thm), inject), nchotomy), case_cong), weak_case_cong) =
+ (tname,
+ {index = i,
+ alt_names = alt_names,
+ descr = descr,
+ sorts = sorts,
+ rec_names = reccomb_names,
+ rec_rewrites = rec_thms,
+ case_name = case_name,
+ case_rewrites = case_thms,
+ induction = induct,
+ exhaustion = exhaustion_thm,
+ distinct = distinct_thm,
+ inject = inject,
+ nchotomy = nchotomy,
+ case_cong = case_cong,
+ weak_case_cong = weak_case_cong});
+
+structure DatatypeInterpretation = InterpretationFun(type T = string list val eq = op =);
+val interpretation = DatatypeInterpretation.interpretation;
+
+
+(******************* definitional introduction of datatypes *******************)
+
+fun add_datatype_def flat_names new_type_names descr sorts types_syntax constr_syntax dt_info
+ case_names_induct case_names_exhausts thy =
+ let
+ val _ = message ("Proofs for datatype(s) " ^ commas_quote new_type_names);
+
+ val ((inject, distinct, dist_rewrites, simproc_dists, induct), thy2) = thy |>
+ DatatypeRepProofs.representation_proofs flat_names dt_info new_type_names descr sorts
+ types_syntax constr_syntax case_names_induct;
+
+ val (casedist_thms, thy3) = DatatypeAbsProofs.prove_casedist_thms new_type_names descr
+ sorts induct case_names_exhausts thy2;
+ val ((reccomb_names, rec_thms), thy4) = DatatypeAbsProofs.prove_primrec_thms
+ flat_names new_type_names descr sorts dt_info inject dist_rewrites
+ (Simplifier.theory_context thy3 dist_ss) induct thy3;
+ val ((case_thms, case_names), thy6) = DatatypeAbsProofs.prove_case_thms
+ flat_names new_type_names descr sorts reccomb_names rec_thms thy4;
+ val (split_thms, thy7) = DatatypeAbsProofs.prove_split_thms new_type_names
+ descr sorts inject dist_rewrites casedist_thms case_thms thy6;
+ val (nchotomys, thy8) = DatatypeAbsProofs.prove_nchotomys new_type_names
+ descr sorts casedist_thms thy7;
+ val (case_congs, thy9) = DatatypeAbsProofs.prove_case_congs new_type_names
+ descr sorts nchotomys case_thms thy8;
+ val (weak_case_congs, thy10) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
+ descr sorts thy9;
+
+ val dt_infos = map
+ (make_dt_info (SOME new_type_names) (flat descr) sorts induct reccomb_names rec_thms)
+ ((0 upto length (hd descr) - 1) ~~ (hd descr) ~~ case_names ~~ case_thms ~~
+ casedist_thms ~~ simproc_dists ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
+
+ val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
+
+ val thy12 =
+ thy10
+ |> add_case_tr' case_names
+ |> Sign.add_path (space_implode "_" new_type_names)
+ |> add_rules simps case_thms rec_thms inject distinct
+ weak_case_congs (Simplifier.attrib (op addcongs))
+ |> put_dt_infos dt_infos
+ |> add_cases_induct dt_infos induct
+ |> Sign.parent_path
+ |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms) |> snd
+ |> DatatypeInterpretation.data (map fst dt_infos);
+ in
+ ({distinct = distinct,
+ inject = inject,
+ exhaustion = casedist_thms,
+ rec_thms = rec_thms,
+ case_thms = case_thms,
+ split_thms = split_thms,
+ induction = induct,
+ simps = simps}, thy12)
+ end;
+
+
+(*********************** declare existing type as datatype *********************)
+
+fun prove_rep_datatype alt_names new_type_names descr sorts induct inject half_distinct thy =
+ let
+ val ((_, [induct']), _) =
+ Variable.importT_thms [induct] (Variable.thm_context induct);
+
+ fun err t = error ("Ill-formed predicate in induction rule: " ^
+ Syntax.string_of_term_global thy t);
+
+ fun get_typ (t as _ $ Var (_, Type (tname, Ts))) =
+ ((tname, map (fst o dest_TFree) Ts) handle TERM _ => err t)
+ | get_typ t = err t;
+ val dtnames = map get_typ (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct')));
+
+ val dt_info = get_datatypes thy;
+
+ val distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct;
+ val (case_names_induct, case_names_exhausts) =
+ (mk_case_names_induct descr, mk_case_names_exhausts descr (map #1 dtnames));
+
+ val _ = message ("Proofs for datatype(s) " ^ commas_quote new_type_names);
+
+ val (casedist_thms, thy2) = thy |>
+ DatatypeAbsProofs.prove_casedist_thms new_type_names [descr] sorts induct
+ case_names_exhausts;
+ val ((reccomb_names, rec_thms), thy3) = DatatypeAbsProofs.prove_primrec_thms
+ false new_type_names [descr] sorts dt_info inject distinct
+ (Simplifier.theory_context thy2 dist_ss) induct thy2;
+ val ((case_thms, case_names), thy4) = DatatypeAbsProofs.prove_case_thms false
+ new_type_names [descr] sorts reccomb_names rec_thms thy3;
+ val (split_thms, thy5) = DatatypeAbsProofs.prove_split_thms
+ new_type_names [descr] sorts inject distinct casedist_thms case_thms thy4;
+ val (nchotomys, thy6) = DatatypeAbsProofs.prove_nchotomys new_type_names
+ [descr] sorts casedist_thms thy5;
+ val (case_congs, thy7) = DatatypeAbsProofs.prove_case_congs new_type_names
+ [descr] sorts nchotomys case_thms thy6;
+ val (weak_case_congs, thy8) = DatatypeAbsProofs.prove_weak_case_congs new_type_names
+ [descr] sorts thy7;
+
+ val ((_, [induct']), thy10) =
+ thy8
+ |> store_thmss "inject" new_type_names inject
+ ||>> store_thmss "distinct" new_type_names distinct
+ ||> Sign.add_path (space_implode "_" new_type_names)
+ ||>> PureThy.add_thms [((Binding.name "induct", induct), [case_names_induct])];
+
+ val dt_infos = map (make_dt_info alt_names descr sorts induct' reccomb_names rec_thms)
+ ((0 upto length descr - 1) ~~ descr ~~ case_names ~~ case_thms ~~ casedist_thms ~~
+ map FewConstrs distinct ~~ inject ~~ nchotomys ~~ case_congs ~~ weak_case_congs);
+
+ val simps = flat (distinct @ inject @ case_thms) @ rec_thms;
+
+ val thy11 =
+ thy10
+ |> add_case_tr' case_names
+ |> add_rules simps case_thms rec_thms inject distinct
+ weak_case_congs (Simplifier.attrib (op addcongs))
+ |> put_dt_infos dt_infos
+ |> add_cases_induct dt_infos induct'
+ |> Sign.parent_path
+ |> store_thmss "splits" new_type_names (map (fn (x, y) => [x, y]) split_thms)
+ |> snd
+ |> DatatypeInterpretation.data (map fst dt_infos);
+ in
+ ({distinct = distinct,
+ inject = inject,
+ exhaustion = casedist_thms,
+ rec_thms = rec_thms,
+ case_thms = case_thms,
+ split_thms = split_thms,
+ induction = induct',
+ simps = simps}, thy11)
+ end;
+
+fun gen_rep_datatype prep_term after_qed alt_names raw_ts thy =
+ let
+ fun constr_of_term (Const (c, T)) = (c, T)
+ | constr_of_term t =
+ error ("Not a constant: " ^ Syntax.string_of_term_global thy t);
+ fun no_constr (c, T) = error ("Bad constructor: "
+ ^ Sign.extern_const thy c ^ "::"
+ ^ Syntax.string_of_typ_global thy T);
+ fun type_of_constr (cT as (_, T)) =
+ let
+ val frees = OldTerm.typ_tfrees T;
+ val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) T
+ handle TYPE _ => no_constr cT
+ val _ = if has_duplicates (eq_fst (op =)) vs then no_constr cT else ();
+ val _ = if length frees <> length vs then no_constr cT else ();
+ in (tyco, (vs, cT)) end;
+
+ val raw_cs = AList.group (op =) (map (type_of_constr o constr_of_term o prep_term thy) raw_ts);
+ val _ = case map_filter (fn (tyco, _) =>
+ if Symtab.defined (get_datatypes thy) tyco then SOME tyco else NONE) raw_cs
+ of [] => ()
+ | tycos => error ("Type(s) " ^ commas (map quote tycos)
+ ^ " already represented inductivly");
+ val raw_vss = maps (map (map snd o fst) o snd) raw_cs;
+ val ms = case distinct (op =) (map length raw_vss)
+ of [n] => 0 upto n - 1
+ | _ => error ("Different types in given constructors");
+ fun inter_sort m = map (fn xs => nth xs m) raw_vss
+ |> Library.foldr1 (Sorts.inter_sort (Sign.classes_of thy))
+ val sorts = map inter_sort ms;
+ val vs = Name.names Name.context Name.aT sorts;
+
+ fun norm_constr (raw_vs, (c, T)) = (c, map_atyps
+ (TFree o (the o AList.lookup (op =) (map fst raw_vs ~~ vs)) o fst o dest_TFree) T);
+
+ val cs = map (apsnd (map norm_constr)) raw_cs;
+ val dtyps_of_typ = map (dtyp_of_typ (map (rpair (map fst vs) o fst) cs))
+ o fst o strip_type;
+ val new_type_names = map Long_Name.base_name (the_default (map fst cs) alt_names);
+
+ fun mk_spec (i, (tyco, constr)) = (i, (tyco,
+ map (DtTFree o fst) vs,
+ (map o apsnd) dtyps_of_typ constr))
+ val descr = map_index mk_spec cs;
+ val injs = DatatypeProp.make_injs [descr] vs;
+ val half_distincts = map snd (DatatypeProp.make_distincts [descr] vs);
+ val ind = DatatypeProp.make_ind [descr] vs;
+ val rules = (map o map o map) Logic.close_form [[[ind]], injs, half_distincts];
+
+ fun after_qed' raw_thms =
+ let
+ val [[[induct]], injs, half_distincts] =
+ unflat rules (map Drule.zero_var_indexes_list raw_thms);
+ (*FIXME somehow dubious*)
+ in
+ ProofContext.theory_result
+ (prove_rep_datatype alt_names new_type_names descr vs induct injs half_distincts)
+ #-> after_qed
+ end;
+ in
+ thy
+ |> ProofContext.init
+ |> Proof.theorem_i NONE after_qed' ((map o map) (rpair []) (flat rules))
+ end;
+
+val rep_datatype = gen_rep_datatype Sign.cert_term;
+val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global (K I);
+
+
+
+(******************************** add datatype ********************************)
+
+fun gen_add_datatype prep_typ err flat_names new_type_names dts thy =
+ let
+ val _ = Theory.requires thy "Datatype" "datatype definitions";
+
+ (* this theory is used just for parsing *)
+
+ val tmp_thy = thy |>
+ Theory.copy |>
+ Sign.add_types (map (fn (tvs, tname, mx, _) =>
+ (tname, length tvs, mx)) dts);
+
+ val (tyvars, _, _, _)::_ = dts;
+ val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) =>
+ let val full_tname = Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname)
+ in (case duplicates (op =) tvs of
+ [] => if eq_set (tyvars, tvs) then ((full_tname, tvs), (tname, mx))
+ else error ("Mutually recursive datatypes must have same type parameters")
+ | dups => error ("Duplicate parameter(s) for datatype " ^ quote (Binding.str_of tname) ^
+ " : " ^ commas dups))
+ end) dts);
+
+ val _ = (case duplicates (op =) (map fst new_dts) @ duplicates (op =) new_type_names of
+ [] => () | dups => error ("Duplicate datatypes: " ^ commas dups));
+
+ fun prep_dt_spec ((tvs, tname, mx, constrs), tname') (dts', constr_syntax, sorts, i) =
+ let
+ fun prep_constr (cname, cargs, mx') (constrs, constr_syntax', sorts') =
+ let
+ val (cargs', sorts'') = Library.foldl (prep_typ tmp_thy) (([], sorts'), cargs);
+ val _ = (case fold (curry OldTerm.add_typ_tfree_names) cargs' [] \\ tvs of
+ [] => ()
+ | vs => error ("Extra type variables on rhs: " ^ commas vs))
+ in (constrs @ [((if flat_names then Sign.full_name tmp_thy else
+ Sign.full_name_path tmp_thy tname')
+ (Binding.map_name (Syntax.const_name mx') cname),
+ map (dtyp_of_typ new_dts) cargs')],
+ constr_syntax' @ [(cname, mx')], sorts'')
+ end handle ERROR msg => cat_error msg
+ ("The error above occured in constructor " ^ quote (Binding.str_of cname) ^
+ " of datatype " ^ quote (Binding.str_of tname));
+
+ val (constrs', constr_syntax', sorts') =
+ fold prep_constr constrs ([], [], sorts)
+
+ in
+ case duplicates (op =) (map fst constrs') of
+ [] =>
+ (dts' @ [(i, (Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname),
+ map DtTFree tvs, constrs'))],
+ constr_syntax @ [constr_syntax'], sorts', i + 1)
+ | dups => error ("Duplicate constructors " ^ commas dups ^
+ " in datatype " ^ quote (Binding.str_of tname))
+ end;
+
+ val (dts', constr_syntax, sorts', i) =
+ fold prep_dt_spec (dts ~~ new_type_names) ([], [], [], 0);
+ val sorts = sorts' @ (map (rpair (Sign.defaultS tmp_thy)) (tyvars \\ map fst sorts'));
+ val dt_info = get_datatypes thy;
+ val (descr, _) = unfold_datatypes tmp_thy dts' sorts dt_info dts' i;
+ val _ = check_nonempty descr handle (exn as Datatype_Empty s) =>
+ if err then error ("Nonemptiness check failed for datatype " ^ s)
+ else raise exn;
+
+ val descr' = flat descr;
+ val case_names_induct = mk_case_names_induct descr';
+ val case_names_exhausts = mk_case_names_exhausts descr' (map #1 new_dts);
+ in
+ add_datatype_def
+ flat_names new_type_names descr sorts types_syntax constr_syntax dt_info
+ case_names_induct case_names_exhausts thy
+ end;
+
+val add_datatype = gen_add_datatype cert_typ;
+val add_datatype_cmd = gen_add_datatype read_typ true;
+
+
+
+(** package setup **)
+
+(* setup theory *)
+
+val setup =
+ DatatypeRepProofs.distinctness_limit_setup #>
+ simproc_setup #>
+ trfun_setup #>
+ DatatypeInterpretation.init;
+
+
+(* outer syntax *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val datatype_decl =
+ Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.binding -- P.opt_infix --
+ (P.$$$ "=" |-- P.enum1 "|" (P.binding -- Scan.repeat P.typ -- P.opt_mixfix));
+
+fun mk_datatype args =
+ let
+ val names = map
+ (fn ((((NONE, _), t), _), _) => Binding.name_of t | ((((SOME t, _), _), _), _) => t) args;
+ val specs = map (fn ((((_, vs), t), mx), cons) =>
+ (vs, t, mx, map (fn ((x, y), z) => (x, y, z)) cons)) args;
+ in snd o add_datatype_cmd false names specs end;
+
+val _ =
+ OuterSyntax.command "datatype" "define inductive datatypes" K.thy_decl
+ (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype));
+
+val _ =
+ OuterSyntax.command "rep_datatype" "represent existing types inductively" K.thy_goal
+ (Scan.option (P.$$$ "(" |-- Scan.repeat1 P.name --| P.$$$ ")") -- Scan.repeat1 P.term
+ >> (fn (alt_names, ts) => Toplevel.print
+ o Toplevel.theory_to_proof (rep_datatype_cmd alt_names ts)));
+
+end;
+
+
+(* document antiquotation *)
+
+val _ = ThyOutput.antiquotation "datatype" Args.tyname
+ (fn {source = src, context = ctxt, ...} => fn dtco =>
+ let
+ val thy = ProofContext.theory_of ctxt;
+ val (vs, cos) = the_datatype_spec thy dtco;
+ val ty = Type (dtco, map TFree vs);
+ fun pretty_typ_bracket (ty as Type (_, _ :: _)) =
+ Pretty.enclose "(" ")" [Syntax.pretty_typ ctxt ty]
+ | pretty_typ_bracket ty =
+ Syntax.pretty_typ ctxt ty;
+ fun pretty_constr (co, tys) =
+ (Pretty.block o Pretty.breaks)
+ (Syntax.pretty_term ctxt (Const (co, tys ---> ty)) ::
+ map pretty_typ_bracket tys);
+ val pretty_datatype =
+ Pretty.block
+ (Pretty.command "datatype" :: Pretty.brk 1 ::
+ Syntax.pretty_typ ctxt ty ::
+ Pretty.str " =" :: Pretty.brk 1 ::
+ flat (separate [Pretty.brk 1, Pretty.str "| "]
+ (map (single o pretty_constr) cos)));
+ in ThyOutput.output (ThyOutput.maybe_pretty_source (K pretty_datatype) src [()]) end);
+
+end;
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/datatype_package/datatype_prop.ML Thu Jun 11 08:04:26 2009 +0200
@@ -0,0 +1,435 @@
+(* Title: HOL/Tools/datatype_prop.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Characteristic properties of datatypes.
+*)
+
+signature DATATYPE_PROP =
+sig
+ val indexify_names: string list -> string list
+ val make_tnames: typ list -> string list
+ val make_injs : DatatypeAux.descr list -> (string * sort) list -> term list list
+ val make_distincts : DatatypeAux.descr list ->
+ (string * sort) list -> (int * term list) list (*no symmetric inequalities*)
+ val make_ind : DatatypeAux.descr list -> (string * sort) list -> term
+ val make_casedists : DatatypeAux.descr list -> (string * sort) list -> term list
+ val make_primrec_Ts : DatatypeAux.descr list -> (string * sort) list ->
+ string list -> typ list * typ list
+ val make_primrecs : string list -> DatatypeAux.descr list ->
+ (string * sort) list -> theory -> term list
+ val make_cases : string list -> DatatypeAux.descr list ->
+ (string * sort) list -> theory -> term list list
+ val make_splits : string list -> DatatypeAux.descr list ->
+ (string * sort) list -> theory -> (term * term) list
+ val make_weak_case_congs : string list -> DatatypeAux.descr list ->
+ (string * sort) list -> theory -> term list
+ val make_case_congs : string list -> DatatypeAux.descr list ->
+ (string * sort) list -> theory -> term list
+ val make_nchotomys : DatatypeAux.descr list ->
+ (string * sort) list -> term list
+end;
+
+structure DatatypeProp : DATATYPE_PROP =
+struct
+
+open DatatypeAux;
+
+fun indexify_names names =
+ let
+ fun index (x :: xs) tab =
+ (case AList.lookup (op =) tab x of
+ NONE => if member (op =) xs x then (x ^ "1") :: index xs ((x, 2) :: tab) else x :: index xs tab
+ | SOME i => (x ^ string_of_int i) :: index xs ((x, i + 1) :: tab))
+ | index [] _ = [];
+ in index names [] end;
+
+fun make_tnames Ts =
+ let
+ fun type_name (TFree (name, _)) = implode (tl (explode name))
+ | type_name (Type (name, _)) =
+ let val name' = Long_Name.base_name name
+ in if Syntax.is_identifier name' then name' else "x" end;
+ in indexify_names (map type_name Ts) end;
+
+
+(************************* injectivity of constructors ************************)
+
+fun make_injs descr sorts =
+ let
+ val descr' = flat descr;
+ fun make_inj T (cname, cargs) =
+ if null cargs then I else
+ let
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val constr_t = Const (cname, Ts ---> T);
+ val tnames = make_tnames Ts;
+ val frees = map Free (tnames ~~ Ts);
+ val frees' = map Free ((map ((op ^) o (rpair "'")) tnames) ~~ Ts);
+ in cons (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (HOLogic.mk_eq (list_comb (constr_t, frees), list_comb (constr_t, frees')),
+ foldr1 (HOLogic.mk_binop "op &")
+ (map HOLogic.mk_eq (frees ~~ frees')))))
+ end;
+ in
+ map2 (fn d => fn T => fold_rev (make_inj T) (#3 (snd d)) [])
+ (hd descr) (Library.take (length (hd descr), get_rec_types descr' sorts))
+ end;
+
+
+(************************* distinctness of constructors ***********************)
+
+fun make_distincts descr sorts =
+ let
+ val descr' = flat descr;
+ val recTs = get_rec_types descr' sorts;
+ val newTs = Library.take (length (hd descr), recTs);
+
+ fun prep_constr (cname, cargs) = (cname, map (typ_of_dtyp descr' sorts) cargs);
+
+ fun make_distincts' _ [] = []
+ | make_distincts' T ((cname, cargs)::constrs) =
+ let
+ val frees = map Free ((make_tnames cargs) ~~ cargs);
+ val t = list_comb (Const (cname, cargs ---> T), frees);
+
+ fun make_distincts'' (cname', cargs') =
+ let
+ val frees' = map Free ((map ((op ^) o (rpair "'"))
+ (make_tnames cargs')) ~~ cargs');
+ val t' = list_comb (Const (cname', cargs' ---> T), frees')
+ in
+ HOLogic.mk_Trueprop (HOLogic.Not $ HOLogic.mk_eq (t, t'))
+ end
+
+ in map make_distincts'' constrs @ make_distincts' T constrs end;
+
+ in
+ map2 (fn ((_, (_, _, constrs))) => fn T =>
+ (length constrs, make_distincts' T (map prep_constr constrs))) (hd descr) newTs
+ end;
+
+
+(********************************* induction **********************************)
+
+fun make_ind descr sorts =
+ let
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val pnames = if length descr' = 1 then ["P"]
+ else map (fn i => "P" ^ string_of_int i) (1 upto length descr');
+
+ fun make_pred i T =
+ let val T' = T --> HOLogic.boolT
+ in Free (List.nth (pnames, i), T') end;
+
+ fun make_ind_prem k T (cname, cargs) =
+ let
+ fun mk_prem ((dt, s), T) =
+ let val (Us, U) = strip_type T
+ in list_all (map (pair "x") Us, HOLogic.mk_Trueprop
+ (make_pred (body_index dt) U $ app_bnds (Free (s, T)) (length Us)))
+ end;
+
+ val recs = List.filter is_rec_type cargs;
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val recTs' = map (typ_of_dtyp descr' sorts) recs;
+ val tnames = Name.variant_list pnames (make_tnames Ts);
+ val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
+ val frees = tnames ~~ Ts;
+ val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs');
+
+ in list_all_free (frees, Logic.list_implies (prems,
+ HOLogic.mk_Trueprop (make_pred k T $
+ list_comb (Const (cname, Ts ---> T), map Free frees))))
+ end;
+
+ val prems = List.concat (map (fn ((i, (_, _, constrs)), T) =>
+ map (make_ind_prem i T) constrs) (descr' ~~ recTs));
+ val tnames = make_tnames recTs;
+ val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
+ (map (fn (((i, _), T), tname) => make_pred i T $ Free (tname, T))
+ (descr' ~~ recTs ~~ tnames)))
+
+ in Logic.list_implies (prems, concl) end;
+
+(******************************* case distinction *****************************)
+
+fun make_casedists descr sorts =
+ let
+ val descr' = List.concat descr;
+
+ fun make_casedist_prem T (cname, cargs) =
+ let
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val frees = Name.variant_list ["P", "y"] (make_tnames Ts) ~~ Ts;
+ val free_ts = map Free frees
+ in list_all_free (frees, Logic.mk_implies (HOLogic.mk_Trueprop
+ (HOLogic.mk_eq (Free ("y", T), list_comb (Const (cname, Ts ---> T), free_ts))),
+ HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT))))
+ end;
+
+ fun make_casedist ((_, (_, _, constrs)), T) =
+ let val prems = map (make_casedist_prem T) constrs
+ in Logic.list_implies (prems, HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT)))
+ end
+
+ in map make_casedist
+ ((hd descr) ~~ Library.take (length (hd descr), get_rec_types descr' sorts))
+ end;
+
+(*************** characteristic equations for primrec combinator **************)
+
+fun make_primrec_Ts descr sorts used =
+ let
+ val descr' = List.concat descr;
+
+ val rec_result_Ts = map TFree (Name.variant_list used (replicate (length descr') "'t") ~~
+ replicate (length descr') HOLogic.typeS);
+
+ val reccomb_fn_Ts = List.concat (map (fn (i, (_, _, constrs)) =>
+ map (fn (_, cargs) =>
+ let
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val recs = List.filter (is_rec_type o fst) (cargs ~~ Ts);
+
+ fun mk_argT (dt, T) =
+ binder_types T ---> List.nth (rec_result_Ts, body_index dt);
+
+ val argTs = Ts @ map mk_argT recs
+ in argTs ---> List.nth (rec_result_Ts, i)
+ end) constrs) descr');
+
+ in (rec_result_Ts, reccomb_fn_Ts) end;
+
+fun make_primrecs new_type_names descr sorts thy =
+ let
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
+
+ val (rec_result_Ts, reccomb_fn_Ts) = make_primrec_Ts descr sorts used;
+
+ val rec_fns = map (uncurry (mk_Free "f"))
+ (reccomb_fn_Ts ~~ (1 upto (length reccomb_fn_Ts)));
+
+ val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
+ val reccomb_names = map (Sign.intern_const thy)
+ (if length descr' = 1 then [big_reccomb_name] else
+ (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
+ (1 upto (length descr'))));
+ val reccombs = map (fn ((name, T), T') => list_comb
+ (Const (name, reccomb_fn_Ts @ [T] ---> T'), rec_fns))
+ (reccomb_names ~~ recTs ~~ rec_result_Ts);
+
+ fun make_primrec T comb_t ((ts, f::fs), (cname, cargs)) =
+ let
+ val recs = List.filter is_rec_type cargs;
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val recTs' = map (typ_of_dtyp descr' sorts) recs;
+ val tnames = make_tnames Ts;
+ val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
+ val frees = map Free (tnames ~~ Ts);
+ val frees' = map Free (rec_tnames ~~ recTs');
+
+ fun mk_reccomb ((dt, T), t) =
+ let val (Us, U) = strip_type T
+ in list_abs (map (pair "x") Us,
+ List.nth (reccombs, body_index dt) $ app_bnds t (length Us))
+ end;
+
+ val reccombs' = map mk_reccomb (recs ~~ recTs' ~~ frees')
+
+ in (ts @ [HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (comb_t $ list_comb (Const (cname, Ts ---> T), frees),
+ list_comb (f, frees @ reccombs')))], fs)
+ end
+
+ in fst (Library.foldl (fn (x, ((dt, T), comb_t)) =>
+ Library.foldl (make_primrec T comb_t) (x, #3 (snd dt)))
+ (([], rec_fns), descr' ~~ recTs ~~ reccombs))
+ end;
+
+(****************** make terms of form t_case f1 ... fn *********************)
+
+fun make_case_combs new_type_names descr sorts thy fname =
+ let
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
+ val newTs = Library.take (length (hd descr), recTs);
+ val T' = TFree (Name.variant used "'t", HOLogic.typeS);
+
+ val case_fn_Ts = map (fn (i, (_, _, constrs)) =>
+ map (fn (_, cargs) =>
+ let val Ts = map (typ_of_dtyp descr' sorts) cargs
+ in Ts ---> T' end) constrs) (hd descr);
+
+ val case_names = map (fn s =>
+ Sign.intern_const thy (s ^ "_case")) new_type_names
+ in
+ map (fn ((name, Ts), T) => list_comb
+ (Const (name, Ts @ [T] ---> T'),
+ map (uncurry (mk_Free fname)) (Ts ~~ (1 upto length Ts))))
+ (case_names ~~ case_fn_Ts ~~ newTs)
+ end;
+
+(**************** characteristic equations for case combinator ****************)
+
+fun make_cases new_type_names descr sorts thy =
+ let
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val newTs = Library.take (length (hd descr), recTs);
+
+ fun make_case T comb_t ((cname, cargs), f) =
+ let
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val frees = map Free ((make_tnames Ts) ~~ Ts)
+ in HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (comb_t $ list_comb (Const (cname, Ts ---> T), frees),
+ list_comb (f, frees)))
+ end
+
+ in map (fn (((_, (_, _, constrs)), T), comb_t) =>
+ map (make_case T comb_t) (constrs ~~ (snd (strip_comb comb_t))))
+ ((hd descr) ~~ newTs ~~ (make_case_combs new_type_names descr sorts thy "f"))
+ end;
+
+
+(*************************** the "split" - equations **************************)
+
+fun make_splits new_type_names descr sorts thy =
+ let
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val used' = List.foldr OldTerm.add_typ_tfree_names [] recTs;
+ val newTs = Library.take (length (hd descr), recTs);
+ val T' = TFree (Name.variant used' "'t", HOLogic.typeS);
+ val P = Free ("P", T' --> HOLogic.boolT);
+
+ fun make_split (((_, (_, _, constrs)), T), comb_t) =
+ let
+ val (_, fs) = strip_comb comb_t;
+ val used = ["P", "x"] @ (map (fst o dest_Free) fs);
+
+ fun process_constr (((cname, cargs), f), (t1s, t2s)) =
+ let
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val frees = map Free (Name.variant_list used (make_tnames Ts) ~~ Ts);
+ val eqn = HOLogic.mk_eq (Free ("x", T),
+ list_comb (Const (cname, Ts ---> T), frees));
+ val P' = P $ list_comb (f, frees)
+ in ((List.foldr (fn (Free (s, T), t) => HOLogic.mk_all (s, T, t))
+ (HOLogic.imp $ eqn $ P') frees)::t1s,
+ (List.foldr (fn (Free (s, T), t) => HOLogic.mk_exists (s, T, t))
+ (HOLogic.conj $ eqn $ (HOLogic.Not $ P')) frees)::t2s)
+ end;
+
+ val (t1s, t2s) = List.foldr process_constr ([], []) (constrs ~~ fs);
+ val lhs = P $ (comb_t $ Free ("x", T))
+ in
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, mk_conj t1s)),
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, HOLogic.Not $ mk_disj t2s)))
+ end
+
+ in map make_split ((hd descr) ~~ newTs ~~
+ (make_case_combs new_type_names descr sorts thy "f"))
+ end;
+
+(************************* additional rules for TFL ***************************)
+
+fun make_weak_case_congs new_type_names descr sorts thy =
+ let
+ val case_combs = make_case_combs new_type_names descr sorts thy "f";
+
+ fun mk_case_cong comb =
+ let
+ val Type ("fun", [T, _]) = fastype_of comb;
+ val M = Free ("M", T);
+ val M' = Free ("M'", T);
+ in
+ Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (M, M')),
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (comb $ M, comb $ M')))
+ end
+ in
+ map mk_case_cong case_combs
+ end;
+
+
+(*---------------------------------------------------------------------------
+ * Structure of case congruence theorem looks like this:
+ *
+ * (M = M')
+ * ==> (!!x1,...,xk. (M' = C1 x1..xk) ==> (f1 x1..xk = g1 x1..xk))
+ * ==> ...
+ * ==> (!!x1,...,xj. (M' = Cn x1..xj) ==> (fn x1..xj = gn x1..xj))
+ * ==>
+ * (ty_case f1..fn M = ty_case g1..gn M')
+ *---------------------------------------------------------------------------*)
+
+fun make_case_congs new_type_names descr sorts thy =
+ let
+ val case_combs = make_case_combs new_type_names descr sorts thy "f";
+ val case_combs' = make_case_combs new_type_names descr sorts thy "g";
+
+ fun mk_case_cong ((comb, comb'), (_, (_, _, constrs))) =
+ let
+ val Type ("fun", [T, _]) = fastype_of comb;
+ val (_, fs) = strip_comb comb;
+ val (_, gs) = strip_comb comb';
+ val used = ["M", "M'"] @ map (fst o dest_Free) (fs @ gs);
+ val M = Free ("M", T);
+ val M' = Free ("M'", T);
+
+ fun mk_clause ((f, g), (cname, _)) =
+ let
+ val (Ts, _) = strip_type (fastype_of f);
+ val tnames = Name.variant_list used (make_tnames Ts);
+ val frees = map Free (tnames ~~ Ts)
+ in
+ list_all_free (tnames ~~ Ts, Logic.mk_implies
+ (HOLogic.mk_Trueprop
+ (HOLogic.mk_eq (M', list_comb (Const (cname, Ts ---> T), frees))),
+ HOLogic.mk_Trueprop
+ (HOLogic.mk_eq (list_comb (f, frees), list_comb (g, frees)))))
+ end
+
+ in
+ Logic.list_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (M, M')) ::
+ map mk_clause (fs ~~ gs ~~ constrs),
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (comb $ M, comb' $ M')))
+ end
+
+ in
+ map mk_case_cong (case_combs ~~ case_combs' ~~ hd descr)
+ end;
+
+(*---------------------------------------------------------------------------
+ * Structure of exhaustion theorem looks like this:
+ *
+ * !v. (? y1..yi. v = C1 y1..yi) | ... | (? y1..yj. v = Cn y1..yj)
+ *---------------------------------------------------------------------------*)
+
+fun make_nchotomys descr sorts =
+ let
+ val descr' = List.concat descr;
+ val recTs = get_rec_types descr' sorts;
+ val newTs = Library.take (length (hd descr), recTs);
+
+ fun mk_eqn T (cname, cargs) =
+ let
+ val Ts = map (typ_of_dtyp descr' sorts) cargs;
+ val tnames = Name.variant_list ["v"] (make_tnames Ts);
+ val frees = tnames ~~ Ts
+ in
+ List.foldr (fn ((s, T'), t) => HOLogic.mk_exists (s, T', t))
+ (HOLogic.mk_eq (Free ("v", T),
+ list_comb (Const (cname, Ts ---> T), map Free frees))) frees
+ end
+
+ in map (fn ((_, (_, _, constrs)), T) =>
+ HOLogic.mk_Trueprop (HOLogic.mk_all ("v", T, mk_disj (map (mk_eqn T) constrs))))
+ (hd descr ~~ newTs)
+ end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/datatype_package/datatype_realizer.ML Thu Jun 11 08:04:26 2009 +0200
@@ -0,0 +1,230 @@
+(* Title: HOL/Tools/datatype_realizer.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Porgram extraction from proofs involving datatypes:
+Realizers for induction and case analysis
+*)
+
+signature DATATYPE_REALIZER =
+sig
+ val add_dt_realizers: string list -> theory -> theory
+ val setup: theory -> theory
+end;
+
+structure DatatypeRealizer : DATATYPE_REALIZER =
+struct
+
+open DatatypeAux;
+
+fun subsets i j = if i <= j then
+ let val is = subsets (i+1) j
+ in map (fn ks => i::ks) is @ is end
+ else [[]];
+
+fun forall_intr_prf (t, prf) =
+ let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p)
+ in Abst (a, SOME T, Proofterm.prf_abstract_over t prf) end;
+
+fun prf_of thm =
+ Reconstruct.reconstruct_proof (Thm.theory_of_thm thm) (Thm.prop_of thm) (Thm.proof_of thm);
+
+fun prf_subst_vars inst =
+ Proofterm.map_proof_terms (subst_vars ([], inst)) I;
+
+fun is_unit t = snd (strip_type (fastype_of t)) = HOLogic.unitT;
+
+fun tname_of (Type (s, _)) = s
+ | tname_of _ = "";
+
+fun mk_realizes T = Const ("realizes", T --> HOLogic.boolT --> HOLogic.boolT);
+
+fun make_ind sorts ({descr, rec_names, rec_rewrites, induction, ...} : datatype_info) is thy =
+ let
+ val recTs = get_rec_types descr sorts;
+ val pnames = if length descr = 1 then ["P"]
+ else map (fn i => "P" ^ string_of_int i) (1 upto length descr);
+
+ val rec_result_Ts = map (fn ((i, _), P) =>
+ if i mem is then TFree ("'" ^ P, HOLogic.typeS) else HOLogic.unitT)
+ (descr ~~ pnames);
+
+ fun make_pred i T U r x =
+ if i mem is then
+ Free (List.nth (pnames, i), T --> U --> HOLogic.boolT) $ r $ x
+ else Free (List.nth (pnames, i), U --> HOLogic.boolT) $ x;
+
+ fun mk_all i s T t =
+ if i mem is then list_all_free ([(s, T)], t) else t;
+
+ val (prems, rec_fns) = split_list (flat (fst (fold_map
+ (fn ((i, (_, _, constrs)), T) => fold_map (fn (cname, cargs) => fn j =>
+ let
+ val Ts = map (typ_of_dtyp descr sorts) cargs;
+ val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts);
+ val recs = filter (is_rec_type o fst o fst) (cargs ~~ tnames ~~ Ts);
+ val frees = tnames ~~ Ts;
+
+ fun mk_prems vs [] =
+ let
+ val rT = nth (rec_result_Ts) i;
+ val vs' = filter_out is_unit vs;
+ val f = mk_Free "f" (map fastype_of vs' ---> rT) j;
+ val f' = Envir.eta_contract (list_abs_free
+ (map dest_Free vs, if i mem is then list_comb (f, vs')
+ else HOLogic.unit));
+ in (HOLogic.mk_Trueprop (make_pred i rT T (list_comb (f, vs'))
+ (list_comb (Const (cname, Ts ---> T), map Free frees))), f')
+ end
+ | mk_prems vs (((dt, s), T) :: ds) =
+ let
+ val k = body_index dt;
+ val (Us, U) = strip_type T;
+ val i = length Us;
+ val rT = nth (rec_result_Ts) k;
+ val r = Free ("r" ^ s, Us ---> rT);
+ val (p, f) = mk_prems (vs @ [r]) ds
+ in (mk_all k ("r" ^ s) (Us ---> rT) (Logic.mk_implies
+ (list_all (map (pair "x") Us, HOLogic.mk_Trueprop
+ (make_pred k rT U (app_bnds r i)
+ (app_bnds (Free (s, T)) i))), p)), f)
+ end
+
+ in (apfst (curry list_all_free frees) (mk_prems (map Free frees) recs), j + 1) end)
+ constrs) (descr ~~ recTs) 1)));
+
+ fun mk_proj j [] t = t
+ | mk_proj j (i :: is) t = if null is then t else
+ if (j: int) = i then HOLogic.mk_fst t
+ else mk_proj j is (HOLogic.mk_snd t);
+
+ val tnames = DatatypeProp.make_tnames recTs;
+ val fTs = map fastype_of rec_fns;
+ val ps = map (fn ((((i, _), T), U), s) => Abs ("x", T, make_pred i U T
+ (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Bound 0) (Bound 0)))
+ (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names);
+ val r = if null is then Extraction.nullt else
+ foldr1 HOLogic.mk_prod (List.mapPartial (fn (((((i, _), T), U), s), tname) =>
+ if i mem is then SOME
+ (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Free (tname, T))
+ else NONE) (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names ~~ tnames));
+ val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
+ (map (fn ((((i, _), T), U), tname) =>
+ make_pred i U T (mk_proj i is r) (Free (tname, T)))
+ (descr ~~ recTs ~~ rec_result_Ts ~~ tnames)));
+ val cert = cterm_of thy;
+ val inst = map (pairself cert) (map head_of (HOLogic.dest_conj
+ (HOLogic.dest_Trueprop (concl_of induction))) ~~ ps);
+
+ val thm = OldGoals.simple_prove_goal_cterm (cert (Logic.list_implies (prems, concl)))
+ (fn prems =>
+ [rewrite_goals_tac (map mk_meta_eq [fst_conv, snd_conv]),
+ rtac (cterm_instantiate inst induction) 1,
+ ALLGOALS ObjectLogic.atomize_prems_tac,
+ rewrite_goals_tac (@{thm o_def} :: map mk_meta_eq rec_rewrites),
+ REPEAT ((resolve_tac prems THEN_ALL_NEW (fn i =>
+ REPEAT (etac allE i) THEN atac i)) 1)]);
+
+ val ind_name = Thm.get_name induction;
+ val vs = map (fn i => List.nth (pnames, i)) is;
+ val (thm', thy') = thy
+ |> Sign.root_path
+ |> PureThy.store_thm
+ (Binding.qualified_name (space_implode "_" (ind_name :: vs @ ["correctness"])), thm)
+ ||> Sign.restore_naming thy;
+
+ val ivs = rev (Term.add_vars (Logic.varify (DatatypeProp.make_ind [descr] sorts)) []);
+ val rvs = rev (Thm.fold_terms Term.add_vars thm' []);
+ val ivs1 = map Var (filter_out (fn (_, T) =>
+ tname_of (body_type T) mem ["set", "bool"]) ivs);
+ val ivs2 = map (fn (ixn, _) => Var (ixn, valOf (AList.lookup (op =) rvs ixn))) ivs;
+
+ val prf = List.foldr forall_intr_prf
+ (List.foldr (fn ((f, p), prf) =>
+ (case head_of (strip_abs_body f) of
+ Free (s, T) =>
+ let val T' = Logic.varifyT T
+ in Abst (s, SOME T', Proofterm.prf_abstract_over
+ (Var ((s, 0), T')) (AbsP ("H", SOME p, prf)))
+ end
+ | _ => AbsP ("H", SOME p, prf)))
+ (Proofterm.proof_combP
+ (prf_of thm', map PBound (length prems - 1 downto 0))) (rec_fns ~~ prems_of thm)) ivs2;
+
+ val r' = if null is then r else Logic.varify (List.foldr (uncurry lambda)
+ r (map Logic.unvarify ivs1 @ filter_out is_unit
+ (map (head_of o strip_abs_body) rec_fns)));
+
+ in Extraction.add_realizers_i [(ind_name, (vs, r', prf))] thy' end;
+
+
+fun make_casedists sorts ({index, descr, case_name, case_rewrites, exhaustion, ...} : datatype_info) thy =
+ let
+ val cert = cterm_of thy;
+ val rT = TFree ("'P", HOLogic.typeS);
+ val rT' = TVar (("'P", 0), HOLogic.typeS);
+
+ fun make_casedist_prem T (cname, cargs) =
+ let
+ val Ts = map (typ_of_dtyp descr sorts) cargs;
+ val frees = Name.variant_list ["P", "y"] (DatatypeProp.make_tnames Ts) ~~ Ts;
+ val free_ts = map Free frees;
+ val r = Free ("r" ^ Long_Name.base_name cname, Ts ---> rT)
+ in (r, list_all_free (frees, Logic.mk_implies (HOLogic.mk_Trueprop
+ (HOLogic.mk_eq (Free ("y", T), list_comb (Const (cname, Ts ---> T), free_ts))),
+ HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $
+ list_comb (r, free_ts)))))
+ end;
+
+ val SOME (_, _, constrs) = AList.lookup (op =) descr index;
+ val T = List.nth (get_rec_types descr sorts, index);
+ val (rs, prems) = split_list (map (make_casedist_prem T) constrs);
+ val r = Const (case_name, map fastype_of rs ---> T --> rT);
+
+ val y = Var (("y", 0), Logic.legacy_varifyT T);
+ val y' = Free ("y", T);
+
+ val thm = OldGoals.prove_goalw_cterm [] (cert (Logic.list_implies (prems,
+ HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $
+ list_comb (r, rs @ [y'])))))
+ (fn prems =>
+ [rtac (cterm_instantiate [(cert y, cert y')] exhaustion) 1,
+ ALLGOALS (EVERY'
+ [asm_simp_tac (HOL_basic_ss addsimps case_rewrites),
+ resolve_tac prems, asm_simp_tac HOL_basic_ss])]);
+
+ val exh_name = Thm.get_name exhaustion;
+ val (thm', thy') = thy
+ |> Sign.root_path
+ |> PureThy.store_thm (Binding.qualified_name (exh_name ^ "_P_correctness"), thm)
+ ||> Sign.restore_naming thy;
+
+ val P = Var (("P", 0), rT' --> HOLogic.boolT);
+ val prf = forall_intr_prf (y, forall_intr_prf (P,
+ List.foldr (fn ((p, r), prf) =>
+ forall_intr_prf (Logic.legacy_varify r, AbsP ("H", SOME (Logic.varify p),
+ prf))) (Proofterm.proof_combP (prf_of thm',
+ map PBound (length prems - 1 downto 0))) (prems ~~ rs)));
+ val r' = Logic.legacy_varify (Abs ("y", Logic.legacy_varifyT T,
+ list_abs (map dest_Free rs, list_comb (r,
+ map Bound ((length rs - 1 downto 0) @ [length rs])))));
+
+ in Extraction.add_realizers_i
+ [(exh_name, (["P"], r', prf)),
+ (exh_name, ([], Extraction.nullt, prf_of exhaustion))] thy'
+ end;
+
+fun add_dt_realizers names thy =
+ if ! Proofterm.proofs < 2 then thy
+ else let
+ val _ = message "Adding realizers for induction and case analysis ..."
+ val infos = map (DatatypePackage.the_datatype thy) names;
+ val info :: _ = infos;
+ in
+ thy
+ |> fold_rev (make_ind (#sorts info) info) (subsets 0 (length (#descr info) - 1))
+ |> fold_rev (make_casedists (#sorts info)) infos
+ end;
+
+val setup = DatatypePackage.interpretation add_dt_realizers;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/datatype_package/datatype_rep_proofs.ML Thu Jun 11 08:04:26 2009 +0200
@@ -0,0 +1,642 @@
+(* Title: HOL/Tools/datatype_rep_proofs.ML
+ Author: Stefan Berghofer, TU Muenchen
+
+Definitional introduction of datatypes
+Proof of characteristic theorems:
+
+ - injectivity of constructors
+ - distinctness of constructors
+ - induction theorem
+*)
+
+signature DATATYPE_REP_PROOFS =
+sig
+ val distinctness_limit : int Config.T
+ val distinctness_limit_setup : theory -> theory
+ val representation_proofs : bool -> DatatypeAux.datatype_info Symtab.table ->
+ string list -> DatatypeAux.descr list -> (string * sort) list ->
+ (binding * mixfix) list -> (binding * mixfix) list list -> attribute
+ -> theory -> (thm list list * thm list list * thm list list *
+ DatatypeAux.simproc_dist list * thm) * theory
+end;
+
+structure DatatypeRepProofs : DATATYPE_REP_PROOFS =
+struct
+
+open DatatypeAux;
+
+(*the kind of distinctiveness axioms depends on number of constructors*)
+val (distinctness_limit, distinctness_limit_setup) =
+ Attrib.config_int "datatype_distinctness_limit" 7;
+
+val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
+
+val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq];
+
+
+(** theory context references **)
+
+val f_myinv_f = thm "f_myinv_f";
+val myinv_f_f = thm "myinv_f_f";
+
+
+fun exh_thm_of (dt_info : datatype_info Symtab.table) tname =
+ #exhaustion (the (Symtab.lookup dt_info tname));
+
+(******************************************************************************)
+
+fun representation_proofs flat_names (dt_info : datatype_info Symtab.table)
+ new_type_names descr sorts types_syntax constr_syntax case_names_induct thy =
+ let
+ val Datatype_thy = ThyInfo.the_theory "Datatype" thy;
+ val node_name = "Datatype.node";
+ val In0_name = "Datatype.In0";
+ val In1_name = "Datatype.In1";
+ val Scons_name = "Datatype.Scons";
+ val Leaf_name = "Datatype.Leaf";
+ val Numb_name = "Datatype.Numb";
+ val Lim_name = "Datatype.Lim";
+ val Suml_name = "Datatype.Suml";
+ val Sumr_name = "Datatype.Sumr";
+
+ val [In0_inject, In1_inject, Scons_inject, Leaf_inject,
+ In0_eq, In1_eq, In0_not_In1, In1_not_In0,
+ Lim_inject, Suml_inject, Sumr_inject] = map (PureThy.get_thm Datatype_thy)
+ ["In0_inject", "In1_inject", "Scons_inject", "Leaf_inject",
+ "In0_eq", "In1_eq", "In0_not_In1", "In1_not_In0",
+ "Lim_inject", "Suml_inject", "Sumr_inject"];
+
+ val descr' = flat descr;
+
+ val big_name = space_implode "_" new_type_names;
+ val thy1 = add_path flat_names big_name thy;
+ val big_rec_name = big_name ^ "_rep_set";
+ val rep_set_names' =
+ (if length descr' = 1 then [big_rec_name] else
+ (map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int)
+ (1 upto (length descr'))));
+ val rep_set_names = map (Sign.full_bname thy1) rep_set_names';
+
+ val tyvars = map (fn (_, (_, Ts, _)) => map dest_DtTFree Ts) (hd descr);
+ val leafTs' = get_nonrec_types descr' sorts;
+ val branchTs = get_branching_types descr' sorts;
+ val branchT = if null branchTs then HOLogic.unitT
+ else BalancedTree.make (fn (T, U) => Type ("+", [T, U])) branchTs;
+ val arities = get_arities descr' \ 0;
+ val unneeded_vars = hd tyvars \\ List.foldr OldTerm.add_typ_tfree_names [] (leafTs' @ branchTs);
+ val leafTs = leafTs' @ (map (fn n => TFree (n, (the o AList.lookup (op =) sorts) n)) unneeded_vars);
+ val recTs = get_rec_types descr' sorts;
+ val newTs = Library.take (length (hd descr), recTs);
+ val oldTs = Library.drop (length (hd descr), recTs);
+ val sumT = if null leafTs then HOLogic.unitT
+ else BalancedTree.make (fn (T, U) => Type ("+", [T, U])) leafTs;
+ val Univ_elT = HOLogic.mk_setT (Type (node_name, [sumT, branchT]));
+ val UnivT = HOLogic.mk_setT Univ_elT;
+ val UnivT' = Univ_elT --> HOLogic.boolT;
+ val Collect = Const ("Collect", UnivT' --> UnivT);
+
+ val In0 = Const (In0_name, Univ_elT --> Univ_elT);
+ val In1 = Const (In1_name, Univ_elT --> Univ_elT);
+ val Leaf = Const (Leaf_name, sumT --> Univ_elT);
+ val Lim = Const (Lim_name, (branchT --> Univ_elT) --> Univ_elT);
+
+ (* make injections needed for embedding types in leaves *)
+
+ fun mk_inj T' x =
+ let
+ fun mk_inj' T n i =
+ if n = 1 then x else
+ let val n2 = n div 2;
+ val Type (_, [T1, T2]) = T
+ in
+ if i <= n2 then
+ Const ("Sum_Type.Inl", T1 --> T) $ (mk_inj' T1 n2 i)
+ else
+ Const ("Sum_Type.Inr", T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2))
+ end
+ in mk_inj' sumT (length leafTs) (1 + find_index_eq T' leafTs)
+ end;
+
+ (* make injections for constructors *)
+
+ fun mk_univ_inj ts = BalancedTree.access
+ {left = fn t => In0 $ t,
+ right = fn t => In1 $ t,
+ init =
+ if ts = [] then Const (@{const_name undefined}, Univ_elT)
+ else foldr1 (HOLogic.mk_binop Scons_name) ts};
+
+ (* function spaces *)
+
+ fun mk_fun_inj T' x =
+ let
+ fun mk_inj T n i =
+ if n = 1 then x else
+ let
+ val n2 = n div 2;
+ val Type (_, [T1, T2]) = T;
+ fun mkT U = (U --> Univ_elT) --> T --> Univ_elT
+ in
+ if i <= n2 then Const (Suml_name, mkT T1) $ mk_inj T1 n2 i
+ else Const (Sumr_name, mkT T2) $ mk_inj T2 (n - n2) (i - n2)
+ end
+ in mk_inj branchT (length branchTs) (1 + find_index_eq T' branchTs)
+ end;
+
+ val mk_lim = List.foldr (fn (T, t) => Lim $ mk_fun_inj T (Abs ("x", T, t)));
+
+ (************** generate introduction rules for representing set **********)
+
+ val _ = message "Constructing representing sets ...";
+
+ (* make introduction rule for a single constructor *)
+
+ fun make_intr s n (i, (_, cargs)) =
+ let
+ fun mk_prem (dt, (j, prems, ts)) = (case strip_dtyp dt of
+ (dts, DtRec k) =>
+ let
+ val Ts = map (typ_of_dtyp descr' sorts) dts;
+ val free_t =
+ app_bnds (mk_Free "x" (Ts ---> Univ_elT) j) (length Ts)
+ in (j + 1, list_all (map (pair "x") Ts,
+ HOLogic.mk_Trueprop
+ (Free (List.nth (rep_set_names', k), UnivT') $ free_t)) :: prems,
+ mk_lim free_t Ts :: ts)
+ end
+ | _ =>
+ let val T = typ_of_dtyp descr' sorts dt
+ in (j + 1, prems, (Leaf $ mk_inj T (mk_Free "x" T j))::ts)
+ end);
+
+ val (_, prems, ts) = List.foldr mk_prem (1, [], []) cargs;
+ val concl = HOLogic.mk_Trueprop
+ (Free (s, UnivT') $ mk_univ_inj ts n i)
+ in Logic.list_implies (prems, concl)
+ end;
+
+ val intr_ts = maps (fn ((_, (_, _, constrs)), rep_set_name) =>
+ map (make_intr rep_set_name (length constrs))
+ ((1 upto (length constrs)) ~~ constrs)) (descr' ~~ rep_set_names');
+
+ val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy2) =
+ InductivePackage.add_inductive_global (serial_string ())
+ {quiet_mode = ! quiet_mode, verbose = false, kind = Thm.internalK,
+ alt_name = Binding.name big_rec_name, coind = false, no_elim = true, no_ind = false,
+ skip_mono = true, fork_mono = false}
+ (map (fn s => ((Binding.name s, UnivT'), NoSyn)) rep_set_names') []
+ (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] thy1;
+
+ (********************************* typedef ********************************)
+
+ val (typedefs, thy3) = thy2 |>
+ parent_path flat_names |>
+ fold_map (fn ((((name, mx), tvs), c), name') =>
+ TypedefPackage.add_typedef false (SOME (Binding.name name')) (name, tvs, mx)
+ (Collect $ Const (c, UnivT')) NONE
+ (rtac exI 1 THEN rtac CollectI 1 THEN
+ QUIET_BREADTH_FIRST (has_fewer_prems 1)
+ (resolve_tac rep_intrs 1)))
+ (types_syntax ~~ tyvars ~~
+ (Library.take (length newTs, rep_set_names)) ~~ new_type_names) ||>
+ add_path flat_names big_name;
+
+ (*********************** definition of constructors ***********************)
+
+ val big_rep_name = (space_implode "_" new_type_names) ^ "_Rep_";
+ val rep_names = map (curry op ^ "Rep_") new_type_names;
+ val rep_names' = map (fn i => big_rep_name ^ (string_of_int i))
+ (1 upto (length (flat (tl descr))));
+ val all_rep_names = map (Sign.intern_const thy3) rep_names @
+ map (Sign.full_bname thy3) rep_names';
+
+ (* isomorphism declarations *)
+
+ val iso_decls = map (fn (T, s) => (Binding.name s, T --> Univ_elT, NoSyn))
+ (oldTs ~~ rep_names');
+
+ (* constructor definitions *)
+
+ fun make_constr_def tname T n ((thy, defs, eqns, i), ((cname, cargs), (cname', mx))) =
+ let
+ fun constr_arg (dt, (j, l_args, r_args)) =
+ let val T = typ_of_dtyp descr' sorts dt;
+ val free_t = mk_Free "x" T j
+ in (case (strip_dtyp dt, strip_type T) of
+ ((_, DtRec m), (Us, U)) => (j + 1, free_t :: l_args, mk_lim
+ (Const (List.nth (all_rep_names, m), U --> Univ_elT) $
+ app_bnds free_t (length Us)) Us :: r_args)
+ | _ => (j + 1, free_t::l_args, (Leaf $ mk_inj T free_t)::r_args))
+ end;
+
+ val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) cargs;
+ val constrT = (map (typ_of_dtyp descr' sorts) cargs) ---> T;
+ val abs_name = Sign.intern_const thy ("Abs_" ^ tname);
+ val rep_name = Sign.intern_const thy ("Rep_" ^ tname);
+ val lhs = list_comb (Const (cname, constrT), l_args);
+ val rhs = mk_univ_inj r_args n i;
+ val def = Logic.mk_equals (lhs, Const (abs_name, Univ_elT --> T) $ rhs);
+ val def_name = Long_Name.base_name cname ^ "_def";
+ val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (Const (rep_name, T --> Univ_elT) $ lhs, rhs));
+ val ([def_thm], thy') =
+ thy
+ |> Sign.add_consts_i [(cname', constrT, mx)]
+ |> (PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)];
+
+ in (thy', defs @ [def_thm], eqns @ [eqn], i + 1) end;
+
+ (* constructor definitions for datatype *)
+
+ fun dt_constr_defs ((thy, defs, eqns, rep_congs, dist_lemmas),
+ ((((_, (_, _, constrs)), tname), T), constr_syntax)) =
+ let
+ val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong;
+ val rep_const = cterm_of thy
+ (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> Univ_elT));
+ val cong' = standard (cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong);
+ val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
+ val (thy', defs', eqns', _) = Library.foldl ((make_constr_def tname T) (length constrs))
+ ((add_path flat_names tname thy, defs, [], 1), constrs ~~ constr_syntax)
+ in
+ (parent_path flat_names thy', defs', eqns @ [eqns'],
+ rep_congs @ [cong'], dist_lemmas @ [dist])
+ end;
+
+ val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) = Library.foldl dt_constr_defs
+ ((thy3 |> Sign.add_consts_i iso_decls |> parent_path flat_names, [], [], [], []),
+ hd descr ~~ new_type_names ~~ newTs ~~ constr_syntax);
+
+ (*********** isomorphisms for new types (introduced by typedef) ***********)
+
+ val _ = message "Proving isomorphism properties ...";
+
+ val newT_iso_axms = map (fn (_, td) =>
+ (collect_simp (#Abs_inverse td), #Rep_inverse td,
+ collect_simp (#Rep td))) typedefs;
+
+ val newT_iso_inj_thms = map (fn (_, td) =>
+ (collect_simp (#Abs_inject td) RS iffD1, #Rep_inject td RS iffD1)) typedefs;
+
+ (********* isomorphisms between existing types and "unfolded" types *******)
+
+ (*---------------------------------------------------------------------*)
+ (* isomorphisms are defined using primrec-combinators: *)
+ (* generate appropriate functions for instantiating primrec-combinator *)
+ (* *)
+ (* e.g. dt_Rep_i = list_rec ... (%h t y. In1 (Scons (Leaf h) y)) *)
+ (* *)
+ (* also generate characteristic equations for isomorphisms *)
+ (* *)
+ (* e.g. dt_Rep_i (cons h t) = In1 (Scons (dt_Rep_j h) (dt_Rep_i t)) *)
+ (*---------------------------------------------------------------------*)
+
+ fun make_iso_def k ks n ((fs, eqns, i), (cname, cargs)) =
+ let
+ val argTs = map (typ_of_dtyp descr' sorts) cargs;
+ val T = List.nth (recTs, k);
+ val rep_name = List.nth (all_rep_names, k);
+ val rep_const = Const (rep_name, T --> Univ_elT);
+ val constr = Const (cname, argTs ---> T);
+
+ fun process_arg ks' ((i2, i2', ts, Ts), dt) =
+ let
+ val T' = typ_of_dtyp descr' sorts dt;
+ val (Us, U) = strip_type T'
+ in (case strip_dtyp dt of
+ (_, DtRec j) => if j mem ks' then
+ (i2 + 1, i2' + 1, ts @ [mk_lim (app_bnds
+ (mk_Free "y" (Us ---> Univ_elT) i2') (length Us)) Us],
+ Ts @ [Us ---> Univ_elT])
+ else
+ (i2 + 1, i2', ts @ [mk_lim
+ (Const (List.nth (all_rep_names, j), U --> Univ_elT) $
+ app_bnds (mk_Free "x" T' i2) (length Us)) Us], Ts)
+ | _ => (i2 + 1, i2', ts @ [Leaf $ mk_inj T' (mk_Free "x" T' i2)], Ts))
+ end;
+
+ val (i2, i2', ts, Ts) = Library.foldl (process_arg ks) ((1, 1, [], []), cargs);
+ val xs = map (uncurry (mk_Free "x")) (argTs ~~ (1 upto (i2 - 1)));
+ val ys = map (uncurry (mk_Free "y")) (Ts ~~ (1 upto (i2' - 1)));
+ val f = list_abs_free (map dest_Free (xs @ ys), mk_univ_inj ts n i);
+
+ val (_, _, ts', _) = Library.foldl (process_arg []) ((1, 1, [], []), cargs);
+ val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (rep_const $ list_comb (constr, xs), mk_univ_inj ts' n i))
+
+ in (fs @ [f], eqns @ [eqn], i + 1) end;
+
+ (* define isomorphisms for all mutually recursive datatypes in list ds *)
+
+ fun make_iso_defs (ds, (thy, char_thms)) =
+ let
+ val ks = map fst ds;
+ val (_, (tname, _, _)) = hd ds;
+ val {rec_rewrites, rec_names, ...} = the (Symtab.lookup dt_info tname);
+
+ fun process_dt ((fs, eqns, isos), (k, (tname, _, constrs))) =
+ let
+ val (fs', eqns', _) = Library.foldl (make_iso_def k ks (length constrs))
+ ((fs, eqns, 1), constrs);
+ val iso = (List.nth (recTs, k), List.nth (all_rep_names, k))
+ in (fs', eqns', isos @ [iso]) end;
+
+ val (fs, eqns, isos) = Library.foldl process_dt (([], [], []), ds);
+ val fTs = map fastype_of fs;
+ val defs = map (fn (rec_name, (T, iso_name)) => (Binding.name (Long_Name.base_name iso_name ^ "_def"),
+ Logic.mk_equals (Const (iso_name, T --> Univ_elT),
+ list_comb (Const (rec_name, fTs @ [T] ---> Univ_elT), fs)))) (rec_names ~~ isos);
+ val (def_thms, thy') =
+ apsnd Theory.checkpoint ((PureThy.add_defs false o map Thm.no_attributes) defs thy);
+
+ (* prove characteristic equations *)
+
+ val rewrites = def_thms @ (map mk_meta_eq rec_rewrites);
+ val char_thms' = map (fn eqn => SkipProof.prove_global thy' [] [] eqn
+ (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])) eqns;
+
+ in (thy', char_thms' @ char_thms) end;
+
+ val (thy5, iso_char_thms) = apfst Theory.checkpoint (List.foldr make_iso_defs
+ (add_path flat_names big_name thy4, []) (tl descr));
+
+ (* prove isomorphism properties *)
+
+ fun mk_funs_inv thy thm =
+ let
+ val prop = Thm.prop_of thm;
+ val _ $ (_ $ ((S as Const (_, Type (_, [U, _]))) $ _ )) $
+ (_ $ (_ $ (r $ (a $ _)) $ _)) = Type.freeze prop;
+ val used = OldTerm.add_term_tfree_names (a, []);
+
+ fun mk_thm i =
+ let
+ val Ts = map (TFree o rpair HOLogic.typeS)
+ (Name.variant_list used (replicate i "'t"));
+ val f = Free ("f", Ts ---> U)
+ in SkipProof.prove_global thy [] [] (Logic.mk_implies
+ (HOLogic.mk_Trueprop (HOLogic.list_all
+ (map (pair "x") Ts, S $ app_bnds f i)),
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (list_abs (map (pair "x") Ts,
+ r $ (a $ app_bnds f i)), f))))
+ (fn _ => EVERY [REPEAT_DETERM_N i (rtac ext 1),
+ REPEAT (etac allE 1), rtac thm 1, atac 1])
+ end
+ in map (fn r => r RS subst) (thm :: map mk_thm arities) end;
+
+ (* prove inj dt_Rep_i and dt_Rep_i x : dt_rep_set_i *)
+
+ val fun_congs = map (fn T => make_elim (Drule.instantiate'
+ [SOME (ctyp_of thy5 T)] [] fun_cong)) branchTs;
+
+ fun prove_iso_thms (ds, (inj_thms, elem_thms)) =
+ let
+ val (_, (tname, _, _)) = hd ds;
+ val {induction, ...} = the (Symtab.lookup dt_info tname);
+
+ fun mk_ind_concl (i, _) =
+ let
+ val T = List.nth (recTs, i);
+ val Rep_t = Const (List.nth (all_rep_names, i), T --> Univ_elT);
+ val rep_set_name = List.nth (rep_set_names, i)
+ in (HOLogic.all_const T $ Abs ("y", T, HOLogic.imp $
+ HOLogic.mk_eq (Rep_t $ mk_Free "x" T i, Rep_t $ Bound 0) $
+ HOLogic.mk_eq (mk_Free "x" T i, Bound 0)),
+ Const (rep_set_name, UnivT') $ (Rep_t $ mk_Free "x" T i))
+ end;
+
+ val (ind_concl1, ind_concl2) = ListPair.unzip (map mk_ind_concl ds);
+
+ val rewrites = map mk_meta_eq iso_char_thms;
+ val inj_thms' = map snd newT_iso_inj_thms @
+ map (fn r => r RS @{thm injD}) inj_thms;
+
+ val inj_thm = SkipProof.prove_global thy5 [] []
+ (HOLogic.mk_Trueprop (mk_conj ind_concl1)) (fn _ => EVERY
+ [(indtac induction [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
+ REPEAT (EVERY
+ [rtac allI 1, rtac impI 1,
+ exh_tac (exh_thm_of dt_info) 1,
+ REPEAT (EVERY
+ [hyp_subst_tac 1,
+ rewrite_goals_tac rewrites,
+ REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
+ (eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1)
+ ORELSE (EVERY
+ [REPEAT (eresolve_tac (Scons_inject ::
+ map make_elim [Leaf_inject, Inl_inject, Inr_inject]) 1),
+ REPEAT (cong_tac 1), rtac refl 1,
+ REPEAT (atac 1 ORELSE (EVERY
+ [REPEAT (rtac ext 1),
+ REPEAT (eresolve_tac (mp :: allE ::
+ map make_elim (Suml_inject :: Sumr_inject ::
+ Lim_inject :: inj_thms') @ fun_congs) 1),
+ atac 1]))])])])]);
+
+ val inj_thms'' = map (fn r => r RS @{thm datatype_injI})
+ (split_conj_thm inj_thm);
+
+ val elem_thm =
+ SkipProof.prove_global thy5 [] [] (HOLogic.mk_Trueprop (mk_conj ind_concl2))
+ (fn _ =>
+ EVERY [(indtac induction [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
+ rewrite_goals_tac rewrites,
+ REPEAT ((resolve_tac rep_intrs THEN_ALL_NEW
+ ((REPEAT o etac allE) THEN' ares_tac elem_thms)) 1)]);
+
+ in (inj_thms'' @ inj_thms, elem_thms @ (split_conj_thm elem_thm))
+ end;
+
+ val (iso_inj_thms_unfolded, iso_elem_thms) = List.foldr prove_iso_thms
+ ([], map #3 newT_iso_axms) (tl descr);
+ val iso_inj_thms = map snd newT_iso_inj_thms @
+ map (fn r => r RS @{thm injD}) iso_inj_thms_unfolded;
+
+ (* prove dt_rep_set_i x --> x : range dt_Rep_i *)
+
+ fun mk_iso_t (((set_name, iso_name), i), T) =
+ let val isoT = T --> Univ_elT
+ in HOLogic.imp $
+ (Const (set_name, UnivT') $ mk_Free "x" Univ_elT i) $
+ (if i < length newTs then Const ("True", HOLogic.boolT)
+ else HOLogic.mk_mem (mk_Free "x" Univ_elT i,
+ Const ("image", [isoT, HOLogic.mk_setT T] ---> UnivT) $
+ Const (iso_name, isoT) $ Const (@{const_name UNIV}, HOLogic.mk_setT T)))
+ end;
+
+ val iso_t = HOLogic.mk_Trueprop (mk_conj (map mk_iso_t
+ (rep_set_names ~~ all_rep_names ~~ (0 upto (length descr' - 1)) ~~ recTs)));
+
+ (* all the theorems are proved by one single simultaneous induction *)
+
+ val range_eqs = map (fn r => mk_meta_eq (r RS @{thm range_ex1_eq}))
+ iso_inj_thms_unfolded;
+
+ val iso_thms = if length descr = 1 then [] else
+ Library.drop (length newTs, split_conj_thm
+ (SkipProof.prove_global thy5 [] [] iso_t (fn _ => EVERY
+ [(indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
+ REPEAT (rtac TrueI 1),
+ rewrite_goals_tac (mk_meta_eq choice_eq ::
+ symmetric (mk_meta_eq @{thm expand_fun_eq}) :: range_eqs),
+ rewrite_goals_tac (map symmetric range_eqs),
+ REPEAT (EVERY
+ [REPEAT (eresolve_tac ([rangeE, ex1_implies_ex RS exE] @
+ maps (mk_funs_inv thy5 o #1) newT_iso_axms) 1),
+ TRY (hyp_subst_tac 1),
+ rtac (sym RS range_eqI) 1,
+ resolve_tac iso_char_thms 1])])));
+
+ val Abs_inverse_thms' =
+ map #1 newT_iso_axms @
+ map2 (fn r_inj => fn r => f_myinv_f OF [r_inj, r RS mp])
+ iso_inj_thms_unfolded iso_thms;
+
+ val Abs_inverse_thms = maps (mk_funs_inv thy5) Abs_inverse_thms';
+
+ (******************* freeness theorems for constructors *******************)
+
+ val _ = message "Proving freeness of constructors ...";
+
+ (* prove theorem Rep_i (Constr_j ...) = Inj_j ... *)
+
+ fun prove_constr_rep_thm eqn =
+ let
+ val inj_thms = map fst newT_iso_inj_thms;
+ val rewrites = @{thm o_def} :: constr_defs @ (map (mk_meta_eq o #2) newT_iso_axms)
+ in SkipProof.prove_global thy5 [] [] eqn (fn _ => EVERY
+ [resolve_tac inj_thms 1,
+ rewrite_goals_tac rewrites,
+ rtac refl 3,
+ resolve_tac rep_intrs 2,
+ REPEAT (resolve_tac iso_elem_thms 1)])
+ end;
+
+ (*--------------------------------------------------------------*)
+ (* constr_rep_thms and rep_congs are used to prove distinctness *)
+ (* of constructors. *)
+ (*--------------------------------------------------------------*)
+
+ val constr_rep_thms = map (map prove_constr_rep_thm) constr_rep_eqns;
+
+ val dist_rewrites = map (fn (rep_thms, dist_lemma) =>
+ dist_lemma::(rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0]))
+ (constr_rep_thms ~~ dist_lemmas);
+
+ fun prove_distinct_thms _ _ (_, []) = []
+ | prove_distinct_thms lim dist_rewrites' (k, ts as _ :: _) =
+ if k >= lim then [] else let
+ (*number of constructors < distinctness_limit : C_i ... ~= C_j ...*)
+ fun prove [] = []
+ | prove (t :: ts) =
+ let
+ val dist_thm = SkipProof.prove_global thy5 [] [] t (fn _ =>
+ EVERY [simp_tac (HOL_ss addsimps dist_rewrites') 1])
+ in dist_thm :: standard (dist_thm RS not_sym) :: prove ts end;
+ in prove ts end;
+
+ val distinct_thms = DatatypeProp.make_distincts descr sorts
+ |> map2 (prove_distinct_thms
+ (Config.get_thy thy5 distinctness_limit)) dist_rewrites;
+
+ val simproc_dists = map (fn ((((_, (_, _, constrs)), rep_thms), congr), dists) =>
+ if length constrs < Config.get_thy thy5 distinctness_limit
+ then FewConstrs dists
+ else ManyConstrs (congr, HOL_basic_ss addsimps rep_thms)) (hd descr ~~
+ constr_rep_thms ~~ rep_congs ~~ distinct_thms);
+
+ (* prove injectivity of constructors *)
+
+ fun prove_constr_inj_thm rep_thms t =
+ let val inj_thms = Scons_inject :: (map make_elim
+ (iso_inj_thms @
+ [In0_inject, In1_inject, Leaf_inject, Inl_inject, Inr_inject,
+ Lim_inject, Suml_inject, Sumr_inject]))
+ in SkipProof.prove_global thy5 [] [] t (fn _ => EVERY
+ [rtac iffI 1,
+ REPEAT (etac conjE 2), hyp_subst_tac 2, rtac refl 2,
+ dresolve_tac rep_congs 1, dtac box_equals 1,
+ REPEAT (resolve_tac rep_thms 1),
+ REPEAT (eresolve_tac inj_thms 1),
+ REPEAT (ares_tac [conjI] 1 ORELSE (EVERY [REPEAT (rtac ext 1),
+ REPEAT (eresolve_tac (make_elim fun_cong :: inj_thms) 1),
+ atac 1]))])
+ end;
+
+ val constr_inject = map (fn (ts, thms) => map (prove_constr_inj_thm thms) ts)
+ ((DatatypeProp.make_injs descr sorts) ~~ constr_rep_thms);
+
+ val ((constr_inject', distinct_thms'), thy6) =
+ thy5
+ |> parent_path flat_names
+ |> store_thmss "inject" new_type_names constr_inject
+ ||>> store_thmss "distinct" new_type_names distinct_thms;
+
+ (*************************** induction theorem ****************************)
+
+ val _ = message "Proving induction rule for datatypes ...";
+
+ val Rep_inverse_thms = (map (fn (_, iso, _) => iso RS subst) newT_iso_axms) @
+ (map (fn r => r RS myinv_f_f RS subst) iso_inj_thms_unfolded);
+ val Rep_inverse_thms' = map (fn r => r RS myinv_f_f) iso_inj_thms_unfolded;
+
+ fun mk_indrule_lemma ((prems, concls), ((i, _), T)) =
+ let
+ val Rep_t = Const (List.nth (all_rep_names, i), T --> Univ_elT) $
+ mk_Free "x" T i;
+
+ val Abs_t = if i < length newTs then
+ Const (Sign.intern_const thy6
+ ("Abs_" ^ (List.nth (new_type_names, i))), Univ_elT --> T)
+ else Const ("Inductive.myinv", [T --> Univ_elT, Univ_elT] ---> T) $
+ Const (List.nth (all_rep_names, i), T --> Univ_elT)
+
+ in (prems @ [HOLogic.imp $
+ (Const (List.nth (rep_set_names, i), UnivT') $ Rep_t) $
+ (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
+ concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i])
+ end;
+
+ val (indrule_lemma_prems, indrule_lemma_concls) =
+ Library.foldl mk_indrule_lemma (([], []), (descr' ~~ recTs));
+
+ val cert = cterm_of thy6;
+
+ val indrule_lemma = SkipProof.prove_global thy6 [] []
+ (Logic.mk_implies
+ (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems),
+ HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY
+ [REPEAT (etac conjE 1),
+ REPEAT (EVERY
+ [TRY (rtac conjI 1), resolve_tac Rep_inverse_thms 1,
+ etac mp 1, resolve_tac iso_elem_thms 1])]);
+
+ val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
+ val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else
+ map (Free o apfst fst o dest_Var) Ps;
+ val indrule_lemma' = cterm_instantiate (map cert Ps ~~ map cert frees) indrule_lemma;
+
+ val dt_induct_prop = DatatypeProp.make_ind descr sorts;
+ val dt_induct = SkipProof.prove_global thy6 []
+ (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop)
+ (fn {prems, ...} => EVERY
+ [rtac indrule_lemma' 1,
+ (indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
+ EVERY (map (fn (prem, r) => (EVERY
+ [REPEAT (eresolve_tac Abs_inverse_thms 1),
+ simp_tac (HOL_basic_ss addsimps ((symmetric r)::Rep_inverse_thms')) 1,
+ DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
+ (prems ~~ (constr_defs @ (map mk_meta_eq iso_char_thms))))]);
+
+ val ([dt_induct'], thy7) =
+ thy6
+ |> Sign.add_path big_name
+ |> PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])]
+ ||> Sign.parent_path
+ ||> Theory.checkpoint;
+
+ in
+ ((constr_inject', distinct_thms', dist_rewrites, simproc_dists, dt_induct'), thy7)
+ end;
+
+end;
--- a/src/HOL/Tools/datatype_prop.ML Wed Jun 10 21:04:36 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,435 +0,0 @@
-(* Title: HOL/Tools/datatype_prop.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Characteristic properties of datatypes.
-*)
-
-signature DATATYPE_PROP =
-sig
- val indexify_names: string list -> string list
- val make_tnames: typ list -> string list
- val make_injs : DatatypeAux.descr list -> (string * sort) list -> term list list
- val make_distincts : DatatypeAux.descr list ->
- (string * sort) list -> (int * term list) list (*no symmetric inequalities*)
- val make_ind : DatatypeAux.descr list -> (string * sort) list -> term
- val make_casedists : DatatypeAux.descr list -> (string * sort) list -> term list
- val make_primrec_Ts : DatatypeAux.descr list -> (string * sort) list ->
- string list -> typ list * typ list
- val make_primrecs : string list -> DatatypeAux.descr list ->
- (string * sort) list -> theory -> term list
- val make_cases : string list -> DatatypeAux.descr list ->
- (string * sort) list -> theory -> term list list
- val make_splits : string list -> DatatypeAux.descr list ->
- (string * sort) list -> theory -> (term * term) list
- val make_weak_case_congs : string list -> DatatypeAux.descr list ->
- (string * sort) list -> theory -> term list
- val make_case_congs : string list -> DatatypeAux.descr list ->
- (string * sort) list -> theory -> term list
- val make_nchotomys : DatatypeAux.descr list ->
- (string * sort) list -> term list
-end;
-
-structure DatatypeProp : DATATYPE_PROP =
-struct
-
-open DatatypeAux;
-
-fun indexify_names names =
- let
- fun index (x :: xs) tab =
- (case AList.lookup (op =) tab x of
- NONE => if member (op =) xs x then (x ^ "1") :: index xs ((x, 2) :: tab) else x :: index xs tab
- | SOME i => (x ^ string_of_int i) :: index xs ((x, i + 1) :: tab))
- | index [] _ = [];
- in index names [] end;
-
-fun make_tnames Ts =
- let
- fun type_name (TFree (name, _)) = implode (tl (explode name))
- | type_name (Type (name, _)) =
- let val name' = Long_Name.base_name name
- in if Syntax.is_identifier name' then name' else "x" end;
- in indexify_names (map type_name Ts) end;
-
-
-(************************* injectivity of constructors ************************)
-
-fun make_injs descr sorts =
- let
- val descr' = flat descr;
- fun make_inj T (cname, cargs) =
- if null cargs then I else
- let
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val constr_t = Const (cname, Ts ---> T);
- val tnames = make_tnames Ts;
- val frees = map Free (tnames ~~ Ts);
- val frees' = map Free ((map ((op ^) o (rpair "'")) tnames) ~~ Ts);
- in cons (HOLogic.mk_Trueprop (HOLogic.mk_eq
- (HOLogic.mk_eq (list_comb (constr_t, frees), list_comb (constr_t, frees')),
- foldr1 (HOLogic.mk_binop "op &")
- (map HOLogic.mk_eq (frees ~~ frees')))))
- end;
- in
- map2 (fn d => fn T => fold_rev (make_inj T) (#3 (snd d)) [])
- (hd descr) (Library.take (length (hd descr), get_rec_types descr' sorts))
- end;
-
-
-(************************* distinctness of constructors ***********************)
-
-fun make_distincts descr sorts =
- let
- val descr' = flat descr;
- val recTs = get_rec_types descr' sorts;
- val newTs = Library.take (length (hd descr), recTs);
-
- fun prep_constr (cname, cargs) = (cname, map (typ_of_dtyp descr' sorts) cargs);
-
- fun make_distincts' _ [] = []
- | make_distincts' T ((cname, cargs)::constrs) =
- let
- val frees = map Free ((make_tnames cargs) ~~ cargs);
- val t = list_comb (Const (cname, cargs ---> T), frees);
-
- fun make_distincts'' (cname', cargs') =
- let
- val frees' = map Free ((map ((op ^) o (rpair "'"))
- (make_tnames cargs')) ~~ cargs');
- val t' = list_comb (Const (cname', cargs' ---> T), frees')
- in
- HOLogic.mk_Trueprop (HOLogic.Not $ HOLogic.mk_eq (t, t'))
- end
-
- in map make_distincts'' constrs @ make_distincts' T constrs end;
-
- in
- map2 (fn ((_, (_, _, constrs))) => fn T =>
- (length constrs, make_distincts' T (map prep_constr constrs))) (hd descr) newTs
- end;
-
-
-(********************************* induction **********************************)
-
-fun make_ind descr sorts =
- let
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val pnames = if length descr' = 1 then ["P"]
- else map (fn i => "P" ^ string_of_int i) (1 upto length descr');
-
- fun make_pred i T =
- let val T' = T --> HOLogic.boolT
- in Free (List.nth (pnames, i), T') end;
-
- fun make_ind_prem k T (cname, cargs) =
- let
- fun mk_prem ((dt, s), T) =
- let val (Us, U) = strip_type T
- in list_all (map (pair "x") Us, HOLogic.mk_Trueprop
- (make_pred (body_index dt) U $ app_bnds (Free (s, T)) (length Us)))
- end;
-
- val recs = List.filter is_rec_type cargs;
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val recTs' = map (typ_of_dtyp descr' sorts) recs;
- val tnames = Name.variant_list pnames (make_tnames Ts);
- val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
- val frees = tnames ~~ Ts;
- val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs');
-
- in list_all_free (frees, Logic.list_implies (prems,
- HOLogic.mk_Trueprop (make_pred k T $
- list_comb (Const (cname, Ts ---> T), map Free frees))))
- end;
-
- val prems = List.concat (map (fn ((i, (_, _, constrs)), T) =>
- map (make_ind_prem i T) constrs) (descr' ~~ recTs));
- val tnames = make_tnames recTs;
- val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
- (map (fn (((i, _), T), tname) => make_pred i T $ Free (tname, T))
- (descr' ~~ recTs ~~ tnames)))
-
- in Logic.list_implies (prems, concl) end;
-
-(******************************* case distinction *****************************)
-
-fun make_casedists descr sorts =
- let
- val descr' = List.concat descr;
-
- fun make_casedist_prem T (cname, cargs) =
- let
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val frees = Name.variant_list ["P", "y"] (make_tnames Ts) ~~ Ts;
- val free_ts = map Free frees
- in list_all_free (frees, Logic.mk_implies (HOLogic.mk_Trueprop
- (HOLogic.mk_eq (Free ("y", T), list_comb (Const (cname, Ts ---> T), free_ts))),
- HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT))))
- end;
-
- fun make_casedist ((_, (_, _, constrs)), T) =
- let val prems = map (make_casedist_prem T) constrs
- in Logic.list_implies (prems, HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT)))
- end
-
- in map make_casedist
- ((hd descr) ~~ Library.take (length (hd descr), get_rec_types descr' sorts))
- end;
-
-(*************** characteristic equations for primrec combinator **************)
-
-fun make_primrec_Ts descr sorts used =
- let
- val descr' = List.concat descr;
-
- val rec_result_Ts = map TFree (Name.variant_list used (replicate (length descr') "'t") ~~
- replicate (length descr') HOLogic.typeS);
-
- val reccomb_fn_Ts = List.concat (map (fn (i, (_, _, constrs)) =>
- map (fn (_, cargs) =>
- let
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val recs = List.filter (is_rec_type o fst) (cargs ~~ Ts);
-
- fun mk_argT (dt, T) =
- binder_types T ---> List.nth (rec_result_Ts, body_index dt);
-
- val argTs = Ts @ map mk_argT recs
- in argTs ---> List.nth (rec_result_Ts, i)
- end) constrs) descr');
-
- in (rec_result_Ts, reccomb_fn_Ts) end;
-
-fun make_primrecs new_type_names descr sorts thy =
- let
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
-
- val (rec_result_Ts, reccomb_fn_Ts) = make_primrec_Ts descr sorts used;
-
- val rec_fns = map (uncurry (mk_Free "f"))
- (reccomb_fn_Ts ~~ (1 upto (length reccomb_fn_Ts)));
-
- val big_reccomb_name = (space_implode "_" new_type_names) ^ "_rec";
- val reccomb_names = map (Sign.intern_const thy)
- (if length descr' = 1 then [big_reccomb_name] else
- (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int)
- (1 upto (length descr'))));
- val reccombs = map (fn ((name, T), T') => list_comb
- (Const (name, reccomb_fn_Ts @ [T] ---> T'), rec_fns))
- (reccomb_names ~~ recTs ~~ rec_result_Ts);
-
- fun make_primrec T comb_t ((ts, f::fs), (cname, cargs)) =
- let
- val recs = List.filter is_rec_type cargs;
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val recTs' = map (typ_of_dtyp descr' sorts) recs;
- val tnames = make_tnames Ts;
- val rec_tnames = map fst (List.filter (is_rec_type o snd) (tnames ~~ cargs));
- val frees = map Free (tnames ~~ Ts);
- val frees' = map Free (rec_tnames ~~ recTs');
-
- fun mk_reccomb ((dt, T), t) =
- let val (Us, U) = strip_type T
- in list_abs (map (pair "x") Us,
- List.nth (reccombs, body_index dt) $ app_bnds t (length Us))
- end;
-
- val reccombs' = map mk_reccomb (recs ~~ recTs' ~~ frees')
-
- in (ts @ [HOLogic.mk_Trueprop (HOLogic.mk_eq
- (comb_t $ list_comb (Const (cname, Ts ---> T), frees),
- list_comb (f, frees @ reccombs')))], fs)
- end
-
- in fst (Library.foldl (fn (x, ((dt, T), comb_t)) =>
- Library.foldl (make_primrec T comb_t) (x, #3 (snd dt)))
- (([], rec_fns), descr' ~~ recTs ~~ reccombs))
- end;
-
-(****************** make terms of form t_case f1 ... fn *********************)
-
-fun make_case_combs new_type_names descr sorts thy fname =
- let
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val used = List.foldr OldTerm.add_typ_tfree_names [] recTs;
- val newTs = Library.take (length (hd descr), recTs);
- val T' = TFree (Name.variant used "'t", HOLogic.typeS);
-
- val case_fn_Ts = map (fn (i, (_, _, constrs)) =>
- map (fn (_, cargs) =>
- let val Ts = map (typ_of_dtyp descr' sorts) cargs
- in Ts ---> T' end) constrs) (hd descr);
-
- val case_names = map (fn s =>
- Sign.intern_const thy (s ^ "_case")) new_type_names
- in
- map (fn ((name, Ts), T) => list_comb
- (Const (name, Ts @ [T] ---> T'),
- map (uncurry (mk_Free fname)) (Ts ~~ (1 upto length Ts))))
- (case_names ~~ case_fn_Ts ~~ newTs)
- end;
-
-(**************** characteristic equations for case combinator ****************)
-
-fun make_cases new_type_names descr sorts thy =
- let
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val newTs = Library.take (length (hd descr), recTs);
-
- fun make_case T comb_t ((cname, cargs), f) =
- let
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val frees = map Free ((make_tnames Ts) ~~ Ts)
- in HOLogic.mk_Trueprop (HOLogic.mk_eq
- (comb_t $ list_comb (Const (cname, Ts ---> T), frees),
- list_comb (f, frees)))
- end
-
- in map (fn (((_, (_, _, constrs)), T), comb_t) =>
- map (make_case T comb_t) (constrs ~~ (snd (strip_comb comb_t))))
- ((hd descr) ~~ newTs ~~ (make_case_combs new_type_names descr sorts thy "f"))
- end;
-
-
-(*************************** the "split" - equations **************************)
-
-fun make_splits new_type_names descr sorts thy =
- let
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val used' = List.foldr OldTerm.add_typ_tfree_names [] recTs;
- val newTs = Library.take (length (hd descr), recTs);
- val T' = TFree (Name.variant used' "'t", HOLogic.typeS);
- val P = Free ("P", T' --> HOLogic.boolT);
-
- fun make_split (((_, (_, _, constrs)), T), comb_t) =
- let
- val (_, fs) = strip_comb comb_t;
- val used = ["P", "x"] @ (map (fst o dest_Free) fs);
-
- fun process_constr (((cname, cargs), f), (t1s, t2s)) =
- let
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val frees = map Free (Name.variant_list used (make_tnames Ts) ~~ Ts);
- val eqn = HOLogic.mk_eq (Free ("x", T),
- list_comb (Const (cname, Ts ---> T), frees));
- val P' = P $ list_comb (f, frees)
- in ((List.foldr (fn (Free (s, T), t) => HOLogic.mk_all (s, T, t))
- (HOLogic.imp $ eqn $ P') frees)::t1s,
- (List.foldr (fn (Free (s, T), t) => HOLogic.mk_exists (s, T, t))
- (HOLogic.conj $ eqn $ (HOLogic.Not $ P')) frees)::t2s)
- end;
-
- val (t1s, t2s) = List.foldr process_constr ([], []) (constrs ~~ fs);
- val lhs = P $ (comb_t $ Free ("x", T))
- in
- (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, mk_conj t1s)),
- HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, HOLogic.Not $ mk_disj t2s)))
- end
-
- in map make_split ((hd descr) ~~ newTs ~~
- (make_case_combs new_type_names descr sorts thy "f"))
- end;
-
-(************************* additional rules for TFL ***************************)
-
-fun make_weak_case_congs new_type_names descr sorts thy =
- let
- val case_combs = make_case_combs new_type_names descr sorts thy "f";
-
- fun mk_case_cong comb =
- let
- val Type ("fun", [T, _]) = fastype_of comb;
- val M = Free ("M", T);
- val M' = Free ("M'", T);
- in
- Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (M, M')),
- HOLogic.mk_Trueprop (HOLogic.mk_eq (comb $ M, comb $ M')))
- end
- in
- map mk_case_cong case_combs
- end;
-
-
-(*---------------------------------------------------------------------------
- * Structure of case congruence theorem looks like this:
- *
- * (M = M')
- * ==> (!!x1,...,xk. (M' = C1 x1..xk) ==> (f1 x1..xk = g1 x1..xk))
- * ==> ...
- * ==> (!!x1,...,xj. (M' = Cn x1..xj) ==> (fn x1..xj = gn x1..xj))
- * ==>
- * (ty_case f1..fn M = ty_case g1..gn M')
- *---------------------------------------------------------------------------*)
-
-fun make_case_congs new_type_names descr sorts thy =
- let
- val case_combs = make_case_combs new_type_names descr sorts thy "f";
- val case_combs' = make_case_combs new_type_names descr sorts thy "g";
-
- fun mk_case_cong ((comb, comb'), (_, (_, _, constrs))) =
- let
- val Type ("fun", [T, _]) = fastype_of comb;
- val (_, fs) = strip_comb comb;
- val (_, gs) = strip_comb comb';
- val used = ["M", "M'"] @ map (fst o dest_Free) (fs @ gs);
- val M = Free ("M", T);
- val M' = Free ("M'", T);
-
- fun mk_clause ((f, g), (cname, _)) =
- let
- val (Ts, _) = strip_type (fastype_of f);
- val tnames = Name.variant_list used (make_tnames Ts);
- val frees = map Free (tnames ~~ Ts)
- in
- list_all_free (tnames ~~ Ts, Logic.mk_implies
- (HOLogic.mk_Trueprop
- (HOLogic.mk_eq (M', list_comb (Const (cname, Ts ---> T), frees))),
- HOLogic.mk_Trueprop
- (HOLogic.mk_eq (list_comb (f, frees), list_comb (g, frees)))))
- end
-
- in
- Logic.list_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (M, M')) ::
- map mk_clause (fs ~~ gs ~~ constrs),
- HOLogic.mk_Trueprop (HOLogic.mk_eq (comb $ M, comb' $ M')))
- end
-
- in
- map mk_case_cong (case_combs ~~ case_combs' ~~ hd descr)
- end;
-
-(*---------------------------------------------------------------------------
- * Structure of exhaustion theorem looks like this:
- *
- * !v. (? y1..yi. v = C1 y1..yi) | ... | (? y1..yj. v = Cn y1..yj)
- *---------------------------------------------------------------------------*)
-
-fun make_nchotomys descr sorts =
- let
- val descr' = List.concat descr;
- val recTs = get_rec_types descr' sorts;
- val newTs = Library.take (length (hd descr), recTs);
-
- fun mk_eqn T (cname, cargs) =
- let
- val Ts = map (typ_of_dtyp descr' sorts) cargs;
- val tnames = Name.variant_list ["v"] (make_tnames Ts);
- val frees = tnames ~~ Ts
- in
- List.foldr (fn ((s, T'), t) => HOLogic.mk_exists (s, T', t))
- (HOLogic.mk_eq (Free ("v", T),
- list_comb (Const (cname, Ts ---> T), map Free frees))) frees
- end
-
- in map (fn ((_, (_, _, constrs)), T) =>
- HOLogic.mk_Trueprop (HOLogic.mk_all ("v", T, mk_disj (map (mk_eqn T) constrs))))
- (hd descr ~~ newTs)
- end;
-
-end;
--- a/src/HOL/Tools/datatype_realizer.ML Wed Jun 10 21:04:36 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,230 +0,0 @@
-(* Title: HOL/Tools/datatype_realizer.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Porgram extraction from proofs involving datatypes:
-Realizers for induction and case analysis
-*)
-
-signature DATATYPE_REALIZER =
-sig
- val add_dt_realizers: string list -> theory -> theory
- val setup: theory -> theory
-end;
-
-structure DatatypeRealizer : DATATYPE_REALIZER =
-struct
-
-open DatatypeAux;
-
-fun subsets i j = if i <= j then
- let val is = subsets (i+1) j
- in map (fn ks => i::ks) is @ is end
- else [[]];
-
-fun forall_intr_prf (t, prf) =
- let val (a, T) = (case t of Var ((a, _), T) => (a, T) | Free p => p)
- in Abst (a, SOME T, Proofterm.prf_abstract_over t prf) end;
-
-fun prf_of thm =
- Reconstruct.reconstruct_proof (Thm.theory_of_thm thm) (Thm.prop_of thm) (Thm.proof_of thm);
-
-fun prf_subst_vars inst =
- Proofterm.map_proof_terms (subst_vars ([], inst)) I;
-
-fun is_unit t = snd (strip_type (fastype_of t)) = HOLogic.unitT;
-
-fun tname_of (Type (s, _)) = s
- | tname_of _ = "";
-
-fun mk_realizes T = Const ("realizes", T --> HOLogic.boolT --> HOLogic.boolT);
-
-fun make_ind sorts ({descr, rec_names, rec_rewrites, induction, ...} : datatype_info) is thy =
- let
- val recTs = get_rec_types descr sorts;
- val pnames = if length descr = 1 then ["P"]
- else map (fn i => "P" ^ string_of_int i) (1 upto length descr);
-
- val rec_result_Ts = map (fn ((i, _), P) =>
- if i mem is then TFree ("'" ^ P, HOLogic.typeS) else HOLogic.unitT)
- (descr ~~ pnames);
-
- fun make_pred i T U r x =
- if i mem is then
- Free (List.nth (pnames, i), T --> U --> HOLogic.boolT) $ r $ x
- else Free (List.nth (pnames, i), U --> HOLogic.boolT) $ x;
-
- fun mk_all i s T t =
- if i mem is then list_all_free ([(s, T)], t) else t;
-
- val (prems, rec_fns) = split_list (flat (fst (fold_map
- (fn ((i, (_, _, constrs)), T) => fold_map (fn (cname, cargs) => fn j =>
- let
- val Ts = map (typ_of_dtyp descr sorts) cargs;
- val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts);
- val recs = filter (is_rec_type o fst o fst) (cargs ~~ tnames ~~ Ts);
- val frees = tnames ~~ Ts;
-
- fun mk_prems vs [] =
- let
- val rT = nth (rec_result_Ts) i;
- val vs' = filter_out is_unit vs;
- val f = mk_Free "f" (map fastype_of vs' ---> rT) j;
- val f' = Envir.eta_contract (list_abs_free
- (map dest_Free vs, if i mem is then list_comb (f, vs')
- else HOLogic.unit));
- in (HOLogic.mk_Trueprop (make_pred i rT T (list_comb (f, vs'))
- (list_comb (Const (cname, Ts ---> T), map Free frees))), f')
- end
- | mk_prems vs (((dt, s), T) :: ds) =
- let
- val k = body_index dt;
- val (Us, U) = strip_type T;
- val i = length Us;
- val rT = nth (rec_result_Ts) k;
- val r = Free ("r" ^ s, Us ---> rT);
- val (p, f) = mk_prems (vs @ [r]) ds
- in (mk_all k ("r" ^ s) (Us ---> rT) (Logic.mk_implies
- (list_all (map (pair "x") Us, HOLogic.mk_Trueprop
- (make_pred k rT U (app_bnds r i)
- (app_bnds (Free (s, T)) i))), p)), f)
- end
-
- in (apfst (curry list_all_free frees) (mk_prems (map Free frees) recs), j + 1) end)
- constrs) (descr ~~ recTs) 1)));
-
- fun mk_proj j [] t = t
- | mk_proj j (i :: is) t = if null is then t else
- if (j: int) = i then HOLogic.mk_fst t
- else mk_proj j is (HOLogic.mk_snd t);
-
- val tnames = DatatypeProp.make_tnames recTs;
- val fTs = map fastype_of rec_fns;
- val ps = map (fn ((((i, _), T), U), s) => Abs ("x", T, make_pred i U T
- (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Bound 0) (Bound 0)))
- (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names);
- val r = if null is then Extraction.nullt else
- foldr1 HOLogic.mk_prod (List.mapPartial (fn (((((i, _), T), U), s), tname) =>
- if i mem is then SOME
- (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Free (tname, T))
- else NONE) (descr ~~ recTs ~~ rec_result_Ts ~~ rec_names ~~ tnames));
- val concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &")
- (map (fn ((((i, _), T), U), tname) =>
- make_pred i U T (mk_proj i is r) (Free (tname, T)))
- (descr ~~ recTs ~~ rec_result_Ts ~~ tnames)));
- val cert = cterm_of thy;
- val inst = map (pairself cert) (map head_of (HOLogic.dest_conj
- (HOLogic.dest_Trueprop (concl_of induction))) ~~ ps);
-
- val thm = OldGoals.simple_prove_goal_cterm (cert (Logic.list_implies (prems, concl)))
- (fn prems =>
- [rewrite_goals_tac (map mk_meta_eq [fst_conv, snd_conv]),
- rtac (cterm_instantiate inst induction) 1,
- ALLGOALS ObjectLogic.atomize_prems_tac,
- rewrite_goals_tac (@{thm o_def} :: map mk_meta_eq rec_rewrites),
- REPEAT ((resolve_tac prems THEN_ALL_NEW (fn i =>
- REPEAT (etac allE i) THEN atac i)) 1)]);
-
- val ind_name = Thm.get_name induction;
- val vs = map (fn i => List.nth (pnames, i)) is;
- val (thm', thy') = thy
- |> Sign.root_path
- |> PureThy.store_thm
- (Binding.qualified_name (space_implode "_" (ind_name :: vs @ ["correctness"])), thm)
- ||> Sign.restore_naming thy;
-
- val ivs = rev (Term.add_vars (Logic.varify (DatatypeProp.make_ind [descr] sorts)) []);
- val rvs = rev (Thm.fold_terms Term.add_vars thm' []);
- val ivs1 = map Var (filter_out (fn (_, T) =>
- tname_of (body_type T) mem ["set", "bool"]) ivs);
- val ivs2 = map (fn (ixn, _) => Var (ixn, valOf (AList.lookup (op =) rvs ixn))) ivs;
-
- val prf = List.foldr forall_intr_prf
- (List.foldr (fn ((f, p), prf) =>
- (case head_of (strip_abs_body f) of
- Free (s, T) =>
- let val T' = Logic.varifyT T
- in Abst (s, SOME T', Proofterm.prf_abstract_over
- (Var ((s, 0), T')) (AbsP ("H", SOME p, prf)))
- end
- | _ => AbsP ("H", SOME p, prf)))
- (Proofterm.proof_combP
- (prf_of thm', map PBound (length prems - 1 downto 0))) (rec_fns ~~ prems_of thm)) ivs2;
-
- val r' = if null is then r else Logic.varify (List.foldr (uncurry lambda)
- r (map Logic.unvarify ivs1 @ filter_out is_unit
- (map (head_of o strip_abs_body) rec_fns)));
-
- in Extraction.add_realizers_i [(ind_name, (vs, r', prf))] thy' end;
-
-
-fun make_casedists sorts ({index, descr, case_name, case_rewrites, exhaustion, ...} : datatype_info) thy =
- let
- val cert = cterm_of thy;
- val rT = TFree ("'P", HOLogic.typeS);
- val rT' = TVar (("'P", 0), HOLogic.typeS);
-
- fun make_casedist_prem T (cname, cargs) =
- let
- val Ts = map (typ_of_dtyp descr sorts) cargs;
- val frees = Name.variant_list ["P", "y"] (DatatypeProp.make_tnames Ts) ~~ Ts;
- val free_ts = map Free frees;
- val r = Free ("r" ^ Long_Name.base_name cname, Ts ---> rT)
- in (r, list_all_free (frees, Logic.mk_implies (HOLogic.mk_Trueprop
- (HOLogic.mk_eq (Free ("y", T), list_comb (Const (cname, Ts ---> T), free_ts))),
- HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $
- list_comb (r, free_ts)))))
- end;
-
- val SOME (_, _, constrs) = AList.lookup (op =) descr index;
- val T = List.nth (get_rec_types descr sorts, index);
- val (rs, prems) = split_list (map (make_casedist_prem T) constrs);
- val r = Const (case_name, map fastype_of rs ---> T --> rT);
-
- val y = Var (("y", 0), Logic.legacy_varifyT T);
- val y' = Free ("y", T);
-
- val thm = OldGoals.prove_goalw_cterm [] (cert (Logic.list_implies (prems,
- HOLogic.mk_Trueprop (Free ("P", rT --> HOLogic.boolT) $
- list_comb (r, rs @ [y'])))))
- (fn prems =>
- [rtac (cterm_instantiate [(cert y, cert y')] exhaustion) 1,
- ALLGOALS (EVERY'
- [asm_simp_tac (HOL_basic_ss addsimps case_rewrites),
- resolve_tac prems, asm_simp_tac HOL_basic_ss])]);
-
- val exh_name = Thm.get_name exhaustion;
- val (thm', thy') = thy
- |> Sign.root_path
- |> PureThy.store_thm (Binding.qualified_name (exh_name ^ "_P_correctness"), thm)
- ||> Sign.restore_naming thy;
-
- val P = Var (("P", 0), rT' --> HOLogic.boolT);
- val prf = forall_intr_prf (y, forall_intr_prf (P,
- List.foldr (fn ((p, r), prf) =>
- forall_intr_prf (Logic.legacy_varify r, AbsP ("H", SOME (Logic.varify p),
- prf))) (Proofterm.proof_combP (prf_of thm',
- map PBound (length prems - 1 downto 0))) (prems ~~ rs)));
- val r' = Logic.legacy_varify (Abs ("y", Logic.legacy_varifyT T,
- list_abs (map dest_Free rs, list_comb (r,
- map Bound ((length rs - 1 downto 0) @ [length rs])))));
-
- in Extraction.add_realizers_i
- [(exh_name, (["P"], r', prf)),
- (exh_name, ([], Extraction.nullt, prf_of exhaustion))] thy'
- end;
-
-fun add_dt_realizers names thy =
- if ! Proofterm.proofs < 2 then thy
- else let
- val _ = message "Adding realizers for induction and case analysis ..."
- val infos = map (DatatypePackage.the_datatype thy) names;
- val info :: _ = infos;
- in
- thy
- |> fold_rev (make_ind (#sorts info) info) (subsets 0 (length (#descr info) - 1))
- |> fold_rev (make_casedists (#sorts info)) infos
- end;
-
-val setup = DatatypePackage.interpretation add_dt_realizers;
-
-end;
--- a/src/HOL/Tools/datatype_rep_proofs.ML Wed Jun 10 21:04:36 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,642 +0,0 @@
-(* Title: HOL/Tools/datatype_rep_proofs.ML
- Author: Stefan Berghofer, TU Muenchen
-
-Definitional introduction of datatypes
-Proof of characteristic theorems:
-
- - injectivity of constructors
- - distinctness of constructors
- - induction theorem
-*)
-
-signature DATATYPE_REP_PROOFS =
-sig
- val distinctness_limit : int Config.T
- val distinctness_limit_setup : theory -> theory
- val representation_proofs : bool -> DatatypeAux.datatype_info Symtab.table ->
- string list -> DatatypeAux.descr list -> (string * sort) list ->
- (binding * mixfix) list -> (binding * mixfix) list list -> attribute
- -> theory -> (thm list list * thm list list * thm list list *
- DatatypeAux.simproc_dist list * thm) * theory
-end;
-
-structure DatatypeRepProofs : DATATYPE_REP_PROOFS =
-struct
-
-open DatatypeAux;
-
-(*the kind of distinctiveness axioms depends on number of constructors*)
-val (distinctness_limit, distinctness_limit_setup) =
- Attrib.config_int "datatype_distinctness_limit" 7;
-
-val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma);
-
-val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq];
-
-
-(** theory context references **)
-
-val f_myinv_f = thm "f_myinv_f";
-val myinv_f_f = thm "myinv_f_f";
-
-
-fun exh_thm_of (dt_info : datatype_info Symtab.table) tname =
- #exhaustion (the (Symtab.lookup dt_info tname));
-
-(******************************************************************************)
-
-fun representation_proofs flat_names (dt_info : datatype_info Symtab.table)
- new_type_names descr sorts types_syntax constr_syntax case_names_induct thy =
- let
- val Datatype_thy = ThyInfo.the_theory "Datatype" thy;
- val node_name = "Datatype.node";
- val In0_name = "Datatype.In0";
- val In1_name = "Datatype.In1";
- val Scons_name = "Datatype.Scons";
- val Leaf_name = "Datatype.Leaf";
- val Numb_name = "Datatype.Numb";
- val Lim_name = "Datatype.Lim";
- val Suml_name = "Datatype.Suml";
- val Sumr_name = "Datatype.Sumr";
-
- val [In0_inject, In1_inject, Scons_inject, Leaf_inject,
- In0_eq, In1_eq, In0_not_In1, In1_not_In0,
- Lim_inject, Suml_inject, Sumr_inject] = map (PureThy.get_thm Datatype_thy)
- ["In0_inject", "In1_inject", "Scons_inject", "Leaf_inject",
- "In0_eq", "In1_eq", "In0_not_In1", "In1_not_In0",
- "Lim_inject", "Suml_inject", "Sumr_inject"];
-
- val descr' = flat descr;
-
- val big_name = space_implode "_" new_type_names;
- val thy1 = add_path flat_names big_name thy;
- val big_rec_name = big_name ^ "_rep_set";
- val rep_set_names' =
- (if length descr' = 1 then [big_rec_name] else
- (map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int)
- (1 upto (length descr'))));
- val rep_set_names = map (Sign.full_bname thy1) rep_set_names';
-
- val tyvars = map (fn (_, (_, Ts, _)) => map dest_DtTFree Ts) (hd descr);
- val leafTs' = get_nonrec_types descr' sorts;
- val branchTs = get_branching_types descr' sorts;
- val branchT = if null branchTs then HOLogic.unitT
- else BalancedTree.make (fn (T, U) => Type ("+", [T, U])) branchTs;
- val arities = get_arities descr' \ 0;
- val unneeded_vars = hd tyvars \\ List.foldr OldTerm.add_typ_tfree_names [] (leafTs' @ branchTs);
- val leafTs = leafTs' @ (map (fn n => TFree (n, (the o AList.lookup (op =) sorts) n)) unneeded_vars);
- val recTs = get_rec_types descr' sorts;
- val newTs = Library.take (length (hd descr), recTs);
- val oldTs = Library.drop (length (hd descr), recTs);
- val sumT = if null leafTs then HOLogic.unitT
- else BalancedTree.make (fn (T, U) => Type ("+", [T, U])) leafTs;
- val Univ_elT = HOLogic.mk_setT (Type (node_name, [sumT, branchT]));
- val UnivT = HOLogic.mk_setT Univ_elT;
- val UnivT' = Univ_elT --> HOLogic.boolT;
- val Collect = Const ("Collect", UnivT' --> UnivT);
-
- val In0 = Const (In0_name, Univ_elT --> Univ_elT);
- val In1 = Const (In1_name, Univ_elT --> Univ_elT);
- val Leaf = Const (Leaf_name, sumT --> Univ_elT);
- val Lim = Const (Lim_name, (branchT --> Univ_elT) --> Univ_elT);
-
- (* make injections needed for embedding types in leaves *)
-
- fun mk_inj T' x =
- let
- fun mk_inj' T n i =
- if n = 1 then x else
- let val n2 = n div 2;
- val Type (_, [T1, T2]) = T
- in
- if i <= n2 then
- Const ("Sum_Type.Inl", T1 --> T) $ (mk_inj' T1 n2 i)
- else
- Const ("Sum_Type.Inr", T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2))
- end
- in mk_inj' sumT (length leafTs) (1 + find_index_eq T' leafTs)
- end;
-
- (* make injections for constructors *)
-
- fun mk_univ_inj ts = BalancedTree.access
- {left = fn t => In0 $ t,
- right = fn t => In1 $ t,
- init =
- if ts = [] then Const (@{const_name undefined}, Univ_elT)
- else foldr1 (HOLogic.mk_binop Scons_name) ts};
-
- (* function spaces *)
-
- fun mk_fun_inj T' x =
- let
- fun mk_inj T n i =
- if n = 1 then x else
- let
- val n2 = n div 2;
- val Type (_, [T1, T2]) = T;
- fun mkT U = (U --> Univ_elT) --> T --> Univ_elT
- in
- if i <= n2 then Const (Suml_name, mkT T1) $ mk_inj T1 n2 i
- else Const (Sumr_name, mkT T2) $ mk_inj T2 (n - n2) (i - n2)
- end
- in mk_inj branchT (length branchTs) (1 + find_index_eq T' branchTs)
- end;
-
- val mk_lim = List.foldr (fn (T, t) => Lim $ mk_fun_inj T (Abs ("x", T, t)));
-
- (************** generate introduction rules for representing set **********)
-
- val _ = message "Constructing representing sets ...";
-
- (* make introduction rule for a single constructor *)
-
- fun make_intr s n (i, (_, cargs)) =
- let
- fun mk_prem (dt, (j, prems, ts)) = (case strip_dtyp dt of
- (dts, DtRec k) =>
- let
- val Ts = map (typ_of_dtyp descr' sorts) dts;
- val free_t =
- app_bnds (mk_Free "x" (Ts ---> Univ_elT) j) (length Ts)
- in (j + 1, list_all (map (pair "x") Ts,
- HOLogic.mk_Trueprop
- (Free (List.nth (rep_set_names', k), UnivT') $ free_t)) :: prems,
- mk_lim free_t Ts :: ts)
- end
- | _ =>
- let val T = typ_of_dtyp descr' sorts dt
- in (j + 1, prems, (Leaf $ mk_inj T (mk_Free "x" T j))::ts)
- end);
-
- val (_, prems, ts) = List.foldr mk_prem (1, [], []) cargs;
- val concl = HOLogic.mk_Trueprop
- (Free (s, UnivT') $ mk_univ_inj ts n i)
- in Logic.list_implies (prems, concl)
- end;
-
- val intr_ts = maps (fn ((_, (_, _, constrs)), rep_set_name) =>
- map (make_intr rep_set_name (length constrs))
- ((1 upto (length constrs)) ~~ constrs)) (descr' ~~ rep_set_names');
-
- val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy2) =
- InductivePackage.add_inductive_global (serial_string ())
- {quiet_mode = ! quiet_mode, verbose = false, kind = Thm.internalK,
- alt_name = Binding.name big_rec_name, coind = false, no_elim = true, no_ind = false,
- skip_mono = true, fork_mono = false}
- (map (fn s => ((Binding.name s, UnivT'), NoSyn)) rep_set_names') []
- (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] thy1;
-
- (********************************* typedef ********************************)
-
- val (typedefs, thy3) = thy2 |>
- parent_path flat_names |>
- fold_map (fn ((((name, mx), tvs), c), name') =>
- TypedefPackage.add_typedef false (SOME (Binding.name name')) (name, tvs, mx)
- (Collect $ Const (c, UnivT')) NONE
- (rtac exI 1 THEN rtac CollectI 1 THEN
- QUIET_BREADTH_FIRST (has_fewer_prems 1)
- (resolve_tac rep_intrs 1)))
- (types_syntax ~~ tyvars ~~
- (Library.take (length newTs, rep_set_names)) ~~ new_type_names) ||>
- add_path flat_names big_name;
-
- (*********************** definition of constructors ***********************)
-
- val big_rep_name = (space_implode "_" new_type_names) ^ "_Rep_";
- val rep_names = map (curry op ^ "Rep_") new_type_names;
- val rep_names' = map (fn i => big_rep_name ^ (string_of_int i))
- (1 upto (length (flat (tl descr))));
- val all_rep_names = map (Sign.intern_const thy3) rep_names @
- map (Sign.full_bname thy3) rep_names';
-
- (* isomorphism declarations *)
-
- val iso_decls = map (fn (T, s) => (Binding.name s, T --> Univ_elT, NoSyn))
- (oldTs ~~ rep_names');
-
- (* constructor definitions *)
-
- fun make_constr_def tname T n ((thy, defs, eqns, i), ((cname, cargs), (cname', mx))) =
- let
- fun constr_arg (dt, (j, l_args, r_args)) =
- let val T = typ_of_dtyp descr' sorts dt;
- val free_t = mk_Free "x" T j
- in (case (strip_dtyp dt, strip_type T) of
- ((_, DtRec m), (Us, U)) => (j + 1, free_t :: l_args, mk_lim
- (Const (List.nth (all_rep_names, m), U --> Univ_elT) $
- app_bnds free_t (length Us)) Us :: r_args)
- | _ => (j + 1, free_t::l_args, (Leaf $ mk_inj T free_t)::r_args))
- end;
-
- val (_, l_args, r_args) = List.foldr constr_arg (1, [], []) cargs;
- val constrT = (map (typ_of_dtyp descr' sorts) cargs) ---> T;
- val abs_name = Sign.intern_const thy ("Abs_" ^ tname);
- val rep_name = Sign.intern_const thy ("Rep_" ^ tname);
- val lhs = list_comb (Const (cname, constrT), l_args);
- val rhs = mk_univ_inj r_args n i;
- val def = Logic.mk_equals (lhs, Const (abs_name, Univ_elT --> T) $ rhs);
- val def_name = Long_Name.base_name cname ^ "_def";
- val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
- (Const (rep_name, T --> Univ_elT) $ lhs, rhs));
- val ([def_thm], thy') =
- thy
- |> Sign.add_consts_i [(cname', constrT, mx)]
- |> (PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)];
-
- in (thy', defs @ [def_thm], eqns @ [eqn], i + 1) end;
-
- (* constructor definitions for datatype *)
-
- fun dt_constr_defs ((thy, defs, eqns, rep_congs, dist_lemmas),
- ((((_, (_, _, constrs)), tname), T), constr_syntax)) =
- let
- val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong;
- val rep_const = cterm_of thy
- (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> Univ_elT));
- val cong' = standard (cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong);
- val dist = standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma);
- val (thy', defs', eqns', _) = Library.foldl ((make_constr_def tname T) (length constrs))
- ((add_path flat_names tname thy, defs, [], 1), constrs ~~ constr_syntax)
- in
- (parent_path flat_names thy', defs', eqns @ [eqns'],
- rep_congs @ [cong'], dist_lemmas @ [dist])
- end;
-
- val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) = Library.foldl dt_constr_defs
- ((thy3 |> Sign.add_consts_i iso_decls |> parent_path flat_names, [], [], [], []),
- hd descr ~~ new_type_names ~~ newTs ~~ constr_syntax);
-
- (*********** isomorphisms for new types (introduced by typedef) ***********)
-
- val _ = message "Proving isomorphism properties ...";
-
- val newT_iso_axms = map (fn (_, td) =>
- (collect_simp (#Abs_inverse td), #Rep_inverse td,
- collect_simp (#Rep td))) typedefs;
-
- val newT_iso_inj_thms = map (fn (_, td) =>
- (collect_simp (#Abs_inject td) RS iffD1, #Rep_inject td RS iffD1)) typedefs;
-
- (********* isomorphisms between existing types and "unfolded" types *******)
-
- (*---------------------------------------------------------------------*)
- (* isomorphisms are defined using primrec-combinators: *)
- (* generate appropriate functions for instantiating primrec-combinator *)
- (* *)
- (* e.g. dt_Rep_i = list_rec ... (%h t y. In1 (Scons (Leaf h) y)) *)
- (* *)
- (* also generate characteristic equations for isomorphisms *)
- (* *)
- (* e.g. dt_Rep_i (cons h t) = In1 (Scons (dt_Rep_j h) (dt_Rep_i t)) *)
- (*---------------------------------------------------------------------*)
-
- fun make_iso_def k ks n ((fs, eqns, i), (cname, cargs)) =
- let
- val argTs = map (typ_of_dtyp descr' sorts) cargs;
- val T = List.nth (recTs, k);
- val rep_name = List.nth (all_rep_names, k);
- val rep_const = Const (rep_name, T --> Univ_elT);
- val constr = Const (cname, argTs ---> T);
-
- fun process_arg ks' ((i2, i2', ts, Ts), dt) =
- let
- val T' = typ_of_dtyp descr' sorts dt;
- val (Us, U) = strip_type T'
- in (case strip_dtyp dt of
- (_, DtRec j) => if j mem ks' then
- (i2 + 1, i2' + 1, ts @ [mk_lim (app_bnds
- (mk_Free "y" (Us ---> Univ_elT) i2') (length Us)) Us],
- Ts @ [Us ---> Univ_elT])
- else
- (i2 + 1, i2', ts @ [mk_lim
- (Const (List.nth (all_rep_names, j), U --> Univ_elT) $
- app_bnds (mk_Free "x" T' i2) (length Us)) Us], Ts)
- | _ => (i2 + 1, i2', ts @ [Leaf $ mk_inj T' (mk_Free "x" T' i2)], Ts))
- end;
-
- val (i2, i2', ts, Ts) = Library.foldl (process_arg ks) ((1, 1, [], []), cargs);
- val xs = map (uncurry (mk_Free "x")) (argTs ~~ (1 upto (i2 - 1)));
- val ys = map (uncurry (mk_Free "y")) (Ts ~~ (1 upto (i2' - 1)));
- val f = list_abs_free (map dest_Free (xs @ ys), mk_univ_inj ts n i);
-
- val (_, _, ts', _) = Library.foldl (process_arg []) ((1, 1, [], []), cargs);
- val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq
- (rep_const $ list_comb (constr, xs), mk_univ_inj ts' n i))
-
- in (fs @ [f], eqns @ [eqn], i + 1) end;
-
- (* define isomorphisms for all mutually recursive datatypes in list ds *)
-
- fun make_iso_defs (ds, (thy, char_thms)) =
- let
- val ks = map fst ds;
- val (_, (tname, _, _)) = hd ds;
- val {rec_rewrites, rec_names, ...} = the (Symtab.lookup dt_info tname);
-
- fun process_dt ((fs, eqns, isos), (k, (tname, _, constrs))) =
- let
- val (fs', eqns', _) = Library.foldl (make_iso_def k ks (length constrs))
- ((fs, eqns, 1), constrs);
- val iso = (List.nth (recTs, k), List.nth (all_rep_names, k))
- in (fs', eqns', isos @ [iso]) end;
-
- val (fs, eqns, isos) = Library.foldl process_dt (([], [], []), ds);
- val fTs = map fastype_of fs;
- val defs = map (fn (rec_name, (T, iso_name)) => (Binding.name (Long_Name.base_name iso_name ^ "_def"),
- Logic.mk_equals (Const (iso_name, T --> Univ_elT),
- list_comb (Const (rec_name, fTs @ [T] ---> Univ_elT), fs)))) (rec_names ~~ isos);
- val (def_thms, thy') =
- apsnd Theory.checkpoint ((PureThy.add_defs false o map Thm.no_attributes) defs thy);
-
- (* prove characteristic equations *)
-
- val rewrites = def_thms @ (map mk_meta_eq rec_rewrites);
- val char_thms' = map (fn eqn => SkipProof.prove_global thy' [] [] eqn
- (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])) eqns;
-
- in (thy', char_thms' @ char_thms) end;
-
- val (thy5, iso_char_thms) = apfst Theory.checkpoint (List.foldr make_iso_defs
- (add_path flat_names big_name thy4, []) (tl descr));
-
- (* prove isomorphism properties *)
-
- fun mk_funs_inv thy thm =
- let
- val prop = Thm.prop_of thm;
- val _ $ (_ $ ((S as Const (_, Type (_, [U, _]))) $ _ )) $
- (_ $ (_ $ (r $ (a $ _)) $ _)) = Type.freeze prop;
- val used = OldTerm.add_term_tfree_names (a, []);
-
- fun mk_thm i =
- let
- val Ts = map (TFree o rpair HOLogic.typeS)
- (Name.variant_list used (replicate i "'t"));
- val f = Free ("f", Ts ---> U)
- in SkipProof.prove_global thy [] [] (Logic.mk_implies
- (HOLogic.mk_Trueprop (HOLogic.list_all
- (map (pair "x") Ts, S $ app_bnds f i)),
- HOLogic.mk_Trueprop (HOLogic.mk_eq (list_abs (map (pair "x") Ts,
- r $ (a $ app_bnds f i)), f))))
- (fn _ => EVERY [REPEAT_DETERM_N i (rtac ext 1),
- REPEAT (etac allE 1), rtac thm 1, atac 1])
- end
- in map (fn r => r RS subst) (thm :: map mk_thm arities) end;
-
- (* prove inj dt_Rep_i and dt_Rep_i x : dt_rep_set_i *)
-
- val fun_congs = map (fn T => make_elim (Drule.instantiate'
- [SOME (ctyp_of thy5 T)] [] fun_cong)) branchTs;
-
- fun prove_iso_thms (ds, (inj_thms, elem_thms)) =
- let
- val (_, (tname, _, _)) = hd ds;
- val {induction, ...} = the (Symtab.lookup dt_info tname);
-
- fun mk_ind_concl (i, _) =
- let
- val T = List.nth (recTs, i);
- val Rep_t = Const (List.nth (all_rep_names, i), T --> Univ_elT);
- val rep_set_name = List.nth (rep_set_names, i)
- in (HOLogic.all_const T $ Abs ("y", T, HOLogic.imp $
- HOLogic.mk_eq (Rep_t $ mk_Free "x" T i, Rep_t $ Bound 0) $
- HOLogic.mk_eq (mk_Free "x" T i, Bound 0)),
- Const (rep_set_name, UnivT') $ (Rep_t $ mk_Free "x" T i))
- end;
-
- val (ind_concl1, ind_concl2) = ListPair.unzip (map mk_ind_concl ds);
-
- val rewrites = map mk_meta_eq iso_char_thms;
- val inj_thms' = map snd newT_iso_inj_thms @
- map (fn r => r RS @{thm injD}) inj_thms;
-
- val inj_thm = SkipProof.prove_global thy5 [] []
- (HOLogic.mk_Trueprop (mk_conj ind_concl1)) (fn _ => EVERY
- [(indtac induction [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
- REPEAT (EVERY
- [rtac allI 1, rtac impI 1,
- exh_tac (exh_thm_of dt_info) 1,
- REPEAT (EVERY
- [hyp_subst_tac 1,
- rewrite_goals_tac rewrites,
- REPEAT (dresolve_tac [In0_inject, In1_inject] 1),
- (eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1)
- ORELSE (EVERY
- [REPEAT (eresolve_tac (Scons_inject ::
- map make_elim [Leaf_inject, Inl_inject, Inr_inject]) 1),
- REPEAT (cong_tac 1), rtac refl 1,
- REPEAT (atac 1 ORELSE (EVERY
- [REPEAT (rtac ext 1),
- REPEAT (eresolve_tac (mp :: allE ::
- map make_elim (Suml_inject :: Sumr_inject ::
- Lim_inject :: inj_thms') @ fun_congs) 1),
- atac 1]))])])])]);
-
- val inj_thms'' = map (fn r => r RS @{thm datatype_injI})
- (split_conj_thm inj_thm);
-
- val elem_thm =
- SkipProof.prove_global thy5 [] [] (HOLogic.mk_Trueprop (mk_conj ind_concl2))
- (fn _ =>
- EVERY [(indtac induction [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
- rewrite_goals_tac rewrites,
- REPEAT ((resolve_tac rep_intrs THEN_ALL_NEW
- ((REPEAT o etac allE) THEN' ares_tac elem_thms)) 1)]);
-
- in (inj_thms'' @ inj_thms, elem_thms @ (split_conj_thm elem_thm))
- end;
-
- val (iso_inj_thms_unfolded, iso_elem_thms) = List.foldr prove_iso_thms
- ([], map #3 newT_iso_axms) (tl descr);
- val iso_inj_thms = map snd newT_iso_inj_thms @
- map (fn r => r RS @{thm injD}) iso_inj_thms_unfolded;
-
- (* prove dt_rep_set_i x --> x : range dt_Rep_i *)
-
- fun mk_iso_t (((set_name, iso_name), i), T) =
- let val isoT = T --> Univ_elT
- in HOLogic.imp $
- (Const (set_name, UnivT') $ mk_Free "x" Univ_elT i) $
- (if i < length newTs then Const ("True", HOLogic.boolT)
- else HOLogic.mk_mem (mk_Free "x" Univ_elT i,
- Const ("image", [isoT, HOLogic.mk_setT T] ---> UnivT) $
- Const (iso_name, isoT) $ Const (@{const_name UNIV}, HOLogic.mk_setT T)))
- end;
-
- val iso_t = HOLogic.mk_Trueprop (mk_conj (map mk_iso_t
- (rep_set_names ~~ all_rep_names ~~ (0 upto (length descr' - 1)) ~~ recTs)));
-
- (* all the theorems are proved by one single simultaneous induction *)
-
- val range_eqs = map (fn r => mk_meta_eq (r RS @{thm range_ex1_eq}))
- iso_inj_thms_unfolded;
-
- val iso_thms = if length descr = 1 then [] else
- Library.drop (length newTs, split_conj_thm
- (SkipProof.prove_global thy5 [] [] iso_t (fn _ => EVERY
- [(indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
- REPEAT (rtac TrueI 1),
- rewrite_goals_tac (mk_meta_eq choice_eq ::
- symmetric (mk_meta_eq @{thm expand_fun_eq}) :: range_eqs),
- rewrite_goals_tac (map symmetric range_eqs),
- REPEAT (EVERY
- [REPEAT (eresolve_tac ([rangeE, ex1_implies_ex RS exE] @
- maps (mk_funs_inv thy5 o #1) newT_iso_axms) 1),
- TRY (hyp_subst_tac 1),
- rtac (sym RS range_eqI) 1,
- resolve_tac iso_char_thms 1])])));
-
- val Abs_inverse_thms' =
- map #1 newT_iso_axms @
- map2 (fn r_inj => fn r => f_myinv_f OF [r_inj, r RS mp])
- iso_inj_thms_unfolded iso_thms;
-
- val Abs_inverse_thms = maps (mk_funs_inv thy5) Abs_inverse_thms';
-
- (******************* freeness theorems for constructors *******************)
-
- val _ = message "Proving freeness of constructors ...";
-
- (* prove theorem Rep_i (Constr_j ...) = Inj_j ... *)
-
- fun prove_constr_rep_thm eqn =
- let
- val inj_thms = map fst newT_iso_inj_thms;
- val rewrites = @{thm o_def} :: constr_defs @ (map (mk_meta_eq o #2) newT_iso_axms)
- in SkipProof.prove_global thy5 [] [] eqn (fn _ => EVERY
- [resolve_tac inj_thms 1,
- rewrite_goals_tac rewrites,
- rtac refl 3,
- resolve_tac rep_intrs 2,
- REPEAT (resolve_tac iso_elem_thms 1)])
- end;
-
- (*--------------------------------------------------------------*)
- (* constr_rep_thms and rep_congs are used to prove distinctness *)
- (* of constructors. *)
- (*--------------------------------------------------------------*)
-
- val constr_rep_thms = map (map prove_constr_rep_thm) constr_rep_eqns;
-
- val dist_rewrites = map (fn (rep_thms, dist_lemma) =>
- dist_lemma::(rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0]))
- (constr_rep_thms ~~ dist_lemmas);
-
- fun prove_distinct_thms _ _ (_, []) = []
- | prove_distinct_thms lim dist_rewrites' (k, ts as _ :: _) =
- if k >= lim then [] else let
- (*number of constructors < distinctness_limit : C_i ... ~= C_j ...*)
- fun prove [] = []
- | prove (t :: ts) =
- let
- val dist_thm = SkipProof.prove_global thy5 [] [] t (fn _ =>
- EVERY [simp_tac (HOL_ss addsimps dist_rewrites') 1])
- in dist_thm :: standard (dist_thm RS not_sym) :: prove ts end;
- in prove ts end;
-
- val distinct_thms = DatatypeProp.make_distincts descr sorts
- |> map2 (prove_distinct_thms
- (Config.get_thy thy5 distinctness_limit)) dist_rewrites;
-
- val simproc_dists = map (fn ((((_, (_, _, constrs)), rep_thms), congr), dists) =>
- if length constrs < Config.get_thy thy5 distinctness_limit
- then FewConstrs dists
- else ManyConstrs (congr, HOL_basic_ss addsimps rep_thms)) (hd descr ~~
- constr_rep_thms ~~ rep_congs ~~ distinct_thms);
-
- (* prove injectivity of constructors *)
-
- fun prove_constr_inj_thm rep_thms t =
- let val inj_thms = Scons_inject :: (map make_elim
- (iso_inj_thms @
- [In0_inject, In1_inject, Leaf_inject, Inl_inject, Inr_inject,
- Lim_inject, Suml_inject, Sumr_inject]))
- in SkipProof.prove_global thy5 [] [] t (fn _ => EVERY
- [rtac iffI 1,
- REPEAT (etac conjE 2), hyp_subst_tac 2, rtac refl 2,
- dresolve_tac rep_congs 1, dtac box_equals 1,
- REPEAT (resolve_tac rep_thms 1),
- REPEAT (eresolve_tac inj_thms 1),
- REPEAT (ares_tac [conjI] 1 ORELSE (EVERY [REPEAT (rtac ext 1),
- REPEAT (eresolve_tac (make_elim fun_cong :: inj_thms) 1),
- atac 1]))])
- end;
-
- val constr_inject = map (fn (ts, thms) => map (prove_constr_inj_thm thms) ts)
- ((DatatypeProp.make_injs descr sorts) ~~ constr_rep_thms);
-
- val ((constr_inject', distinct_thms'), thy6) =
- thy5
- |> parent_path flat_names
- |> store_thmss "inject" new_type_names constr_inject
- ||>> store_thmss "distinct" new_type_names distinct_thms;
-
- (*************************** induction theorem ****************************)
-
- val _ = message "Proving induction rule for datatypes ...";
-
- val Rep_inverse_thms = (map (fn (_, iso, _) => iso RS subst) newT_iso_axms) @
- (map (fn r => r RS myinv_f_f RS subst) iso_inj_thms_unfolded);
- val Rep_inverse_thms' = map (fn r => r RS myinv_f_f) iso_inj_thms_unfolded;
-
- fun mk_indrule_lemma ((prems, concls), ((i, _), T)) =
- let
- val Rep_t = Const (List.nth (all_rep_names, i), T --> Univ_elT) $
- mk_Free "x" T i;
-
- val Abs_t = if i < length newTs then
- Const (Sign.intern_const thy6
- ("Abs_" ^ (List.nth (new_type_names, i))), Univ_elT --> T)
- else Const ("Inductive.myinv", [T --> Univ_elT, Univ_elT] ---> T) $
- Const (List.nth (all_rep_names, i), T --> Univ_elT)
-
- in (prems @ [HOLogic.imp $
- (Const (List.nth (rep_set_names, i), UnivT') $ Rep_t) $
- (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))],
- concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i])
- end;
-
- val (indrule_lemma_prems, indrule_lemma_concls) =
- Library.foldl mk_indrule_lemma (([], []), (descr' ~~ recTs));
-
- val cert = cterm_of thy6;
-
- val indrule_lemma = SkipProof.prove_global thy6 [] []
- (Logic.mk_implies
- (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems),
- HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY
- [REPEAT (etac conjE 1),
- REPEAT (EVERY
- [TRY (rtac conjI 1), resolve_tac Rep_inverse_thms 1,
- etac mp 1, resolve_tac iso_elem_thms 1])]);
-
- val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma)));
- val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else
- map (Free o apfst fst o dest_Var) Ps;
- val indrule_lemma' = cterm_instantiate (map cert Ps ~~ map cert frees) indrule_lemma;
-
- val dt_induct_prop = DatatypeProp.make_ind descr sorts;
- val dt_induct = SkipProof.prove_global thy6 []
- (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop)
- (fn {prems, ...} => EVERY
- [rtac indrule_lemma' 1,
- (indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1,
- EVERY (map (fn (prem, r) => (EVERY
- [REPEAT (eresolve_tac Abs_inverse_thms 1),
- simp_tac (HOL_basic_ss addsimps ((symmetric r)::Rep_inverse_thms')) 1,
- DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)]))
- (prems ~~ (constr_defs @ (map mk_meta_eq iso_char_thms))))]);
-
- val ([dt_induct'], thy7) =
- thy6
- |> Sign.add_path big_name
- |> PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])]
- ||> Sign.parent_path
- ||> Theory.checkpoint;
-
- in
- ((constr_inject', distinct_thms', dist_rewrites, simproc_dists, dt_induct'), thy7)
- end;
-
-end;
--- a/src/HOL/Tools/quickcheck_generators.ML Wed Jun 10 21:04:36 2009 +0200
+++ b/src/HOL/Tools/quickcheck_generators.ML Thu Jun 11 08:04:26 2009 +0200
@@ -12,7 +12,11 @@
-> seed -> (('a -> 'b) * (unit -> Term.term)) * seed
val ensure_random_typecopy: string -> theory -> theory
val random_aux_specification: string -> term list -> local_theory -> local_theory
- val eval_ref: (unit -> int -> int * int -> term list option * (int * int)) option ref
+ val mk_random_aux_eqs: theory -> DatatypeAux.descr -> (string * sort) list
+ -> string list -> string list * string list -> typ list * typ list
+ -> string * (term list * (term * term) list)
+ val ensure_random_datatype: string list -> theory -> theory
+ val eval_ref: (unit -> int -> seed -> term list option * seed) option ref
val setup: theory -> theory
end;
@@ -61,10 +65,7 @@
type seed = Random_Engine.seed;
-fun random_fun (T1 : typ) (T2 : typ) (eq : 'a -> 'a -> bool) (term_of : 'a -> term)
- (random : seed -> ('b * (unit -> term)) * seed)
- (random_split : seed -> seed * seed)
- (seed : seed) =
+fun random_fun T1 T2 eq term_of random random_split seed =
let
val (seed', seed'') = random_split seed;
val state = ref (seed', [], Const (@{const_name undefined}, T1 --> T2));
@@ -137,36 +138,53 @@
(* definitional scheme for random instances on datatypes *)
(*FIXME avoid this low-level proving*)
-val rct = Thm.cprop_of @{thm random_aux_rec} |> Thm.dest_arg |> Thm.dest_arg
- |> Thm.dest_arg |> Thm.dest_arg1 |> Thm.dest_fun;
+local
+
fun dest_ctyp_nth k cT = nth (Thm.dest_ctyp cT) k;
-val aT = rct |> Thm.ctyp_of_term |> dest_ctyp_nth 1;
+val eq = Thm.cprop_of @{thm random_aux_rec} |> Thm.dest_arg |> Thm.dest_arg |> Thm.dest_arg;
+val lhs = eq |> Thm.dest_arg1;
+val pt_random_aux = lhs |> Thm.dest_fun;
+val ct_k = lhs |> Thm.dest_arg;
+val pt_rhs = eq |> Thm.dest_arg |> Thm.dest_fun;
+val aT = pt_random_aux |> Thm.ctyp_of_term |> dest_ctyp_nth 1;
+
+val rew_thms = map mk_meta_eq [@{thm code_numeral_zero_minus_one},
+ @{thm Suc_code_numeral_minus_one}, @{thm select_weight_cons_zero}, @{thm beyond_zero}];
+val rew_ts = map (Logic.dest_equals o Thm.prop_of) rew_thms;
+val rew_ss = HOL_ss addsimps rew_thms;
+
+in
fun random_aux_primrec eq lthy =
let
val thy = ProofContext.theory_of lthy;
- val rews = map mk_meta_eq [@{thm code_numeral_zero_minus_one},
- @{thm Suc_code_numeral_minus_one}, @{thm select_weight_cons_zero}, @{thm beyond_zero}];
- val (rt as Free (random_aux, T)) $ (vt as Free (v, _)) =
- (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) eq;
+ val ((t_random_aux as Free (random_aux, T)) $ (t_k as Free (v, _)), proto_t_rhs) =
+ (HOLogic.dest_eq o HOLogic.dest_Trueprop) eq;
val Type (_, [_, iT]) = T;
val icT = Thm.ctyp_of thy iT;
+ val cert = Thm.cterm_of thy;
+ val inst = Thm.instantiate_cterm ([(aT, icT)], []);
fun subst_v t' = map_aterms (fn t as Free (w, _) => if v = w then t' else t | t => t);
- val eqs0 = [subst_v @{term "0::code_numeral"} eq, subst_v (@{term "Suc_code_numeral"} $ vt) eq];
- val eqs1 = map (Pattern.rewrite_term thy (map (Logic.dest_equals o Thm.prop_of) rews) []) eqs0
+ val t_rhs = lambda t_k proto_t_rhs;
+ val eqs0 = [subst_v @{term "0::code_numeral"} eq, subst_v (@{term "Suc_code_numeral"} $ t_k) eq];
+ val eqs1 = map (Pattern.rewrite_term thy rew_ts []) eqs0;
val ((_, eqs2), lthy') = PrimrecPackage.add_primrec_simple
[((Binding.name random_aux, T), NoSyn)] eqs1 lthy;
- val eq_tac = ALLGOALS Goal.conjunction_tac THEN ALLGOALS (simp_tac (HOL_ss addsimps rews))
+ val eq_tac = ALLGOALS Goal.conjunction_tac THEN ALLGOALS (simp_tac rew_ss)
THEN (ALLGOALS (ProofContext.fact_tac (flat eqs2)));
val eqs3 = Goal.prove_multi lthy' [v] [] eqs0 (K eq_tac);
- val rct' = Thm.instantiate_cterm ([(aT, icT)], []) rct
+ val cT_random_aux = inst pt_random_aux;
+ val cT_rhs = inst pt_rhs;
val rule = @{thm random_aux_rec}
- |> Drule.instantiate ([(aT, icT)], [(rct', Thm.cterm_of thy rt)])
- |> (fn thm => thm OF eqs3)
+ |> Drule.instantiate ([(aT, icT)],
+ [(cT_random_aux, cert t_random_aux), (cT_rhs, cert t_rhs)])
+ |> (fn thm => thm OF eqs3);
val tac = ALLGOALS (rtac rule);
val simp = Goal.prove lthy' [v] [] eq (K tac);
in (simp, lthy') end;
+end;
+
fun random_aux_primrec_multi prefix [eq] lthy =
lthy
|> random_aux_primrec eq
@@ -208,19 +226,19 @@
fun random_aux_specification prefix eqs lthy =
let
- val _ $ Free (v, _) $ Free (w, _) =
- (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o hd) eqs;
+ val vs = fold Term.add_free_names ((snd o strip_comb o fst o HOLogic.dest_eq
+ o HOLogic.dest_Trueprop o hd) eqs) [];
fun mk_proto_eq eq =
let
- val (head $ arg, rhs) = (HOLogic.dest_eq o HOLogic.dest_Trueprop) eq;
- in ((HOLogic.mk_Trueprop o HOLogic.mk_eq) (head, lambda arg rhs)) end;
+ val (head $ t $ u, rhs) = (HOLogic.dest_eq o HOLogic.dest_Trueprop) eq;
+ in ((HOLogic.mk_Trueprop o HOLogic.mk_eq) (head, lambda t (lambda u rhs))) end;
val proto_eqs = map mk_proto_eq eqs;
fun prove_simps proto_simps lthy =
let
- val ext_simps = map (fn thm => fun_cong OF [thm]) proto_simps;
+ val ext_simps = map (fn thm => fun_cong OF [fun_cong OF [thm]]) proto_simps;
val tac = ALLGOALS Goal.conjunction_tac
THEN ALLGOALS (ProofContext.fact_tac ext_simps);
- in (Goal.prove_multi lthy [v, w] [] eqs (K tac), lthy) end;
+ in (Goal.prove_multi lthy vs [] eqs (K tac), lthy) end;
val b = Binding.qualify true prefix (Binding.name "simps");
in
lthy
@@ -236,13 +254,135 @@
(* constructing random instances on datatypes *)
-(*still under construction*)
+exception Datatype_Fun; (*FIXME*)
+
+fun mk_random_aux_eqs thy descr vs tycos (names, auxnames) (Ts, Us) =
+ let
+ val mk_const = curry (Sign.mk_const thy);
+ val i = @{term "i\<Colon>code_numeral"};
+ val i1 = @{term "(i\<Colon>code_numeral) - 1"};
+ val j = @{term "j\<Colon>code_numeral"};
+ val seed = @{term "s\<Colon>Random.seed"};
+ val random_auxN = "random_aux";
+ val random_auxsN = map (prefix (random_auxN ^ "_")) (names @ auxnames);
+ fun termifyT T = HOLogic.mk_prodT (T, @{typ "unit \<Rightarrow> term"});
+ val rTs = Ts @ Us;
+ fun random_resultT T = @{typ Random.seed}
+ --> HOLogic.mk_prodT (termifyT T,@{typ Random.seed});
+ val pTs = map random_resultT rTs;
+ fun sizeT T = @{typ code_numeral} --> @{typ code_numeral} --> T;
+ val random_auxT = sizeT o random_resultT;
+ val random_auxs = map2 (fn s => fn rT => Free (s, random_auxT rT))
+ random_auxsN rTs;
+
+ fun mk_random_call T = (NONE, (HOLogic.mk_random T j, T));
+ fun mk_random_aux_call fTs (k, _) (tyco, Ts) =
+ let
+ val _ = if null fTs then () else raise Datatype_Fun; (*FIXME*)
+ val random = nth random_auxs k;
+ val size = Option.map snd (DatatypeCodegen.find_shortest_path descr k)
+ |> the_default 0;
+ in (SOME size, (random $ i1 $ j, Type (tyco, Ts))) end;
+
+ val tss = DatatypeAux.interpret_construction descr vs
+ { atyp = mk_random_call, dtyp = mk_random_aux_call };
+
+ fun mk_consexpr simpleT (c, xs) =
+ let
+ val (ks, simple_tTs) = split_list xs;
+ val T = termifyT simpleT;
+ val tTs = (map o apsnd) termifyT simple_tTs;
+ val is_rec = exists is_some ks;
+ val k = fold (fn NONE => I | SOME k => curry Int.max k) ks 0;
+ val vs = Name.names Name.context "x" (map snd simple_tTs);
+ val vs' = (map o apsnd) termifyT vs;
+ val tc = HOLogic.mk_return T @{typ Random.seed}
+ (HOLogic.mk_valtermify_app c vs simpleT);
+ val t = HOLogic.mk_ST (map (fn (t, _) => (t, @{typ Random.seed})) tTs ~~ map SOME vs')
+ tc @{typ Random.seed} (SOME T, @{typ Random.seed});
+ val tk = if is_rec
+ then if k = 0 then i
+ else @{term "Quickcheck.beyond :: code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"}
+ $ HOLogic.mk_number @{typ code_numeral} k $ i
+ else @{term "1::code_numeral"}
+ in (is_rec, HOLogic.mk_prod (tk, t)) end;
+ fun sort_rec xs =
+ map_filter (fn (true, t) => SOME t | _ => NONE) xs
+ @ map_filter (fn (false, t) => SOME t | _ => NONE) xs;
+ val gen_exprss = tss
+ |> (map o apfst) Type
+ |> map (fn (T, cs) => (T, (sort_rec o map (mk_consexpr T)) cs));
+ fun mk_select (rT, xs) =
+ mk_const @{const_name Quickcheck.collapse} [@{typ "Random.seed"}, termifyT rT]
+ $ (mk_const @{const_name Random.select_weight} [random_resultT rT]
+ $ HOLogic.mk_list (HOLogic.mk_prodT (@{typ code_numeral}, random_resultT rT)) xs)
+ $ seed;
+ val auxs_lhss = map (fn t => t $ i $ j $ seed) random_auxs;
+ val auxs_rhss = map mk_select gen_exprss;
+ val prefix = space_implode "_" (random_auxN :: names);
+ in (prefix, (random_auxs, auxs_lhss ~~ auxs_rhss)) end;
+
+fun mk_random_datatype descr vs tycos (names, auxnames) (Ts, Us) thy =
+ let
+ val i = @{term "i\<Colon>code_numeral"};
+ val mk_prop_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq;
+ fun mk_size_arg k = case DatatypeCodegen.find_shortest_path descr k
+ of SOME (_, l) => if l = 0 then i
+ else @{term "max :: code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"}
+ $ HOLogic.mk_number @{typ code_numeral} l $ i
+ | NONE => i;
+ val (prefix, (random_auxs, auxs_eqs)) = (apsnd o apsnd o map) mk_prop_eq
+ (mk_random_aux_eqs thy descr vs tycos (names, auxnames) (Ts, Us));
+ val random_defs = map_index (fn (k, T) => mk_prop_eq
+ (HOLogic.mk_random T i, nth random_auxs k $ mk_size_arg k $ i)) Ts;
+ in
+ thy
+ |> TheoryTarget.instantiation (tycos, vs, @{sort random})
+ |> random_aux_specification prefix auxs_eqs
+ |> `(fn lthy => map (Syntax.check_term lthy) random_defs)
+ |-> (fn random_defs' => fold_map (fn random_def =>
+ Specification.definition (NONE, (Attrib.empty_binding,
+ random_def))) random_defs')
+ |> snd
+ |> Class.prove_instantiation_exit (K (Class.intro_classes_tac []))
+ end;
+
+fun ensure_random_datatype raw_tycos thy =
+ let
+ val pp = Syntax.pp_global thy;
+ val algebra = Sign.classes_of thy;
+ val (descr, raw_vs, tycos, (names, auxnames), raw_TUs) =
+ DatatypePackage.the_datatype_descr thy raw_tycos;
+ val aTs = (flat o maps snd o maps snd) (DatatypeAux.interpret_construction descr raw_vs
+ { atyp = single, dtyp = K o K });
+ fun meet_random T = Sorts.meet_sort (Sign.classes_of thy) (Logic.varifyT T, @{sort random});
+ val vtab = (Vartab.empty
+ |> fold (fn (v, sort) => Vartab.update ((v, 0), sort)) raw_vs
+ |> fold meet_random aTs
+ |> SOME) handle CLASS_ERROR => NONE;
+ val vconstrain = case vtab of SOME vtab => (fn (v, _) =>
+ (v, (the o Vartab.lookup vtab) (v, 0)))
+ | NONE => I;
+ val vs = map vconstrain raw_vs;
+ val constrain = map_atyps (fn TFree v => TFree (vconstrain v));
+ val has_inst = exists (fn tyco =>
+ can (Sorts.mg_domain algebra tyco) @{sort random}) tycos;
+ in if is_some vtab andalso not has_inst then
+ (mk_random_datatype descr vs tycos (names, auxnames)
+ ((pairself o map) constrain raw_TUs) thy
+ (*FIXME ephemeral handles*)
+ handle Datatype_Fun => thy
+ | e as TERM (msg, ts) => (tracing (cat_lines (msg :: map (Syntax.string_of_term_global thy) ts)); raise e)
+ | e as TYPE (msg, _, _) => (tracing msg; raise e)
+ | e as ERROR msg => (tracing msg; raise e))
+ else thy end;
(** setup **)
val setup = Code_Target.extend_target (target, (Code_ML.target_Eval, K I))
#> Quickcheck.add_generator ("code", compile_generator_expr o ProofContext.theory_of)
- #> TypecopyPackage.interpretation ensure_random_typecopy;
+ #> TypecopyPackage.interpretation ensure_random_typecopy
+ #> DatatypePackage.interpretation ensure_random_datatype;
end;
--- a/src/HOL/Tools/typedef_codegen.ML Wed Jun 10 21:04:36 2009 +0200
+++ b/src/HOL/Tools/typedef_codegen.ML Thu Jun 11 08:04:26 2009 +0200
@@ -1,6 +1,5 @@
(* Title: HOL/Tools/typedef_codegen.ML
- ID: $Id$
- Author: Stefan Berghofer and Florian Haftmann, TU Muenchen
+ Author: Stefan Berghofer, TU Muenchen
Code generators for trivial typedefs.
*)
--- a/src/HOL/ex/Quickcheck_Generators.thy Wed Jun 10 21:04:36 2009 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,274 +0,0 @@
-(* Author: Florian Haftmann, TU Muenchen *)
-
-header {* Experimental counterexample generators *}
-
-theory Quickcheck_Generators
-imports Quickcheck State_Monad
-begin
-
-subsection {* Datatypes *}
-
-definition collapse :: "('a \<Rightarrow> ('a \<Rightarrow> 'b \<times> 'a) \<times> 'a) \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'a" where
- "collapse f = (do g \<leftarrow> f; g done)"
-
-lemma random'_if:
- fixes random' :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> Random.seed \<Rightarrow> ('a \<times> (unit \<Rightarrow> term)) \<times> Random.seed"
- assumes "random' 0 j = (\<lambda>s. undefined)"
- and "\<And>i. random' (Suc_code_numeral i) j = rhs2 i"
- shows "random' i j s = (if i = 0 then undefined else rhs2 (i - 1) s)"
- by (cases i rule: code_numeral.exhaust) (insert assms, simp_all)
-
-setup {*
-let
- fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT);
- fun scomp T1 T2 sT f g = Const (@{const_name scomp},
- liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g;
- exception REC of string;
- exception TYP of string;
- fun mk_collapse thy ty = Sign.mk_const thy
- (@{const_name collapse}, [@{typ Random.seed}, ty]);
- fun term_ty ty = HOLogic.mk_prodT (ty, @{typ "unit \<Rightarrow> term"});
- fun mk_split thy ty ty' = Sign.mk_const thy
- (@{const_name split}, [ty, @{typ "unit \<Rightarrow> term"}, liftT (term_ty ty') @{typ Random.seed}]);
- fun mk_scomp_split thy ty ty' t t' =
- scomp (term_ty ty) (term_ty ty') @{typ Random.seed} t
- (mk_split thy ty ty' $ Abs ("", ty, Abs ("", @{typ "unit \<Rightarrow> term"}, t')))
- fun mk_cons thy this_ty (c, args) =
- let
- val tys = map (fst o fst) args;
- val c_ty = tys ---> this_ty;
- val c = Const (c, tys ---> this_ty);
- val t_indices = map (curry ( op * ) 2) (length tys - 1 downto 0);
- val c_indices = map (curry ( op + ) 1) t_indices;
- val c_t = list_comb (c, map Bound c_indices);
- val t_t = Abs ("", @{typ unit}, HOLogic.reflect_term
- (list_comb (c, map (fn k => Bound (k + 1)) t_indices))
- |> map_aterms (fn t as Bound _ => t $ @{term "()"} | t => t));
- val return = HOLogic.mk_return (term_ty this_ty) @{typ Random.seed}
- (HOLogic.mk_prod (c_t, t_t));
- val t = fold_rev (fn ((ty, _), random) =>
- mk_scomp_split thy ty this_ty random)
- args return;
- val is_rec = exists (snd o fst) args;
- in (is_rec, t) end;
- fun mk_conss thy ty [] = NONE
- | mk_conss thy ty [(_, t)] = SOME t
- | mk_conss thy ty ts = SOME (mk_collapse thy (term_ty ty) $
- (Sign.mk_const thy (@{const_name Random.select}, [liftT (term_ty ty) @{typ Random.seed}]) $
- HOLogic.mk_list (liftT (term_ty ty) @{typ Random.seed}) (map snd ts)));
- fun mk_clauses thy ty (tyco, (ts_rec, ts_atom)) =
- let
- val SOME t_atom = mk_conss thy ty ts_atom;
- in case mk_conss thy ty ts_rec
- of SOME t_rec => mk_collapse thy (term_ty ty) $
- (Sign.mk_const thy (@{const_name Random.select_default}, [liftT (term_ty ty) @{typ Random.seed}]) $
- @{term "i\<Colon>code_numeral"} $ t_rec $ t_atom)
- | NONE => t_atom
- end;
- fun mk_random_eqs thy vs tycos =
- let
- val this_ty = Type (hd tycos, map TFree vs);
- val this_ty' = liftT (term_ty this_ty) @{typ Random.seed};
- val random_name = Long_Name.base_name @{const_name random};
- val random'_name = random_name ^ "_" ^ Class.type_name (hd tycos) ^ "'";
- fun random ty = Sign.mk_const thy (@{const_name random}, [ty]);
- val random' = Free (random'_name,
- @{typ code_numeral} --> @{typ code_numeral} --> this_ty');
- fun atom ty = if Sign.of_sort thy (ty, @{sort random})
- then ((ty, false), random ty $ @{term "j\<Colon>code_numeral"})
- else raise TYP
- ("Will not generate random elements for type(s) " ^ quote (hd tycos));
- fun dtyp tyco = ((this_ty, true), random' $ @{term "i\<Colon>code_numeral"} $ @{term "j\<Colon>code_numeral"});
- fun rtyp (tyco, Ts) _ = raise REC
- ("Will not generate random elements for mutual recursive type " ^ quote (hd tycos));
- val rhss = DatatypePackage.construction_interpretation thy
- { atom = atom, dtyp = dtyp, rtyp = rtyp } vs tycos
- |> fst
- |> (map o apsnd o map) (mk_cons thy this_ty)
- |> (map o apsnd) (List.partition fst)
- |> map (mk_clauses thy this_ty)
- val eqss = map ((apsnd o map) (HOLogic.mk_Trueprop o HOLogic.mk_eq) o (fn rhs => ((this_ty, random'), [
- (random' $ @{term "0\<Colon>code_numeral"} $ @{term "j\<Colon>code_numeral"}, Abs ("s", @{typ Random.seed},
- Const (@{const_name undefined}, HOLogic.mk_prodT (term_ty this_ty, @{typ Random.seed})))),
- (random' $ @{term "Suc_code_numeral i"} $ @{term "j\<Colon>code_numeral"}, rhs)
- ]))) rhss;
- in eqss end;
- fun random_inst [tyco] thy =
- let
- val (raw_vs, _) = DatatypePackage.the_datatype_spec thy tyco;
- val vs = (map o apsnd)
- (curry (Sorts.inter_sort (Sign.classes_of thy)) @{sort random}) raw_vs;
- val ((this_ty, random'), eqs') = singleton (mk_random_eqs thy vs) tyco;
- val eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq)
- (Sign.mk_const thy (@{const_name random}, [this_ty]) $ @{term "i\<Colon>code_numeral"},
- random' $ @{term "max (i\<Colon>code_numeral) 1"} $ @{term "i\<Colon>code_numeral"})
- val del_func = Attrib.internal (fn _ => Thm.declaration_attribute
- (fn thm => Context.mapping (Code.del_eqn thm) I));
- fun add_code simps lthy =
- let
- val thy = ProofContext.theory_of lthy;
- val thm = @{thm random'_if}
- |> Drule.instantiate' [SOME (Thm.ctyp_of thy this_ty)] [SOME (Thm.cterm_of thy random')]
- |> (fn thm => thm OF simps)
- |> singleton (ProofContext.export lthy (ProofContext.init thy));
- val c = (fst o dest_Const o fst o strip_comb o fst
- o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) thm;
- in
- lthy
- |> LocalTheory.theory (Code.del_eqns c
- #> PureThy.add_thm ((Binding.name (fst (dest_Free random') ^ "_code"), thm), [Thm.kind_internal])
- #-> Code.add_eqn)
- end;
- in
- thy
- |> TheoryTarget.instantiation ([tyco], vs, @{sort random})
- |> PrimrecPackage.add_primrec
- [(Binding.name (fst (dest_Free random')), SOME (snd (dest_Free random')), NoSyn)]
- (map (fn eq => ((Binding.empty, [del_func]), eq)) eqs')
- |-> add_code
- |> `(fn lthy => Syntax.check_term lthy eq)
- |-> (fn eq => Specification.definition (NONE, (Attrib.empty_binding, eq)))
- |> snd
- |> Class.prove_instantiation_instance (K (Class.intro_classes_tac []))
- |> LocalTheory.exit_global
- end
- | random_inst tycos thy = raise REC
- ("Will not generate random elements for mutual recursive type(s) " ^ commas (map quote tycos));
- fun add_random_inst [@{type_name bool}] thy = thy
- | add_random_inst [@{type_name nat}] thy = thy
- | add_random_inst [@{type_name char}] thy = thy
- | add_random_inst [@{type_name String.literal}] thy = thy
- | add_random_inst tycos thy = random_inst tycos thy
- handle REC msg => (warning msg; thy)
- | TYP msg => (warning msg; thy)
-in DatatypePackage.interpretation add_random_inst end
-*}
-
-
-subsection {* Examples *}
-
-theorem "map g (map f xs) = map (g o f) xs"
- quickcheck [generator = code]
- by (induct xs) simp_all
-
-theorem "map g (map f xs) = map (f o g) xs"
- quickcheck [generator = code]
- oops
-
-theorem "rev (xs @ ys) = rev ys @ rev xs"
- quickcheck [generator = code]
- by simp
-
-theorem "rev (xs @ ys) = rev xs @ rev ys"
- quickcheck [generator = code]
- oops
-
-theorem "rev (rev xs) = xs"
- quickcheck [generator = code]
- by simp
-
-theorem "rev xs = xs"
- quickcheck [generator = code]
- oops
-
-primrec app :: "('a \<Rightarrow> 'a) list \<Rightarrow> 'a \<Rightarrow> 'a" where
- "app [] x = x"
- | "app (f # fs) x = app fs (f x)"
-
-lemma "app (fs @ gs) x = app gs (app fs x)"
- quickcheck [generator = code]
- by (induct fs arbitrary: x) simp_all
-
-lemma "app (fs @ gs) x = app fs (app gs x)"
- quickcheck [generator = code]
- oops
-
-primrec occurs :: "'a \<Rightarrow> 'a list \<Rightarrow> nat" where
- "occurs a [] = 0"
- | "occurs a (x#xs) = (if (x=a) then Suc(occurs a xs) else occurs a xs)"
-
-primrec del1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
- "del1 a [] = []"
- | "del1 a (x#xs) = (if (x=a) then xs else (x#del1 a xs))"
-
-lemma "Suc (occurs a (del1 a xs)) = occurs a xs"
- -- {* Wrong. Precondition needed.*}
- quickcheck [generator = code]
- oops
-
-lemma "xs ~= [] \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs"
- quickcheck [generator = code]
- -- {* Also wrong.*}
- oops
-
-lemma "0 < occurs a xs \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs"
- quickcheck [generator = code]
- by (induct xs) auto
-
-primrec replace :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
- "replace a b [] = []"
- | "replace a b (x#xs) = (if (x=a) then (b#(replace a b xs))
- else (x#(replace a b xs)))"
-
-lemma "occurs a xs = occurs b (replace a b xs)"
- quickcheck [generator = code]
- -- {* Wrong. Precondition needed.*}
- oops
-
-lemma "occurs b xs = 0 \<or> a=b \<longrightarrow> occurs a xs = occurs b (replace a b xs)"
- quickcheck [generator = code]
- by (induct xs) simp_all
-
-
-subsection {* Trees *}
-
-datatype 'a tree = Twig | Leaf 'a | Branch "'a tree" "'a tree"
-
-primrec leaves :: "'a tree \<Rightarrow> 'a list" where
- "leaves Twig = []"
- | "leaves (Leaf a) = [a]"
- | "leaves (Branch l r) = (leaves l) @ (leaves r)"
-
-primrec plant :: "'a list \<Rightarrow> 'a tree" where
- "plant [] = Twig "
- | "plant (x#xs) = Branch (Leaf x) (plant xs)"
-
-primrec mirror :: "'a tree \<Rightarrow> 'a tree" where
- "mirror (Twig) = Twig "
- | "mirror (Leaf a) = Leaf a "
- | "mirror (Branch l r) = Branch (mirror r) (mirror l)"
-
-theorem "plant (rev (leaves xt)) = mirror xt"
- quickcheck [generator = code]
- --{* Wrong! *}
- oops
-
-theorem "plant (leaves xt @ leaves yt) = Branch xt yt"
- quickcheck [generator = code]
- --{* Wrong! *}
- oops
-
-datatype 'a ntree = Tip "'a" | Node "'a" "'a ntree" "'a ntree"
-
-primrec inOrder :: "'a ntree \<Rightarrow> 'a list" where
- "inOrder (Tip a)= [a]"
- | "inOrder (Node f x y) = (inOrder x)@[f]@(inOrder y)"
-
-primrec root :: "'a ntree \<Rightarrow> 'a" where
- "root (Tip a) = a"
- | "root (Node f x y) = f"
-
-theorem "hd (inOrder xt) = root xt"
- quickcheck [generator = code]
- --{* Wrong! *}
- oops
-
-lemma "int (f k) = k"
- quickcheck [generator = code]
- oops
-
-lemma "int (nat k) = k"
- quickcheck [generator = code]
- oops
-
-end
--- a/src/HOL/ex/ROOT.ML Wed Jun 10 21:04:36 2009 +0200
+++ b/src/HOL/ex/ROOT.ML Thu Jun 11 08:04:26 2009 +0200
@@ -9,7 +9,6 @@
"FuncSet",
"Word",
"Eval_Examples",
- "Quickcheck_Generators",
"Codegenerator_Test",
"Codegenerator_Pretty_Test",
"NormalForm",
--- a/src/Pure/Isar/class_target.ML Wed Jun 10 21:04:36 2009 +0200
+++ b/src/Pure/Isar/class_target.ML Thu Jun 11 08:04:26 2009 +0200
@@ -127,22 +127,21 @@
};
-fun rep_class_data (ClassData data) = data;
-fun mk_class_data ((consts, base_sort, base_morph, assm_intro, of_class, axiom),
+fun make_class_data ((consts, base_sort, base_morph, assm_intro, of_class, axiom),
(defs, operations)) =
ClassData { consts = consts, base_sort = base_sort,
base_morph = base_morph, assm_intro = assm_intro, of_class = of_class, axiom = axiom,
defs = defs, operations = operations };
fun map_class_data f (ClassData { consts, base_sort, base_morph, assm_intro,
of_class, axiom, defs, operations }) =
- mk_class_data (f ((consts, base_sort, base_morph, assm_intro, of_class, axiom),
+ make_class_data (f ((consts, base_sort, base_morph, assm_intro, of_class, axiom),
(defs, operations)));
fun merge_class_data _ (ClassData { consts = consts,
base_sort = base_sort, base_morph = base_morph, assm_intro = assm_intro,
of_class = of_class, axiom = axiom, defs = defs1, operations = operations1 },
ClassData { consts = _, base_sort = _, base_morph = _, assm_intro = _,
of_class = _, axiom = _, defs = defs2, operations = operations2 }) =
- mk_class_data ((consts, base_sort, base_morph, assm_intro, of_class, axiom),
+ make_class_data ((consts, base_sort, base_morph, assm_intro, of_class, axiom),
(Thm.merge_thms (defs1, defs2),
AList.merge (op =) (K true) (operations1, operations2)));
@@ -158,7 +157,9 @@
(* queries *)
-val lookup_class_data = Option.map rep_class_data oo try o Graph.get_node o ClassData.get;
+fun lookup_class_data thy class = case try (Graph.get_node (ClassData.get thy)) class
+ of SOME (ClassData data) => SOME data
+ | NONE => NONE;
fun the_class_data thy class = case lookup_class_data thy class
of NONE => error ("Undeclared class " ^ quote class)
@@ -188,8 +189,8 @@
in (axiom, of_class) end;
fun all_assm_intros thy =
- Graph.fold (fn (_, (data, _)) => fold (insert Thm.eq_thm)
- ((the_list o #assm_intro o rep_class_data) data)) (ClassData.get thy) [];
+ Graph.fold (fn (_, (ClassData { assm_intro, ... }, _)) => fold (insert Thm.eq_thm)
+ (the_list assm_intro)) (ClassData.get thy) [];
fun these_defs thy = maps (#defs o the_class_data thy) o ancestry thy;
fun these_operations thy = maps (#operations o the_class_data thy) o ancestry thy;
@@ -240,7 +241,7 @@
val operations = map (fn (v_ty as (_, ty), (c, _)) =>
(c, (class, (ty, Free v_ty)))) params;
val add_class = Graph.new_node (class,
- mk_class_data (((map o pairself) fst params, base_sort,
+ make_class_data (((map o pairself) fst params, base_sort,
morph, assm_intro, of_class, axiom), ([], operations)))
#> fold (curry Graph.add_edge class) sups;
in ClassData.map add_class thy end;
--- a/src/Pure/Isar/code.ML Wed Jun 10 21:04:36 2009 +0200
+++ b/src/Pure/Isar/code.ML Thu Jun 11 08:04:26 2009 +0200
@@ -580,13 +580,11 @@
cases: (int * (int * string list)) Symtab.table * unit Symtab.table
};
-fun mk_spec ((concluded_history, eqns), (dtyps, cases)) =
+fun make_spec ((concluded_history, eqns), (dtyps, cases)) =
Spec { concluded_history = concluded_history, eqns = eqns, dtyps = dtyps, cases = cases };
-val empty_spec =
- mk_spec ((false, Symtab.empty), (Symtab.empty, (Symtab.empty, Symtab.empty)));
fun map_spec f (Spec { concluded_history = concluded_history, eqns = eqns,
dtyps = dtyps, cases = cases }) =
- mk_spec (f ((concluded_history, eqns), (dtyps, cases)));
+ make_spec (f ((concluded_history, eqns), (dtyps, cases)));
fun merge_spec (Spec { concluded_history = _, eqns = eqns1, dtyps = dtyps1, cases = (cases1, undefs1) },
Spec { concluded_history = _, eqns = eqns2, dtyps = dtyps2, cases = (cases2, undefs2) }) =
let
@@ -602,15 +600,16 @@
val dtyps = Symtab.join (K (AList.merge (op =) (K true))) (dtyps1, dtyps2);
val cases = (Symtab.merge (K true) (cases1, cases2),
Symtab.merge (K true) (undefs1, undefs2));
- in mk_spec ((false, eqns), (dtyps, cases)) end;
+ in make_spec ((false, eqns), (dtyps, cases)) end;
(* code setup data *)
fun the_spec (Spec x) = x;
-val the_eqns = #eqns o the_spec;
-val the_dtyps = #dtyps o the_spec;
-val the_cases = #cases o the_spec;
+fun the_eqns (Spec { eqns, ... }) = eqns;
+fun the_dtyps (Spec { dtyps, ... }) = dtyps;
+fun the_cases (Spec { cases, ... }) = cases;
+fun history_concluded (Spec { concluded_history, ... }) = concluded_history;
val map_concluded_history = map_spec o apfst o apfst;
val map_eqns = map_spec o apfst o apsnd;
val map_dtyps = map_spec o apsnd o apfst;
@@ -665,7 +664,8 @@
structure Code_Data = TheoryDataFun
(
type T = spec * data ref;
- val empty = (empty_spec, ref empty_data);
+ val empty = (make_spec ((false, Symtab.empty),
+ (Symtab.empty, (Symtab.empty, Symtab.empty))), ref empty_data);
fun copy (spec, data) = (spec, ref (! data));
val extend = copy;
fun merge pp ((spec1, data1), (spec2, data2)) =
@@ -706,13 +706,13 @@
|> Option.map (Lazy.force o snd o snd o fst)
|> these;
-fun continue_history thy = if (#concluded_history o the_spec o the_exec) thy
+fun continue_history thy = if (history_concluded o the_exec) thy
then thy
|> (Code_Data.map o apfst o map_concluded_history) (K false)
|> SOME
else NONE;
-fun conclude_history thy = if (#concluded_history o the_spec o the_exec) thy
+fun conclude_history thy = if (history_concluded o the_exec) thy
then NONE
else thy
|> (Code_Data.map o apfst)
--- a/src/Tools/code/code_ml.ML Wed Jun 10 21:04:36 2009 +0200
+++ b/src/Tools/code/code_ml.ML Thu Jun 11 08:04:26 2009 +0200
@@ -311,13 +311,13 @@
let
fun pr_superclass (_, (classrel, dss)) =
concat [
- (str o deresolve) classrel,
+ (str o Long_Name.base_name o deresolve) classrel,
str "=",
pr_dicts NOBR [DictConst dss]
];
fun pr_classparam ((classparam, c_inst), (thm, _)) =
concat [
- (str o deresolve) classparam,
+ (str o Long_Name.base_name o deresolve) classparam,
str "=",
pr_app (K false) thm reserved_names NOBR (c_inst, [])
];
--- a/src/Tools/code/code_preproc.ML Wed Jun 10 21:04:36 2009 +0200
+++ b/src/Tools/code/code_preproc.ML Thu Jun 11 08:04:26 2009 +0200
@@ -44,22 +44,22 @@
functrans: (string * (serial * (theory -> (thm * bool) list -> (thm * bool) list option))) list
};
-fun mk_thmproc ((pre, post), functrans) =
+fun make_thmproc ((pre, post), functrans) =
Thmproc { pre = pre, post = post, functrans = functrans };
fun map_thmproc f (Thmproc { pre, post, functrans }) =
- mk_thmproc (f ((pre, post), functrans));
+ make_thmproc (f ((pre, post), functrans));
fun merge_thmproc (Thmproc { pre = pre1, post = post1, functrans = functrans1 },
Thmproc { pre = pre2, post = post2, functrans = functrans2 }) =
let
val pre = Simplifier.merge_ss (pre1, pre2);
val post = Simplifier.merge_ss (post1, post2);
val functrans = AList.merge (op =) (eq_fst (op =)) (functrans1, functrans2);
- in mk_thmproc ((pre, post), functrans) end;
+ in make_thmproc ((pre, post), functrans) end;
structure Code_Preproc_Data = TheoryDataFun
(
type T = thmproc;
- val empty = mk_thmproc ((Simplifier.empty_ss, Simplifier.empty_ss), []);
+ val empty = make_thmproc ((Simplifier.empty_ss, Simplifier.empty_ss), []);
fun copy spec = spec;
val extend = copy;
fun merge pp = merge_thmproc;
--- a/src/Tools/code/code_target.ML Wed Jun 10 21:04:36 2009 +0200
+++ b/src/Tools/code/code_target.ML Thu Jun 11 08:04:26 2009 +0200
@@ -128,11 +128,11 @@
module_alias: string Symtab.table
};
-fun mk_target ((serial, serializer), ((reserved, includes), (name_syntax_table, module_alias))) =
+fun make_target ((serial, serializer), ((reserved, includes), (name_syntax_table, module_alias))) =
Target { serial = serial, serializer = serializer, reserved = reserved,
includes = includes, name_syntax_table = name_syntax_table, module_alias = module_alias };
fun map_target f ( Target { serial, serializer, reserved, includes, name_syntax_table, module_alias } ) =
- mk_target (f ((serial, serializer), ((reserved, includes), (name_syntax_table, module_alias))));
+ make_target (f ((serial, serializer), ((reserved, includes), (name_syntax_table, module_alias))));
fun merge_target strict target (Target { serial = serial1, serializer = serializer,
reserved = reserved1, includes = includes1,
name_syntax_table = name_syntax_table1, module_alias = module_alias1 },
@@ -140,7 +140,7 @@
reserved = reserved2, includes = includes2,
name_syntax_table = name_syntax_table2, module_alias = module_alias2 }) =
if serial1 = serial2 orelse not strict then
- mk_target ((serial1, serializer),
+ make_target ((serial1, serializer),
((merge (op =) (reserved1, reserved2), Symtab.merge (op =) (includes1, includes2)),
(merge_name_syntax_table (name_syntax_table1, name_syntax_table2),
Symtab.join (K snd) (module_alias1, module_alias2))
@@ -190,7 +190,7 @@
in
thy
|> (CodeTargetData.map o apfst oo Symtab.map_default)
- (target, mk_target ((serial (), seri), (([], Symtab.empty),
+ (target, make_target ((serial (), seri), (([], Symtab.empty),
(mk_name_syntax_table ((Symtab.empty, Symreltab.empty), (Symtab.empty, Symtab.empty)),
Symtab.empty))))
((map_target o apfst o apsnd o K) seri)
--- a/src/Tools/quickcheck.ML Wed Jun 10 21:04:36 2009 +0200
+++ b/src/Tools/quickcheck.ML Thu Jun 11 08:04:26 2009 +0200
@@ -42,13 +42,13 @@
fun dest_test_params (Test_Params { size, iterations, default_type }) =
((size, iterations), default_type);
-fun mk_test_params ((size, iterations), default_type) =
+fun make_test_params ((size, iterations), default_type) =
Test_Params { size = size, iterations = iterations, default_type = default_type };
fun map_test_params f (Test_Params { size, iterations, default_type}) =
- mk_test_params (f ((size, iterations), default_type));
+ make_test_params (f ((size, iterations), default_type));
fun merge_test_params (Test_Params { size = size1, iterations = iterations1, default_type = default_type1 },
Test_Params { size = size2, iterations = iterations2, default_type = default_type2 }) =
- mk_test_params ((Int.max (size1, size2), Int.max (iterations1, iterations2)),
+ make_test_params ((Int.max (size1, size2), Int.max (iterations1, iterations2)),
case default_type1 of NONE => default_type2 | _ => default_type1);
structure Data = TheoryDataFun(