src/HOL/Tools/recdef.ML
changeset 31723 f5cafe803b55
parent 31243 4c34331a88f9
child 31902 862ae16a799d
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/recdef.ML	Fri Jun 19 17:23:21 2009 +0200
@@ -0,0 +1,331 @@
+(*  Title:      HOL/Tools/recdef.ML
+    Author:     Markus Wenzel, TU Muenchen
+
+Wrapper module for Konrad Slind's TFL package.
+*)
+
+signature RECDEF =
+sig
+  val get_recdef: theory -> string
+    -> {simps: thm list, rules: thm list list, induct: thm, tcs: term list} option
+  val get_hints: Proof.context -> {simps: thm list, congs: (string * thm) list, wfs: thm list}
+  val simp_add: attribute
+  val simp_del: attribute
+  val cong_add: attribute
+  val cong_del: attribute
+  val wf_add: attribute
+  val wf_del: attribute
+  val add_recdef: bool -> xstring -> string -> ((binding * string) * Attrib.src list) list ->
+    Attrib.src option -> theory -> theory
+      * {simps: thm list, rules: thm list list, induct: thm, tcs: term list}
+  val add_recdef_i: bool -> xstring -> term -> ((binding * term) * attribute list) list ->
+    theory -> theory * {simps: thm list, rules: thm list list, induct: thm, tcs: term list}
+  val defer_recdef: xstring -> string list -> (Facts.ref * Attrib.src list) list
+    -> theory -> theory * {induct_rules: thm}
+  val defer_recdef_i: xstring -> term list -> thm list -> theory -> theory * {induct_rules: thm}
+  val recdef_tc: bstring * Attrib.src list -> xstring -> int option -> bool ->
+    local_theory -> Proof.state
+  val recdef_tc_i: bstring * Attrib.src list -> string -> int option -> bool ->
+    local_theory -> Proof.state
+  val setup: theory -> theory
+end;
+
+structure Recdef: RECDEF =
+struct
+
+
+(** recdef hints **)
+
+(* type hints *)
+
+type hints = {simps: thm list, congs: (string * thm) list, wfs: thm list};
+
+fun mk_hints (simps, congs, wfs) = {simps = simps, congs = congs, wfs = wfs}: hints;
+fun map_hints f ({simps, congs, wfs}: hints) = mk_hints (f (simps, congs, wfs));
+
+fun map_simps f = map_hints (fn (simps, congs, wfs) => (f simps, congs, wfs));
+fun map_congs f = map_hints (fn (simps, congs, wfs) => (simps, f congs, wfs));
+fun map_wfs f = map_hints (fn (simps, congs, wfs) => (simps, congs, f wfs));
+
+fun pretty_hints ({simps, congs, wfs}: hints) =
+ [Pretty.big_list "recdef simp hints:" (map Display.pretty_thm simps),
+  Pretty.big_list "recdef cong hints:" (map Display.pretty_thm (map snd congs)),
+  Pretty.big_list "recdef wf hints:" (map Display.pretty_thm wfs)];
+
+
+(* congruence rules *)
+
+local
+
+val cong_head =
+  fst o Term.dest_Const o Term.head_of o fst o Logic.dest_equals o Thm.concl_of;
+
+fun prep_cong raw_thm =
+  let val thm = safe_mk_meta_eq raw_thm in (cong_head thm, thm) end;
+
+in
+
+fun add_cong raw_thm congs =
+  let
+    val (c, thm) = prep_cong raw_thm;
+    val _ = if AList.defined (op =) congs c
+      then warning ("Overwriting recdef congruence rule for " ^ quote c)
+      else ();
+  in AList.update (op =) (c, thm) congs end;
+
+fun del_cong raw_thm congs =
+  let
+    val (c, thm) = prep_cong raw_thm;
+    val _ = if AList.defined (op =) congs c
+      then ()
+      else warning ("No recdef congruence rule for " ^ quote c);
+  in AList.delete (op =) c congs end;
+
+end;
+
+
+
+(** global and local recdef data **)
+
+(* theory data *)
+
+type recdef_info = {simps: thm list, rules: thm list list, induct: thm, tcs: term list};
+
+structure GlobalRecdefData = TheoryDataFun
+(
+  type T = recdef_info Symtab.table * hints;
+  val empty = (Symtab.empty, mk_hints ([], [], [])): T;
+  val copy = I;
+  val extend = I;
+  fun merge _
+   ((tab1, {simps = simps1, congs = congs1, wfs = wfs1}),
+    (tab2, {simps = simps2, congs = congs2, wfs = wfs2})) : T =
+      (Symtab.merge (K true) (tab1, tab2),
+        mk_hints (Thm.merge_thms (simps1, simps2),
+          AList.merge (op =) Thm.eq_thm (congs1, congs2),
+          Thm.merge_thms (wfs1, wfs2)));
+);
+
+val get_recdef = Symtab.lookup o #1 o GlobalRecdefData.get;
+
+fun put_recdef name info thy =
+  let
+    val (tab, hints) = GlobalRecdefData.get thy;
+    val tab' = Symtab.update_new (name, info) tab
+      handle Symtab.DUP _ => error ("Duplicate recursive function definition " ^ quote name);
+  in GlobalRecdefData.put (tab', hints) thy end;
+
+val get_global_hints = #2 o GlobalRecdefData.get;
+
+
+(* proof data *)
+
+structure LocalRecdefData = ProofDataFun
+(
+  type T = hints;
+  val init = get_global_hints;
+);
+
+val get_hints = LocalRecdefData.get;
+fun map_hints f = Context.mapping (GlobalRecdefData.map (apsnd f)) (LocalRecdefData.map f);
+
+
+(* attributes *)
+
+fun attrib f = Thm.declaration_attribute (map_hints o f);
+
+val simp_add = attrib (map_simps o Thm.add_thm);
+val simp_del = attrib (map_simps o Thm.del_thm);
+val cong_add = attrib (map_congs o add_cong);
+val cong_del = attrib (map_congs o del_cong);
+val wf_add = attrib (map_wfs o Thm.add_thm);
+val wf_del = attrib (map_wfs o Thm.del_thm);
+
+
+(* modifiers *)
+
+val recdef_simpN = "recdef_simp";
+val recdef_congN = "recdef_cong";
+val recdef_wfN = "recdef_wf";
+
+val recdef_modifiers =
+ [Args.$$$ recdef_simpN -- Args.colon >> K ((I, simp_add): Method.modifier),
+  Args.$$$ recdef_simpN -- Args.add -- Args.colon >> K (I, simp_add),
+  Args.$$$ recdef_simpN -- Args.del -- Args.colon >> K (I, simp_del),
+  Args.$$$ recdef_congN -- Args.colon >> K (I, cong_add),
+  Args.$$$ recdef_congN -- Args.add -- Args.colon >> K (I, cong_add),
+  Args.$$$ recdef_congN -- Args.del -- Args.colon >> K (I, cong_del),
+  Args.$$$ recdef_wfN -- Args.colon >> K (I, wf_add),
+  Args.$$$ recdef_wfN -- Args.add -- Args.colon >> K (I, wf_add),
+  Args.$$$ recdef_wfN -- Args.del -- Args.colon >> K (I, wf_del)] @
+  Clasimp.clasimp_modifiers;
+
+
+
+(** prepare_hints(_i) **)
+
+fun prepare_hints thy opt_src =
+  let
+    val ctxt0 = ProofContext.init thy;
+    val ctxt =
+      (case opt_src of
+        NONE => ctxt0
+      | SOME src => #2 (Method.syntax (Method.sections recdef_modifiers) src ctxt0));
+    val {simps, congs, wfs} = get_hints ctxt;
+    val cs = local_claset_of ctxt;
+    val ss = local_simpset_of ctxt addsimps simps;
+  in (cs, ss, rev (map snd congs), wfs) end;
+
+fun prepare_hints_i thy () =
+  let
+    val ctxt0 = ProofContext.init thy;
+    val {simps, congs, wfs} = get_global_hints thy;
+  in (local_claset_of ctxt0, local_simpset_of ctxt0 addsimps simps, rev (map snd congs), wfs) end;
+
+
+
+(** add_recdef(_i) **)
+
+fun requires_recdef thy = Theory.requires thy "Recdef" "recursive functions";
+
+fun gen_add_recdef tfl_fn prep_att prep_hints not_permissive raw_name R eq_srcs hints thy =
+  let
+    val _ = requires_recdef thy;
+
+    val name = Sign.intern_const thy raw_name;
+    val bname = Long_Name.base_name name;
+    val _ = writeln ("Defining recursive function " ^ quote name ^ " ...");
+
+    val ((eq_names, eqs), raw_eq_atts) = apfst split_list (split_list eq_srcs);
+    val eq_atts = map (map (prep_att thy)) raw_eq_atts;
+
+    val (cs, ss, congs, wfs) = prep_hints thy hints;
+    (*We must remove imp_cong to prevent looping when the induction rule
+      is simplified. Many induction rules have nested implications that would
+      give rise to looping conditional rewriting.*)
+    val (thy, {rules = rules_idx, induct, tcs}) =
+        tfl_fn not_permissive thy cs (ss delcongs [imp_cong])
+               congs wfs name R eqs;
+    val rules = (map o map) fst (partition_eq (eq_snd (op = : int * int -> bool)) rules_idx);
+    val simp_att = if null tcs then [Simplifier.simp_add, Nitpick_Const_Simp_Thms.add,
+      Code.add_default_eqn_attribute, Quickcheck_RecFun_Simp_Thms.add] else [];
+
+    val ((simps' :: rules', [induct']), thy) =
+      thy
+      |> Sign.add_path bname
+      |> PureThy.add_thmss
+        (((Binding.name "simps", List.concat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts))
+      ||>> PureThy.add_thms [((Binding.name "induct", induct), [])];
+    val result = {simps = simps', rules = rules', induct = induct', tcs = tcs};
+    val thy =
+      thy
+      |> put_recdef name result
+      |> Sign.parent_path;
+  in (thy, result) end;
+
+val add_recdef = gen_add_recdef Tfl.define Attrib.attribute prepare_hints;
+fun add_recdef_i x y z w = gen_add_recdef Tfl.define_i (K I) prepare_hints_i x y z w ();
+
+
+
+(** defer_recdef(_i) **)
+
+fun gen_defer_recdef tfl_fn eval_thms raw_name eqs raw_congs thy =
+  let
+    val name = Sign.intern_const thy raw_name;
+    val bname = Long_Name.base_name name;
+
+    val _ = requires_recdef thy;
+    val _ = writeln ("Deferred recursive function " ^ quote name ^ " ...");
+
+    val congs = eval_thms (ProofContext.init thy) raw_congs;
+    val (thy2, induct_rules) = tfl_fn thy congs name eqs;
+    val ([induct_rules'], thy3) =
+      thy2
+      |> Sign.add_path bname
+      |> PureThy.add_thms [((Binding.name "induct_rules", induct_rules), [])]
+      ||> Sign.parent_path;
+  in (thy3, {induct_rules = induct_rules'}) end;
+
+val defer_recdef = gen_defer_recdef Tfl.defer Attrib.eval_thms;
+val defer_recdef_i = gen_defer_recdef Tfl.defer_i (K I);
+
+
+
+(** recdef_tc(_i) **)
+
+fun gen_recdef_tc prep_att prep_name (bname, raw_atts) raw_name opt_i int lthy =
+  let
+    val thy = ProofContext.theory_of lthy;
+    val name = prep_name thy raw_name;
+    val atts = map (prep_att thy) raw_atts;
+    val tcs =
+      (case get_recdef thy name of
+        NONE => error ("No recdef definition of constant: " ^ quote name)
+      | SOME {tcs, ...} => tcs);
+    val i = the_default 1 opt_i;
+    val tc = nth tcs (i - 1) handle Subscript =>
+      error ("No termination condition #" ^ string_of_int i ^
+        " in recdef definition of " ^ quote name);
+  in
+    Specification.theorem Thm.internalK NONE (K I) (Binding.name bname, atts)
+      [] (Element.Shows [(Attrib.empty_binding, [(HOLogic.mk_Trueprop tc, [])])]) int lthy
+  end;
+
+val recdef_tc = gen_recdef_tc Attrib.intern_src Sign.intern_const;
+val recdef_tc_i = gen_recdef_tc (K I) (K I);
+
+
+
+(** package setup **)
+
+(* setup theory *)
+
+val setup =
+  Attrib.setup @{binding recdef_simp} (Attrib.add_del simp_add simp_del)
+    "declaration of recdef simp rule" #>
+  Attrib.setup @{binding recdef_cong} (Attrib.add_del cong_add cong_del)
+    "declaration of recdef cong rule" #>
+  Attrib.setup @{binding recdef_wf} (Attrib.add_del wf_add wf_del)
+    "declaration of recdef wf rule";
+
+
+(* outer syntax *)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val _ = List.app OuterKeyword.keyword ["permissive", "congs", "hints"];
+
+val hints =
+  P.$$$ "(" |-- P.!!! (P.position (P.$$$ "hints" -- Args.parse) --| P.$$$ ")") >> Args.src;
+
+val recdef_decl =
+  Scan.optional (P.$$$ "(" -- P.!!! (P.$$$ "permissive" -- P.$$$ ")") >> K false) true --
+  P.name -- P.term -- Scan.repeat1 (SpecParse.opt_thm_name ":" -- P.prop)
+    -- Scan.option hints
+  >> (fn ((((p, f), R), eqs), src) => #1 o add_recdef p f R (map P.triple_swap eqs) src);
+
+val _ =
+  OuterSyntax.command "recdef" "define general recursive functions (TFL)" K.thy_decl
+    (recdef_decl >> Toplevel.theory);
+
+
+val defer_recdef_decl =
+  P.name -- Scan.repeat1 P.prop --
+  Scan.optional (P.$$$ "(" |-- P.$$$ "congs" |-- P.!!! (SpecParse.xthms1 --| P.$$$ ")")) []
+  >> (fn ((f, eqs), congs) => #1 o defer_recdef f eqs congs);
+
+val _ =
+  OuterSyntax.command "defer_recdef" "defer general recursive functions (TFL)" K.thy_decl
+    (defer_recdef_decl >> Toplevel.theory);
+
+val _ =
+  OuterSyntax.local_theory_to_proof' "recdef_tc" "recommence proof of termination condition (TFL)"
+    K.thy_goal
+    ((SpecParse.opt_thm_name ":" >> apfst Binding.name_of) -- P.xname --
+        Scan.option (P.$$$ "(" |-- P.nat --| P.$$$ ")")
+      >> (fn ((thm_name, name), i) => recdef_tc thm_name name i));
+
+end;
+
+end;