tuned signature;
authorwenzelm
Mon, 01 Aug 2005 19:20:32 +0200
changeset 16978 e35b518bffc9
parent 16977 7c04742abac3
child 16979 4d4d42ea3096
tuned signature;
TFL/casesplit.ML
src/Provers/eqsubst.ML
--- a/TFL/casesplit.ML	Mon Aug 01 19:20:31 2005 +0200
+++ b/TFL/casesplit.ML	Mon Aug 01 19:20:32 2005 +0200
@@ -1,51 +1,49 @@
-(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
+(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *)
 (*  Title:      TFL/casesplit.ML
     Author:     Lucas Dixon, University of Edinburgh
                 lucas.dixon@ed.ac.uk
     Date:       17 Aug 2004
 *)
-(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
+(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *)
 (*  DESCRIPTION:
 
-    A structure that defines a tactic to program case splits. 
+    A structure that defines a tactic to program case splits.
 
     casesplit_free :
-      string * Term.type -> int -> Thm.thm -> Thm.thm Seq.seq
+      string * typ -> int -> thm -> thm Seq.seq
 
-    casesplit_name : 
-      string -> int -> Thm.thm -> Thm.thm Seq.seq
+    casesplit_name :
+      string -> int -> thm -> thm Seq.seq
 
     These use the induction theorem associated with the recursive data
-    type to be split. 
+    type to be split.
 
     The structure includes a function to try and recursively split a
-    conjecture into a list sub-theorems: 
+    conjecture into a list sub-theorems:
 
-    splitto : Thm.thm list -> Thm.thm -> Thm.thm
+    splitto : thm list -> thm -> thm
 *)
-(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
+(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *)
 
 (* logic-specific *)
 signature CASE_SPLIT_DATA =
 sig
-  val dest_Trueprop : Term.term -> Term.term
-  val mk_Trueprop : Term.term -> Term.term
-  val read_cterm : Sign.sg -> string -> Thm.cterm
+  val dest_Trueprop : term -> term
+  val mk_Trueprop : term -> term
 
-  val localize : Thm.thm list
-  val local_impI : Thm.thm
-  val atomize : Thm.thm list
-  val rulify1 : Thm.thm list
-  val rulify2 : Thm.thm list
+  val localize : thm list
+  val local_impI : thm
+  val atomize : thm list
+  val rulify1 : thm list
+  val rulify2 : thm list
 
 end;
 
 (* for HOL *)
-structure CaseSplitData_HOL : CASE_SPLIT_DATA = 
+structure CaseSplitData_HOL : CASE_SPLIT_DATA =
 struct
 val dest_Trueprop = HOLogic.dest_Trueprop;
 val mk_Trueprop = HOLogic.mk_Trueprop;
-val read_cterm = HOLogic.read_cterm;
 
 val localize = [Thm.symmetric (thm "induct_implies_def")];
 val local_impI = thm "induct_impliesI";
@@ -62,29 +60,29 @@
   exception find_split_exp of string
 
   (* getting a case split thm from the induction thm *)
-  val case_thm_of_ty : Sign.sg -> Term.typ -> Thm.thm
-  val cases_thm_of_induct_thm : Thm.thm -> Thm.thm
+  val case_thm_of_ty : theory -> typ -> thm
+  val cases_thm_of_induct_thm : thm -> thm
 
   (* case split tactics *)
   val casesplit_free :
-      string * Term.typ -> int -> Thm.thm -> Thm.thm Seq.seq
-  val casesplit_name : string -> int -> Thm.thm -> Thm.thm Seq.seq
+      string * typ -> int -> thm -> thm Seq.seq
+  val casesplit_name : string -> int -> thm -> thm Seq.seq
 
   (* finding a free var to split *)
   val find_term_split :
-      Term.term * Term.term -> (string * Term.typ) option
+      term * term -> (string * typ) option
   val find_thm_split :
-      Thm.thm -> int -> Thm.thm -> (string * Term.typ) option
+      thm -> int -> thm -> (string * typ) option
   val find_thms_split :
-      Thm.thm list -> int -> Thm.thm -> (string * Term.typ) option
+      thm list -> int -> thm -> (string * typ) option
 
   (* try to recursively split conjectured thm to given list of thms *)
-  val splitto : Thm.thm list -> Thm.thm -> Thm.thm
+  val splitto : thm list -> thm -> thm
 
   (* for use with the recdef package *)
   val derive_init_eqs :
-      Sign.sg ->
-      (Thm.thm * int) list -> Term.term list -> (Thm.thm * int) list
+      theory ->
+      (thm * int) list -> term list -> (thm * int) list
 end;
 
 functor CaseSplitFUN(Data : CASE_SPLIT_DATA) =
@@ -93,7 +91,7 @@
 val rulify_goals = Tactic.rewrite_goals_rule Data.rulify1;
 val atomize_goals = Tactic.rewrite_goals_rule Data.atomize;
 
-(* 
+(*
 val localize = Tactic.norm_hhf_rule o Tactic.simplify false Data.localize;
 fun atomize_term sg =
   ObjectLogic.drop_judgment sg o MetaSimplifier.rewrite_term sg Data.atomize [];
@@ -108,26 +106,26 @@
 *)
 
 (* beta-eta contract the theorem *)
-fun beta_eta_contract thm = 
+fun beta_eta_contract thm =
     let
       val thm2 = equal_elim (Thm.beta_conversion true (Thm.cprop_of thm)) thm
       val thm3 = equal_elim (Thm.eta_conversion (Thm.cprop_of thm2)) thm2
     in thm3 end;
 
 (* make a casethm from an induction thm *)
-val cases_thm_of_induct_thm = 
+val cases_thm_of_induct_thm =
      Seq.hd o (ALLGOALS (fn i => REPEAT (etac Drule.thin_rl i)));
 
 (* get the case_thm (my version) from a type *)
-fun case_thm_of_ty sgn ty  = 
-    let 
+fun case_thm_of_ty sgn ty  =
+    let
       val dtypestab = DatatypePackage.get_datatypes sgn;
-      val ty_str = case ty of 
+      val ty_str = case ty of
                      Type(ty_str, _) => ty_str
-                   | TFree(s,_)  => raise ERROR_MESSAGE 
-                                            ("Free type: " ^ s)   
-                   | TVar((s,i),_) => raise ERROR_MESSAGE 
-                                            ("Free variable: " ^ s)   
+                   | TFree(s,_)  => raise ERROR_MESSAGE
+                                            ("Free type: " ^ s)
+                   | TVar((s,i),_) => raise ERROR_MESSAGE
+                                            ("Free variable: " ^ s)
       val dt = case (Symtab.lookup (dtypestab,ty_str))
                 of SOME dt => dt
                  | NONE => raise ERROR_MESSAGE ("Not a Datatype: " ^ ty_str)
@@ -135,15 +133,15 @@
       cases_thm_of_induct_thm (#induction dt)
     end;
 
-(* 
- val ty = (snd o hd o map Term.dest_Free o Term.term_frees) t;  
+(*
+ val ty = (snd o hd o map Term.dest_Free o Term.term_frees) t;
 *)
 
 
 (* for use when there are no prems to the subgoal *)
 (* does a case split on the given variable *)
-fun mk_casesplit_goal_thm sgn (vstr,ty) gt = 
-    let 
+fun mk_casesplit_goal_thm sgn (vstr,ty) gt =
+    let
       val x = Free(vstr,ty)
       val abst = Abs(vstr, ty, Term.abstract_over (x, gt));
 
@@ -155,12 +153,12 @@
       val free_ct = ctermify x;
 
       val casethm_vars = rev (Term.term_vars (Thm.concl_of case_thm));
-       
+
       val casethm_tvars = Term.term_tvars (Thm.concl_of case_thm);
-      val (Pv, Dv, type_insts) = 
-          case (Thm.concl_of case_thm) of 
-            (_ $ ((Pv as Var(P,Pty)) $ (Dv as Var(D, Dty)))) => 
-            (Pv, Dv, 
+      val (Pv, Dv, type_insts) =
+          case (Thm.concl_of case_thm) of
+            (_ $ ((Pv as Var(P,Pty)) $ (Dv as Var(D, Dty)))) =>
+            (Pv, Dv,
              Sign.typ_match sgn (Dty, ty) Vartab.empty)
           | _ => raise ERROR_MESSAGE ("not a valid case thm");
       val type_cinsts = map (fn (ixn, (S, T)) => (ctypify (TVar (ixn, S)), ctypify T))
@@ -168,17 +166,17 @@
       val cPv = ctermify (Envir.subst_TVars type_insts Pv);
       val cDv = ctermify (Envir.subst_TVars type_insts Dv);
     in
-      (beta_eta_contract 
+      (beta_eta_contract
          (case_thm
-            |> Thm.instantiate (type_cinsts, []) 
+            |> Thm.instantiate (type_cinsts, [])
             |> Thm.instantiate ([], [(cPv, abs_ct), (cDv, free_ct)])))
     end;
 
 
 (* for use when there are no prems to the subgoal *)
 (* does a case split on the given variable (Free fv) *)
-fun casesplit_free fv i th = 
-    let 
+fun casesplit_free fv i th =
+    let
       val (subgoalth, exp) = IsaND.fix_alls i th;
       val subgoalth' = atomize_goals subgoalth;
       val gt = Data.dest_Trueprop (Logic.get_goal (Thm.prop_of subgoalth') 1);
@@ -189,27 +187,27 @@
 
       val split_goal_th = splitter_thm RS subgoalth';
       val rulified_split_goal_th = rulify_goals split_goal_th;
-    in 
+    in
       IsaND.export_back exp rulified_split_goal_th
     end;
 
 
 (* for use when there are no prems to the subgoal *)
 (* does a case split on the given variable *)
-fun casesplit_name vstr i th = 
-    let 
+fun casesplit_name vstr i th =
+    let
       val (subgoalth, exp) = IsaND.fix_alls i th;
       val subgoalth' = atomize_goals subgoalth;
       val gt = Data.dest_Trueprop (Logic.get_goal (Thm.prop_of subgoalth') 1);
 
       val freets = Term.term_frees gt;
-      fun getter x = 
-          let val (n,ty) = Term.dest_Free x in 
-            (if vstr = n orelse vstr = Syntax.dest_skolem n 
+      fun getter x =
+          let val (n,ty) = Term.dest_Free x in
+            (if vstr = n orelse vstr = Syntax.dest_skolem n
              then SOME (n,ty) else NONE )
             handle Fail _ => NONE (* dest_skolem *)
           end;
-      val (n,ty) = case Library.get_first getter freets 
+      val (n,ty) = case Library.get_first getter freets
                 of SOME (n, ty) => (n, ty)
                  | _ => raise ERROR_MESSAGE ("no such variable " ^ vstr);
       val sgn = Thm.sign_of_thm th;
@@ -220,12 +218,12 @@
       val split_goal_th = splitter_thm RS subgoalth';
 
       val rulified_split_goal_th = rulify_goals split_goal_th;
-    in 
+    in
       IsaND.export_back exp rulified_split_goal_th
     end;
 
 
-(* small example: 
+(* small example:
 Goal "P (x :: nat) & (C y --> Q (y :: nat))";
 by (rtac (thm "conjI") 1);
 val th = topthm();
@@ -251,21 +249,21 @@
   | find_term_split (Free v, Const _) = SOME v
   | find_term_split (Free v, Abs _) = SOME v (* do we really want this case? *)
   | find_term_split (Free v, Var _) = NONE (* keep searching *)
-  | find_term_split (a $ b, a2 $ b2) = 
-    (case find_term_split (a, a2) of 
-       NONE => find_term_split (b,b2)  
+  | find_term_split (a $ b, a2 $ b2) =
+    (case find_term_split (a, a2) of
+       NONE => find_term_split (b,b2)
      | vopt => vopt)
-  | find_term_split (Abs(_,ty,t1), Abs(_,ty2,t2)) = 
+  | find_term_split (Abs(_,ty,t1), Abs(_,ty2,t2)) =
     find_term_split (t1, t2)
-  | find_term_split (Const (x,ty), Const(x2,ty2)) = 
+  | find_term_split (Const (x,ty), Const(x2,ty2)) =
     if x = x2 then NONE else (* keep searching *)
     raise find_split_exp (* stop now *)
             "Terms are not identical upto a free varaible! (Consts)"
-  | find_term_split (Bound i, Bound j) =     
+  | find_term_split (Bound i, Bound j) =
     if i = j then NONE else (* keep searching *)
     raise find_split_exp (* stop now *)
             "Terms are not identical upto a free varaible! (Bound)"
-  | find_term_split (a, b) = 
+  | find_term_split (a, b) =
     raise find_split_exp (* stop now *)
             "Terms are not identical upto a free varaible! (Other)";
 
@@ -273,7 +271,7 @@
 then look for a free variable to split, breaking the subgoal closer to
 splitth. *)
 fun find_thm_split splitth i genth =
-    find_term_split (Logic.get_goal (Thm.prop_of genth) i, 
+    find_term_split (Logic.get_goal (Thm.prop_of genth) i,
                      Thm.concl_of splitth) handle find_split_exp _ => NONE;
 
 (* as above but searches "splitths" for a theorem that suggest a case split *)
@@ -292,33 +290,33 @@
 (* Note: This should not be a separate tactic but integrated into the
 case split done during recdef's case analysis, this would avoid us
 having to (re)search for variables to split. *)
-fun splitto splitths genth = 
-    let 
+fun splitto splitths genth =
+    let
       val _ = assert (not (null splitths)) "splitto: no given splitths";
       val sgn = Thm.sign_of_thm genth;
 
-      (* check if we are a member of splitths - FIXME: quicker and 
+      (* check if we are a member of splitths - FIXME: quicker and
       more flexible with discrim net. *)
-      fun solve_by_splitth th split = 
+      fun solve_by_splitth th split =
           Thm.biresolution false [(false,split)] 1 th;
 
-      fun split th = 
-          (case find_thms_split splitths 1 th of 
-             NONE => 
+      fun split th =
+          (case find_thms_split splitths 1 th of
+             NONE =>
              (writeln "th:";
               Display.print_thm th; writeln "split ths:";
               Display.print_thms splitths; writeln "\n--";
               raise ERROR_MESSAGE "splitto: cannot find variable to split on")
-            | SOME v => 
-             let 
+            | SOME v =>
+             let
                val gt = Data.dest_Trueprop (List.nth(Thm.prems_of th, 0));
                val split_thm = mk_casesplit_goal_thm sgn v gt;
                val (subthms, expf) = IsaND.fixed_subgoal_thms split_thm;
-             in 
+             in
                expf (map recsplitf subthms)
              end)
 
-      and recsplitf th = 
+      and recsplitf th =
           (* note: multiple unifiers! we only take the first element,
              probably fine -- there is probably only one anyway. *)
           (case Library.get_first (Seq.pull o solve_by_splitth th) splitths of
@@ -339,27 +337,27 @@
 (* derive eqs, assuming strict, ie the rules have no assumptions = all
    the well-foundness conditions have been solved. *)
 local
-  fun get_related_thms i = 
+  fun get_related_thms i =
       List.mapPartial ((fn (r,x) => if x = i then SOME r else NONE));
-      
-  fun solve_eq (th, [], i) = 
+
+  fun solve_eq (th, [], i) =
       raise ERROR_MESSAGE "derive_init_eqs: missing rules"
     | solve_eq (th, [a], i) = (a, i)
     | solve_eq (th, splitths as (_ :: _), i) = (splitto splitths th,i);
 in
-fun derive_init_eqs sgn rules eqs = 
-    let 
-      val eqths = map (Thm.trivial o (Thm.cterm_of sgn) o Data.mk_Trueprop) 
+fun derive_init_eqs sgn rules eqs =
+    let
+      val eqths = map (Thm.trivial o (Thm.cterm_of sgn) o Data.mk_Trueprop)
                       eqs
     in
       (rev o map solve_eq)
-        (Library.foldln 
-           (fn (e,i) => 
-               (curry (op ::)) (e, (get_related_thms (i - 1) rules), i - 1)) 
+        (Library.foldln
+           (fn (e,i) =>
+               (curry (op ::)) (e, (get_related_thms (i - 1) rules), i - 1))
            eqths [])
     end;
 end;
-(* 
+(*
     val (rs_hwfc, unhidefs) = Library.split_list (map hide_prems rules)
     (map2 (op |>) (ths, expfs))
 *)
--- a/src/Provers/eqsubst.ML	Mon Aug 01 19:20:31 2005 +0200
+++ b/src/Provers/eqsubst.ML	Mon Aug 01 19:20:32 2005 +0200
@@ -1,12 +1,12 @@
-(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
+(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *)
 (*  Title:      Provers/eqsubst.ML
     ID:         $Id$
     Author:     Lucas Dixon, University of Edinburgh
                 lucas.dixon@ed.ac.uk
-    Modified:   18 Feb 2005 - Lucas - 
+    Modified:   18 Feb 2005 - Lucas -
     Created:    29 Jan 2005
 *)
-(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) 
+(* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *)
 (*  DESCRIPTION:
 
     A Tactic to perform a substiution using an equation.
@@ -27,81 +27,81 @@
 
 
 (* the signature of an instance of the SQSUBST tactic *)
-signature EQSUBST_TAC = 
+signature EQSUBST_TAC =
 sig
 
-  exception eqsubst_occL_exp of 
-            string * (int list) * (Thm.thm list) * int * Thm.thm;
+  exception eqsubst_occL_exp of
+            string * (int list) * (thm list) * int * thm;
 
 
-  type match = 
-       ((Term.indexname * (Term.sort * Term.typ)) list (* type instantiations *)
-        * (Term.indexname * (Term.typ * Term.term)) list) (* term instantiations *)
-       * (string * Term.typ) list (* fake named type abs env *)
-       * (string * Term.typ) list (* type abs env *)
-       * Term.term (* outer term *)
+  type match =
+       ((indexname * (sort * typ)) list (* type instantiations *)
+        * (indexname * (typ * term)) list) (* term instantiations *)
+       * (string * typ) list (* fake named type abs env *)
+       * (string * typ) list (* type abs env *)
+       * term (* outer term *)
 
-  type searchinfo = 
-       Sign.sg (* sign for matching *)
+  type searchinfo =
+       theory (* sign for matching *)
        * int (* maxidx *)
        * BasicIsaFTerm.FcTerm (* focusterm to search under *)
 
   val prep_subst_in_asm :
          int (* subgoal to subst in *)
-      -> Thm.thm (* target theorem with subgoals *)
+      -> thm (* target theorem with subgoals *)
       -> int (* premise to subst in *)
-      -> (Thm.cterm list (* certified free var placeholders for vars *) 
+      -> (cterm list (* certified free var placeholders for vars *)
           * int (* premice no. to subst *)
           * int (* number of assumptions of premice *)
-          * Thm.thm) (* premice as a new theorem for forward reasoning *)
+          * thm) (* premice as a new theorem for forward reasoning *)
          * searchinfo (* search info: prem id etc *)
 
   val prep_subst_in_asms :
          int (* subgoal to subst in *)
-      -> Thm.thm (* target theorem with subgoals *)
-      -> ((Thm.cterm list (* certified free var placeholders for vars *) 
+      -> thm (* target theorem with subgoals *)
+      -> ((cterm list (* certified free var placeholders for vars *)
           * int (* premice no. to subst *)
           * int (* number of assumptions of premice *)
-          * Thm.thm) (* premice as a new theorem for forward reasoning *)
+          * thm) (* premice as a new theorem for forward reasoning *)
          * searchinfo) list
 
   val apply_subst_in_asm :
       int (* subgoal *)
-      -> Thm.thm (* overall theorem *)
-      -> Thm.thm (* rule *)
-      -> (Thm.cterm list (* certified free var placeholders for vars *) 
+      -> thm (* overall theorem *)
+      -> thm (* rule *)
+      -> (cterm list (* certified free var placeholders for vars *)
           * int (* assump no being subst *)
-          * int (* num of premises of asm *) 
-          * Thm.thm) (* premthm *)
+          * int (* num of premises of asm *)
+          * thm) (* premthm *)
       * match
-      -> Thm.thm Seq.seq
+      -> thm Seq.seq
 
   val prep_concl_subst :
          int (* subgoal *)
-      -> Thm.thm (* overall goal theorem *)
-      -> (Thm.cterm list * Thm.thm) * searchinfo (* (cvfs, conclthm), matchf *)
+      -> thm (* overall goal theorem *)
+      -> (cterm list * thm) * searchinfo (* (cvfs, conclthm), matchf *)
 
   val apply_subst_in_concl :
         int (* subgoal *)
-        -> Thm.thm (* thm with all goals *)
-        -> Thm.cterm list (* certified free var placeholders for vars *)
-           * Thm.thm  (* trivial thm of goal concl *)
+        -> thm (* thm with all goals *)
+        -> cterm list (* certified free var placeholders for vars *)
+           * thm  (* trivial thm of goal concl *)
             (* possible matches/unifiers *)
-        -> Thm.thm (* rule *)
+        -> thm (* rule *)
         -> match
-        -> Thm.thm Seq.seq (* substituted goal *)
+        -> thm Seq.seq (* substituted goal *)
 
   (* basic notion of search *)
-  val searchf_tlr_unify_all : 
-      (searchinfo 
-       -> Term.term (* lhs *)
+  val searchf_tlr_unify_all :
+      (searchinfo
+       -> term (* lhs *)
        -> match Seq.seq Seq.seq)
-  val searchf_tlr_unify_valid : 
-      (searchinfo 
-       -> Term.term (* lhs *)
+  val searchf_tlr_unify_valid :
+      (searchinfo
+       -> term (* lhs *)
        -> match Seq.seq Seq.seq)
 
-  (* specialise search constructor for conclusion skipping occurnaces. *)
+  (* specialise search constructor for conclusion skipping occurrences. *)
      val skip_first_occs_search :
         int -> ('a -> 'b -> 'c Seq.seq Seq.seq) -> 'a -> 'b -> 'c Seq.seq
   (* specialised search constructor for assumptions using skips *)
@@ -110,98 +110,98 @@
         'a -> int -> 'b -> 'c IsaPLib.skipseqT
 
   (* tactics and methods *)
-  val eqsubst_asm_meth : int list -> Thm.thm list -> Proof.method
-  val eqsubst_asm_tac : 
-      int list -> Thm.thm list -> int -> Thm.thm -> Thm.thm Seq.seq
-  val eqsubst_asm_tac' : 
+  val eqsubst_asm_meth : int list -> thm list -> Proof.method
+  val eqsubst_asm_tac :
+      int list -> thm list -> int -> thm -> thm Seq.seq
+  val eqsubst_asm_tac' :
       (* search function with skips *)
-      (searchinfo -> int -> Term.term -> match IsaPLib.skipseqT) 
+      (searchinfo -> int -> term -> match IsaPLib.skipseqT)
       -> int (* skip to *)
-      -> Thm.thm (* rule *)
+      -> thm (* rule *)
       -> int (* subgoal number *)
-      -> Thm.thm (* tgt theorem with subgoals *)
-      -> Thm.thm Seq.seq (* new theorems *)
+      -> thm (* tgt theorem with subgoals *)
+      -> thm Seq.seq (* new theorems *)
 
-  val eqsubst_meth : int list -> Thm.thm list -> Proof.method
-  val eqsubst_tac : 
-      int list -> Thm.thm list -> int -> Thm.thm -> Thm.thm Seq.seq
-  val eqsubst_tac' : 
-      (searchinfo -> Term.term -> match Seq.seq) 
-      -> Thm.thm -> int -> Thm.thm -> Thm.thm Seq.seq
+  val eqsubst_meth : int list -> thm list -> Proof.method
+  val eqsubst_tac :
+      int list -> thm list -> int -> thm -> thm Seq.seq
+  val eqsubst_tac' :
+      (searchinfo -> term -> match Seq.seq)
+      -> thm -> int -> thm -> thm Seq.seq
 
-  val meth : (bool * int list) * Thm.thm list -> Proof.context -> Proof.method
+  val meth : (bool * int list) * thm list -> Proof.context -> Proof.method
   val setup : (Theory.theory -> Theory.theory) list
 end;
 
 
-functor EQSubstTacFUN (structure EqRuleData : EQRULE_DATA) 
+functor EQSubstTacFUN (structure EqRuleData : EQRULE_DATA)
   : EQSUBST_TAC
 = struct
 
   (* a type abriviation for match information *)
-  type match = 
-       ((Term.indexname * (Term.sort * Term.typ)) list (* type instantiations *)
-        * (Term.indexname * (Term.typ * Term.term)) list) (* term instantiations *)
-       * (string * Term.typ) list (* fake named type abs env *)
-       * (string * Term.typ) list (* type abs env *)
-       * Term.term (* outer term *)
+  type match =
+       ((indexname * (sort * typ)) list (* type instantiations *)
+        * (indexname * (typ * term)) list) (* term instantiations *)
+       * (string * typ) list (* fake named type abs env *)
+       * (string * typ) list (* type abs env *)
+       * term (* outer term *)
 
-  type searchinfo = 
+  type searchinfo =
        Sign.sg (* sign for matching *)
        * int (* maxidx *)
        * BasicIsaFTerm.FcTerm (* focusterm to search under *)
 
 (* FOR DEBUGGING...
 type trace_subst_errT = int (* subgoal *)
-        * Thm.thm (* thm with all goals *)
+        * thm (* thm with all goals *)
         * (Thm.cterm list (* certified free var placeholders for vars *)
-           * Thm.thm)  (* trivial thm of goal concl *)
+           * thm)  (* trivial thm of goal concl *)
             (* possible matches/unifiers *)
-        * Thm.thm (* rule *)
-        * (((Term.indexname * Term.typ) list (* type instantiations *)
-              * (Term.indexname * Term.term) list ) (* term instantiations *)
-             * (string * Term.typ) list (* Type abs env *)
-             * Term.term) (* outer term *);
+        * thm (* rule *)
+        * (((indexname * typ) list (* type instantiations *)
+              * (indexname * term) list ) (* term instantiations *)
+             * (string * typ) list (* Type abs env *)
+             * term) (* outer term *);
 
 val trace_subst_err = (ref NONE : trace_subst_errT option ref);
 val trace_subst_search = ref false;
 exception trace_subst_exp of trace_subst_errT;
  *)
 
-(* also defined in /HOL/Tools/inductive_codegen.ML, 
+(* also defined in /HOL/Tools/inductive_codegen.ML,
    maybe move this to seq.ML ? *)
 infix 5 :->;
 fun s :-> f = Seq.flat (Seq.map f s);
 
 (* search from top, left to right, then down *)
-fun search_tlr_all_f f ft = 
+fun search_tlr_all_f f ft =
     let
-      fun maux ft = 
-          let val t' = (IsaFTerm.focus_of_fcterm ft) 
-            (* val _ = 
-                if !trace_subst_search then 
+      fun maux ft =
+          let val t' = (IsaFTerm.focus_of_fcterm ft)
+            (* val _ =
+                if !trace_subst_search then
                   (writeln ("Examining: " ^ (TermLib.string_of_term t'));
                    TermLib.writeterm t'; ())
                 else (); *)
-          in 
-          (case t' of 
-            (_ $ _) => Seq.append(maux (IsaFTerm.focus_left ft), 
-                       Seq.cons(f ft, 
+          in
+          (case t' of
+            (_ $ _) => Seq.append(maux (IsaFTerm.focus_left ft),
+                       Seq.cons(f ft,
                                   maux (IsaFTerm.focus_right ft)))
           | (Abs _) => Seq.cons(f ft, maux (IsaFTerm.focus_abs ft))
           | leaf => Seq.single (f ft)) end
     in maux ft end;
 
 (* search from top, left to right, then down *)
-fun search_tlr_valid_f f ft = 
+fun search_tlr_valid_f f ft =
     let
-      fun maux ft = 
-          let 
+      fun maux ft =
+          let
             val hereseq = if IsaFTerm.valid_match_start ft then f ft else Seq.empty
-          in 
-          (case (IsaFTerm.focus_of_fcterm ft) of 
-            (_ $ _) => Seq.append(maux (IsaFTerm.focus_left ft), 
-                       Seq.cons(hereseq, 
+          in
+          (case (IsaFTerm.focus_of_fcterm ft) of
+            (_ $ _) => Seq.append(maux (IsaFTerm.focus_left ft),
+                       Seq.cons(hereseq,
                                   maux (IsaFTerm.focus_right ft)))
           | (Abs _) => Seq.cons(hereseq, maux (IsaFTerm.focus_abs ft))
           | leaf => Seq.single (hereseq))
@@ -209,16 +209,16 @@
     in maux ft end;
 
 (* search all unifications *)
-fun searchf_tlr_unify_all (sgn, maxidx, ft) lhs = 
-    IsaFTerm.find_fcterm_matches 
-      search_tlr_all_f 
+fun searchf_tlr_unify_all (sgn, maxidx, ft) lhs =
+    IsaFTerm.find_fcterm_matches
+      search_tlr_all_f
       (IsaFTerm.clean_unify_ft sgn maxidx lhs)
       ft;
 
 (* search only for 'valid' unifiers (non abs subterms and non vars) *)
-fun searchf_tlr_unify_valid (sgn, maxidx, ft) lhs  = 
-    IsaFTerm.find_fcterm_matches 
-      search_tlr_valid_f 
+fun searchf_tlr_unify_valid (sgn, maxidx, ft) lhs  =
+    IsaFTerm.find_fcterm_matches
+      search_tlr_valid_f
       (IsaFTerm.clean_unify_ft sgn maxidx lhs)
       ft;
 
@@ -228,7 +228,7 @@
 (* conclthm is a theorem of for just the conclusion *)
 (* m is instantiation/match information *)
 (* rule is the equation for substitution *)
-fun apply_subst_in_concl i th (cfvs, conclthm) rule m = 
+fun apply_subst_in_concl i th (cfvs, conclthm) rule m =
     (RWInst.rw m rule conclthm)
       |> IsaND.unfix_frees cfvs
       |> RWInst.beta_eta_contract
@@ -236,8 +236,8 @@
 
 (* substitute within the conclusion of goal i of gth, using a meta
 equation rule. Note that we assume rule has var indicies zero'd *)
-fun prep_concl_subst i gth = 
-    let 
+fun prep_concl_subst i gth =
+    let
       val th = Thm.incr_indexes 1 gth;
       val tgt_term = Thm.prop_of th;
 
@@ -251,20 +251,20 @@
       val conclterm = Logic.strip_imp_concl fixedbody;
       val conclthm = trivify conclterm;
       val maxidx = Term.maxidx_of_term conclterm;
-      val ft = ((IsaFTerm.focus_right  
+      val ft = ((IsaFTerm.focus_right
                  o IsaFTerm.focus_left
-                 o IsaFTerm.fcterm_of_term 
+                 o IsaFTerm.fcterm_of_term
                  o Thm.prop_of) conclthm)
     in
       ((cfvs, conclthm), (sgn, maxidx, ft))
     end;
 
 (* substitute using an object or meta level equality *)
-fun eqsubst_tac' searchf instepthm i th = 
-    let 
+fun eqsubst_tac' searchf instepthm i th =
+    let
       val (cvfsconclthm, searchinfo) = prep_concl_subst i th;
-      val stepthms = 
-          Seq.map Drule.zero_var_indexes 
+      val stepthms =
+          Seq.map Drule.zero_var_indexes
                   (Seq.of_list (EqRuleData.prep_meta_eq instepthm));
       fun rewrite_with_thm r =
           let val (lhs,_) = Logic.dest_equals (Thm.concl_of r);
@@ -275,43 +275,43 @@
 
 (* Tactic.distinct_subgoals_tac -- fails to free type variables *)
 
-(* custom version of distinct subgoals that works with term and 
-   type variables. Asssumes th is in beta-eta normal form. 
+(* custom version of distinct subgoals that works with term and
+   type variables. Asssumes th is in beta-eta normal form.
    Also, does nothing if flexflex contraints are present. *)
 fun distinct_subgoals th =
     if List.null (Thm.tpairs_of th) then
       let val (fixes,fixedthm) = IsaND.fix_vars_and_tvars th
         val (fixedthconcl,cprems) = IsaND.hide_prems fixedthm
       in
-        IsaND.unfix_frees_and_tfrees 
+        IsaND.unfix_frees_and_tfrees
           fixes
-          (Drule.implies_intr_list 
-             (Library.gen_distinct 
+          (Drule.implies_intr_list
+             (Library.gen_distinct
                 (fn (x, y) => Thm.term_of x = Thm.term_of y)
                 cprems) fixedthconcl)
       end
     else th;
 
-(* General substiuttion of multiple occurances using one of 
+(* General substiuttion of multiple occurances using one of
    the given theorems*)
-exception eqsubst_occL_exp of 
-          string * (int list) * (Thm.thm list) * int * Thm.thm;
-fun skip_first_occs_search occ srchf sinfo lhs = 
-    case (IsaPLib.skipto_seqseq occ (srchf sinfo lhs)) of 
+exception eqsubst_occL_exp of
+          string * (int list) * (thm list) * int * thm;
+fun skip_first_occs_search occ srchf sinfo lhs =
+    case (IsaPLib.skipto_seqseq occ (srchf sinfo lhs)) of
       IsaPLib.skipmore _ => Seq.empty
     | IsaPLib.skipseq ss => Seq.flat ss;
 
-fun eqsubst_tac occL thms i th = 
+fun eqsubst_tac occL thms i th =
     let val nprems = Thm.nprems_of th in
       if nprems < i then Seq.empty else
-      let val thmseq = (Seq.of_list thms) 
-        fun apply_occ occ th = 
-            thmseq :-> 
-                    (fn r => eqsubst_tac' (skip_first_occs_search 
+      let val thmseq = (Seq.of_list thms)
+        fun apply_occ occ th =
+            thmseq :->
+                    (fn r => eqsubst_tac' (skip_first_occs_search
                                     occ searchf_tlr_unify_valid) r
                                  (i + ((Thm.nprems_of th) - nprems))
                                  th);
-        val sortedoccL = 
+        val sortedoccL =
             Library.sort (Library.rev_order o Library.int_ord) occL;
       in
         Seq.map distinct_subgoals (Seq.EVERY (map apply_occ sortedoccL) th)
@@ -323,21 +323,21 @@
 (* inthms are the given arguments in Isar, and treated as eqstep with
    the first one, then the second etc *)
 fun eqsubst_meth occL inthms =
-    Method.METHOD 
+    Method.METHOD
       (fn facts =>
           HEADGOAL ( Method.insert_tac facts THEN' eqsubst_tac occL inthms ));
 
 (* apply a substitution inside assumption j, keeps asm in the same place *)
-fun apply_subst_in_asm i th rule ((cfvs, j, ngoalprems, pth),m) = 
-    let 
+fun apply_subst_in_asm i th rule ((cfvs, j, ngoalprems, pth),m) =
+    let
       val th2 = Thm.rotate_rule (j - 1) i th; (* put premice first *)
-      val preelimrule = 
+      val preelimrule =
           (RWInst.rw m rule pth)
             |> (Seq.hd o Tactic.prune_params_tac)
             |> Thm.permute_prems 0 ~1 (* put old asm first *)
             |> IsaND.unfix_frees cfvs (* unfix any global params *)
             |> RWInst.beta_eta_contract; (* normal form *)
-  (*    val elimrule = 
+  (*    val elimrule =
           preelimrule
             |> Tactic.make_elim (* make into elim rule *)
             |> Thm.lift_rule (th2, i); (* lift into context *)
@@ -347,7 +347,7 @@
       Seq.map (Thm.rotate_rule (~j) ((Thm.nprems_of rule) + i))
       (Tactic.dtac preelimrule i th2)
 
-      (* (Thm.bicompose 
+      (* (Thm.bicompose
                  false (* use unification *)
                  (true, (* elim resolution *)
                   elimrule, (2 + (Thm.nprems_of rule)) - ngoalprems)
@@ -361,8 +361,8 @@
 subgoal i of gth. Note the repetition of work done for each
 assumption, i.e. this can be made more efficient for search over
 multiple assumptions.  *)
-fun prep_subst_in_asm i gth j = 
-    let 
+fun prep_subst_in_asm i gth j =
+    let
       val th = Thm.incr_indexes 1 gth;
       val tgt_term = Thm.prop_of th;
 
@@ -379,58 +379,58 @@
       val pth = trivify asmt;
       val maxidx = Term.maxidx_of_term asmt;
 
-      val ft = ((IsaFTerm.focus_right 
-                 o IsaFTerm.fcterm_of_term 
+      val ft = ((IsaFTerm.focus_right
+                 o IsaFTerm.fcterm_of_term
                  o Thm.prop_of) pth)
     in ((cfvs, j, asm_nprems, pth), (sgn, maxidx, ft)) end;
 
 (* prepare subst in every possible assumption *)
-fun prep_subst_in_asms i gth = 
+fun prep_subst_in_asms i gth =
     map (prep_subst_in_asm i gth)
-        ((rev o IsaPLib.mk_num_list o length) 
+        ((rev o IsaPLib.mk_num_list o length)
            (Logic.prems_of_goal (Thm.prop_of gth) i));
 
 
 (* substitute in an assumption using an object or meta level equality *)
-fun eqsubst_asm_tac' searchf skipocc instepthm i th = 
-    let 
+fun eqsubst_asm_tac' searchf skipocc instepthm i th =
+    let
       val asmpreps = prep_subst_in_asms i th;
-      val stepthms = 
-          Seq.map Drule.zero_var_indexes 
+      val stepthms =
+          Seq.map Drule.zero_var_indexes
               (Seq.of_list (EqRuleData.prep_meta_eq instepthm))
       fun rewrite_with_thm r =
           let val (lhs,_) = Logic.dest_equals (Thm.concl_of r)
             fun occ_search occ [] = Seq.empty
               | occ_search occ ((asminfo, searchinfo)::moreasms) =
-                (case searchf searchinfo occ lhs of 
+                (case searchf searchinfo occ lhs of
                    IsaPLib.skipmore i => occ_search i moreasms
-                 | IsaPLib.skipseq ss => 
+                 | IsaPLib.skipseq ss =>
                    Seq.append (Seq.map (Library.pair asminfo)
-                                       (Seq.flat ss), 
+                                       (Seq.flat ss),
                                occ_search 1 moreasms))
                               (* find later substs also *)
-          in 
+          in
             (occ_search skipocc asmpreps) :-> (apply_subst_in_asm i th r)
           end;
     in stepthms :-> rewrite_with_thm end;
 
 
-fun skip_first_asm_occs_search searchf sinfo occ lhs = 
+fun skip_first_asm_occs_search searchf sinfo occ lhs =
     IsaPLib.skipto_seqseq occ (searchf sinfo lhs);
 
-fun eqsubst_asm_tac occL thms i th = 
-    let val nprems = Thm.nprems_of th 
+fun eqsubst_asm_tac occL thms i th =
+    let val nprems = Thm.nprems_of th
     in
       if nprems < i then Seq.empty else
-      let val thmseq = (Seq.of_list thms) 
-        fun apply_occ occK th = 
-            thmseq :-> 
-                    (fn r => 
-                        eqsubst_asm_tac' (skip_first_asm_occs_search 
+      let val thmseq = (Seq.of_list thms)
+        fun apply_occ occK th =
+            thmseq :->
+                    (fn r =>
+                        eqsubst_asm_tac' (skip_first_asm_occs_search
                                             searchf_tlr_unify_valid) occK r
                                          (i + ((Thm.nprems_of th) - nprems))
                                          th);
-        val sortedoccs = 
+        val sortedoccs =
             Library.sort (Library.rev_order o Library.int_ord) occL
       in
         Seq.map distinct_subgoals
@@ -442,14 +442,14 @@
 (* inthms are the given arguments in Isar, and treated as eqstep with
    the first one, then the second etc *)
 fun eqsubst_asm_meth occL inthms =
-    Method.METHOD 
+    Method.METHOD
       (fn facts =>
           HEADGOAL (Method.insert_tac facts THEN' eqsubst_asm_tac occL inthms ));
 
 (* combination method that takes a flag (true indicates that subst
 should be done to an assumption, false = apply to the conclusion of
 the goal) as well as the theorems to use *)
-fun meth ((asmflag, occL), inthms) ctxt = 
+fun meth ((asmflag, occL), inthms) ctxt =
     if asmflag then eqsubst_asm_meth occL inthms else eqsubst_meth occL inthms;
 
 (* syntax for options, given "(asm)" will give back true, without
@@ -464,13 +464,13 @@
 
 (* method syntax, first take options, then theorems *)
 fun meth_syntax meth src ctxt =
-    meth (snd (Method.syntax ((Scan.lift options_syntax) 
-                                -- (Scan.lift ith_syntax) 
-                                -- Attrib.local_thms) src ctxt)) 
+    meth (snd (Method.syntax ((Scan.lift options_syntax)
+                                -- (Scan.lift ith_syntax)
+                                -- Attrib.local_thms) src ctxt))
          ctxt;
 
 (* setup function for adding method to theory. *)
-val setup = 
+val setup =
     [Method.add_method ("subst", meth_syntax meth, "Substiution with an equation. Use \"(asm)\" option to substitute in an assumption.")];
 
-end;
\ No newline at end of file
+end;