src/Tools/isac/Interpret/script.sml
author Walther Neuper <wneuper@ist.tugraz.at>
Sat, 12 Nov 2016 17:21:43 +0100
changeset 59257 a1daf71787b1
parent 59253 f0bb15a046ae
child 59258 6b1aad933adb
permissions -rw-r--r--
--- polished LUCAS_INTERPRETER

Notes
# "---" marks changesets where Test_Isac does not work
# outcommented tests in script.sml moved from src/ to test/
     1 (* Title:  interpreter for scripts
     2    Author: Walther Neuper 2000
     3    (c) due to copyright terms
     4 12345678901234567890123456789012345678901234567890123456789012345678901234567890
     5         10        20        30        40        50        60        70        80
     6 *)
     7 
     8 signature LUCAS_INTERPRETER =
     9 sig
    10 
    11   type step = tac_ * mout * ptree * pos' * pos' list
    12   datatype locate = NotLocatable | Steps of istate * step list
    13   
    14   val next_tac : (*diss: next-tactic-function*)
    15     theory' * rls -> ptree * pos' -> scr -> istate * 'a -> tac_ * (istate * 'a) * (term * safe)
    16   val locate_gen : (*diss: locate-function*)
    17     theory' * rls -> tac_ -> ptree * pos' -> scr * 'a -> istate * Proof.context -> locate
    18   val sel_rules : ptree -> pos' -> tac list 
    19   val init_form : 'a -> scr -> (term * term) list -> term option
    20   val formal_args : term -> term list
    21 
    22   val tac_2tac : tac_ -> tac
    23   val init_scrstate : theory -> itm list -> metID -> istate * Proof.context * scr
    24   val from_pblobj' : theory' -> pos * pos_ -> ptree -> rls * (istate * Proof.context) * scr
    25   val from_pblobj_or_detail' : theory' -> pos * pos_ -> ptree -> 
    26     rls * (istate * Proof.context) * scr
    27   val rule2thm'' : rule -> thm''
    28   val rule2rls' : rule -> string
    29 
    30   val itms2args : 'a -> metID -> itm list -> term list
    31 end 
    32 
    33 (**)
    34 structure Lucin: LUCAS_INTERPRETER(**) =
    35 struct
    36 (**)
    37 (* traces the leaves (ie. non-tactical nodes) of Prog found by next_tac, see "and scr" *)   
    38 val trace_script = Unsynchronized.ref false;
    39 
    40 (* data for creating a new node in ctree; designed for use as:
    41    fun ass* scrstate steps = / ... case ass* scrstate steps of /
    42    Assoc (scrstate, steps) => ... ass* scrstate steps *)
    43 type step =
    44     tac_         (*transformed from associated tac                   *)
    45     * mout       (*result with indentation etc.                      *)
    46     * ptree      (*containing node created by tac_ + resp. scrstate  *)
    47     * pos'       (*position in ptree; ptree * pos' is the proofstate *)
    48     * pos' list; (*of ptree-nodes probably cut (by fst tac_)         *)
    49 
    50 fun rule2thm'' (Thm (id, thm)) = (id, thm) : thm''
    51   | rule2thm'' r = error ("rule2thm': not defined for " ^ rule2str r);
    52 fun rule2rls' (Rls_ rls) = id_rls rls
    53   | rule2rls' r = error ("rule2rls': not defined for " ^ rule2str r);
    54 
    55 (*.makes a (rule,term) list to a Step (m, mout, pt', p', cid) for solve;
    56    complicated with current t in rrlsstate.*)
    57 fun rts2steps steps ((pt, p), (f, f'', rss, rts), (thy', ro, er, pa)) [(r, (f', am))] =
    58       let
    59         val thy = assoc_thy thy'
    60         val ctxt = get_ctxt pt p |> insert_assumptions am
    61 	      val m = Rewrite' (thy', ro, er, pa, rule2thm'' r, f, (f', am))
    62 	      val is = RrlsState (f', f'', rss, rts)
    63 	      val p = case p of (_, Frm) => p | (p', Res) => (lev_on p', Res) | _ => error "rts2steps: p1"
    64 	      val (p', cid, mout, pt') = generate1 thy m (is, ctxt) p pt
    65       in (is, (m, mout, pt', p', cid) :: steps) end
    66   | rts2steps steps ((pt, p) ,(f, f'', rss, rts), (thy', ro, er, pa)) ((r, (f', am)) :: rts') =
    67       let
    68         val thy = assoc_thy thy'
    69         val ctxt = get_ctxt pt p |> insert_assumptions am
    70 	      val m = Rewrite' (thy', ro, er, pa, rule2thm'' r, f, (f', am))
    71 	      val is = RrlsState (f', f'', rss, rts)
    72 	      val p = case p of (_, Frm) => p | (p', Res) => (lev_on p', Res) | _ => error "rts2steps: p1"
    73 	      val (p', cid, mout, pt') = generate1 thy m (is, ctxt) p pt
    74       in rts2steps ((m, mout, pt', p', cid)::steps) 
    75 		    ((pt', p'), (f', f'', rss, rts), (thy', ro, er, pa)) rts'
    76 		  end
    77   | rts2steps _ _ _ = error "rts2steps: uncovered fun-def"
    78 
    79 (* functions for the environment stack: NOT YET IMPLEMENTED
    80 fun accessenv id es = the (assoc ((top es) : env, id))
    81     handle _ => error ("accessenv: " ^ free2str id ^ " not in env");
    82 fun updateenv id vl (es : env stack) = 
    83     (push (overwrite(top es, (id, vl))) (pop es)) : env stack;
    84 fun pushenv id vl (es : env stack) = 
    85     (push (overwrite(top es, (id, vl))) es) : env stack;
    86 val popenv = pop : env stack -> env stack;
    87 *)
    88 
    89 fun de_esc_underscore str =
    90   let
    91     fun scan [] = []
    92     | scan (s :: ss) = if s = "'" then (scan ss) else (s :: (scan ss))
    93   in (implode o scan o Symbol.explode) str end;
    94 
    95 (*go at a location in a script and fetch the contents*)
    96 fun go [] t = t
    97   | go (D::p) (Abs(_, _, t0)) = go (p : loc_) t0
    98   | go (L::p) (t1 $ _) = go p t1
    99   | go (R::p) (_ $ t2) = go p t2
   100   | go l _ = error ("go: no " ^ loc_2str l);
   101 
   102 (*.get argument of first stactic in a script for init_form.*)
   103 fun get_stac thy (_ $ body) =
   104   let
   105     fun get_t y (Const ("Script.Seq",_) $ e1 $ e2) a = 
   106     	  (case get_t y e1 a of NONE => get_t y e2 a | la => la)
   107       | get_t y (Const ("Script.Seq",_) $ e1 $ e2 $ a) _ = 
   108     	  (case get_t y e1 a of NONE => get_t y e2 a | la => la)
   109       | get_t y (Const ("Script.Try",_) $ e) a = get_t y e a
   110       | get_t y (Const ("Script.Try",_) $ e $ a) _ = get_t y e a
   111       | get_t y (Const ("Script.Repeat",_) $ e) a = get_t y e a
   112       | get_t y (Const ("Script.Repeat",_) $ e $ a) _ = get_t y e a
   113       | get_t y (Const ("Script.Or",_) $e1 $ e2) a =
   114     	  (case get_t y e1 a of NONE => get_t y e2 a | la => la)
   115       | get_t y (Const ("Script.Or",_) $e1 $ e2 $ a) _ =
   116     	  (case get_t y e1 a of NONE => get_t y e2 a | la => la)
   117       | get_t y (Const ("Script.While",_) $ _ $ e) a = get_t y e a
   118       | get_t y (Const ("Script.While",_) $ _ $ e $ a) _ = get_t y e a
   119       | get_t y (Const ("Script.Letpar",_) $ e1 $ Abs (_, _, e2)) a = 
   120     	  (case get_t y e1 a of NONE => get_t y e2 a | la => la)
   121     (*| get_t y (Const ("HOL.Let",_) $ e1 $ Abs (_,_,e2)) a =
   122 	      (case get_t y e1 a of NONE => get_t y e2 a | la => la)
   123       | get_t y (Abs (_,_,e)) a = get_t y e a*)
   124       | get_t y (Const ("HOL.Let",_) $ e1 $ Abs (_, _, _)) a =
   125     	get_t y e1 a (*don't go deeper without evaluation !*)
   126       | get_t _ (Const ("If", _) $ _ $ _ $ _) _ = NONE
   127     	(*(case get_t y e1 a of NONE => get_t y e2 a | la => la)*)
   128     
   129       | get_t _ (Const ("Script.Rewrite",_) $ _ $ _ $ a) _ = SOME a
   130       | get_t _ (Const ("Script.Rewrite",_) $ _ $ _    ) a = SOME a
   131       | get_t _ (Const ("Script.Rewrite'_Inst",_) $ _ $ _ $ _ $ a) _ = SOME a
   132       | get_t _ (Const ("Script.Rewrite'_Inst",_) $ _ $ _ $ _ )    a = SOME a
   133       | get_t _ (Const ("Script.Rewrite'_Set",_) $ _ $ _ $ a) _ = SOME a
   134       | get_t _ (Const ("Script.Rewrite'_Set",_) $ _ $ _ )    a = SOME a
   135       | get_t _ (Const ("Script.Rewrite'_Set'_Inst",_) $ _ $ _ $ _ $a)_ =SOME a
   136       | get_t _ (Const ("Script.Rewrite'_Set'_Inst",_) $ _ $ _ $ _ )  a =SOME a
   137       | get_t _ (Const ("Script.Calculate",_) $ _ $ a) _ = SOME a
   138       | get_t _ (Const ("Script.Calculate",_) $ _ )    a = SOME a
   139     
   140       | get_t _ (Const ("Script.Substitute",_) $ _ $ a) _ = SOME a
   141       | get_t _ (Const ("Script.Substitute",_) $ _ )    a = SOME a
   142     
   143       | get_t _ (Const ("Script.SubProblem",_) $ _ $ _) _ = NONE
   144 
   145       | get_t _ _ _ = ((*tracing ("### get_t yac: list-expr "^(term2str x));*) NONE)
   146     in get_t thy body e_term end
   147   | get_stac _ t = error ("get_stac: no fun-def. for " ^ term2str t);
   148     
   149 fun init_form thy (Prog sc) env =
   150     (case get_stac thy sc of NONE => NONE | SOME stac => SOME (subst_atomic env stac))
   151   | init_form _ _ _ = error "init_form: no match";
   152 
   153 (* get the arguments of the script out of the scripts parsetree *)
   154 fun formal_args scr = (fst o split_last o snd o strip_comb) scr;
   155 
   156 (* get the identifier of the script out of the scripts parsetree *)
   157 fun id_of_scr sc = (id_of o fst o strip_comb) sc;
   158 
   159 (*WN020526: not clear, when a is available in ass_up for eval_true*)
   160 (*WN060906: in "fun handle_leaf" eg. uses "SOME M__"(from some PREVIOUS
   161   curried Rewrite) for CURRENT value (which may be different from PREVIOUS);
   162   thus "NONE" must be set at the end of currying (ill designed anyway)*)
   163 fun upd_env_opt env (SOME a, v) = upd_env env (a, v)
   164   | upd_env_opt env (NONE, _) = 
   165       ((*tracing ("*** upd_env_opt: (NONE," ^ term2str v ^ ")");*) env);
   166 
   167 type dsc = typ; (* <-> nam..unknow in Descript.thy *)
   168 
   169 (*.create the actual parameters (args) of script: their order 
   170   is given by the order in met.pat .*)
   171 (*WN.5.5.03: ?: does this allow for different descriptions ???
   172              ?: why not taken from formal args of script ???
   173 !: FIXXXME penv: push it here in itms2args into script-evaluation*)
   174 (* val (thy, mI, itms) = (thy, metID, itms);
   175    *)
   176 val errmsg = "ERROR: the guard is missing (#ppc in 'type met' added in prep_met)."
   177 fun itms2args _ mI (itms : itm list) =
   178   let
   179     val mvat = max_vt itms
   180     fun okv mvat (_, vats, b, _, _) = member op = vats mvat andalso b
   181     val itms = filter (okv mvat) itms
   182     fun test_dsc d (_, _, _, _, itm_) = (d = d_in itm_)
   183     fun itm2arg itms (_,(d,_)) =
   184         case find_first (test_dsc d) itms of
   185           NONE => error ("itms2args: '" ^ term2str d ^ "' not in itms")
   186         | SOME (_, _, _, _, itm_) => penvval_in itm_
   187       (*| SOME (_,_,_,_,itm_) => mk_arg thy (d_in itm_) (ts_in itm_);
   188             penv postponed; presently penv holds already env for script*)
   189     val pats = (#ppc o get_met) mI
   190     val _ = if pats = [] then raise ERROR errmsg else ()
   191   in (flat o (map (itm2arg itms))) pats end;
   192 
   193 (* convert a script-tac 'stac' to a tactic 'tac';
   194    if stac is an initac, then convert to a 'tac_' (as required in appy).
   195    arg ptree for pushing the thy specified in rootpbl into subpbls    *)
   196 fun stac2tac_ _ thy (Const ("Script.Rewrite", _) $ Free (thmID, _) $ _ $ _) =
   197     let
   198       val tid = (de_esc_underscore o strip_thy) thmID
   199     in (Rewrite (tid, assoc_thm'' thy tid), Empty_Tac_) end
   200   | stac2tac_ _ thy (Const ("Script.Rewrite'_Inst", _) $ sub $ Free (thmID, _) $ _ $ _) =
   201     let
   202       val subML = ((map isapair2pair) o isalist2list) sub
   203       val subStr = subst2subs subML
   204       val tid = (de_esc_underscore o strip_thy) thmID (*4.10.02 unnoetig*)
   205     in (Rewrite_Inst (subStr, (tid, assoc_thm'' thy tid)), Empty_Tac_) end
   206   | stac2tac_ _ _ (Const ("Script.Rewrite'_Set",_) $ Free (rls, _) $ _ $ _) =
   207      (Rewrite_Set ((de_esc_underscore o strip_thy) rls), Empty_Tac_)
   208   | stac2tac_ _ _ (Const ("Script.Rewrite'_Set'_Inst", _) $ sub $ Free (rls, _) $ _ $ _) =
   209     let
   210       val subML = ((map isapair2pair) o isalist2list) sub;
   211       val subStr = subst2subs subML;
   212     in (Rewrite_Set_Inst (subStr, rls), Empty_Tac_) end
   213   | stac2tac_ _ _ (Const ("Script.Calculate", _) $ Free (op_, _) $ _) = (Calculate op_, Empty_Tac_)
   214   | stac2tac_ _ _ (Const ("Script.Take", _) $ t) = (Take (term2str t), Empty_Tac_)
   215   | stac2tac_ _ _ (Const ("Script.Substitute", _) $ isasub $ _) =
   216     (Substitute ((subte2sube o isalist2list) isasub), Empty_Tac_)
   217   | stac2tac_ _ thy (Const("Script.Check'_elementwise", _) $ _ $ 
   218     (Const ("Set.Collect", _) $ Abs (_, _, pred))) =
   219       (Check_elementwise (term_to_string''' thy pred), Empty_Tac_)
   220   | stac2tac_ _ _ (Const("Script.Or'_to'_List", _) $ _ ) = (Or_to_List, Empty_Tac_)
   221   | stac2tac_ _ _ (Const ("Script.Tac", _) $ Free (str, _)) = 
   222     (Tac ((de_esc_underscore o strip_thy) str),  Empty_Tac_) 
   223 
   224     (*compare "| assod _ (Subproblem'"*)
   225   | stac2tac_ pt _ (stac as Const ("Script.SubProblem",_) $
   226 	  (Const ("Product_Type.Pair",_) $Free (dI', _) $ (Const ("Product_Type.Pair", _) $ pI' $ mI')) $ 
   227 	    ags') =
   228     let
   229       val dI = ((implode o drop_last(*.."'"*) o Symbol.explode) dI')(*^""*);
   230       val thy = maxthy (assoc_thy dI) (rootthy pt);
   231 	    val pI = ((map (de_esc_underscore o free2str)) o isalist2list) pI';
   232 	    val mI = ((map (de_esc_underscore o free2str)) o isalist2list) mI';
   233 	    val ags = isalist2list ags';
   234 	    val (pI, pors, mI) = 
   235 	      if mI = ["no_met"] 
   236 	      then
   237           let
   238             val pors = (match_ags thy ((#ppc o get_pbt) pI) ags)
   239 		          handle ERROR "actual args do not match formal args" 
   240 			        => (match_ags_msg pI stac ags(*raise exn*); [])
   241 		        val pI' = refine_ori' pors pI;
   242 		      in (pI', pors (* refinement over models with diff.prec only *), 
   243 		          (hd o #met o get_pbt) pI') end
   244 	      else (pI, (match_ags thy ((#ppc o get_pbt) pI) ags)
   245 		      handle ERROR "actual args do not match formal args"
   246 		      => (match_ags_msg pI stac ags(*raise exn*); []), mI);
   247       val (fmz_, vals) = oris2fmz_vals pors;
   248 	    val {cas,ppc,thy,...} = get_pbt pI
   249 	    val dI = theory2theory' thy (*.take dI from _refined_ pbl.*)
   250 	    val dI = theory2theory' (maxthy (assoc_thy dI) (rootthy pt));
   251       val ctxt = dI |> Thy_Info.get_theory |> Proof_Context.init_global |> declare_constraints' vals
   252 	    val hdl =
   253         case cas of
   254 		      NONE => pblterm dI pI
   255 		    | SOME t => subst_atomic ((vars_of_pbl_' ppc) ~~~ vals) t
   256       val f = subpbl (strip_thy dI) pI
   257     in (Subproblem (dI, pI),	Subproblem' ((dI, pI, mI), pors, hdl, fmz_, ctxt, f))
   258     end
   259   | stac2tac_ _ thy t = error ("stac2tac_ TODO: no match for " ^ term_to_string''' thy t);
   260 
   261 fun stac2tac pt thy t = (fst o stac2tac_ pt thy) t;
   262 
   263 datatype ass = 
   264     Ass of
   265       tac_ *   (* SubProblem gets args instantiated in assod *)
   266   	  term     (* for itr_arg, result in ets *)
   267   | AssWeak of
   268       tac_ *
   269   	  term     (*for itr_arg,result in ets*)
   270   | NotAss;
   271 
   272 (* check if tac_ is associated with stac.
   273    Additional task: check if term t (the result has been calculated from) in tac_
   274    has been changed (see "datatype tac_"); if yes, recalculate result
   275    TODO.WN120106 recalculate impl.only for Substitute'
   276 args
   277   pt     : ptree for pushing the thy specified in rootpbl into subpbls
   278   d      : unused (planned for data for comparison)
   279   tac_   : from user (via applicable_in); to be compared with ...
   280   stac   : found in Script
   281 returns
   282   Ass    : associated: e.g. thmID in stac = thmID in m
   283                        +++ arg   in stac = arg   in m
   284   AssWeak: weakly ass.:e.g. thmID in stac = thmID in m, //arg//
   285   NotAss :             e.g. thmID in stac/=/thmID in m (not =)
   286 *)
   287 fun assod _ _ (m as Rewrite_Inst' (_, _, _, _, _, thm'' as (thmID, _), f, (f', _))) stac =
   288     (case stac of
   289 	    (Const ("Script.Rewrite'_Inst", _) $ _ $ Free (thmID_, _) $ _ $ f_) =>
   290 	      if thmID = thmID_
   291         then 
   292 	        if f = f_ 
   293           then ((*tracing"3### assod ..Ass";*) Ass (m,f')) 
   294 	        else ((*tracing"3### assod ..AssWeak";*) AssWeak(m, f'))
   295 	      else ((*tracing"3### assod ..NotAss";*) NotAss)
   296     | (Const ("Script.Rewrite'_Set'_Inst",_) $ _ $ Free (rls_, _) $ _ $ f_) =>
   297 	      if contains_rule (Thm thm'') (assoc_rls rls_)
   298         then if f = f_ then Ass (m,f') else AssWeak (m,f')
   299 	      else NotAss
   300     | _ => NotAss)
   301   | assod _ _ (m as Rewrite' (_, _, _, _, thm'' as (thmID, _), f, (f', _))) stac =
   302     (case stac of
   303 	    (Const ("Script.Rewrite", _) $ Free (thmID_, _) $ _ $ f_) =>
   304 	      ((*tracing ("3### assod: stac = " ^ ter2str t);
   305 	       tracing ("3### assod: f(m)= " ^ term2str f);*)
   306 	      if thmID = thmID_
   307         then 
   308 	        if f = f_
   309           then ((*tracing"3### assod ..Ass";*) Ass (m,f')) 
   310 	        else 
   311             ((*tracing"### assod ..AssWeak";
   312 		         tracing("### assod: f(m)  = " ^ term2str f);
   313 		         tracing("### assod: f(stac)= " ^ term2str f_)*)
   314 		         AssWeak (m,f'))
   315 	      else ((*tracing"3### assod ..NotAss";*) NotAss))
   316     | (Const ("Script.Rewrite'_Set", _) $ Free (rls_, _) $ _ $ f_) =>
   317 	       if contains_rule (Thm thm'') (assoc_rls rls_)
   318          then if f = f_ then Ass (m, f') else AssWeak (m, f')
   319 	       else NotAss
   320     | _ => NotAss)
   321   | assod _ _ (m as Rewrite_Set_Inst' (_, _, _, rls, f, (f', _))) 
   322       (Const ("Script.Rewrite'_Set'_Inst", _) $ _ $ Free (rls_, _) $ _ $ f_) = 
   323     if id_rls rls = rls_ 
   324     then if f = f_ then Ass (m, f') else AssWeak (m ,f')
   325     else NotAss
   326   | assod _ _ (m as Detail_Set_Inst' (_, _, _, rls, f, (f',_))) 
   327       (Const ("Script.Rewrite'_Set'_Inst", _) $ _ $ Free (rls_, _) $ _ $ f_) = 
   328     if id_rls rls = rls_
   329     then if f = f_ then Ass (m, f') else AssWeak (m, f')
   330     else NotAss
   331   | assod _ _ (m as Rewrite_Set' (_, _, rls, f, (f', _))) 
   332       (Const ("Script.Rewrite'_Set", _) $ Free (rls_, _) $ _ $ f_) = 
   333     if id_rls rls = rls_
   334     then if f = f_ then Ass (m, f') else AssWeak (m, f')
   335     else NotAss
   336   | assod _ _ (m as Detail_Set' (_, _, rls, f, (f', _))) 
   337       (Const ("Script.Rewrite'_Set", _) $ Free (rls_, _) $ _ $ f_) = 
   338     if id_rls rls = rls_
   339     then if f = f_ then Ass (m, f') else AssWeak (m, f')
   340     else NotAss
   341   | assod _ _ (m as Calculate' (_, op_, f, (f', _))) stac =
   342     (case stac of
   343 	    (Const ("Script.Calculate",_) $ Free (op__,_) $ f_) =>
   344 	      if op_ = op__
   345         then if f = f_ then Ass (m, f') else AssWeak (m, f')
   346 	      else NotAss
   347     | (Const ("Script.Rewrite'_Set'_Inst", _) $ _ $ Free(rls_,_) $_$f_)  =>
   348         let val thy = assoc_thy "Isac";
   349         in
   350           if contains_rule (Calc (assoc_calc' thy op_ |> snd)) (assoc_rls rls_)
   351           then if f = f_ then Ass (m, f') else AssWeak (m, f')
   352           else NotAss
   353         end
   354     | (Const ("Script.Rewrite'_Set",_) $ Free (rls_, _) $ _ $ f_) =>
   355         let val thy = assoc_thy "Isac";
   356         in
   357           if contains_rule (Calc (assoc_calc' thy op_ |> snd)) (assoc_rls rls_)
   358           then if f = f_ then Ass (m,f') else AssWeak (m,f')
   359           else NotAss
   360         end
   361     | _ => NotAss)
   362   | assod _ _ (m as Check_elementwise' (consts, _, (consts_chkd, _)))
   363       (Const ("Script.Check'_elementwise",_) $ consts' $ _) =
   364     if consts = consts'
   365     then Ass (m, consts_chkd)
   366     else NotAss
   367   | assod _ _ (m as Or_to_List' (_, list)) (Const ("Script.Or'_to'_List", _) $ _) = Ass (m, list) 
   368   | assod _ _ (m as Take' term) (Const ("Script.Take", _) $ _) = Ass (m, term)
   369   | assod _ _ (m as Substitute' (ro, erls, subte, f, f')) (Const ("Script.Substitute", _) $ _ $ t) =
   370 	  if f = t then Ass (m, f')
   371 	  else (*compare | applicable_in (p,p_) pt (m as Substitute sube)*)
   372 		  if foldl and_ (true, map contains_Var subte)
   373 		  then
   374 		    let val t' = subst_atomic (map HOLogic.dest_eq subte (*TODO subte2subst*)) t
   375 		    in if t = t' then error "assod: Substitute' not applicable to val of Expr"
   376 		       else Ass (Substitute' (ro, erls, subte, t, t'), t')
   377 		    end
   378 		  else (case rewrite_terms_ (Isac()) ro erls subte t of
   379 		         SOME (t', _) =>  Ass (Substitute' (ro, erls, subte, t, t'), t')
   380 		       | NONE => error "assod: Substitute' not applicable to val of Expr")
   381   | assod _ _ (m as Tac_ (thy, _, id, f')) (Const ("Script.Tac",_) $ Free (id', _)) =
   382     if id = id'
   383     then Ass (m, ((Thm.term_of o the o (parse thy)) f'))
   384     else NotAss
   385 
   386     (*compare "| stac2tac_ thy (Const ("Script.SubProblem",_)"*)
   387   | assod pt _ (Subproblem' ((domID, pblID, _), _, _, _, _, _))
   388 	    (stac as Const ("Script.SubProblem",_) $ (Const ("Product_Type.Pair",_) $
   389 		    Free (dI',_) $ (Const ("Product_Type.Pair",_) $ pI' $ mI')) $ ags') =
   390     let 
   391       val dI = ((implode o drop_last(*.."'"*) o Symbol.explode) dI')(*^""*);
   392       val thy = maxthy (assoc_thy dI) (rootthy pt);
   393 	    val pI = ((map (de_esc_underscore o free2str)) o isalist2list) pI';
   394 	    val mI = ((map (de_esc_underscore o free2str)) o isalist2list) mI';
   395 	    val ags = isalist2list ags';
   396 	    val (pI, pors, mI) = 
   397 	      if mI = ["no_met"] 
   398 	      then
   399           let
   400             val pors = (match_ags thy ((#ppc o get_pbt) pI) ags)
   401 		          handle ERROR "actual args do not match formal args"
   402 			          => (match_ags_msg pI stac ags(*raise exn*);[]);
   403 		        val pI' = refine_ori' pors pI;
   404 		      in (pI', pors (*refinement over models with diff.prec only*), (hd o #met o get_pbt) pI')
   405           end
   406 	      else (pI, (match_ags thy ((#ppc o get_pbt) pI) ags)
   407 		      handle ERROR "actual args do not match formal args"
   408 		      => (match_ags_msg pI stac ags(*raise exn*); []), mI);
   409       val (fmz_, vals) = oris2fmz_vals pors;
   410 	    val {cas, ppc, thy, ...} = get_pbt pI
   411 	    val dI = theory2theory' thy (*take dI from _refined_ pbl*)
   412 	    val dI = theory2theory' (maxthy (assoc_thy dI) (rootthy pt))
   413 	    val ctxt = dI |> Thy_Info.get_theory |> Proof_Context.init_global |> declare_constraints' vals
   414 	    val hdl = 
   415         case cas of
   416 		      NONE => pblterm dI pI
   417 		    | SOME t => subst_atomic ((vars_of_pbl_' ppc) ~~~ vals) t
   418       val f = subpbl (strip_thy dI) pI
   419     in 
   420       if domID = dI andalso pblID = pI
   421       then Ass (Subproblem' ((dI, pI, mI), pors, hdl, fmz_, ctxt, f), f) 
   422       else NotAss
   423     end
   424   | assod _ _ m _ = 
   425     (if (!trace_script) 
   426      then tracing("@@@ the 'tac_' proposed to apply does NOT match the leaf found in the script:\n"
   427 		   ^ "@@@ tac_ = " ^ tac_2str m)
   428      else ();
   429     NotAss);
   430 
   431 fun tac_2tac (Refine_Tacitly' (pI, _, _, _, _)) = Refine_Tacitly pI
   432   | tac_2tac (Model_Problem' (_, _, _)) = Model_Problem
   433   | tac_2tac (Add_Given' (t, _)) = Add_Given t
   434   | tac_2tac (Add_Find' (t, _)) = Add_Find t
   435   | tac_2tac (Add_Relation' (t, _)) = Add_Relation t
   436  
   437   | tac_2tac (Specify_Theory' dI) = Specify_Theory dI
   438   | tac_2tac (Specify_Problem' (dI, _)) = Specify_Problem dI
   439   | tac_2tac (Specify_Method' (dI, _, _)) = Specify_Method dI
   440   
   441   | tac_2tac (Rewrite' (_, _, _, _, thm, _, _)) = Rewrite thm
   442   | tac_2tac (Rewrite_Inst' (_, _, _, _, sub, thm, _, _)) = Rewrite_Inst (subst2subs sub, thm)
   443 
   444   | tac_2tac (Rewrite_Set' (_, _, rls, _, _)) = Rewrite_Set (id_rls rls)
   445   | tac_2tac (Detail_Set' (_, _, rls, _, _)) = Detail_Set (id_rls rls)
   446 
   447   | tac_2tac (Rewrite_Set_Inst' (_, _, sub, rls, _, _)) = 
   448     Rewrite_Set_Inst (subst2subs sub,id_rls rls)
   449   | tac_2tac (Detail_Set_Inst' (_, _, sub, rls, _, _)) = 
   450     Detail_Set_Inst (subst2subs sub,id_rls rls)
   451 
   452   | tac_2tac (Calculate' (_, op_, _, _)) = Calculate (op_)
   453   | tac_2tac (Check_elementwise' (_, pred, _)) = Check_elementwise pred
   454 
   455   | tac_2tac (Or_to_List' _) = Or_to_List
   456   | tac_2tac (Take' term) = Take (term2str term)
   457   | tac_2tac (Substitute' (_, _, subte, _, _)) = Substitute (subte2sube subte) 
   458   | tac_2tac (Tac_ (_, _, id, _)) = Tac id
   459 
   460   | tac_2tac (Subproblem' ((domID, pblID, _), _, _, _,_ ,_)) = Subproblem (domID, pblID)
   461   | tac_2tac (Check_Postcond' (pblID, _)) = Check_Postcond pblID
   462   | tac_2tac Empty_Tac_ = Empty_Tac
   463   | tac_2tac m = error ("tac_2tac: not impl. for "^(tac_2str m));
   464 
   465 val idT = Type ("Script.ID", []);
   466 
   467 fun make_rule thy t =
   468   let val ct = Thm.global_cterm_of thy (Trueprop $ t)
   469   in Thm (term_to_string''' thy (Thm.term_of ct), Thm.make_thm ct) end;
   470 
   471 fun rep_tac_ (Rewrite_Inst' (thy', _, _, put, subs, (thmID, _), f, (f', _))) = 
   472     let val fT = type_of f;
   473       val b = if put then @{term True} else @{term False};
   474       val sT = (type_of o fst o hd) subs;
   475       val subs' = list2isalist (HOLogic.mk_prodT (sT, sT)) (map HOLogic.mk_prod subs);
   476       val sT' = type_of subs';
   477       val lhs = Const ("Script.Rewrite'_Inst", [sT', idT, bool, fT] ---> fT) 
   478         $ subs' $ Free (thmID, idT) $ b $ f;
   479     in (((make_rule (assoc_thy thy')) o HOLogic.mk_eq) (lhs, f'), (lhs, f')) end
   480   | rep_tac_ (Rewrite' (thy', _, _, put, (thmID, _), f, (f', _)))=
   481     let 
   482       val fT = type_of f;
   483       val b = if put then @{term True} else @{term False};
   484       val lhs = Const ("Script.Rewrite", [idT, HOLogic.boolT, fT] ---> fT)
   485         $ Free (thmID, idT) $ b $ f;
   486     in (((make_rule (assoc_thy thy')) o HOLogic.mk_eq) (lhs, f'), (lhs, f')) end
   487   | rep_tac_ (Rewrite_Set_Inst' (_, _, _, _, _, (f', _))) = (e_rule, (e_term, f'))
   488   | rep_tac_ (Rewrite_Set' (thy', put, rls, f, (f', _))) =
   489     let 
   490       val fT = type_of f;
   491       val b = if put then @{term True} else @{term False};
   492       val lhs = Const ("Script.Rewrite'_Set", [idT, bool, fT] ---> fT) 
   493         $ Free (id_rls rls, idT) $ b $ f;
   494     in (((make_rule (assoc_thy thy')) o HOLogic.mk_eq) (lhs,f'),(lhs,f')) end
   495   | rep_tac_ (Calculate' (thy', op_, f, (f', _)))=
   496     let
   497       val fT = type_of f;
   498       val lhs = Const ("Script.Calculate",[idT,fT] ---> fT) $ Free (op_,idT) $ f
   499     in (((make_rule (assoc_thy thy')) o HOLogic.mk_eq) (lhs,f'),(lhs,f')) end
   500   | rep_tac_ (Check_elementwise' (_, _, (t', _))) = (Erule, (e_term, t'))
   501   | rep_tac_ (Subproblem' (_, _, _, _, _, t')) = (Erule, (e_term, t'))
   502   | rep_tac_ (Take' t') = (Erule, (e_term, t'))
   503   | rep_tac_ (Substitute' (_, _, _, t, t')) = (Erule, (t, t'))
   504   | rep_tac_ (Or_to_List' (t, t')) = (Erule, (t, t'))
   505   | rep_tac_ m = error ("rep_tac_: not impl.for " ^ tac_2str m)
   506 
   507 fun tac_2res m = (snd o snd o rep_tac_) m;
   508 
   509 (* handle a leaf at the end of recursive descent:
   510    a leaf is either a tactic or an 'expr' in "let v = expr"
   511    where "expr" does not contain a tactic.
   512    Handling a leaf comprises
   513    (1) 'subst_stacexpr' substitute env and complete curried tactic
   514    (2) rewrite the leaf by 'srls'
   515 *)
   516 fun handle_leaf call thy srls E a v t =
   517       (*WN050916 'upd_env_opt' is a blind copy from previous version*)
   518     case subst_stacexpr E a v t of
   519 	    (a', STac stac) => (*script-tactic*)
   520 	      let val stac' =
   521             eval_listexpr_ (assoc_thy thy) srls (subst_atomic (upd_env_opt E (a,v)) stac)
   522 	      in
   523           (if (! trace_script) 
   524 	         then tracing ("@@@ "^call^" leaf '"^term2str t^"' ---> STac '"^term2str stac ^"'")
   525 	         else ();
   526 	         (a', STac stac'))
   527 	      end
   528     | (a', Expr lexpr) => (*leaf-expression*)
   529 	      let val lexpr' =
   530             eval_listexpr_ (assoc_thy thy) srls (subst_atomic (upd_env_opt E (a,v)) lexpr)
   531 	      in
   532           (if (! trace_script) 
   533 	         then tracing("@@@ "^call^" leaf '"^term2str t^"' ---> Expr '"^term2str lexpr'^"'")
   534 	         else ();
   535 	         (a', Expr lexpr')) (*lexpr' is the value of the Expr*)
   536 	      end;
   537 
   538 (** locate an applicable stac in a script **)
   539 datatype assoc = (* ExprVal in the sense of denotational semantics               *)
   540   Assoc of       (* the stac is associated, strongly or weakly                   *)
   541   scrstate *     (* the current; returned for next_tac etc. outside ass*         *)  
   542   (step list)    (* list of steps done until associated stac found;
   543 	                  initiated with the data for doing the 1st step,
   544                     thus the head holds these data further on,
   545 		                while the tail holds steps finished (incl.scrstate in ptree) *)
   546 | NasApp of      (* stac not associated, but applicable, ptree-node generated    *)
   547   scrstate * (step list)
   548 | NasNap of      (* stac not associated, not applicable, nothing generated;
   549 	                  for distinction in Or, for leaving iterations, leaving Seq,
   550 		                evaluate scriptexpressions                                   *)
   551   term * env;
   552 fun assoc2str (Assoc _) = "Assoc"
   553   | assoc2str (NasNap _) = "NasNap"
   554   | assoc2str (NasApp _) = "NasApp";
   555 
   556 datatype asap = (* arg. of assy _only_ for distinction w.r.t. Or                 *)
   557   Aundef        (* undefined: set only by (topmost) Or                           *)
   558 | AssOnly       (* do not execute appl stacs - there could be an associated
   559 	                 in parallel Or-branch                                         *)
   560 | AssGen;       (* no Ass(Weak) found within Or, thus 
   561                    search for _applicable_ stacs, execute and generate pt        *)
   562 (*this constructions doesnt allow arbitrary nesting of Or !!!                    *)
   563 
   564 (* assy, ass_up, astep_up scan for locate_gen in a script.
   565    search is clearly separated into (1)-(2):
   566    (1) assy is recursive descent;
   567    (2) ass_up resumes interpretation at a location somewhere in the script;
   568        astep_up does only get to the parentnode of the scriptexpr.
   569    consequence:
   570    * call of (2) means _always_ that in this branch below
   571      there was an appl.stac (Repeat, Or e1, ...) found by the previous step.
   572 *)
   573 (*WN161112 blanks between list elements left as is until istate is introduced here*)
   574 fun assy ya ((E,l,a,v,S,b),ss) (Const ("HOL.Let",_) $ e $ (Abs (id,T,body))) =
   575     (case assy ya ((E , l @ [L, R], a,v,S,b),ss) e of
   576        NasApp ((E',l,a,v,S,_),ss) => 
   577          let
   578            val id' = mk_Free (id, T);
   579            val E' = upd_env E' (id', v);
   580          in assy ya ((E', l @ [R, D], a,v,S,b),ss) body end
   581      | NasNap (v,E) =>
   582          let
   583            val id' = mk_Free (id, T);
   584            val E' = upd_env E (id', v);
   585          in assy ya ((E', l @ [R, D], a,v,S,b),ss) body end
   586      | ay => ay)
   587   | assy (ya as (thy,_,srls,_,_)) ((E,l,_,v,S,b),ss) (Const ("Script.While",_) $ c $ e $ a) =
   588     if eval_true_ thy srls (subst_atomic (upd_env E (a,v)) c) 
   589     then assy ya ((E, l @ [L, R], SOME a,v,S,b),ss)  e
   590     else NasNap (v, E)
   591   | assy (ya as (thy,_,srls,_,_)) ((E,l,a,v,S,b),ss) (Const ("Script.While",_) $ c $ e) =
   592     if eval_true_ thy srls (subst_atomic (upd_env_opt E (a,v)) c) 
   593     then assy ya ((E, l @ [R], a,v,S,b),ss) e
   594     else NasNap (v, E)
   595   | assy (ya as (thy,_,srls,_,_)) ((E,l,a,v,S,b),ss) (Const ("If",_) $ c $ e1 $ e2) =
   596     if eval_true_ thy srls (subst_atomic (upd_env_opt E (a,v)) c) 
   597     then assy ya ((E, l @ [L, R], a,v,S,b),ss) e1
   598     else assy ya ((E, l @ [R], a,v,S,b),ss) e2 
   599   | assy ya ((E,l,_,v,S,b),ss) (Const ("Script.Try",_) $ e $ a) =
   600     (case assy ya ((E, l @ [L, R], SOME a,v,S,b),ss) e of ay => ay) 
   601   | assy ya ((E,l,a,v,S,b),ss) (Const ("Script.Try",_) $ e) =
   602     (case assy ya ((E, l @ [R], a,v,S,b),ss) e of ay => ay)
   603   | assy ya ((E,l,_,v,S,b),ss) (Const ("Script.Seq",_) $e1 $ e2 $ a) =
   604     (case assy ya ((E, l @ [L, L, R], SOME a,v,S,b),ss) e1 of
   605 	     NasNap (v, E) => assy ya ((E, l @ [L, R], SOME a,v,S,b),ss) e2
   606      | NasApp ((E,_,_,v,_,_),ss) => assy ya ((E, l @ [L, R], SOME a,v,S,b),ss) e2
   607      | ay => ay)
   608   | assy ya ((E,l,a,v,S,b),ss) (Const ("Script.Seq",_) $e1 $ e2) =
   609     (case assy ya ((E, l @ [L, R], a,v,S,b),ss) e1 of
   610 	     NasNap (v, E) => assy ya ((E, l @ [R], a,v,S,b),ss) e2
   611      | NasApp ((E,_,_,v,_,_),ss) => assy ya ((E, l @ [R], a,v,S,b),ss) e2
   612      | ay => ay)
   613   | assy ya ((E,l,_,v,S,b),ss) (Const ("Script.Repeat",_) $ e $ a) =
   614     assy ya ((E,(l @ [L, R]),SOME a,v,S,b),ss) e
   615   | assy ya ((E,l,a,v,S,b),ss) (Const ("Script.Repeat",_) $ e) =
   616     assy ya ((E,(l @ [R]),a,v,S,b),ss) e
   617   | assy (y,x,s,sc,Aundef) ((E,l,_,v,S,b),ss) (Const ("Script.Or",_) $e1 $ e2 $ a) =
   618     (case assy (y,x,s,sc,AssOnly) ((E,(l @ [L, L, R]),SOME a,v,S,b),ss) e1 of
   619 	     NasNap (v, E) => 
   620 	       (case assy (y,x,s,sc,AssOnly) ((E,(l @ [L, R]),SOME a,v,S,b),ss) e2 of
   621 	          NasNap (v, E) => 
   622 	            (case assy (y,x,s,sc,AssGen) ((E,(l @ [L, L, R]),SOME a,v,S,b),ss) e1 of
   623 	               NasNap (v, E) => 
   624 	                 assy (y,x,s,sc,AssGen) ((E, (l @ [L, R]), SOME a,v,S,b),ss) e2
   625 	             | ay => ay)
   626 	        | ay =>ay)
   627      | NasApp _ => error ("assy: FIXXXME ///must not return NasApp///")
   628      | ay => (ay))
   629   | assy ya ((E,l,a,v,S,b),ss) (Const ("Script.Or",_) $e1 $ e2) =
   630     (case assy ya ((E,(l @ [L, R]),a,v,S,b),ss) e1 of
   631 	     NasNap (v, E) => assy ya ((E,(l @ [R]),a,v,S,b),ss) e2
   632      | ay => (ay))
   633     (*here is not a tactical like TRY etc, but a tactic creating a step in calculation*)
   634   | assy (thy',ctxt,sr,d,ap) ((E,l,a,v,S,_), (m,_,pt,(p,p_),c)::ss) t =
   635     (case handle_leaf "locate" thy' sr E a v t of
   636 	     (a', Expr _) => 
   637 	        (NasNap (eval_listexpr_ (assoc_thy thy') sr
   638 		     (subst_atomic (upd_env_opt E (a',v)) t), E))
   639      | (a', STac stac) =>
   640 	       let
   641            val p' = 
   642              case p_ of Frm => p 
   643              | Res => lev_on p
   644 		         | _ => error ("assy: call by " ^ pos'2str (p,p_));
   645 	       in
   646            case assod pt d m stac of
   647 	         Ass (m,v') =>
   648 	           let val (p'',c',f',pt') =
   649                  generate1 (assoc_thy thy') m (ScrState (E,l,a',v',S,true), ctxt) (p',p_) pt;
   650 	           in Assoc ((E,l,a',v',S,true), (m,f',pt',p'',c @ c')::ss) end
   651            | AssWeak (m,v') => 
   652 	           let val (p'',c',f',pt') =
   653                generate1 (assoc_thy thy') m (ScrState (E,l,a',v',S,false), ctxt) (p',p_) pt;
   654 	           in Assoc ((E,l,a',v',S,false), (m,f',pt',p'',c @ c')::ss) end
   655            | NotAss =>
   656              (case ap of   (*switch for Or: 1st AssOnly, 2nd AssGen*)
   657                 AssOnly => (NasNap (v, E))
   658               | _ =>
   659                   (case applicable_in (p,p_) pt (stac2tac pt (assoc_thy thy') stac) of
   660 		                 Appl m' =>
   661 		                   let
   662                          val is = (E,l,a',tac_2res m',S,false(*FIXXXME.WN0?*))
   663 		                     val (p'',c',f',pt') =
   664 		                       generate1 (assoc_thy thy') m' (ScrState is, ctxt) (p',p_) pt;
   665 		                   in NasApp (is,(m,f',pt',p'',c @ c')::ss) end
   666 		               | Notappl _ => (NasNap (v, E))
   667 		              )
   668 		         )
   669          end)
   670   | assy _ (_, []) t = error ("assy: uncovered fun-def with " ^ term2str t);
   671 
   672 (*WN161112 blanks between list elements left as is until istate is introduced here*)
   673 fun ass_up (ys as (y,ctxt,s,Prog sc,d)) ((E,l,a,v,S,b),ss) (Const ("HOL.Let",_) $ _) =
   674     let 
   675 	    val l = drop_last l; (*comes from e, goes to Abs*)
   676       val (i, T, body) =
   677         (case go l sc of
   678            Const ("HOL.Let",_) $ _ $ (Abs (i, T, body)) => (i, T, body)
   679          | t => error ("ass_up..HOL.Let $ _ with " ^ term2str t))
   680       val i = mk_Free (i, T);
   681       val E = upd_env E (i, v);
   682     in case assy (y,ctxt,s,d,Aundef) ((E, l @ [R, D], a,v,S,b),ss) body of
   683 	       Assoc iss => Assoc iss
   684 	     | NasApp iss => astep_up ys iss 
   685 	     | NasNap (v, E) => astep_up ys ((E,l,a,v,S,b),ss) 
   686 	  end
   687   | ass_up ys iss (Abs (_,_,_)) = astep_up ys iss (*TODO 5.9.00: env ?*)
   688   | ass_up ys iss (Const ("HOL.Let",_) $ _ $ (Abs _)) = astep_up ys iss (*TODO 5.9.00: env ?*)
   689   | ass_up ysa iss (Const ("Script.Seq",_) $ _ $ _ $ _) =
   690     astep_up ysa iss (*all has been done in (*2*) below*)
   691   | ass_up ysa iss (Const ("Script.Seq",_) $ _ $ _) =
   692     astep_up ysa iss (*2*: comes from e2*)
   693 
   694   | ass_up (ysa as (y,ctxt,s,Prog sc,d)) ((E,l,a,v,S,b),ss)
   695 	  (Const ("Script.Seq",_) $ _ ) = (*2*: comes from e1, goes to e2*)
   696      let 
   697        val up = drop_last l;
   698      val e2 =
   699        (case go up sc of
   700           Const ("Script.Seq",_) $ _ $ e2 => e2
   701         | t => error ("ass_up..Script.Seq $ _ with " ^ term2str t))
   702      in case assy (y,ctxt,s,d,Aundef) ((E, up @ [R], a,v,S,b),ss) e2 of
   703          NasNap (v,E) => astep_up ysa ((E,up,a,v,S,b),ss)
   704        | NasApp iss => astep_up ysa iss
   705        | ay => ay 
   706      end
   707   | ass_up ysa iss (Const ("Script.Try",_) $ _ $ _) = astep_up ysa iss
   708   | ass_up ysa iss (Const ("Script.Try",_) $ _) = astep_up ysa iss
   709   | ass_up (ys as (y,ctxt,s,_,d)) ((E,l,_,v,S,b),ss)
   710 	   (*(Const ("Script.While",_) $ c $ e $ a) = WN050930 blind fix*)
   711 	    (t as Const ("Script.While",_) $ c $ e $ a) =
   712     if eval_true_ y s (subst_atomic (upd_env E (a,v)) c)
   713     then case assy (y,ctxt,s,d,Aundef) ((E, l @ [L, R], SOME a,v,S,b),ss) e of 
   714       NasNap (v,E') => astep_up ys ((E',l, SOME a,v,S,b),ss)
   715     | NasApp ((E',l,a,v,S,b),ss) =>
   716       ass_up ys ((E',l,a,v,S,b),ss) t (*WN050930 't' was not assigned*)
   717     | ay => ay
   718     else astep_up ys ((E,l, SOME a,v,S,b),ss)
   719   | ass_up (ys as (y,ctxt,s,_,d)) ((E,l,a,v,S,b),ss)
   720 	   (*(Const ("Script.While",_) $ c $ e) = WN050930 blind fix*)
   721 	     (t as Const ("Script.While",_) $ c $ e) =
   722     if eval_true_ y s (subst_atomic (upd_env_opt E (a,v)) c)
   723     then case assy (y,ctxt,s,d,Aundef) ((E, l @ [R], a,v,S,b),ss) e of 
   724        NasNap (v,E') => astep_up ys ((E',l, a,v,S,b),ss)
   725      | NasApp ((E',l,a,v,S,b),ss) =>
   726        ass_up ys ((E',l,a,v,S,b),ss) t (*WN050930 't' was not assigned*)
   727      | ay => ay
   728     else astep_up ys ((E,l, a,v,S,b),ss)
   729   | ass_up y iss (Const ("If",_) $ _ $ _ $ _) = astep_up y iss
   730   | ass_up (ys as (y,ctxt,s,_,d)) ((E,l,_,v,S,b),ss)
   731 	    (t as Const ("Script.Repeat",_) $ e $ a) =
   732     (case assy (y,ctxt,s,d, Aundef) ((E, (l @ [L, R]), SOME a,v,S,b),ss) e of 
   733       NasNap (v,E') => astep_up ys ((E',l, SOME a,v,S,b),ss)
   734     | NasApp ((E',l,a,v,S,b),ss) =>
   735       ass_up ys ((E',l,a,v,S,b),ss) t
   736     | ay => ay)
   737   | ass_up (ys as (y,ctxt,s,_,d)) ((E,l,a,v,S,b),ss)
   738 	    (t as Const ("Script.Repeat",_) $ e) =
   739     (case assy (y,ctxt,s,d,Aundef) ((E, (l @ [R]), a,v,S,b),ss) e of 
   740        NasNap (v', E') => astep_up ys ((E',l,a,v',S,b),ss)
   741      | NasApp ((E',l,a,v',S,_),ss) => ass_up ys ((E',l,a,v',S,b),ss) t
   742      | ay => ay)
   743   | ass_up y iss (Const ("Script.Or",_) $ _ $ _ $ _) = astep_up y iss
   744   | ass_up y iss (Const ("Script.Or",_) $ _ $ _) = astep_up y iss
   745   | ass_up y ((E,l,a,v,S,b),ss) (Const ("Script.Or",_) $ _ ) = 
   746     astep_up y ((E, (drop_last l), a,v,S,b),ss)
   747   | ass_up _ _ t =
   748     error ("ass_up not impl for t= " ^ term2str t)
   749 and astep_up (ys as (_,_,_,Prog sc,_)) ((E,l,a,v,S,b),ss) =
   750   if 1 < length l
   751   then 
   752     let val up = drop_last l;
   753     in ass_up ys ((E,up,a,v,S,b),ss) (go up sc) end
   754   else (NasNap (v, E))
   755   | astep_up _ ((_,l,_,_,_,_),_) = error ("astep_up: uncovered fun-def with " ^ loc_2str l)
   756 
   757 (*check if there are tacs for rewriting only*)
   758 fun rew_only ([]:step list) = true
   759   | rew_only (((Rewrite' _          ,_,_,_,_))::ss) = rew_only ss
   760   | rew_only (((Rewrite_Inst' _     ,_,_,_,_))::ss) = rew_only ss
   761   | rew_only (((Rewrite_Set' _      ,_,_,_,_))::ss) = rew_only ss
   762   | rew_only (((Rewrite_Set_Inst' _ ,_,_,_,_))::ss) = rew_only ss
   763   | rew_only (((Calculate' _        ,_,_,_,_))::ss) = rew_only ss
   764   | rew_only (((Begin_Trans' _      ,_,_,_,_))::ss) = rew_only ss
   765   | rew_only (((End_Trans' _        ,_,_,_,_))::ss) = rew_only ss
   766   | rew_only _ = false; 
   767 
   768 datatype locate =
   769   Steps of istate (* producing hd of step list (which was latest)
   770 	                   for next_tac, for reporting Safe|Unsafe to DG  *)
   771 	   * step       (* (scrstate producing this step is in ptree !)   *) 
   772 		 list         (* locate_gen may produce intermediate steps      *)
   773 | NotLocatable;   (* no (m Ass m') or (m AssWeak m') found          *)
   774 
   775 (* locate_gen tries to locate an input tac m in the script. 
   776    pursuing this goal the script is executed until an (m' equiv m) is found,
   777    or the end of the script
   778 args
   779    m   : input by the user, already checked by applicable_in,
   780          (to be searched within Or; and _not_ an m doing the step on ptree !)
   781    p,pt: (incl ets) at the time of input
   782    scr : the script
   783    d   : canonical simplifier for locating Take, Substitute, Subproblems etc.
   784    ets : ets at the time of input
   785    l   : the location (in scr) of the stac which generated the current formula
   786 returns
   787    Steps: pt,p (incl. ets) with m done
   788           pos' list of proofobjs cut (from generate)
   789           safe: implied from last proofobj
   790 	  ets:
   791    ///ToDo : ets contains a list of tacs to be done before m can be done
   792           NOT IMPL. -- "error: do other step before"
   793    NotLocatable: thus generate_hard
   794 *)
   795 (*WN161112 blanks between list elements left as is until istate is introduced here*)
   796 fun locate_gen (thy', _) (Rewrite' (_, ro, er, pa, thm, f, _)) (pt, p) 
   797 	    (Rfuns {locate_rule=lo,...}, _) (RrlsState (_,f'',rss,rts), _) = 
   798     (case lo rss f (Thm thm) of
   799 	    [] => NotLocatable
   800     | rts' => Steps (rts2steps [] ((pt,p),(f,f'',rss,rts),(thy',ro,er,pa)) rts'))
   801   | locate_gen (thy',srls) (m:tac_) ((pt,p):ptree * pos') 
   802 	    (scr as Prog (_ $ body),d) (ScrState (E,l,a,v,S,b), ctxt)  = 
   803     let val thy = assoc_thy thy';
   804     in case if l = [] orelse (
   805 		       (*init.in solve..Apply_Method...*)(last_elem o fst) p = 0 andalso snd p = Res)
   806 	       then (assy (thy',ctxt,srls,d,Aundef) ((E,[R],a,v,S,b), [(m,EmptyMout,pt,p,[])]) body)
   807 	       else (astep_up (thy',ctxt,srls,scr,d) ((E,l,a,v,S,b), [(m,EmptyMout,pt,p,[])]) ) of
   808 	    Assoc ((is as (_,_,_,_,_,strong_ass), ss as (_ :: _))) =>
   809 	      (if strong_ass
   810          then (Steps (ScrState is, ss))
   811 	       else
   812            if rew_only ss (*andalso 'not strong_ass'= associated weakly*)
   813 	         then
   814              let
   815                val (po,p_) = p
   816                val po' = case p_ of Frm => po | Res => lev_on po | _ => error ("locate_gen " ^ pos_2str p_)
   817                val (p'',c'',f'',pt'') = generate1 thy m (ScrState is, ctxt) (po',p_) pt
   818 	           in Steps (ScrState is, [(m, f'',pt'',p'',c'')]) end
   819 	         else Steps (ScrState is, ss))
   820 	  
   821     | NasApp _ => NotLocatable
   822     | err => error ("not-found-in-script: NotLocatable from " ^ PolyML.makestring err)
   823     end
   824   | locate_gen _ m _ (sc,_) (is, _) = 
   825     error ("locate_gen: wrong arguments,\n tac= " ^ tac_2str m ^ ",\n " ^
   826       "scr= " ^ scr2str sc ^ ",\n istate= " ^ istate2str is);
   827 
   828 (** find the next stactic in a script **)
   829 
   830 (*appy, nxt_up, nstep_up scanning for next_tac.
   831   search is clearly separated into (1)-(2):
   832   (1) appy is recursive descent;
   833   (2) nxt_up resumes interpretation at a location somewhere in the script;
   834       nstep_up does only get to the parentnode of the scriptexpr.
   835   consequence:
   836   * call of (2) means _always_ that in this branch below
   837     there was an applicable stac (Repeat, Or e1, ...)
   838 *)
   839 datatype appy =  (* ExprVal in the sense of denotational semantics  *)
   840     Appy of      (* applicable stac found, search stalled           *)
   841     tac_ *       (* tac_ associated (fun assod) with stac           *)
   842     scrstate     (* after determination of stac WN.18.8.03          *)
   843   | Napp of      (* stac found was not applicable; 
   844 	                  this mode may become Skip in Repeat, Try and Or *)
   845     env (*stack*)(* popped while nxt_up                             *)
   846   | Skip of      (* for restart after Appy, for leaving iterations,
   847 	                  for passing the value of scriptexpressions,
   848 		                and for finishing the script successfully       *)
   849     term * env (*stack*);
   850 
   851 datatype appy_ = (* as argument in nxt_up, nstep_up, from appy               *)
   852 (*Appy              is only (final) returnvalue, not argument during search  *)
   853   Napp_          (* ev. detects 'script is not appropriate for this example' *)
   854 | Skip_;         (* detects 'script successfully finished'
   855 		                also used as init-value for resuming; this works,
   856 	                  because 'nxt_up Or e1' treats as Appy                    *)
   857 
   858 fun appy thy ptp E l (Const ("HOL.Let",_) $ e $ (Abs (i,T,b))) a v =
   859     (case appy thy ptp E (l @ [L, R]) e a v of
   860       Skip (res, E) => 
   861         let val E' = upd_env E (Free (i,T), res);
   862         in appy thy ptp E' (l @ [R, D]) b a v end
   863     | ay => ay)
   864   | appy (thy as (th,sr)) ptp E l (Const ("Script.While"(*1*),_) $ c $ e $ a) _ v =
   865     (if eval_true_ th sr (subst_atomic (upd_env E (a,v)) c)
   866      then appy thy ptp E (l @ [L, R]) e (SOME a) v
   867      else Skip (v, E))
   868   | appy (thy as (th,sr)) ptp E l (Const ("Script.While"(*2*),_) $ c $ e) a v =
   869     (if eval_true_ th sr (subst_atomic (upd_env_opt E (a,v)) c)
   870      then appy thy ptp E (l @ [R]) e a v
   871      else Skip (v, E))
   872   | appy (thy as (th,sr)) ptp E l (Const ("If",_) $ c $ e1 $ e2) a v =
   873     (if eval_true_ th sr (subst_atomic (upd_env_opt E (a,v)) c)
   874      then appy thy ptp E (l @ [L, R]) e1 a v
   875      else appy thy ptp E (l @ [R]) e2 a v)
   876   | appy thy ptp E l (Const ("Script.Repeat"(*1*),_) $ e $ a) _ v = 
   877     appy thy ptp E (l @ [L, R]) e (SOME a) v
   878   | appy thy ptp E l (Const ("Script.Repeat"(*2*),_) $ e) a v = appy thy ptp E (l @ [R]) e a v
   879   | appy thy ptp E l (Const ("Script.Try",_) $ e $ a) _ v =
   880     (case appy thy ptp E (l @ [L, R]) e (SOME a) v of
   881       Napp E => (Skip (v, E))
   882     | ay => ay)
   883   | appy thy ptp E l(Const ("Script.Try",_) $ e) a v =
   884     (case appy thy ptp E (l @ [R]) e a v of
   885       Napp E => (Skip (v, E))
   886     | ay => ay)
   887   | appy thy ptp E l (Const ("Script.Or"(*1*),_) $e1 $ e2 $ a) _ v =
   888     (case appy thy ptp E (l @ [L, L, R]) e1 (SOME a) v of
   889 	    Appy lme => Appy lme
   890     | _ => appy thy ptp E (*env*) (l @ [L, R]) e2 (SOME a) v)
   891   | appy thy ptp E l (Const ("Script.Or"(*2*),_) $e1 $ e2) a v =
   892     (case appy thy ptp E (l @ [L, R]) e1 a v of
   893 	    Appy lme => Appy lme
   894     | _ => appy thy ptp E (l @ [R]) e2 a v)
   895   | appy thy ptp E l (Const ("Script.Seq"(*1*),_) $ e1 $ e2 $ a) _ v =
   896     (case appy thy ptp E (l @ [L, L, R]) e1 (SOME a) v of
   897 	    Skip (v,E) => appy thy ptp E (l @ [L, R]) e2 (SOME a) v
   898     | ay => ay)
   899   | appy thy ptp E l (Const ("Script.Seq",_) $ e1 $ e2) a v =
   900     (case appy thy ptp E (l @ [L,R]) e1 a v of
   901 	    Skip (v,E) => appy thy ptp E (l @ [R]) e2 a v
   902     | ay => ay)
   903   (* a leaf has been found *)   
   904   | appy ((th,sr)) (pt, p) E l t a v =
   905     case handle_leaf "next  " th sr E a v t of
   906 	    (_, Expr s) => Skip (s, E)
   907     | (a', STac stac) =>
   908 	    let val (m,m') = stac2tac_ pt (assoc_thy th) stac
   909       in case m of 
   910 	      Subproblem _ => Appy (m', (E,l,a',tac_2res m',Sundef,false))
   911 	    | _ =>
   912         (case applicable_in p pt m of
   913 		       Appl m' => (Appy (m', (E,l,a',tac_2res m',Sundef,false)))
   914 		     | _ => Napp E)
   915 	    end
   916 (*GOON*)
   917 fun nxt_up thy ptp (scr as (Prog sc)) E l ay (Const ("HOL.Let", _) $ _) a v = (*comes from let=...*)
   918     if ay = Napp_
   919     then nstep_up thy ptp scr E (drop_last l) Napp_ a v
   920     else (*Skip_*)
   921 	    let
   922         val up = drop_last l
   923         val (i, T, body) =
   924           (case go up sc of
   925              Const ("HOL.Let",_) $ _ $ (Abs aa) => aa
   926            | t => error ("nxt_up..HOL.Let $ _ with " ^ term2str t))
   927         val i = mk_Free (i, T)
   928         val E = upd_env E (i, v)
   929       in
   930         case appy thy ptp E (up @ [R,D]) body a v  of
   931 	        Appy lre => Appy lre
   932 	      | Napp E => nstep_up thy ptp scr E up Napp_ a v
   933 	      | Skip (v,E) => nstep_up thy ptp scr E up Skip_ a v
   934 	    end
   935   | nxt_up thy ptp scr E l ay (Abs _) a v =  nstep_up thy ptp scr E l ay a v
   936   | nxt_up thy ptp scr E l ay (Const ("HOL.Let",_) $ _ $ (Abs _)) a v =
   937     nstep_up thy ptp scr E l ay a v
   938   (*no appy_: never causes Napp -> Helpless*)
   939   | nxt_up (thy as (th, sr)) ptp scr E l _ (Const ("Script.While"(*1*), _) $ c $ e $ _) a v = 
   940     if eval_true_ th sr (subst_atomic (upd_env_opt E (a, v)) c) 
   941     then case appy thy ptp E (l @ [L,R]) e a v of
   942 	     Appy lr => Appy lr
   943 	  | Napp E => nstep_up thy ptp scr E l Skip_ a v
   944 	  | Skip (v,E) => nstep_up thy ptp scr E l Skip_ a v
   945     else nstep_up thy ptp scr E l Skip_ a v
   946   (*no appy_: never causes Napp - Helpless*)
   947   | nxt_up (thy as (th, sr)) ptp scr E l _ (Const ("Script.While"(*2*), _) $ c $ e) a v = 
   948     if eval_true_ th sr (subst_atomic (upd_env_opt E (a, v)) c) 
   949     then case appy thy ptp E (l @ [R]) e a v of
   950 	    Appy lr => Appy lr
   951 	  | Napp E => nstep_up thy ptp scr E l Skip_ a v
   952 	  | Skip (v,E) => nstep_up thy ptp scr E l Skip_ a v
   953     else nstep_up thy ptp scr E l Skip_ a v
   954   | nxt_up thy ptp scr E l ay (Const ("If", _) $ _ $ _ $ _) a v = nstep_up thy ptp scr E l ay a v
   955   | nxt_up thy ptp scr E l _ (*no appy_: there was already a stac below*)
   956       (Const ("Script.Repeat"(*1*), _) $ e $ _) a v =
   957     (case appy thy ptp (*upd_env*) E (*a,v)*) ((l @ [L, R]):loc_) e a v  of
   958       Appy lr => Appy lr
   959     | Napp E =>  nstep_up thy ptp scr E l Skip_ a v
   960     | Skip (v,E) =>  nstep_up thy ptp scr E l Skip_ a v)
   961   | nxt_up thy ptp scr E l _ (*no appy_: there was already a stac below*)
   962       (Const ("Script.Repeat"(*2*), _) $ e) a v =
   963     (case appy thy ptp (*upd_env*) E (*a,v)*) (l @ [R]) e a v  of
   964       Appy lr => Appy lr
   965     | Napp E => nstep_up thy ptp scr E l Skip_ a v
   966     | Skip (v,E) => nstep_up thy ptp scr E l Skip_ a v)
   967   | nxt_up thy ptp scr E l _ (Const ("Script.Try",_) $ _ $ _) a v = (*makes Napp to Skip*)
   968     nstep_up thy ptp scr E l Skip_ a v
   969 
   970   | nxt_up thy ptp scr E l _ (Const ("Script.Try"(*2*), _) $ _) a v = (*makes Napp to Skip*)
   971     nstep_up thy ptp scr E l Skip_ a v
   972   | nxt_up thy ptp scr E l ay (Const ("Script.Or",_) $ _ $ _ $ _) a v =
   973     nstep_up thy ptp scr E l ay a v
   974   | nxt_up thy ptp scr E l ay (Const ("Script.Or",_) $ _ $ _) a v = nstep_up thy ptp scr E l ay a v
   975   | nxt_up thy ptp scr E l ay (Const ("Script.Or",_) $ _ ) a v = 
   976     nstep_up thy ptp scr E (drop_last l) ay a v
   977   | nxt_up thy ptp scr E l ay (Const ("Script.Seq"(*1*),_) $ _ $ _ $ _) a v =
   978     (*all has been done in (*2*) below*) nstep_up thy ptp scr E l ay a v
   979   | nxt_up thy ptp scr E l ay (Const ("Script.Seq"(*2*),_) $ _ $ _) a v = (*comes from e2*)
   980     nstep_up thy ptp scr E l ay a v
   981   | nxt_up thy ptp (scr as Prog sc) E l ay (Const ("Script.Seq",_) $ _) a v = (*comes from e1*)
   982     if ay = Napp_
   983     then nstep_up thy ptp scr E (drop_last l) Napp_ a v
   984     else (*Skip_*)
   985       let val up = drop_last l;
   986           val e2 =
   987             (case go up sc of
   988                Const ("Script.Seq"(*2*), _) $ _ $ e2 => e2
   989              | t => error ("nxt_up..Script.Seq $ _ with " ^ term2str t))
   990       in case appy thy ptp E (up @ [R]) e2 a v  of
   991         Appy lr => Appy lr
   992       | Napp E => nstep_up thy ptp scr E up Napp_ a v
   993       | Skip (v,E) => nstep_up thy ptp scr E up Skip_ a v end
   994   | nxt_up _ _ _ _ _ _ t _ _ = error ("nxt_up not impl for " ^ term2str t)
   995 and nstep_up thy ptp (Prog sc) E l ay a v = 
   996     if 1 < length l
   997     then 
   998       let val up = drop_last l; 
   999       in (nxt_up thy ptp (Prog sc) E up ay (go up sc) a v ) end
  1000     else (*interpreted to end*)
  1001       if ay = Skip_ then Skip (v, E) else Napp E 
  1002   | nstep_up _ _ _ _ l _ _ _ = error ("nstep_up: uncovered fun-def at " ^ loc_2str l)
  1003 
  1004 (* decide for the next applicable stac in the script;
  1005    returns (stactic, value) - the value in case the script is finished 
  1006    12.8.02:         ~~~~~ and no assumptions ??? FIXME ???
  1007    20.8.02: must return p in case of finished, because the next script
  1008             consulted need not be the calling script:
  1009             in case of detail ie. _inserted_ PrfObjs, the next stac
  1010             has to searched in a script with PblObj.status<>Complete !
  1011             (.. not true for other details ..PrfObj ??????????????????
  1012    20.8.02: do NOT return safe (is only changed in locate !!!)
  1013 *)
  1014 fun next_tac (thy,_) _ (Rfuns {next_rule, ...}) (RrlsState(f, f', rss, _), ctxt) =
  1015     if f = f'
  1016     then (End_Detail' (f',[])(*8.6.03*), (Uistate, ctxt), 
  1017     	(f', Sundef(*FIXME is no value of next_tac! vor 8.6.03*)))  (*finished*)
  1018     else
  1019       (case next_rule rss f of
  1020     	  NONE => (Empty_Tac_, (Uistate, ctxt), (e_term, Sundef))   (*helpless*)
  1021     	| SOME (Thm thm'')(*8.6.03: muss auch f' liefern ?!!*) => 
  1022     	    (Rewrite' (thy, "e_rew_ord", e_rls, false, thm'', f, (e_term, [(*!?!8.6.03*)])),
  1023   	         (Uistate, ctxt), (e_term, Sundef)))                  (*next stac*)
  1024   | next_tac thy (ptp as (pt, (p, _)):ptree * pos') (sc as Prog (_ $ body)) 
  1025 	    (ScrState (E,l,a,v,s,_), ctxt) =
  1026     (case if l = [] then appy thy ptp E [R] body NONE v
  1027           else nstep_up thy ptp sc E l Skip_ a v of
  1028        Skip (v, _) =>                                              (*finished*)
  1029          (case par_pbl_det pt p of
  1030 	          (true, p', _) => 
  1031 	             let
  1032 	               val (_,pblID,_) = get_obj g_spec pt p';
  1033 	              in (Check_Postcond' (pblID, (v, [(*assigned in next step*)])), 
  1034 	                   (e_istate, ctxt), (v,s)) 
  1035                 end
  1036 	        | _ => (End_Detail' (e_term,[])(*8.6.03*), (e_istate, ctxt), (v,s)))
  1037      | Napp _ => (Empty_Tac_, (e_istate, ctxt), (e_term, Sundef))   (*helpless*)
  1038      | Appy (m', scrst as (_,_,_,v,_,_)) => (m', (ScrState scrst, ctxt), (v, Sundef))) (*next stac*)
  1039   | next_tac _ _ _ (is, _) = error ("next_tac: not impl for " ^ (istate2str is));
  1040 
  1041 (*.create the initial interpreter state from the items of the guard.*)
  1042 local
  1043 val errmsg = "ERROR: found no actual arguments for prog. of "
  1044 fun msg_miss (sc, metID, formals, actuals) =
  1045   "ERROR in creating the environment for '" ^ id_of_scr sc ^ 
  1046 	"' from \nthe items of the guard of " ^ metID2str metID ^ ",\n" ^
  1047 	"formal arg(s), from the script, miss actual arg(s), from the guards env:\n" ^
  1048 	(string_of_int o length) formals ^ " formals: " ^ terms2str formals ^ "\n" ^
  1049 	(string_of_int o length) actuals ^ " actuals: " ^ terms2str actuals
  1050 fun msg_type (sc, metID, a, f, formals, actuals) =
  1051   "ERROR in creating the environment for '" ^
  1052 	id_of_scr sc ^ "' from \nthe items of the guard of " ^
  1053 	metID2str metID ^ ",\n" ^
  1054 	"different types of formal arg, from the script, " ^
  1055 	"and actual arg, from the guards env:'\n" ^
  1056 	"formal: '" ^ term2str a ^ "::" ^ (type2str o type_of) a ^ "'\n" ^
  1057 	"actual: '" ^ term2str f ^ "::" ^ (type2str o type_of) f ^ "'\n" ^
  1058 	"in\n" ^
  1059 	"formals: " ^ terms2str formals ^ "\n" ^
  1060 	"actuals: " ^ terms2str actuals
  1061 in
  1062 fun init_scrstate thy itms metID =
  1063   let
  1064     val actuals = itms2args thy metID itms
  1065     val _ = if actuals <> [] then () else raise ERROR (errmsg ^ strs2str' metID)
  1066     val (scr, sc) = (case (#scr o get_met) metID of
  1067        scr as Prog sc => (scr, sc) | _ => raise ERROR ("init_scrstate with " ^ metID2str metID))
  1068     val formals = formal_args sc    
  1069 	  (*expects same sequence of (actual) args in itms and (formal) args in met*)
  1070 	  fun relate_args env [] [] = env
  1071 	    | relate_args _ _ [] = error (msg_miss (sc, metID, formals, actuals))
  1072 	    | relate_args env [] _ = env (*may drop Find!*)
  1073 	    | relate_args env (a::aa) (f::ff) = 
  1074 	      if type_of a = type_of f 
  1075 	      then relate_args (env @ [(a, f)]) aa ff
  1076         else error (msg_type (sc, metID, a, f, formals, actuals))
  1077     val env = relate_args [] formals actuals;
  1078     val ctxt = Proof_Context.init_global thy |> declare_constraints' actuals
  1079     val {pre, prls, ...} = get_met metID;
  1080     val pres = check_preconds thy prls pre itms |> map snd;
  1081     val ctxt = ctxt |> insert_assumptions pres;
  1082   in (ScrState (env, [], NONE, e_term, Safe, true), ctxt, scr) : istate * Proof.context * scr end;
  1083 end (*local*)
  1084 
  1085 (* decide, where to get script/istate from:
  1086    (* 1 *) from PblObj.env: at begin of script if no init_form
  1087    (* 2 *) from PblObj/PrfObj: if stac is in the middle of the script
  1088    (* 3 *) from rls/PrfObj: in case of detail a ruleset *)
  1089 fun from_pblobj_or_detail' _ (p, p_) pt =
  1090   if member op = [Pbl, Met] p_
  1091   then case get_obj g_env pt p of
  1092     NONE => error "from_pblobj_or_detail': no istate"
  1093   | SOME is =>
  1094       let
  1095         val metID = get_obj g_metID pt p
  1096         val {srls, ...} = get_met metID
  1097       in (srls, is, (#scr o get_met) metID) end
  1098   else
  1099     let val (pbl, p', rls') = par_pbl_det pt p
  1100     in if pbl 
  1101        then (*if last_elem p = 0 nothing written to pt yet*)                                (* 2 *)
  1102          let
  1103 	         val metID = get_obj g_metID pt p'
  1104 	         val {srls,...} = get_met metID
  1105 	       in (srls, get_loc pt (p,p_), (#scr o get_met) metID) end
  1106        else (*FIXME.WN0?: get from pbl or met !!! unused for Rrls in locate_gen, next_tac*) (* 3 *)
  1107 	       (e_rls, get_loc pt (p,p_),
  1108 	          case rls' of
  1109 		          Rls {scr = scr,...} => scr
  1110 	          | Seq {scr = scr,...} => scr
  1111 	          | Rrls {scr=rfuns,...} => rfuns
  1112 	          | Erls => error "from_pblobj_or_detail' with Erls")
  1113     end
  1114 
  1115 (*.get script and istate from PblObj, see                                                  ( * 1 *)
  1116 fun from_pblobj' thy' (p,p_) pt =
  1117   let
  1118     val p' = par_pblobj pt p
  1119 	  val thy = assoc_thy thy'
  1120 	  val itms =
  1121 	    (case get_obj I pt p' of
  1122 	      PblObj {meth = itms, ...} => itms
  1123 	    | PrfObj _ => error "from_pblobj' NOT with PrfObj")
  1124 	  val metID = get_obj g_metID pt p'
  1125 	  val {srls, scr, ...} = get_met metID
  1126   in
  1127     if last_elem p = 0 (*nothing written to pt yet*)
  1128     then
  1129        let val (is, ctxt, scr) = init_scrstate thy itms metID
  1130 	     in (srls, (is, ctxt), scr) end
  1131     else (srls, get_loc pt (p,p_), scr)
  1132   end;
  1133     
  1134 (*.get the stactics and problems of a script as tacs
  1135   instantiated with the current environment;
  1136   l is the location which generated the given formula.*)
  1137 (*WN.12.5.03: quick-and-dirty repair for listexpressions*)
  1138 fun is_spec_pos Pbl = true
  1139   | is_spec_pos Met = true
  1140   | is_spec_pos _ = false;
  1141 
  1142 (*. fetch _all_ tactics from script .*)
  1143 fun sel_rules _ (([],Res):pos') = 
  1144     raise PTREE "no tactics applicable at the end of a calculation"
  1145   | sel_rules pt (p,p_) =
  1146     if is_spec_pos p_ 
  1147     then [get_obj g_tac pt p]
  1148     else
  1149       let
  1150         val pp = par_pblobj pt p;
  1151         val thy' = (get_obj g_domID pt pp):theory';
  1152         val thy = assoc_thy thy';
  1153         val metID = get_obj g_metID pt pp;
  1154         val metID' = if metID =e_metID then (thd3 o snd3) (get_obj g_origin pt pp) else metID
  1155         val (sc, srls) = (case get_met metID' of
  1156             {scr = Prog sc, srls, ...} => (sc, srls) | _ => error "sel_rules 1")
  1157         val (env, a, v) = (case get_istate pt (p, p_) of
  1158             ScrState (env, _, a, v, _, _) => (env, a, v) | _ => error "sel_rules 2")
  1159       in map ((stac2tac pt thy) o rep_stacexpr o #2 o
  1160   	    (handle_leaf "selrul" thy' srls env a v)) (stacpbls sc)
  1161   	  end;
  1162 
  1163 (*. fetch tactics from script and filter _applicable_ tactics;
  1164     in case of Rewrite_Set* go down to _atomic_ rewrite-tactics .*)
  1165 fun sel_appl_atomic_tacs _ (([], Res) : pos') = 
  1166     raise PTREE "no tactics applicable at the end of a calculation"
  1167   | sel_appl_atomic_tacs pt (p, p_) =
  1168     if is_spec_pos p_ 
  1169     then [get_obj g_tac pt p]
  1170     else
  1171       let
  1172         val pp = par_pblobj pt p
  1173         val thy' = (get_obj g_domID pt pp):theory'
  1174         val thy = assoc_thy thy'
  1175         val metID = get_obj g_metID pt pp
  1176         val metID' =
  1177           if metID = e_metID 
  1178           then (thd3 o snd3) (get_obj g_origin pt pp)
  1179           else metID
  1180         val (sc, srls, erls, ro) = (case get_met metID' of
  1181             {scr = Prog sc, srls, erls, rew_ord' = ro, ...} => (sc, srls, erls, ro)
  1182           | _ => error "sel_appl_atomic_tacs 1")
  1183         val (env, a, v) = (case get_istate pt (p, p_) of
  1184             ScrState (env, _, a, v, _, _) => (env, a, v) | _ => error "sel_appl_atomic_tacs 2")
  1185         val alltacs = (*we expect at least 1 stac in a script*)
  1186           map ((stac2tac pt thy) o rep_stacexpr o #2 o
  1187            (handle_leaf "selrul" thy' srls env a v)) (stacpbls sc)
  1188         val f =
  1189           (case p_ of Frm => get_obj g_form pt p | Res => (fst o (get_obj g_result pt)) p
  1190           | _ => error "")
  1191           (*WN071231 ? replace atomic_appl_tacs with applicable_in (ineff!) ?*)
  1192       in ((gen_distinct eq_tac) o flat o (map (atomic_appl_tacs thy ro erls f))) alltacs end;
  1193 (**)
  1194 end
  1195 (**)
  1196 (*open Lucin;*)