# HG changeset patch # User bulwahn # Date 1264006942 -3600 # Node ID bd7e347eb76816151e4b258d6c92d1cc592db287 # Parent 1dfb1df1d9d0ea768e492293dc297e5e178f315c added registration of equational theorems from prim_rec and rec_def to Spec_Rules diff -r 1dfb1df1d9d0 -r bd7e347eb768 src/HOL/Tools/old_primrec.ML --- a/src/HOL/Tools/old_primrec.ML Wed Jan 20 14:09:46 2010 +0100 +++ b/src/HOL/Tools/old_primrec.ML Wed Jan 20 18:02:22 2010 +0100 @@ -281,11 +281,14 @@ thy' |> fold_map note ((map fst eqns ~~ atts) ~~ map single simps); val simps'' = maps snd simps'; + fun dest_eqn th = fst (strip_comb (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of th))))) + val specs = (distinct (op =) (map dest_eqn simps''), simps'') in thy'' |> note (("simps", [Simplifier.simp_add, Nitpick_Simps.add, Code.add_default_eqn_attribute]), simps'') |> snd + |> Spec_Rules.add_global Spec_Rules.Equational specs |> note (("induct", []), [prepare_induct (#2 (hd dts)) rec_eqns]) |> snd |> Sign.parent_path diff -r 1dfb1df1d9d0 -r bd7e347eb768 src/HOL/Tools/primrec.ML --- a/src/HOL/Tools/primrec.ML Wed Jan 20 14:09:46 2010 +0100 +++ b/src/HOL/Tools/primrec.ML Wed Jan 20 18:02:22 2010 +0100 @@ -265,6 +265,15 @@ local +fun specs_of simps = + let + val eqns = maps snd simps + fun dest_eqn th = fst (strip_comb (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of th))))) + val t = distinct (op =) (map dest_eqn eqns) + in + (t, eqns) + end + fun gen_primrec prep_spec raw_fixes raw_spec lthy = let val (fixes, spec) = fst (prep_spec raw_fixes raw_spec lthy); @@ -277,7 +286,8 @@ lthy |> add_primrec_simple fixes (map snd spec) |-> (fn (prefix, simps) => fold_map Local_Theory.note (attr_bindings prefix ~~ simps) - #-> (fn simps' => Local_Theory.note (simp_attr_binding prefix, maps snd simps'))) + #-> (fn simps' => Local_Theory.note (simp_attr_binding prefix, maps snd simps') + ##> (Spec_Rules.add Spec_Rules.Equational (specs_of simps')))) |>> snd end; diff -r 1dfb1df1d9d0 -r bd7e347eb768 src/HOL/Tools/recdef.ML --- a/src/HOL/Tools/recdef.ML Wed Jan 20 14:09:46 2010 +0100 +++ b/src/HOL/Tools/recdef.ML Wed Jan 20 18:02:22 2010 +0100 @@ -204,13 +204,18 @@ val simp_att = if null tcs then [Simplifier.simp_add, Nitpick_Simps.add, Code.add_default_eqn_attribute] else []; - + fun specs_of simps = + let + fun dest_eqn th = fst (strip_comb (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of th))))) + val ts = distinct (op =) (map dest_eqn simps) + in (ts, simps) end val ((simps' :: rules', [induct']), thy) = thy |> Sign.add_path bname |> PureThy.add_thmss (((Binding.name "simps", flat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts)) - ||>> PureThy.add_thms [((Binding.name "induct", induct), [])]; + ||>> PureThy.add_thms [((Binding.name "induct", induct), [])] + ||> Spec_Rules.add_global Spec_Rules.Equational (specs_of (flat rules)); val result = {simps = simps', rules = rules', induct = induct', tcs = tcs}; val thy = thy