proper handling of hints;
authorwenzelm
Tue, 05 Sep 2000 18:53:21 +0200
changeset 9866 90cbf68b9227
parent 9865 9a39eabfa17b
child 9867 bf8300fa4238
proper handling of hints; tuned;
TFL/post.sml
TFL/tfl.sml
--- a/TFL/post.sml	Tue Sep 05 18:52:29 2000 +0200
+++ b/TFL/post.sml	Tue Sep 05 18:53:21 2000 +0200
@@ -1,223 +1,201 @@
-(*  Title:      TFL/post
+(*  Title:      TFL/post.sml
     ID:         $Id$
     Author:     Konrad Slind, Cambridge University Computer Laboratory
     Copyright   1997  University of Cambridge
 
-Postprocessing of TFL definitions
+Second part of main module (postprocessing of TFL definitions).
 *)
 
-signature TFL = 
-  sig
-
-   val trace : bool ref
-
-   structure Prim : TFL_sig
-   val quiet_mode : bool ref
-   val message : string -> unit
-
-   val tgoalw : theory -> thm list -> thm list -> thm list
-   val tgoal: theory -> thm list -> thm list
-
-   val std_postprocessor : simpset -> theory 
-                           -> {induction:thm, rules:thm, TCs:term list list} 
-                           -> {induction:thm, rules:thm, nested_tcs:thm list}
-
-   val define_i : theory -> xstring -> term 
-                  -> simpset * thm list (*allows special simplication*)
-                  -> term list
-                  -> theory * {rules:(thm*int)list, induct:thm, tcs:term list}
-
-   val define   : theory -> xstring -> string 
-                  -> simpset * thm list 
-                  -> string list 
-                  -> theory * {rules:(thm*int)list, induct:thm, tcs:term list}
-
-   val defer_i : theory -> xstring
-                  -> thm list (* congruence rules *)
-                  -> term list -> theory * thm
-
-   val defer : theory -> xstring
-                  -> thm list 
-                  -> string list -> theory * thm
-
-  end;
-
+signature TFL =
+sig
+  val trace: bool ref
+  val quiet_mode: bool ref
+  val message: string -> unit
+  val tgoalw: theory -> thm list -> thm list -> thm list
+  val tgoal: theory -> thm list -> thm list
+  val std_postprocessor: claset -> simpset -> thm list -> theory ->
+    {induction: thm, rules: thm, TCs: term list list} ->
+    {induction: thm, rules: thm, nested_tcs: thm list}
+  val define_i: theory -> claset -> simpset -> thm list -> thm list -> xstring ->
+    term -> term list -> theory * {rules: (thm * int) list, induct: thm, tcs: term list}
+  val define: theory -> claset -> simpset -> thm list -> thm list -> xstring ->
+    string -> string list -> theory * {rules: (thm * int) list, induct: thm, tcs: term list}
+  val defer_i: theory -> thm list -> xstring -> term list -> theory * thm
+  val defer: theory -> thm list -> xstring -> string list -> theory * thm
+end;
 
 structure Tfl: TFL =
 struct
- structure Prim = Prim
- structure S = USyntax
 
- fun read_term thy = Sign.simple_read_term (Theory.sign_of thy) HOLogic.termT;
+structure S = USyntax
 
 
- (* messages *)
+(* messages *)
+
+val trace = Prim.trace
 
- val quiet_mode = ref false;
- fun message s = if ! quiet_mode then () else writeln s;
+val quiet_mode = ref false;
+fun message s = if ! quiet_mode then () else writeln s;
+
 
- val trace = Prim.trace
+(* misc *)
+
+fun read_term thy = Sign.simple_read_term (Theory.sign_of thy) HOLogic.termT;
 
 
- (*---------------------------------------------------------------------------
-  * Extract termination goals so that they can be put it into a goalstack, or 
-  * have a tactic directly applied to them.
-  *--------------------------------------------------------------------------*)
- fun termination_goals rules = 
-     map (#1 o Type.freeze_thaw o HOLogic.dest_Trueprop)
-       (foldr (fn (th,A) => union_term (prems_of th, A)) (rules, []));
+(*---------------------------------------------------------------------------
+ * Extract termination goals so that they can be put it into a goalstack, or
+ * have a tactic directly applied to them.
+ *--------------------------------------------------------------------------*)
+fun termination_goals rules =
+    map (#1 o Type.freeze_thaw o HOLogic.dest_Trueprop)
+      (foldr (fn (th,A) => union_term (prems_of th, A)) (rules, []));
 
- (*---------------------------------------------------------------------------
-  * Finds the termination conditions in (highly massaged) definition and 
-  * puts them into a goalstack.
-  *--------------------------------------------------------------------------*)
- fun tgoalw thy defs rules = 
-   case termination_goals rules of
-       [] => error "tgoalw: no termination conditions to prove"
-     | L  => goalw_cterm defs 
-	       (Thm.cterm_of (Theory.sign_of thy) 
-			 (HOLogic.mk_Trueprop(USyntax.list_mk_conj L)));
+(*---------------------------------------------------------------------------
+ * Finds the termination conditions in (highly massaged) definition and
+ * puts them into a goalstack.
+ *--------------------------------------------------------------------------*)
+fun tgoalw thy defs rules =
+  case termination_goals rules of
+      [] => error "tgoalw: no termination conditions to prove"
+    | L  => goalw_cterm defs
+              (Thm.cterm_of (Theory.sign_of thy)
+                        (HOLogic.mk_Trueprop(USyntax.list_mk_conj L)));
 
- fun tgoal thy = tgoalw thy [];
+fun tgoal thy = tgoalw thy [];
 
- (*---------------------------------------------------------------------------
+(*---------------------------------------------------------------------------
  * Three postprocessors are applied to the definition.  It
  * attempts to prove wellfoundedness of the given relation, simplifies the
- * non-proved termination conditions, and finally attempts to prove the 
+ * non-proved termination conditions, and finally attempts to prove the
  * simplified termination conditions.
  *--------------------------------------------------------------------------*)
- fun std_postprocessor ss =
-   Prim.postprocess
-    {WFtac      = REPEAT (ares_tac [wf_empty, wf_pred_nat, 
-				    wf_measure, wf_inv_image, 
-				    wf_lex_prod, wf_less_than, wf_trancl] 1),
-     terminator = asm_simp_tac ss 1
-		  THEN TRY(CLASET' (fn cs => fast_tac
-			   (cs addSDs [not0_implies_Suc] addss ss)) 1),
-     simplifier = Rules.simpl_conv ss []};
+fun std_postprocessor cs ss wfs =
+  Prim.postprocess
+   {wf_tac     = REPEAT (ares_tac wfs 1),
+    terminator = asm_simp_tac ss 1
+                 THEN TRY (fast_tac (cs addSDs [not0_implies_Suc] addss ss) 1),
+    simplifier = Rules.simpl_conv ss []};
 
 
 
- val concl = #2 o Rules.dest_thm;
+val concl = #2 o Rules.dest_thm;
 
 (*---------------------------------------------------------------------------
- * Postprocess a definition made by "define". This is a separate stage of 
+ * Postprocess a definition made by "define". This is a separate stage of
  * processing from the definition stage.
  *---------------------------------------------------------------------------*)
- local 
- structure R = Rules
- structure U = Utils
+local
+structure R = Rules
+structure U = Utils
 
- (* The rest of these local definitions are for the tricky nested case *)
- val solved = not o U.can S.dest_eq o #2 o S.strip_forall o concl
+(* The rest of these local definitions are for the tricky nested case *)
+val solved = not o U.can S.dest_eq o #2 o S.strip_forall o concl
 
- fun id_thm th = 
-    let val {lhs,rhs} = S.dest_eq(#2(S.strip_forall(#2 (R.dest_thm th))))
-    in lhs aconv rhs
-    end handle _ => false
+fun id_thm th =
+   let val {lhs,rhs} = S.dest_eq(#2(S.strip_forall(#2 (R.dest_thm th))))
+   in lhs aconv rhs
+   end handle _ => false
 
- fun prover s = prove_goal HOL.thy s (fn _ => [fast_tac HOL_cs 1]);
- val P_imp_P_iff_True = prover "P --> (P= True)" RS mp;
- val P_imp_P_eq_True = P_imp_P_iff_True RS eq_reflection;
- fun mk_meta_eq r = case concl_of r of
-      Const("==",_)$_$_ => r
-   |   _ $(Const("op =",_)$_$_) => r RS eq_reflection
-   |   _ => r RS P_imp_P_eq_True
+fun prover s = prove_goal HOL.thy s (fn _ => [fast_tac HOL_cs 1]);
+val P_imp_P_iff_True = prover "P --> (P= True)" RS mp;
+val P_imp_P_eq_True = P_imp_P_iff_True RS eq_reflection;
+fun mk_meta_eq r = case concl_of r of
+     Const("==",_)$_$_ => r
+  |   _ $(Const("op =",_)$_$_) => r RS eq_reflection
+  |   _ => r RS P_imp_P_eq_True
 
- (*Is this the best way to invoke the simplifier??*)
- fun rewrite L = rewrite_rule (map mk_meta_eq (filter(not o id_thm) L))
+(*Is this the best way to invoke the simplifier??*)
+fun rewrite L = rewrite_rule (map mk_meta_eq (filter(not o id_thm) L))
 
- fun join_assums th = 
-   let val {sign,...} = rep_thm th
-       val tych = cterm_of sign
-       val {lhs,rhs} = S.dest_eq(#2 (S.strip_forall (concl th)))
-       val cntxtl = (#1 o S.strip_imp) lhs  (* cntxtl should = cntxtr *)
-       val cntxtr = (#1 o S.strip_imp) rhs  (* but union is solider *)
-       val cntxt = gen_union (op aconv) (cntxtl, cntxtr)
-   in 
-     R.GEN_ALL 
-       (R.DISCH_ALL 
-	  (rewrite (map (R.ASSUME o tych) cntxt) (R.SPEC_ALL th)))
-   end
-   val gen_all = S.gen_all
- in
- fun proof_stage ss theory {f, R, rules, full_pats_TCs, TCs} =
-   let val dummy = message "Proving induction theorem ..."
-       val ind = Prim.mk_induction theory 
-                    {fconst=f, R=R, SV=[], pat_TCs_list=full_pats_TCs}
-       val dummy = message "Postprocessing ...";
-       val {rules, induction, nested_tcs} = 
-	   std_postprocessor ss theory {rules=rules, induction=ind, TCs=TCs}
-   in
-   case nested_tcs
-   of [] => {induction=induction, rules=rules,tcs=[]}
-   | L  => let val dummy = message "Simplifying nested TCs ..."
-	       val (solved,simplified,stubborn) =
-		U.itlist (fn th => fn (So,Si,St) =>
-		      if (id_thm th) then (So, Si, th::St) else
-		      if (solved th) then (th::So, Si, St) 
-		      else (So, th::Si, St)) nested_tcs ([],[],[])
-	       val simplified' = map join_assums simplified
-	       val rewr = full_simplify (ss addsimps (solved @ simplified'));
-	       val induction' = rewr induction
-	       and rules'     = rewr rules
-	   in
-	   {induction = induction',
-		rules = rules',
-		  tcs = map (gen_all o S.rhs o #2 o S.strip_forall o concl)
-			    (simplified@stubborn)}
-	   end
-   end;
+fun join_assums th =
+  let val {sign,...} = rep_thm th
+      val tych = cterm_of sign
+      val {lhs,rhs} = S.dest_eq(#2 (S.strip_forall (concl th)))
+      val cntxtl = (#1 o S.strip_imp) lhs  (* cntxtl should = cntxtr *)
+      val cntxtr = (#1 o S.strip_imp) rhs  (* but union is solider *)
+      val cntxt = gen_union (op aconv) (cntxtl, cntxtr)
+  in
+    R.GEN_ALL
+      (R.DISCH_ALL
+         (rewrite (map (R.ASSUME o tych) cntxt) (R.SPEC_ALL th)))
+  end
+  val gen_all = S.gen_all
+in
+fun proof_stage cs ss wfs theory {f, R, rules, full_pats_TCs, TCs} =
+  let
+    val _ = message "Proving induction theorem ..."
+    val ind = Prim.mk_induction theory {fconst=f, R=R, SV=[], pat_TCs_list=full_pats_TCs}
+    val _ = message "Postprocessing ...";
+    val {rules, induction, nested_tcs} =
+      std_postprocessor cs ss wfs theory {rules=rules, induction=ind, TCs=TCs}
+  in
+  case nested_tcs
+  of [] => {induction=induction, rules=rules,tcs=[]}
+  | L  => let val dummy = message "Simplifying nested TCs ..."
+              val (solved,simplified,stubborn) =
+               U.itlist (fn th => fn (So,Si,St) =>
+                     if (id_thm th) then (So, Si, th::St) else
+                     if (solved th) then (th::So, Si, St)
+                     else (So, th::Si, St)) nested_tcs ([],[],[])
+              val simplified' = map join_assums simplified
+              val rewr = full_simplify (ss addsimps (solved @ simplified'));
+              val induction' = rewr induction
+              and rules'     = rewr rules
+          in
+          {induction = induction',
+               rules = rules',
+                 tcs = map (gen_all o S.rhs o #2 o S.strip_forall o concl)
+                           (simplified@stubborn)}
+          end
+  end;
 
 
- (*lcp: curry the predicate of the induction rule*)
- fun curry_rule rl = split_rule_var
-			 (head_of (HOLogic.dest_Trueprop (concl_of rl)), 
-			  rl);
+(*lcp: curry the predicate of the induction rule*)
+fun curry_rule rl = split_rule_var
+                        (head_of (HOLogic.dest_Trueprop (concl_of rl)),
+                         rl);
 
- (*lcp: put a theorem into Isabelle form, using meta-level connectives*)
- val meta_outer = 
-     curry_rule o standard o 
-     rule_by_tactic (REPEAT 
-		     (FIRSTGOAL (resolve_tac [allI, impI, conjI]
-				 ORELSE' etac conjE)));
+(*lcp: put a theorem into Isabelle form, using meta-level connectives*)
+val meta_outer =
+    curry_rule o standard o
+    rule_by_tactic (REPEAT
+                    (FIRSTGOAL (resolve_tac [allI, impI, conjI]
+                                ORELSE' etac conjE)));
 
- (*Strip off the outer !P*)
- val spec'= read_instantiate [("x","P::?'b=>bool")] spec;
+(*Strip off the outer !P*)
+val spec'= read_instantiate [("x","P::?'b=>bool")] spec;
 
- fun simplify_defn thy (ss, tflCongs) id pats def0 =
-    let	val def = freezeT def0 RS meta_eq_to_obj_eq
-	val {theory,rules,rows,TCs,full_pats_TCs} = 
-		 Prim.post_definition (Prim.congs thy tflCongs)
-		    (thy, (def,pats))
-	val {lhs=f,rhs} = S.dest_eq(concl def)
-	val (_,[R,_]) = S.strip_comb rhs
-	val ss' = ss addsimps Prim.default_simps
-	val {induction, rules, tcs} = 
-	      proof_stage ss' theory 
-		{f = f, R = R, rules = rules,
-		 full_pats_TCs = full_pats_TCs,
-		 TCs = TCs}
-	val rules' = map (standard o rulify) (R.CONJUNCTS rules)
-    in  {induct = meta_outer (rulify (induction RS spec')), 
-	 rules = ListPair.zip(rules', rows),
-	 tcs = (termination_goals rules') @ tcs}
-    end
-   handle Utils.ERR {mesg,func,module} => 
-		error (mesg ^ 
-		       "\n    (In TFL function " ^ module ^ "." ^ func ^ ")");
+fun simplify_defn thy cs ss congs wfs id pats def0 =
+   let val def = freezeT def0 RS meta_eq_to_obj_eq
+       val {theory,rules,rows,TCs,full_pats_TCs} = Prim.post_definition congs (thy, (def,pats))
+       val {lhs=f,rhs} = S.dest_eq (concl def)
+       val (_,[R,_]) = S.strip_comb rhs
+       val {induction, rules, tcs} =
+             proof_stage cs ss wfs theory
+               {f = f, R = R, rules = rules,
+                full_pats_TCs = full_pats_TCs,
+                TCs = TCs}
+       val rules' = map (standard o rulify) (R.CONJUNCTS rules)
+   in  {induct = meta_outer (rulify (induction RS spec')),
+        rules = ListPair.zip(rules', rows),
+        tcs = (termination_goals rules') @ tcs}
+   end
+  handle Utils.ERR {mesg,func,module} =>
+               error (mesg ^
+                      "\n    (In TFL function " ^ module ^ "." ^ func ^ ")");
 
 (*---------------------------------------------------------------------------
- * Defining a function with an associated termination relation. 
+ * Defining a function with an associated termination relation.
  *---------------------------------------------------------------------------*)
- fun define_i thy fid R ss_congs eqs = 
-   let val {functional,pats} = Prim.mk_functional thy eqs
-       val (thy, def) = Prim.wfrec_definition0 thy (Sign.base_name fid) R functional
-   in (thy, simplify_defn thy ss_congs fid pats def) end;
+fun define_i thy cs ss congs wfs fid R eqs =
+  let val {functional,pats} = Prim.mk_functional thy eqs
+      val (thy, def) = Prim.wfrec_definition0 thy (Sign.base_name fid) R functional
+  in (thy, simplify_defn thy cs ss congs wfs fid pats def) end;
 
- fun define thy fid R ss_congs seqs = 
-   define_i thy fid (read_term thy R) ss_congs (map (read_term thy) seqs)
-     handle Utils.ERR {mesg,...} => error mesg;
+fun define thy cs ss congs wfs fid R seqs =
+  define_i thy cs ss congs wfs fid (read_term thy R) (map (read_term thy) seqs)
+    handle Utils.ERR {mesg,...} => error mesg;
 
 
 (*---------------------------------------------------------------------------
@@ -226,28 +204,28 @@
  *
  *---------------------------------------------------------------------------*)
 
- local open USyntax
- in 
- fun func_of_cond_eqn tm =
-   #1(strip_comb(#lhs(dest_eq(#2 (strip_forall(#2(strip_imp tm)))))))
- end;
+local open USyntax
+in
+fun func_of_cond_eqn tm =
+  #1(strip_comb(#lhs(dest_eq(#2 (strip_forall(#2(strip_imp tm)))))))
+end;
 
- fun defer_i thy fid congs eqs = 
-  let val {rules,R,theory,full_pats_TCs,SV,...} = 
-	      Prim.lazyR_def thy (Sign.base_name fid) congs eqs
-      val f = func_of_cond_eqn (concl(R.CONJUNCT1 rules handle _ => rules))
-      val dummy = message "Proving induction theorem ...";
-      val induction = Prim.mk_induction theory 
-                         {fconst=f, R=R, SV=SV, pat_TCs_list=full_pats_TCs}
-  in (theory, 
-      (*return the conjoined induction rule and recursion equations, 
-	with assumptions remaining to discharge*)
-      standard (induction RS (rules RS conjI)))
-  end
+fun defer_i thy congs fid eqs =
+ let val {rules,R,theory,full_pats_TCs,SV,...} =
+             Prim.lazyR_def thy (Sign.base_name fid) congs eqs
+     val f = func_of_cond_eqn (concl(R.CONJUNCT1 rules handle _ => rules))
+     val dummy = message "Proving induction theorem ...";
+     val induction = Prim.mk_induction theory
+                        {fconst=f, R=R, SV=SV, pat_TCs_list=full_pats_TCs}
+ in (theory,
+     (*return the conjoined induction rule and recursion equations,
+       with assumptions remaining to discharge*)
+     standard (induction RS (rules RS conjI)))
+ end
 
- fun defer thy fid congs seqs = 
-   defer_i thy fid congs (map (read_term thy) seqs)
-     handle Utils.ERR {mesg,...} => error mesg;
- end;
+fun defer thy congs fid seqs =
+  defer_i thy congs fid (map (read_term thy) seqs)
+    handle Utils.ERR {mesg,...} => error mesg;
+end;
 
 end;
--- a/TFL/tfl.sml	Tue Sep 05 18:52:29 2000 +0200
+++ b/TFL/tfl.sml	Tue Sep 05 18:53:21 2000 +0200
@@ -1,29 +1,62 @@
-(*  Title:      TFL/tfl
+(*  Title:      TFL/tfl.sml
     ID:         $Id$
     Author:     Konrad Slind, Cambridge University Computer Laboratory
     Copyright   1997  University of Cambridge
 
-Main module
+First part of main module.
 *)
 
-structure Prim : TFL_sig =
+signature TFL_sig =
+sig
+  val trace: bool ref
+  type pattern
+  val mk_functional: theory -> term list -> {functional: term, pats: pattern list}
+  val wfrec_definition0: theory -> string -> term -> term -> theory * thm
+  val post_definition: thm list -> theory * (thm * pattern list) ->
+   {theory: theory,
+    rules: thm,
+    rows: int list,
+    TCs: term list list,
+    full_pats_TCs: (term * term list) list}
+  val wfrec_eqns: theory -> xstring -> thm list -> term list ->
+   {WFR: term,
+    SV: term list,
+    proto_def: term,
+    extracta: (thm * term list) list,
+    pats: pattern list}
+  val lazyR_def: theory -> xstring -> thm list -> term list ->
+   {theory: theory,
+    rules: thm,
+    R: term,
+    SV: term list,
+    full_pats_TCs: (term * term list) list,
+    patterns : pattern list}
+  val mk_induction: theory ->
+    {fconst: term, R: term, SV: term list, pat_TCs_list: (term * term list) list} -> thm
+  val postprocess: {wf_tac: tactic, terminator: tactic, simplifier: cterm -> thm} -> theory ->
+    {rules: thm, induction: thm, TCs: term list list} ->
+    {rules: thm, induction: thm, nested_tcs: thm list}
+end;
+
+
+structure Prim: TFL_sig =
 struct
 
 val trace = ref false;
 
-open BasisLibrary; (*restore original structures*)
+open BasisLibrary;
 
-(* Abbreviations *)
 structure R = Rules;
 structure S = USyntax;
 structure U = S.Utils;
 
-fun TFL_ERR{func,mesg} = U.ERR{module = "Tfl", func = func, mesg = mesg};
+
+fun TFL_ERR {func, mesg} = U.ERR {module = "Tfl", func = func, mesg = mesg};
 
 val concl = #2 o R.dest_thm;
 val hyp = #1 o R.dest_thm;
 
-val list_mk_type = U.end_itlist (curry(op -->));
+val list_mk_type = U.end_itlist (curry (op -->));
 
 fun enumerate xs = ListPair.zip(xs, 0 upto (length xs - 1));
 
@@ -36,79 +69,6 @@
      end;
 
 
-
-(*---------------------------------------------------------------------------
-          handling of user-supplied congruence rules: tnn *)
-
-(*Convert conclusion from = to ==*)
-val eq_reflect_list = map (fn th => (th RS eq_reflection) handle _ => th);
-
-val cong_hd = fst o dest_Const o head_of o fst o Logic.dest_equals o concl_of;
-
-fun add_cong(congs,thm) =
-  let val c = cong_hd thm
-  in overwrite_warn (congs,(c,thm))
-       ("Overwriting congruence rule for " ^ quote c)
-  end
-
-fun del_cong(congs,thm) =
-  let val c = cong_hd thm
-      val (del, rest) = partition (fn (n, _) => n = c) congs
-  in if null del
-     then (warning ("No congruence rule for " ^ quote c ^ " present"); congs)
-     else rest
-  end;
-
-fun add_congs(congs,thms) = foldl add_cong (congs, eq_reflect_list thms);
-fun del_congs(congs,thms) = foldl del_cong (congs, eq_reflect_list thms);
-
-(** recdef data **)
-
-(* theory data kind 'Provers/simpset' *)
-
-structure RecdefCongsArgs =
-struct
-  val name = "HOL/recdef-congs";
-  type T = (string * thm) list ref;
-
-  val empty = ref(add_congs([], [Thms.LET_CONG, if_cong]));
-  fun copy (ref congs) = (ref congs): T;            (*create new reference!*)
-  val prep_ext = copy;
-  fun merge (ref congs1, ref congs2) = ref (merge_alists congs1 congs2);
-  fun print _ (ref congs) = print_thms(map snd congs);
-end;
-
-structure RecdefCongs = TheoryDataFun(RecdefCongsArgs);
-val print_recdef_congs = RecdefCongs.print;
-val recdef_congs_ref_of_sg = RecdefCongs.get_sg;
-val recdef_congs_ref_of = RecdefCongs.get;
-val init = RecdefCongs.init;
-
-
-(* access global recdef_congs *)
-val recdef_congs_of_sg = ! o recdef_congs_ref_of_sg;
-val recdef_congs_of = recdef_congs_of_sg o Theory.sign_of;
-
-(* change global recdef_congs *)
-
-fun change_recdef_congs_of f x thy =
-  let val r = recdef_congs_ref_of thy
-  in r := f (! r, x); thy end;
-
-fun change_recdef_congs f x =
- (change_recdef_congs_of f x (Context.the_context ()); ());
-
-val Add_recdef_congs = change_recdef_congs add_congs;
-val Del_recdef_congs = change_recdef_congs del_congs;
-
-
-fun congs thy ths = map snd (recdef_congs_of thy) @ eq_reflect_list ths;
-
-val default_simps =
-    [less_Suc_eq RS iffD2, lex_prod_def, measure_def, inv_image_def];
-
-
-
 (*---------------------------------------------------------------------------
  * The next function is common to pattern-match translation and
  * proof of completeness of cases for the induction theorem.
@@ -477,7 +437,6 @@
 
 fun givens pats = map pat_of (filter given pats);
 
-(*called only by Tfl.simplify_defn*)
 fun post_definition meta_tflCongs (theory, (def, pats)) =
  let val tych = Thry.typecheck theory
      val f = #lhs(S.dest_eq(concl def))
@@ -569,7 +528,7 @@
 
 fun lazyR_def thy fid tflCongs eqns =
  let val {proto_def,WFR,pats,extracta,SV} =
-           wfrec_eqns thy fid (congs thy tflCongs) eqns
+           wfrec_eqns thy fid tflCongs eqns
      val R1 = S.rand WFR
      val f = #lhs(S.dest_eq proto_def)
      val (extractants,TCl) = ListPair.unzip extracta
@@ -952,7 +911,7 @@
    (R.MP rule tcthm, R.PROVE_HYP tcthm induction)
 
 
-fun postprocess{WFtac, terminator, simplifier} theory {rules,induction,TCs} =
+fun postprocess{wf_tac, terminator, simplifier} theory {rules,induction,TCs} =
 let val tych = Thry.typecheck theory
 
    (*---------------------------------------------------------------------
@@ -961,7 +920,7 @@
    val (rules1,induction1)  =
        let val thm = R.prove(tych(HOLogic.mk_Trueprop
                                   (hd(#1(R.dest_thm rules)))),
-                             WFtac)
+                             wf_tac)
        in (R.PROVE_HYP thm rules,  R.PROVE_HYP thm induction)
        end handle _ => (rules,induction)
 
@@ -1044,7 +1003,4 @@
   {induction = ind2, rules = R.LIST_CONJ rules2, nested_tcs = extras}
 end;
 
-end; (* TFL *)
-
-val Add_recdef_congs = Prim.Add_recdef_congs;
-val Del_recdef_congs = Prim.Del_recdef_congs;
+end;