src/HOL/Tools/TFL/tfl.ML
author wenzelm
Sat, 16 Apr 2011 16:15:37 +0200
changeset 43232 23f352990944
parent 41739 a2ad5b824051
child 44206 2b47822868e4
permissions -rw-r--r--
modernized structure Proof_Context;
     1 (*  Title:      HOL/Tools/TFL/tfl.ML
     2     Author:     Konrad Slind, Cambridge University Computer Laboratory
     3 
     4 First part of main module.
     5 *)
     6 
     7 signature PRIM =
     8 sig
     9   val trace: bool Unsynchronized.ref
    10   val trace_thms: string -> thm list -> unit
    11   val trace_cterm: string -> cterm -> unit
    12   type pattern
    13   val mk_functional: theory -> term list -> {functional: term, pats: pattern list}
    14   val wfrec_definition0: theory -> string -> term -> term -> theory * thm
    15   val post_definition: thm list -> theory * (thm * pattern list) ->
    16    {rules: thm,
    17     rows: int list,
    18     TCs: term list list,
    19     full_pats_TCs: (term * term list) list}
    20   val wfrec_eqns: theory -> xstring -> thm list -> term list ->
    21    {WFR: term,
    22     SV: term list,
    23     proto_def: term,
    24     extracta: (thm * term list) list,
    25     pats: pattern list}
    26   val lazyR_def: theory -> xstring -> thm list -> term list ->
    27    {theory: theory,
    28     rules: thm,
    29     R: term,
    30     SV: term list,
    31     full_pats_TCs: (term * term list) list,
    32     patterns : pattern list}
    33   val mk_induction: theory ->
    34     {fconst: term, R: term, SV: term list, pat_TCs_list: (term * term list) list} -> thm
    35   val postprocess: bool -> {wf_tac: tactic, terminator: tactic, simplifier: cterm -> thm}
    36     -> theory -> {rules: thm, induction: thm, TCs: term list list}
    37     -> {rules: thm, induction: thm, nested_tcs: thm list}
    38 end;
    39 
    40 structure Prim: PRIM =
    41 struct
    42 
    43 val trace = Unsynchronized.ref false;
    44 
    45 
    46 fun TFL_ERR func mesg = Utils.ERR {module = "Tfl", func = func, mesg = mesg};
    47 
    48 val concl = #2 o Rules.dest_thm;
    49 val hyp = #1 o Rules.dest_thm;
    50 
    51 val list_mk_type = Utils.end_itlist (curry (op -->));
    52 
    53 fun front_last [] = raise TFL_ERR "front_last" "empty list"
    54   | front_last [x] = ([],x)
    55   | front_last (h::t) =
    56      let val (pref,x) = front_last t
    57      in
    58         (h::pref,x)
    59      end;
    60 
    61 
    62 (*---------------------------------------------------------------------------
    63  * The next function is common to pattern-match translation and
    64  * proof of completeness of cases for the induction theorem.
    65  *
    66  * The curried function "gvvariant" returns a function to generate distinct
    67  * variables that are guaranteed not to be in names.  The names of
    68  * the variables go u, v, ..., z, aa, ..., az, ...  The returned
    69  * function contains embedded refs!
    70  *---------------------------------------------------------------------------*)
    71 fun gvvariant names =
    72   let val slist = Unsynchronized.ref names
    73       val vname = Unsynchronized.ref "u"
    74       fun new() =
    75          if member (op =) (!slist) (!vname)
    76          then (vname := Symbol.bump_string (!vname);  new())
    77          else (slist := !vname :: !slist;  !vname)
    78   in
    79   fn ty => Free(new(), ty)
    80   end;
    81 
    82 
    83 (*---------------------------------------------------------------------------
    84  * Used in induction theorem production. This is the simple case of
    85  * partitioning up pattern rows by the leading constructor.
    86  *---------------------------------------------------------------------------*)
    87 fun ipartition gv (constructors,rows) =
    88   let fun pfail s = raise TFL_ERR "partition.part" s
    89       fun part {constrs = [],   rows = [],   A} = rev A
    90         | part {constrs = [],   rows = _::_, A} = pfail"extra cases in defn"
    91         | part {constrs = _::_, rows = [],   A} = pfail"cases missing in defn"
    92         | part {constrs = c::crst, rows,     A} =
    93           let val (c, T) = dest_Const c
    94               val L = binder_types T
    95               val (in_group, not_in_group) =
    96                fold_rev (fn (row as (p::rst, rhs)) =>
    97                          fn (in_group,not_in_group) =>
    98                   let val (pc,args) = USyntax.strip_comb p
    99                   in if (#1(dest_Const pc) = c)
   100                      then ((args@rst, rhs)::in_group, not_in_group)
   101                      else (in_group, row::not_in_group)
   102                   end)      rows ([],[])
   103               val col_types = Utils.take type_of (length L, #1(hd in_group))
   104           in
   105           part{constrs = crst, rows = not_in_group,
   106                A = {constructor = c,
   107                     new_formals = map gv col_types,
   108                     group = in_group}::A}
   109           end
   110   in part{constrs = constructors, rows = rows, A = []}
   111   end;
   112 
   113 
   114 
   115 (*---------------------------------------------------------------------------
   116  * Each pattern carries with it a tag (i,b) where
   117  * i is the clause it came from and
   118  * b=true indicates that clause was given by the user
   119  * (or is an instantiation of a user supplied pattern)
   120  * b=false --> i = ~1
   121  *---------------------------------------------------------------------------*)
   122 
   123 type pattern = term * (int * bool)
   124 
   125 fun pattern_map f (tm,x) = (f tm, x);
   126 
   127 fun pattern_subst theta = pattern_map (subst_free theta);
   128 
   129 val pat_of = fst;
   130 fun row_of_pat x = fst (snd x);
   131 fun given x = snd (snd x);
   132 
   133 (*---------------------------------------------------------------------------
   134  * Produce an instance of a constructor, plus genvars for its arguments.
   135  *---------------------------------------------------------------------------*)
   136 fun fresh_constr ty_match colty gv c =
   137   let val (_,Ty) = dest_Const c
   138       val L = binder_types Ty
   139       and ty = body_type Ty
   140       val ty_theta = ty_match ty colty
   141       val c' = USyntax.inst ty_theta c
   142       val gvars = map (USyntax.inst ty_theta o gv) L
   143   in (c', gvars)
   144   end;
   145 
   146 
   147 (*---------------------------------------------------------------------------
   148  * Goes through a list of rows and picks out the ones beginning with a
   149  * pattern with constructor = name.
   150  *---------------------------------------------------------------------------*)
   151 fun mk_group name rows =
   152   fold_rev (fn (row as ((prfx, p::rst), rhs)) =>
   153             fn (in_group,not_in_group) =>
   154                let val (pc,args) = USyntax.strip_comb p
   155                in if ((#1 (Term.dest_Const pc) = name) handle TERM _ => false)
   156                   then (((prfx,args@rst), rhs)::in_group, not_in_group)
   157                   else (in_group, row::not_in_group) end)
   158       rows ([],[]);
   159 
   160 (*---------------------------------------------------------------------------
   161  * Partition the rows. Not efficient: we should use hashing.
   162  *---------------------------------------------------------------------------*)
   163 fun partition _ _ (_,_,_,[]) = raise TFL_ERR "partition" "no rows"
   164   | partition gv ty_match
   165               (constructors, colty, res_ty, rows as (((prfx,_),_)::_)) =
   166 let val fresh = fresh_constr ty_match colty gv
   167      fun part {constrs = [],      rows, A} = rev A
   168        | part {constrs = c::crst, rows, A} =
   169          let val (c',gvars) = fresh c
   170              val (in_group, not_in_group) = mk_group (#1 (dest_Const c')) rows
   171              val in_group' =
   172                  if (null in_group)  (* Constructor not given *)
   173                  then [((prfx, #2(fresh c)), (USyntax.ARB res_ty, (~1,false)))]
   174                  else in_group
   175          in
   176          part{constrs = crst,
   177               rows = not_in_group,
   178               A = {constructor = c',
   179                    new_formals = gvars,
   180                    group = in_group'}::A}
   181          end
   182 in part{constrs=constructors, rows=rows, A=[]}
   183 end;
   184 
   185 (*---------------------------------------------------------------------------
   186  * Misc. routines used in mk_case
   187  *---------------------------------------------------------------------------*)
   188 
   189 fun mk_pat (c,l) =
   190   let val L = length (binder_types (type_of c))
   191       fun build (prfx,tag,plist) =
   192           let val (args, plist') = chop L plist
   193           in (prfx,tag,list_comb(c,args)::plist') end
   194   in map build l end;
   195 
   196 fun v_to_prfx (prfx, v::pats) = (v::prfx,pats)
   197   | v_to_prfx _ = raise TFL_ERR "mk_case" "v_to_prfx";
   198 
   199 fun v_to_pats (v::prfx,tag, pats) = (prfx, tag, v::pats)
   200   | v_to_pats _ = raise TFL_ERR "mk_case" "v_to_pats";
   201 
   202 
   203 (*----------------------------------------------------------------------------
   204  * Translation of pattern terms into nested case expressions.
   205  *
   206  * This performs the translation and also builds the full set of patterns.
   207  * Thus it supports the construction of induction theorems even when an
   208  * incomplete set of patterns is given.
   209  *---------------------------------------------------------------------------*)
   210 
   211 fun mk_case ty_info ty_match usednames range_ty =
   212  let
   213  fun mk_case_fail s = raise TFL_ERR "mk_case" s
   214  val fresh_var = gvvariant usednames
   215  val divide = partition fresh_var ty_match
   216  fun expand constructors ty ((_,[]), _) = mk_case_fail"expand_var_row"
   217    | expand constructors ty (row as ((prfx, p::rst), rhs)) =
   218        if (is_Free p)
   219        then let val fresh = fresh_constr ty_match ty fresh_var
   220                 fun expnd (c,gvs) =
   221                   let val capp = list_comb(c,gvs)
   222                   in ((prfx, capp::rst), pattern_subst[(p,capp)] rhs)
   223                   end
   224             in map expnd (map fresh constructors)  end
   225        else [row]
   226  fun mk{rows=[],...} = mk_case_fail"no rows"
   227    | mk{path=[], rows = ((prfx, []), (tm,tag))::_} =  (* Done *)
   228         ([(prfx,tag,[])], tm)
   229    | mk{path=[], rows = _::_} = mk_case_fail"blunder"
   230    | mk{path as u::rstp, rows as ((prfx, []), rhs)::rst} =
   231         mk{path = path,
   232            rows = ((prfx, [fresh_var(type_of u)]), rhs)::rst}
   233    | mk{path = u::rstp, rows as ((_, p::_), _)::_} =
   234      let val (pat_rectangle,rights) = ListPair.unzip rows
   235          val col0 = map(hd o #2) pat_rectangle
   236      in
   237      if (forall is_Free col0)
   238      then let val rights' = map (fn(v,e) => pattern_subst[(v,u)] e)
   239                                 (ListPair.zip (col0, rights))
   240               val pat_rectangle' = map v_to_prfx pat_rectangle
   241               val (pref_patl,tm) = mk{path = rstp,
   242                                       rows = ListPair.zip (pat_rectangle',
   243                                                            rights')}
   244           in (map v_to_pats pref_patl, tm)
   245           end
   246      else
   247      let val pty as Type (ty_name,_) = type_of p
   248      in
   249      case (ty_info ty_name)
   250      of NONE => mk_case_fail("Not a known datatype: "^ty_name)
   251       | SOME{case_const,constructors} =>
   252         let
   253             val case_const_name = #1(dest_Const case_const)
   254             val nrows = maps (expand constructors pty) rows
   255             val subproblems = divide(constructors, pty, range_ty, nrows)
   256             val groups      = map #group subproblems
   257             and new_formals = map #new_formals subproblems
   258             and constructors' = map #constructor subproblems
   259             val news = map (fn (nf,rows) => {path = nf@rstp, rows=rows})
   260                            (ListPair.zip (new_formals, groups))
   261             val rec_calls = map mk news
   262             val (pat_rect,dtrees) = ListPair.unzip rec_calls
   263             val case_functions = map USyntax.list_mk_abs
   264                                   (ListPair.zip (new_formals, dtrees))
   265             val types = map type_of (case_functions@[u]) @ [range_ty]
   266             val case_const' = Const(case_const_name, list_mk_type types)
   267             val tree = list_comb(case_const', case_functions@[u])
   268             val pat_rect1 = flat (ListPair.map mk_pat (constructors', pat_rect))
   269         in (pat_rect1,tree)
   270         end
   271      end end
   272  in mk
   273  end;
   274 
   275 
   276 (* Repeated variable occurrences in a pattern are not allowed. *)
   277 fun FV_multiset tm =
   278    case (USyntax.dest_term tm)
   279      of USyntax.VAR{Name = c, Ty = T} => [Free(c, T)]
   280       | USyntax.CONST _ => []
   281       | USyntax.COMB{Rator, Rand} => FV_multiset Rator @ FV_multiset Rand
   282       | USyntax.LAMB _ => raise TFL_ERR "FV_multiset" "lambda";
   283 
   284 fun no_repeat_vars thy pat =
   285  let fun check [] = true
   286        | check (v::rst) =
   287          if member (op aconv) rst v then
   288             raise TFL_ERR "no_repeat_vars"
   289                           (quote (#1 (dest_Free v)) ^
   290                           " occurs repeatedly in the pattern " ^
   291                           quote (Syntax.string_of_term_global thy pat))
   292          else check rst
   293  in check (FV_multiset pat)
   294  end;
   295 
   296 fun dest_atom (Free p) = p
   297   | dest_atom (Const p) = p
   298   | dest_atom  _ = raise TFL_ERR "dest_atom" "function name not an identifier";
   299 
   300 fun same_name (p,q) = #1(dest_atom p) = #1(dest_atom q);
   301 
   302 local fun mk_functional_err s = raise TFL_ERR "mk_functional" s
   303       fun single [_$_] =
   304               mk_functional_err "recdef does not allow currying"
   305         | single [f] = f
   306         | single fs  =
   307               (*multiple function names?*)
   308               if length (distinct same_name fs) < length fs
   309               then mk_functional_err
   310                    "The function being declared appears with multiple types"
   311               else mk_functional_err
   312                    (string_of_int (length fs) ^
   313                     " distinct function names being declared")
   314 in
   315 fun mk_functional thy clauses =
   316  let val (L,R) = ListPair.unzip (map HOLogic.dest_eq clauses
   317                    handle TERM _ => raise TFL_ERR "mk_functional"
   318                         "recursion equations must use the = relation")
   319      val (funcs,pats) = ListPair.unzip (map (fn (t$u) =>(t,u)) L)
   320      val atom = single (distinct (op aconv) funcs)
   321      val (fname,ftype) = dest_atom atom
   322      val dummy = map (no_repeat_vars thy) pats
   323      val rows = ListPair.zip (map (fn x => ([]:term list,[x])) pats,
   324                               map_index (fn (i, t) => (t,(i,true))) R)
   325      val names = List.foldr OldTerm.add_term_names [] R
   326      val atype = type_of(hd pats)
   327      and aname = Name.variant names "a"
   328      val a = Free(aname,atype)
   329      val ty_info = Thry.match_info thy
   330      val ty_match = Thry.match_type thy
   331      val range_ty = type_of (hd R)
   332      val (patts, case_tm) = mk_case ty_info ty_match (aname::names) range_ty
   333                                     {path=[a], rows=rows}
   334      val patts1 = map (fn (_,tag,[pat]) => (pat,tag)) patts
   335           handle Match => mk_functional_err "error in pattern-match translation"
   336      val patts2 = Library.sort (Library.int_ord o Library.pairself row_of_pat) patts1
   337      val finals = map row_of_pat patts2
   338      val originals = map (row_of_pat o #2) rows
   339      val dummy = case (subtract (op =) finals originals)
   340              of [] => ()
   341           | L => mk_functional_err
   342  ("The following clauses are redundant (covered by preceding clauses): " ^
   343                    commas (map (fn i => string_of_int (i + 1)) L))
   344  in {functional = Abs(Long_Name.base_name fname, ftype,
   345                       abstract_over (atom,
   346                                      absfree(aname,atype, case_tm))),
   347      pats = patts2}
   348 end end;
   349 
   350 
   351 (*----------------------------------------------------------------------------
   352  *
   353  *                    PRINCIPLES OF DEFINITION
   354  *
   355  *---------------------------------------------------------------------------*)
   356 
   357 
   358 (*For Isabelle, the lhs of a definition must be a constant.*)
   359 fun const_def sign (c, Ty, rhs) =
   360   singleton (Syntax.check_terms (Proof_Context.init_global sign))
   361     (Const("==",dummyT) $ Const(c,Ty) $ rhs);
   362 
   363 (*Make all TVars available for instantiation by adding a ? to the front*)
   364 fun poly_tvars (Type(a,Ts)) = Type(a, map (poly_tvars) Ts)
   365   | poly_tvars (TFree (a,sort)) = TVar (("?" ^ a, 0), sort)
   366   | poly_tvars (TVar ((a,i),sort)) = TVar (("?" ^ a, i+1), sort);
   367 
   368 local val f_eq_wfrec_R_M =
   369            #ant(USyntax.dest_imp(#2(USyntax.strip_forall (concl Thms.WFREC_COROLLARY))))
   370       val {lhs=f, rhs} = USyntax.dest_eq f_eq_wfrec_R_M
   371       val (fname,_) = dest_Free f
   372       val (wfrec,_) = USyntax.strip_comb rhs
   373 in
   374 fun wfrec_definition0 thy fid R (functional as Abs(x, Ty, _)) =
   375  let val def_name = Long_Name.base_name fid ^ "_def"
   376      val wfrec_R_M =  map_types poly_tvars
   377                           (wfrec $ map_types poly_tvars R)
   378                       $ functional
   379      val def_term = const_def thy (fid, Ty, wfrec_R_M)
   380      val ([def], thy') =
   381       Global_Theory.add_defs false [Thm.no_attributes (Binding.name def_name, def_term)] thy
   382  in (thy', def) end;
   383 end;
   384 
   385 
   386 
   387 (*---------------------------------------------------------------------------
   388  * This structure keeps track of congruence rules that aren't derived
   389  * from a datatype definition.
   390  *---------------------------------------------------------------------------*)
   391 fun extraction_thms thy =
   392  let val {case_rewrites,case_congs} = Thry.extract_info thy
   393  in (case_rewrites, case_congs)
   394  end;
   395 
   396 
   397 (*---------------------------------------------------------------------------
   398  * Pair patterns with termination conditions. The full list of patterns for
   399  * a definition is merged with the TCs arising from the user-given clauses.
   400  * There can be fewer clauses than the full list, if the user omitted some
   401  * cases. This routine is used to prepare input for mk_induction.
   402  *---------------------------------------------------------------------------*)
   403 fun merge full_pats TCs =
   404 let fun insert (p,TCs) =
   405       let fun insrt ((x as (h,[]))::rst) =
   406                  if (p aconv h) then (p,TCs)::rst else x::insrt rst
   407             | insrt (x::rst) = x::insrt rst
   408             | insrt[] = raise TFL_ERR "merge.insert" "pattern not found"
   409       in insrt end
   410     fun pass ([],ptcl_final) = ptcl_final
   411       | pass (ptcs::tcl, ptcl) = pass(tcl, insert ptcs ptcl)
   412 in
   413   pass (TCs, map (fn p => (p,[])) full_pats)
   414 end;
   415 
   416 
   417 fun givens pats = map pat_of (filter given pats);
   418 
   419 fun post_definition meta_tflCongs (theory, (def, pats)) =
   420  let val tych = Thry.typecheck theory
   421      val f = #lhs(USyntax.dest_eq(concl def))
   422      val corollary = Rules.MATCH_MP Thms.WFREC_COROLLARY def
   423      val pats' = filter given pats
   424      val given_pats = map pat_of pats'
   425      val rows = map row_of_pat pats'
   426      val WFR = #ant(USyntax.dest_imp(concl corollary))
   427      val R = #Rand(USyntax.dest_comb WFR)
   428      val corollary' = Rules.UNDISCH corollary  (* put WF R on assums *)
   429      val corollaries = map (fn pat => Rules.SPEC (tych pat) corollary')
   430                            given_pats
   431      val (case_rewrites,context_congs) = extraction_thms theory
   432      (*case_ss causes minimal simplification: bodies of case expressions are
   433        not simplified. Otherwise large examples (Red-Black trees) are too
   434        slow.*)
   435      val case_ss = Simplifier.global_context theory
   436        (HOL_basic_ss addcongs
   437          (map (#weak_case_cong o snd) o Symtab.dest o Datatype.get_all) theory addsimps case_rewrites)
   438      val corollaries' = map (Simplifier.simplify case_ss) corollaries
   439      val extract = Rules.CONTEXT_REWRITE_RULE
   440                      (f, [R], @{thm cut_apply}, meta_tflCongs@context_congs)
   441      val (rules, TCs) = ListPair.unzip (map extract corollaries')
   442      val rules0 = map (rewrite_rule [Thms.CUT_DEF]) rules
   443      val mk_cond_rule = Rules.FILTER_DISCH_ALL(not o curry (op aconv) WFR)
   444      val rules1 = Rules.LIST_CONJ(map mk_cond_rule rules0)
   445  in
   446  {rules = rules1,
   447   rows = rows,
   448   full_pats_TCs = merge (map pat_of pats) (ListPair.zip (given_pats, TCs)),
   449   TCs = TCs}
   450  end;
   451 
   452 
   453 (*---------------------------------------------------------------------------
   454  * Perform the extraction without making the definition. Definition and
   455  * extraction commute for the non-nested case.  (Deferred recdefs)
   456  *
   457  * The purpose of wfrec_eqns is merely to instantiate the recursion theorem
   458  * and extract termination conditions: no definition is made.
   459  *---------------------------------------------------------------------------*)
   460 
   461 fun wfrec_eqns thy fid tflCongs eqns =
   462  let val {lhs,rhs} = USyntax.dest_eq (hd eqns)
   463      val (f,args) = USyntax.strip_comb lhs
   464      val (fname,fty) = dest_atom f
   465      val (SV,a) = front_last args    (* SV = schematic variables *)
   466      val g = list_comb(f,SV)
   467      val h = Free(fname,type_of g)
   468      val eqns1 = map (subst_free[(g,h)]) eqns
   469      val {functional as Abs(x, Ty, _),  pats} = mk_functional thy eqns1
   470      val given_pats = givens pats
   471      (* val f = Free(x,Ty) *)
   472      val Type("fun", [f_dty, f_rty]) = Ty
   473      val dummy = if x<>fid then
   474                         raise TFL_ERR "wfrec_eqns"
   475                                       ("Expected a definition of " ^
   476                                       quote fid ^ " but found one of " ^
   477                                       quote x)
   478                  else ()
   479      val (case_rewrites,context_congs) = extraction_thms thy
   480      val tych = Thry.typecheck thy
   481      val WFREC_THM0 = Rules.ISPEC (tych functional) Thms.WFREC_COROLLARY
   482      val Const(@{const_name All},_) $ Abs(Rname,Rtype,_) = concl WFREC_THM0
   483      val R = Free (Name.variant (List.foldr OldTerm.add_term_names [] eqns) Rname,
   484                    Rtype)
   485      val WFREC_THM = Rules.ISPECL [tych R, tych g] WFREC_THM0
   486      val ([proto_def, WFR],_) = USyntax.strip_imp(concl WFREC_THM)
   487      val dummy =
   488            if !trace then
   489                writeln ("ORIGINAL PROTO_DEF: " ^
   490                           Syntax.string_of_term_global thy proto_def)
   491            else ()
   492      val R1 = USyntax.rand WFR
   493      val corollary' = Rules.UNDISCH (Rules.UNDISCH WFREC_THM)
   494      val corollaries = map (fn pat => Rules.SPEC (tych pat) corollary') given_pats
   495      val corollaries' = map (rewrite_rule case_rewrites) corollaries
   496      fun extract X = Rules.CONTEXT_REWRITE_RULE
   497                        (f, R1::SV, @{thm cut_apply}, tflCongs@context_congs) X
   498  in {proto_def = proto_def,
   499      SV=SV,
   500      WFR=WFR,
   501      pats=pats,
   502      extracta = map extract corollaries'}
   503  end;
   504 
   505 
   506 (*---------------------------------------------------------------------------
   507  * Define the constant after extracting the termination conditions. The
   508  * wellfounded relation used in the definition is computed by using the
   509  * choice operator on the extracted conditions (plus the condition that
   510  * such a relation must be wellfounded).
   511  *---------------------------------------------------------------------------*)
   512 
   513 fun lazyR_def thy fid tflCongs eqns =
   514  let val {proto_def,WFR,pats,extracta,SV} =
   515            wfrec_eqns thy fid tflCongs eqns
   516      val R1 = USyntax.rand WFR
   517      val f = #lhs(USyntax.dest_eq proto_def)
   518      val (extractants,TCl) = ListPair.unzip extracta
   519      val dummy = if !trace
   520                  then writeln (cat_lines ("Extractants =" ::
   521                   map (Display.string_of_thm_global thy) extractants))
   522                  else ()
   523      val TCs = fold_rev (union (op aconv)) TCl []
   524      val full_rqt = WFR::TCs
   525      val R' = USyntax.mk_select{Bvar=R1, Body=USyntax.list_mk_conj full_rqt}
   526      val R'abs = USyntax.rand R'
   527      val proto_def' = subst_free[(R1,R')] proto_def
   528      val dummy = if !trace then writeln ("proto_def' = " ^
   529                                          Syntax.string_of_term_global
   530                                          thy proto_def')
   531                            else ()
   532      val {lhs,rhs} = USyntax.dest_eq proto_def'
   533      val (c,args) = USyntax.strip_comb lhs
   534      val (name,Ty) = dest_atom c
   535      val defn = const_def thy (name, Ty, USyntax.list_mk_abs (args,rhs))
   536      val ([def0], theory) =
   537        thy
   538        |> Global_Theory.add_defs false
   539             [Thm.no_attributes (Binding.name (fid ^ "_def"), defn)]
   540      val def = Thm.unvarify_global def0;
   541      val dummy =
   542        if !trace then writeln ("DEF = " ^ Display.string_of_thm_global theory def)
   543        else ()
   544      (* val fconst = #lhs(USyntax.dest_eq(concl def))  *)
   545      val tych = Thry.typecheck theory
   546      val full_rqt_prop = map (Dcterm.mk_prop o tych) full_rqt
   547          (*lcp: a lot of object-logic inference to remove*)
   548      val baz = Rules.DISCH_ALL
   549                  (fold_rev Rules.DISCH full_rqt_prop
   550                   (Rules.LIST_CONJ extractants))
   551      val dum = if !trace then writeln ("baz = " ^ Display.string_of_thm_global theory baz)
   552                            else ()
   553      val f_free = Free (fid, fastype_of f)  (*'cos f is a Const*)
   554      val SV' = map tych SV;
   555      val SVrefls = map Thm.reflexive SV'
   556      val def0 = (fold (fn x => fn th => Rules.rbeta(Thm.combination th x))
   557                    SVrefls def)
   558                 RS meta_eq_to_obj_eq
   559      val def' = Rules.MP (Rules.SPEC (tych R') (Rules.GEN (tych R1) baz)) def0
   560      val body_th = Rules.LIST_CONJ (map Rules.ASSUME full_rqt_prop)
   561      val SELECT_AX = (*in this way we hope to avoid a STATIC dependence upon
   562                        theory Hilbert_Choice*)
   563          ML_Context.thm "Hilbert_Choice.tfl_some"
   564          handle ERROR msg => cat_error msg
   565     "defer_recdef requires theory Main or at least Hilbert_Choice as parent"
   566      val bar = Rules.MP (Rules.ISPECL[tych R'abs, tych R1] SELECT_AX) body_th
   567  in {theory = theory, R=R1, SV=SV,
   568      rules = fold (Utils.C Rules.MP) (Rules.CONJUNCTS bar) def',
   569      full_pats_TCs = merge (map pat_of pats) (ListPair.zip (givens pats, TCl)),
   570      patterns = pats}
   571  end;
   572 
   573 
   574 
   575 (*----------------------------------------------------------------------------
   576  *
   577  *                           INDUCTION THEOREM
   578  *
   579  *---------------------------------------------------------------------------*)
   580 
   581 
   582 (*------------------------  Miscellaneous function  --------------------------
   583  *
   584  *           [x_1,...,x_n]     ?v_1...v_n. M[v_1,...,v_n]
   585  *     -----------------------------------------------------------
   586  *     ( M[x_1,...,x_n], [(x_i,?v_1...v_n. M[v_1,...,v_n]),
   587  *                        ...
   588  *                        (x_j,?v_n. M[x_1,...,x_(n-1),v_n])] )
   589  *
   590  * This function is totally ad hoc. Used in the production of the induction
   591  * theorem. The nchotomy theorem can have clauses that look like
   592  *
   593  *     ?v1..vn. z = C vn..v1
   594  *
   595  * in which the order of quantification is not the order of occurrence of the
   596  * quantified variables as arguments to C. Since we have no control over this
   597  * aspect of the nchotomy theorem, we make the correspondence explicit by
   598  * pairing the incoming new variable with the term it gets beta-reduced into.
   599  *---------------------------------------------------------------------------*)
   600 
   601 fun alpha_ex_unroll (xlist, tm) =
   602   let val (qvars,body) = USyntax.strip_exists tm
   603       val vlist = #2 (USyntax.strip_comb (USyntax.rhs body))
   604       val plist = ListPair.zip (vlist, xlist)
   605       val args = map (the o AList.lookup (op aconv) plist) qvars
   606                    handle Option => raise Fail "TFL.alpha_ex_unroll: no correspondence"
   607       fun build ex      []   = []
   608         | build (_$rex) (v::rst) =
   609            let val ex1 = Term.betapply(rex, v)
   610            in  ex1 :: build ex1 rst
   611            end
   612      val (nex::exl) = rev (tm::build tm args)
   613   in
   614   (nex, ListPair.zip (args, rev exl))
   615   end;
   616 
   617 
   618 
   619 (*----------------------------------------------------------------------------
   620  *
   621  *             PROVING COMPLETENESS OF PATTERNS
   622  *
   623  *---------------------------------------------------------------------------*)
   624 
   625 fun mk_case ty_info usednames thy =
   626  let
   627  val divide = ipartition (gvvariant usednames)
   628  val tych = Thry.typecheck thy
   629  fun tych_binding(x,y) = (tych x, tych y)
   630  fun fail s = raise TFL_ERR "mk_case" s
   631  fun mk{rows=[],...} = fail"no rows"
   632    | mk{path=[], rows = [([], (thm, bindings))]} =
   633                          Rules.IT_EXISTS (map tych_binding bindings) thm
   634    | mk{path = u::rstp, rows as (p::_, _)::_} =
   635      let val (pat_rectangle,rights) = ListPair.unzip rows
   636          val col0 = map hd pat_rectangle
   637          val pat_rectangle' = map tl pat_rectangle
   638      in
   639      if (forall is_Free col0) (* column 0 is all variables *)
   640      then let val rights' = map (fn ((thm,theta),v) => (thm,theta@[(u,v)]))
   641                                 (ListPair.zip (rights, col0))
   642           in mk{path = rstp, rows = ListPair.zip (pat_rectangle', rights')}
   643           end
   644      else                     (* column 0 is all constructors *)
   645      let val Type (ty_name,_) = type_of p
   646      in
   647      case (ty_info ty_name)
   648      of NONE => fail("Not a known datatype: "^ty_name)
   649       | SOME{constructors,nchotomy} =>
   650         let val thm' = Rules.ISPEC (tych u) nchotomy
   651             val disjuncts = USyntax.strip_disj (concl thm')
   652             val subproblems = divide(constructors, rows)
   653             val groups      = map #group subproblems
   654             and new_formals = map #new_formals subproblems
   655             val existentials = ListPair.map alpha_ex_unroll
   656                                    (new_formals, disjuncts)
   657             val constraints = map #1 existentials
   658             val vexl = map #2 existentials
   659             fun expnd tm (pats,(th,b)) = (pats, (Rules.SUBS [Rules.ASSUME (tych tm)] th, b))
   660             val news = map (fn (nf,rows,c) => {path = nf@rstp,
   661                                                rows = map (expnd c) rows})
   662                            (Utils.zip3 new_formals groups constraints)
   663             val recursive_thms = map mk news
   664             val build_exists = Library.foldr
   665                                 (fn((x,t), th) =>
   666                                  Rules.CHOOSE (tych x, Rules.ASSUME (tych t)) th)
   667             val thms' = ListPair.map build_exists (vexl, recursive_thms)
   668             val same_concls = Rules.EVEN_ORS thms'
   669         in Rules.DISJ_CASESL thm' same_concls
   670         end
   671      end end
   672  in mk
   673  end;
   674 
   675 
   676 fun complete_cases thy =
   677  let val tych = Thry.typecheck thy
   678      val ty_info = Thry.induct_info thy
   679  in fn pats =>
   680  let val names = List.foldr OldTerm.add_term_names [] pats
   681      val T = type_of (hd pats)
   682      val aname = Name.variant names "a"
   683      val vname = Name.variant (aname::names) "v"
   684      val a = Free (aname, T)
   685      val v = Free (vname, T)
   686      val a_eq_v = HOLogic.mk_eq(a,v)
   687      val ex_th0 = Rules.EXISTS (tych (USyntax.mk_exists{Bvar=v,Body=a_eq_v}), tych a)
   688                            (Rules.REFL (tych a))
   689      val th0 = Rules.ASSUME (tych a_eq_v)
   690      val rows = map (fn x => ([x], (th0,[]))) pats
   691  in
   692  Rules.GEN (tych a)
   693        (Rules.RIGHT_ASSOC
   694           (Rules.CHOOSE(tych v, ex_th0)
   695                 (mk_case ty_info (vname::aname::names)
   696                  thy {path=[v], rows=rows})))
   697  end end;
   698 
   699 
   700 (*---------------------------------------------------------------------------
   701  * Constructing induction hypotheses: one for each recursive call.
   702  *
   703  * Note. R will never occur as a variable in the ind_clause, because
   704  * to do so, it would have to be from a nested definition, and we don't
   705  * allow nested defns to have R variable.
   706  *
   707  * Note. When the context is empty, there can be no local variables.
   708  *---------------------------------------------------------------------------*)
   709 (*
   710 local infix 5 ==>
   711       fun (tm1 ==> tm2) = USyntax.mk_imp{ant = tm1, conseq = tm2}
   712 in
   713 fun build_ih f P (pat,TCs) =
   714  let val globals = USyntax.free_vars_lr pat
   715      fun nested tm = is_some (USyntax.find_term (curry (op aconv) f) tm)
   716      fun dest_TC tm =
   717          let val (cntxt,R_y_pat) = USyntax.strip_imp(#2(USyntax.strip_forall tm))
   718              val (R,y,_) = USyntax.dest_relation R_y_pat
   719              val P_y = if (nested tm) then R_y_pat ==> P$y else P$y
   720          in case cntxt
   721               of [] => (P_y, (tm,[]))
   722                | _  => let
   723                     val imp = USyntax.list_mk_conj cntxt ==> P_y
   724                     val lvs = gen_rems (op aconv) (USyntax.free_vars_lr imp, globals)
   725                     val locals = #2(Utils.pluck (curry (op aconv) P) lvs) handle Utils.ERR _ => lvs
   726                     in (USyntax.list_mk_forall(locals,imp), (tm,locals)) end
   727          end
   728  in case TCs
   729     of [] => (USyntax.list_mk_forall(globals, P$pat), [])
   730      |  _ => let val (ihs, TCs_locals) = ListPair.unzip(map dest_TC TCs)
   731                  val ind_clause = USyntax.list_mk_conj ihs ==> P$pat
   732              in (USyntax.list_mk_forall(globals,ind_clause), TCs_locals)
   733              end
   734  end
   735 end;
   736 *)
   737 
   738 local infix 5 ==>
   739       fun (tm1 ==> tm2) = USyntax.mk_imp{ant = tm1, conseq = tm2}
   740 in
   741 fun build_ih f (P,SV) (pat,TCs) =
   742  let val pat_vars = USyntax.free_vars_lr pat
   743      val globals = pat_vars@SV
   744      fun nested tm = is_some (USyntax.find_term (curry (op aconv) f) tm)
   745      fun dest_TC tm =
   746          let val (cntxt,R_y_pat) = USyntax.strip_imp(#2(USyntax.strip_forall tm))
   747              val (R,y,_) = USyntax.dest_relation R_y_pat
   748              val P_y = if (nested tm) then R_y_pat ==> P$y else P$y
   749          in case cntxt
   750               of [] => (P_y, (tm,[]))
   751                | _  => let
   752                     val imp = USyntax.list_mk_conj cntxt ==> P_y
   753                     val lvs = subtract (op aconv) globals (USyntax.free_vars_lr imp)
   754                     val locals = #2(Utils.pluck (curry (op aconv) P) lvs) handle Utils.ERR _ => lvs
   755                     in (USyntax.list_mk_forall(locals,imp), (tm,locals)) end
   756          end
   757  in case TCs
   758     of [] => (USyntax.list_mk_forall(pat_vars, P$pat), [])
   759      |  _ => let val (ihs, TCs_locals) = ListPair.unzip(map dest_TC TCs)
   760                  val ind_clause = USyntax.list_mk_conj ihs ==> P$pat
   761              in (USyntax.list_mk_forall(pat_vars,ind_clause), TCs_locals)
   762              end
   763  end
   764 end;
   765 
   766 (*---------------------------------------------------------------------------
   767  * This function makes good on the promise made in "build_ih".
   768  *
   769  * Input  is tm = "(!y. R y pat ==> P y) ==> P pat",
   770  *           TCs = TC_1[pat] ... TC_n[pat]
   771  *           thm = ih1 /\ ... /\ ih_n |- ih[pat]
   772  *---------------------------------------------------------------------------*)
   773 fun prove_case f thy (tm,TCs_locals,thm) =
   774  let val tych = Thry.typecheck thy
   775      val antc = tych(#ant(USyntax.dest_imp tm))
   776      val thm' = Rules.SPEC_ALL thm
   777      fun nested tm = is_some (USyntax.find_term (curry (op aconv) f) tm)
   778      fun get_cntxt TC = tych(#ant(USyntax.dest_imp(#2(USyntax.strip_forall(concl TC)))))
   779      fun mk_ih ((TC,locals),th2,nested) =
   780          Rules.GENL (map tych locals)
   781             (if nested then Rules.DISCH (get_cntxt TC) th2 handle Utils.ERR _ => th2
   782              else if USyntax.is_imp (concl TC) then Rules.IMP_TRANS TC th2
   783              else Rules.MP th2 TC)
   784  in
   785  Rules.DISCH antc
   786  (if USyntax.is_imp(concl thm') (* recursive calls in this clause *)
   787   then let val th1 = Rules.ASSUME antc
   788            val TCs = map #1 TCs_locals
   789            val ylist = map (#2 o USyntax.dest_relation o #2 o USyntax.strip_imp o
   790                             #2 o USyntax.strip_forall) TCs
   791            val TClist = map (fn(TC,lvs) => (Rules.SPEC_ALL(Rules.ASSUME(tych TC)),lvs))
   792                             TCs_locals
   793            val th2list = map (Utils.C Rules.SPEC th1 o tych) ylist
   794            val nlist = map nested TCs
   795            val triples = Utils.zip3 TClist th2list nlist
   796            val Pylist = map mk_ih triples
   797        in Rules.MP thm' (Rules.LIST_CONJ Pylist) end
   798   else thm')
   799  end;
   800 
   801 
   802 (*---------------------------------------------------------------------------
   803  *
   804  *         x = (v1,...,vn)  |- M[x]
   805  *    ---------------------------------------------
   806  *      ?v1 ... vn. x = (v1,...,vn) |- M[x]
   807  *
   808  *---------------------------------------------------------------------------*)
   809 fun LEFT_ABS_VSTRUCT tych thm =
   810   let fun CHOOSER v (tm,thm) =
   811         let val ex_tm = USyntax.mk_exists{Bvar=v,Body=tm}
   812         in (ex_tm, Rules.CHOOSE(tych v, Rules.ASSUME (tych ex_tm)) thm)
   813         end
   814       val [veq] = filter (can USyntax.dest_eq) (#1 (Rules.dest_thm thm))
   815       val {lhs,rhs} = USyntax.dest_eq veq
   816       val L = USyntax.free_vars_lr rhs
   817   in  #2 (fold_rev CHOOSER L (veq,thm))  end;
   818 
   819 
   820 (*----------------------------------------------------------------------------
   821  * Input : f, R,  and  [(pat1,TCs1),..., (patn,TCsn)]
   822  *
   823  * Instantiates WF_INDUCTION_THM, getting Sinduct and then tries to prove
   824  * recursion induction (Rinduct) by proving the antecedent of Sinduct from
   825  * the antecedent of Rinduct.
   826  *---------------------------------------------------------------------------*)
   827 fun mk_induction thy {fconst, R, SV, pat_TCs_list} =
   828 let val tych = Thry.typecheck thy
   829     val Sinduction = Rules.UNDISCH (Rules.ISPEC (tych R) Thms.WF_INDUCTION_THM)
   830     val (pats,TCsl) = ListPair.unzip pat_TCs_list
   831     val case_thm = complete_cases thy pats
   832     val domain = (type_of o hd) pats
   833     val Pname = Name.variant (List.foldr (Library.foldr OldTerm.add_term_names)
   834                               [] (pats::TCsl)) "P"
   835     val P = Free(Pname, domain --> HOLogic.boolT)
   836     val Sinduct = Rules.SPEC (tych P) Sinduction
   837     val Sinduct_assumf = USyntax.rand ((#ant o USyntax.dest_imp o concl) Sinduct)
   838     val Rassums_TCl' = map (build_ih fconst (P,SV)) pat_TCs_list
   839     val (Rassums,TCl') = ListPair.unzip Rassums_TCl'
   840     val Rinduct_assum = Rules.ASSUME (tych (USyntax.list_mk_conj Rassums))
   841     val cases = map (fn pat => Term.betapply (Sinduct_assumf, pat)) pats
   842     val tasks = Utils.zip3 cases TCl' (Rules.CONJUNCTS Rinduct_assum)
   843     val proved_cases = map (prove_case fconst thy) tasks
   844     val v = Free (Name.variant (List.foldr OldTerm.add_term_names [] (map concl proved_cases))
   845                     "v",
   846                   domain)
   847     val vtyped = tych v
   848     val substs = map (Rules.SYM o Rules.ASSUME o tych o (curry HOLogic.mk_eq v)) pats
   849     val proved_cases1 = ListPair.map (fn (th,th') => Rules.SUBS[th]th')
   850                           (substs, proved_cases)
   851     val abs_cases = map (LEFT_ABS_VSTRUCT tych) proved_cases1
   852     val dant = Rules.GEN vtyped (Rules.DISJ_CASESL (Rules.ISPEC vtyped case_thm) abs_cases)
   853     val dc = Rules.MP Sinduct dant
   854     val Parg_ty = type_of(#Bvar(USyntax.dest_forall(concl dc)))
   855     val vars = map (gvvariant[Pname]) (USyntax.strip_prod_type Parg_ty)
   856     val dc' = fold_rev (Rules.GEN o tych) vars
   857                        (Rules.SPEC (tych(USyntax.mk_vstruct Parg_ty vars)) dc)
   858 in
   859    Rules.GEN (tych P) (Rules.DISCH (tych(concl Rinduct_assum)) dc')
   860 end
   861 handle Utils.ERR _ => raise TFL_ERR "mk_induction" "failed derivation";
   862 
   863 
   864 
   865 
   866 (*---------------------------------------------------------------------------
   867  *
   868  *                        POST PROCESSING
   869  *
   870  *---------------------------------------------------------------------------*)
   871 
   872 
   873 fun simplify_induction thy hth ind =
   874   let val tych = Thry.typecheck thy
   875       val (asl,_) = Rules.dest_thm ind
   876       val (_,tc_eq_tc') = Rules.dest_thm hth
   877       val tc = USyntax.lhs tc_eq_tc'
   878       fun loop [] = ind
   879         | loop (asm::rst) =
   880           if (can (Thry.match_term thy asm) tc)
   881           then Rules.UNDISCH
   882                  (Rules.MATCH_MP
   883                      (Rules.MATCH_MP Thms.simp_thm (Rules.DISCH (tych asm) ind))
   884                      hth)
   885          else loop rst
   886   in loop asl
   887 end;
   888 
   889 
   890 (*---------------------------------------------------------------------------
   891  * The termination condition is an antecedent to the rule, and an
   892  * assumption to the theorem.
   893  *---------------------------------------------------------------------------*)
   894 fun elim_tc tcthm (rule,induction) =
   895    (Rules.MP rule tcthm, Rules.PROVE_HYP tcthm induction)
   896 
   897 
   898 fun trace_thms s L =
   899   if !trace then writeln (cat_lines (s :: map Display.string_of_thm_without_context L))
   900   else ();
   901 
   902 fun trace_cterm s ct =
   903   if !trace then
   904     writeln (cat_lines [s, Syntax.string_of_term_global (Thm.theory_of_cterm ct) (Thm.term_of ct)])
   905   else ();
   906 
   907 
   908 fun postprocess strict {wf_tac, terminator, simplifier} theory {rules,induction,TCs} =
   909 let val tych = Thry.typecheck theory
   910     val prove = Rules.prove strict;
   911 
   912    (*---------------------------------------------------------------------
   913     * Attempt to eliminate WF condition. It's the only assumption of rules
   914     *---------------------------------------------------------------------*)
   915    val (rules1,induction1)  =
   916        let val thm = prove(tych(HOLogic.mk_Trueprop
   917                                   (hd(#1(Rules.dest_thm rules)))),
   918                              wf_tac)
   919        in (Rules.PROVE_HYP thm rules, Rules.PROVE_HYP thm induction)
   920        end handle Utils.ERR _ => (rules,induction);
   921 
   922    (*----------------------------------------------------------------------
   923     * The termination condition (tc) is simplified to |- tc = tc' (there
   924     * might not be a change!) and then 3 attempts are made:
   925     *
   926     *   1. if |- tc = T, then eliminate it with eqT; otherwise,
   927     *   2. apply the terminator to tc'. If |- tc' = T then eliminate; else
   928     *   3. replace tc by tc' in both the rules and the induction theorem.
   929     *---------------------------------------------------------------------*)
   930 
   931    fun simplify_tc tc (r,ind) =
   932        let val tc1 = tych tc
   933            val _ = trace_cterm "TC before simplification: " tc1
   934            val tc_eq = simplifier tc1
   935            val _ = trace_thms "result: " [tc_eq]
   936        in
   937        elim_tc (Rules.MATCH_MP Thms.eqT tc_eq) (r,ind)
   938        handle Utils.ERR _ =>
   939         (elim_tc (Rules.MATCH_MP(Rules.MATCH_MP Thms.rev_eq_mp tc_eq)
   940                   (prove(tych(HOLogic.mk_Trueprop(USyntax.rhs(concl tc_eq))),
   941                            terminator)))
   942                  (r,ind)
   943          handle Utils.ERR _ =>
   944           (Rules.UNDISCH(Rules.MATCH_MP (Rules.MATCH_MP Thms.simp_thm r) tc_eq),
   945            simplify_induction theory tc_eq ind))
   946        end
   947 
   948    (*----------------------------------------------------------------------
   949     * Nested termination conditions are harder to get at, since they are
   950     * left embedded in the body of the function (and in induction
   951     * theorem hypotheses). Our "solution" is to simplify them, and try to
   952     * prove termination, but leave the application of the resulting theorem
   953     * to a higher level. So things go much as in "simplify_tc": the
   954     * termination condition (tc) is simplified to |- tc = tc' (there might
   955     * not be a change) and then 2 attempts are made:
   956     *
   957     *   1. if |- tc = T, then return |- tc; otherwise,
   958     *   2. apply the terminator to tc'. If |- tc' = T then return |- tc; else
   959     *   3. return |- tc = tc'
   960     *---------------------------------------------------------------------*)
   961    fun simplify_nested_tc tc =
   962       let val tc_eq = simplifier (tych (#2 (USyntax.strip_forall tc)))
   963       in
   964       Rules.GEN_ALL
   965        (Rules.MATCH_MP Thms.eqT tc_eq
   966         handle Utils.ERR _ =>
   967           (Rules.MATCH_MP(Rules.MATCH_MP Thms.rev_eq_mp tc_eq)
   968                       (prove(tych(HOLogic.mk_Trueprop (USyntax.rhs(concl tc_eq))),
   969                                terminator))
   970             handle Utils.ERR _ => tc_eq))
   971       end
   972 
   973    (*-------------------------------------------------------------------
   974     * Attempt to simplify the termination conditions in each rule and
   975     * in the induction theorem.
   976     *-------------------------------------------------------------------*)
   977    fun strip_imp tm = if USyntax.is_neg tm then ([],tm) else USyntax.strip_imp tm
   978    fun loop ([],extras,R,ind) = (rev R, ind, extras)
   979      | loop ((r,ftcs)::rst, nthms, R, ind) =
   980         let val tcs = #1(strip_imp (concl r))
   981             val extra_tcs = subtract (op aconv) tcs ftcs
   982             val extra_tc_thms = map simplify_nested_tc extra_tcs
   983             val (r1,ind1) = fold simplify_tc tcs (r,ind)
   984             val r2 = Rules.FILTER_DISCH_ALL(not o USyntax.is_WFR) r1
   985         in loop(rst, nthms@extra_tc_thms, r2::R, ind1)
   986         end
   987    val rules_tcs = ListPair.zip (Rules.CONJUNCTS rules1, TCs)
   988    val (rules2,ind2,extras) = loop(rules_tcs,[],[],induction1)
   989 in
   990   {induction = ind2, rules = Rules.LIST_CONJ rules2, nested_tcs = extras}
   991 end;
   992 
   993 
   994 end;